diff --git a/.travis.yml b/.travis.yml index 42b133a..b6947db 100644 --- a/.travis.yml +++ b/.travis.yml @@ -1,4 +1,4 @@ -dist: bionic +dist: focal language: minimal services: @@ -9,8 +9,8 @@ env: - NJOBS="2" - CONTRIB_NAME="demo" matrix: - - COQ_IMAGE="mathcomp/mathcomp:1.10.0-coq-8.10" - - COQ_IMAGE="mathcomp/mathcomp:1.10.0-coq-8.11" + - COQ_IMAGE="mathcomp/mathcomp:1.11.0-coq-8.11" + - COQ_IMAGE="mathcomp/mathcomp:1.11.0-coq-8.12" - COQ_IMAGE="coqorg/coq:dev" # "default" switch: $COMPILER # this config uses: $COMPILER_EDGE (cf. "opam switch" command below) diff --git a/_CoqProject b/_CoqProject index 1365b3a..686f5e2 100644 --- a/_CoqProject +++ b/_CoqProject @@ -2,16 +2,20 @@ -arg -w -arg -notation-overridden -arg -w -arg -redundant-canonical-projection -finmap/finmap.v -finmap/ordtype.v +options.v +core/axioms.v +core/prelude.v +core/pred.v +core/ordtype.v +core/seqperm.v +core/finmap.v +core/rtc.v +pcm/pcm.v +pcm/morphism.v +pcm/invertible.v +pcm/unionmap.v pcm/automap.v -pcm/axioms.v pcm/heap.v pcm/mutex.v pcm/lift.v pcm/natmap.v -pcm/pcm.v -pcm/pred.v -pcm/prelude.v -pcm/unionmap.v - diff --git a/pcm/axioms.v b/core/axioms.v similarity index 77% rename from pcm/axioms.v rename to core/axioms.v index 6bc48fb..d0ab7bd 100644 --- a/pcm/axioms.v +++ b/core/axioms.v @@ -22,9 +22,7 @@ limitations under the License. (******************************************************************************) From Coq Require Import ssreflect ssrfun Eqdep ClassicalFacts. -Set Implicit Arguments. -Unset Strict Implicit. -Unset Printing Implicit Defensive. +From fcsl Require Import options. (*****************************) (* Axioms and extensionality *) @@ -38,9 +36,6 @@ Axiom fext : forall A (B : A -> Type) (f1 f2 : forall x, B x), Lemma pf_irr (P : Prop) (p1 p2 : P) : p1 = p2. Proof. by apply/ext_prop_dep_proof_irrel_cic/pext. Qed. -Lemma eta A (B : A -> Type) (f : forall x, B x) : f = [eta f]. -Proof. by apply: fext. Qed. - Lemma sval_inj A P : injective (@sval A P). Proof. move=>[x Hx][y Hy] /= H; move: Hx Hy; rewrite H=>*. @@ -60,8 +55,10 @@ Proof. by apply: fext. Qed. (* Cast and John Major Equality via cast *) (*****************************************) +(* depends on StreicherK axiom *) + Section Cast. -Variable (C : Type) (interp : C -> Type). +Variable (T : Type) (interp : T -> Type). Definition cast A B (pf : A = B) (v : interp B) : interp A := ecast _ _ (esym pf) v. @@ -69,29 +66,30 @@ Definition cast A B (pf : A = B) (v : interp B) : interp A := Lemma eqc A (pf : A = A) (v : interp A) : cast pf v = v. Proof. by move: pf; apply: Streicher_K. Qed. -Definition jmeq A B (v : interp A) (w : interp B) := forall pf, v = cast pf w. +Definition jmeq A B (v : interp A) (w : interp B) := exists pf, v = cast pf w. -Lemma jmrefl A (v : interp A) : jmeq v v. -Proof. by move=>pf; rewrite eqc. Qed. +Lemma jm_refl A (v : interp A) : jmeq v v. +Proof. by exists (erefl _); rewrite eqc. Qed. -Lemma jmsym A B (v : interp A) (w : interp B) : jmeq v w -> jmeq w v. -Proof. -move=> H pf; rewrite (H (esym pf)). -by move: (pf); rewrite pf in w H * => {pf}-pf; rewrite !eqc. -Qed. +Lemma jm_sym A B (v : interp A) (w : interp B) : jmeq v w -> jmeq w v. +Proof. by case=>? ->; subst B; rewrite eqc; apply: jm_refl. Qed. + +Lemma jm_trans A B C (u : interp A) (v : interp B) (w : interp C) : + jmeq u v -> jmeq v w -> jmeq u w. +Proof. by case=>? -> [? ->]; subst B C; rewrite !eqc; apply: jm_refl. Qed. Lemma jmE A (v w : interp A) : jmeq v w <-> v = w. -Proof. by split=>[/(_ erefl) //|->]; apply: jmrefl. Qed. +Proof. by split=>[[?]|] ->; [rewrite eqc | apply: jm_refl]. Qed. Lemma castE A B (pf1 pf2 : A = B) (v1 v2 : interp B) : v1 = v2 <-> cast pf1 v1 = cast pf2 v2. -Proof. by move: (pf1) pf2; rewrite pf1 =>*; rewrite !eqc. Qed. +Proof. by subst B; rewrite !eqc. Qed. End Cast. -Arguments cast {C} interp [A][B] pf v. -Arguments jmeq {C} interp [A][B] v w. -Hint Resolve jmrefl : core. +Arguments cast {T} interp [A][B] pf v. +Arguments jmeq {T} interp [A][B] v w. +Hint Resolve jm_refl : core. (* special notation for the common case when interp = id *) Notation icast pf v := (@cast _ id _ _ pf v). Notation ijmeq v w := (@jmeq _ id _ _ v w). @@ -115,10 +113,9 @@ Definition dyn_inj := @inj_pair2 _ P. End Dynamic. Prenex Implicits dyn_tp dyn_val dyn_injT dyn_inj. -Arguments dyn {C} interp {A} _ : rename. +Arguments dyn {T} interp {A} _ : rename. Notation idyn v := (@dyn _ id _ v). -Lemma dynE (A B : Type) interp (pf : A = B) (v : interp A) (w : interp B) : +Lemma dynE (A B : Type) interp (v : interp A) (w : interp B) : jmeq interp v w <-> dyn interp v = dyn interp w. -Proof. by rewrite -pf in w *; rewrite jmE; split => [->|/dyn_inj]. Qed. - +Proof. by split=>[[pf ->]|[pf]]; subst B; [rewrite !eqc | move/dyn_inj=>->]. Qed. diff --git a/finmap/finmap.v b/core/finmap.v similarity index 90% rename from finmap/finmap.v rename to core/finmap.v index 114c727..24b7775 100644 --- a/finmap/finmap.v +++ b/core/finmap.v @@ -17,13 +17,10 @@ limitations under the License. (******************************************************************************) From Coq Require Import ssreflect ssrbool ssrfun. -From mathcomp Require Import ssrnat eqtype seq path. -From fcsl Require Import ordtype. From Coq Require Setoid. - -Set Implicit Arguments. -Unset Strict Implicit. -Unset Printing Implicit Defensive. +From mathcomp Require Import ssrnat eqtype seq path. +From fcsl Require Export ordtype seqperm. +From fcsl Require Import options. Section Def. Variables (K : ordType) (V : Type). @@ -59,7 +56,7 @@ Lemma fmapE (s1 s2 : fmap) : s1 = s2 <-> seq_of s1 = seq_of s2. Proof. split=>[->|] //. move: s1 s2 => [s1 H1] [s2 H2] /= H. -by rewrite H in H1 H2 *; rewrite (bool_irrelevance H1 H2). +by rewrite H in H1 H2 *; rewrite (eq_irrelevance H1 H2). Qed. Lemma sorted_nil : sorted ord (map key [::]). Proof. by []. Qed. @@ -82,7 +79,7 @@ Proof. elim: s k1 k2 v=>[|[k' v'] s IH] k1 k2 v H1 /=; first by rewrite H1. case/andP=>H2 H3; case: ifP=>/= H4; first by rewrite H1 H3 H4. case: ifP=>H5 /=; first by rewrite H1 (eqP H5) H3. -by rewrite H2 IH //; move: (total k2 k'); rewrite H4 H5. +by rewrite H2 IH //; move: (ord_total k2 k'); rewrite H4 H5. Qed. Lemma sorted_ins' s k v : @@ -91,7 +88,7 @@ Proof. case: s=>// [[k' v']] s /= H. case: ifP=>//= H1; first by rewrite H H1. case: ifP=>//= H2; first by rewrite (eqP H2). -by rewrite path_ins' //; move: (total k k'); rewrite H1 H2. +by rewrite path_ins' //; move: (ord_total k k'); rewrite H1 H2. Qed. Definition ins k v s := let: FinMap s' p' := s in FinMap (sorted_ins' k v p'). @@ -104,6 +101,8 @@ Definition rem k s := let: FinMap s' p' := s in FinMap (sorted_filter' k p'). Definition supp s := map key (seq_of s). +Definition cosupp s := map value (seq_of s). + End Ops. Prenex Implicits fnd ins rem supp. @@ -129,7 +128,7 @@ Lemma first_ins' (k : K) (v : V) s : Proof. elim: s=>[|[k1 v1] s IH] H //=. move: (H k1); rewrite inE eq_refl; move/(_ (erefl _)). -case: totalP=>// O _; rewrite IH //. +case: ordP=>// O _; rewrite IH //. by move=>x H'; apply: H; rewrite inE /= H' orbT. Qed. @@ -144,7 +143,7 @@ Qed. Lemma path_supp_ord (s : fmap) k : path ord k (supp s) -> forall m, m \in supp s -> ord k m. Proof. -case: s=>s H; rewrite /supp /= => H1 m H2; case: totalP H1 H2=>//. +case: s=>s H; rewrite /supp /= => H1 m H2; case: ordP H1 H2=>//. - by move=>H1 H2; move: (notin_path (ord_path H1 H2)); case: (m \in _). by move/eqP=>->; move/notin_path; case: (k \in _). Qed. @@ -157,6 +156,19 @@ rewrite inE negb_or; case/andP=>H1 H2. by rewrite eq_sym (negbTE H1); apply: IH. Qed. +Lemma size_ins' (x : K) v (s : seq (K*V)) : + sorted ord (map key s) -> + size (ins' x v s) = + if x \in map key s then size s else (size s).+1. +Proof. +elim: s x v=>[|[k' v'] s IH] //= k v P. +case: ifP=>O /=; rewrite inE eq_sym; +move/path_sorted/(IH k v): (P)=>{IH} E; +case: ordP O=>//= O _. +- by move/notin_path/negbTE: (ord_path O P)=>->. +by rewrite E; case: ifP. +Qed. + Lemma fmapP (s1 s2 : fmap) : (forall k, fnd k s1 = fnd k s2) -> s1 = s2. Proof. rewrite /fnd; move: s1 s2 => [s1 P1][s2 P2] H; rewrite fmapE /=. @@ -165,7 +177,7 @@ elim: s1 P1 s2 P2 H=>[|[k v] s1 IH] /= P1. have S1: sorted ord (map key s1) by apply: path_sorted P1. case=>[|[k2 v2] s2] /= P2; first by move/(_ k); rewrite eq_refl. have S2: sorted ord (map key s2) by apply: path_sorted P2. -move: (IH S1 s2 S2)=>{IH} /= IH H. +move: (IH S1 s2 S2)=>/= {}IH H. move: (notin_path P1) (notin_path P2)=>N1 N2. case E: (k == k2). - rewrite -{k2 E}(eqP E) in P2 H N2 *. @@ -174,7 +186,7 @@ case E: (k == k2). case E: (k == k'); first by rewrite -(eqP E) !notin_filter. by move: (H k'); rewrite E. move: (H k); rewrite eq_refl eq_sym E notin_filter //. -move: (total k k2); rewrite E /=; case/orP=>L1. +move: (ord_total k k2); rewrite E /=; case/orP=>L1. - by apply: notin_path; apply: ord_path P2. move: (H k2); rewrite E eq_refl notin_filter //. by apply: notin_path; apply: ord_path P1. @@ -215,9 +227,9 @@ case: s1 s2=>s1 H1 [s2 H2]; rewrite /supp /=. elim: s1 s2 H1 H2=>[|[k1 _] s1 IH][|[k2 _] s2] //=. - by move=>_ _; move/(_ k2); rewrite inE eq_refl. - by move=>_ _; move/(_ k1); rewrite inE eq_refl. -case: (totalP k1 k2)=>/= O H1 H2. +case: (ordP k1 k2)=>/= O H1 H2. - move/(_ k1); rewrite !inE eq_refl /= eq_sym. - case: totalP O => //= O _. + case: ordP O => //= O _. by move/(ord_path O)/notin_path/negbTE: H2=>->. - rewrite -{k2 O}(eqP O) in H1 H2 *. move=>H; congr (_ :: _); apply: IH. @@ -226,7 +238,7 @@ case: (totalP k1 k2)=>/= O H1 H2. move=>x; move: (H x); rewrite !inE /=. case: eqP=>// -> /= _. by move/notin_path/negbTE: H1=>->; move/notin_path/negbTE: H2=>->. move/(_ k2); rewrite !inE eq_refl /= eq_sym. -case: totalP O=>//= O _. +case: ordP O=>//= O _. by move/(ord_path O)/notin_path/negbTE: H1=>->. Qed. @@ -326,32 +338,32 @@ rewrite /ins; case: s => s H; case H1: (k1 == k2); rewrite fmapE /=. - rewrite (eqP H1) {H1}. elim: s H k2 v1 v2=>[|[k3 v3] s IH] /= H k2 v1 v2; first by rewrite irr eq_refl. - case: (totalP k2 k3)=>H1 /=; rewrite ?irr ?eq_refl //. - case: (totalP k2 k3) H1=>H2 _ //. + case: (ordP k2 k3)=>H1 /=; rewrite ?irr ?eq_refl //. + case: (ordP k2 k3) H1=>H2 _ //. by rewrite IH //; apply: path_sorted H. elim: s H k1 k2 H1 v1 v2=>[|[k3 v3] s IH] H k1 k2 H1 v1 v2 /=. - rewrite H1 eq_sym H1. - by case: (totalP k1 k2) H1=>H2 H1. -case: (totalP k2 k3)=>H2 /=. -- case: (totalP k1 k2) (H1)=>H3 _ //=; last first. - - by case: (totalP k1 k3)=>//= H4; rewrite ?H2 ?H3. - case: (totalP k1 k3)=>H4 /=. - - case: (totalP k2 k1) H3=>//= H3. - by case: (totalP k2 k3) H2=>//=. + by case: (ordP k1 k2) H1=>H2 H1. +case: (ordP k2 k3)=>H2 /=. +- case: (ordP k1 k2) (H1)=>H3 _ //=; last first. + - by case: (ordP k1 k3)=>//= H4; rewrite ?H2 ?H3. + case: (ordP k1 k3)=>H4 /=. + - case: (ordP k2 k1) H3=>//= H3. + by case: (ordP k2 k3) H2=>//=. - rewrite (eqP H4) in H3. - by case: (totalP k2 k3) H2 H3. - by case: (totalP k1 k3) (trans H3 H2) H4. + by case: (ordP k2 k3) H2 H3. + by case: (ordP k1 k3) (trans H3 H2) H4. - rewrite -(eqP H2) {H2} (H1). - case: (totalP k1 k2) (H1)=>//= H2 _; rewrite ?irr ?eq_refl //. + case: (ordP k1 k2) (H1)=>//= H2 _; rewrite ?irr ?eq_refl //. rewrite eq_sym H1. - by case: (totalP k2 k1) H1 H2. -case: (totalP k1 k3)=>H3 /=. + by case: (ordP k2 k1) H1 H2. +case: (ordP k1 k3)=>H3 /=. - rewrite eq_sym H1. - case: (totalP k2 k1) H1 (trans H3 H2)=>//. - by case: (totalP k2 k3) H2=>//=. + case: (ordP k2 k1) H1 (trans H3 H2)=>//. + by case: (ordP k2 k3) H2=>//=. - rewrite (eqP H3). - by case: (totalP k2 k3) H2. -case: (totalP k2 k3)=>H4 /=. + by case: (ordP k2 k3) H2. +case: (ordP k2 k3)=>H4 /=. - by move: (trans H4 H2); rewrite irr. - by rewrite (eqP H4) irr in H2. by rewrite IH //; apply: path_sorted H. @@ -376,24 +388,24 @@ Proof. rewrite /rem; case: s => s H /=; case H1: (k1 == k2); rewrite /= fmapE /=. - rewrite (eqP H1) {H1}. elim: s k2 H=>[|[k3 v3] s IH] k2 /= H; rewrite ?eq_refl 1?eq_sym //. - case: (totalP k3 k2)=>H1 /=; rewrite ?eq_refl //=; - case: (totalP k3 k2) H1=>//= H1 _. + case: (ordP k3 k2)=>H1 /=; rewrite ?eq_refl //=; + case: (ordP k3 k2) H1=>//= H1 _. by rewrite IH //; apply: path_sorted H. elim: s k1 k2 H1 H=>[|[k3 v3] s IH] k1 k2 H1 /= H; first by rewrite eq_sym H1. -case: (totalP k2 k3)=>H2 /=. +case: (ordP k2 k3)=>H2 /=. - rewrite eq_sym H1 /=. - case: (totalP k3 k1)=>H3 /=; case: (totalP k2 k3) (H2)=>//=. + case: (ordP k3 k1)=>H3 /=; case: (ordP k2 k3) (H2)=>//=. rewrite -(eqP H3) in H1 *. rewrite -IH //; last by apply: path_sorted H. rewrite last_ins' /= 1?eq_sym ?H1 //. by apply: ord_path H. - by move: H1; rewrite (eqP H2) /= eq_sym => -> /=; rewrite irr eq_refl. -case: (totalP k3 k1)=>H3 /=. -- case: (totalP k2 k3) H2=>//= H2 _. +case: (ordP k3 k1)=>H3 /=. +- case: (ordP k2 k3) H2=>//= H2 _. by rewrite IH //; apply: path_sorted H. - rewrite -(eqP H3) in H1 *. by rewrite IH //; apply: path_sorted H. -case: (totalP k2 k3) H2=>//= H2 _. +case: (ordP k2 k3) H2=>//= H2 _. by rewrite IH //; apply: path_sorted H. Qed. @@ -446,6 +458,18 @@ Lemma seqof_ins k v (s : fmap) : path ord k (supp s) -> seq_of (ins k v s) = (k, v) :: seq_of s. Proof. by case: s; elim=>[|[k1 v1] s IH] //= _; case/andP=>->. Qed. +Lemma seqof_ins_perm k v (s : fmap) : + k \notin supp s -> perm (seq_of (ins k v s)) ((k, v) :: seq_of s). +Proof. +case: s; elim=>[|[k1 v1] s IH] //= H; case: ifP=>E //. +case: eqP=>[->|N]; first by rewrite /supp /= inE eq_refl. +rewrite /supp /= inE negb_or; case/andP=>K1 {}/IH IH. +suff S : perm (ins' k v s) [:: (k, v) & s]. +- move/(pperm_cons (k1, v1))/pperm_trans: S. + by apply; apply: permutation_swap. +by apply: IH; apply: path_sorted H. +Qed. + Lemma path_supp_ins k1 k v (s : fmap) : ord k1 k -> path ord k1 (supp s) -> path ord k1 (supp (ins k v s)). Proof. @@ -453,7 +477,7 @@ case: s=>s p. elim: s p k1 k v=>[| [k2 v2] s IH] //= p k1 k v H2; first by rewrite H2. case/andP=>H3 H4. have H5: path ord k1 (map key s) by apply: ord_path H4. -rewrite /supp /=; case: (totalP k k2)=>H /=. +rewrite /supp /=; case: (ordP k k2)=>H /=. - by rewrite H2 H H4. - by rewrite H2 (eqP H) H4. rewrite H3 /=. @@ -535,16 +559,16 @@ Lemma fcat_sins k v s1 s2 : fcat s1 (ins k v s2) = ins k v (fcat s1 s2). Proof. elim/fmap_ind': s2 k v s1=>[|k1 v1 s1 H1 IH k2 v2 s2] //. -case: (totalP k2 k1)=>//= H2. +case: (ordP k2 k1)=>//= H2. - have H: path ord k2 (supp (ins k1 v1 s1)). - by apply: (path_supp_ins _ H2); apply: ord_path H1. by rewrite {1}/fcat seqof_ins //= fcat_ins' ?notin_path. - by rewrite IH ins_ins H2 IH ins_ins H2. have H: path ord k1 (supp (ins k2 v2 s1)) by apply: (path_supp_ins _ H2). rewrite ins_ins. -case: (totalP k2 k1) H2 => // H2 _. +case: (ordP k2 k1) H2 => // H2 _. rewrite {1}/fcat seqof_ins //= fcat_ins' ?notin_path // IH ?notin_path //. -rewrite ins_ins; case: (totalP k2 k1) H2 => // H2 _; congr (ins _ _ _). +rewrite ins_ins; case: (ordP k2 k1) H2 => // H2 _; congr (ins _ _ _). by rewrite -/(fcat s2 (ins k2 v2 s1)) IH. Qed. @@ -612,7 +636,7 @@ Lemma supp_eq_ins (s1 s2 : fmap) k1 k2 v1 v2 : Proof. move=>H1 H2 H; move: (H k1) (H k2). rewrite !supp_ins !inE /= !eq_refl (eq_sym k2). -case: totalP=>/= E; last 1 first. +case: ordP=>/= E; last 1 first. - by move: H1; move/(ord_path E); move/notin_path; move/negbTE=>->. - by move: H2; move/(ord_path E); move/notin_path; move/negbTE=>->. rewrite (eqP E) in H1 H2 H * => _ _; split=>// x; move: (H x). @@ -637,10 +661,6 @@ Qed. End FMapInd. - - - - Section Filtering. Variables (K : ordType) (V : Type). Notation fmap := (finMap K V). @@ -660,7 +680,7 @@ Lemma supp_kfilt (p : pred K) (s : fmap) : supp (kfilter p s) = filter p (supp s). Proof. case: s; rewrite /supp /kfilter /kfilter' /=. -elim=>[|[k v] s IH] //= /path_sorted /IH {IH} H. +elim=>[|[k v] s IH] //= /path_sorted {}/IH H. by case E: (p k)=>//=; rewrite H. Qed. @@ -673,7 +693,7 @@ Lemma fnd_kfilt (p : pred K) k (s : fmap) : Proof. case: s; rewrite /fnd /kfilter /kfilter' /=. elim=>[|[k1 v] s IH] /=; first by case: ifP. -move/path_sorted=>/IH {IH} H. +move/path_sorted=>{}/IH H. case: ifP=>E1 /=; first by case: ifP=>E2 //; rewrite -(eqP E2) E1. case: ifP H=>E2 H //=; rewrite H; case: eqP=>// E3. by rewrite -E3 E1 in E2. @@ -764,8 +784,7 @@ Proof. rewrite /disj; case E: (all _ _); first by apply/disj_true/allP. move: E; rewrite all_predC=>/negbFE H. case E: {-1}(supp s1) (H)=>[|k ?]; first by rewrite E. -move/(nth_find k); move: H; rewrite has_find=>/(mem_nth k). -by apply: disj_false. +by move/(nth_find k); move: H; rewrite has_find=>/(mem_nth k); apply: disj_false. Qed. Lemma disjC s1 s2 : disj s1 s2 = disj s2 s1. @@ -854,10 +873,11 @@ Lemma fcatsK (s s1 s2 : fmap) : disj s1 s && disj s2 s -> fcat s1 s = fcat s2 s -> s1 = s2. Proof. elim/fmap_ind': s s1 s2=>// k v s. -move/notin_path=>H IH s1 s2; rewrite !disj_ins -!andbA. -case/and4P=> H1 H2 H3 H4; rewrite !fcat_sins => H5. +move/notin_path=>H IH s1 s2; rewrite !disj_ins. +case/andP; case/andP=>H1 H2; case/andP=>H3 H4. +rewrite !fcat_sins // => H5. apply: IH; first by rewrite H2 H4. -by apply: cancel_ins H5; rewrite supp_fcat negb_or /= ?H1 ?H3 H. +by apply: cancel_ins H5; rewrite supp_fcat negb_or /= ?H1?H3 H. Qed. Lemma fcatKs (s s1 s2 : fmap) : @@ -894,6 +914,19 @@ rewrite !kfilt_ins; case: ifP=>E1; case: ifP=>E2. by apply: IH=>x S; apply: H; rewrite supp_ins inE /= S orbT. Qed. +Lemma seqof_fcat_perm s1 s2 : + disj s1 s2 -> perm (seq_of (fcat s1 s2)) (seq_of s1 ++ seq_of s2). +Proof. +elim/fmap_ind': s2 s1=>[|k v s L IH] s1. +- by rewrite fcats0 cats0. +rewrite fcat_sins disj_ins; case/andP=>D1 D2. +apply: pperm_trans (seqof_ins_perm _ _) _. +- by rewrite supp_fcat inE /= negb_or D1 (notin_path L). +move/(pperm_cons (k, v))/pperm_trans: {IH} (IH _ D2); apply. +apply: pperm_trans (pperm_cons_catCA _ _ _) _; apply: pperm_catL=>//. +by apply/pperm_sym; apply: seqof_ins_perm (notin_path L). +Qed. + End DisjointUnion. Section EqType. @@ -908,33 +941,45 @@ case: eqP; first by move/fmapE=>->; apply: ReflectT. by move=>H; apply: ReflectF; move/fmapE; move/H. Qed. -Canonical Structure fmap_eqMixin := EqMixin feqP. -Canonical Structure fmap_eqType := EqType (finMap K V) fmap_eqMixin. +Definition fmap_eqMixin := EqMixin feqP. +Canonical Structure fmap_eqType := Eval hnf in EqType (finMap K V) fmap_eqMixin. End EqType. (* mapping a function over a contents of a finite map *) Section Map. -Variables (K : ordType) (U V : Type) (f : U -> V). +Variables (K : ordType) (U V : Type) (f : K -> U -> V). Definition mapf' (m : seq (K * U)) : seq (K * V) := - map (fun kv => (key kv, f (value kv))) m. + map (fun kv => (key kv, f (key kv) (value kv))) m. Lemma map_key_mapf (m : seq (K * U)) : map key (mapf' m) = map key m. Proof. by elim: m=>[|[k v] m IH] //=; rewrite IH. Qed. Lemma sorted_map (m : seq (K * U)) : sorted ord (map key m) -> sorted ord (map key (mapf' m)). -Proof. by rewrite map_key_mapf. Qed. +Proof. +elim: m=>[|[k v] m IH] //= H. +rewrite path_min_sorted; first by apply: IH; apply: path_sorted H. +rewrite map_key_mapf. +by apply/(order_path_min _ H);apply/trans. +Qed. Definition mapf (m : finMap K U) : finMap K V := let: FinMap _ pf := m in FinMap (sorted_map pf). -Lemma mapf_ins k v s : mapf (ins k v s) = ins k (f v) (mapf s). +Lemma supp_mapf (s : finMap K U) : + supp (mapf s) = supp s. +Proof. +case: s; rewrite /supp /mapf /mapf' /=. +by elim=>[|[k v] s IH] //= /path_sorted {}/IH ->. +Qed. + +Lemma mapf_ins k v s : mapf (ins k v s) = ins k (f k v) (mapf s). Proof. case: s=>s H; apply/fmapE=>/=. elim: s k v H=>[|[k1 v1] s IH] //= k v H. -rewrite eq_sym; case: totalP=>O //=. +rewrite eq_sym; case: ordP=>O //=. by rewrite IH // (path_sorted H). Qed. @@ -1045,13 +1090,13 @@ Arguments mapk {A B V} f m. Section Zip. Variables (K : ordType) (V : Type) (zip_f : V -> V -> option V). -Variable (unit_f : V -> V). +Variable (unit_f : K -> V -> V). Variable (comm : commutative zip_f). Variable (assoc : forall x y z, obind (zip_f x) (zip_f y z) = obind (zip_f^~ z) (zip_f x y)). -Variable (unitL : forall x, zip_f (unit_f x) x = Some x). -Variable (unitE : forall x y, - (exists z, zip_f x y = Some z) <-> unit_f x = unit_f y). +Variable (unitL : forall k x, zip_f (unit_f k x) x = Some x). +Variable (unitE : forall k k' x y, + (exists z, zip_f x y = Some z) <-> unit_f k x = unit_f k' y). Fixpoint zip' (s1 s2 : seq (K * V)) := match s1, s2 with @@ -1071,7 +1116,7 @@ Definition zip_unit' (s : seq (K * V)) := mapf' unit_f s. Lemma zipC' s1 s2 : zip' s1 s2 = zip' s2 s1. Proof. elim: s1 s2=>[|[k1 v1] s1 IH]; first by case=>//; case. -case=>[|[k2 v2] s2] //=; case: eqVneq => // ->. +case=>[|[k2 v2] s2] //=; rewrite eq_sym; case: eqP=>// ->{k2}. by rewrite comm IH. Qed. @@ -1243,11 +1288,11 @@ Qed. Lemma zunit0 : zip_unit (nil K V) = nil K V. Proof. by apply/fmapE. Qed. -Lemma zunit_ins f k v : zip_unit (ins k v f) = ins k (unit_f v) (zip_unit f). +Lemma zunit_ins f k v : zip_unit (ins k v f) = ins k (unit_f k v) (zip_unit f). Proof. case: f=>s H; apply/fmapE=>/=; rewrite /zip_unit'. elim: s k v H=>[|[k1 v1] s IH] //= k v H. -rewrite eq_sym; case: totalP=>//= O. +rewrite eq_sym; case: ordP=>//= O. by rewrite IH // (path_sorted H). Qed. @@ -1274,4 +1319,3 @@ by move=>H x; move/H/negbTE=>->. Qed. End Zip. - diff --git a/core/ordtype.v b/core/ordtype.v new file mode 100644 index 0000000..6d59309 --- /dev/null +++ b/core/ordtype.v @@ -0,0 +1,568 @@ +(* +Copyright 2010 IMDEA Software Institute +Licensed under the Apache License, Version 2.0 (the "License"); +you may not use this file except in compliance with the License. +You may obtain a copy of the License at + http://www.apache.org/licenses/LICENSE-2.0 +Unless required by applicable law or agreed to in writing, software +distributed under the License is distributed on an "AS IS" BASIS, +WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. +See the License for the specific language governing permissions and +limitations under the License. +*) + +(******************************************************************************) +(* This file defines ordType - the structure for types with a decidable *) +(* (strict) order relation. *) +(* ordType is a subclass of mathcomp's eqType *) +(* This file also defines some important instances of ordType *) +(******************************************************************************) + +From Coq Require Import ssreflect ssrbool ssrfun. +From mathcomp Require Import ssrnat eqtype seq path fintype. +From fcsl Require Import options. + +Module Ordered. + +Section RawMixin. + +Structure mixin_of (T : eqType) := + Mixin {ordering : rel T; + _ : irreflexive ordering; + _ : transitive ordering; + _ : forall x y, x != y -> ordering x y || ordering y x}. + +End RawMixin. + +(* the class takes a naked type T and returns all the *) +(* related mixins; the inherited ones and the added ones *) +Section ClassDef. + +Record class_of (T : Type) := Class { + base : Equality.class_of T; + mixin : mixin_of (EqType T base)}. + +Local Coercion base : class_of >-> Equality.class_of. + +Structure type : Type := Pack {sort : Type; _ : class_of sort}. +Local Coercion sort : type >-> Sortclass. + +Variables (T : Type) (cT : type). +Definition class := let: Pack _ c as cT' := cT return class_of cT' in c. +Definition clone c of phant_id class c := @Pack T c. + +(* produce an ordered type out of the inherited mixins *) +(* equalize m0 and m by means of a phantom; will be exploited *) +(* further down in the definition of OrdType *) +Definition pack b (m0 : mixin_of (EqType T b)) := + fun m & phant_id m0 m => Pack (@Class T b m). + +Definition eqType := Eval hnf in EqType cT class. + +End ClassDef. + +Module Exports. +Coercion base : class_of >-> Equality.class_of. +Coercion sort : type >-> Sortclass. +Coercion eqType : type >-> Equality.type. +Canonical Structure eqType. +Notation ordType := Ordered.type. +Notation OrdMixin := Mixin. +Notation OrdType T m := (@pack T _ m _ id). +Definition ord T : rel (sort T) := (ordering (mixin (class T))). +Notation "[ 'ordType' 'of' T 'for' cT ]" := (@clone T cT _ idfun) + (at level 0, format "[ 'ordType' 'of' T 'for' cT ]") : form_scope. +Notation "[ 'ordType' 'of' T ]" := (@clone T _ _ id) + (at level 0, format "[ 'ordType' 'of' T ]") : form_scope. +End Exports. +End Ordered. +Export Ordered.Exports. + +Definition oleq (T : ordType) (t1 t2 : T) := ord t1 t2 || (t1 == t2). + +Prenex Implicits ord oleq. + +Section Lemmas. +Variable T : ordType. +Implicit Types x y : T. + +Lemma irr : irreflexive (@ord T). +Proof. by case: T=>s [b [m]]. Qed. + +Lemma trans : transitive (@ord T). +Proof. by case: T=>s [b [m]]. Qed. + +Lemma semiconnex x y : x != y -> ord x y || ord y x. +Proof. by case: T x y=>s [b [m]]. Qed. + +Lemma ord_total x y : [|| ord x y, x == y | ord y x]. +Proof. +case E: (x == y)=>/=; first by rewrite orbT. +by apply: semiconnex; rewrite negbT. +Qed. + +Lemma nsym x y : ord x y -> ord y x -> False. +Proof. by move=>E1 E2; move: (trans E1 E2); rewrite irr. Qed. + +Lemma orefl x : oleq x x. +Proof. by rewrite /oleq eq_refl orbT. Qed. + +Lemma otrans : transitive (@oleq T). +Proof. +move=>x y z /=; case/orP; last by move/eqP=>->. +rewrite /oleq; move=>T1; case/orP; first by move/(trans T1)=>->. +by move/eqP=><-; rewrite T1. +Qed. + +Lemma oantisym : antisymmetric (@oleq T). +Proof. +move=>x y; rewrite /oleq (eq_sym x). +case: eqP=>// _; rewrite !orbF=>/andP [H1 H2]. +by move: (trans H1 H2); rewrite irr. +Qed. + +Lemma subrel_ord : subrel (@ord T) oleq. +Proof. exact: subrelUl. Qed. + +Lemma sorted_oleq s : sorted (@ord T) s -> sorted (@oleq T) s. +Proof. by case: s=>//= x s; apply/sub_path/subrel_ord. Qed. + +Lemma path_filter (r : rel T) (tr : transitive r) s (p : pred T) x : + path r x s -> path r x (filter p s). +Proof. exact: (subseq_order_path tr (filter_subseq p s)). Qed. + +Lemma eq_sorted_ord (s1 s2 : seq T) : + sorted ord s1 -> sorted ord s2 -> s1 =i s2 -> s1 = s2. +Proof. by apply: eq_sorted_irr; [apply: trans | apply: irr]. Qed. + +End Lemmas. + +Hint Resolve orefl irr trans otrans oantisym : core. + +Section Totality. +Variable K : ordType. + +Variant total_spec (x y : K) : bool -> bool -> bool -> Type := +| total_spec_lt of ord x y : total_spec x y true false false +| total_spec_eq of x == y : total_spec x y false true false +| total_spec_gt of ord y x : total_spec x y false false true. + +Lemma ordP x y : total_spec x y (ord x y) (x == y) (ord y x). +Proof. +case H1: (x == y). +- by rewrite (eqP H1) irr; apply: total_spec_eq. +case H2: (ord x y); case H3: (ord y x). +- by case: (nsym H2 H3). +- by apply: total_spec_lt H2. +- by apply: total_spec_gt H3. +by move: (ord_total x y); rewrite H1 H2 H3. +Qed. +End Totality. + +Lemma eq_oleq (K : ordType) (x y : K) : (x == y) = oleq x y && oleq y x. +Proof. by rewrite /oleq (eq_sym x); case: ordP. Qed. + +(* Monotone (i.e. strictly increasing) functions for Ord Types *) +Section Mono. +Variables (A B :ordType). + +Definition strictly_increasing f x y := @ord A x y -> @ord B (f x) (f y). + +Structure mono : Type := Mono + {fun_of: A -> B; _: forall x y, strictly_increasing fun_of x y}. + +End Mono. +Arguments strictly_increasing {A B} f x y. +Arguments Mono {A B _} _. + +Section Weakening. +Variable T : ordType. +Implicit Types x y : T. + +Lemma ordW x y: ord x y -> oleq x y. +Proof. by rewrite /oleq =>->//. Qed. + +Lemma oleqNord x y: oleq x y = ~~ (ord y x). +Proof. by rewrite/oleq; case:(ordP x y)=>//. Qed. + +Variant oleq_spec x y : bool -> bool -> Type := +| oleq_spec_le of oleq x y : oleq_spec x y true false +| oleq_spec_gt of ord y x : oleq_spec x y false true. + +Lemma oleqP x y : oleq_spec x y (oleq x y) (ord y x). +Proof. +case:(ordP x y). +- by move=>/ordW O; rewrite O; apply: oleq_spec_le. +- by move=>/eqP E; rewrite E orefl; apply: oleq_spec_le; apply: orefl. +by move=>O; rewrite oleqNord O //=; apply: oleq_spec_gt. +Qed. + +Lemma oleq_total x y: oleq x y || oleq y x. +Proof. by case:oleqP=>// /ordW ->//. Qed. + +End Weakening. + +Section NatOrd. +Lemma irr_ltn_nat : irreflexive ltn. Proof. by move=>x; rewrite /= ltnn. Qed. +Lemma trans_ltn_nat : transitive ltn. Proof. by apply: ltn_trans. Qed. +Lemma semiconn_ltn_nat x y : x != y -> (x < y) || (y < x). +Proof. by case: ltngtP. Qed. + +Definition nat_ordMixin := OrdMixin irr_ltn_nat trans_ltn_nat semiconn_ltn_nat. +Canonical Structure nat_ordType := Eval hnf in OrdType nat nat_ordMixin. +End NatOrd. + +Section ProdOrd. +Variables K T : ordType. + +(* lexicographic ordering *) +Definition lex : rel (K * T) := + fun x y => if x.1 == y.1 then ord x.2 y.2 else ord x.1 y.1. + +Lemma irr_lex : irreflexive lex. +Proof. by move=>x; rewrite /lex eq_refl irr. Qed. + +Lemma trans_lex : transitive lex. +Proof. +move=>[x1 x2][y1 y2][z1 z2]; rewrite /lex /=. +case: ifP=>H1; first by rewrite (eqP H1); case: eqP=>// _; apply: trans. +case: ifP=>H2; first by rewrite (eqP H2) in H1 *; rewrite H1. +case: ifP=>H3; last by apply: trans. +by rewrite (eqP H3)=>R1; move/(nsym R1). +Qed. + +Lemma semiconn_lex : forall x y, x != y -> lex x y || lex y x. +Proof. +move=>[x1 x2][y1 y2]; rewrite /lex /=. +by case: ifP=>H1 H2; rewrite eq_sym H1 semiconnex //; move: H2; rewrite -pair_eqE /= H1 //. +Qed. + +Definition prod_ordMixin := OrdMixin irr_lex trans_lex semiconn_lex. +Canonical Structure prod_ordType := Eval hnf in OrdType (K * T) prod_ordMixin. +End ProdOrd. + +Section FinTypeOrd. +Variable T : finType. + +Definition ordf : rel T := + fun x y => index x (enum T) < index y (enum T). + +Lemma irr_ordf : irreflexive ordf. +Proof. by move=>x; rewrite /ordf ltnn. Qed. + +Lemma trans_ordf : transitive ordf. +Proof. by move=>x y z; rewrite /ordf; apply: ltn_trans. Qed. + +Lemma semiconn_ordf x y : x != y -> ordf x y || ordf y x. +Proof. +rewrite /ordf; case: ltngtP=>//= H. +have [H1 H2]: x \in enum T /\ y \in enum T by rewrite !mem_enum. +by rewrite -(nth_index x H1) -(nth_index x H2) H eq_refl. +Qed. + +Definition fin_ordMixin := OrdMixin irr_ordf trans_ordf semiconn_ordf. +End FinTypeOrd. + +(* notation to let us write I_n instead of (ordinal_finType n) *) +Notation "[ 'fin_ordMixin' 'of' T ]" := + (fin_ordMixin _ : Ordered.mixin_of [eqType of T]) (at level 0). + +Definition ordinal_ordMixin n := [fin_ordMixin of 'I_n]. +Canonical Structure ordinal_ordType n := Eval hnf in OrdType 'I_n (ordinal_ordMixin n). + +Section SeqOrd. +Variable (T : ordType). + +Fixpoint ords x : pred (seq T) := + fun y => match x , y with + | [::] , [::] => false + | [::] , t :: ts => true + | x :: xs , y :: ys => if x == y then ords xs ys + else ord x y + | _ :: _ , [::] => false + end. + +Lemma irr_ords : irreflexive ords. +Proof. by elim=>//= a l ->; rewrite irr; case:eqP=> //=. Qed. + +Lemma trans_ords : transitive ords. +Proof. +elim=>[|y ys IHy][|x xs][|z zs]//=. +case:eqP=>//[->|H0];case:eqP=>//H; first by move/IHy; apply. +- by case:eqP=>//; rewrite -H; first (by move/H0). +case:eqP=>//[->|H1] H2; first by move/(nsym H2). +by move/(trans H2). +Qed. + +Lemma semiconn_ords : forall x y, x != y -> ords x y || ords y x. +Proof. +elim=>[|x xs IH][|y ys]//=. +case:ifP => H1 H2; rewrite eq_sym H1. +- by apply: IH; move: H2; rewrite -eqseqE /= H1 andTb. +by rewrite semiconnex // H1. +Qed. + +Definition seq_ordMixin := OrdMixin irr_ords trans_ords semiconn_ords. +Canonical Structure seq_ordType := Eval hnf in OrdType (seq T) seq_ordMixin. +End SeqOrd. + +(* A trivial total ordering for Unit *) +Section unitOrd. +Let ordtt (x y : unit) := false. + +Lemma irr_tt : irreflexive ordtt. +Proof. by []. Qed. + +Lemma trans_tt : transitive ordtt. +Proof. by []. Qed. + +Lemma semiconn_tt x y : x != y -> ordtt x y || ordtt y x. +Proof. by []. Qed. + +Let unit_ordMixin := OrdMixin irr_tt trans_tt semiconn_tt. +Canonical Structure unit_ordType := Eval hnf in OrdType unit unit_ordMixin. +End unitOrd. + + +(* ordering with path, seq and last *) + +Lemma eq_last (A : eqType) (s : seq A) x y : + x \in s -> last y s = last x s. +Proof. by elim: s x y=>[|w s IH]. Qed. + +Lemma seq_last_in (A : eqType) (s : seq A) x : + last x s \notin s -> s = [::]. +Proof. +case: (lastP s)=>{s} // s y; case: negP=>//; elim; rewrite last_rcons. +by elim: s=>[|y' s IH]; rewrite /= inE // IH orbT. +Qed. + +Lemma path_last (A : eqType) (s : seq A) leT x : + transitive leT -> path leT x s -> + (x == last x s) || leT x (last x s). +Proof. +move=>T /(order_path_min T) /allP; case: s=>[|a s] H /=. +- by rewrite eq_refl. +by rewrite (H (last a s)) ?orbT // mem_last. +Qed. + +Lemma path_lastR (A : eqType) (s : seq A) leT x : + reflexive leT -> transitive leT -> + path leT x s -> leT x (last x s). +Proof. by move=>R T P; case: eqP (path_last T P)=>// <- _; apply: R. Qed. + +(* in a sorted list, the last element is maximal *) +(* and the maximal element is last *) + +Lemma sorted_last_key_max (A : eqType) (s : seq A) leT x y : + transitive leT -> sorted leT s -> x \in s -> + (x == last y s) || leT x (last y s). +Proof. +move=>T; elim: s x y=>[|z s IH] //= x y H; rewrite inE. +case: eqP=>[->|] /= _; first by apply: path_last. +by apply: IH (path_sorted H). +Qed. + +Lemma sorted_last_key_maxR (A : eqType) (s : seq A) leT x y : + reflexive leT -> transitive leT -> + sorted leT s -> x \in s -> leT x (last y s). +Proof. +move=>R T S X; case/orP: (sorted_last_key_max y T S X)=>// /eqP <-. +by apply: R. +Qed. + +Lemma sorted_max_key_last (A : eqType) (s : seq A) leT x y : + transitive leT -> antisymmetric leT -> + sorted leT s -> x \in s -> + (forall z, z \in s -> leT z x) -> last y s = x. +Proof. +move=>T S; elim: s x y => [|w s IH] //= x y; rewrite inE /=. +case: eqP=>[<- /= H1 _ H2 | _ H /= H1 H2]; last first. +- apply: IH (path_sorted H) H1 _ => z H3; apply: H2. + by rewrite inE /= H3 orbT. +case/orP: (path_last T H1)=>[/eqP //|] X. +by apply: S; rewrite X H2 ?mem_last. +Qed. + +Lemma max_key_last_notin (A : eqType) (s : seq A) (leT : rel A) x y : + leT y x -> (forall z, z \in s -> leT z x) -> leT (last y s) x. +Proof. +elim: s x y=>[|w s IH] //= x y H1 H2; apply: IH. +- by apply: (H2 w); rewrite inE eq_refl. +by move=>z D; apply: H2; rewrite inE D orbT. +Qed. + +Lemma seq_last_mono (A : eqType) (s1 s2 : seq A) leT x : + transitive leT -> path leT x s1 -> path leT x s2 -> + {subset s1 <= s2} -> + (last x s1 == last x s2) || leT (last x s1) (last x s2). +Proof. +move=>T; case: s1=>/= [_ H1 _|a s]; first by apply: path_last H1. +case/andP=>H1 H2 H3 H; apply: sorted_last_key_max (path_sorted H3) _=>//. +apply: {x s2 H1 H3} H; rewrite inE orbC -implyNb. +by case E: (_ \notin _) (@seq_last_in _ s a)=>//= ->. +Qed. + +Lemma seq_last_monoR (A : eqType) (s1 s2 : seq A) leT x : + reflexive leT -> transitive leT -> + path leT x s1 -> path leT x s2 -> + {subset s1 <= s2} -> + leT (last x s1) (last x s2). +Proof. by move=>R T P1 P2 S; case: eqP (seq_last_mono T P1 P2 S)=>[->|]. Qed. + +Lemma ord_path (A : eqType) (s : seq A) leT (x y : A) : + transitive leT -> + leT x y -> path leT y s -> path leT x s. +Proof. +move=>T; elim: s x y=>[|k s IH] x y //= H1 /andP [H2 ->]. +by rewrite (T _ _ _ H1 H2). +Qed. + +Lemma path_mem (A : eqType) (s : seq A) leT x y : + transitive leT -> + path leT x s -> y \in s -> leT x y. +Proof. +move=>T; elim: s x=>[|z s IH] x //= /andP [O P]. +rewrite inE; case/orP=>[/eqP -> //|]. +by apply: IH; apply: ord_path O P. +Qed. + +Lemma sorted_rcons (A : eqType) (s : seq A) leT (y : A) : + sorted leT s -> (forall x, x \in s -> leT x y) -> + sorted leT (rcons s y). +Proof. +elim: s=>[|a s IH] //= P H; rewrite rcons_path P /=. +by apply: H (mem_last _ _). +Qed. + +Lemma sorted_subset_subseq (A : eqType) (s1 s2 : seq A) leT : + irreflexive leT -> transitive leT -> + sorted leT s1 -> sorted leT s2 -> + {subset s1 <= s2} -> subseq s1 s2. +Proof. +move=>R T S1 S2 H. +suff -> : s1 = filter (fun x => x \in s1) s2 by apply: filter_subseq. +apply: eq_sorted_irr S1 _ _=>//; first by rewrite sorted_filter. +by move=>k; rewrite mem_filter; case S : (_ \in _)=>//; rewrite (H _ S). +Qed. + +Lemma sorted_ord_index (A : eqType) (s : seq A) leT x y : + irreflexive leT -> transitive leT -> + sorted leT s -> x \in s -> leT x y -> index x s < index y s. +Proof. +move=>I T S P H; elim: s S P=>[|z s IH] //= P; rewrite !inE !(eq_sym z). +case: eqP H P=>[<-{z} H P _|_ H P /= X]; first by case: eqP H=>[<-|] //; rewrite I. +case: eqP H P=>[<-{z} H|_ H]; last first. +- by move/path_sorted=>S; rewrite ltnS; apply: IH. +by move/(path_mem T)/(_ X)=>/(T _ _ _ H); rewrite I. +Qed. + +Lemma sorted_index_ord (A : eqType) (s : seq A) leT x y : + irreflexive leT -> transitive leT -> + sorted leT s -> y \in s -> index x s < index y s -> leT x y. +Proof. +move=>I T; elim: s=>[|z s IH] //= P; rewrite !inE !(eq_sym z). +case: eqP=>//= /eqP N; case: eqP P=>[-> P /(path_mem T P)|_ P] //. +by rewrite ltnS; apply: IH; apply: path_sorted P. +Qed. + +Lemma nth_cons A (a x : A) (s : seq A) (n : nat) : + nth a (x :: s) n = if n == 0 then x else nth a s n.-1. +Proof. by case: n. Qed. + +Lemma nth_base A (s : seq A) k1 k2 i : + i < size s -> nth k1 s i = nth k2 s i. +Proof. +elim: s i=>[|x xs IH] //= i K; rewrite !nth_cons. +by case: eqP=>//; case: i K=>// i; rewrite ltnS=>/IH ->. +Qed. + +Lemma nth_path_head (A : eqType) (s : seq A) leT x0 k i : + transitive leT -> + i <= size s -> path leT k s -> + (k == nth x0 (k::s) i) || leT k (nth x0 (k::s) i). +Proof. +move=>T; case: (posnP i)=>[->|N S P]; first by rewrite eq_refl. +apply/orP; right; elim: i N S P=>[|i] //; case: s=>//= x xs IH _. +rewrite ltnS=>N /andP [H1 H2]; case: i IH N=>//= i /(_ (erefl _)) IH N. +rewrite !ltnS in IH; move: (IH (ltnW N)); rewrite H1 H2=>/(_ (erefl _)). +by move/T; apply; apply/pathP. +Qed. + +Lemma nth_path_last (A : eqType) (s : seq A) leT x0 k i : + transitive leT -> + i < size s -> path leT k s -> + (nth x0 s i == last k s) || leT (nth x0 s i) (last k s). +Proof. +move=>T S P. +suff : forall z, z \in s -> (z == last k s) || leT z (last k s). +- by apply; rewrite mem_nth. +move=>z; apply: sorted_last_key_max=>//. +by apply: path_sorted P. +Qed. + +Lemma nth_consS A (s : seq A) x0 k i : nth x0 s i = nth x0 (k::s) i.+1. +Proof. by []. Qed. + +Lemma nth_leT A (s : seq A) leT x0 k i : + i < size s -> path leT k s -> + leT (nth x0 (k::s) i) (nth x0 s i). +Proof. +elim: i k s=>[|i IH] k s; first by case: s=>[|x xs] //= _ /andP []. +by case: s IH=>[|x xs] //= IH N /andP [P1 P2]; apply: IH. +Qed. + +Lemma nth_ltn_mono A (s : seq A) leT x0 k i j : + transitive leT -> + i <= size s -> j <= size s -> + path leT k s -> + i < j -> leT (nth x0 (k::s) i) (nth x0 (k::s) j). +Proof. +move=>T S1 S2 P; elim: j S2=>[|j IH] //= S2. +move: (nth_leT x0 S2 P)=>L. +rewrite ltnS leq_eqVlt=>/orP; case=>[/eqP -> //|]. +by move/(IH (ltnW S2))/T; apply. +Qed. + +Lemma nth_mono_ltn A (s : seq A) leT x0 k i j : + irreflexive leT -> + transitive leT -> + i <= size s -> j <= size s -> + path leT k s -> + leT (nth x0 (k::s) i) (nth x0 (k::s) j) -> i < j. +Proof. +move=>I T S1 S2 P; case: ltngtP=>//; last by move=>->; rewrite I. +by move/(nth_ltn_mono x0 T S2 S1 P)/T=>X /X; rewrite I. +Qed. + +Lemma nth_between (A : eqType) (s : seq A) leT x0 k z i : + irreflexive leT -> + transitive leT -> + path leT k s -> + leT (nth x0 (k::s) i) z -> leT z (nth x0 s i) -> z \notin s. +Proof. +move=>I T P H1 H2; apply/negP=>Z; move: H1 H2. +case: (leqP i (size s))=>N; last first. +- rewrite !nth_default ?(ltnW N) //= => H. + by move/(T _ _ _ H); rewrite I. +have S : index z s < size s by rewrite index_mem. +rewrite -(nth_index x0 Z) !(nth_consS s x0 k). +move/(nth_mono_ltn I T N S P)=>K1. +move/(nth_mono_ltn I T S (leq_trans K1 S) P); rewrite ltnS. +by case: ltngtP K1. +Qed. + +(* how to prove that something's sorted via index? *) + +Lemma index_sorted (A : eqType) (s : seq A) (leT : rel A) : + uniq s -> + (forall a b, a \in s -> b \in s -> index a s < index b s -> leT a b) -> + sorted leT s. +Proof. +elim: s=>[|x xs IH] //= U H; have P : all (leT x) xs. +- apply/allP=>z Z; apply: H; rewrite ?(inE,eq_refl,Z,orbT) //. + by case: ifP U=>// /eqP ->; rewrite Z. +rewrite (path_min_sorted P); apply: IH=>[|a b Xa Xb N]; first by case/andP: U. +apply: H; rewrite ?(inE,Xa,Xb,orbT) //. +by case: eqP U=>[->|]; case: eqP=>[->|]; rewrite ?(Xa,Xb). +Qed. diff --git a/core/pred.v b/core/pred.v new file mode 100644 index 0000000..55ca2e4 --- /dev/null +++ b/core/pred.v @@ -0,0 +1,1238 @@ +(* +Copyright 2010 IMDEA Software Institute +Licensed under the Apache License, Version 2.0 (the "License"); +you may not use this file except in compliance with the License. +You may obtain a copy of the License at + http://www.apache.org/licenses/LICENSE-2.0 +Unless required by applicable law or agreed to in writing, software +distributed under the License is distributed on an "AS IS" BASIS, +WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. +See the License for the specific language governing permissions and +limitations under the License. +*) + +(******************************************************************************) +(* This file re-implements some of ssrbool's entities in Prop universe *) +(******************************************************************************) + +From Coq Require Import ssreflect ssrbool ssrfun Setoid. +From mathcomp Require Import ssrnat seq eqtype. +From fcsl Require Import prelude. +From fcsl Require Import options. +Set Warnings "-projection-no-head-constant". + +(* First some basic propositional equalities *) + +Lemma andTp p : True /\ p <-> p. Proof. by intuition. Qed. +Lemma andpT p : p /\ True <-> p. Proof. by intuition. Qed. +Lemma andFp p : False /\ p <-> False. Proof. by intuition. Qed. +Lemma andpF p : p /\ False <-> False. Proof. by intuition. Qed. +Lemma orTp p : True \/ p <-> True. Proof. by intuition. Qed. +Lemma orpT p : p \/ True <-> True. Proof. by intuition. Qed. +Lemma orFp p : False \/ p <-> p. Proof. by intuition. Qed. +Lemma orpF p : p \/ False <-> p. Proof. by intuition. Qed. + +Lemma iffC p q : (p <-> q) <-> (q <-> p). Proof. by intuition. Qed. + +Declare Scope rel_scope. +Delimit Scope rel_scope with rel. +Open Scope rel_scope. + +(**************************************************************************) +(* We follow ssrbool, and provide four different types of predicates. *) +(* *) +(* (1) Pred is the type of propositional functions *) +(* (2) Simpl_Pred is the type of predicates that automatically simplify *) +(* when used in an applicative position. *) +(* (3) Mem_Pred is for predicates that support infix notation x \In P *) +(* (4) PredType is the structure for interpreting various types, such as *) +(* lists, tuples, etc. as predicates. *) +(* *) +(* Important point is that custom lemmas over predicates can be stated in *) +(* terms of Pred, while Simpl_Pred, Mem_Pred and PredType are for *) +(* technical developments used in this file only. More on this point *) +(* can be found in ssrbool. *) +(**************************************************************************) + +Definition Pred T := T -> Prop. +Identity Coercion fun_of_Pred : Pred >-> Funclass. + +Notation xPred0 := (fun _ => False). +Notation xPred1 := (fun x y => x = y). +Notation xPredT := (fun _ => True). +Notation xPredI := (fun (p1 p2 : Pred _) x => p1 x /\ p2 x). +Notation xPredU := (fun (p1 p2 : Pred _) x => p1 x \/ p2 x). +Notation xPredC := (fun (p : Pred _) x => ~ p x). +Notation xPredD := (fun (p1 p2 : Pred _) x => ~ p2 x /\ p1 x). +Notation xPreim := (fun f (p : Pred _) x => p (f x)). + +Section Predicates. +Variable T : Type. + +(* simple predicates *) + +Definition Simpl_Pred := simpl_fun T Prop. +Definition SimplPred (p : Pred T) : Simpl_Pred := SimplFun p. +Coercion Pred_of_Simpl (p : Simpl_Pred) : Pred T := p : T -> Prop. + +(* it's useful to declare the operations as simple predicates, so that *) +(* complex expressions automatically reduce when used in applicative *) +(* positions *) + +Definition Pred0 := SimplPred xPred0. +Definition Pred1 x := SimplPred (xPred1 x). +Definition PredT := SimplPred xPredT. +Definition PredI p1 p2 := SimplPred (xPredI p1 p2). +Definition PredU p1 p2 := SimplPred (xPredU p1 p2). +Definition PredC p := SimplPred (xPredC p). +Definition PredD p1 p2 := SimplPred (xPredD p1 p2). +Definition Preim rT f (d : Pred rT) := SimplPred (xPreim f d). + +(* membership predicates *) + +Variant Mem_Pred : Type := MemProp of Pred T. +Definition isMem pT toPred mem := mem = (fun p : pT => MemProp [eta toPred p]). + +(* the general structure for predicates *) + +Structure PredType : Type := PropPredType { + Pred_Sort :> Type; + toPred : Pred_Sort -> Pred T; + _ : {mem | isMem toPred mem}}. + +Definition mkPredType pT toP := PropPredType (exist (@isMem pT toP) _ (erefl _)). + +(* Pred, SimplPred, Mem_Pred, pred and simpl_pred are PredType's *) +Canonical Structure PredPredType := Eval hnf in @mkPredType (Pred T) id. +Canonical Structure SimplPredPredType := Eval hnf in mkPredType Pred_of_Simpl. +Canonical Structure PropfunPredType := Eval hnf in @mkPredType (T -> Prop) id. +Coercion Pred_of_Mem mp : Pred_Sort PredPredType := + let: MemProp p := mp in [eta p]. +Canonical Structure MemPredType := Eval hnf in mkPredType Pred_of_Mem. +Canonical Structure predPredType := Eval hnf in @mkPredType (pred T) id. +Canonical Structure simplpredPredType := + Eval hnf in @mkPredType (simpl_pred T) (fun p x => p x). + +End Predicates. + +Arguments Pred0 {T}. +Arguments PredT {T}. +Prenex Implicits Pred0 PredT PredI PredU PredC PredD Preim. + +Notation "r1 +p r2" := (PredU r1 r2 : Pred _) + (at level 55, right associativity) : rel_scope. +Notation "r1 *p r2" := (xPredI r1 r2 : Pred _) + (at level 45, right associativity) : rel_scope. + +Notation "[ 'Pred' : T | E ]" := (SimplPred (fun _ : T => E)) + (at level 0, format "[ 'Pred' : T | E ]") : fun_scope. +Notation "[ 'Pred' x | E ]" := (SimplPred (fun x => E)) + (at level 0, x ident, format "[ 'Pred' x | E ]") : fun_scope. +Notation "[ 'Pred' x : T | E ]" := (SimplPred (fun x : T => E)) + (at level 0, x ident, only parsing) : fun_scope. +Notation "[ 'Pred' x y | E ]" := (SimplPred (fun t => let: (x, y) := t in E)) + (at level 0, x ident, y ident, format "[ 'Pred' x y | E ]") : fun_scope. +Notation "[ 'Pred' x y : T | E ]" := + (SimplPred (fun t : (T*T) => let: (x, y) := t in E)) + (at level 0, x ident, y ident, only parsing) : fun_scope. + +Definition repack_Pred T pT := + let: PropPredType _ a mP := pT return {type of @PropPredType T for pT} -> _ in + fun k => k a mP. + +Notation "[ 'PredType' 'of' T ]" := (repack_Pred (fun a => @PropPredType _ T a)) + (at level 0, format "[ 'PredType' 'of' T ]") : form_scope. + +Notation Pred_Class := (Pred_Sort (PredPredType _)). +Coercion Sort_of_Simpl_Pred T (p : Simpl_Pred T) : Pred_Class := p : Pred T. + +Definition PredArgType := Type. +Coercion Pred_of_argType (T : PredArgType) : Simpl_Pred T := PredT. + +Notation "{ :: T }" := (T%type : PredArgType) + (at level 0, format "{ :: T }") : type_scope. + +(* These must be defined outside a Section because "cooking" kills the *) +(* nosimpl tag. *) +Definition Mem T (pT : PredType T) : pT -> Mem_Pred T := + nosimpl (let: PropPredType _ _ (exist mem _) := pT return pT -> _ in mem). +Definition InMem T x mp := nosimpl Pred_of_Mem T mp x. + +Prenex Implicits Mem. + +(* Membership Predicates can be used as simple ones *) +Coercion Pred_of_Mem_Pred T mp := [Pred x : T | InMem x mp]. + +(* equality and subset *) + +Definition EqPredType T (pT : PredType T) (p1 p2 : pT) := + forall x : T, toPred p1 x <-> toPred p2 x. + +Definition SubPredType T (pT : PredType T) (p1 p2 : pT) := + forall x : T, toPred p1 x -> toPred p2 x. + +Definition EqSimplPred T (p1 p2 : Simpl_Pred T) := EqPredType p1 p2. +Definition SubSimplPred T (p1 p2 : Simpl_Pred T) := SubPredType p1 p2. + +Definition EqPredFun T1 T2 (pT2 : PredType T2) p1 p2 := + forall x : T1, @EqPredType T2 pT2 (p1 x) (p2 x). +Definition SubPredFun T1 T2 (pT2 : PredType T2) p1 p2 := + forall x : T1, @SubPredType T2 pT2 (p1 x) (p2 x). + +Definition EqMem T p1 p2 := forall x : T, InMem x p1 <-> InMem x p2. +Definition SubMem T p1 p2 := forall x : T, InMem x p1 -> InMem x p2. + +Notation "A <~> B" := (@EqPredType _ _ A B) + (at level 70, no associativity) : rel_scope. +Notation "A ~> B" := (@SubPredType _ _ A B) + (at level 70, no associativity) : rel_scope. +Notation "A <~1> B" := (@EqPredFun _ _ _ A B) + (at level 70, no associativity) : rel_scope. +Notation "A ~1> B" := (@SubPredFun _ _ _ A B) + (at level 70, no associativity) : rel_scope. + +Notation "x \In A" := (InMem x (Mem A)) + (at level 70, no associativity) : rel_scope. +Notation "x \Notin A" := (~ (x \In A)) + (at level 70, no associativity) : rel_scope. +Notation "A =p B" := (EqMem (Mem A) (Mem B)) + (at level 70, no associativity) : type_scope. +Notation "A <=p B" := (SubMem (Mem A) (Mem B)) + (at level 70, no associativity) : type_scope. + +(* Some notation for turning PredTypes into Pred or Simple Pred *) +Notation "[ 'Mem' A ]" := (Pred_of_Simpl (Pred_of_Mem_Pred (Mem A))) + (at level 0, only parsing) : fun_scope. +Notation "[ 'PredI' A & B ]" := (PredI [Mem A] [Mem B]) + (at level 0, format "[ 'PredI' A & B ]") : fun_scope. +Notation "[ 'PredU' A & B ]" := (PredU [Mem A] [Mem B]) + (at level 0, format "[ 'PredU' A & B ]") : fun_scope. +Notation "[ 'PredD' A & B ]" := (PredD [Mem A] [Mem B]) + (at level 0, format "[ 'PredD' A & B ]") : fun_scope. +Notation "[ 'PredC' A ]" := (PredC [Mem A]) + (at level 0, format "[ 'PredC' A ]") : fun_scope. +Notation "[ 'Preim' f 'of' A ]" := (Preim f [Mem A]) + (at level 0, format "[ 'Preim' f 'of' A ]") : fun_scope. + +Notation "[ 'Pred' x \In A ]" := [Pred x | x \In A] + (at level 0, x ident, format "[ 'Pred' x \In A ]") : fun_scope. +Notation "[ 'Pred' x \In A | E ]" := [Pred x | (x \In A) /\ E] + (at level 0, x ident, format "[ 'Pred' x \In A | E ]") : fun_scope. +Notation "[ 'Pred' x y \In A & B | E ]" := + [Pred x y | (x \In A) /\ (y \In B) /\ E] + (at level 0, x ident, y ident, + format "[ 'Pred' x y \In A & B | E ]") : fun_scope. +Notation "[ 'Pred' x y \In A & B ]" := [Pred x y | (x \In A) /\ (y \In B)] + (at level 0, x ident, y ident, + format "[ 'Pred' x y \In A & B ]") : fun_scope. +Notation "[ 'Pred' x y \In A | E ]" := [Pred x y \In A & A | E] + (at level 0, x ident, y ident, + format "[ 'Pred' x y \In A | E ]") : fun_scope. +Notation "[ 'Pred' x y \In A ]" := [Pred x y \In A & A] + (at level 0, x ident, y ident, + format "[ 'Pred' x y \In A ]") : fun_scope. + +Section Simplifications. +Variables (T : Type) (pT : PredType T). + +Lemma Mem_toPred : forall (p : pT), Mem (toPred p) = Mem p. +Proof. by rewrite /Mem; case: pT => T1 app1 [mem1 /= ->]. Qed. + +Lemma toPredE x (p : pT) : toPred p x = (x \In p). +Proof. by rewrite -Mem_toPred. Qed. + +Lemma In_Simpl x (p : Simpl_Pred T) : (x \In p) = p x. +Proof. by []. Qed. + +Lemma Simpl_PredE (p : Pred T) : p <~> [Pred x | p x]. +Proof. by []. Qed. + +Lemma Mem_Simpl (p : Simpl_Pred T) : Mem p = p :> Pred T. +Proof. by []. Qed. + +Definition MemE := Mem_Simpl. (* could be extended *) + +Lemma Mem_Mem (p : pT) : (Mem (Mem p) = Mem p) * (Mem [Mem p] = Mem p). +Proof. by rewrite -Mem_toPred. Qed. + +End Simplifications. + +(**************************************) +(* Definitions and lemmas for setoids *) +(**************************************) + +Section RelProperties. +Variables (T : Type) (pT : PredType T). + +Lemma EqPredType_refl (r : pT) : EqPredType r r. Proof. by []. Qed. +Lemma SubPredType_refl (r : pT) : SubPredType r r. Proof. by []. Qed. + +Lemma EqPredType_sym (r1 r2 : pT) : EqPredType r1 r2 -> EqPredType r2 r1. +Proof. by move=>H1 x; split; move/H1. Qed. + +Lemma EqPredType_trans' (r1 r2 r3 : pT) : + EqPredType r1 r2 -> EqPredType r2 r3 -> EqPredType r1 r3. +Proof. by move=>H1 H2 x; split; [move/H1; move/H2 | move/H2; move/H1]. Qed. + +Lemma SubPredType_trans' (r1 r2 r3 : pT) : + SubPredType r1 r2 -> SubPredType r2 r3 -> SubPredType r1 r3. +Proof. by move=>H1 H2 x; move/H1; move/H2. Qed. + +Definition EqPredType_trans r2 r1 r3 := @EqPredType_trans' r1 r2 r3. +Definition SubPredType_trans r2 r1 r3 := @SubPredType_trans' r1 r2 r3. +End RelProperties. + +Hint Resolve EqPredType_refl SubPredType_refl : core. + +(* Declaration of relations *) + +(* Unfortunately, Coq setoids don't seem to understand implicit coercions *) +(* and canonical structures so we have to repeat relation declarations *) +(* for all instances. This is really annoying, but at least we don't have *) +(* to reprove the lemmas on refl, sym and trans. *) + +Add Parametric Relation T (pT : PredType T) : pT (@EqPredType _ pT) + reflexivity proved by (@EqPredType_refl _ _) + symmetry proved by (@EqPredType_sym _ _) + transitivity proved by (@EqPredType_trans' _ _) as EqPredType_rel. + +Add Parametric Relation T : (Simpl_Pred T) (@EqSimplPred _) + reflexivity proved by (@EqPredType_refl _ _) + symmetry proved by (@EqPredType_sym _ _) + transitivity proved by (@EqPredType_trans' _ _) as EqSimplPred_rel. + +Add Parametric Relation T : (Simpl_Pred T) (@SubSimplPred _) + reflexivity proved by (@SubPredType_refl _ _) + transitivity proved by (@SubPredType_trans' _ _) as SubSimplPred_rel. + +Add Parametric Relation T : (Mem_Pred T) (@EqMem T) + reflexivity proved by (@EqPredType_refl _ _) + symmetry proved by (@EqPredType_sym _ _) + transitivity proved by (@EqPredType_trans' _ _) as EqMem_rel. + +Add Parametric Relation T : (Mem_Pred T) (@SubMem _) + reflexivity proved by (@SubPredType_refl _ _) + transitivity proved by (@SubPredType_trans' _ _) as SubMem_rel. + +(* Declaring morphisms. *) + +Add Parametric Morphism T : (@Pred_of_Simpl T) with signature + @EqSimplPred _ ==> @EqPredType T (PredPredType T) as Pred_of_Simpl_morph. +Proof. by []. Qed. + +Add Parametric Morphism T (pT : PredType T) : (@EqPredType T pT) with signature + @EqPredType _ _ ==> @EqPredType _ _ ==> iff as EqPredType_morph. +Proof. by move=>r1 s1 H1 r2 s2 H2; rewrite H1 H2. Qed. + +Add Parametric Morphism T (pT : PredType T) : (@SubPredType T pT) with signature + @EqPredType _ _ ==> @EqPredType _ _ ==> iff as SubPred_morph. +Proof. by move=>r1 s1 H1 r2 s2 H2; split=>H x; move/H1; move/H; move/H2. Qed. + +Add Parametric Morphism T : (@InMem T) with signature + @eq _ ==> @EqMem _ ==> iff as InMem_morph. +Proof. by move=>x r s H; split; move/H. Qed. + +Add Parametric Morphism T (pT : PredType T) : (@Mem T pT) with signature + @EqPredType _ _ ==> @EqMem _ as Mem_morhp. +Proof. by move=>x y H p; rewrite /EqPredType -!toPredE in H *; rewrite H. Qed. + +Add Parametric Morphism T : (@PredU T) with signature + @EqPredType _ _ ==> @EqPredType _ _ ==> @EqSimplPred _ as predU_morph. +Proof. +move=>r1 s1 H1 r2 h2 H2 x; split; +by case; [move/H1 | move/H2]=>/=; auto. +Qed. + +Add Parametric Morphism T : (@PredI T) with signature + @EqPredType _ _ ==> @EqPredType _ _ ==> @EqPredType _ _ as predI_morph. +Proof. +move=>r1 s1 H1 r2 s2 H2 x; split; +by case; move/H1=>T1; move/H2=>T2. +Qed. + +Add Parametric Morphism T : (@PredC T) with signature + @EqPredType _ _ ==> @EqPredType _ _ as predC_morph. +Proof. by move=>r s H x; split=>H1; apply/H. Qed. + +Section RelLaws. +Variable (T : Type). +Implicit Type r : Pred T. + +Lemma orrI r : r +p r <~> r. +Proof. by move=>x; split; [case | left]. Qed. + +Lemma orrC r1 r2 : r1 +p r2 <~> r2 +p r1. +Proof. move=>x; split=>/=; tauto. Qed. + +Lemma orr0 r : r +p Pred0 <~> r. +Proof. by move=>x; split; [case | left]. Qed. + +Lemma or0r r : Pred0 +p r <~> r. +Proof. by rewrite orrC orr0. Qed. + +Lemma orrCA r1 r2 r3 : r1 +p r2 +p r3 <~> r2 +p r1 +p r3. +Proof. by move=>x; split=>/=; intuition. Qed. + +Lemma orrAC r1 r2 r3 : (r1 +p r2) +p r3 <~> (r1 +p r3) +p r2. +Proof. by move=>?; split=>/=; intuition. Qed. + +Lemma orrA r1 r2 r3 : (r1 +p r2) +p r3 <~> r1 +p r2 +p r3. +Proof. by rewrite (orrC r2) orrCA orrC. Qed. + +(* absorption *) +Lemma orrAb r1 r2 : r1 <~> r1 +p r2 <-> r2 ~> r1. +Proof. +split; first by move=>-> x /=; auto. +move=>H x /=; split; first by auto. +by case=>//; move/H. +Qed. + +Lemma sub_orl r1 r2 : r1 ~> r1 +p r2. Proof. by left. Qed. +Lemma sub_orr r1 r2 : r2 ~> r1 +p r2. Proof. by right. Qed. + +End RelLaws. + +Section SubMemLaws. +Variable T : Type. +Implicit Type p q : Pred T. + +Lemma subp_refl p : p <=p p. +Proof. by []. Qed. + +Lemma subp_asym p1 p2 : p1 <=p p2 -> p2 <=p p1 -> p1 =p p2. +Proof. by move=>H1 H2 x; split; [move/H1 | move/H2]. Qed. + +Lemma subp_trans p2 p1 p3 : p1 <=p p2 -> p2 <=p p3 -> p1 <=p p3. +Proof. by move=>H1 H2 x; move/H1; move/H2. Qed. + +Lemma subp_or p1 p2 q : p1 <=p q /\ p2 <=p q <-> p1 +p p2 <=p q. +Proof. +split=>[[H1] H2 x|H1]; first by case; [move/H1 | move/H2]. +by split=>x H2; apply: H1; [left | right]. +Qed. + +Lemma subp_and p1 p2 q : q <=p p1 /\ q <=p p2 <-> q <=p p1 *p p2. +Proof. +split=>[[H1] H2 x|] H; last by split=>x; case/H. +by split; [apply: H1 | apply: H2]. +Qed. + +Lemma subp_orl p1 p2 q : p1 <=p p2 -> p1 +p q <=p p2 +p q. +Proof. by move=>H x; case; [move/H; left|right]. Qed. + +Lemma subp_orr p1 p2 q : p1 <=p p2 -> q +p p1 <=p q +p p2. +Proof. by move=>H x; case; [left | move/H; right]. Qed. + +Lemma subp_andl p1 p2 q : p1 <=p p2 -> p1 *p q <=p p2 *p q. +Proof. by by move=>H x [H1 H2]; split; [apply: H|]. Qed. + +Lemma subp_andr p1 p2 q : p1 <=p p2 -> q *p p1 <=p q *p p2. +Proof. by move=>H x [H1 H2]; split; [|apply: H]. Qed. + +End SubMemLaws. + +Hint Resolve subp_refl : core. + +Section ListMembership. +Variable T : Type. + +Fixpoint Mem_Seq (s : seq T) := + if s is y::s' then (fun x => x = y \/ Mem_Seq s' x) else xPred0. + +Definition EqSeq_Class := seq T. +Identity Coercion seq_of_EqSeq : EqSeq_Class >-> seq. + +Coercion Pred_of_Eq_Seq (s : EqSeq_Class) : Pred_Class := [eta Mem_Seq s]. + +Canonical Structure seq_PredType := Eval hnf in @mkPredType T (seq T) Pred_of_Eq_Seq. +(* The line below makes Mem_Seq a canonical instance of topred. *) +Canonical Structure Mem_Seq_PredType := Eval hnf in mkPredType Mem_Seq. + +Lemma In_cons y s x : (x \In y :: s) <-> (x = y) \/ (x \In s). +Proof. by []. Qed. + +Lemma In_nil x : (x \In [::]) <-> False. +Proof. by []. Qed. + +Lemma Mem_Seq1 x y : (x \In [:: y]) <-> (x = y). +Proof. by rewrite In_cons orpF. Qed. + +Definition InE := (Mem_Seq1, In_cons, In_Simpl). + +Lemma Mem_cat x : forall s1 s2, (x \In s1 ++ s2) <-> x \In s1 \/ x \In s2. +Proof. +elim=>[|y s1 IH] s2 /=; first by split; [right | case]. +rewrite !InE /=. +split. +- case=>[->|/IH]; first by left; left. + by case; [left; right | right]. +case; first by case; [left | move=>H; right; apply/IH; left]. +by move=>H; right; apply/IH; right. +Qed. + +Lemma In_split x s : x \In s -> exists s1 s2, s = s1 ++ x :: s2. +Proof. +elim:s=>[|y s IH] //=; rewrite InE. +case=>[<-|]; first by exists [::], s. +by case/IH=>s1 [s2 ->]; exists (y :: s1), s2. +Qed. + +End ListMembership. + +Lemma Mem_map T T' (f : T -> T') x (s : seq T) : + x \In s -> f x \In (map f s). +Proof. +elim: s=>[|y s IH] //; rewrite InE /=. +by case=>[<-|/IH]; [left | right]. +Qed. + +Lemma Mem_map_inv T T' (f : T -> T') x (s : seq T) : + x \In (map f s) -> exists y, x = f y /\ y \In s. +Proof. +elim: s=>[|y s IH] //=; rewrite InE /=. +case; first by move=>->; exists y; split=>//; left. +by case/IH=>z [->]; exists z; split=>//; right. +Qed. + +Prenex Implicits Mem_map_inv. + +(* for equality types, membership predicates coincide *) +Lemma mem_seqP (A : eqType) x (s : seq A) : reflect (x \In s) (x \in s). +Proof. +elim: s=>[|y s IH]; first by constructor. +rewrite inE; case: eqP=>[<-|H /=]; first by constructor; left. +by apply: equivP IH _; rewrite InE; split; [right | case]. +Qed. + +(* Setoids for extensional equality of functions *) +Add Parametric Relation A B : (A -> B) (@eqfun _ _) + reflexivity proved by (@frefl B A) + symmetry proved by (@fsym B A) + transitivity proved by (@ftrans B A) as eqfun_morph. + +(***********************************) +(* Image of a collective predicate *) +(***********************************) + +Section Image. +Variables (A B : Type) (P : Pred A) (f : A -> B). + +Inductive image_spec b : Prop := Im_mem a of b = f a & a \In P. + +Definition Image' : Pred B := image_spec. + +End Image. + +(* swap to make the notation consider P before E; helps inference *) +Notation Image f P := (Image' P f). + +Notation "[ 'Image' E | i <- s ]" := (Image (fun i => E) s) + (at level 0, E at level 99, i ident, + format "[ '[hv' 'Image' E '/ ' | i <- s ] ']'") : rel_scope. + +Notation "[ 'Image' E | i <- s & C ]" := [Image E | i <- [PredI s & C]] + (at level 0, E at level 99, i ident, + format "[ '[hv' 'Image' E '/ ' | i <- s '/ ' & C ] ']'") : rel_scope. + +Notation "[ 'Image' E | i : T <- s ]" := (Image (fun i : T => E) s) + (at level 0, E at level 99, i ident, only parsing) : rel_scope. + +Notation "[ 'Image' E | i : T <- s & C ]" := + [Image E | i : T <- [PredI s & C]] + (at level 0, E at level 99, i ident, only parsing) : rel_scope. + +Lemma Image_mem A B (f : A -> B) (P : Pred A) x : x \In P -> f x \In Image f P. +Proof. by apply: Im_mem. Qed. + +Lemma Image_inj_sub A B (f : A -> B) (X1 X2 : Pred A) : + injective f -> Image f X1 <=p Image f X2 -> X1 <=p X2. +Proof. by move=>H E x /(Image_mem f) /E [y] /H ->. Qed. + +Lemma Image_inj_eqmem A B (f : A -> B) (X1 X2 : Pred A) : + injective f -> Image f X1 =p Image f X2 -> X1 =p X2. +Proof. by move=>H E; split; apply: Image_inj_sub H _ _; rewrite E. Qed. + +Lemma ImageU A B (f : A -> B) (X1 X2 : Pred A) : + Image f (PredU X1 X2) =p [PredU Image f X1 & Image f X2]. +Proof. +move=>x; split. +- by case=>y -> [H|H]; [left | right]; apply: Image_mem. +by case; case=>y -> H; apply: Image_mem; [left | right]. +Qed. + +Lemma ImageIm A B C (f1 : B -> C) (f2 : A -> B) (X : Pred A) : + Image f1 (Image f2 X) =p Image (f1 \o f2) X. +Proof. +move=>x; split; first by case=>_ -> [x' ->] H; exists x'. +by case=>a -> H; exists (f2 a)=>//; exists a. +Qed. + +Lemma ImageEq A B (f1 f2 : A -> B) (X : Pred A) : + f1 =1 f2 -> Image f1 X =p Image f2 X. +Proof. by move=>H x; split; case=>a ->; exists a. Qed. + + +(**********************************) +(**********************************) +(* Theory of relations *) +(**********************************) +(**********************************) + +(* Not a definition to avoid universe inconsistencies. *) + +Local Notation Rel A := (A -> A -> Prop). + +(**********************************) +(* Special cases of relations *) +(**********************************) + +Section RelationProperties. + +Variables (A : Type) (R : Rel A). + +Definition Total := forall x y, R x y \/ R y x. +Definition Transitive := forall y x z, R x y -> R y z -> R x z. + +Definition Symmetric := forall x y, R x y <-> R y x. +Definition pre_Symmetric := forall x y, R x y -> R y x. +Definition Antisymmetric := forall x y, R x y -> R y x -> x = y. + +Lemma sym_iff_pre_sym : pre_Symmetric <-> Symmetric. +Proof. by split=> S x y; [split; apply: S | rewrite S]. Qed. + +(* Do NOT use Reflexive for actual lemmas, because auto does not unfold definitions, + so Hint Resolve lemma_refl won't work *) +Definition Reflexive := forall x, R x x. +Definition Irreflexive := forall x, R x x -> False. + +Definition left_Transitive := forall x y, R x y -> forall z, R x z <-> R y z. +Definition right_Transitive := forall x y, R x y -> forall z, R z x <-> R z y. + +(** Partial equivalence relation *) +Section PER. + +Hypotheses (symR : Symmetric) (trR : Transitive). + +Lemma sym_left_Transitive : left_Transitive. +Proof. by move=> x y Rxy z; split; apply: trR; rewrite // symR. Qed. + +Lemma sym_right_Transitive : right_Transitive. +Proof. by move=> x y Rxy z; rewrite !(symR z); apply: sym_left_Transitive. Qed. + +End PER. + +(* We define the equivalence property with prenex quantification so that it *) +(* can be localized using the {in ..., ..} form defined below. *) + +Definition Equivalence_rel := forall x y z, R z z /\ (R x y -> R x z <-> R y z). + +Lemma Equivalence_relP : Equivalence_rel <-> Reflexive /\ left_Transitive. +Proof. +split=> [eqiR | [Rxx ltrR] x y z]; last by split=>// Rxy; apply: ltrR. +by split=> [x | x y Rxy z]; [case: (eqiR x x x)| case: (eqiR x y z)=> _ /(_ Rxy)]. +Qed. + +Lemma Equivalence_relS : Equivalence_rel <-> [/\ Reflexive, Symmetric & Transitive]. +Proof. +split; last first. +- case=>r s t; split; first by apply: r. + by move=>Rxy; split; apply: t=>//; apply/s. +case/Equivalence_relP=>r t; split=>//; last first. +- by move=>x y z Ryx Rxz; rewrite (t y x Ryx). +move=>x y; split. +- by move=>Rxy; rewrite -(t x y Rxy); apply: r. +by move=>Ryx; rewrite -(t y x Ryx); apply: r. +Qed. + +(** _Functional_ (a.k.a. deterministic) relation *) +Definition functional := forall x y1 y2, R x y1 -> R x y2 -> y1 = y2. +(* a relation preserves the resource invariant *) +Definition preserved (P : A -> Prop) R := forall x y, R x y -> P x -> P y. +Definition bpreserved (P : A -> Prop) R := forall x y, R x y -> P y -> P x. + +End RelationProperties. + +Arguments sym_iff_pre_sym {A R}. + + +(* Lifting equivalence relation to option type *) + +Section OptionRel. +Variables (A : Type) (R : Rel A). + +Definition optionR (a1 a2 : option A) := + match a1, a2 with + Some e1, Some e2 => R e1 e2 + | None, None => True + | _, _ => False + end. + +Lemma Reflexive_optionR : Reflexive R -> Reflexive optionR. +Proof. by move=>r; case. Qed. + +Lemma Symmetric_optionR : Symmetric R -> Symmetric optionR. +Proof. by move=>s; case=>[x|]; case=>[y|] //=. Qed. + +Lemma Transitive_optionR : Transitive R -> Transitive optionR. +Proof. by move=>t; case=>[x|]; case=>[y|]; case=>[z|] //=; apply: t. Qed. + +Lemma Equivalence_optionR : Equivalence_rel R -> Equivalence_rel optionR. +Proof. +case/Equivalence_relS=>r s t; apply/Equivalence_relS; split. +- by apply: Reflexive_optionR. +- by apply: Symmetric_optionR. +by apply: Transitive_optionR. +Qed. + +End OptionRel. + +(* Lifting equivalence relation to sum type *) +Section SumRel. +Variables (A1 A2 : Type) (R1 : Rel A1) (R2 : Rel A2). + +Definition sumR (x y : A1 + A2) := + match x, y with + inl x1, inl y1 => R1 x1 y1 + | inr x2, inr y2 => R2 x2 y2 + | _, _ => False + end. + +Lemma Reflexive_sumR : Reflexive R1 -> Reflexive R2 -> Reflexive sumR. +Proof. by move=>r1 r2; case. Qed. + +Lemma Symmetric_sumR : Symmetric R1 -> Symmetric R2 -> Symmetric sumR. +Proof. by move=>s1 s2; case=>x; case=>y //=. Qed. + +Lemma Transitive_sumR : Transitive R1 -> Transitive R2 -> Transitive sumR. +Proof. by move=>t1 t2; case=>x; case=>y; case=>z //; [apply:t1 | apply:t2]. Qed. + +Lemma Equivalence_sumR : + Equivalence_rel R1 -> Equivalence_rel R2 -> Equivalence_rel sumR. +Proof. +case/Equivalence_relS=>r1 s1 t1 /Equivalence_relS [r2 s2 t2]. +apply/Equivalence_relS; split. +- by apply: Reflexive_sumR. +- by apply: Symmetric_sumR. +by apply: Transitive_sumR. +Qed. + +End SumRel. + + +(****************) +(* Transitivity *) +(****************) + +Section Transitivity. +Variables (A : Type) (R : Rel A). + +Lemma trans_imp (F : A -> Prop) : Transitive (fun x y => F x -> F y). +Proof. by move=>x y z H1 H2 /H1. Qed. + +Lemma trans_eq B (f : A -> B) : Transitive (fun x y => f x = f y). +Proof. by move=>x y z ->. Qed. + +Lemma rev_Trans : Transitive R -> Transitive (fun x y => R y x). +Proof. by move=> trR x y z Ryx Rzy; apply: trR Rzy Ryx. Qed. +End Transitivity. + +Hint Resolve trans_imp trans_eq : core. + +(**********************************************) +(* Reflexive-Transitive closure of a relation *) +(**********************************************) + +Fixpoint iter' T (g : T -> T -> Prop) n s1 s2 := + if n is n'.+1 then exists s, g s1 s /\ iter' g n' s s2 else s1 = s2. +Definition iter T (g : T -> T -> Prop) s1 s2 := exists n, iter' g n s1 s2. +(* curry the iteration *) +Definition iterc A T (g : A -> T -> A -> T -> Prop) a1 s1 a2 s2 := + iter (fun '(a1, a2) '(b1, b2) => g a1 a2 b1 b2) (a1, s1) (a2, s2). + +Section IteratedRels. +Variable T : Type. +Implicit Type g : T -> T -> Prop. + +Lemma iter_refl g s : iter g s s. +Proof. by exists 0. Qed. + +Lemma iter_trans g : Transitive (iter g). +Proof. +move=> s2 s1 s3; case=>n; elim: n s1 =>[|n IH] s1 /=; first by move=>->. +by case=>s [H1 H2] /(IH _ H2) [m]; exists m.+1, s. +Qed. + +Lemma iterS g n s1 s2 : iter' g n.+1 s1 s2 <-> exists s, iter' g n s1 s /\ g s s2. +Proof. +elim: n s1=>[|n IH] s1. +- by split; [case=>s [H <-]; exists s1 | case=>s [<- H]; exists s2]. +split; first by case=>s [H1] /IH [s'][H G]; exists s'; split=>//; exists s. +by case=>s [[s'][H1 H G]]; exists s'; split=>//; apply/IH; exists s. +Qed. + +Lemma iter'_sub g g' n s1 s2 : + (forall s1 s2, g s1 s2 -> g' s1 s2) -> + iter' g n s1 s2 -> iter' g' n s1 s2. +Proof. by move=>H; elim: n s1=>[|n IH] s1 //= [s][/H G] /IH; exists s. Qed. + +Lemma iter_sub g g' s1 s2 : + (forall s1 s2, g s1 s2 -> g' s1 s2) -> iter g s1 s2 -> iter g' s1 s2. +Proof. by move=>H [n]; exists n; apply: iter'_sub H _. Qed. + +Lemma iter1 g s1 s2 : g s1 s2 -> iter g s1 s2. +Proof. by exists 1, s2. Qed. + +Lemma iter'_add g n1 n2 s1 s2 s3 : + iter' g n1 s1 s2 -> iter' g n2 s2 s3 -> + iter' g (n1 + n2) s1 s3. +Proof. +elim: n1 s1 s2=>[|n1 IH] s1 s2 /=; first by rewrite add0n=>->. +by case=>s [H1 H2] /(IH _ _ H2); exists s. +Qed. + +Lemma iter'_split g n1 n2 s1 s2 : + iter' g (n1 + n2) s1 s2 -> + exists s, iter' g n1 s1 s /\ iter' g n2 s s2. +Proof. +elim: n1 s1 s2=>[|n1 IH] s1 s2 /=; first by rewrite add0n; exists s1. +by case=>s [H1] {}/IH [s'][H2 H3]; exists s'; split=>//; exists s. +Qed. + +End IteratedRels. + +Lemma iter2 A T (g : A -> T -> A -> T -> Prop) x1 s1 x2 s2 : + g x1 s1 x2 s2 -> iterc g x1 s1 x2 s2. +Proof. by apply: iter1. Qed. + +Lemma iter_pres' A (g : Rel A) n C : preserved C g -> preserved C (iter' g n). +Proof. +move=>coh; elim: n=>[|n IH] x y /=; first by move=><-. +by case=>z [/coh H1] H2 /H1; apply: IH. +Qed. + +Lemma iter_pres A (g : Rel A) C : preserved C g -> preserved C (iter g). +Proof. by move=>pres x y [n] /(iter_pres' pres). Qed. + +Prenex Implicits iter1 iter2. +Arguments iter_refl {T g s}. +Hint Resolve iter_refl : core. + +(*****************************************************) +(* Induction lemmas for Reflexive Transitive closure *) +(*****************************************************) + +Section ReflexiveTransitiveClosureInductions. +Variables (A : Type) (R : Rel A). +Implicit Types (C P Q : A -> Prop) (F : Rel A). + +Lemma iterf_ind C F : + (forall x, C x -> F x x) -> + Transitive F -> + (forall x y, R x y -> C x -> C y /\ F x y) -> + forall x y, iter R x y -> C x -> F x y. +Proof. +move=>X Y H x y [n]; elim: n x=>[|n IH] x; first by move=>->; apply: X. +by case=>z [O] {}/IH Z /(H _ _ O) [/Z] H1 H2; apply: Y H2 H1. +Qed. + +Lemma iterb_ind C F : + (forall x, C x -> F x x) -> + Transitive F -> + (forall x y, R x y -> C y -> C x /\ F x y) -> + forall x y, iter R x y -> C y -> F x y. +Proof. +move=>X Y H x y [n]; elim: n x y=>[|n IH] x y; first by move=>->; apply: X. +by case/iterS=>z[]{}/IH Z O /(H _ _ O) [/Z]; apply: Y. +Qed. + +Lemma iter_ind C : + (forall x y, R x y -> C x -> C y) -> + forall x y, iter R x y -> C x -> C y. +Proof. +move=>H x y /(@iterf_ind (fun => True) (fun x y => C x -> C y)). +by apply=>// t1 t2 /H X _. +Qed. + +(* induction under a stable condition; this is always what we have *) +Lemma iters_ind C F : + (forall x, C x -> F x x) -> Transitive F -> + (forall x y, iter R x y -> C x -> C y) -> + (forall x y, R x y -> C x -> F x y) -> + forall x y, iter R x y -> C x -> F x y. +Proof. +move=>H1 H2 H3 H x y; apply: iterf_ind x y=>// x y. +by move=>H4 H5; split; [apply: H3 (iter1 H4) H5 | apply: H H5]. +Qed. + +(* can relax the transitivity condition *) +Lemma iters_ind' C F : + (forall x, C x -> F x x) -> + (forall x y z, C x -> C y -> C z -> F x y -> F y z -> F x z) -> + (forall x y, iter R x y -> C x -> C y) -> + (forall x y, R x y -> C x -> F x y) -> + forall x y, iter R x y -> C x -> F x y. +Proof. +move=>H1 H2 H3 H4 x y [n]; elim: n x=>[|n IH] x. +- by move=>->; apply: H1. +case=>z [O] H Cx; move: (H3 _ _ (iter1 O) Cx)=>Cz. +by apply: H2 (IH _ H Cz)=>//; [apply: H3 Cz; exists n | apply: H4]. +Qed. + +Lemma pres_iterf_ind P Q F : + preserved P R -> + (forall x, P x -> Q x -> F x x) -> + Transitive F -> + (forall x y, R x y -> P x -> Q x -> Q y /\ F x y) -> + forall x y, iter R x y -> P x -> Q x -> F x y. +Proof. +move=>pres H1 H2 H3 x y O X1 X2; move: x y {X1 X2} O (conj X1 X2). +apply: iterf_ind=>//; first by move=>s []; apply: H1. +by move=>x y O [X Y]; case: (H3 _ _ O X Y) (pres _ _ O X). +Qed. + +Lemma pres_iterb_ind P Q F : + bpreserved P R -> + (forall x, P x -> Q x -> F x x) -> + Transitive F -> + (forall x y, R x y -> P y -> Q y -> Q x /\ F x y) -> + forall x y, iter R x y -> P y -> Q y -> F x y. +Proof. +move=>pres H1 H2 H3 x y O X1 X2; move: x y {X1 X2} O (conj X1 X2). +apply: iterb_ind=>//; first by move=>s []; apply: H1. +by move=>x y O [X Y]; case: (H3 _ _ O X Y) (pres _ _ O X). +Qed. + +Lemma pres_iters_ind P Q F : + preserved P R -> + (forall x, P x -> Q x -> F x x) -> + Transitive F -> + (forall x y, iter R x y -> P x -> Q x -> Q y) -> + (forall x y, R x y -> P x -> Q x -> F x y) -> + forall x y, iter R x y -> P x -> Q x -> F x y. +Proof. +move=>pres H1 H2 H3 H4 x y O X1 X2; move: x y {X1 X2} O (conj X1 X2). +apply: iterf_ind=>//; first by move=>s []; apply: H1. +by move=>x y O [X Y]; move: (H3 _ _ (iter1 O) X Y) (H4 _ _ O X Y) (pres _ _ O X). +Qed. + +End ReflexiveTransitiveClosureInductions. + +Section ReflexiveTransitiveClosureInductionsCurry. +Variables (A T : Type) (R : A -> T -> A -> T -> Prop) . +Variables (C P Q : A -> T -> Prop) (F : A -> T -> A -> T -> Prop). + +Notation Rc := (fun '(a, x) '(b, y) => R a x b y). +Notation Cc := (fun '(a, x) => C a x). +Notation Fc := (fun '(a, x) '(b, y)=> F a x b y). +Notation Pc := (fun '(a, x) => P a x). +Notation Qc := (fun '(a, x) => Q a x). + +Lemma itercf_ind : + (forall a x, C a x -> F a x a x) -> + Transitive Fc -> + (forall a x b y, R a x b y -> C a x -> C b y /\ F a x b y) -> + forall a x b y, iterc R a x b y -> C a x -> F a x b y. +Proof. +move=>H1 H2 H3 a x b y; apply: (@iterf_ind _ Rc Cc Fc)=>//. +- by case=>a' x'; apply: H1. +by case=>a' x' [b' y']; apply: H3. +Qed. + +Lemma itercb_ind : + (forall a x, C a x -> F a x a x) -> + Transitive Fc -> + (forall a x b y, R a x b y -> C b y -> C a x /\ F a x b y) -> + forall a x b y, iterc R a x b y -> C b y -> F a x b y. +Proof. +move=>H1 H2 H3 a x b y; apply: (@iterb_ind _ Rc Cc Fc)=>//. +- by case=>a' x'; apply: H1. +by case=>a' x' [b' y']; apply: H3. +Qed. + +Lemma iterc_ind : + (forall a x b y, R a x b y -> C a x -> C b y) -> + forall a x b y, iterc R a x b y -> C a x -> C b y. +Proof. +move=>H a x b y; apply: (@iter_ind _ Rc Cc)=>//. +by case=>a' x' [b' y']; apply: H. +Qed. + +Lemma itercs_ind : + (forall a x, C a x -> F a x a x) -> Transitive Fc -> + (forall a x b y, iterc R a x b y -> C a x -> C b y) -> + (forall a x b y, R a x b y -> C a x -> F a x b y) -> + forall a x b y, iterc R a x b y -> C a x -> F a x b y. +Proof. +move=>H1 H2 H3 H4 a x b y; apply: (@iters_ind _ Rc Cc Fc)=>//. +- by case=>a' x'; apply: H1. +- by case=>a' x' [b' y']; apply: H3. +by case=>a' x' [b' y']; apply: H4. +Qed. + +Lemma pres_itercf_ind : + preserved Pc Rc -> + (forall a x, P a x -> Q a x -> F a x a x) -> + Transitive Fc -> + (forall a x b y, R a x b y -> P a x -> Q a x -> Q b y /\ F a x b y) -> + forall a x b y, iterc R a x b y -> P a x -> Q a x -> F a x b y. +Proof. +move=>H1 H2 H3 H4 a x b y; apply: (@pres_iterf_ind _ Rc Pc Qc Fc)=>//. +- by case=>a' x' [b' y']; apply: (H1 (a', x') (b', y')). +- by case=>a' x'; apply: H2. +by case=>a' x' [b' y']; apply: H4. +Qed. + +Lemma pres_itercs_ind : + preserved Pc Rc -> + (forall a x, P a x -> Q a x -> F a x a x) -> + Transitive Fc -> + (forall a x b y, iterc R a x b y -> P a x -> Q a x -> Q b y) -> + (forall a x b y, R a x b y -> P a x -> Q a x -> F a x b y) -> + forall a x b y, iterc R a x b y -> P a x -> Q a x -> F a x b y. +Proof. +move=>H1 H2 H3 H4 H5 a x b y; apply: (@pres_iters_ind _ Rc Pc Qc Fc)=>//. +- by case=>a' x' [b' y']; apply: (H1 (a', x') (b', y')). +- by case=>a' x'; apply: H2. +- by case=>a' x' [b' y']; apply: H4. +by case=>a' x' [b' y']; apply: H5. +Qed. + +End ReflexiveTransitiveClosureInductionsCurry. + + +Section SomeBasicConstructions. +Variable (A : Type). +Implicit Types (P : A -> Prop) (R : Rel A). + +(** _Empty_ relation *) +Section EmptyRelation. +Definition emp_rel : Rel A := fun x y => False. + +Lemma emp_rel_func : functional emp_rel. +Proof. by []. Qed. +Lemma emp_rel_pres P : preserved P emp_rel. +Proof. by []. Qed. +End EmptyRelation. + +(** _Full_ (a.k.a _unversal_) relation *) +Section FullRelation. +Definition full_rel : Rel A := fun x y => True. + +Lemma full_rel_refl : forall x, full_rel x x. +Proof. by []. Qed. +Lemma full_rel_sym : Symmetric full_rel. +Proof. by []. Qed. +Lemma full_rel_trans : Transitive full_rel. +Proof. by []. Qed. +Lemma full_rel_tot : Total full_rel. +Proof. by move=> ??; left. Qed. +End FullRelation. + +(**_Identity_ relation *) +Section IdentityRelation. +Definition id_rel : Rel A := fun x y => y = x. + +Lemma id_rel_refl : forall x, id_rel x x. +Proof. by []. Qed. +Lemma id_rel_sym : Symmetric id_rel. +Proof. by []. Qed. +Lemma id_rel_trans : Transitive id_rel. +Proof. by move=> x y z ->->. Qed. +Lemma id_rel_func : functional id_rel. +Proof. by move=> x y1 y2 ->->. Qed. +Lemma id_rel_pres P : preserved P id_rel. +Proof. by move=>x y ->. Qed. +End IdentityRelation. + +(** Imposing "precondition" on a relation *) +Section PreConditionalRelation. +Definition pre_rel P R := fun x y => P x /\ R x y. + +Lemma pre_rel_func R P : functional R -> functional (pre_rel P R). +Proof. by move=> funcR x y1 y2 [_ Rxy1] [_ /(funcR _ _ _ Rxy1)]. Qed. +End PreConditionalRelation. + + +(** Imposing "postcondition" on a relation *) +Section PostConditionalRelation. +Definition post_rel R Q := fun x y => R x y /\ Q y. + +Lemma reinv_rel_func R Q : functional R -> functional (post_rel R Q). +Proof. by move=> funcR x y1 y2 [Rxy1 _] [/(funcR _ _ _ Rxy1)]. Qed. +End PostConditionalRelation. + + +Section ProductRelation. +Variables (B : Type) (R : Rel A) (S : Rel B). + +Definition prod_rel : Rel (A * B) := + fun '(a1, b1) '(a2, b2) => R a1 a2 /\ S b1 b2. + +Lemma prod_rel_refl : (forall a, R a a) -> (forall b, S b b) -> forall p, prod_rel p p. +Proof. by move=> pR pS [a b] /=. Qed. + +Lemma prod_rel_sym : Symmetric R -> Symmetric S -> Symmetric prod_rel. +Proof. by move=> pR pS [a1 b1] [a2 b2] /=; rewrite pR pS. Qed. + +Lemma prod_rel_trans : Transitive R -> Transitive S -> Transitive prod_rel. +Proof. +move=> pR pS [a2 b2] [a1 b1] [a3 b3] [??] [??] /=. +by split; [apply: (pR a2) | apply: (pS b2)]. +Qed. + +Lemma prod_rel_asym : Antisymmetric R -> Antisymmetric S -> Antisymmetric prod_rel. +Proof. +move=> pR pS [a1 b1] [a2 b2] [R12 S12] [R21 S21]. +by rewrite (pR _ _ R12 R21) (pS _ _ S12 S21). +Qed. + +Lemma prod_rel_pre_sym : pre_Symmetric R -> pre_Symmetric S -> pre_Symmetric prod_rel. +Proof. by rewrite !sym_iff_pre_sym; apply: prod_rel_sym. Qed. + +Lemma prod_rel_irrefl : Irreflexive R -> Irreflexive S -> Irreflexive prod_rel. +Proof. by move=> pR _ [a b] /= [] /pR. Qed. + +Lemma prod_rel_ltrans : left_Transitive R -> left_Transitive S -> left_Transitive prod_rel. +Proof. +move=> pR pS [a1 b1] [a2 b2] [R12 S12] [a3 b3] /=. +by rewrite (pR _ _ R12 a3) (pS _ _ S12 b3). +Qed. + +Lemma prod_rel_rtrans : right_Transitive R -> right_Transitive S -> right_Transitive prod_rel. +Proof. +move=> pR pS [a1 b1] [a2 b2] [R12 S12] [a3 b3] /=. +by rewrite (pR _ _ R12 a3) (pS _ _ S12 b3). +Qed. + +Lemma prod_rel_func : functional R -> functional S -> functional prod_rel. +Proof. +move=> pR pS [a2 b2] [a1 b1] [a3 b3] [R21 S21] [R23 S23] /=. +by rewrite (pR _ _ _ R21 R23) (pS _ _ _ S21 S23). +Qed. + +End ProductRelation. + + +Section UnionRelation. +Variables (R S : Rel A). + +Definition Union_rel : Rel A := fun x y => R x y \/ S x y. + +Definition fcompatible := forall x y1 y2, R x y1 -> S x y2 -> y1 = y2. + +Lemma union_rel_func : functional R -> functional S -> + fcompatible -> functional Union_rel. +Proof. +move=> fR fS fC x y1 y2; case=> [hR1 | hS1]; case=> [hR2 | hS2]. +- by apply: fR hR1 hR2. +- apply: fC hR1 hS2. +- apply/esym; apply: fC hR2 hS1. +by apply: fS hS1 hS2. +Qed. + +End UnionRelation. + +End SomeBasicConstructions. + +Arguments id_rel [A] _ _ /. +Arguments pre_rel [A] P R _ _ /. +Arguments post_rel [A] R Q _ _ /. + + +(*************************) +(* Property localization *) +(*************************) + +Local Notation "{ 'All1' P }" := (forall x, P x : Prop) (at level 0). +Local Notation "{ 'All2' P }" := (forall x y, P x y : Prop) (at level 0). +Local Notation "{ 'All3' P }" := (forall x y z, P x y z: Prop) (at level 0). +Local Notation ph := (phantom _). + +Section LocalProperties. + +Variables T1 T2 T3 : Type. + +Variables (d1 : T1 -> Prop) (d2 : T2 -> Prop) (d3 : T3 -> Prop). +Local Notation ph := (phantom Prop). + +Definition Prop_in1 P & ph {All1 P} := + forall x, d1 x -> P x. + +Definition Prop_in11 P & ph {All2 P} := + forall x y, d1 x -> d2 y -> P x y. + +Definition Prop_in2 P & ph {All2 P} := + forall x y, d1 x -> d1 y -> P x y. + +Definition Prop_in111 P & ph {All3 P} := + forall x y z, d1 x -> d2 y -> d3 z -> P x y z. + +Definition Prop_in12 P & ph {All3 P} := + forall x y z, d1 x -> d2 y -> d2 z -> P x y z. + +Definition Prop_in21 P & ph {All3 P} := + forall x y z, d1 x -> d1 y -> d2 z -> P x y z. + +Definition Prop_in3 P & ph {All3 P} := + forall x y z, d1 x -> d1 y -> d1 z -> P x y z. + +End LocalProperties. + +Definition inPhantom := Phantom Prop. + +Notation "{ 'In' d , P }" := + (Prop_in1 d (inPhantom P)) + (at level 0, format "{ 'In' d , P }") : type_scope. + +Notation "{ 'In' d1 & d2 , P }" := + (Prop_in11 d1 d2 (inPhantom P)) + (at level 0, format "{ 'In' d1 & d2 , P }") : type_scope. + +Notation "{ 'In' d & , P }" := + (Prop_in2 d (inPhantom P)) + (at level 0, format "{ 'In' d & , P }") : type_scope. + +Notation "{ 'In' d1 & d2 & d3 , P }" := + (Prop_in111 d1 d2 d3 (inPhantom P)) + (at level 0, format "{ 'In' d1 & d2 & d3 , P }") : type_scope. + +Notation "{ 'In' d1 & & d3 , P }" := + (Prop_in21 d1 d3 (inPhantom P)) + (at level 0, format "{ 'In' d1 & & d3 , P }") : type_scope. + +Notation "{ 'In' d1 & d2 & , P }" := + (Prop_in12 d1 d2 (inPhantom P)) + (at level 0, format "{ 'In' d1 & d2 & , P }") : type_scope. + +Notation "{ 'In' d & & , P }" := + (Prop_in3 d (inPhantom P)) + (at level 0, format "{ 'In' d & & , P }") : type_scope. + + +(* Weakening and monotonicity lemmas for localized predicates. *) +(* Note that using these lemmas in backward reasoning will force expansion of *) +(* the predicate definition, as Coq needs to expose the quantifier to apply *) +(* these lemmas. We define a few specialized variants to avoid this for some *) +(* of the ssrfun predicates. *) + +Section LocalGlobal. + +Variables T1 T2 T3 : Type. +Variables (D1 : T1 -> Prop) (D2 : T2 -> Prop) (D3 : T3 -> Prop). +Variables (P1 : T1 -> Prop) (P2 : T1 -> T2 -> Prop). +Variable P3 : T1 -> T2 -> T3 -> Prop. +Variables (d1 d1' : T1 -> Prop). + +Local Notation "{ 'All1' P }" := (forall x, P x : Prop) (at level 0). +Local Notation "{ 'All2' P }" := (forall x y, P x y : Prop) (at level 0). +Local Notation "{ 'All3' P }" := (forall x y z, P x y z: Prop) (at level 0). +Local Notation ph := (phantom _). + +Lemma In1W : {All1 P1} -> {In D1, {All1 P1}}. +Proof. by move=> ? ?. Qed. +Lemma In2W : {All2 P2} -> {In D1 & D2, {All2 P2}}. +Proof. by move=> ? ?. Qed. +Lemma In3W : {All3 P3} -> {In D1 & D2 & D3, {All3 P3}}. +Proof. by move=> ? ?. Qed. + +End LocalGlobal. diff --git a/core/prelude.v b/core/prelude.v new file mode 100644 index 0000000..4fe1e98 --- /dev/null +++ b/core/prelude.v @@ -0,0 +1,578 @@ +(* +Copyright 2010 IMDEA Software Institute +Licensed under the Apache License, Version 2.0 (the "License"); +you may not use this file except in compliance with the License. +You may obtain a copy of the License at + http://www.apache.org/licenses/LICENSE-2.0 +Unless required by applicable law or agreed to in writing, software +distributed under the License is distributed on an "AS IS" BASIS, +WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. +See the License for the specific language governing permissions and +limitations under the License. +*) + +(******************************************************************************) +(* This file is Prelude -- often used notation definitions and lemmas that *) +(* are not included in the other libraries. *) +(******************************************************************************) + +From Coq Require Import ssreflect ssrbool ssrfun Eqdep. +From mathcomp Require Import ssrnat seq eqtype. +From fcsl Require Import axioms. +From fcsl Require Import options. + +(***********) +(* Prelude *) +(***********) + +(* turn off the automation of Program *) +Obligation Tactic := auto. + +(* often used notation definitions and lemmas that are *) +(* not included in the other libraries *) + +(* export inj_pair without exporting the whole Eqdep library *) +Definition inj_pair2 := @inj_pair2. +Arguments inj_pair2 [U P p x y] _. +Prenex Implicits inj_pair2. + +(* Because of a bug in inversion and injection tactics *) +(* we occasionally have to destruct pair by hand, else we *) +(* lose the second equation. *) +Lemma inj_pair A B (a1 a2 : A) (b1 b2 : B) : + (a1, b1) = (a2, b2) -> (a1 = a2) * (b1 = b2). +Proof. by case. Qed. +Arguments inj_pair [A B a1 a2 b1 b2]. +Prenex Implicits inj_pair. + +(* eta laws for pairs and units *) +Notation prod_eta := surjective_pairing. + +(* eta law often used with injection *) +Lemma prod_inj A B (x y : A * B) : x = y <-> (x.1, x.2) = (y.1, y.2). +Proof. by case: x y=>x1 x2 []. Qed. + +Lemma idfunE (U : Type) (x : U) : idfun x = x. +Proof. by []. Qed. + +(* idfun is a left and right unit for composition *) +Lemma idfun0E (U V : Type) (f : U -> V): + (idfun \o f = f) * (f \o idfun = f). +Proof. by []. Qed. + +(* Triples *) +Section TripleLemmas. +Variables (A B C : Type). +Implicit Types (a : A) (b : B) (c : C). + +Lemma t1P a1 a2 b1 b2 c1 c2 : (a1, b1, c1) = (a2, b2, c2) -> a1 = a2. +Proof. by case. Qed. + +Lemma t2P a1 a2 b1 b2 c1 c2 : (a1, b1, c1) = (a2, b2, c2) -> b1 = b2. +Proof. by case. Qed. + +Lemma t3P a1 a2 b1 b2 c1 c2 : (a1, b1, c1) = (a2, b2, c2) -> c1 = c2. +Proof. by case. Qed. + +Lemma t12P a1 a2 b1 b2 c1 c2 : (a1, b1, c1) = (a2, b2, c2) -> (a1 = a2) * (b1 = b2). +Proof. by case. Qed. + +Lemma t13P a1 a2 b1 b2 c1 c2 : (a1, b1, c1) = (a2, b2, c2) -> (a1 = a2) * (c1 = c2). +Proof. by case. Qed. + +Lemma t23P a1 a2 b1 b2 c1 c2 : (a1, b1, c1) = (a2, b2, c2) -> (b1 = b2) * (c1 = c2). +Proof. by case. Qed. + +End TripleLemmas. +Prenex Implicits t1P t2P t3P t12P t13P t23P. + +(***************************) +(* operations on functions *) +(***************************) + +Lemma eta A (B : A -> Type) (f : forall x, B x) : f = [eta f]. +Proof. by []. Qed. + +Lemma ext A (B : A -> Type) (f1 f2 : forall x, B x) : + f1 = f2 -> forall x, f1 x = f2 x. +Proof. by move=>->. Qed. + +Lemma compA A B C D (h : A -> B) (g : B -> C) (f : C -> D) : + (f \o g) \o h = f \o (g \o h). +Proof. by []. Qed. + +Definition fprod A1 A2 B1 B2 (f1 : A1 -> B1) (f2 : A2 -> B2) := + fun (x : A1 * A2) => (f1 x.1, f2 x.2). + +Notation "f1 \* f2" := (fprod f1 f2) (at level 42, left associativity) : fun_scope. + +(* product morphism, a.k.a. fork or fanout or fsplice *) +Definition pmorphism A B1 B2 (f1 : A -> B1) (f2 : A -> B2) := + fun x : A => (f1 x, f2 x). +Arguments pmorphism {A B1 B2} f1 f2 x /. +Notation "f1 \** f2 " := (pmorphism f1 f2) + (at level 50, left associativity, format "f1 \** '/ ' f2") : fun_scope. + +(* product with functions *) +Lemma prod_feta (A B : Type) : @idfun (A * B) = fst \** snd. +Proof. by apply: fext=>x; rewrite /pmorphism -prod_eta. Qed. + +Reserved Notation "[ ** f1 & f2 ]" (at level 0). +Reserved Notation "[ ** f1 , f2 & f3 ]" (at level 0, format + "'[hv' [ ** '[' f1 , '/' f2 ']' '/ ' & f3 ] ']'"). +Reserved Notation "[ ** f1 , f2 , f3 & f4 ]" (at level 0, format + "'[hv' [ ** '[' f1 , '/' f2 , '/' f3 ']' '/ ' & f4 ] ']'"). +Reserved Notation "[ ** f1 , f2 , f3 , f4 & f5 ]" (at level 0, format + "'[hv' [ ** '[' f1 , '/' f2 , '/' f3 , '/' f4 ']' '/ ' & f5 ] ']'"). + +Notation "[ ** f1 & f2 ]" := (f1 \** f2) (only parsing) : fun_scope. +Notation "[ ** f1 , f2 & f3 ]" := ((f1 \** f2) \** f3) : fun_scope. +Notation "[ ** f1 , f2 , f3 & f4 ]" := (((f1 \** f2) \** f3) \** f4) : fun_scope. +Notation "[ ** f1 , f2 , f3 , f4 & f5 ]" := ((((f1 \** f2) \** f3) \** f4) \** f5) : fun_scope. + +(* composing relation and function *) + +Definition relfuncomp A B (D : rel A) (f : B -> A) : rel B := + fun x y => D (f x) (f y). + +Notation "D \\o f" := (@relfuncomp _ _ D f) (at level 42, left associativity). + +(************************) +(* extension to ssrbool *) +(************************) + +Reserved Notation "[ /\ P1 , P2 , P3 , P4 , P5 & P6 ]" (at level 0, format + "'[hv' [ /\ '[' P1 , '/' P2 , '/' P3 , '/' P4 , '/' P5 ']' '/ ' & P6 ] ']'"). + +Reserved Notation "[ /\ P1 , P2 , P3 , P4 , P5 , P6 & P7 ]" (at level 0, format + "'[hv' [ /\ '[' P1 , '/' P2 , '/' P3 , '/' P4 , '/' P5 , '/' P6 ']' '/ ' & P7 ] ']'"). + +Reserved Notation "[ /\ P1 , P2 , P3 , P4 , P5 , P6 & P7 ]" (at level 0, format + "'[hv' [ /\ '[' P1 , '/' P2 , '/' P3 , '/' P4 , '/' P5 , '/' P6 ']' '/ ' & P7 ] ']'"). + +Reserved Notation "[ /\ P1 , P2 , P3 , P4 , P5 , P6 , P7 & P8 ]" (at level 0, format + "'[hv' [ /\ '[' P1 , '/' P2 , '/' P3 , '/' P4 , '/' P5 , '/' P6 , '/' P7 ']' '/ ' & P8 ] ']'"). + +Reserved Notation "[ /\ P1 , P2 , P3 , P4 , P5 , P6 , P7 , P8 & P9 ]" (at level 0, format + "'[hv' [ /\ '[' P1 , '/' P2 , '/' P3 , '/' P4 , '/' P5 , '/' P6 , '/' P7 , '/' P8 ']' '/ ' & P9 ] ']'"). + +Reserved Notation "[ /\ P1 , P2 , P3 , P4 , P5 , P6 , P7 , P8 , P9 & P10 ]" (at level 0, format + "'[hv' [ /\ '[' P1 , '/' P2 , '/' P3 , '/' P4 , '/' P5 , '/' P6 , '/' P7 , '/' P8 '/' , P9 ']' '/ ' & P10 ] ']'"). + + +Reserved Notation "[ \/ P1 , P2 , P3 , P4 | P5 ]" (at level 0, format + "'[hv' [ \/ '[' P1 , '/' P2 , '/' P3 , '/' P4 ']' '/ ' | P5 ] ']'"). +Reserved Notation "[ \/ P1 , P2 , P3 , P4 , P5 | P6 ]" (at level 0, format + "'[hv' [ \/ '[' P1 , '/' P2 , '/' P3 , '/' P4 , '/' P5 ']' '/ ' | P6 ] ']'"). + +Inductive and6 (P1 P2 P3 P4 P5 P6 : Prop) : Prop := + And6 of P1 & P2 & P3 & P4 & P5 & P6. +Inductive and7 (P1 P2 P3 P4 P5 P6 P7 : Prop) : Prop := + And7 of P1 & P2 & P3 & P4 & P5 & P6 & P7. +Inductive and8 (P1 P2 P3 P4 P5 P6 P7 P8 : Prop) : Prop := + And8 of P1 & P2 & P3 & P4 & P5 & P6 & P7 & P8. +Inductive and9 (P1 P2 P3 P4 P5 P6 P7 P8 P9 : Prop) : Prop := + And9 of P1 & P2 & P3 & P4 & P5 & P6 & P7 & P8 & P9. +Inductive and10 (P1 P2 P3 P4 P5 P6 P7 P8 P9 P10 : Prop) : Prop := + And10 of P1 & P2 & P3 & P4 & P5 & P6 & P7 & P8 & P9 & P10. + +Inductive or5 (P1 P2 P3 P4 P5 : Prop) : Prop := + Or51 of P1 | Or52 of P2 | Or53 of P3 | Or54 of P4 | Or55 of P5. +Inductive or6 (P1 P2 P3 P4 P5 P6 : Prop) : Prop := + Or61 of P1 | Or62 of P2 | Or63 of P3 | Or64 of P4 | Or65 of P5 | Or66 of P6. +Inductive or7 (P1 P2 P3 P4 P5 P6 P7 : Prop) : Prop := + Or71 of P1 | Or72 of P2 | Or73 of P3 | Or74 of P4 | Or75 of P5 | Or76 of P6 | Or77 of P7. + +Notation "[ /\ P1 , P2 , P3 , P4 , P5 & P6 ]" := (and6 P1 P2 P3 P4 P5 P6) : type_scope. +Notation "[ /\ P1 , P2 , P3 , P4 , P5 , P6 & P7 ]" := (and7 P1 P2 P3 P4 P5 P6 P7) : type_scope. +Notation "[ /\ P1 , P2 , P3 , P4 , P5 , P6 , P7 & P8 ]" := (and8 P1 P2 P3 P4 P5 P6 P7 P8) : type_scope. +Notation "[ /\ P1 , P2 , P3 , P4 , P5 , P6 , P7 , P8 & P9 ]" := (and9 P1 P2 P3 P4 P5 P6 P7 P8 P9) : type_scope. +Notation "[ /\ P1 , P2 , P3 , P4 , P5 , P6 , P7 , P8 , P9 & P10 ]" := (and10 P1 P2 P3 P4 P5 P6 P7 P8 P9 P10) : type_scope. + +Notation "[ \/ P1 , P2 , P3 , P4 | P5 ]" := (or5 P1 P2 P3 P4 P5) : type_scope. +Notation "[ \/ P1 , P2 , P3 , P4 , P5 | P6 ]" := (or6 P1 P2 P3 P4 P5 P6) : type_scope. +Notation "[ \/ P1 , P2 , P3 , P4 , P5 , P6 | P7 ]" := (or7 P1 P2 P3 P4 P5 P6 P7) : type_scope. + +(** Add the ability to rewrite with [<->] for the custom logical connectives *) + +From Coq Require Import Classes.Morphisms Program.Basics Program.Tactics. +From Coq Require Import Relations. + +Local Obligation Tactic := try solve [simpl_relation | firstorder auto]. + +Definition iter_arrow_type (n : nat) (A : Type) := iter n (fun T => A -> T). + +Fixpoint iter_respectful {S T} (A : relation S) (Z : relation T) (n : nat) + : relation (iter_arrow_type n S T) := + if n is n'.+1 then respectful A (@iter_respectful _ _ A Z n') + else Z. +Arguments iter_respectful {S T} A Z n. + +(** Logical conjunction *) +Program Instance and3_impl_morphism : Proper (iter_respectful impl impl 3) and3 | 1. +Program Instance and3_iff_morphism : Proper (iter_respectful iff iff 3) and3 | 1. + +Program Instance and4_impl_morphism : Proper (iter_respectful impl impl 4) and4 | 1. +Program Instance and4_iff_morphism : Proper (iter_respectful iff iff 4) and4 | 1. + +Program Instance and5_impl_morphism : Proper (iter_respectful impl impl 5) and5 | 1. +Program Instance and5_iff_morphism : Proper (iter_respectful iff iff 5) and5 | 1. + +Program Instance and6_impl_morphism : Proper (iter_respectful impl impl 6) and6 | 1. +Program Instance and6_iff_morphism : Proper (iter_respectful iff iff 6) and6 | 1. + +Program Instance and7_impl_morphism : Proper (iter_respectful impl impl 7) and7 | 1. +Program Instance and7_iff_morphism : Proper (iter_respectful iff iff 7) and7 | 1. + +Program Instance and8_impl_morphism : Proper (iter_respectful impl impl 8) and8 | 1. +Program Instance and8_iff_morphism : Proper (iter_respectful iff iff 8) and8 | 1. + +Program Instance and9_impl_morphism : Proper (iter_respectful impl impl 9) and9 | 1. +Program Instance and9_iff_morphism : Proper (iter_respectful iff iff 9) and9 | 1. + +Program Instance and10_impl_morphism : Proper (iter_respectful impl impl 10) and10 | 1. +Program Instance and10_iff_morphism : Proper (iter_respectful iff iff 10) and10 | 1. + +(** Logical disjunction *) +Program Instance or3_impl_morphism : Proper (iter_respectful impl impl 3) or3 | 1. +Program Instance or3_iff_morphism : Proper (iter_respectful iff iff 3) or3 | 1. + +Program Instance or4_impl_morphism : Proper (iter_respectful impl impl 4) or4 | 1. +Program Instance or4_iff_morphism : Proper (iter_respectful iff iff 4) or4 | 1. + +Program Instance or5_impl_morphism : Proper (iter_respectful impl impl 5) or5 | 1. +Program Instance or5_iff_morphism : Proper (iter_respectful iff iff 5) or5 | 1. + +Program Instance or6_impl_morphism : Proper (iter_respectful impl impl 6) or6 | 1. +Program Instance or6_iff_morphism : Proper (iter_respectful iff iff 6) or6 | 1. + +Program Instance or7_impl_morphism : Proper (iter_respectful impl impl 7) or7 | 1. +Program Instance or7_iff_morphism : Proper (iter_respectful iff iff 7) or7 | 1. + + +Section ReflectConnectives. +Variable b1 b2 b3 b4 b5 b6 b7 b8 b9 b10 : bool. + +Lemma and6P : reflect [/\ b1, b2, b3, b4, b5 & b6] [&& b1, b2, b3, b4, b5 & b6]. +Proof. +by case b1; case b2; case b3; case b4; case b5; case b6; constructor; try by case. +Qed. + +Lemma and7P : reflect [/\ b1, b2, b3, b4, b5, b6 & b7] [&& b1, b2, b3, b4, b5, b6 & b7]. +Proof. +by case b1; case b2; case b3; case b4; case b5; case b6; case: b7; constructor; try by case. +Qed. + +Lemma and8P : reflect [/\ b1, b2, b3, b4, b5, b6, b7 & b8] [&& b1, b2, b3, b4, b5, b6, b7 & b8]. +Proof. +by case b1; case b2; case b3; case b4; case b5; case b6; case: b7; case: b8; constructor; try by case. +Qed. + +Lemma and9P : reflect [/\ b1, b2, b3, b4, b5, b6, b7, b8 & b9] [&& b1, b2, b3, b4, b5, b6, b7, b8 & b9]. +Proof. +by case b1; case b2; case b3; case b4; case b5; case b6; case: b7; case: b8; case: b9; constructor; try by case. +Qed. + +Lemma and10P : reflect [/\ b1, b2, b3, b4, b5, b6, b7, b8, b9 & b10] [&& b1, b2, b3, b4, b5, b6, b7, b8, b9 & b10]. +Proof. +by case b1; case b2; case b3; case b4; case b5; case b6; + case: b7; case: b8; case: b9; case: b10; constructor; try by case. +Qed. + +Lemma or5P : reflect [\/ b1, b2, b3, b4 | b5] [|| b1, b2, b3, b4 | b5]. +Proof. +case b1; first by constructor; constructor 1. +case b2; first by constructor; constructor 2. +case b3; first by constructor; constructor 3. +case b4; first by constructor; constructor 4. +case b5; first by constructor; constructor 5. +by constructor; case. +Qed. + +Lemma or6P : reflect [\/ b1, b2, b3, b4, b5 | b6] [|| b1, b2, b3, b4, b5 | b6]. +Proof. +case b1; first by constructor; constructor 1. +case b2; first by constructor; constructor 2. +case b3; first by constructor; constructor 3. +case b4; first by constructor; constructor 4. +case b5; first by constructor; constructor 5. +case b6; first by constructor; constructor 6. +by constructor; case. +Qed. + +Lemma or7P : reflect [\/ b1, b2, b3, b4, b5, b6 | b7] [|| b1, b2, b3, b4, b5, b6 | b7]. +Proof. +case b1; first by constructor; constructor 1. +case b2; first by constructor; constructor 2. +case b3; first by constructor; constructor 3. +case b4; first by constructor; constructor 4. +case b5; first by constructor; constructor 5. +case b6; first by constructor; constructor 6. +case b7; first by constructor; constructor 7. +by constructor; case. +Qed. + +End ReflectConnectives. + +Arguments and6P {b1 b2 b3 b4 b5 b6}. +Arguments and7P {b1 b2 b3 b4 b5 b6 b7}. +Arguments and8P {b1 b2 b3 b4 b5 b6 b7 b8}. +Arguments and9P {b1 b2 b3 b4 b5 b6 b7 b8 b9}. +Arguments and10P {b1 b2 b3 b4 b5 b6 b7 b8 b9 b10}. + +Arguments or5P {b1 b2 b3 b4 b5}. +Arguments or6P {b1 b2 b3 b4 b5 b6}. +Arguments or7P {b1 b2 b3 b4 b5 b6 b7}. +Prenex Implicits and6P and7P or5P or6P or7P. + +Lemma andX (a b : bool) : reflect (a * b) (a && b). +Proof. by case: a; case: b; constructor=>//; case. Qed. + +Arguments andX {a b}. + +(********************) +(* extension to seq *) +(********************) + +Lemma all_mem (A : eqType) (s1 s2 : seq A) : + reflect {subset s2 <= s1} (all [mem s1] s2). +Proof. by case: allP=>H; constructor; [move=>x /H | move=>X; apply: H=>x /X]. Qed. + +Lemma all_predC_sym (A : eqType) (s1 s2 : seq A) : + all [predC s1] s2 = all [predC s2] s1. +Proof. by rewrite all_predC has_sym -all_predC. Qed. + +(**************) +(* empty type *) +(**************) + +Lemma false_eqP : Equality.axiom (fun _ _ : False => true). +Proof. by case. Qed. + +Definition false_EqMixin := EqMixin false_eqP. +Canonical false_EqType := Eval hnf in EqType False false_EqMixin. + +(*************) +(* sum types *) +(*************) + +Section InvertingSumTags. +Variables A B : Type. + +Definition lft : A + B -> option A := + fun x => if x is inl x' then Some x' else None. +Definition rgt : A + B -> option B := + fun x => if x is inr x' then Some x' else None. + +Lemma lft_inl_ocanc : ocancel lft inl. Proof. by case. Qed. +Lemma rgt_inr_ocanc : ocancel rgt inr. Proof. by case. Qed. +Lemma inl_lft_pcanc : pcancel inl lft. Proof. by []. Qed. +Lemma inr_rgt_pcanc : pcancel inr rgt. Proof. by []. Qed. + +End InvertingSumTags. + +Prenex Implicits lft rgt. + +Hint Extern 0 (ocancel _ _) => + (apply: lft_inl_ocanc || apply: rgt_inr_ocanc) : core. + +(********) +(* nats *) +(********) + + Lemma gt0 m n : m < n -> 0 < n. + Proof. by case: n. Qed. + +(*****************) +(* filter + last *) +(*****************) + + Lemma filter_nilp (A : eqType) (p : pred A) (s : seq A) : + reflect (forall x, p x -> x \in s -> false) + ([seq x <- s | p x] == [::]). + Proof. + case: eqP=>E; constructor. + - move=>x H1 H2; suff : x \in [seq x <- s | p x] by rewrite E. + by rewrite mem_filter H1 H2. + move=>H; apply: E; apply: size0nil; apply/eqP; rewrite size_filter. + by rewrite eqn0Ngt -has_count; apply/hasPn=>x /H; case: (p x)=>//; apply. + Qed. + + Lemma index_inj (A : eqType) (s : seq A) (x y : A) : + x \in s -> index x s = index y s -> x = y. + Proof. + elim: s=>[|k s IH] //=; rewrite inE eq_sym. + case: eqP=>[->{k} _|_ /= S]; first by case: eqP. + by case: eqP=>// _ []; apply: IH S. + Qed. + + Section LastFilter. + Variables (A : eqType). + + (* if s has an element, last returns one of them *) + Lemma last_in x k (s : seq A) : x \in s -> last k s \in s. + Proof. + elim: s k=>[|k s IH] k' //=; rewrite !inE. + case/orP=>[/eqP <-|/IH ->]; first by apply: mem_last. + by rewrite orbT. + Qed. + + Arguments last_in x [k s]. + + Lemma last_notin x k (s : seq A) : x \in s -> k \notin s -> last k s != k. + Proof. by move/(last_in _ (k:=k))=>H /negbTE; case: eqP H=>// ->->. Qed. + + (* last either returns a default, or one of s's elements *) + Lemma last_change k (s : seq A) : last k s != k -> last k s \in s. + Proof. by move: (mem_last k s); rewrite inE; case: eqP. Qed. + + Lemma last_changeE1 k (s : seq A) : + last k s != k -> forall x, last x s = last k s. + Proof. by elim: s k=>[|k s IH] y //=; rewrite eq_refl. Qed. + + Lemma last_changeE2 k (s : seq A) : + last k s != k -> forall x, x \notin s -> last x s != x. + Proof. by move/last_change/last_notin. Qed. + + (* common formats of last_change *) + Lemma last_nochange k (s : seq A) : last k s = k -> (k \in s) || (s == [::]). + Proof. + case: s k=>[|k s] //= k'; rewrite inE; case: eqP=>[->|N L] //. + by move: (@last_change k s); rewrite L=>-> //; case: eqP N. + Qed. + + Lemma last_nochange_nil k (s : seq A) : last k s = k -> k \notin s -> s = [::]. + Proof. by move/last_nochange; case/orP=>[/negbF ->|/eqP]. Qed. + + (* last has bigger index than anything in x *) + Lemma index_last_mono x k (s : seq A) : + uniq (k :: s) -> x \in s -> index x s <= index (last k s) s. + Proof. + elim: s k=>[|k s IH] //= k'; rewrite !inE negb_or (eq_sym x). + case: eqP=>//= _; case: eqP=>//= _ /and3P [_ H2 H3 H]. + case: eqP=>[/esym E|_]; last by apply: IH=>//=; rewrite H2 H3. + by rewrite (last_nochange_nil E H2) in H. + Qed. + + (* if it has bigger index, and is in the list, then it's last *) + Lemma max_index_last (s : seq A) (x y : A) : + uniq s -> x \in s -> + (forall z, z \in s -> index z s <= index x s) -> last y s = x. + Proof. + elim: s y=>[|k s IH] y //= /andP [Nk U]; rewrite inE (eq_sym k). + case: (x =P k) Nk=>[<-{k} Nk _|_ Nk /= S] /= D; last first. + - apply: IH=>// z Z; move: (D z); rewrite inE Z orbT=>/(_ (erefl _)). + by case: ifP Z Nk=>// /eqP ->->. + suff : size s == 0 by move/eqP/size0nil=>->. + rewrite eqn0Ngt -has_predT; apply/hasPn=>z Z. + move: (D z); rewrite inE Z orbT=>/(_ (erefl _)). + by case: ifP Z Nk=>// /eqP ->->. + Qed. + + (* last_filter either returns default or a p-element of ks *) + Lemma last_filter_change k p (ks : seq A) : + last k (filter p ks) != k -> + p (last k (filter p ks)) && (last k (filter p ks) \in ks). + Proof. by move/last_change; rewrite mem_filter. Qed. + + Lemma index_filter_mono (p : pred A) (ks : seq A) x y : + p x -> index x ks <= index y ks -> + index x (filter p ks) <= index y (filter p ks). + Proof. + move=>Px; elim: ks=>[|k ks IH] //=; case P : (p k)=>/=; + by case: ifP Px; case: ifP=>// _ /eqP <-; rewrite P. + Qed. + + Lemma filter_sub (p1 p2 : pred A) (s : seq A) : + subpred p1 p2 -> {subset filter p1 s <= filter p2 s}. + Proof. + move=>S; rewrite (_ : filter p1 s = filter p1 (filter p2 s)). + - by apply: mem_subseq; apply: filter_subseq. + rewrite -filter_predI; apply: eq_in_filter=>x X /=. + by case E : (p1 x)=>//=; rewrite (S _ E). + Qed. + + Lemma subset_nil (s : seq A) : {subset s <= [::]} -> s = [::]. + Proof. + move=>H; apply: size0nil; apply/eqP. + rewrite eqn0Ngt -(filter_predT s) size_filter -has_count. + by case: hasP=>//; case=>x /H. + Qed. + + Lemma last_filter_neq (p1 p2 : pred A) x (s : seq A) : + subpred p1 p2 -> x \notin s -> + last x (filter p1 s) != x -> last x (filter p2 s) != x. + Proof. + move=>S N /last_filter_change /andP [H1 H2]. + apply: (@last_notin (last x [seq x <-s | p1 x])). + - by rewrite mem_filter H2 andbT; apply: S. + by rewrite mem_filter negb_and N orbT. + Qed. + + Lemma last_filter_eq (p1 p2 : pred A) x (s : seq A) : + subpred p1 p2 -> x \notin s -> + last x (filter p2 s) = x -> last x (filter p1 s) = x. + Proof. + move=>S N /eqP E; apply/eqP. + by apply: contraTT E; apply: last_filter_neq. + Qed. + + Lemma index_last_sub (p1 p2 : pred A) x (s : seq A) : + subpred p1 p2 -> uniq (x :: s) -> + index (last x (filter p1 s)) (x :: s) <= + index (last x (filter p2 s)) (x :: s). + Proof. + move=>S; elim: s x=>[|k s IH] //= x; rewrite !inE negb_or -andbA. + rewrite -(eq_sym k) -!(eq_sym (last _ _)); case/and4P=>N Sx Sk U. + have [Ux Uk] : uniq (x :: s) /\ uniq (k :: s) by rewrite /= Sx Sk U. + case P1 : (p1 k)=>/=. + - rewrite (S _ P1) /=; case: (last k _ =P k). + - move/last_nochange; rewrite mem_filter (negbTE Sk) andbF /=. + move/eqP=>-> /=; rewrite (negbTE N). + case: (last k _ =P k); first by move=>->; rewrite (negbTE N). + by case/eqP/last_filter_change/andP; case: eqP Sx=>// <- /negbTE ->. + move/eqP=>N1; move: (last_filter_neq S Sk N1)=>N2. + move: (IH _ Uk); rewrite /= !(eq_sym k). + rewrite (negbTE N1) (negbTE N2) -(last_changeE1 N1 x) -(last_changeE1 N2 x). + rewrite (negbTE (last_changeE2 N1 _)) ?(mem_filter,negb_and,Sx,orbT) //. + by rewrite (negbTE (last_changeE2 N2 _)) ?(mem_filter,negb_and,Sx,orbT). + case P2 : (p2 k)=>/=. + - case: (last x _ =P x)=>// /eqP N1; move: (last_filter_neq S Sx N1)=>N2. + move: (IH _ Ux); rewrite /= !(eq_sym x) (negbTE N1) (negbTE N2). + rewrite -(last_changeE1 N1 k) {1 3}(last_changeE1 N2 k). + rewrite (negbTE (last_changeE2 N1 _)) ?(mem_filter,negb_and,Sk,orbT) //. + by rewrite !(negbTE (last_changeE2 N2 _)) ?(mem_filter,negb_and,Sk,Sx,orbT). + case: (last x _ =P x)=>// /eqP N1; move: (last_filter_neq S Sx N1)=>N2. + move: (IH _ Ux); rewrite /= !(eq_sym x) (negbTE N1) (negbTE N2). + rewrite -(last_changeE1 N1 k) -(last_changeE1 N2 k). + rewrite (negbTE (last_changeE2 N1 _)) ?(mem_filter,negb_and,Sk,orbT) //. + by rewrite !(negbTE (last_changeE2 N2 _)) ?(mem_filter,negb_and,Sk,orbT). + Qed. + + Lemma last_filter_last_helper (p : pred A) x (s : seq A) y : + uniq (x :: s) -> p y -> y \in s -> + index y s <= index (last x (filter p s)) s. + Proof. + elim: s x=>[|k s IH] x //=; rewrite !inE !negb_or !(eq_sym _ k). + case/andP=>/andP [H1 H2] /andP [H3 H4] Px. + case: eqP=> [->|_] //= Ks; case P: (p k)=>/=. + - case: eqP=>E; last by apply: IH=>//=; rewrite H3 H4. + move: (@last_in y k (filter p s)); rewrite -E !mem_filter. + by rewrite Px Ks P (negbTE H3); move/(_ (erefl _)). + case: eqP=>E; last by apply: IH=>//=; rewrite H2 H4. + by move: H1; rewrite E; move/last_filter_change; rewrite -E P. + Qed. + + Lemma last_filter_last (p : pred A) x (s : seq A) y : + uniq (x :: s) -> p y -> y \in s -> + index y (x :: s) <= index (last x (filter p s)) (x :: s). + Proof. + move=>/= /andP [Sx U] H Sy /=; case: (x =P y)=>//= _. + have Hy : y \in [seq x <- s | p x] by rewrite mem_filter H Sy. + rewrite eq_sym; case: (last x _ =P x); last first. + - by move=>_; apply: last_filter_last_helper=>//=; rewrite Sx U. + move/last_nochange; rewrite mem_filter (negbTE Sx) andbF /=. + by move/eqP=>E; rewrite E in Hy. + Qed. + + End LastFilter. \ No newline at end of file diff --git a/core/rtc.v b/core/rtc.v new file mode 100644 index 0000000..b8d6c21 --- /dev/null +++ b/core/rtc.v @@ -0,0 +1,274 @@ +(* +Copyright 2020 IMDEA Software Institute +Licensed under the Apache License, Version 2.0 (the "License"); +you may not use this file except in compliance with the License. +You may obtain a copy of the License at + http://www.apache.org/licenses/LICENSE-2.0 +Unless required by applicable law or agreed to in writing, software +distributed under the License is distributed on an "AS IS" BASIS, +WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. +See the License for the specific language governing permissions and +limitations under the License. +*) + +From Coq Require Import ssreflect ssrbool ssrfun. +From mathcomp Require Import ssrnat eqtype seq choice fintype path fingraph. +From fcsl Require Import axioms pred prelude finmap. +From fcsl Require Import options. + +(* Reflexive and Transitive closures of a relation with finite domain *) +(* The domain is currently taken to a finite set of nats *) + +Section TransitiveClosure. +Variables (h : seq nat) (R : rel nat). +Hypothesis Rclosed : forall x y, R x y -> (x \in h) = (y \in h). + +Local Definition tp := [finType of seq_sub h]. + +Definition Rtp (x y : tp) : bool := R (eqtype.val x) (eqtype.val y). + +(* reflexive transitive closure *) +Definition rtc (x y : nat) : bool := + if decP (@idP (x \in h)) is left pfx then + if decP (@idP (y \in h)) is left pfy then + connect Rtp (Sub x pfx) (Sub y pfy) + else false + else false. + +(* irreflexive transitive closure *) +(* i.e., we make at least one step by R *) +Definition tc (x y : nat) := + (x \in h) && has (fun z => R x z && rtc z y) h. + +(* lemmas about rtc *) + +Lemma path_closed x p : path R x p -> (x \in h) = (last x p \in h). +Proof. by elim: p x=>[|z p IH] x //= /andP [Rxz /IH <-]; apply: Rclosed. Qed. + +Lemma iter'_path n x y : + iter' R n x y <-> exists p, [/\ size p = n, path R x p & y = last x p]. +Proof. +split. +- elim: n x=>[|n IH] x /=; first by move=>->; exists [::]. + by case=>z [Rxz] /IH [p][Sp Pzp Ly]; exists (z::p)=>//=; rewrite Sp Rxz. +elim: n x=>[|n IH] x /=; first by case=>p [/size0nil ->]. +case=>p []; case: p=>//= z p [Sp] /andP [Rxz Pzp] Ly. +by exists z; split=>//; apply: IH; exists p. +Qed. + +Lemma iter_path x y : iter R x y <-> exists2 p, path R x p & y = last x p. +Proof. +split; first by case=>n /iter'_path [p][Sp Pxp Ly]; exists p. +by case=>p Pxp Ly; exists (size p); apply/iter'_path; exists p. +Qed. + +Lemma rtc_pathP x y : + reflect (x \in h /\ exists2 p, path R x p & y = last x p) + (rtc x y). +Proof. +case V : (rtc x y); constructor. +- move: V; rewrite /rtc; case: decP=>// pfx; case: decP=>// pfy. + case/connectP=>p P E; split=>//; rewrite (_ : x = ssval (Sub x pfx)) //. + have {E} Ey : y = ssval (last (Sub x pfx) p) by rewrite -E. + exists (map (@ssval _ _) p); last by rewrite last_map. + by rewrite (map_path (e':=Rtp) (b:=pred0)) //; apply/hasP; case. +case=>Mx [p P Ey]; move: V; rewrite /rtc. +case: decP=>// {Mx} pfx; case: decP=>[pfy|]; last first. +- move: (mem_last x p) pfx; rewrite -Ey inE=>/orP [/eqP ->->//|]. + elim: p x P Ey=>[|p ps IH] //= x /andP [Rxp P] Ey. + by rewrite inE (Rclosed Rxp); case: eqP=>[-> _|_ My] //; apply: IH. +case/connectP; elim: p x pfx P Ey=>[|p ps IH] /= x pfx P Ey. +- by rewrite Ey /= in pfy *; rewrite (bool_irrelevance pfx); exists [::]. +case/andP: P=>Rxp P; move: {-}(pfx); rewrite (Rclosed Rxp)=>pfp. +by case: (IH p pfp P Ey)=>q; exists (Sub p pfp :: q)=>//; apply/andP. +Qed. + +Lemma rtcP x y : reflect (x \in h /\ iter R x y) (rtc x y). +Proof. by case: rtc_pathP=>[[pfx]|] H; constructor; rewrite iter_path. Qed. + +Lemma rtc_closed x y : rtc x y -> (x \in h) && (y \in h). +Proof. by case/rtc_pathP=>Px [p] /path_closed Mx ->; rewrite -Mx Px. Qed. + +Lemma rtc_cond x y : rtc x y -> (x \in h) = (y \in h). +Proof. by case/rtc_closed/andP=>->->. Qed. + +(* sometimes we start from the end of the path *) + +Lemma rtc_pathP_last x y : + reflect (y \in h /\ exists2 p, path R x p & y = last x p) + (rtc x y). +Proof. +case: rtc_pathP=>[[pfx]|] H; constructor. +- by split=>//; case: H pfx=>p /path_closed ->->. +by case=>pfy X; apply: H; split=>//; case: X pfy=>p /path_closed ->->. +Qed. + +Lemma rtcP_last x y : reflect (y \in h /\ iter R x y) (rtc x y). +Proof. by case: rtc_pathP_last=>[[pfx]|] H; constructor; rewrite iter_path. Qed. + +(* lemmas about tc *) + +Lemma tc_pathP x y : + reflect (x \in h /\ exists z p, [/\ R x z, path R z p & y = last z p]) + (tc x y). +Proof. +rewrite /tc; case: andP=>[[Mx /hasP]|] H; constructor. +- by case: H=>z ? /andP [?] /rtc_pathP [_][p]; split=>//; exists z, p. +case=>Mx [z][p][Rxz P E]; have Mz : z \in h by rewrite -(Rclosed Rxz). +apply: H; case: hasP=>//; case; exists z=>//; rewrite Rxz. +by apply/rtc_pathP; split=>//; exists p. +Qed. + +Lemma tcP x y : reflect (x \in h /\ exists n, iter' R n.+1 x y) (tc x y). +Proof. +case: tc_pathP=>H; constructor. +- case: H=>Mx [z][p][Rxz P E]; split=>//; exists (size p). + by exists z; split=>//; apply/iter'_path; exists p. +case=>Mx [n /= [z][Rxz] /iter'_path [p][Sp Pzp Ly]]. +by apply: H; split=>//; exists z, p. +Qed. + +Lemma tc_closed x y : tc x y -> (x \in h) && (y \in h). +Proof. +case/tc_pathP=>Mx [z][p][Rxz /path_closed Px ->]. +by rewrite -Px -(Rclosed Rxz) Mx. +Qed. + +Lemma tc_cond x y : tc x y -> (x \in h) = (y \in h). +Proof. by case/tc_closed/andP=>->->. Qed. + +Lemma tc_pathP_last x y : + reflect (y \in h /\ exists z p, [/\ R z y, path R x p & z = last x p]) + (tc x y). +Proof. +case: tc_pathP=>H; constructor. +- case: H=>Mx [z][p][Rxz Pzp Ly]; split. + - by rewrite {1}Ly -(path_closed Pzp) -(Rclosed Rxz). + case: (lastP p) Pzp Ly=>[|q w]; first by move=>_ ->; exists x, [::]. + rewrite rcons_path last_rcons=>/andP [Pzw Rwq ->{y}]. + by exists (last z q), (z::q); rewrite /= Rxz. +case=>My [z][p][Rzy Pxp Lz]; apply: H; split. +- by rewrite (path_closed Pxp) -Lz (Rclosed Rzy). +case: p Pxp Lz=>[|w q] /=; first by move=>_ <-; exists y, [::]. +case/andP=>Rxw Pwq E; exists w, (rcons q y). +by rewrite rcons_path last_rcons -E Pwq. +Qed. + +Lemma tcP_last x y : reflect (y \in h /\ exists n, iter' R n.+1 x y) (tc x y). +Proof. +case: tc_pathP_last=>H; constructor. + case: H=>My [z][p][Rzy Pxp Lz]; split=>//; exists (size p); apply/iterS. + by exists z; split=>//; apply/iter'_path; exists p. +case=>My [n] /iterS [z][/iter'_path [p][Sp Pxp Lz] Rzy]. +by apply: H; split=>//; exists z, p. +Qed. + +Lemma rtc1 x y : x \in h -> y \in h -> R x y -> rtc x y. +Proof. +rewrite /rtc=>Dx Dy Rxy; case: decP=>// Dx'; case: decP=>// Dy'. +by apply/connectP; exists [:: Sub y Dy']=>//=; rewrite andbT. +Qed. + +Lemma rtc_refl x : x \in h -> rtc x x. +Proof. by rewrite /rtc; case: decP. Qed. + +Lemma tc1 x y : x \in h -> y \in h -> R x y -> tc x y. +Proof. +by rewrite /tc=>-> Dy Rxy; apply/hasP; exists y=>//; rewrite Rxy rtc_refl. +Qed. + +Lemma rtc_trans : transitive rtc. +Proof. +rewrite /rtc=>x y z; case: decP=>// Dy; case: decP=>// Dx. +by case: decP=>// Dz; apply: connect_trans. +Qed. + +Lemma tc_trans : transitive tc. +Proof. +rewrite /tc=>x y z /andP [Dy /hasP [y' Dy' /andP [Ry Ry'x]]]. +case/andP=>Dx /hasP [x' Dx' /andP [Rx Rx'z]]. +rewrite Dy; apply/hasP; exists y'=>//; rewrite Ry. +by apply: rtc_trans (rtc_trans Ry'x (rtc1 _ _ _)) Rx'z. +Qed. + +Lemma rtc_tc x y : tc x y -> rtc x y. +Proof. +case/andP=>Dx /hasP [z Dz] /andP [Rxz]. +by apply: rtc_trans (rtc1 Dx Dz Rxz). +Qed. + +End TransitiveClosure. + + +(* some helper lemmas *) + +Lemma connect_idemp (T : finType) (R : rel T) : + reflexive R -> transitive R -> connect R =2 R. +Proof. +move=>Rr Tr x y; apply/connectP; case: ifP=>[Rxy|H [p P Ly]]. +- by exists [:: y]=>//=; rewrite Rxy. +by move: (path_lastR Rr Tr P); rewrite -Ly H. +Qed. + +Lemma rtcT (h : seq nat) (R : rel nat) x y : + (forall x y, R x y -> (x \in h) && (y \in h)) -> + transitive R -> rtc h R x y -> R x y || (x == y) && (x \in h). +Proof. +move=>Rcond Tr /rtcP [|Dx [n]]; first by move=>?? /Rcond/andP [->->]. +elim: n x Dx=>[|n IH] x Dx /=; first by move=><-; rewrite eq_refl Dx orbT. +case=>z [Rxz]; case/Rcond/andP: (Rxz)=>-> Dz /(IH _ Dz). +by case: eqP Rxz=>[<- ->//|_ Rxz]; rewrite orbF=>/(Tr _ _ _ Rxz)=>->. +Qed. + +Lemma rtc_idemp (h : seq nat) (R : rel nat) : + (forall x y, R x y -> (x \in h) && (y \in h)) -> + rtc h (rtc h R) =2 rtc h R. +Proof. +move=>Rcond x y. +have X : forall x y, R x y -> (x \in h) = (y \in h). +- by move=>?? /Rcond/andP [->->]. +apply/idP/idP; last first. +- by move=>H; case/andP: (rtc_closed X H)=>Dx Dy; apply: rtc1. +case/(rtcT (rtc_closed X) (@rtc_trans h R))/orP=>//. +by case/andP=>/eqP <-; apply: rtc_refl. +Qed. + +Lemma rtc_sub (h : seq nat) (R1 R2 : rel nat) x y : + (forall x y, R2 x y -> (x \in h) && (y \in h)) -> + (forall x y, R1 x y -> R2 x y) -> + rtc h R1 x y -> rtc h R2 x y. +Proof. +move=>Rcond Rsub. +have X : forall x y, R2 x y -> (x \in h) = (y \in h). +- by move=>?? /Rcond /andP [->->]. +case/rtcP=>[?? /Rsub/X //|Dx H]; apply/rtcP=>//. +by split=>//; apply: iter_sub Rsub H. +Qed. + +Lemma tcT (h : seq nat) (R : rel nat) : + (forall x y, R x y -> (x \in h) && (y \in h)) -> + transitive R -> tc h R =2 R. +Proof. +move=>Rcond Tr x y; apply/idP/idP; last first. +- by move=>H; case/andP: (Rcond _ _ H)=>Dx Dy; apply: tc1. +case/andP=>Dx /hasP [z Dz] /andP [Rxz] /(rtcT Rcond Tr). +by case: eqP=>[<-|_] //; rewrite orbF; apply: Tr Rxz. +Qed. + +Lemma tc_idemp (h : seq nat) (R : rel nat) : + (forall x y, R x y -> (x \in h) && (y \in h)) -> + tc h (tc h R) =2 tc h R. +Proof. +move=>Rcond x y; rewrite (tcT _ (@tc_trans _ _)) //. +by apply: tc_closed=>?? /Rcond/andP [->->]. +Qed. + +Lemma tc_sub (h : seq nat) (R1 R2 : rel nat) x y : + (forall x y, R2 x y -> (x \in h) && (y \in h)) -> + (forall x y, R1 x y -> R2 x y) -> + tc h R1 x y -> tc h R2 x y. +Proof. +move=>Rcond Rsub /andP [Dx] /hasP [z Dz] /andP [Rxz R]. +rewrite /tc Dx; apply/hasP; exists z=>//. +by rewrite (Rsub _ _ Rxz) (rtc_sub _ _ R). +Qed. diff --git a/core/seqperm.v b/core/seqperm.v new file mode 100644 index 0000000..fd6435c --- /dev/null +++ b/core/seqperm.v @@ -0,0 +1,209 @@ +(* +Copyright 2017 IMDEA Software Institute +Licensed under the Apache License, Version 2.0 (the "License"); +you may not use this file except in compliance with the License. +You may obtain a copy of the License at + http://www.apache.org/licenses/LICENSE-2.0 +Unless required by applicable law or agreed to in writing, software +distributed under the License is distributed on an "AS IS" BASIS, +WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. +See the License for the specific language governing permissions and +limitations under the License. +*) + +From Coq Require Import ssreflect ssrbool ssrfun. +From mathcomp Require Import ssrnat seq path eqtype. +From fcsl Require Import pred. +From fcsl Require Import options. + +(****************************************************) +(* A theory of permutations over non-equality types *) +(****************************************************) + +Section Permutations. +Variable A : Type. + +Inductive perm : seq A -> seq A -> Prop := +| permutation_nil : perm [::] [::] +| permutation_skip x s1 s2 of perm s1 s2 : perm (x :: s1) (x :: s2) +| permutation_swap x y s1 s2 of perm s1 s2 : perm [:: x, y & s1] [:: y, x & s2] +| permutation_trans t s1 s2 of perm s1 t & perm t s2 : perm s1 s2. + +Lemma pperm_refl (s : seq A) : perm s s. +Proof. by elim: s=>*; [apply: permutation_nil | apply: permutation_skip]. Qed. + +Hint Resolve pperm_refl : core. + +Lemma pperm_nil (s : seq A) : perm [::] s <-> s = [::]. +Proof. +split; last by move=>->; apply: permutation_nil. +move E: {1}[::]=>l H; move: H {E}(esym E). +by elim=>//??? _ IH1 _ IH2 /IH1/IH2. +Qed. + +Lemma pperm_sym s1 s2 : perm s1 s2 <-> perm s2 s1. +Proof. +suff {s1 s2} L : forall s1 s2, perm s1 s2 -> perm s2 s1 by split; apply: L. +apply: perm_ind=>[|||??? _ H1 _ H2] *; +by [apply: permutation_nil | apply: permutation_skip | + apply: permutation_swap | apply: permutation_trans H2 H1]. +Qed. + +Lemma pperm_trans s2 s1 s3 : perm s1 s2 -> perm s2 s3 -> perm s1 s3. +Proof. by apply: permutation_trans. Qed. + +Lemma pperm_in s1 s2 x : perm s1 s2 -> x \In s1 -> x \In s2. +Proof. elim=>//??? =>[|?|_ I1 _ I2 /I1/I2]; rewrite ?InE; tauto. Qed. + +Lemma pperm_catC s1 s2 : perm (s1 ++ s2) (s2 ++ s1). +Proof. +elim: s1 s2=>[|x s1 IH1] s2 /=; first by rewrite cats0. +apply: (@pperm_trans (x::s2++s1)); first by apply: permutation_skip. +elim: s2=>[|y s2 IH2] //=. +apply: (@pperm_trans (y::x::s2++s1)); first by apply: permutation_swap. +by apply: permutation_skip. +Qed. + +Hint Resolve pperm_catC : core. + +Lemma pperm_cat2lL s s1 s2 : perm s1 s2 -> perm (s ++ s1) (s ++ s2). +Proof. by elim: s=>[//|e s IH /IH]; apply: permutation_skip. Qed. + +Lemma pperm_cat2rL s s1 s2 : perm s1 s2 -> perm (s1 ++ s) (s2 ++ s). +Proof. +move=>?. +apply: (@pperm_trans (s ++ s1)); first by apply: pperm_catC. +apply: (@pperm_trans (s ++ s2)); last by apply: pperm_catC. +by apply: pperm_cat2lL. +Qed. + +Lemma pperm_catL s1 t1 s2 t2 : + perm s1 s2 -> perm t1 t2 -> perm (s1 ++ t1) (s2 ++ t2). +Proof. by move/(pperm_cat2rL t1)=>H1/(pperm_cat2lL s2); apply: pperm_trans. Qed. + +Lemma pperm_cat_consL s1 t1 s2 t2 x : + perm s1 s2 -> perm t1 t2 -> perm (s1 ++ x :: t1) (s2 ++ x :: t2). +Proof. by move=>*; apply: pperm_catL=>//; apply: permutation_skip. Qed. + +Lemma pperm_cons_catCA s1 s2 x : perm (x :: s1 ++ s2) (s1 ++ x :: s2). +Proof. +rewrite -cat1s -(cat1s _ s2) !catA. +by apply/pperm_cat2rL/pperm_catC. +Qed. + +Lemma pperm_cons_catAC s1 s2 x : perm (s1 ++ x :: s2) (x :: s1 ++ s2). +Proof. by apply/pperm_sym/pperm_cons_catCA. Qed. + +Hint Resolve pperm_cons_catCA pperm_cons_catAC : core. + +Lemma pperm_cons_cat_consL s1 s2 s x : + perm s (s1 ++ s2) -> perm (x :: s) (s1 ++ x :: s2). +Proof. +move=>?. +apply: (@pperm_trans (x :: (s1 ++ s2))); first by apply: permutation_skip. +by apply: pperm_cons_catCA. +Qed. + +Lemma pperm_size l1 l2: perm l1 l2 -> size l1 = size l2. +Proof. by elim=>//=???? =>[|?|]->. Qed. + +Lemma pperm_cat_consR s1 s2 t1 t2 x : + perm (s1 ++ x :: t1) (s2 ++ x :: t2) -> perm (s1 ++ t1) (s2 ++ t2). +Proof. +move: s1 t1 s2 t2 x. +suff H: + forall r1 r2, perm r1 r2 -> forall x s1 t1 s2 t2, + r1 = s1 ++ x :: t1 -> r2 = s2 ++ x :: t2 -> perm (s1 ++ t1) (s2 ++ t2). +- by move=>s1 t1 s2 t2 x /H; apply. +apply: perm_ind; last 1 first. +- move=>s2 s1 s3 H1 IH1 H2 IH2 x r1 t1 r2 t2 E1 E2. + case: (@In_split _ x s2). + - by apply: pperm_in H1 _; rewrite E1 Mem_cat; right; left. + move=>s4 [s5] E; apply: (@pperm_trans (s4++s5)); first by apply: IH1 E1 E. + by apply: IH2 E E2. +- by move=>x []. +- move=>x t1 t2 H IH y [|b s1] s2 [|c p1] p2 /= [E1 E2] [E3 E4]; subst x; + rewrite ?E1 ?E2 ?E3 ?E4 in H * =>//. + - by subst y; apply: pperm_trans H _. + - by apply: pperm_trans H. + by apply: permutation_skip=>//; apply: IH E2 E4. +move=>x y p1 p2 H IH z [|b s1] t1 [|c s2] t2 /= [E1 E2] [E3 E4]; subst x y; + rewrite -?E2 -?E4 in H IH * =>//. +- by apply: permutation_skip. +- case: s2 E4=>/=[|a s2][<-]=>[|E4]; apply: permutation_skip=>//. + by subst p2; apply: pperm_trans H _; apply pperm_cons_catAC. +- case: s1 E2=>/=[|a s1][<-]=>[|E2]; apply: permutation_skip=>//. + by subst p1; apply: pperm_trans H; apply pperm_cons_catCA. +case: s1 E2=>/=[|a s1][->]=>E2; case: s2 E4=>/=[|d s2][->]=>E4; + rewrite ?E2 ?E4 in H IH *. +- by apply: permutation_skip. +- apply: (@pperm_trans [:: d, z & s2 ++ t2]); last by apply: permutation_swap. + by apply: permutation_skip=>//; apply/(pperm_trans H _ )/pperm_cons_catAC. +- apply: (@pperm_trans [:: a, z & s1 ++ t1]); first by apply: permutation_swap. + by apply: permutation_skip=>//; apply/pperm_trans/H/pperm_cons_catCA. +by apply: permutation_swap; apply: IH. +Qed. + +Lemma pperm_cons x s1 s2 : perm (x :: s1) (x :: s2) <-> perm s1 s2. +Proof. +by split; [apply/(@pperm_cat_consR [::] [::]) | apply: permutation_skip]. +Qed. + +Lemma pperm_cat2l s s1 s2: perm (s ++ s1) (s ++ s2) <-> perm s1 s2. +Proof. by split; [elim: s=>// ??? /pperm_cons | apply: pperm_cat2lL]. Qed. + +Lemma pperm_cat2r s s1 s2 : perm (s1 ++ s) (s2 ++ s) <-> perm s1 s2. +Proof. +split; last by apply: pperm_cat2rL. +by elim: s=>[|??? /pperm_cat_consR]; rewrite ?cats0. +Qed. + +Lemma pperm_catAC s1 s2 s3 : perm ((s1 ++ s2) ++ s3) ((s1 ++ s3) ++ s2). +Proof. by move=>*; rewrite -!catA pperm_cat2l. Qed. + +Lemma pperm_catCA s1 s2 s3 : perm (s1 ++ s2 ++ s3) (s2 ++ s1 ++ s3). +Proof. by move=>*; rewrite !catA pperm_cat2r. Qed. + +Lemma pperm_cons_cat_cons x s1 s2 s : + perm (x :: s) (s1 ++ x :: s2) <-> perm s (s1 ++ s2). +Proof. +by split; [apply: (@pperm_cat_consR [::]) | apply: pperm_cons_cat_consL]. +Qed. + +Lemma pperm_cat_cons x s1 s2 t1 t2 : + perm (s1 ++ x :: t1) (s2 ++ x :: t2) <-> perm (s1 ++ t1) (s2 ++ t2). +Proof. +split=>[|H]; first by apply: pperm_cat_consR. +apply: (@pperm_trans (x::s1++t1))=>//; apply: (@pperm_trans (x::s2++t2))=>//. +by apply/pperm_cons. +Qed. + +End Permutations. + +Hint Resolve pperm_refl pperm_catC pperm_cons_catCA + pperm_cons_catAC pperm_catAC pperm_catCA : core. + +(* perm and map *) +Lemma pperm_map A B (f : A -> B) (s1 s2 : seq A) : + perm s1 s2 -> perm (map f s1) (map f s2). +Proof. +elim=>[//|||??? _ IH1 _ IH2]*; +by [apply/pperm_cons | apply/permutation_swap | apply/(pperm_trans IH1 IH2)]. +Qed. + +(* mapping to ssreflect decidable perm *) +Lemma perm_eq_perm (A : eqType) (s1 s2 : seq A) : + reflect (perm s1 s2) (perm_eq s1 s2). +Proof. +apply: (iffP idP); last first. +- elim=>[|||??? _ H1 _ H2]*. + - by apply seq.perm_refl. + - by rewrite seq.perm_cons. + - by rewrite -![[:: _, _ & _]]/([::_] ++ [::_] ++ _) seq.perm_catCA; + rewrite !seq.perm_cat2l. + by apply: seq.perm_trans H1 H2. +elim: s2 s1 =>[s1 /seq.perm_size/size0nil->// | x s2 IH s1 H]. +move: (seq.perm_mem H x); rewrite mem_head=>H'; move: H' H. +move/splitPr=>[p1 p2]; rewrite -cat1s seq.perm_catCA seq.perm_cons=>/IH. +by rewrite -[_::s2]cat0s pperm_cat_cons. +Qed. diff --git a/finmap/ordtype.v b/finmap/ordtype.v deleted file mode 100644 index 85a2578..0000000 --- a/finmap/ordtype.v +++ /dev/null @@ -1,331 +0,0 @@ -(* -Copyright 2010 IMDEA Software Institute -Licensed under the Apache License, Version 2.0 (the "License"); -you may not use this file except in compliance with the License. -You may obtain a copy of the License at - http://www.apache.org/licenses/LICENSE-2.0 -Unless required by applicable law or agreed to in writing, software -distributed under the License is distributed on an "AS IS" BASIS, -WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. -See the License for the specific language governing permissions and -limitations under the License. -*) - -(******************************************************************************) -(* This file defines ordType - the structure for types with a decidable *) -(* (strict) order relation. *) -(* ordType is a subclass of mathcomp's eqType *) -(* This file also defines some important instances of ordType *) -(******************************************************************************) - -From Coq Require Import ssreflect ssrbool ssrfun. -From mathcomp Require Import ssrnat eqtype seq path fintype. -Set Implicit Arguments. -Unset Strict Implicit. -Unset Printing Implicit Defensive. - -Module Ordered. - -Section RawMixin. - -Structure mixin_of (T : eqType) := - Mixin {ordering : rel T; - _ : irreflexive ordering; - _ : transitive ordering; - _ : forall x y, [|| ordering x y, x == y | ordering y x]}. - -End RawMixin. - -(* the class takes a naked type T and returns all the *) -(* relatex mixins; the inherited ones and the added ones *) -Section ClassDef. - -Record class_of (T : Type) := Class { - base : Equality.class_of T; - mixin : mixin_of (EqType T base)}. - -Local Coercion base : class_of >-> Equality.class_of. - -Structure type : Type := Pack {sort : Type; _ : class_of sort}. -Local Coercion sort : type >-> Sortclass. - -Variables (T : Type) (cT : type). -Definition class := let: Pack _ c as cT' := cT return class_of cT' in c. -Definition clone c of phant_id class c := @Pack T c. - -(* produce an ordered type out of the inherited mixins *) -(* equalize m0 and m by means of a phantom; will be exploited *) -(* further down in the definition of OrdType *) -Definition pack b (m0 : mixin_of (EqType T b)) := - fun m & phant_id m0 m => Pack (@Class T b m). - -Definition eqType := Eval hnf in EqType cT class. - -End ClassDef. - -Module Exports. -Coercion sort : type >-> Sortclass. -Coercion eqType : type >-> Equality.type. -Canonical Structure eqType. -Notation ordType := Ordered.type. -Notation OrdMixin := Mixin. -Notation OrdType T m := (@pack T _ m _ id). -Definition ord T : rel (sort T) := (ordering (mixin (class T))). -Notation "[ 'ordType' 'of' T 'for' cT ]" := (@clone T cT _ id) - (at level 0, format "[ 'ordType' 'of' T 'for' cT ]") : form_scope. -Notation "[ 'ordType' 'of' T ]" := (@clone T _ _ id) - (at level 0, format "[ 'ordType' 'of' T ]") : form_scope. -End Exports. -End Ordered. -Export Ordered.Exports. - -Definition oleq (T : ordType) (t1 t2 : T) := ord t1 t2 || (t1 == t2). - -Prenex Implicits ord oleq. - -Section Lemmas. -Variable T : ordType. -Implicit Types x y : T. - -Lemma irr : irreflexive (@ord T). -Proof. by case: T=>s [b [m]]. Qed. - -Lemma trans : transitive (@ord T). -Proof. by case: T=>s [b [m]]. Qed. - -Lemma total x y : [|| ord x y, x == y | ord y x]. -Proof. by case: T x y=>s [b [m]]. Qed. - -Lemma nsym x y : ord x y -> ord y x -> False. -Proof. by move=>E1 E2; move: (trans E1 E2); rewrite irr. Qed. - -Lemma orefl x : oleq x x. -Proof. by rewrite /oleq eq_refl orbT. Qed. - -Lemma otrans : transitive (@oleq T). -Proof. -move=>x y z /=; case/orP; last by move/eqP=>->. -rewrite /oleq; move=>T1; case/orP; first by move/(trans T1)=>->. -by move/eqP=><-; rewrite T1. -Qed. - -Lemma sorted_oleq s : sorted (@ord T) s -> sorted (@oleq T) s. -Proof. by elim: s=>[|x s IH] //=; apply: sub_path=>z y; rewrite /oleq=>->. Qed. - -End Lemmas. - -Hint Resolve orefl : core. - -Section Totality. -Variable K : ordType. - -Variant total_spec (x y : K) : bool -> bool -> bool -> Type := -| total_spec_lt of ord x y : total_spec x y true false false -| total_spec_eq of x == y : total_spec x y false true false -| total_spec_gt of ord y x : total_spec x y false false true. - -Lemma totalP x y : total_spec x y (ord x y) (x == y) (ord y x). -Proof. -case H1: (x == y). -- by rewrite (eqP H1) irr; apply: total_spec_eq. -case H2: (ord x y); case H3: (ord y x). -- by case: (nsym H2 H3). -- by apply: total_spec_lt H2. -- by apply: total_spec_gt H3. -by move: (total x y); rewrite H1 H2 H3. -Qed. -End Totality. - - -(* Monotone (i.e. strictly increasing) functions for Ord Types *) -Section Mono. -Variables (A B :ordType). - -Definition strictly_increasing f x y := @ord A x y -> @ord B (f x) (f y). - -Structure mono : Type := Mono - {fun_of: A -> B; _: forall x y, strictly_increasing fun_of x y}. - -End Mono. -Arguments strictly_increasing {A B} f x y. -Arguments Mono {A B _} _. - -Section NatOrd. -Lemma irr_ltn_nat : irreflexive ltn. Proof. by move=>x; rewrite /= ltnn. Qed. -Lemma trans_ltn_nat : transitive ltn. Proof. by apply: ltn_trans. Qed. -Lemma total_ltn_nat x y : [|| x < y, x == y | y < x]. -Proof. by case: ltngtP. Qed. - -Definition nat_ordMixin := OrdMixin irr_ltn_nat trans_ltn_nat total_ltn_nat. -Canonical Structure nat_ordType := OrdType nat nat_ordMixin. -End NatOrd. - -Section ProdOrd. -Variables K T : ordType. - -(* lexicographic ordering *) -Definition lex : rel (K * T) := - fun x y => if x.1 == y.1 then ord x.2 y.2 else ord x.1 y.1. - -Lemma irr_lex : irreflexive lex. -Proof. by move=>x; rewrite /lex eq_refl irr. Qed. - -Lemma trans_lex : transitive lex. -Proof. -move=>[x1 x2][y1 y2][z1 z2]; rewrite /lex /=. -case: ifP=>H1; first by rewrite (eqP H1); case: eqP=>// _; apply: trans. -case: ifP=>H2; first by rewrite (eqP H2) in H1 *; rewrite H1. -case: ifP=>H3; last by apply: trans. -by rewrite (eqP H3)=>R1; move/(nsym R1). -Qed. - -Lemma total_lex : forall x y, [|| lex x y, x == y | lex y x]. -Proof. -move=>[x1 x2][y1 y2]; rewrite /lex /=. -case: ifP=>H1. -- rewrite (eqP H1) eq_refl -pair_eqE /= eq_refl /=; exact: total. -rewrite (eq_sym y1) -pair_eqE /= H1 /=. -by move: (total x1 y1); rewrite H1. -Qed. - -Definition prod_ordMixin := OrdMixin irr_lex trans_lex total_lex. -Canonical Structure prod_ordType := Eval hnf in OrdType (K * T) prod_ordMixin. -End ProdOrd. - -Section FinTypeOrd. -Variable T : finType. - -Definition ordf : rel T := - fun x y => index x (enum T) < index y (enum T). - -Lemma irr_ordf : irreflexive ordf. -Proof. by move=>x; rewrite /ordf ltnn. Qed. - -Lemma trans_ordf : transitive ordf. -Proof. by move=>x y z; rewrite /ordf; apply: ltn_trans. Qed. - -Lemma total_ordf x y : [|| ordf x y, x == y | ordf y x]. -Proof. -rewrite /ordf; case: ltngtP=>//= H; rewrite ?orbT ?orbF //. -have [H1 H2]: x \in enum T /\ y \in enum T by rewrite !mem_enum. -by rewrite -(nth_index x H1) -(nth_index x H2) H eq_refl. -Qed. - -Definition fin_ordMixin := OrdMixin irr_ordf trans_ordf total_ordf. -End FinTypeOrd. - -(* notation to let us write I_n instead of (ordinal_finType n) *) -Notation "[ 'fin_ordMixin' 'of' T ]" := - (fin_ordMixin _ : Ordered.mixin_of [eqType of T]) (at level 0). - -Definition ordinal_ordMixin n := [fin_ordMixin of 'I_n]. -Canonical Structure ordinal_ordType n := OrdType 'I_n (ordinal_ordMixin n). - -Section SeqOrd. -Variable (T : ordType). - -Fixpoint ords x : pred (seq T) := - fun y => match x , y with - | [::] , [::] => false - | [::] , t :: ts => true - | x :: xs , y :: ys => if x == y then ords xs ys - else ord x y - | _ :: _ , [::] => false - end. - -Lemma irr_ords : irreflexive ords. -Proof. by elim=>//= a l ->; rewrite irr; case:eqP=> //=. Qed. - -Lemma trans_ords : transitive ords. -Proof. -elim=>[|y ys IHy][|x xs][|z zs]//=. -case:eqP=>//[->|H0];case:eqP=>//H; first by move/IHy; apply. -- by case:eqP=>//; rewrite -H; first (by move/H0). -case:eqP=>//[->|H1] H2; first by move/(nsym H2). -by move/(trans H2). -Qed. - -Lemma total_ords : forall x y, [|| ords x y, x == y | ords y x]. -Proof. -elim=>[|x xs IH][|y ys]//=; case:eqP=>//[->|H1]; - (case:eqP=>//= H; first (by rewrite orbT //=)). -- by case:eqP=>//H3 ; case: (or3P (IH ys))=> [-> | /eqP H0 | ->]; - [ rewrite orTb // | apply: False_ind; apply: H; rewrite H0 | rewrite orbT //]. -case:eqP; first by move/(esym)/H1. -by move=>_ ;case: (or3P (total x y))=>[-> //| /eqP /H1 //| -> //]; rewrite orbT. -Qed. - -Definition seq_ordMixin := OrdMixin irr_ords trans_ords total_ords. -Canonical Structure seq_ordType := Eval hnf in OrdType (seq T) seq_ordMixin. -End SeqOrd. - -(* A trivial total ordering for Unit *) -Section unitOrd. -Let ordtt (x y : unit ) := false. - -Lemma irr_tt : irreflexive ordtt. -Proof. by []. Qed. - -Lemma trans_tt : transitive ordtt. -Proof. by []. Qed. - -Lemma total_tt x y : [|| ordtt x y, x == y | ordtt y x ]. -Proof. by []. Qed. - -Let unit_ordMixin := OrdMixin irr_tt trans_tt total_tt. -Canonical Structure unit_ordType := Eval hnf in OrdType unit unit_ordMixin. -End unitOrd. - - -(* ordering with path, seq and last *) - -Lemma seq_last_in (A : eqType) (s : seq A) x : - last x s \notin s -> s = [::]. -Proof. by case: s=> // h s; rewrite last_cons mem_last. Qed. - -Lemma path_last (A : ordType) (s : seq A) x : - path oleq x s -> oleq x (last x s). -Proof. -move/(order_path_min (@otrans _)); rewrite -nth_last. -by case: s =>// h s' /all_nthP->. -Qed. - -(* in a sorted list, the last element is maximal *) -(* and the maximal element is last *) - -Lemma sorted_last_key_max (A : ordType) (s : seq A) x y : - sorted oleq s -> x \in s -> oleq x (last y s). -Proof. -elim: s x y=>[|z s IH] //= x y H; rewrite inE /=. -case: eqP=>[->|] /= _; first by apply: path_last. -by apply: IH (path_sorted H). -Qed. - -Lemma sorted_max_key_last (A : ordType) (s : seq A) x y : - sorted oleq s -> x \in s -> - (forall z, z \in s -> oleq z x) -> last y s = x. -Proof. -elim: s x y => [|w s IH] //= x y; rewrite inE /=. -case: eqP=>[<- /= H1 _ H2 | _ H /= H1 H2]; last first. -- apply: IH (path_sorted H) H1 _ => z H3; apply: H2. - by rewrite inE /= H3 orbT. -apply/eqP; move: (H2 (last x s)) (path_last H1); rewrite inE /= /oleq eq_sym. -case: totalP=>//=; case E: (last x s \in s)=>//. -by move/negbT/seq_last_in: E=>->; rewrite irr. -Qed. - -Lemma seq_last_mono (A : ordType) (s1 s2 : seq A) x : - path oleq x s1 -> path oleq x s2 -> - {subset s1 <= s2} -> - oleq (last x s1) (last x s2). -Proof. -case: s1=>/= [_ H1 _|a s]; first by apply: path_last H1. -case/andP=>H1 H2 H3 H; apply: sorted_last_key_max (path_sorted H3) _. -apply: {x s2 H1 H3} H; rewrite inE orbC -implyNb. -by case E: (_ \notin _) (@seq_last_in _ s a)=>//= ->. -Qed. - - - - - diff --git a/options.v b/options.v new file mode 100644 index 0000000..f41c1c8 --- /dev/null +++ b/options.v @@ -0,0 +1,4 @@ +Export Set Implicit Arguments. +Export Unset Strict Implicit. +Export Unset Printing Implicit Defensive. +Export Unset Program Cases. diff --git a/pcm/automap.v b/pcm/automap.v index 36c5ccf..1ec2286 100644 --- a/pcm/automap.v +++ b/pcm/automap.v @@ -13,38 +13,33 @@ limitations under the License. From Coq Require Import ssreflect ssrbool ssrfun. From mathcomp Require Import ssrnat eqtype seq. -From fcsl Require Import pred ordtype pcm unionmap. -Set Implicit Arguments. -Unset Strict Implicit. -Unset Printing Implicit Defensive. - -(******************************************************************************) -(* This file implementations of Canonical structure lemmas (a.k.a. overloaded *) -(* lemmas) for automating three tasks: *) -(* *) -(* 1. checking if implications of the form valid e1 -> valid e2 hold by *) -(* deciding if the terms in e2 are all contained in e1 *) -(* *) -(* 2. checking if dom_eq e1 e2 holds by cancelling the common terms, to *) -(* obtain residuals rs1 and rs2, and then issuing a subgoal dom_eq rs1 *) -(* rs2. *) -(* *) -(* 3. checking if the union e is undef, because it contains duplicate *) -(* pointers or an undef *) -(******************************************************************************) - -(* For each task, we have two implementations: a naive and a *) -(* sophisticated one. The lemmas validO, domeqO, invalidO are the naive *) -(* ones, and validX, domeqX, invalidX are the sophisticated ones. I keep *) -(* both O/X versions for now, for experimentation purposes, but *) -(* eventually should retain only validX and domeqX. *) - -(* First, two general helper concepts for searching in sequences. They will *) -(* be useful in syntactifying the expressions e1 and e2 for both tasks. *) -(* The concepts are: *) -(* *) -(* - onth: returns None to signal that an element is not found prefix: *) -(* - will be used for growing interpretation contexts *) +From fcsl Require Import pred pcm unionmap. +From fcsl Require Import options. + +(**************************************************************************) +(**************************************************************************) +(* Canonical structure lemmas for automating three tasks: *) +(* *) +(* 1. checking if implications of the form valid e1 -> valid e2 hold by *) +(* deciding if the terms in e2 are all contained in e1 *) +(* *) +(* 2. checking if dom_eq e1 e2 holds by cancelling the common terms, to *) +(* obtain residuals rs1 and rs2, and then issuing a subgoal dom_eq rs1 *) +(* rs2. *) +(* *) +(* 3. checking if the union e is undef, because it contains duplicate *) +(* pointers or an undef *) +(* *) +(**************************************************************************) +(**************************************************************************) + + +(* First, two general helper concepts for searching in sequences. They will be *) +(* useful in syntactifying the expressions e1 and e2 for both tasks. The *) +(* contepts are: *) +(* *) +(* - onth: returns None to signal that an element is not found *) +(* - prefix: will be used for growing interpretation contexts *) Section Helpers. Variable A : Type. @@ -100,15 +95,15 @@ Proof. by elim: s n=>//= a s IH [[->]|n /IH]; rewrite inE ?eq_refl // orbC => ->. Qed. -(* Context structure for reflection of unionmap expressions. We *) -(* reflect the keys and the variables of the map expression. (The *) +(* Context structure for reflection of unionmap expressions. We *) +(* reflect the keys and the variables of the map expression. (The *) (* variables are all expressions that are not recognized as a key, or as *) -(* a disjoint union). We reflect disjoint union as a sequence. *) +(* a disjoint union). We reflect disjoint union as a sequence. *) (* *) (* The context of keys is thus seq K. The context of vars is seq U. *) Section ReflectionContexts. -Variables (K : ordType) (T : Type) (U : union_map_class K T). +Variables (K : ordType) (C : pred K) (T : Type) (U : union_map_class C T). Structure ctx := Context {keyx : seq K; varx : seq U}. @@ -167,7 +162,7 @@ End OneShotFilter. (* now for reflection *) Section Reflection. -Variables (K : ordType) (T : Type) (U : union_map_class K T). +Variables (K : ordType) (C : pred K) (T : Type) (U : union_map_class C T). Implicit Type i : ctx U. Inductive term := Pts of nat & T | Var of nat. @@ -175,8 +170,8 @@ Inductive term := Pts of nat & T | Var of nat. (* interpretation function for elements *) Definition interp' i t := match t with - Pts n v => if onth (keyx i) n is Some k then pts k v else um_undef - | Var n => if onth (varx i) n is Some f then f else um_undef + Pts n v => if onth (keyx i) n is Some k then pts k v else undef + | Var n => if onth (varx i) n is Some f then f else undef end. (* main interpretation function *) @@ -280,18 +275,18 @@ Proof. by elim: ts=>//= t ts IH; case: t=>[m v|m] //; rewrite inE IH. Qed. End Reflection. -(**************************************************************************) -(**************************************************************************) -(* Purely functional decision procedures for the three tasks. Further *) -(* below, they will be embedded into the canonical programs validO/validX *) -(* and domeqO/domeqX and invalidO/invalidX respectively. *) -(**************************************************************************) -(**************************************************************************) +(**********************************************************************) +(**********************************************************************) +(* Purely functional decision procedures for the three tasks. Further *) +(* below, they will be embedded into the canonical programs validX *) +(* and domeqX and invalidX respectively. *) +(**********************************************************************) +(**********************************************************************) -(* subterm is purely functional version of validO/validX *) +(* subterm is purely functional version of validX *) Section Subterm. -Variables (K : ordType) (T : Type) (U : union_map_class K T). +Variables (K : ordType) (C : pred K) (T : Type) (U : union_map_class C T). Implicit Types (i : ctx U) (ts : seq (term T)). Fixpoint subterm ts1 ts2 := @@ -320,10 +315,10 @@ Qed. End Subterm. -(* subtract is purely functional version of domeqO/domeqX *) +(* subtract is purely functional version of domeqX *) Section Subtract. -Variables (K : ordType) (T : Type) (U : union_map_class K T). +Variables (K : ordType) (C : pred K) (T : Type) (U : union_map_class C T). Implicit Types (i : ctx U) (ts : seq (term T)). (* We need a subterm lemma that returns the uncancelled stuff from *) @@ -361,10 +356,10 @@ Qed. End Subtract. -(* invalid is a purely functional test of invalidO/invalidX *) +(* invalid is a purely functional test of invalidX *) Section Invalid. -Variables (K : ordType) (T : Type) (U : union_map_class K T). +Variables (K : ordType) (C : pred K) (T : Type) (U : union_map_class C T). Implicit Types (i : ctx U) (t : term T) (ts : seq (term T)). Definition undefx i t := @@ -383,7 +378,7 @@ case: t W=>[n v|n] /= /size_onth [k] X; rewrite /isundef /= X; last first. by case/orP: H V=>// /undefbE ->; rewrite valid_undef. rewrite negb_and negbK has_getkeys -orbA /=. case/orP=>// V; last by rewrite validPtUn andbCA (negbTE (IH A _)). -by case: (keyP V X)=>u ->; rewrite joinA pts_undef join_undefL valid_undef. +by case: (keyP V X)=>u ->; rewrite joinA pts_undef2 undef_join valid_undef. Qed. End Invalid. @@ -402,8 +397,6 @@ Variable A : Type. Structure tagged_elem := XTag {xuntag :> A}. -(* Local Coercion untag : tagged_elem >-> A. *) - Definition extend_tag := XTag. Definition recurse_tag := extend_tag. Canonical found_tag x := recurse_tag x. @@ -437,18 +430,18 @@ Canonical extend_form x := Form (@extend_pf x). End XFind. -(* Next, we syntactify a unionmap into a seq term as follows. *) -(* *) -(* - if the map is f1 \+ f2, then recurse over both and concat. the results *) -(* - if the map is the empty map, return [::] *) -(* - if the map is k \-> v then add k to the context, and return [Pts x v], *) -(* where x is the index for l in the context *) -(* if the map is whatever else, add the map to the context and return *) -(* [Var n], where n is the index for the map in the context *) +(* Next, we syntactify a unionmap into a seq term as follows. *) +(* *) +(* - if the map is f1 \+ f2, then recurse over both and concatenate the results *) +(* - if the map is the empty map, return [::] *) +(* - if the map is k \-> v then add k to the context, and return [Pts x v], *) +(* where x is the index for l in the context *) +(* if the map is whatever else, add the map to the context and return *) +(* [Var n], where n is the index for the map in the context *) Module Syntactify. Section Syntactify. -Variables (K : ordType) (T : Type) (U : union_map_class K T). +Variables (K : ordType) (C : pred K) (T : Type) (U : union_map_class C T). Implicit Types (i : ctx U) (ts : seq (term T)). (* a tagging structure to control the flow of computation *) @@ -528,45 +521,16 @@ End Syntactify. Export Syntactify.Exports. -(****************************) -(* Automating validO/validX *) -(****************************) - -(* validO is a naive lemma that automates subterm checking *) -(* it leaves a goal of the form subterm ts2 ts1 *) -(* this one evaluates to a boolean, so it's easy to discharge *) -(* but it takes the space for the lemma's argument *) -(* which, thus, can't be picked from the goal *) +(*********************) +(* Automating validX *) +(*********************) -Section ValidO. -Variables (K : ordType) (T : Type) (U : union_map_class K T). -Implicit Types (i : ctx U) (ts : seq (term T)). -Notation form := Syntactify.form. -Notation untag := Syntactify.untag. - -Lemma validO j k ts1 ts2 (f1 : form (@empx K T U) j ts1) - (f2 : form j k ts2) : - valid (untag f1) -> subterm ts2 ts1 -> valid (untag f2). -Proof. -case: f1 f2=>f1 [<- _ A1][f2][<- S2 A2] /= V; rewrite (sc_interp S2 A1) in V. -by case/(subterm_sound A2 (sc_wf S2 A1))=>xs /domeqP []; rewrite V=>/validL ->. -Qed. - -End ValidO. - -Arguments validO [K T U j k ts1 ts2 f1 f2] _ _. - -Example ex0 (x y z : nat) (v1 v2 : nat) h: - valid (Unit \+ y \\-> v1 \+ h \+ x \\-> v1) -> - valid (x \\-> v2 \+ Unit). -Proof. by move=>V; rewrite (validO V). Abort. - -(* validX is a more refined lemma for subterm checking which *) +(* validX is a refined lemma for subterm checking which *) (* automatically discharges the spurious argument from above *) Module ValidX. Section ValidX. -Variables (K : ordType) (T : Type) (U : union_map_class K T). +Variables (K : ordType) (C : pred K) (T : Type) (U : union_map_class C T). Implicit Types (j : ctx U) (ts : seq (term T)). Notation form := Syntactify.form. Notation untag := Syntactify.untag. @@ -607,12 +571,33 @@ Canonical start j k ts1 ts2 f2 := RForm (@start_pf j k ts1 ts2 f2). End ValidX. +(* Wrappers for automated versions of joinKx(xK), cancPL(PR) lemmas *) +Section WrappersForCancellationLemmas. +Variable U : cpcm. + +Lemma joinKx' (x1 x2 x : U) : x1 \+ x = x2 \+ x -> valid (x1 \+ x) -> x1 = x2. +Proof. by move=>E V; apply: joinKx V E. Qed. + +Lemma joinxK' (x x1 x2 : U) : x \+ x1 = x \+ x2 -> valid (x \+ x1) -> x1 = x2. +Proof. by move=>E V; apply: joinxK V E. Qed. + +Lemma cancPL' (P : U -> Prop) (s1 s2 t1 t2 : U) : + precise P -> s1 \+ t1 = s2 \+ t2 -> P s1 -> P s2 -> valid (s1 \+ t1) -> + (s1 = s2) * (t1 = t2). +Proof. by move=>H E H1 H2 V; apply: cancPL H V H1 H2 E. Qed. + +Lemma cancPR' (P : U -> Prop) (s1 s2 t1 t2 : U) : + precise P -> s1 \+ t1 = s2 \+ t2 -> P t1 -> P t2 -> valid (s1 \+ t1) -> + (s1 = s2) * (t1 = t2). +Proof. by move=>H E H1 H2 V; apply: cancPR H V H1 H2 E. Qed. +End WrappersForCancellationLemmas. + Module Exports. Canonical equate. Canonical start. Section Exports. -Variables (K : ordType) (T : Type) (U : union_map_class K T). +Variables (K : ordType) (C : pred K) (T : Type) (U : union_map_class C T). Implicit Types (j : ctx U) (ts : seq (term T)). Notation form := Syntactify.form. Notation untag := Syntactify.untag. @@ -625,53 +610,31 @@ Proof. by case: g f1; case=>pivot H [f1][<- Sc A] /(H A); apply. Qed. End Exports. -Arguments validX [K T U m j ts1 f1 g] _. +Arguments validX [K C T U m j ts1 f1 g] _. Example ex0 (x y z : nat) (v1 v2 : nat) h: valid (Unit \+ y \\-> v1 \+ h \+ x \\-> v1) -> valid (x \\-> v2 \+ Unit). Proof. apply: validX. Abort. +(* Automated versions of joinKx(xK), cancPL(PR) lemmas *) +Notation joinKX V E := (joinKx' E (validX V)). +Notation joinXK V E := (joinxK' E (validX V)). +Notation cancPLX pf V H1 H2 E := (cancPL' pf E H1 H2 (validX V)). +Notation cancPRX pf V H1 H2 E := (cancPR' pf E H1 H2 (validX V)). + End Exports. End ValidX. Export ValidX.Exports. -(****************************) -(* Automating domeqO/domeqX *) -(****************************) - -(* naive lemma first *) - -Section DomeqO. -Variables (K : ordType) (T : Type) (U : union_map_class K T). -Implicit Types (j : ctx U) (ts : seq (term T)). -Notation form := Syntactify.form. -Notation untag := Syntactify.untag. - -Lemma domeqO j k rs1 rs2 ts1 ts2 - (f1 : form (empx U) j ts1) (f2 : form j k ts2) : - subtract ts1 ts2 [::] = (rs1, rs2) -> - dom_eq (pprint k (rev rs1)) (pprint k rs2) -> - dom_eq (untag f1) (untag f2). -Proof. -case: f1 f2=>f1 [<- _ A1][f2][<- S A2]. -case/(subtract_sound (sc_wf S A1) A2)=>// ys [/= D1 D2]. -rewrite unitR in D1; rewrite (sc_interp S A1). -rewrite !pp_interp interp_rev => D; apply: domeq_trans D1 _. -rewrite domeq_sym; apply: domeq_trans D2 _. -by rewrite domeq_sym; apply: domeqUn. -Qed. - -End DomeqO. - -Example ex0 (x y z : nat) (v1 v2 : nat) h: - dom_eq (Unit \+ y \\-> v1 \+ h \+ x \\-> v1) (x \\-> v2 \+ Unit). -Proof. apply: domeqO=>//=. Abort. +(*********************) +(* Automating domeqX *) +(*********************) Module DomeqX. Section DomeqX. -Variables (K : ordType) (T : Type) (U : union_map_class K T). +Variables (K : ordType) (C : pred K) (T : Type) (U : union_map_class C T). Implicit Types (j : ctx U) (ts : seq (term T)). Notation form := Syntactify.form. Notation untag := Syntactify.untag. @@ -710,7 +673,7 @@ Canonical equate. Canonical start. Section Exports. -Variables (K : ordType) (T : Type) (U : union_map_class K T). +Variables (K : ordType) (C : pred K) (T : Type) (U : union_map_class C T). Implicit Types (j : ctx U) (ts : seq (term T)). Notation form := Syntactify.form. Notation untag := Syntactify.untag. @@ -728,7 +691,7 @@ Qed. End Exports. -Arguments domeqX [K T U m j k rs1 rs2 ts1 f1 g] _. +Arguments domeqX [K C T U m j k rs1 rs2 ts1 f1 g] _. Example ex0 (x y z : nat) (v1 v2 : nat) h: dom_eq (Unit \+ y \\-> v1 \+ h \+ x \\-> v1) (x \\-> v2 \+ Unit). @@ -740,37 +703,13 @@ End DomeqX. Export DomeqX.Exports. -(********************************) -(* Automating invalidO/invalidX *) -(********************************) - -(* naive lemma first *) - -Section InvalidO. -Variables (K : ordType) (T : Type) (U : union_map_class K T). -Implicit Types (i : ctx U) (ts : seq (term T)). -Notation form := Syntactify.form. -Notation untag := Syntactify.untag. - -Lemma undefO i ts (f : form (empx U) i ts) : - isundef i ts -> untag f = um_undef. -Proof. by case: f=>f [<- _ A] /(isundef_sound A)/invalidE. Qed. - -Lemma invalidO i ts (f : form (empx U) i ts) : - isundef i ts -> valid (untag f) = false. -Proof. by move/undefO=>->; rewrite valid_undef. Qed. - -End InvalidO. - -Example ex0 (x y z : nat) (v1 v2 : nat) h: - (Unit \+ y \\-> v1 \+ h \+ y \\-> v1) = um_undef. -Proof. by apply: undefO. Abort. - -(* now the sophisticated one *) +(***********************) +(* Automating invalidX *) +(***********************) Module InvalidX. Section InvalidX. -Variables (K : ordType) (T : Type) (U : union_map_class K T). +Variables (K : ordType) (C : pred K) (T : Type) (U : union_map_class C T). Implicit Types (i : ctx U) (ts : seq (term T)). Notation form := Syntactify.form. Notation untag := Syntactify.untag. @@ -801,14 +740,14 @@ Canonical equate. Canonical start. Section Exports. -Variables (K : ordType) (T : Type) (U : union_map_class K T). +Variables (K : ordType) (C : pred K) (T : Type) (U : union_map_class C T). Implicit Types (i : ctx U) (ts : seq (term T)). Notation form := Syntactify.form. Notation untag := Syntactify.untag. (* the main lemmas *) -Lemma undefX m i ts (g : rform i ts m true) : unpack (pivot g) = um_undef. +Lemma undefX m i ts (g : rform i ts m true) : unpack (pivot g) = undef. Proof. by case: g; case=>pivot /= /(_ (erefl _))/negbT/invalidE. Qed. Lemma invalidX m i ts (g : rform i ts m true) : @@ -817,7 +756,7 @@ Proof. by rewrite undefX valid_undef. Qed. End Exports. -Arguments invalidX {K T U m i ts g}. +Arguments invalidX {K C T U m i ts g}. Example ex0 (x y z : nat) (v1 v2 : nat) h: valid (Unit \+ y \\-> v1 \+ h \+ y \\-> v1). @@ -827,8 +766,3 @@ End Exports. End InvalidX. Export InvalidX.Exports. - - - - - diff --git a/pcm/heap.v b/pcm/heap.v index 21e6a79..af7c33e 100644 --- a/pcm/heap.v +++ b/pcm/heap.v @@ -19,10 +19,8 @@ limitations under the License. From Coq Require Import ssreflect ssrbool ssrfun Eqdep. From mathcomp Require Import ssrnat eqtype seq path. -From fcsl Require Import axioms ordtype pcm finmap unionmap. -Set Implicit Arguments. -Unset Strict Implicit. -Unset Printing Implicit Defensive. +From fcsl Require Import axioms pcm finmap unionmap. +From fcsl Require Import options. (*************) (* Locations *) @@ -41,14 +39,14 @@ Lemma eq_ptrP : Equality.axiom eq_ptr. Proof. by case=>x [y] /=; case: eqP=>[->|*]; constructor=>//; case. Qed. Definition ptr_eqMixin := EqMixin eq_ptrP. -Canonical ptr_eqType := EqType ptr ptr_eqMixin. +Canonical ptr_eqType := Eval hnf in EqType ptr ptr_eqMixin. (* some pointer arithmetic: offsetting from a base *) Definition ptr_offset x i := ptr_nat (nat_ptr x + i). Notation "x .+ i" := (ptr_offset x i) - (at level 3, format "x .+ i"). + (at level 5, format "x .+ i"). Lemma ptrE x y : (x == y) = (nat_ptr x == nat_ptr y). Proof. by move: x y=>[x][y]. Qed. @@ -81,11 +79,11 @@ Proof. by case=>x /=; rewrite ltnn. Qed. Lemma ltn_ptr_trans : transitive ltn_ptr. Proof. by case=>x [y][z]; apply: ltn_trans. Qed. -Lemma ltn_ptr_total : forall x y : ptr, [|| ltn_ptr x y, x == y | ltn_ptr y x]. +Lemma ltn_ptr_semiconn : forall x y : ptr, x != y -> ltn_ptr x y || ltn_ptr y x. Proof. by case=>x [y]; rewrite ptrE /=; case: ltngtP. Qed. -Definition ptr_ordMixin := OrdMixin ltn_ptr_irr ltn_ptr_trans ltn_ptr_total. -Canonical ptr_ordType := OrdType ptr ptr_ordMixin. +Definition ptr_ordMixin := OrdMixin ltn_ptr_irr ltn_ptr_trans ltn_ptr_semiconn. +Canonical ptr_ordType := Eval hnf in OrdType ptr ptr_ordMixin. (*********) (* Heaps *) @@ -133,7 +131,7 @@ Proof. split; first by move=>->; case: h2. case: h1; case: h2=>// f1 pf1 f2 pf2 E. rewrite {f2}E in pf1 pf2 *. -by congr Def; apply: bool_irrelevance. +by congr Def; apply: eq_irrelevance. Qed. End NullLemmas. @@ -162,7 +160,10 @@ Definition dom_eq h1 h2 := | _, _ => false end. -Definition free x h := +Definition assocs h : seq (ptr_ordType * dynamic id) := + if h is Def f _ then seq_of f else [::]. + +Definition free h x := if h is Def hs ns then Def (free_nullP x ns) else Undef. Definition find (x : ptr) h := @@ -175,9 +176,6 @@ Definition union h1 h2 := else Undef else Undef. -Definition um_filter q f := - if f is Def fs pf then Def (@filt_nullP fs q pf) else Undef. - Definition pts (x : ptr) v := upd x v empty. Definition empb h := if h is Def hs _ then supp hs == [::] else false. @@ -188,7 +186,7 @@ Definition keys_of h : seq ptr := if h is Def f _ then supp f else [::]. Local Notation base := - (@UM.base ptr_ordType (dynamic id) (fun k => k != null)). + (@UM.base ptr_ordType (fun k => k != null) (dynamic id)). Definition from (f : heap) : base := if f is Def hs ns then UM.Def (heap_base ns) else UM.Undef _ _. @@ -220,7 +218,10 @@ Proof. by case: f. Qed. Lemma dom_eqE f1 f2 : dom_eq f1 f2 = UM.dom_eq (from f1) (from f2). Proof. by case: f1 f2=>[|f1 H1][|f2 H2]. Qed. -Lemma freeE k f : free k f = to (UM.free k (from f)). +Lemma assocsE f : assocs f = UM.assocs (from f). +Proof. by case: f. Qed. + +Lemma freeE f k : free f k = to (UM.free (from f) k). Proof. by case: f=>[|f H] //; rewrite heapE. Qed. Lemma findE k f : find k f = UM.find k (from f). @@ -232,9 +233,6 @@ case: f1 f2=>[|f1 H1][|f2 H2] //; rewrite /union /UM.union /=. by case: ifP=>D //; rewrite heapE. Qed. -Lemma umfiltE q f : um_filter q f = to (UM.um_filter q (from f)). -Proof. by case: f=>[|f H] //; rewrite heapE. Qed. - Lemma empbE f : empb f = UM.empb (from f). Proof. by case: f. Qed. @@ -251,8 +249,8 @@ Module Exports. (* the inheritance from PCMs *) Definition heapUMCMix := - UMCMixin ftE tfE defE undefE emptyE updE domE dom_eqE freeE - findE unionE umfiltE empbE undefbE ptsE. + UMCMixin ftE tfE defE undefE emptyE updE domE dom_eqE assocsE freeE + findE unionE empbE undefbE ptsE. Canonical heapUMC := Eval hnf in UMC heap heapUMCMix. Definition heapPCMMix := union_map_classPCMMix heapUMC. @@ -273,7 +271,7 @@ Export Heap.Exports. Notation heap := Heap.heap. Definition heap_pts A (x : ptr) (v : A) := - @UMC.pts _ _ heapUMC x (idyn v). + @UMC.pts _ _ _ heapUMC x (idyn v). Notation "x :-> v" := (@heap_pts _ x v) (at level 30). @@ -302,11 +300,11 @@ Proof. by move=>V /(cancelPt2 V) [->] /dyn_inj ->. Qed. Lemma heap_eta x h : x \in dom h -> exists A (v : A), - find x h = Some (idyn v) /\ h = x :-> v \+ free x h. + find x h = Some (idyn v) /\ h = x :-> v \+ free h x. Proof. by case/um_eta; case=>A v H; exists A, v. Qed. Lemma heap_eta2 A x h (v : A) : - find x h = Some (idyn v) -> h = x :-> v \+ free x h. + find x h = Some (idyn v) -> h = x :-> v \+ free h x. Proof. move=>E; case: (heap_eta (find_some E))=>B [w][]. rewrite {}E; case=>E; rewrite -E in w *. @@ -327,19 +325,26 @@ Lemma hcancel2V A x1 x2 (v1 v2 : A) h1 h2 : valid (x1 :-> v1 \+ h1) -> x1 :-> v1 \+ h1 = x2 :-> v2 \+ h2 -> if x1 == x2 then v1 = v2 /\ h1 = h2 - else [/\ free x1 h2 = free x2 h1, - h1 = x2 :-> v2 \+ free x1 h2 & - h2 = x1 :-> v1 \+ free x2 h1]. + else [/\ free h2 x1 = free h1 x2, + h1 = x2 :-> v2 \+ free h2 x1 & + h2 = x1 :-> v1 \+ free h1 x2]. Proof. by move=>V /(cancel2 V); case: ifP=>// _ [/dyn_inj]. Qed. +Lemma heap_ind' (P : heap -> Prop) : + P undef -> P Unit -> + (forall A x (v : A) f, P f -> valid (x :-> v \+ f) -> P (x :-> v \+ f)) -> + forall f, P f. +Proof. +move=>H1 H2 H3; apply: um_ind' H1 H2 _. +by move=>/= k [A v] f; apply: H3. +Qed. + End HeapPointsToLemmas. Prenex Implicits heap_eta2. (******************************************) (* additional stuff, on top of union maps *) -(* mostly random stuff, kept for legacy *) -(* should be removed/redone eventually *) (******************************************) Definition fresh (h : heap) := @@ -362,19 +367,9 @@ elim=>[|y s IH x] /=; first by case=>x; rewrite /ord /= addn1. by case/andP=>H1; move/IH; apply: trans H1. Qed. -Lemma path_filter (A : ordType) (s : seq A) (p : pred A) x : - path ord x s -> path ord x (filter p s). -Proof. -elim: s x=>[|y s IH] x //=. -case/andP=>H1 H2. -case: ifP=>E; first by rewrite /= H1 IH. -apply: IH; elim: s H2=>[|z s IH] //=. -by case/andP=>H2 H3; rewrite (@trans _ y). -Qed. - Lemma dom_fresh h n : (fresh h).+n \notin dom h. Proof. -suff L2: forall h x, x \in dom h -> ord x (fresh h). +suff L2: forall (h : heap) x, x \in dom h -> ord x (fresh h). - by apply: (contra (L2 _ _)); rewrite -leqNgt leq_addr. case=>[|[s H1]] //; rewrite /supp => /= H2 x. rewrite /dom /fresh /supp /=. @@ -395,15 +390,15 @@ Hint Resolve dom_fresh fresh_null : core. (* pick *) (********) -Lemma emp_pick (h : heap) : (pick h == null) = (~~ valid h || empb h). +Lemma unit_pick (h : heap) : (pick h == null) = (~~ valid h || unitb h). Proof. -rewrite /empb; case: h=>[|h] //=; case: (supp h)=>[|x xs] //=. +rewrite /unitb /= /empb !pcmE; case: h=>[|h] //=; case: (supp h)=>[|x xs] //=. by rewrite inE negb_or eq_sym; case/andP; move/negbTE=>->. Qed. -Lemma pickP h : valid h && ~~ empb h = (pick h \in dom h). +Lemma pickP h : valid h && ~~ unitb h = (pick h \in dom h). Proof. -rewrite /dom /empb; case: h=>[|h] //=. +rewrite /dom /unitb /= /empb pcmE; case: h=>[|h] //=. by case: (supp h)=>// *; rewrite inE eq_refl. Qed. @@ -455,14 +450,14 @@ Proof. by move=>->; rewrite updi_cat. Qed. (* helper lemma *) Lemma updiVm' x m xs : m > 0 -> x \notin dom (updi x.+m xs). Proof. -elim: xs x m=>[|v vs IH] x m //= H. -rewrite ptrA domPtUn inE /= negb_and negb_or -{4}(ptr0 x) ptrK -lt0n H /=. +elim: xs x m=>[|v vs IH] x m H; first by rewrite dom0. +rewrite /= ptrA domPtUn inE /= negb_and negb_or -{4}(ptr0 x) ptrK -lt0n H /=. by rewrite orbC IH // addn1. Qed. Lemma updiD x xs : valid (updi x xs) = (x != null) || (size xs == 0). Proof. -elim: xs x=>[|v xs IH] x //=; first by rewrite orbC. +elim: xs x=>[|v xs IH] x; first by rewrite valid_unit orbC. by rewrite validPtUn updiVm' // orbF IH ptr_null andbF andbC. Qed. @@ -471,7 +466,7 @@ Lemma updiVm x m xs : Proof. case: m=>[|m] /=; last first. - by rewrite andbF; apply/negbTE/updiVm'. -case: xs=>[|v xs]; rewrite ptr0 ?andbF ?andbT //=. +case: xs=>[|v xs]; rewrite ptr0 ?andbF ?andbT ?dom0 //=. by rewrite domPtUn inE /= eq_refl -updiS updiD orbF andbT /=. Qed. @@ -479,9 +474,10 @@ Lemma updimV x m xs : x.+m \in dom (updi x xs) = (x != null) && (m < size xs). Proof. case H: (x == null)=>/=. -- by case: xs=>// a s; rewrite (eqP H). -elim: xs x m H=>[|v vs IH] x m H //; case: m=>[|m]. -- by rewrite ptr0 /= domPtUn inE /= eq_refl andbT -updiS updiD H. +- case: xs=>[|a s]; first by rewrite dom0. + by rewrite domPtUn inE validPtUn /= H. +elim: xs x m H=>[|v vs IH] x m H //; case: m=>[|m]; try by rewrite /= dom0. +-by rewrite ptr0 domPtUn inE /= eq_refl andbT -updiS updiD H. rewrite -addn1 addnC -ptrA updiS domPtUn inE /= IH; last first. - by rewrite ptrE /= addn1. by rewrite -updiS updiD H /= -{1}(ptr0 x) ptrA ptrK. @@ -492,7 +488,9 @@ Lemma updiP x y xs : (x \in dom (updi y xs)). Proof. case H: (y == null)=>/=. -- by rewrite (eqP H); elim: xs=>[|z xs IH] //=; constructor; case. +- rewrite (eqP H); elim: xs=>[|z xs IH]. + - by rewrite dom0; constructor; case. + by rewrite domPtUn inE validPtUn; constructor; case. case E: (x \in _); constructor; last first. - by move=>[_][m][H1] H2; rewrite H1 updimV H2 H in E. case: (ptrT x y) E=>m; case/orP; move/eqP=>->. @@ -506,8 +504,8 @@ Lemma updi_inv x xs1 xs2 : valid (updi x xs1) -> updi x xs1 = updi x xs2 -> xs1 = xs2. Proof. elim: xs1 x xs2 =>[|v1 xs1 IH] x /=; case=>[|v2 xs2] //= D; -[move/esym| |]; try by rewrite empbE empbUn empbPt. -by move=> {D}/(hcancelV D) [<- D /(IH _ _ D) <-]. +[move/esym| |]; try by rewrite unitbE um_unitbUn um_unitbPt. +by case/(hcancelV D)=><- {}D /(IH _ _ D) <-. Qed. Lemma updi_iinv x xs1 xs2 h1 h2 : @@ -516,7 +514,7 @@ Lemma updi_iinv x xs1 xs2 h1 h2 : Proof. elim: xs1 x xs2 h1 h2=>[|v1 xs1 IH] x /=; case=>[|v2 xs2] //= h1 h2. - by rewrite !unitL. -move=>[E]; rewrite -!joinA=> D {D}/(hcancelV D)[<- D]. +move=>[E]; rewrite -!joinA=>D; case/(hcancelV D)=><-{}D. by case/(IH _ _ _ _ E D)=>->->. Qed. @@ -577,7 +575,7 @@ Lemma defPtUnO A h x (v : A) (f : partition (x :-> v) h) : valid (untag f) = [&& x != null, valid h & x \notin dom h]. Proof. by rewrite partitionE validPtUn. Qed. -Arguments defPtUnO [A][h] x {v f}. +Arguments defPtUnO [A h] x {v f}. Lemma defPt_nullO A h x (v : A) (f : partition (x :-> v) h) : valid (untag f) -> x != null. @@ -587,7 +585,7 @@ Arguments defPt_nullO [A h x v f] _. Lemma defPt_defO A h x (v : A) (f : partition (x :-> v) h) : valid (untag f) -> valid h. -Proof. by rewrite partitionE; apply: validPtUnV. Qed. +Proof. by rewrite partitionE => /validR. Qed. Arguments defPt_defO [A][h] x [v][f] _. @@ -611,7 +609,7 @@ Proof. by rewrite partitionE; apply: findPtUn. Qed. Arguments lookPtUnO [A h x v f] _. Lemma freePtUnO A h x (v : A) (f : partition (x :-> v) h) : - valid (untag f) -> free x (untag f) = h. + valid (untag f) -> free (untag f) x = h. Proof. by rewrite partitionE; apply: freePtUn. Qed. Arguments freePtUnO [A h x v f] _. @@ -763,4 +761,3 @@ by move=> eq_r12; apply/ffunP=> x; rewrite -[x]enum_rankK -!tnth_fgraph eq_r12. Qed. End Array. - diff --git a/pcm/invertible.v b/pcm/invertible.v new file mode 100644 index 0000000..138fecb --- /dev/null +++ b/pcm/invertible.v @@ -0,0 +1,125 @@ +(* +Copyright 2013 IMDEA Software Institute +Licensed under the Apache License, Version 2.0 (the "License"); +you may not use this file except in compliance with the License. +You may obtain a copy of the License at + http://www.apache.org/licenses/LICENSE-2.0 +Unless required by applicable law or agreed to in writing, software +distributed under the License is distributed on an "AS IS" BASIS, +WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. +See the License for the specific language governing permissions and +limitations under the License. +*) + +(******************************************************************************) +(* This file contains an implementation of invertible PCM morphisms and their *) +(* separating relations. *) +(******************************************************************************) + +From Coq Require Import ssreflect ssrbool ssrfun. +From mathcomp Require Import eqtype. +From fcsl Require Import axioms prelude pcm morphism. +From fcsl Require Import options. + +Section Invertible. + Variables (U V : pcm) (D : rel U). + +Definition inv_rel := + forall a1 a2 a', + valid (a1 \+ a2 \+ a') -> + D a1 (a2 \+ a') -> D a2 (a1 \+ a') -> + D a1 a2 && D (a1 \+ a2) a'. + +Definition inv_morph_axiom (f : morphism V D) := + forall (a : U) (b1 b2 : V), + valid a -> D a Unit -> f a = b1 \+ b2 -> + exists a1 a2, [/\ a = a1 \+ a2, valid (a1 \+ a2), + D a1 a2, f a1 = b1 & f a2 = b2 ]. + +Definition inv_morph f := inv_rel /\ inv_morph_axiom f. + +End Invertible. + +(* trivial seprel in invertible *) +Lemma inv_sepT U : inv_rel (@sepT_seprel U). Proof. by []. Qed. + +Lemma inv_idfun (U : pcm) : inv_morph [morphism of @idfun U]. +Proof. by split=>// a b1 b2 V _ E; exists b1, b2; rewrite -E. Qed. + +(* composition of invertible morhpsisms is invertible *) +Lemma inv_comp (U V W : pcm) (C : sep_rel V) (D : sep_rel U) + (f : morphism V D) (g : morphism W C) : + inv_morph g -> inv_morph f -> inv_morph [morphism of g \o f]. +Proof. +case=>HC Hg [HD Hf]; split=>[a1 a2 a' V1|a b1 b2 V1] /andP []. +- move=>/(HD _ _ _ V1) H1 H2 /andP [/H1] /andP [D1 D2]. + rewrite !sepE !pfjoin ?(validLE3 V1, sepAxx V1 D1 D2) //= in H2 *. + by apply: HC H2; apply: pfV3. +move=>D1; rewrite pfunit=>C1. +case/(Hg _ _ _ (pfV f V1 D1) C1)=>c1 [c2][Ef Vc Xc <-<-]. +case/(Hf _ _ _ V1 D1): Ef Xc=>a1 [a2][-> Va Xc <-<- Xd]. +by exists a1, a2; rewrite !sepE Xc Xd. +Qed. + +(* morphism product is invertible *) + +Lemma inv_cprod (U1 U2 V1 V2 : pcm) (D1 : sep_rel U1) (D2 : sep_rel U2) + (f : morphism V1 D1) (g : morphism V2 D2) : + inv_morph f -> inv_morph g -> inv_morph [morphism of f \* g]. +Proof. +case=>HD1 Hf [HD2 Hg]; split=>[a1 a2 a'|]; rewrite /sep_prod /=. +- case/andP=>Va Vb /andP [/(HD1 _ _ _ Va) H1 /(HD2 _ _ _ Vb) H2]. + by case/andP=>/H1 /andP [->->] /H2 /andP [->->]. +move=>/= a1 b1 b2 /andP [Va Vb] /andP [H1 H2][]. +case/(Hf _ _ _ Va H1)=>c1 [c2][E1 Vc1 Y1 Fc1 Fc2]. +case/(Hg _ _ _ Vb H2)=>d1 [d2][E2 Vd1 Y2 Fd1 Fd2]. +exists (c1, d1), (c2, d2); rewrite /fprod /= Vc1 Vd1. +by rewrite Fc1 Fc2 Fd1 Fd2 Y1 Y2 pcmPE -E1 -E2 -!prod_eta. +Qed. + +(* the construction of ker requires the range to be topped *) + +Lemma inv_ker U (V : tpcm) (D : sep_rel U) (f : morphism V D) : + inv_rel D -> inv_rel ('ker f). +Proof. +move=>HD a1 a2 a' W; rewrite !sepE /=. +case/and3P=>D1 /unitbP Eq1 _ /and3P [D2 /unitbP Eq2 /unitbP]. +case/andP: (HD _ _ _ W D1 D2)=>{}D1 {}D2. +rewrite !pfjoin ?(validLE3 W) //; last by rewrite (sepAxx W). +by rewrite Eq1 Eq2 D1 D2 !unitL =>->; rewrite tpcmE. +Qed. + +(* the construction of eqlz requires the range to be eqpcm *) +(* for now, we just assume cancellativity as a hypothesis *) + +Lemma inv_eqlz U (V : eqpcm) (D1 D2 : sep_rel U) + (f : morphism V D1) (g : morphism V D2) : + (forall x x1 x2 : EQPCM.pcm V, valid (x \+ x1) -> + x \+ x1 = x \+ x2 -> x1 = x2) -> + inv_rel D1 -> inv_rel D2 -> inv_rel ('eqlz f g). +Proof. +move=>K HD1 HD2 x y z W; rewrite !sepE /=. +case/and4P=>X1 X2 Ex _ /and4P [Y1 Y2] Ey /eqP. +case/andP: (HD1 _ _ _ W X1 Y1)=>E1 E1'. +case/andP: (HD2 _ _ _ W X2 Y2)=>E2 E2'. +rewrite E1 E2 Ex Ey E1' E2' !pfjoin ?(validLE3 W) 2?(sepAxx W) //=. +rewrite -(eqP Ex) (eqP Ey) eq_refl=>/K ->; first by rewrite eq_refl. +by rewrite pfV2 // ?(validLE3 W)// (sepAxx W). +Qed. + +(* this theorem generalizes the one from the POPL21 paper *) +(* to work with arbitrary sub-pcm over matching seprel *) +(* (the one in the paper worked with xsub explicitly) *) +(* can be further generalised to any D' that is compatible with D *) + +Lemma inv_comp_sub (U : tpcm) (V : pcm) (D : sep_rel U) (W : sub_pcm D) + (f : {morphism D >-> V}) : + inv_morph f -> inv_morph [morphism of f \o @pval U D W]. +Proof. +case=>HD Hf; split=>[a1 a2 a' V1|]. +- by rewrite !sepE !pfjoin ?(validLE3 V1) //; apply: HD; apply: pfV3. +move=>a b1 b2 V1; rewrite !sepE pfunit /= => H1. +case/(Hf _ _ _ (pfV _ V1 (erefl _)) H1)=>c1 [c2][Eq Vc H2 F1 F2]. +exists (psub _ c1), (psub _ c2); rewrite !sepE -!pfjoin // -Eq psub_pval //=. +by rewrite !pval_psub ?(validL Vc, validR Vc, sep0E Vc). +Qed. diff --git a/pcm/lift.v b/pcm/lift.v index a4285f0..105dcd5 100644 --- a/pcm/lift.v +++ b/pcm/lift.v @@ -11,19 +11,17 @@ See the License for the specific language governing permissions and limitations under the License. *) -From Coq Require Import ssreflect ssrbool ssrfun. -From mathcomp Require Import ssrnat eqtype. -From fcsl Require Import prelude pcm. -Set Implicit Arguments. -Unset Strict Implicit. -Unset Printing Implicit Defensive. - (******************************************************************************) (* This file contains an implementation of lift type adding an undef element *) (* a structure. *) (* Lifting turns an unlifted structure into a PCM and preserves equality. *) (******************************************************************************) +From Coq Require Import ssreflect ssrbool ssrfun. +From mathcomp Require Import ssrnat eqtype. +From fcsl Require Import prelude pcm morphism. +From fcsl Require Import options. + Module Unlifted. Record mixin_of (T : Type) := Mixin { @@ -204,7 +202,6 @@ End Lift. Export Lift.Exports. - Module NatUnlift. Local Definition ojoin (x y : nat) := Some (x + y). @@ -226,16 +223,43 @@ Definition natUnliftedMix := UnliftedMixin NatUnlift.ojoinC NatUnlift.ojoinA NatUnlift.ounitL. Canonical natUnlifted := Eval hnf in Unlifted nat natUnliftedMix. +Definition lnat := lift nat. +Canonical lnatPCM := Eval hnf in [pcm of lnat]. + +Section LNatTPCM. + +Lemma lnat_unitb (x : lnat) : reflect (x = Unit) (x == up 0). +Proof. by case: eqP=>H; constructor. Qed. + +Lemma lnat_valid_undef : ~~ valid (down : lnat). +Proof. by []. Qed. + +Lemma lnat_undef_join x : down \+ x = down :> lnat. +Proof. by []. Qed. + +Definition lnatTPCMMix := TPCMMixin lnat_unitb lnat_valid_undef lnat_undef_join. +Canonical lnatTPCM := Eval hnf in TPCM lnat lnatTPCMMix. +End LNatTPCM. + (* some lemmas for lifted nats *) -Lemma nxV (m1 m2 : lift natUnlifted) : +Lemma nxV (m1 m2 : lnat) : valid (m1 \+ m2) -> exists n1 n2, m1 = up n1 /\ m2 = up n2. Proof. by case: m1=>// n1; case: m2=>// n2; exists n1, n2. Qed. - -Lemma nxE0 (n1 n2 : lift natUnlifted) : +Lemma nxE0 (n1 n2 : lnat) : n1 \+ n2 = up 0 -> (n1 = up 0) * (n2 = up 0). Proof. case: n1 n2=>[|n1][|n2] //; rewrite upE /ojoin /=. by case=>/eqP; rewrite addn_eq0=>/andP [/eqP -> /eqP ->]. Qed. + +(* a morphism over lifted nat to get the underlying nat *) + +Definition nat_unlift (n : lift nat) : nat := if n is up m then m else 0. + +Lemma natunlift_morph_ax : morph_axiom (@sepT _) nat_unlift. +Proof. by rewrite /nat_unlift; split=>//; case=>// x; case. Qed. + +Canonical natunlift_morph := + Morphism' nat_unlift natunlift_morph_ax. diff --git a/pcm/morphism.v b/pcm/morphism.v new file mode 100644 index 0000000..4a3c9f9 --- /dev/null +++ b/pcm/morphism.v @@ -0,0 +1,1708 @@ +(* +Copyright 2017 IMDEA Software Institute +Licensed under the Apache License, Version 2.0 (the "License"); +you may not use this file except in compliance with the License. +You may obtain a copy of the License at + http://www.apache.org/licenses/LICENSE-2.0 +Unless required by applicable law or agreed to in writing, software +distributed under the License is distributed on an "AS IS" BASIS, +WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. +See the License for the specific language governing permissions and +limitations under the License. +*) + +From Coq Require Import ssreflect ssrbool ssrfun. +From mathcomp Require Import eqtype. +From fcsl Require Import axioms prelude pcm. +From fcsl Require Import options. + +(*****************) +(*****************) +(* PCM Morphisms *) +(*****************) +(*****************) + +(* separating relations are generalization of disjointness *) +(* Bart Jacobs calls this orthogonality *) + +Definition orth_axiom (U : pcm) (orth : rel U) := + [/\ (* unit is separated from unit *) + orth Unit Unit, + (* sep is commutative *) + forall x y, orth x y = orth y x, + (* sep implies validity *) + forall x y, orth x y -> valid (x \+ y), + (* is x is in the domain (i.e., x is considered) *) + (* then it is separate from at least Unit *) + forall x y, orth x y -> orth x Unit & + (* associativity *) + forall x y z, orth x y -> orth (x \+ y) z -> orth y z && orth x (y \+ z)]. + + +(* separating relation doesn't have the validity requirement *) +(* this is just a pragmatic issue; removing validity makes proof obligations *) +(* much simpler. Nothing is lost by removing validity, as we always *) +(* obtain it in other ways. Though, the axioms require validity as hypothesis *) +Definition seprel_axiom (U : pcm) (sep : rel U) := + [/\ (* unit is separated from unit *) + sep Unit Unit, + (* sep is commutative *) + (* the validity pre is necessary to get equivalence with orth *) + (* will it be bothersome in practice? *) + forall x y, valid (x \+ y) -> sep x y = sep y x, + (* is x is in the domain (i.e., x is considered) *) + (* then it is separate from at least Unit *) + forall x y, valid (x \+ y) -> sep x y -> sep x Unit & + (* associativity *) + forall x y z, valid (x \+ y \+ z) -> + sep x y -> sep (x \+ y) z -> sep y z && sep x (y \+ z)]. + +Lemma orth_sep (U : pcm) (sep : rel U) : + seprel_axiom sep <-> orth_axiom (fun x y => valid (x \+ y) && sep x y). +Proof. +split. +- case=>H1 H2 H3 H4; split=>[|x y|x y|x y|x y z]. + - by rewrite unitL valid_unit H1. + - by apply/andP/andP; case=> V; rewrite H2 (validE2 V). + - by case/andP. + - by case/andP=> V S; rewrite unitR (H3 x y) ?(validE2 V). + case/andP=>_ S1 /andP [V2 S2]. + by rewrite !(andX (H4 _ _ _ V2 S1 S2)) !(validLE3 V2). +case=>H1 H2 H3 H4 H5; split=>[|x y V|x y V S|x y z V H S]. +- by case/andP: H1. +- by move: (H2 x y); rewrite !(validE2 V). +- by rewrite (andX (H4 x y _)) // V S. +by move: (@H5 x y z); rewrite !(validLE3 V) H S => /(_ erefl erefl). +Qed. + +(* seprel equality *) + +(* because we have stripped compatability relations of the requirement *) +(* that valid (x \+ y), we have to put it back before we can prove *) +(* that cinv equals the equalizer between the given morphisms *) +Definition seprel_eq (U : pcm) (D1 D2 : rel U) := + forall x y, valid (x \+ y) -> D1 x y = D2 x y. + +Notation "D1 =s D2" := (seprel_eq D1 D2) (at level 30). + +Lemma sepEQ (U : pcm) (D1 D2 : rel U) : + D1 =s D2 <-> + (fun x y => valid (x \+ y) && D1 x y) =2 + (fun x y => valid (x \+ y) && D2 x y). +Proof. +split; first by move=>S=>x y; case: (valid (x \+ y)) (S x y)=>// ->. +by move=>S x y V; move: (S x y); rewrite V. +Qed. + +Lemma orthXEQ (U : pcm) (D1 D2 : rel U) : + D1 =2 D2 -> orth_axiom D1 <-> orth_axiom D2. +Proof. +move=>H; split; case=>H1 H2 H3 H4 H5. +- by split=>[|x y|x y|x y|x y z]; rewrite -!H; + [apply: H1 | apply: H2 | apply: H3 | apply: H4 | apply: H5]. +by split=>[|x y|x y|x y|x y z]; rewrite !H; + [apply: H1 | apply: H2 | apply: H3 | apply: H4 | apply: H5]. +Qed. + +Lemma sepXEQ (U : pcm) (D1 D2 : rel U) : + D1 =s D2 -> seprel_axiom D1 <-> seprel_axiom D2. +Proof. by move/sepEQ=>H; rewrite !orth_sep; apply: orthXEQ. Qed. + +(* Hence, given a separating relation R *) +(* we say that x \orth_R y iff valid (x \+ y) and R x y *) +(* Similarly, when R is the sep relatio of a morphism f *) +(* we will write x \orth_f y iff valid (x \+ y) and sep f x y *) + +Module SepRel. +Section SepRelDef. +Variable U : pcm. + +Structure mixin_of (sep : rel U) := Mixin { + _ : seprel_axiom sep}. + +Section ClassDef. + +Notation class_of := mixin_of (only parsing). + +Structure type : Type := Pack {sort; _ : class_of sort}. +Local Coercion sort : type >-> rel. + +Variable (p : rel U) (cp : type). +Definition class := let: Pack _ c as cT' := cp return class_of cT' in c. +Definition clone c of phant_id class c := @Pack p c. +Definition pack c := @Pack p c. + +End ClassDef. +End SepRelDef. + +Module Exports. +Coercion sort : type >-> rel. +Notation sep_rel := type. + +Notation "[ 'seprelMixin' 'of' R ]" := (class _ : mixin_of R) + (at level 0, format "[ 'seprelMixin' 'of' R ]") : pcm_scope. +Notation "[ 'seprel' 'of' R 'for' S ]" := (@clone _ R S _ idfun) + (at level 0, format "[ 'seprel' 'of' R 'for' S ]") : pcm_scope. +Notation "[ 'seprel' 'of' R ]" := (@clone _ R _ _ id) + (at level 0, format "[ 'seprel' 'of' R ]") : pcm_scope. + +Definition seprel (U : pcm) (sep : rel U) := + fun pf : seprel_axiom sep => @pack U sep (Mixin pf). +Arguments seprel {U}. + +Section Laws. +Variables (U : pcm) (sep : sep_rel U). + +Lemma sepax : seprel_axiom sep. +Proof. by case: sep=>s []. Qed. + +Lemma sep00 : sep Unit Unit. +Proof. by case: sep=>r [[H ???]]; apply: H. Qed. + +Lemma sepC x y : valid (x \+ y) -> sep x y = sep y x. +Proof. by case: sep=>r [[? H ??]]; apply: H. Qed. + +Lemma sepx0 x y : valid (x \+ y) -> sep x y -> sep x Unit. +Proof. by case: sep=>r [[?? H ?]]; apply: H. Qed. + +(* The order of the rewrite rules in the conclusion is important for + backwards reasoning: [sep x (y \+ z)] is more specific than [sep y z] and + hence [rewrite] should be able to do its work; + had we chosen to put [sep y z] first, [rewrite] would fail because of + the indefinite pattern *) +Lemma sepAx x y z : + valid (x \+ y \+ z) -> + sep x y -> sep (x \+ y) z -> sep x (y \+ z) * sep y z. +Proof. +by case: sep=>r [[??? H]] V R1 R2 /=; rewrite !(andX (H _ _ _ V R1 R2)). +Qed. + +(* derived lemmas *) + +Lemma sep0x x y : valid (x \+ y) -> sep x y -> sep Unit y. +Proof. +rewrite joinC=>V. +by rewrite -!(@sepC y) ?unitR ?(validE2 V) //; apply: sepx0. +Qed. + +Lemma sep0E x y : valid (x \+ y) -> sep x y -> sep x Unit * sep y Unit. +Proof. +by move=>V S; rewrite (sepx0 V S) sepC ?unitR ?(validE2 V) // (sep0x V S). +Qed. + +(* This is a helper lemma -- in most cases you may want to use + sepAxx or sepxxA instead *) +Lemma sepAx23_helper x y z : + valid (x \+ y \+ z) -> + sep x y -> sep (x \+ y) z -> + ((sep x z * sep z x) * (sep y z * sep z y)) * + ((sep x (y \+ z) * sep x (z \+ y)) * + (sep y (x \+ z) * sep y (z \+ x))). +Proof. +move=>V S1 S2. +rewrite !(@sepC z) ?(validLE3 V) // !(joinC z) !(sepAx V S1 S2). +rewrite sepC ?(validLE3 V) // in S1; rewrite (joinC x) in V S2. +by rewrite !(sepAx V S1 S2). +Qed. + +(* This is useful for backward reasoning, because we don't want to + depend on the exact permutation of x, y, z the rewrite rules (see below) + will choose *) +Lemma sepxA x y z : + valid (x \+ (y \+ z)) -> + sep y z -> sep x (y \+ z) -> sep (x \+ y) z * sep x y. +Proof. +move=>V S1; rewrite sepC // => S2. +rewrite (@sepC _ z) -?joinA //; rewrite joinC in V. +by rewrite (@sepC _ y) ?(validLE3 V) // !(sepAx23_helper V S1 S2). +Qed. + +(* nested pairs are a workaround for https://github.com/coq/coq/issues/8304 *) +Lemma sepAxx x y z : + valid (x \+ y \+ z) -> + sep x y -> sep (x \+ y) z -> + (((sep x (y \+ z) * sep x (z \+ y)) * + (sep y (x \+ z) * sep y (z \+ x))) * + ((sep z (x \+ y) * sep z (y \+ x)) * + (sep (y \+ z) x * sep (z \+ y) x))) * + (((sep (x \+ z) y * sep (z \+ x) y) * + (sep (x \+ y) z * sep (y \+ x) z)) * + (((sep x y * sep y x) * + (sep x z * sep z x)))) * + (sep y z * sep z y). +Proof. +move=>V S1 S2; rewrite S1 S2 !(sepAx23_helper V S1 S2). +rewrite -(sepC (validL V)) S1. +rewrite (joinC y) -sepC // S2; +rewrite -(joinC y) sepC 1?joinC ?joinA // (sepAx23_helper V S1 S2). +by rewrite -(joinC x) sepC 1?joinAC // (sepAx23_helper V S1 S2). +Qed. + +(* same results, flipped hypotheses *) +(* nested pairs are a workaround for https://github.com/coq/coq/issues/8304 *) +Lemma sepxxA x y z : + valid (x \+ (y \+ z)) -> + sep y z -> sep x (y \+ z) -> + (((sep x (y \+ z) * sep x (z \+ y)) * + (sep y (x \+ z) * sep y (z \+ x))) * + ((sep z (x \+ y) * sep z (y \+ x)) * + (sep (y \+ z) x * sep (z \+ y) x))) * + (((sep (x \+ z) y * sep (z \+ x) y) * + (sep (x \+ y) z * sep (y \+ x) z)) * + (((sep x y * sep y x) * + (sep x z * sep z x)))) * + (sep y z * sep z y). +Proof. +rewrite joinA=> V; rewrite {1}(@sepC x) ?(validLE3 V) // => S1 S2. +by apply: (sepAxx V); rewrite (joinC x) joinAC in V; rewrite !(sepAxx V S1 S2). +Qed. + +Lemma sepU0 x y : valid (x \+ y) -> sep x y -> sep (x \+ y) Unit. +Proof. +move=>V S1. move/(sepx0 V): S1 (S1)=>S1 S2. +rewrite -[x]unitR in V S2. +by rewrite sepC ?(sepAxx V S1 S2) // joinAC. +Qed. + +Lemma sep0U x y : valid (x \+ y) -> sep x y -> sep Unit (x \+ y). +Proof. by move=>V S; rewrite sepC ?unitL //; apply: sepU0. Qed. + +End Laws. + +(* some example sep relations *) + +(* always-true relation *) +Definition sepT U (x y : U) := true. + +Lemma sepT_seprel_ax (U : pcm) : seprel_axiom (@sepT U). +Proof. by []. Qed. + +Canonical sepT_seprel (U : pcm) := + Eval hnf in seprel (@sepT U) (@sepT_seprel_ax U). + +(* always-unit relation *) +(* always-false relation is not seprel, because we need sep0 Unit Unit *) +(* i.e., sep0 is really the smallest seprel we can have *) +Definition sep0 (U : tpcm) (x y : U) := unitb x && unitb y. + +Prenex Implicits sep0. + +Lemma sep0_seprel_ax (U : tpcm) : seprel_axiom (@sep0 U). +Proof. +rewrite /sep0; split=>[|x y|x y|x y z]. +- by rewrite tpcmE. +- by rewrite andbC. +- by move=>V /andP [->]; rewrite tpcmE. +move=>V /andP [/unitbP -> /unitbP ->] /andP [_ /unitbP ->]. +by rewrite unitL !tpcmE. +Qed. + +Canonical sepU_seprel U := Eval hnf in seprel (@sep0 U) (@sep0_seprel_ax U). + +(* conjunction of seprels is seprel *) + +Definition sepI U (R1 R2 : rel U) (x y : U) := R1 x y && R2 x y. + +Lemma sepI_seprel_ax (U : pcm) (R1 R2 : sep_rel U) : seprel_axiom (sepI R1 R2). +Proof. +rewrite /sepI; split=>[|x y|x y|x y z]. +- by rewrite !sep00. +- by move=>D; rewrite sepC //= [in X in _ && X]sepC. +- by move=>D /andP [X H]; rewrite (sepx0 D X) (sepx0 D H). +move=>D /andP [S1 H1] /andP [S2 H2]. +by rewrite !(sepAxx D S1 S2) !(sepAxx D H1 H2). +Qed. + +Canonical sepI_seprel U (R1 R2 : sep_rel U) := + Eval hnf in seprel (sepI R1 R2) (@sepI_seprel_ax U R1 R2). + +(* pairwise product of seprels is seprel *) + +Definition sep_prod U1 U2 (R1 : rel U1) (R2 : rel U2) (x y : U1 * U2) := + R1 x.1 y.1 && R2 x.2 y.2. + +Lemma sep_prod_seprel_ax U1 U2 (R1 : sep_rel U1) (R2 : sep_rel U2) : + seprel_axiom (sep_prod R1 R2). +Proof. +rewrite /sep_prod; split=>[|x y|x y|x y z]. +- by rewrite !sep00. +- by case/andP=>/sepC -> /sepC ->. +- by case/andP=>?? /andP [/sepx0 -> // /sepx0 ->]. +case/andP=>V1 V2 /andP [D1 D2] /andP [/= Z1 Z2]. +by rewrite !(sepAxx V1 D1 Z1) !(sepAxx V2 D2 Z2). +Qed. + +Canonical sep_prod_seprel U1 U2 (R1 : sep_rel U1) (R2 : sep_rel U2) := + Eval hnf in seprel (sep_prod R1 R2) (sep_prod_seprel_ax R1 R2). + +End Exports. +End SepRel. + +Export SepRel.Exports. + + +(* Now we can define a morphism to come with a sep relation *) +(* on which it is valid *) + +Definition morph_axiom (U V : pcm) (sep : rel U) (f : U -> V) := + [/\ (* f preserves unit *) + f Unit = Unit & + (* f is defined on the domain and preserves joins on the domain *) + forall x y, valid (x \+ y) -> sep x y -> + valid (f x \+ f y) /\ f (x \+ y) = f x \+ f y]. + +Section MorphismStructure. +Variables U V : pcm. + +Structure morphism (D : rel U) : Type := Morphism' { + mfun :> U -> PCM.sort V; + _ : morph_axiom D mfun}. + +Definition morphism_for (D : rel U) of phant V := morphism D. + +Definition clone_morphism (D : rel U) f := + let: @Morphism' _ _ pf := f + return {type of @Morphism' D for f} -> + morphism_for D (Phant V) + in fun k => k pf. + +End MorphismStructure. + +Arguments Morphism' [U V D] mfun _. + +Notation "{ 'morphism' D >-> V }" := (morphism_for D (Phant V)) + (at level 0, format "{ 'morphism' D >-> V }") : pcm_scope. +Notation "[ 'morphism' D 'of' f ]" := + (clone_morphism (@Morphism' _ _ D f)) + (at level 0, format "[ 'morphism' D 'of' f ]") : pcm_scope. +Notation "[ 'morphism' 'of' f ]" := + (clone_morphism (@Morphism' _ _ _ f)) + (at level 0, format "[ 'morphism' 'of' f ]") : pcm_scope. + +Definition Morphism (U V : pcm) (D : rel U) f := + fun (p1 : f Unit = Unit) + (p2 : forall x y, valid (x \+ y) -> D x y -> + valid (f x \+ f y) /\ f (x \+ y) = f x \+ f y) => + @Morphism' U V D f (conj p1 p2). +Arguments Morphism [U V]. + +Section Laws. +Variables (U V : pcm) (D : rel U) (f : {morphism D >-> V}). + +Lemma pfunit : f Unit = Unit. +Proof. by case: f=>g [H1 H2]; apply: H1. Qed. + +Lemma pfjoinV x y : + valid (x \+ y) -> D x y -> + valid (f x \+ f y) /\ f (x \+ y) = f x \+ f y. +Proof. by case: f=>g [H1 H2]; apply: H2. Qed. + +(* derived laws *) + +Lemma pfV2 x y : valid (x \+ y) -> D x y -> valid (f x \+ f y). +Proof. by move=>H S; case: (pfjoinV H S). Qed. + +Lemma pfjoin x y : valid (x \+ y) -> D x y -> f (x \+ y) = f x \+ f y. +Proof. by move=>H S; case: (pfjoinV H S). Qed. + +Lemma joinpf x y : valid (x \+ y) -> D x y -> f x \+ f y = f (x \+ y). +Proof. by move=>*; rewrite pfjoin. Qed. + +End Laws. + +(* some lemmas when D is a seprel *) + +Lemma pfV (U : pcm) V (D : rel U) (f : @morphism U V D) x : + valid x -> D x Unit -> valid (f x). +Proof. by rewrite -{1 3}[x]unitR=>W S; rewrite pfjoin //; apply: pfV2. Qed. + +Lemma pfV3 (U : pcm) V (D : rel U) (f : @morphism U V D) x y z: + valid (x \+ y \+ z) -> D x y -> D (x \+ y) z -> valid (f x \+ f y \+ f z). +Proof. by move=>W D1 D2;rewrite -pfjoin ?(validL W) // pfV2 //=. Qed. + +(* a version of pfV2 where the morphism structure *) +(* is inferred from phantoms *) + +Lemma pfV2_phant U V D (f : @morphism U V D) of phantom (_ -> _) f : + forall x y, valid (x \+ y) -> D x y -> valid (f x \+ f y). +Proof. by move=>x y; apply: pfV2. Qed. + +Notation pfVf f := (@pfV2_phant _ _ _ _ (Phantom (_ -> _) f)). + +(* Domain sep_rel, preimage, kernel, restriction of morphism *) +(* use phantoms to infer the morphism structure *) +(* and uncover D *) +(* we make these constructions only when D is a seprel *) +(* because it's strange in lemmas when they are not *) + +Definition sep (U : pcm) V D (f : @morphism U V D) + of phantom (U -> V) f := D : rel U. + +Definition preim (U : pcm) V (D : rel U) (f : @morphism U V D) (R : rel V) + of phantom (U -> V) f : rel U := + fun x y => D x y && R (f x) (f y). + +(* kernel is given with a tpcm range because we need decidable test for unit *) +(* we may have gone with eqpcms too, which have decidable tests for everything, *) +(* unit included *) + +Definition ker (U : pcm) (V : tpcm) (D : rel U) (f : @morphism U V D) + of phantom (U -> V) f : rel U := + fun x y => D x y && sep0 (f x) (f y). + +(* res requires a tpcm range because it needs to make a result undefined *) +Definition res (U : pcm) (V : tpcm) (f : U -> V) (S : rel U) := + fun x : U => if S x Unit then f x else undef. + +(* equalizer *) +Definition eqlz (U : pcm) (V : eqpcm) (D1 D2 : rel U) + (f1 : @morphism U V D1) (f2 : @morphism U V D2) + of phantom (U -> V) f1 & phantom (U -> V) f2 : rel U := + fun x y => [&& D1 x y, D2 x y, f1 x == f2 x & f1 y == f2 y]. + +(* sep-rel of a join *) +Definition sep_join (U V : pcm) (D1 D2 : rel U) + (f1 : @morphism U V D1) (f2 : @morphism U V D2) + of phantom (U -> V) f1 & phantom (U -> V) f2 := + fun x y => [&& D1 x y, D2 x y & valid ((f1 x \+ f2 x) \+ (f1 y \+ f2 y))]. + +(* join *) +Definition join (U V : pcm) (D1 D2 : rel U) + (f1 : @morphism U V D1) (f2 : @morphism U V D2) + of phantom (U -> V) f1 & phantom (U -> V) f2 := + fun x : U => f1 x \+ f2 x. + +(* Notation for the ops *) +Notation "''sep' f" := (sep (Phantom (_ -> _) f)) + (at level 10, f at level 8, format "''sep' f") : pcm_scope. + +Notation "''preim' f R" := (preim R (Phantom (_ -> _) f)) + (at level 10, f at level 2, R at level 8, format "''preim' f R") : pcm_scope. + +Notation "''ker' f" := (ker (Phantom (_ -> _) f)) + (at level 10, f at level 8, format "''ker' f") : pcm_scope. + +Notation "''res_' S f" := (res f S) + (at level 10, S at level 2, f at level 8, format "''res_' S f") : pcm_scope. + +Notation "''eqlz' f1 f2" := (eqlz (Phantom (_ -> _) f1) (Phantom (_ -> _) f2)) + (at level 10, f1 at level 2, f2 at level 2, + format "''eqlz' f1 f2") : pcm_scope. + +Notation "''sep_join' f1 f2" := + (sep_join (Phantom (_ -> _) f1) (Phantom (_ -> _) f2)) + (at level 10, f1 at level 2, f2 at level 2, + format "''sep_join' f1 f2") : pcm_scope. + +Notation "''join' f1 f2" := (join (Phantom (_ -> _) f1) (Phantom (_ -> _) f2)) + (at level 10, f1 at level 2, f2 at level 2, + format "''join' f1 f2") : pcm_scope. + +(* unfolding the defs to get rid of phantoms *) +Lemma sepX (U : pcm) V (D : rel U) (f : @morphism U V D) : 'sep f = D. +Proof. by []. Qed. + +Lemma preimX (U : pcm) V (D : rel U) (f : @morphism U V D) (R : rel V) x y : + 'preim f R x y = D x y && R (f x) (f y). +Proof. by []. Qed. + +Lemma kerX (U : pcm) (V : tpcm) (D : rel U) (f : @morphism U V D) x y : + 'ker f x y = [&& D x y, unitb (f x) & unitb (f y)]. +Proof. by []. Qed. + +Lemma resX U (V : tpcm) D (f : @morphism U V D) S x : + 'res_S f x = if S x Unit then f x else undef. +Proof. by []. Qed. + +Lemma eqlzX (U : pcm) (V : eqpcm) D1 D2 + (f1 : @morphism U V D1) (f2 : @morphism U V D2) x y : + 'eqlz f1 f2 x y = [&& D1 x y, D2 x y, f1 x == f2 x & f1 y == f2 y]. +Proof. by []. Qed. + +Lemma sep_joinX (U V : pcm) D1 D2 + (f1 : @morphism U V D1) (f2 : @morphism U V D2) x y : + 'sep_join f1 f2 x y = + [&& D1 x y, D2 x y & valid ((f1 x \+ f2 x) \+ (f1 y \+ f2 y))]. +Proof. by []. Qed. + +Lemma joinX (U V : pcm) D1 D2 + (f1 : @morphism U V D1) (f2 : @morphism U V D2) x : + 'join f1 f2 x = f1 x \+ f2 x. +Proof. by []. Qed. + +(* most of the time we just want to unfold everything *) +Definition sepE := (((sepX, preimX), (kerX, resX)), ((eqlzX, sep_joinX), joinX)). + +(* lemmas that obtain when D's are seprels *) +Section MorphPCMLemmas. +Variables (U V : pcm) (D : sep_rel U) (f : {morphism D >-> V}). + +Lemma sep_seprel_ax : seprel_axiom D. +Proof. by case: D=>r []. Qed. + +Canonical sep_seprel := Eval hnf in seprel ('sep f) sep_seprel_ax. + +Variable R : sep_rel V. + +Lemma preim_seprel_ax : seprel_axiom ('preim f R). +Proof. +rewrite /preim; split=>[|x y|x y|x y z]; first by rewrite !pfunit !sep00. +- move=>H; rewrite sepC //=; case S: (D y x)=>//=. + by rewrite -sepC // pfV2 // joinC. +- move=>H /andP [H1 H2]. + by rewrite !pfunit (sep0E H H1) (sep0E _ H2) // pfV2. +move=>H /andP [G1 F1] /andP [G2 F2]. +rewrite pfjoin ?(validLE3 H) // in F2. +rewrite pfjoin ?(validLE3 H) ?(sepAxx H G1 G2) //=. +move: (pfVf f H G2); rewrite pfjoin ?(validLE3 H) // => D2. +by rewrite !(sepAxx D2 F1 F2). +Qed. + +Canonical preim_seprel := + Eval hnf in seprel ('preim f R) preim_seprel_ax. + +End MorphPCMLemmas. + + +Section MorphTPCMLemmas. +Variable (U : pcm) (V : tpcm) (D : sep_rel U) (f : {morphism D >-> V}). + +Lemma ker_seprel_ax : seprel_axiom ('ker f). +Proof. by apply: preim_seprel_ax. Qed. + +Canonical ker_seprel := Eval hnf in seprel ('ker f) ker_seprel_ax. + +Variable S : sep_rel U. + +Lemma res_morph_ax : morph_axiom (sepI S D) ('res_S f). +Proof. +rewrite /res; split=>[|x y]; first by rewrite sep00 pfunit. +by move=>W /andP [X H]; rewrite !(sep0E W X) (sepU0 W X) pfjoin // pfV2. +Qed. + +Canonical res_morph := Morphism' ('res_S f) res_morph_ax. + +End MorphTPCMLemmas. + + +(* equalizer is a sep_rel (i.e., compatibility relation) *) +Section MorphEQLZLemmas. +Variables (U : pcm) (V : eqpcm) (D1 D2 : sep_rel U). +Variables (f1 : @morphism U V D1) (f2 : @morphism U V D2). + +Lemma eqlz_seprel_ax : seprel_axiom ('eqlz f1 f2). +Proof. +rewrite /eqlz; split=>[|x y W|x y W|x y z W]. +- by rewrite !sep00 !pfunit eq_refl. +- by rewrite !andbA andbAC (sepC D1) // (sepC D2). +- by case/and4P=>V1 V2 -> _; rewrite (sepx0 W V1) (sepx0 W V2) !pfunit eq_refl. +case/and4P=>V1 V2 Ex Ey /and4P [W1 W2 _ Ez]. +set j1 := (sepAxx W V1 W1); set j2 := (sepAxx W V2 W2). +by rewrite !pfjoin ?j1 ?j2 ?(validLE3 W) //= Ex (eqP Ey) (eqP Ez) !eq_refl. +Qed. + +Canonical eqlz_seprel := Eval hnf in seprel ('eqlz f1 f2) eqlz_seprel_ax. + +End MorphEQLZLemmas. + +(* join of morphisms is a morphism *) +Section MorphJoinLemmas. +Variables (U V : pcm) (D1 D2 : sep_rel U). +Variables (f1 : @morphism U V D1) (f2 : @morphism U V D2). + +Lemma sepjoin_seprel_ax : seprel_axiom ('sep_join f1 f2). +Proof. +rewrite /sep_join; split=>[|x y W|x y W|x y z W]. +- by rewrite !sep00 ?unitL !pfunit ?unitL valid_unit. +- by rewrite (sepC D1) // (sepC D2) // (joinC (f1 x \+ f2 x)). +- case/and3P=>V1 V2; rewrite (sepx0 W V1) (sepx0 W V2) !pfunit // !unitR. + by apply: validL. +case/and3P=>V1 V2 Wa /and3P [W1 W2]. +set j1 := (sepAxx W V1 W1); set j2 := (sepAxx W V2 W2). +rewrite !pfjoin ?j1 ?j2 ?(validLE3 W) //=. +rewrite !(pull (f2 y)) !joinA !(pull (f1 y)) !joinA. +rewrite !(pull (f2 x)) !joinA !(pull (f1 x)) -!joinA=>Wb. +by rewrite !(validRE3 Wb). +Qed. + +Canonical sepjoin_seprel := + Eval hnf in seprel ('sep_join f1 f2) sepjoin_seprel_ax. + +Lemma join_morph_ax : morph_axiom ('sep_join f1 f2) ('join f1 f2). +Proof. +rewrite /join; split=>[|x y]; first by rewrite !pfunit unitL. +move=>W /and3P [V1 V2]; rewrite !pfjoin // !joinA. +by rewrite !(pull (f2 x)) !joinA !(pull (f1 x)). +Qed. + +Canonical join_morph := Morphism' ('join f1 f2) join_morph_ax. + +End MorphJoinLemmas. + + +(* morphisms compose *) +Section CategoricalDefs. +Variables (U V W : pcm) (Du : rel U) (Dv : rel V). +Variable f : {morphism Du >-> V}. +Variable g : {morphism Dv >-> W}. + +(* identity is a morphism *) +Lemma id_morph_ax : morph_axiom (@sepT U) (@idfun U). +Proof. by []. Qed. + +Canonical id_morph := Eval hnf in Morphism' (@idfun U) id_morph_ax. + +Lemma comp_morph_ax : morph_axiom ('preim f Dv) (g \o f). +Proof. +split=>[|x y] /=; first by rewrite !pfunit. +by rewrite /preim => D /andP [H1 H2]; rewrite !pfjoin // !pfV2. +Qed. + +Canonical comp_morph := Morphism' (g \o f) comp_morph_ax. + +End CategoricalDefs. + +(* morphism equality is function equality *) +(* ie. we don't take domains into consideration *) +(* thus, the categorical laws trivially hold *) +Section CategoricalLaws. +Variables U V W X : pcm. +Variables (Df : rel V) (f : {morphism Df >-> U}). +Variables (Dg : rel W) (g : {morphism Dg >-> V}). +Variables (Dh : rel W) (h : {morphism Dh >-> W}). + +Lemma morph0L : f \o idfun = f. Proof. by []. Qed. +Lemma morph0R : idfun \o f = f. Proof. by []. Qed. +Lemma morphA : f \o (g \o h) = (f \o g) \o h. Proof. by []. Qed. + +End CategoricalLaws. + +(* pairwise product of morphisms is a morphism *) +Section ProdMorph. +Variables U1 U2 V1 V2 : pcm. +Variables (D1 : rel U1) (f1 : {morphism D1 >-> V1}). +Variables (D2 : rel U2) (f2 : {morphism D2 >-> V2}). + +Lemma fprod_morph_ax : morph_axiom (sep_prod D1 D2) (f1 \* f2). +Proof. +rewrite /fprod; split=>[|x y]; first by rewrite !pfunit. +by case/andP=>/= W1 W2 /andP [H1 H2]; rewrite !pfV2 /= ?(pfjoin,W1,W2). +Qed. + +Canonical fprod_morph := Morphism' (f1 \* f2) fprod_morph_ax. + +End ProdMorph. + + +(* projections are morphisms *) +Section ProjectionsMorph. +Variables (U1 U2 : pcm). + +Lemma fst_morph_ax : morph_axiom (@sepT (prodPCM U1 U2)) fst. +Proof. by split=>[|x y]; rewrite ?pcmPE //= =>/andP []. Qed. + +Canonical fst_morph := Morphism' fst fst_morph_ax. + +Lemma snd_morph_ax : morph_axiom (@sepT (prodPCM U1 U2)) snd. +Proof. by split=>[|x y]; rewrite ?pcmPE //= =>/andP []. Qed. + +Canonical snd_morph := Morphism' snd snd_morph_ax. +End ProjectionsMorph. + +(* also for explicit triples *) +Section ProjectionsMorph3. +Variables (U1 U2 U3 : pcm). + +Lemma pcm31_morph_ax : morph_axiom (@sepT (prod3_PCM U1 U2 U3)) pcm31. +Proof. by split=>[|x y]; rewrite !pcmE /= !pcmE //= =>/and3P []. Qed. +Canonical pcm31_morph := Morphism' pcm31 pcm31_morph_ax. + +Lemma pcm32_morph_ax : morph_axiom (@sepT (prod3_PCM U1 U2 U3)) pcm32. +Proof. by split=>[|x y]; rewrite !pcmE /= !pcmE // =>/and3P []. Qed. +Canonical pcm32_morph := Morphism' pcm32 pcm32_morph_ax. + +Lemma pcm33_morph_ax : morph_axiom (@sepT (prod3_PCM U1 U2 U3)) pcm33. +Proof. by split=>[|x y]; rewrite !pcmE /= !pcmE // =>/and3P []. Qed. +Canonical pcm3_morph := Morphism' pcm33 pcm33_morph_ax. +End ProjectionsMorph3. + +(* also for explicit 4-tuples *) +Section ProjectionsMorph4. +Variables (U1 U2 U3 U4 : pcm). + +Lemma pcm41_morph_ax : morph_axiom (@sepT (prod4_PCM U1 U2 U3 U4)) pcm41. +Proof. by split=>[|x y]; rewrite !pcmE /= !pcmE //= =>/and4P []. Qed. +Canonical pcm41_morph := Morphism' pcm41 pcm41_morph_ax. + +Lemma pcm42_morph_ax : morph_axiom (@sepT (prod4_PCM U1 U2 U3 U4)) pcm42. +Proof. by split=>[|x y]; rewrite !pcmE /= !pcmE //= =>/and4P []. Qed. +Canonical pcm42_morph := Morphism' pcm42 pcm42_morph_ax. + +Lemma pcm43_morph_ax : morph_axiom (@sepT (prod4_PCM U1 U2 U3 U4)) pcm43. +Proof. by split=>[|x y]; rewrite !pcmE /= !pcmE //= =>/and4P []. Qed. +Canonical pcm43_morph := Morphism' pcm43 pcm43_morph_ax. + +Lemma pcm44_morph_ax : morph_axiom (@sepT (prod4_PCM U1 U2 U3 U4)) pcm44. +Proof. by split=>[|x y]; rewrite !pcmE /= !pcmE //= =>/and4P []. Qed. +Canonical pcm44_morph := Morphism' pcm44 pcm44_morph_ax. +End ProjectionsMorph4. + +(* and 5-tuples *) +Section ProjectionsMorph5. +Variables (U1 U2 U3 U4 U5 : pcm). + +Lemma pcm51_morph_ax : morph_axiom (@sepT (prod5_PCM U1 U2 U3 U4 U5)) pcm51. +Proof. by split=>[|x y]; rewrite !pcmE /= !pcmE //= =>/and4P []. Qed. +Canonical pcm51_morph := Morphism' pcm51 pcm51_morph_ax. + +Lemma pcm52_morph_ax : morph_axiom (@sepT (prod5_PCM U1 U2 U3 U4 U5)) pcm52. +Proof. by split=>[|x y]; rewrite !pcmE /= !pcmE //= =>/and4P []. Qed. +Canonical pcm52_morph := Morphism' pcm52 pcm52_morph_ax. + +Lemma pcm53_morph_ax : morph_axiom (@sepT (prod5_PCM U1 U2 U3 U4 U5)) pcm53. +Proof. by split=>[|x y]; rewrite !pcmE /= !pcmE //= =>/and4P []. Qed. +Canonical pcm53_morph := Morphism' pcm53 pcm53_morph_ax. + +Lemma pcm54_morph_ax : morph_axiom (@sepT (prod5_PCM U1 U2 U3 U4 U5)) pcm54. +Proof. by split=>[|x y]; rewrite !pcmE /= !pcmE //= =>/and5P []. Qed. +Canonical pcm54_morph := Morphism' pcm54 pcm54_morph_ax. + +Lemma pcm55_morph_ax : morph_axiom (@sepT (prod5_PCM U1 U2 U3 U4 U5)) pcm55. +Proof. by split=>[|x y]; rewrite !pcmE /= !pcmE //= =>/and5P []. Qed. +Canonical pcm55_morph := Morphism' pcm55 pcm55_morph_ax. + +End ProjectionsMorph5. + +(* and 6-tuples *) +Section ProjectionsMorph6. +Variables (U1 U2 U3 U4 U5 U6 : pcm). + +Lemma pcm61_morph_ax : morph_axiom (@sepT (prod6_PCM U1 U2 U3 U4 U5 U6)) pcm61. +Proof. by split=>[|x y]; rewrite !pcmE /= !pcmE //= =>/and6P []. Qed. +Canonical pcm61_morph := Morphism' pcm61 pcm61_morph_ax. + +Lemma pcm62_morph_ax : morph_axiom (@sepT (prod6_PCM U1 U2 U3 U4 U5 U6)) pcm62. +Proof. by split=>[|x y]; rewrite !pcmE /= !pcmE //= =>/and6P []. Qed. +Canonical pcm62_morph := Morphism' pcm62 pcm62_morph_ax. + +Lemma pcm63_morph_ax : morph_axiom (@sepT (prod6_PCM U1 U2 U3 U4 U5 U6)) pcm63. +Proof. by split=>[|x y]; rewrite !pcmE /= !pcmE //= =>/and6P []. Qed. +Canonical pcm63_morph := Morphism' pcm63 pcm63_morph_ax. + +Lemma pcm64_morph_ax : morph_axiom (@sepT (prod6_PCM U1 U2 U3 U4 U5 U6)) pcm64. +Proof. by split=>[|x y]; rewrite !pcmE /= !pcmE //= =>/and6P []. Qed. +Canonical pcm64_morph := Morphism' pcm64 pcm64_morph_ax. + +Lemma pcm65_morph_ax : morph_axiom (@sepT (prod6_PCM U1 U2 U3 U4 U5 U6)) pcm65. +Proof. by split=>[|x y]; rewrite !pcmE /= !pcmE //= =>/and6P []. Qed. +Canonical pcm65_morph := Morphism' pcm65 pcm65_morph_ax. + +Lemma pcm66_morph_ax : morph_axiom (@sepT (prod6_PCM U1 U2 U3 U4 U5 U6)) pcm66. +Proof. by split=>[|x y]; rewrite !pcmE /= !pcmE //= =>/and6P []. Qed. +Canonical pcm66_morph := Morphism' pcm66 pcm66_morph_ax. + +End ProjectionsMorph6. + +(* and 7-tuples *) +Section ProjectionsMorph7. +Variables (U1 U2 U3 U4 U5 U6 U7 : pcm). + +Lemma pcm71_morph_ax : morph_axiom (@sepT (prod7_PCM U1 U2 U3 U4 U5 U6 U7)) pcm71. +Proof. by split=>[|x y]; rewrite !pcmE /= !pcmE //= =>/and7P []. Qed. +Canonical pcm71_morph := Morphism' pcm71 pcm71_morph_ax. + +Lemma pcm72_morph_ax : morph_axiom (@sepT (prod7_PCM U1 U2 U3 U4 U5 U6 U7)) pcm72. +Proof. by split=>[|x y]; rewrite !pcmE /= !pcmE //= =>/and7P []. Qed. +Canonical pcm72_morph := Morphism' pcm72 pcm72_morph_ax. + +Lemma pcm73_morph_ax : morph_axiom (@sepT (prod7_PCM U1 U2 U3 U4 U5 U6 U7)) pcm73. +Proof. by split=>[|x y]; rewrite !pcmE /= !pcmE //= =>/and7P []. Qed. +Canonical pcm73_morph := Morphism' pcm73 pcm73_morph_ax. + +Lemma pcm74_morph_ax : morph_axiom (@sepT (prod7_PCM U1 U2 U3 U4 U5 U6 U7)) pcm74. +Proof. by split=>[|x y]; rewrite !pcmE /= !pcmE //= =>/and7P []. Qed. +Canonical pcm74_morph := Morphism' pcm74 pcm74_morph_ax. + +Lemma pcm75_morph_ax : morph_axiom (@sepT (prod7_PCM U1 U2 U3 U4 U5 U6 U7)) pcm75. +Proof. by split=>[|x y]; rewrite !pcmE /= !pcmE //= =>/and7P []. Qed. +Canonical pcm75_morph := Morphism' pcm75 pcm75_morph_ax. + +Lemma pcm76_morph_ax : morph_axiom (@sepT (prod7_PCM U1 U2 U3 U4 U5 U6 U7)) pcm76. +Proof. by split=>[|x y]; rewrite !pcmE /= !pcmE //= =>/and7P []. Qed. +Canonical pcm76_morph := Morphism' pcm76 pcm76_morph_ax. + +Lemma pcm77_morph_ax : morph_axiom (@sepT (prod7_PCM U1 U2 U3 U4 U5 U6 U7)) pcm77. +Proof. by split=>[|x y]; rewrite !pcmE /= !pcmE //= =>/and7P []. Qed. +Canonical pcm77_morph := Morphism' pcm77 pcm77_morph_ax. + +End ProjectionsMorph7. + +(* product morphism of morphisms is a morphism *) +Section PMorphMorph. +Variables (U V1 V2 : pcm). +Variables (D1 : rel U) (f1 : {morphism D1 >-> V1}). +Variables (D2 : rel U) (f2 : {morphism D2 >-> V2}). + +Lemma pmorph_morph_ax : morph_axiom (sepI D1 D2) (f1 \** f2). +Proof. +split=>[|x y] /=; first by rewrite !pcmPE !pfunit. +by move=>V /andP [S1 S2]; rewrite !pcmPE (pfV2 f1 V S1) (pfV2 f2 V S2) !pfjoin. +Qed. + +Canonical pmorph_morph := + Morphism' (f1 \** f2) pmorph_morph_ax. + +End PMorphMorph. + +(* Can we say anything about pairing as a morphism *) +(* (-, -) : U -> V -> U * V *) +(* Because of currying, we first need to define a PCM of functions *) +(* That's easy, as join and unit can be defined pointwise *) +(* In that PCM, we can ask if pairing is a morphism in either argument *) +(* e.g. if we focus on the first argument, is *) +(* (x \+ y, _) = (x, _) \+ (y, _) *) +(* It isn't, since _ has to be replaced by the same value everywhere *) + +(***********************************************************) +(* Subpcm comes with an injection pval and retraction psub *) +(***********************************************************) + +Module SubPCM. +Section ClassDef. +Variables (V : pcm) (D : rel V). + +(* we demand a sep-rel here, because I couldn't prove *) +(* the sep_pers lemmas in world.v otherwise *) +(* i.e., for constructing sub-resources, we really want this relation *) +(* to be a seprel *) +Record mixin_of (U : pcm) : Type := Mixin { + pval_op : {morphism (@sepT U) >-> V}; + psub_op : {morphism D >-> U}; + (* separated in V iff images separated in U *) + _ : forall x y, valid (psub_op x \+ psub_op y) -> + valid (x \+ y) && D x y; + (* inject then retract is id *) + _ : forall x, psub_op (pval_op x) = x; + (* retract then inject is id on valids *) + _ : forall x, valid x -> D x Unit -> pval_op (psub_op x) = x; + (* injection separates valids from invalids in U *) + (* this axiom needed so that transitions in subresources *) + (* don't need to add side-conditions on validity *) + _ : forall x, valid (pval_op x) -> valid x}. + +(* we base the inheritance on type instead of pcm *) +(* thus, we can infer the sub_pcm structure out of the type *) +(* used in the sub_pcm construction *) +Record class_of (U : Type) := Class { + base : PCM.mixin_of U; + mixin: mixin_of (PCM.Pack base)}. +Local Coercion base : class_of >-> PCM.mixin_of. + +Structure type : Type := Pack {sort; _ : class_of sort}. +Local Coercion sort : type >-> Sortclass. + +Variables (T : Type) (cT : type). +Definition class := let: Pack _ c as cT' := cT return class_of cT' in c. +Let xT := let: Pack T _ := cT in T. +Notation xclass := (class : class_of xT). +Definition clone c of phant_id class c := @Pack T c. +Definition pack (b0 : PCM.mixin_of T) (m0 : mixin_of (PCM.Pack b0)) := + fun m & phant_id m0 m => Pack (@Class T b0 m). + +Definition pcmType := @PCM.Pack cT xclass. + +End ClassDef. + +Module Exports. +Coercion base : class_of >-> PCM.mixin_of. +Coercion mixin : class_of >-> mixin_of. +Coercion sort : type >-> Sortclass. +Bind Scope pcm_scope with sort. +Coercion pcmType : type >-> PCM.type. +Canonical pcmType. +Notation sub_pcm := type. +Notation subPCMMix := (@Mixin _ _). +Notation subPCM T m := (@pack _ _ T _ m _ id). + +Definition pval V D (U : @sub_pcm V D) : {morphism (@sepT U) >-> V} := + pval_op (mixin (class U)). +Definition psub V D (U : @sub_pcm V D) : {morphism D >-> U} := + psub_op (mixin (class U)). + +Arguments pval {V D U}. +Arguments psub {V D}. +Prenex Implicits pval psub. + +Section Lemmas. +Variables (V : pcm) (D : rel V) (U : sub_pcm D). +Notation pval := (@pval V D U). +Notation psub := (@psub V D U). + +Lemma valid_psubU (x y : V) : + valid (psub x \+ psub y) = valid (x \+ y) && D x y. +Proof. +case: U x y=>U' [b][v s H1] *. +by apply/idP/idP; [apply: H1 | case/andP; apply: pfV2]. +Qed. + +Lemma psub_pval (x : U) : psub (pval x) = x. +Proof. by case: U x=>U' [b][v s H1 H2 H3 H4] T; apply: H2. Qed. + +Lemma pval_psub x : valid x -> D x Unit -> pval (psub x) = x. +Proof. by case: U x=>U' [b][v s H1 H2 H3 H4] T; apply: H3. Qed. + +(* derived lemmas *) + +Lemma pval_inj : injective pval. +Proof. by move=>x y E; rewrite -[x]psub_pval -[y]psub_pval E. Qed. + +(* this is a weaker version of valid_psubU *) +(* we can prove valid_sep from valid_psubU, but not vice-versa *) +Lemma valid_sep (x y : U) : + valid (x \+ y) = + valid (pval x \+ pval y) && D (pval x) (pval y). +Proof. by rewrite -valid_psubU !psub_pval. Qed. + +Lemma valid_sep1 (x : U) : valid x = valid (pval x) && D (pval x) Unit. +Proof. by rewrite -[x]unitR valid_sep pfunit !unitR. Qed. + +(* thus, if we limit to single variable, we get the following *) +Lemma valid_pval x : valid (pval x) = valid x. +Proof. +apply/idP/idP; first by case: U x=>U' [b][v s H1 H2 H3 H4] T; apply: H4. +move=> Vx; move: (Vx); rewrite -[x]unitL pfjoin //; last by rewrite unitL. +by rewrite valid_sep => /andP[]. +Qed. + +(* however, valid_sep only really extracts the part of D that *) +(* applies to both arguments. if D doesn't have such part *) +(* but simply restricts x and y separately, then valid_sep *) +(* can be weakened still *) +Lemma valid_pval2 (x y : U) : + D (pval x) (pval y) = D (pval x) Unit && D (pval y) Unit -> + valid (x \+ y) = valid (pval x \+ pval y). +Proof. +move=>E; rewrite valid_sep {}E; case W : (valid _)=>//=. +move: (validL W) (validR W); rewrite !valid_pval !valid_sep1. +by case/andP=>_ -> /andP [_ ->]. +Qed. + +Lemma valid_psub x : valid (psub x) = valid x && D x Unit. +Proof. by rewrite -{2}[x]unitR -valid_psubU pfunit unitR. Qed. + +Lemma valid_psub1 x : valid (psub x) -> valid x. +Proof. by rewrite valid_psub; case/andP. Qed. + +Lemma valid_psub2 x : valid (psub x) -> D x Unit. +Proof. by rewrite valid_psub; case/andP. Qed. + +Lemma pval_psub1 x : valid (psub x) -> pval (psub x) = x. +Proof. by rewrite valid_psub=>/andP [H1 H2]; apply: pval_psub. Qed. + +(* this is a convenient composition of the above *) +Lemma valid_sep1P (x : U) : valid x -> D (pval x) Unit. +Proof. by rewrite valid_sep1=>/andP []. Qed. + +(* this just uses iff instead of equality *) +(* i keep both lemmas for convenience (valid_pvalE can use views) *) +Lemma valid_pvalE x : valid (pval x) <-> valid x. +Proof. by split; rewrite valid_pval. Qed. + +End Lemmas. + +Arguments valid_psubU [V D] U. +Arguments valid_psub [V D] U. +Arguments valid_psub1 [V D] U. +Arguments valid_psub2 [V D] U. +Prenex Implicits valid_psubU valid_psub valid_psub1 valid_psub2. + +End Exports. +End SubPCM. + +Export SubPCM.Exports. + +(* In the case of TPCMs *) + +Lemma psub_undef (V : tpcm) (D : rel V) (U : sub_pcm D) : ~~ valid (psub U undef). +Proof. by rewrite valid_psub tpcmE. Qed. + + +(* specific construction of a sub-pcm obtained *) +(* by modding out with a separating relation *) +(* but requires starting with a tpcm *) + +Module SepSubPCM. +Section SepSubPCM. +Variables (U : tpcm) (D : sep_rel U). + +Notation orth1 x := (valid x && D x Unit). +Notation orth2 x y := (valid (x \+ y) && D x y). + +(* Constructing the subset type that we care about *) + +Definition xsepP x := orth1 x \/ x = undef. +Inductive xsep := ex_sep x of xsepP x. + +Definition xval x := let: ex_sep v _ := x in v. + +Lemma xsep_eq x y pfx pfy : x = y -> @ex_sep x pfx = @ex_sep y pfy. +Proof. by move=>E; subst y; congr ex_sep; apply: pf_irr. Qed. + +Lemma xval_inj x y : xval x = xval y -> x = y. +Proof. by case: x y=>x Hx [y Hy] /= E; subst y; rewrite (pf_irr Hx). Qed. + +Lemma xsep_unitP : xsepP (Unit : U). +Proof. by rewrite /xsepP valid_unit sep00; left. Qed. + +Definition xsep_valid x := locked (orth1 (xval x)). +Definition xsep_unit := locked (ex_sep xsep_unitP). + +Definition xsep_join' x y := + if orth2 (xval x) (xval y) then xval x \+ xval y else undef. + +Lemma xsep_joinP x y : xsepP (xsep_join' x y). +Proof. +rewrite /xsepP /xsep_join'; case: ifP; last by right. +by case/andP=>V /(sepU0 V) ->; rewrite V; left. +Qed. + +Definition xsep_join x y := locked (ex_sep (xsep_joinP x y)). + +(* the subset type is actually a pcm *) + +Lemma xsep_joinC : commutative xsep_join. +Proof. +case=>x Hx [y Hy]; apply: xval_inj; rewrite /= /xsep_join -!lock /xsep_join' /=. +by rewrite joinC; case V: (valid _)=>//=; rewrite -sepC. +Qed. + +Lemma xsep_joinAC : right_commutative xsep_join. +Proof. +case=>a Ha [b Hb][c Hc]; apply: xval_inj; rewrite /= /xsep_join -!lock /xsep_join' /=. +case Sab: (orth2 a b); case Sac: (orth2 a c); rewrite ?tpcmE //=; last first. +- case/andP: Sac=>_ Sac; case: andP=>//; case=>V Sacb. + by rewrite (sepAxx V Sac Sacb) andbT (validLE3 V) in Sab. +- case/andP: Sab=>_ Sab; case: andP=>//; case=>V Sabc. + by rewrite (sepAxx V Sab Sabc) andbT (validLE3 V) in Sac. +case/andP: Sab=>_ Sab; case/andP: Sac=>_ Sac. +case Sabc: (orth2 (a \+ b) c). +- case/andP: Sabc=>V Sabc; rewrite sepC (joinAC a c b) V //. + by rewrite (sepAxx V Sab Sabc). +case Sacb: (orth2 (a \+ c) b)=>//. +case/andP: Sacb=>V Sacb; rewrite sepC (joinAC a b c) V // in Sabc. +by rewrite (sepAxx V Sac Sacb) in Sabc. +Qed. + +Lemma xsep_joinA : associative xsep_join. +Proof. by move=>a b c; rewrite !(xsep_joinC a) xsep_joinAC. Qed. + +Lemma xsep_unitL : left_id xsep_unit xsep_join. +Proof. +case=>x qf; apply: xval_inj; rewrite /= /xsep_join/xsep_join'/xsep_unit -!lock /=. +rewrite unitL; case: qf=>[|->]; last by rewrite tpcmE. +by case/andP=>V E; rewrite sepC ?unitL // V E. +Qed. + +Lemma xsep_validL a b : xsep_valid (xsep_join a b) -> xsep_valid a. +Proof. +rewrite /xsep_valid /xsep_join -!lock /= /xsep_join'. +case: ifP; last by rewrite tpcmE. +by case/andP=>V E; rewrite !(validE2 V) (sepx0 V E). +Qed. + +Lemma xsep_valid_unit : xsep_valid xsep_unit. +Proof. by rewrite /xsep_valid/xsep_unit /= -!lock valid_unit sep00. Qed. + +(* the pcm *) +Definition xsepPCMMix := + PCMMixin xsep_joinC xsep_joinA xsep_unitL xsep_validL xsep_valid_unit. +Canonical xsepPCM := Eval hnf in PCM _ xsepPCMMix. + +(* the topped pcm *) + +Definition xsep_unitb x := unitb (xval x). + +Lemma xsep_undefP : xsepP undef. +Proof. by right. Qed. + +Definition xsep_undef : xsepPCM := locked (ex_sep xsep_undefP). + +Lemma xsep_unitbP x : reflect (x = Unit) (xsep_unitb x). +Proof. +rewrite /Unit /= /xsep_unit -!lock /xsep_unitb; case: x=>x H /=. +case: unitbP=>X; constructor; last by case=>/X. +by rewrite X in H *; rewrite (pf_irr H xsep_unitP). +Qed. + +Lemma xsep_valid_undef : ~~ valid xsep_undef. +Proof. by rewrite pcmE /= /xsep_valid /= /xsep_undef -!lock /= tpcmE. Qed. + +Lemma xsep_undef_join x : xsep_undef \+ x = xsep_undef. +Proof. +apply: xval_inj; rewrite /= pcm_joinE /xsep_undef /=. +by rewrite /xsep_join -!lock /= /xsep_join' !tpcmE. +Qed. + +Definition xsepTPCMMix := + TPCMMixin xsep_unitbP xsep_valid_undef xsep_undef_join. +Canonical xsepTPCM := Eval hnf in TPCM xsep xsepTPCMMix. + + +(* and we have the sub_pcm instance *) + +Lemma xval_morphP : morph_axiom (sepT_seprel xsepPCM) xval. +Proof. +split=>[|x y]; first by rewrite pcmE /= /xsep_unit /= -!lock. +rewrite {1}/valid /= /xsep_valid /= pcm_joinE /= +/xsep_join -!lock /xsep_join' /=; +by case: ifP; rewrite ?tpcmE //; case/andP. +Qed. + +Canonical xval_pmorph := Morphism' xval xval_morphP. + +Definition xpsub x := + if decP (@idP (orth1 x)) is left pf + then ex_sep (or_introl pf) else ex_sep (or_intror (erefl _)). + +Lemma xpsub_morphP : morph_axiom D xpsub. +Proof. +rewrite /xpsub; split=>[|x y V E]. +- by apply: xval_inj; case: decP=>//; rewrite pfunit /= valid_unit sep00. +case: decP=>Hx; last by rewrite (sep0E V E) (validE2 V) in Hx. +case: decP=>Hy; last by rewrite (sep0E V E) (validE2 V) in Hy. +case: decP=>H; last by rewrite V (sepU0 V E) in H. +rewrite /valid /= /xsep_valid /= !pcm_joinE /= /xsep_join /= -!lock /= /xsep_join' /=. +rewrite {1 2}V {1 2}E {1}V {1}(sepU0 V E) /=. +by split=>//; apply: xval_inj; rewrite /= /xsep_join' V E. +Qed. + +Canonical xpsub_pmorph := Morphism' xpsub xpsub_morphP. + +Lemma valid_xpsubU x y : valid (xpsub x \+ xpsub y) -> orth2 x y. +Proof. +rewrite /xpsub {1}/valid /= -?lock /xsep_valid /= !pcm_joinE /= /xsep_join + -!lock /xsep_join' /= -pcmE. +by case: decP=>Hx; case: decP=>Hy; case H: (orth2 x y) Hx Hy; +rewrite /= ?tpcmE //=; case/andP: H=>V H. +Qed. + +Lemma xpsub_xval x : xpsub (xval x) = x. +Proof. by apply: xval_inj; rewrite /xpsub; case: decP; case: x=>// x []. Qed. + +Lemma xval_xpsub x : valid x -> D x Unit -> xval (xpsub x) = x. +Proof. by rewrite /xpsub /= => V H; case: decP=>//=; rewrite V H. Qed. + +Lemma xvalid_pval x : valid (xval x) -> valid x. +Proof. +rewrite {2}/valid /= /xsep_valid /= -!lock. +by case: x=>x /= [/andP [->->]|] // ->; rewrite tpcmE. +Qed. + +Definition xsepSubMix := subPCMMix valid_xpsubU xpsub_xval xval_xpsub xvalid_pval. +Canonical xsepSubPCM : sub_pcm D := Eval hnf in subPCM xsepPCM xsepSubMix. + +End SepSubPCM. + +Module Exports. +Notation xsepPCM := xsepPCM. +Notation xsepTPCM := xsepTPCM. +Notation xsepSubPCM := xsepSubPCM. +Canonical xsepPCM. +Canonical xsepTPCM. +Canonical xsepSubPCM. +Canonical xval_pmorph. +Canonical xpsub_pmorph. + +Prenex Implicits xval xpsub. +Notation xval := xval. +Notation xpsub := xpsub. + +Section Extras. +Variables (U : tpcm) (D : sep_rel U). + +Lemma xval_undef : xval (@undef (xsepTPCM D)) = undef. +Proof. by case: U D=>U' H' D'; rewrite /undef /= /xsep_undef /= -!lock. Qed. + +Lemma xsub_undef : xpsub D undef = undef. +Proof. by rewrite -xval_undef xpsub_xval. Qed. + +Lemma xsepP (x : xsepSubPCM D) : x = undef \/ valid x. +Proof. +rewrite (valid_sep1 x); case: x=>x [H|H]; [right=>//| left]. +by subst x; apply: pval_inj; rewrite /= xval_undef. +Qed. + +End Extras. + +End Exports. +End SepSubPCM. + + +(* We want to keep the definition of xsep abstract to improve performance *) +Module Type SepSubSig. +Parameter xsep : forall U : tpcm, sep_rel U -> Type. +Parameter xsepPCMMix : forall (U : tpcm) (D : sep_rel U), + PCM.mixin_of (xsep D). +Canonical xsepPCM U D := + Eval hnf in PCM (@xsep U D) (@xsepPCMMix U D). +Parameter xsepTPCMMix : forall (U : tpcm) (D : sep_rel U), + TPCM.mixin_of (@xsepPCM U D). +Canonical xsepTPCM U D := + Eval hnf in TPCM (@xsep U D) (@xsepTPCMMix U D). +Parameter xsepSubMix : forall (U : tpcm) (D : sep_rel U), + SubPCM.mixin_of D (xsepPCM D). +Canonical xsepSubPCM (U : tpcm) (D : sep_rel U) : sub_pcm D := + Eval hnf in subPCM (@xsepPCM U D) (@xsepSubMix U D). +Parameter xsep_undef : forall (U : tpcm) (D : sep_rel U), + @pval U D (@xsepSubPCM U D) (@undef (@xsepTPCM U D)) = undef. + +Parameter xsepP : forall (U : tpcm) (D : sep_rel U) + (x : @xsepSubPCM U D), x = undef \/ valid x. +End SepSubSig. + +Module SepSub : SepSubSig. +Section SepSub. +Variables (U : tpcm) (D : sep_rel U). +Definition xsep := @SepSubPCM.xsep U D. +Definition xsepPCMMix := @SepSubPCM.xsepPCMMix U D. +Canonical xsepPCM := Eval hnf in PCM xsep xsepPCMMix. +Definition xsepTPCMMix := @SepSubPCM.xsepTPCMMix U D. +Canonical xsepTPCM := Eval hnf in TPCM xsep xsepTPCMMix. +Definition xsepSubMix := @SepSubPCM.xsepSubMix U D. +Canonical xsepSubPCM : sub_pcm D := + Eval hnf in subPCM xsepPCM xsepSubMix. +Lemma xsep_undef : pval (U:=xsepSubPCM) undef = undef. +Proof. by apply: SepSubPCM.Exports.xval_undef. Qed. +Lemma xsepP (x : xsepSubPCM) : x = undef \/ valid x. +Proof. by apply: SepSubPCM.Exports.xsepP. Qed. +End SepSub. +End SepSub. + +(* Furthermore, locking U and D improves performance still *) +(* The following is a version which defines canonicals for *) +(* the xsep constructor *) + +Module SepSub2. + +Notation xsep D := + (@SepSub.xsep (locked _) + (cast (fun U : tpcm => sep_rel U) (esym (lock _)) D)). + +Section SepSubExports. +Variables (U : tpcm) (D : sep_rel U). + +Lemma eq1 : xsep D = SepSub.xsep D. +Proof. by move: (esym _); rewrite -lock=>eq1'; rewrite eqc. Qed. + +Definition xsepPCMMix := cast PCM.mixin_of eq1 (SepSub.xsepPCMMix D). +Canonical xsepPCM := Eval hnf in PCM (xsep D) xsepPCMMix. + +Lemma eq2 : xsepPCM = SepSub.xsepPCM D. +Proof. +rewrite /xsepPCM/xsepPCMMix; move: eq1. +by rewrite eq1=>eq1'; rewrite !eqc. +Qed. + +Definition xsepTPCMMix := cast TPCM.mixin_of eq2 (SepSub.xsepTPCMMix D). +Canonical xsepTPCM := Eval hnf in TPCM (xsep D) xsepTPCMMix. + +Lemma eq3 : (D, xsepTPCM) = (D, SepSub.xsepTPCM D). +Proof. +rewrite /xsepTPCM/SepSub.xsepTPCM/xsepTPCMMix/xsepPCMMix; move: eq2. +rewrite /xsepPCM/xsepPCMMix; move: eq1. +by rewrite !eq1=>eq1'; rewrite !eqc=>eq2'; rewrite !eqc. +Qed. + +Definition xsepSubMix := + cast (fun W : sep_rel U * tpcm => SubPCM.mixin_of W.1 W.2) + eq3 + (SepSub.xsepSubMix D). +Canonical xsepSubPCM := Eval hnf in subPCM xsepPCM xsepSubMix. + +Lemma xsep_undef : pval (U:=xsepSubPCM) undef = undef. +Proof. +rewrite -(SepSub.xsep_undef D)/pval/undef/xsepSubPCM/xsepSubMix; move: eq3. +rewrite /xsepTPCM/xsepTPCMMix/xsepPCMMix; move: eq2. +rewrite /xsepPCM/xsepPCMMix; move: eq1. +by rewrite !eq1=>eq1'; rewrite !eqc=>eq2'; rewrite !eqc=>eq3'; rewrite !eqc. +Qed. + +Lemma xsepP (x : xsepSubPCM) : x = undef \/ valid x. +Proof. +move: x; rewrite /xsepSubPCM/undef/valid /= /xsepTPCMMix /xsepPCMMix /=; move: eq2. +rewrite /xsepPCM/xsepPCMMix; move: eq1. +rewrite !eq1=>eq1'; rewrite !eqc=>eq2'; rewrite !eqc. +by apply: SepSub.xsepP. +Qed. + +End SepSubExports. +End SepSub2. + + +Module SepSub3. + +Notation xsep D := (locked (@SepSub.xsep _ D)). + +Section SepSubExports. +Variables (U : tpcm) (D : sep_rel U). + +Lemma eq1 : xsep D = SepSub.xsep D. +Proof. by rewrite -lock. Qed. + +Definition xsepPCMMix := cast PCM.mixin_of eq1 (SepSub.xsepPCMMix D). +Definition xsepPCM := Eval hnf in PCM (xsep D) xsepPCMMix. + +Lemma eq2 : xsepPCM = SepSub.xsepPCM D. +Proof. +rewrite /xsepPCM/xsepPCMMix; move: eq1. +by rewrite eq1=>eq1'; rewrite !eqc. +Qed. + +Definition xsepTPCMMix := cast TPCM.mixin_of eq2 (SepSub.xsepTPCMMix D). +Definition xsepTPCM := Eval hnf in TPCM (xsep D) xsepTPCMMix. + +Lemma eq3 : (D, xsepTPCM) = (D, SepSub.xsepTPCM D). +Proof. +rewrite /xsepTPCM/SepSub.xsepTPCM/xsepTPCMMix/xsepPCMMix; move: eq2. +rewrite /xsepPCM/xsepPCMMix; move: eq1. +by rewrite !eq1=>eq1'; rewrite !eqc=>eq2'; rewrite !eqc. +Qed. + +Definition xsepSubMix := + cast (fun W : sep_rel U * tpcm => SubPCM.mixin_of W.1 W.2) + eq3 + (SepSub.xsepSubMix D). +Definition xsepSubPCM := Eval hnf in subPCM xsepPCM xsepSubMix. + +Lemma xsep_undef : pval (U:=xsepSubPCM) (@undef xsepTPCM) = undef. +Proof. +rewrite -(SepSub.xsep_undef D)/pval/undef/xsepSubPCM/xsepSubMix; move: eq3. +rewrite /xsepTPCM/xsepTPCMMix/xsepPCMMix; move: eq2. +rewrite /xsepPCM/xsepPCMMix; move: eq1. +by rewrite !eq1=>eq1'; rewrite !eqc=>eq2'; rewrite !eqc=>eq3'; rewrite !eqc. +Qed. + +Lemma xsepP (x : xsepSubPCM) : x = @undef xsepTPCM \/ valid x. +Proof. +move: x; rewrite /xsepSubPCM/undef/valid /= /xsepTPCMMix /xsepPCMMix /=; move: eq2. +rewrite /xsepPCM/xsepPCMMix; move: eq1. +rewrite !eq1=>eq1'; rewrite !eqc=>eq2'; rewrite !eqc. +by apply: SepSub.xsepP. +Qed. + +End SepSubExports. +End SepSub3. + +(* In general, typechecking is fastest if each individual *) +(* type defined using either of the 3 versions of xsep above, *) +(* is itself hidden using the module system. Thus, the xsep *) +(* type should probably never be used directly itself, without the hiding *) +(* so it doesn't matter which of the three versions we export *) + +(* We will export SepSub2 for those rare (or unexpected) cases *) +(* when we want to have canonicals on xsep, but we want to be *) +(* faster than the nakes SepSub *) + +Export SepSub2. + + +(* We can use the xsep construction to provide a product TPCM *) +(* which is *normal*; that is, has only undef as the invalid element. *) +(* The way to do so is to mod out the plain product (T)PCM by a *) +(* trivial sep relation, via the xsep construction. *) + +Module Type PairingSig. + +Parameter pprod : tpcm -> tpcm -> Type. + +Section PairingSig. +Variables U1 U2 : tpcm. +Parameter pprodPCMMix : PCM.mixin_of (pprod U1 U2). +Canonical pprodPCM := Eval hnf in PCM (pprod U1 U2) pprodPCMMix. +Parameter pprodTPCMMix : TPCM.mixin_of pprodPCM. +Canonical pprodTPCM := Eval hnf in TPCM (pprod U1 U2) pprodTPCMMix. +Parameter pprodSubMix : + SubPCM.mixin_of [seprel of @sepT (U1 * U2)] pprodPCM. +Canonical pprodSubPCM := + Eval hnf in subPCM (pprod U1 U2) pprodSubMix. +Parameter pprod_pval_undef : pval (@undef pprodTPCM) = undef. +Parameter pprodP : forall x : pprod U1 U2, x = undef \/ valid x. +End PairingSig. +End PairingSig. + + +Module Pairing : PairingSig. +Section Pairing. +Variables U1 U2 : tpcm. + +Local Definition Uraw := [tpcm of U1 * U2]. +Local Definition Draw := [seprel of @sepT Uraw]. + +Definition pprod := @SepSubPCM.xsep Uraw Draw. +Definition pprodPCMMix := @SepSubPCM.xsepPCMMix Uraw Draw. +Canonical pprodPCM := Eval hnf in PCM pprod pprodPCMMix. +Definition pprodTPCMMix := @SepSubPCM.xsepTPCMMix Uraw Draw. +Canonical pprodTPCM := Eval hnf in TPCM pprod pprodTPCMMix. +Definition pprodSubMix := @SepSubPCM.xsepSubMix Uraw Draw. +Canonical pprodSubPCM : sub_pcm Draw := + Eval hnf in subPCM pprod pprodSubMix. + +Lemma pprod_pval_undef : pval (@undef pprodTPCM) = undef. +Proof. by apply: SepSubPCM.Exports.xval_undef. Qed. + +Lemma pprodP (x : pprod) : x = undef \/ valid x. +Proof. by apply: SepSubPCM.Exports.xsepP. Qed. + +End Pairing. +End Pairing. + +Export Pairing. + +Section PairingLemmas. +Variables U1 U2 : tpcm. + +Notation U := (pprod U1 U2). + +Definition pfst : U -> _ := fst \o pval. +Canonical pfst_morph := [morphism of pfst]. + +Definition psnd : U -> _ := snd \o pval. +Canonical psnd_morph := [morphism of psnd]. + +Definition ppair : U1 * U2 -> U := psub (pprodSubPCM U1 U2). +Lemma ppair_morph_ax : morph_axiom (@sepT _) ppair. +Proof. by split; [apply: pfunit | apply: pfjoinV]. Qed. +Canonical ppair_morph := Morphism' ppair ppair_morph_ax. + +Lemma pairing_undef : + (pfst undef = undef) * + (psnd undef = undef) * + (forall x, ~~ valid x -> ppair x = undef). +Proof. +split; first by rewrite /pfst/psnd /= pprod_pval_undef. +case=>a b Vab; case: (pprodP (ppair (a, b)))=>//. +by rewrite valid_psub (negbTE Vab). +Qed. + +Lemma pairing_valid : + (forall x, valid (pfst x) = valid x) * + (forall x, valid (psnd x) = valid x) * + (forall x, valid (ppair x) = valid x). +Proof. +split=>[|x]. +- split=>x; (case: (pprodP x)=>[->|V]; first by rewrite pairing_undef !tpcmE); + by case/valid_pvalE/andP: V (V)=>/= E1 E2 ->; rewrite ?(E1,E2). +by rewrite valid_psub andbT. +Qed. + +Lemma pprojPV1 x : valid x = valid (pfst x) && valid (psnd x). +Proof. by rewrite valid_sep1 pcmPV andbT. Qed. + +Lemma pprojPV2 x y : + valid (x \+ y) = valid (pfst x \+ pfst y) && valid (psnd x \+ psnd y). +Proof. by rewrite valid_sep pcmPV andbT. Qed. + +Lemma ppairPV1 x : valid (ppair x) = valid x. +Proof. by rewrite valid_psub andbT. Qed. + +Lemma ppairPV2 x y : valid (ppair x \+ ppair y) = valid (x \+ y). +Proof. by rewrite valid_psubU andbT. Qed. + +Definition pprodPV := (pprojPV1, pprojPV2, ppairPV1, ppairPV2). + +Lemma pfst_ppair x : valid x -> pfst (ppair x) = x.1. +Proof. by move=>V; rewrite /pfst/= pval_psub. Qed. + +Lemma psnd_ppair x : valid x -> psnd (ppair x) = x.2. +Proof. by move=>V; rewrite /psnd/= pval_psub. Qed. + +Lemma pprod_eta x : x = ppair (pfst x, psnd x). +Proof. by rewrite -prod_eta psub_pval. Qed. + +Definition pprodPE := (pfst_ppair, psnd_ppair, pprodPV). + +End PairingLemmas. + +Prenex Implicits pfst psnd ppair. + + +(****************************************) +(* deprecated, to be removed eventually *) +(****************************************) + +Module SubPCMDeprecated. +Section ClassDef. +Variable V : pcm. + +Record mixin_of (sort : pcm) : Type := Mixin { + pval_op : sort -> V; + _ : forall x, valid x -> valid (pval_op x); + _ : pval_op Unit = Unit; + _ : forall x y, valid (x \+ y) -> pval_op (x \+ y) = pval_op x \+ pval_op y; + _ : forall x y, pval_op x = pval_op y -> x = y}. + +Notation class_of := mixin_of (only parsing). + +Structure sub_pcm : Type := Pack {sort : pcm; _ : class_of sort}. +Local Coercion sort : sub_pcm >-> pcm. + +Variables (U : pcm) (sU : sub_pcm). + +Definition class := let: Pack _ c as sU' := sU return class_of sU' in c. +Definition pack c := @Pack U c. +Definition clone := fun c & sU -> U & phant_id (pack c) sU => pack c. + +Definition pval := locked (pval_op class). + +End ClassDef. + +Module Exports. +Coercion sort : sub_pcm >-> pcm. +Notation sub_pcm_deprec := sub_pcm. +Notation pval_deprec := pval. +Notation subPCMMix_deprec := Mixin. +Notation subPCM_deprec m := (@pack _ _ m). + +Prenex Implicits pval_deprec. + +Section Lemmas. +Variables (V : pcm) (U : sub_pcm V). + +Lemma pvalid_deprec (x : U) : valid x -> valid (pval x). +Proof. by rewrite /pval -lock; case: U x=>U' [pval H ???] x; apply: H. Qed. + +Lemma punit_deprec : pval (Unit : U) = Unit. +Proof. by rewrite /pval -lock; case: U=>U' [pval ? H ??]; apply: H. Qed. + +Lemma pjoin_deprec (x y : U) : valid (x \+ y) -> pval (x \+ y) = pval x \+ pval y. +Proof. by rewrite /pval -lock; case: U x y=>U' [pval ?? H ?] x y; apply: H. Qed. + +Lemma pval_inj_deprec (x y : U) : pval x = pval y -> x = y. +Proof. by rewrite /pval -lock; case: U x y=>U' [pval ???]; apply. Qed. + +Lemma pvalE_deprec (x : U) : pval x = pval_op (class U) x. +Proof. by rewrite /pval -lock. Qed. + +End Lemmas. +End Exports. +End SubPCMDeprecated. + +Export SubPCMDeprecated.Exports. + +(*************************************************************************) +(* Because the above interface does not specify the retraction, it's not *) +(* quite canonical for working with subPCM's. That is, we can have *) +(* different subPCM constructions, which will expose additional lemmas *) +(* that the subPCM interface doesn't. *) +(*************************************************************************) + +(* Here are two *) + +(* First, a construction that takes a closed subset of a PCM *) +(* leaving the join operation untouched *) + +(* do we want decidable P's? That can be a separate construction *) +Definition pcm_closed (U : pcm) (P : U -> Prop) := + P Unit /\ forall u v, P u -> P v -> P (u \+ v). + +Lemma pcl0 (U : pcm) (P : U -> Prop) : pcm_closed P -> P Unit. +Proof. by case. Qed. + +Lemma pcl_join (U : pcm) (P : U -> Prop) u v : + pcm_closed P -> P u -> P v -> P (u \+ v). +Proof. by case=>_; apply. Qed. + +Module CSubPCMDeprecated. +Section CSubPCM. +Variables (U : pcm) (P : U -> Prop) (pf : pcm_closed P). +Local Definition tp := {x : U | P x}. + +Definition svalid (x : tp) := let: exist v pf := x in valid v. + +Definition sjoin (x y : tp) : tp := + match x, y with exist u pfu, exist v pfv => + exist P _ (proj2 pf u v pfu pfv) + end. + +Definition sunit : tp := exist P _ (proj1 pf). + +Lemma sjoinC x y : sjoin x y = sjoin y x. +Proof. +case: x y=>x pfx [y pfy] /=; do 2![move: (proj2 pf _ _ _ _)]. +by rewrite joinC=>pfx2 pfy2; rewrite (pf_irr pfx2 pfy2). +Qed. + +Lemma sjoinA x y z : sjoin x (sjoin y z) = sjoin (sjoin x y) z. +Proof. +case: x y z=>x pfx [y pfy][z pfz] /=; do 2![move: (proj2 pf _ _ _ _)]. +by rewrite joinA=>pf1 pf2; rewrite (pf_irr pf1 pf2). +Qed. + +Lemma svalidL x y : svalid (sjoin x y) -> svalid x. +Proof. by case: x y=>x pfx [y pfy]; apply: validL. Qed. + +Lemma sunitL x : sjoin sunit x = x. +Proof. +case: x=>x pfx /=; move: pfx (proj2 pf _ _ _ _). +by rewrite unitL=>pf1 pf2; rewrite (pf_irr pf1 pf2). +Qed. + +Lemma svalidU : svalid sunit. +Proof. by apply: valid_unit. Qed. + +End CSubPCM. + +Module Exports. +Section Exports. +Variables (U : pcm) (P : U -> Prop) (pf : pcm_closed P). + +Definition csubPCMMixin := + PCMMixin (CSubPCMDeprecated.sjoinC pf) (CSubPCMDeprecated.sjoinA pf) + (CSubPCMDeprecated.sunitL pf) (@CSubPCMDeprecated.svalidL _ _ pf) + (CSubPCMDeprecated.svalidU pf). +Definition csub_pcm := Eval hnf in PCM {x : U | P x} csubPCMMixin. + +(* sub_pcm instance *) + +Lemma csubvalid (x : csub_pcm) : valid x -> valid (sval x). +Proof. by rewrite pcmE; case: x. Qed. + +Lemma csubunit : sval (Unit : csub_pcm) = Unit. +Proof. by rewrite pcmE. Qed. + +Lemma csubjoin (x y : csub_pcm) : + valid (x \+ y) -> sval (x \+ y) = sval x \+ sval y. +Proof. by rewrite 2!pcmE /=; case: x; case: y. Qed. + +Lemma csubpinj (x y : csub_pcm) : sval x = sval y -> x = y. +Proof. apply: sval_inj. Qed. + +Definition csubMix := subPCMMix_deprec csubvalid csubunit csubjoin csubpinj. +Canonical csub_subPCM : sub_pcm_deprec U := Eval hnf in subPCM_deprec csubMix. + +End Exports. +End Exports. +End CSubPCMDeprecated. + +Export CSubPCMDeprecated.Exports. + +(* products and closed subset subPCMs *) + +Section SubPCMProdDeprec. +Variables (U V : pcm) (P : prodPCM U V -> Prop) (pc : pcm_closed P). + +Lemma pcl1 : pcm_closed (fun u => exists v, P (u, v)). +Proof. +split=>[|u v [u1] Hu [v1] Hv]. +- by exists Unit; rewrite -pcmPE; apply: pcl0 pc. +by exists (u1 \+ v1); rewrite -pcmPE; apply: (pcl_join pc Hu Hv). +Qed. + +Lemma pcl2 : pcm_closed (fun v => exists u, P (u, v)). +Proof. +split=>[|u v [u1] Hu [v1] Hv]. +- by exists Unit; rewrite -pcmPE; apply: pcl0 pc. +by exists (u1 \+ v1); rewrite -pcmPE; apply: (pcl_join pc Hu Hv). +Qed. + +End SubPCMProdDeprec. diff --git a/pcm/mutex.v b/pcm/mutex.v index 2e23b65..13b502c 100644 --- a/pcm/mutex.v +++ b/pcm/mutex.v @@ -20,13 +20,10 @@ limitations under the License. (******************************************************************************) From Coq Require Import ssreflect ssrbool ssrfun. -From mathcomp Require Import ssrnat eqtype. -From fcsl Require Import prelude pcm. From Coq Require Setoid. - -Set Implicit Arguments. -Unset Strict Implicit. -Unset Printing Implicit Defensive. +From mathcomp Require Import ssrnat eqtype. +From fcsl Require Import prelude pcm morphism. +From fcsl Require Import options. (***********************) (* Generalized mutexes *) @@ -36,34 +33,32 @@ Module GeneralizedMutex. Section GeneralizedMutex. Variable T : Type. (* the mutex stages, excluding undef and unit *) -Inductive mutex := undef | nown | own | mx of T. +Inductive mutex := undef | nown | mx of T. -Definition join x y := +Definition join' x y := match x, y with undef, _ => undef | _, undef => undef | nown, x => x | x, nown => x - | own, _ => undef - | _, own => undef | mx _, mx _ => undef end. -Definition valid x := if x is undef then false else true. +Definition valid' x := if x is undef then false else true. -Lemma joinC : commutative join. -Proof. by case=>[|||x]; case=>[|||y]. Qed. +Lemma joinC : commutative join'. +Proof. by case=>[||x]; case=>[||y]. Qed. -Lemma joinA : associative join. -Proof. by case=>[|||x]; case=>[|||y]; case=>[|||z]. Qed. +Lemma joinA : associative join'. +Proof. by case=>[||x]; case=>[||y]; case=>[||z]. Qed. -Lemma unitL : left_id nown join. +Lemma unitL : left_id nown join'. Proof. by case. Qed. -Lemma validL x y : valid (join x y) -> valid x. +Lemma validL x y : valid' (join' x y) -> valid' x. Proof. by case: x. Qed. -Lemma valid_unit : valid nown. +Lemma valid_unit : valid' nown. Proof. by []. Qed. Definition mutexPCMMix := PCMMixin joinC joinA unitL validL valid_unit. @@ -71,24 +66,23 @@ Canonical mutexPCM := Eval hnf in PCM mutex mutexPCMMix. (* cancelativity *) -Lemma joinmK m1 m2 m : valid (m1 \+ m) -> m1 \+ m = m2 \+ m -> m1 = m2. -Proof. by case: m m1 m2=>[|||m][|||m1][|||m2]. Qed. +Lemma joinmK (m1 m2 m : mutexPCM) : valid (m1 \+ m) -> m1 \+ m = m2 \+ m -> m1 = m2. +Proof. by case: m m1 m2=>[||m][||m1][||m2]; rewrite !pcmE. Qed. Definition mutexCPCMMix := CPCMMixin joinmK. Canonical mutexCPCM := Eval hnf in CPCM mutex mutexCPCMMix. (* topped structure *) -Lemma unitb (x : mutex) : - reflect (x = Unit) (if x is nown then true else false). -Proof. by case: x; constructor. Qed. +Lemma unitb (x : mutex) : reflect (x = Unit) (if x is nown then true else false). +Proof. by case: x=>[||x]; constructor; rewrite !pcmE. Qed. Lemma join0E (x y : mutex) : x \+ y = Unit -> x = Unit /\ y = Unit. -Proof. by case: x y=>[|||x][|||y]. Qed. +Proof. by case: x y=>[||x][||y]. Qed. Lemma valid3 (x y z : mutex) : valid (x \+ y \+ z) = [&& valid (x \+ y), valid (y \+ z) & valid (x \+ z)]. -Proof. by case: x y z=>[|||x][|||y][|||z]. Qed. +Proof. by case: x y z=>[||x][||y][||z]. Qed. Lemma valid_undef : ~~ valid undef. Proof. by []. Qed. @@ -96,7 +90,7 @@ Proof. by []. Qed. Lemma undef_join x : undef \+ x = undef. Proof. by []. Qed. -Definition mutexTPCMMix := TPCMMixin unitb join0E valid_undef undef_join. +Definition mutexTPCMMix := TPCMMixin unitb valid_undef undef_join. Canonical mutexTPCM := Eval hnf in TPCM mutex mutexTPCMMix. End GeneralizedMutex. @@ -108,19 +102,19 @@ Definition mutex_eq (x y : mutex T) := match x, y with undef, undef => true | nown, nown => true - | own, own => true | mx x', mx y' => x' == y' | _, _ => false end. Lemma mutex_eqP : Equality.axiom mutex_eq. Proof. -case=>[|||x]; case=>[|||y] /=; try by constructor. +case=>[||x]; case=>[||y] /=; try by constructor. by case: eqP=>[->|H]; constructor=>//; case=>/H. Qed. Definition mutexEqMix := EqMixin mutex_eqP. Canonical mutexEqType := Eval hnf in EqType (mutex T) mutexEqMix. +Canonical mutexEQPCM := Eval hnf in EQPCM (mutex T) (mutexPCMMix T). End Equality. @@ -129,20 +123,24 @@ Canonical mutexPCM. Canonical mutexCPCM. Canonical mutexEqType. Canonical mutexTPCM. +Canonical mutexEQPCM. Notation mutex := mutex. -Notation mx := mx. Notation mx_undef := undef. Notation nown := nown. -Notation own := own. +Notation mx := mx. Notation mutexPCM := mutexPCM. Notation mutexTPCM := mutexTPCM. Arguments mx_undef {T}. Arguments nown {T}. -Arguments own {T}. - -(* binary mutexes can be obtained with T = Empty_set *) -Notation mtx := (mutex Empty_set). - +Arguments mx [T]. + +(* mutexes with distingusihed own element *) +Notation mtx T := (mutex (option T)). +Notation mtx2 := (mtx False). +Notation mtx3 := (mtx unit). +Notation own := (mx None). +Notation auth x := (mx (Some x)). +Notation auth1 := (mx (Some tt)). End Exports. End GeneralizedMutex. @@ -157,18 +155,17 @@ Implicit Types (t : T) (x y : mutex T). Variant mutex_spec x : mutex T -> Type := | mutex_undef of x = mx_undef : mutex_spec x mx_undef | mutex_nown of x = Unit : mutex_spec x Unit -| mutex_own of x = own : mutex_spec x own | mutex_mx t of x = mx t : mutex_spec x (mx t). Lemma mxP x : mutex_spec x x. -Proof. by case: x=>[|||t]; constructor. Qed. +Proof. by case: x=>[||t]; constructor. Qed. Lemma mxE0 x y : x \+ y = Unit -> (x = Unit) * (y = Unit). -Proof. by case: x=>[|||t1]; case: y=>[|||t2]. Qed. +Proof. exact: GeneralizedMutex.join0E. Qed. (* a form of cancelativity, more useful than the usual form *) Lemma cancelMx t1 t2 x : (mx t1 \+ x = mx t2) <-> (t1 = t2) * (x = Unit). -Proof. by case: x=>[|||x] //=; split=>//; case=>->. Qed. +Proof. by case: x=>[||x] //=; split=>//; case=>->. Qed. Lemma cancelxM t1 t2 x : (x \+ mx t1 = mx t2) <-> (t1 = t2) * (x = Unit). Proof. by rewrite joinC cancelMx. Qed. @@ -180,7 +177,7 @@ Lemma mxxM t x : valid (x \+ mx t) -> x = Unit. Proof. by rewrite joinC=>/mxMx. Qed. Lemma mxxyM t x y : valid (x \+ (y \+ mx t)) -> x \+ y = Unit. -Proof. by rewrite joinC joinAC=>/mxMx; rewrite joinC. Qed. +Proof. by rewrite joinA=>/mxxM. Qed. Lemma mxMxy t x y : valid (mx t \+ x \+ y) -> x \+ y = Unit. Proof. by rewrite -joinA=>/mxMx. Qed. @@ -191,52 +188,53 @@ Proof. by rewrite joinCA=>/mxMx. Qed. Lemma mxyMx t x y : valid (y \+ mx t \+ x) -> y \+ x = Unit. Proof. by rewrite joinAC=>/mxMx. Qed. -(* and the same for own *) +(* inversion principle for join *) +(* own and mx are prime elements, and unit does not factor *) +Variant mxjoin_spec (x y : mutex T) : mutex T -> mutex T -> mutex T -> Type := +| bothnown of x = nown & y = nown : mxjoin_spec x y nown nown nown +| leftmx t of x = mx t & y = nown : mxjoin_spec x y (mx t) (mx t) nown +| rightmx t of x = nown & y = mx t : mxjoin_spec x y (mx t) nown (mx t) +| invalid of ~~ valid (x \+ y) : mxjoin_spec x y undef x y. + +Lemma mxPJ x y : mxjoin_spec x y (x \+ y) x y. +Proof. by case: x y=>[||x][||y]; constructor. Qed. + +End MutexLemmas. + +Prenex Implicits mxMx mxxM mxxyM mxMxy mxxMy mxyMx. + +Section OwnMutex. +Variables T : Type. +Implicit Types x y : mtx T. Lemma mxOx x : valid (own \+ x) -> x = Unit. -Proof. by case: x. Qed. +Proof. by apply: mxMx. Qed. Lemma mxxO x : valid (x \+ own) -> x = Unit. -Proof. by rewrite joinC=>/mxOx. Qed. +Proof. by apply: mxxM. Qed. Lemma mxxyO x y : valid (x \+ (y \+ own)) -> x \+ y = Unit. -Proof. by rewrite joinC joinAC=>/mxOx; rewrite joinC. Qed. +Proof. by apply: mxxyM. Qed. Lemma mxOxy x y : valid (own \+ x \+ y) -> x \+ y = Unit. -Proof. by rewrite -joinA=>/mxOx. Qed. +Proof. by apply: mxMxy. Qed. Lemma mxxOy x y : valid (x \+ (own \+ y)) -> x \+ y = Unit. -Proof. by rewrite joinCA=>/mxOx. Qed. +Proof. by apply: mxxMy. Qed. Lemma mxyOx x y : valid (y \+ own \+ x) -> y \+ x = Unit. -Proof. by rewrite joinAC=>/mxOx. Qed. - -(* inversion principle for join *) -(* own and mx are prime elements, and unit does not factor *) -Variant mxjoin_spec (x y : mutex T) : - mutex T -> mutex T -> mutex T -> Type := -| bothnown of x = nown & y = nown : mxjoin_spec x y nown nown nown -| leftown of x = own & y = nown : mxjoin_spec x y own own nown -| rightown of x = nown & y = own : mxjoin_spec x y own nown own -| leftmx t of x = mx t & y = nown : mxjoin_spec x y (mx t) (mx t) nown -| rightmx t of x = nown & y = mx t : mxjoin_spec x y (mx t) nown (mx t) -| invalid of ~~ valid (x \+ y) : mxjoin_spec x y undef x y. - -Lemma mxPJ x y : mxjoin_spec x y (x \+ y) x y. -Proof. by case: x y=>[|||x][|||y]; constructor. Qed. - -End MutexLemmas. +Proof. by apply: mxyMx. Qed. +End OwnMutex. -Prenex Implicits mxMx mxxM mxxyM mxMxy mxxMy mxyMx - mxOx mxxO mxxyO mxOxy mxxOy mxyOx. +Prenex Implicits mxOx mxxO mxxyO mxOxy mxxOy mxyOx. (* specific lemmas for binary mutexes *) -Lemma mxON (x : mtx) : valid x -> x != own -> x = nown. -Proof. by case: x. Qed. +Lemma mxON (x : mtx2) : valid x -> x != own -> x = nown. +Proof. by case: x=>//; case. Qed. -Lemma mxNN (x : mtx) : valid x -> x != nown -> x = own. -Proof. by case: x. Qed. +Lemma mxNN (x : mtx2) : valid x -> x != nown -> x = own. +Proof. by case: x=>//; case=>//; case. Qed. (* the next batch of lemmas is for automatic simplification *) @@ -244,63 +242,121 @@ Section MutexRewriting. Variable T : eqType. Implicit Types (t : T) (x : mutex T). -Lemma mxE1 : ((@mx_undef T == nown) = false) * - ((@nown T == mx_undef) = false) * - ((@mx_undef T == own) = false) * - ((@own T == mx_undef) = false) * - ((@nown T == own) = false) * - ((@own T == nown) = false). +Lemma mxE1 : (((@mx_undef T == nown) = false) * + ((@nown T == mx_undef) = false)). Proof. by []. Qed. -Lemma mxE2 t : ((mx t == nown) = false) * - ((nown == mx t) = false) * - ((mx t == mx_undef) = false) * - ((mx_undef == mx t) = false) * - ((mx t == own) = false) * - ((own == mx t) = false). +Lemma mxE2 t : (((mx t == nown) = false) * + ((nown == mx t) = false)) * + (((mx t == mx_undef) = false) * + ((mx_undef == mx t) = false)). Proof. by []. Qed. -Lemma mxE3 t x : ((mx t \+ x == nown) = false) * - ((x \+ mx t == nown) = false) * - ((nown == mx t \+ x) = false) * - ((nown == x \+ mx t) = false) * - ((mx t \+ x == own) = false) * - ((x \+ mx t == own) = false) * - ((own \+ x == mx t) = false) * - ((x \+ own == mx t) = false). -Proof. by case: x. Qed. - -Lemma mxE4 x : - ((own \+ x == nown) = false) * - ((x \+ own == nown) = false) * - ((own \+ x == own) = (x == Unit)) * - ((x \+ own == own) = (x == Unit)). +Lemma mxE3 t x : ((((mx t \+ x == nown) = false) * + ((x \+ mx t == nown) = false)) * + (((nown == mx t \+ x) = false) * + ((nown == x \+ mx t) = false))). Proof. by case: x. Qed. Lemma mxE5 t1 t2 x : - ((mx t1 \+ x == mx t2) = (t1 == t2) && (x == Unit)) * - ((x \+ mx t1 == mx t2) = (t1 == t2) && (x == Unit)) * - ((mx t1 == mx t2 \+ x) = (t1 == t2) && (x == Unit)) * - ((mx t1 == x \+ mx t2) = (t1 == t2) && (x == Unit)). + (((mx t1 \+ x == mx t2) = (t1 == t2) && (x == Unit)) * + ((x \+ mx t1 == mx t2) = (t1 == t2) && (x == Unit))) * + (((mx t1 == mx t2 \+ x) = (t1 == t2) && (x == Unit)) * + ((mx t1 == x \+ mx t2) = (t1 == t2) && (x == Unit))). Proof. have L : forall t1 t2 x, (mx t1 \+ x == mx t2) = (t1 == t2) && (x == Unit). -- by move=>*; apply/idP/andP; [case/eqP/cancelMx=>->->|case=>/eqP->/eqP->]. +- by move=>*; apply/eqP/andP=>[/cancelMx[]->->|[]/eqP->/eqP->]. by do !split=>//; rewrite ?L // eq_sym L eq_sym. Qed. Lemma mx_valid t : valid (mx t). Proof. by []. Qed. -Lemma mx_valid_own : valid (@own T). -Proof. by []. Qed. - Lemma mx_injE t1 t2 : (mx t1 == mx t2) = (t1 == t2). Proof. by []. Qed. -Definition mxE := (mxE1, mxE2, mxE3, mxE4, mxE5, mx_injE, - mx_valid, mx_valid_own, +Definition mxE := ((((mxE1, mxE2), (mxE3)), ((mxE5, mx_injE), + (mx_valid))), (* plus a bunch of safe simplifications *) - unitL, unitR, valid_unit, eq_refl, - valid_undef, undef_join, join_undef). + (((unitL, unitR), (valid_unit, eq_refl)), + ((valid_undef, undef_join), join_undef))). End MutexRewriting. + +(* function mapping all mx's to own *) +Definition mxown T (x : mutex T) : mtx2 := + match x with + mx_undef => mx_undef + | nown => nown + | _ => own + end. + +(* this is a morphism under full domain *) +Lemma mxown_morph_ax T : morph_axiom (@sepT _) (@mxown T). +Proof. by split=>[|x y] //; case: x=>//; case: y. Qed. + +Canonical mxown_pmorph T := Morphism' (@mxown T) (mxown_morph_ax T). + +(* the key inversion property is the following *) +(* we could use simple case analysis in practice *) +(* but this appears often, so we might just as well have a lemma for it *) +Lemma mxownP T (x : mutex T) : mxown x = nown -> x = nown. +Proof. by case: x. Qed. + +Definition mxundef T (x : mtx T) : mtx2 := + match x with + | nown => nown + | mx None => own + | _ => undef + end. + +Definition sep_mxundef T (x y : mtx T) := + if x \+ y is (nown | mx None) then true else false. + +Lemma mxundef_seprel_ax T : seprel_axiom (@sep_mxundef T). +Proof. by split=>//; case=>[||x] // [||y] // [||[]] //=; case: y. Qed. + +Canonical mxundef_seprel T := + Eval hnf in seprel (@sep_mxundef T) (mxundef_seprel_ax T). + +Lemma mxundef_ax T : morph_axiom (@sep_mxundef T) (@mxundef T). +Proof. by split=>//; case=>[||x] // [||[]] //; case: x. Qed. + +Canonical mxundef_pmorph T := Morphism' (@mxundef T) (mxundef_ax T). + +Prenex Implicits mxown mxundef. + +(* inversion principle for mxundef *) +Variant mxundef_spec T (x : mtx T) : mtx2 -> mtx T -> Type := +| mxundef_nown of x = nown : mxundef_spec x nown nown +| mxundef_own of x = own : mxundef_spec x own own +| mxundef_undef of x = undef : mxundef_spec x undef undef +| mxundef_mx t of x = auth t : mxundef_spec x undef (auth t). + +Lemma mxundefP T (x : mtx T) : mxundef_spec x (mxundef x) x. +Proof. by case: x=>[||[t|]]; constructor. Qed. + +(* nats into mtx *) +(* notice this is not a morphism *) +Definition nxown n : mtx2 := if n is 0 then nown else own. + +Variant nxown_spec n : mtx2 -> Type := +| nxZ of n = 0 : nxown_spec n nown +| nxS m of n = m.+1 : nxown_spec n own. + +Lemma nxP n : nxown_spec n (nxown n). +Proof. by case: n=>[|n]; [apply: nxZ | apply: (@nxS _ n)]. Qed. + +Variant nxownjoin_spec n1 n2 : mtx2 -> Type := +| nxjZ of n1 = 0 & n2 = 0 : nxownjoin_spec n1 n2 nown +| nxjS m of n1 + n2 = m.+1 : nxownjoin_spec n1 n2 own. + +Lemma nxPJ n1 n2 : nxownjoin_spec n1 n2 (nxown (n1 \+ n2)). +Proof. +case: n1 n2=>[|n1][|n2] /=. +- by constructor. +- by apply: (@nxjS _ _ n2). +- by apply: (@nxjS _ _ n1); rewrite addn0 //. +apply: (@nxjS _ _ (n1 + n2).+1). +by rewrite addSn addnS. +Qed. diff --git a/pcm/natmap.v b/pcm/natmap.v index ea4dfcc..45c7b93 100644 --- a/pcm/natmap.v +++ b/pcm/natmap.v @@ -20,10 +20,9 @@ limitations under the License. From Coq Require Import ssreflect ssrbool ssrfun. From mathcomp Require Import ssrnat eqtype seq path. -From fcsl Require Import prelude ordtype pcm finmap unionmap. -Set Implicit Arguments. -Unset Strict Implicit. -Unset Printing Implicit Defensive. +From fcsl Require Import prelude pcm pred finmap mutex morphism. +From fcsl Require Export unionmap automap rtc. +From fcsl Require Import options. (************************) (************************) @@ -31,7 +30,10 @@ Unset Printing Implicit Defensive. (************************) (************************) +Notation nat_pred := (fun x : nat_ordType => x != 0). + Module Type NMSig. + Parameter tp : Type -> Type. Section Params. @@ -44,80 +46,76 @@ Parameter empty : tp. Parameter upd : nat -> A -> tp -> tp. Parameter dom : tp -> seq nat. Parameter dom_eq : tp -> tp -> bool. -Parameter free : nat -> tp -> tp. +Parameter assocs : tp -> seq (nat * A). +Parameter free : tp -> nat -> tp. Parameter find : nat -> tp -> option A. Parameter union : tp -> tp -> tp. -Parameter nm_filter : pred nat -> tp -> tp. Parameter empb : tp -> bool. Parameter undefb : tp -> bool. Parameter pts : nat -> A -> tp. -Parameter from : tp -> @UM.base nat_ordType A (fun x => x != 0). -Parameter to : @UM.base nat_ordType A (fun x => x != 0) -> tp. +Parameter from : tp -> @UM.base nat_ordType nat_pred A. +Parameter to : @UM.base nat_ordType nat_pred A -> tp. Axiom ftE : forall b, from (to b) = b. Axiom tfE : forall f, to (from f) = f. -Axiom undefE : nm_undef = to (@UM.Undef nat_ordType A (fun x => x != 0)). +Axiom undefE : nm_undef = to (@UM.Undef nat_ordType nat_pred A). Axiom defE : forall f, defined f = UM.valid (from f). -Axiom emptyE : empty = to (@UM.empty nat_ordType A (fun x => x != 0)). +Axiom emptyE : empty = to (@UM.empty nat_ordType nat_pred A). Axiom updE : forall k v f, upd k v f = to (UM.upd k v (from f)). Axiom domE : forall f, dom f = UM.dom (from f). Axiom dom_eqE : forall f1 f2, dom_eq f1 f2 = UM.dom_eq (from f1) (from f2). -Axiom freeE : forall k f, free k f = to (UM.free k (from f)). +Axiom assocsE : forall f, assocs f = UM.assocs (from f). +Axiom freeE : forall f k, free f k = to (UM.free (from f) k). Axiom findE : forall k f, find k f = UM.find k (from f). Axiom unionE : forall f1 f2, union f1 f2 = to (UM.union (from f1) (from f2)). -Axiom nmfiltE : forall q f, nm_filter q f = to (UM.um_filter q (from f)). Axiom empbE : forall f, empb f = UM.empb (from f). Axiom undefbE : forall f, undefb f = UM.undefb (from f). -Axiom ptsE : forall k v, pts k v = to (@UM.pts nat_ordType A (fun x => x != 0) k v). +Axiom ptsE : forall k v, pts k v = to (@UM.pts nat_ordType nat_pred A k v). End Params. End NMSig. - Module NMDef : NMSig. Section NMDef. Variable A : Type. -Definition nonz x := x != 0. - -Definition tp : Type := @UM.base nat_ordType A nonz. - -Definition nm_undef := @UM.Undef nat_ordType A nonz. -Definition defined f := @UM.valid nat_ordType A nonz f. -Definition empty := @UM.empty nat_ordType A nonz. -Definition upd k v f := @UM.upd nat_ordType A nonz k v f. -Definition dom f := @UM.dom nat_ordType A nonz f. -Definition dom_eq f1 f2 := @UM.dom_eq nat_ordType A nonz f1 f2. -Definition free k f := @UM.free nat_ordType A nonz k f. -Definition find k f := @UM.find nat_ordType A nonz k f. -Definition union f1 f2 := @UM.union nat_ordType A nonz f1 f2. -Definition nm_filter q f := @UM.um_filter nat_ordType A nonz q f. -Definition empb f := @UM.empb nat_ordType A nonz f. -Definition undefb f := @UM.undefb nat_ordType A nonz f. -Definition pts k v := @UM.pts nat_ordType A nonz k v. - -Definition from (f : tp) : @UM.base nat_ordType A nonz := f. -Definition to (b : @UM.base nat_ordType A nonz) : tp := b. +Definition tp : Type := @UM.base nat_ordType nat_pred A. + +Definition nm_undef := @UM.Undef nat_ordType nat_pred A. +Definition defined f := @UM.valid nat_ordType nat_pred A f. +Definition empty := @UM.empty nat_ordType nat_pred A. +Definition upd k v f := @UM.upd nat_ordType nat_pred A k v f. +Definition dom f := @UM.dom nat_ordType nat_pred A f. +Definition dom_eq f1 f2 := @UM.dom_eq nat_ordType nat_pred A f1 f2. +Definition assocs f := @UM.assocs nat_ordType nat_pred A f. +Definition free f k := @UM.free nat_ordType nat_pred A f k. +Definition find k f := @UM.find nat_ordType nat_pred A k f. +Definition union f1 f2 := @UM.union nat_ordType nat_pred A f1 f2. +Definition empb f := @UM.empb nat_ordType nat_pred A f. +Definition undefb f := @UM.undefb nat_ordType nat_pred A f. +Definition pts k v := @UM.pts nat_ordType nat_pred A k v. + +Definition from (f : tp) : @UM.base nat_ordType nat_pred A := f. +Definition to (b : @UM.base nat_ordType nat_pred A) : tp := b. Lemma ftE b : from (to b) = b. Proof. by []. Qed. Lemma tfE f : to (from f) = f. Proof. by []. Qed. -Lemma undefE : nm_undef = to (@UM.Undef nat_ordType A nonz). Proof. by []. Qed. +Lemma undefE : nm_undef = to (@UM.Undef nat_ordType nat_pred A). Proof. by []. Qed. Lemma defE f : defined f = UM.valid (from f). Proof. by []. Qed. -Lemma emptyE : empty = to (@UM.empty nat_ordType A nonz). Proof. by []. Qed. +Lemma emptyE : empty = to (@UM.empty nat_ordType nat_pred A). Proof. by []. Qed. Lemma updE k v f : upd k v f = to (UM.upd k v (from f)). Proof. by []. Qed. Lemma domE f : dom f = UM.dom (from f). Proof. by []. Qed. Lemma dom_eqE f1 f2 : dom_eq f1 f2 = UM.dom_eq (from f1) (from f2). Proof. by []. Qed. -Lemma freeE k f : free k f = to (UM.free k (from f)). Proof. by []. Qed. +Lemma assocsE f : assocs f = UM.assocs (from f). Proof. by []. Qed. +Lemma freeE f k : free f k = to (UM.free (from f) k). Proof. by []. Qed. Lemma findE k f : find k f = UM.find k (from f). Proof. by []. Qed. Lemma unionE f1 f2 : union f1 f2 = to (UM.union (from f1) (from f2)). Proof. by []. Qed. -Lemma nmfiltE q f : nm_filter q f = to (UM.um_filter q (from f)). -Proof. by []. Qed. Lemma empbE f : empb f = UM.empb (from f). Proof. by []. Qed. Lemma undefbE f : undefb f = UM.undefb (from f). Proof. by []. Qed. -Lemma ptsE k v : pts k v = to (@UM.pts nat_ordType A nonz k v). +Lemma ptsE k v : pts k v = to (@UM.pts nat_ordType nat_pred A k v). Proof. by []. Qed. End NMDef. @@ -128,8 +126,8 @@ Notation natmap := NMDef.tp. Definition natmapMixin A := UMCMixin (@NMDef.ftE A) (@NMDef.tfE A) (@NMDef.defE A) (@NMDef.undefE A) (@NMDef.emptyE A) (@NMDef.updE A) - (@NMDef.domE A) (@NMDef.dom_eqE A) (@NMDef.freeE A) - (@NMDef.findE A) (@NMDef.unionE A) (@NMDef.nmfiltE A) + (@NMDef.domE A) (@NMDef.dom_eqE A) (@NMDef.assocsE A) + (@NMDef.freeE A) (@NMDef.findE A) (@NMDef.unionE A) (@NMDef.empbE A) (@NMDef.undefbE A) (@NMDef.ptsE A). Canonical nat_mapUMC A := @@ -152,40 +150,59 @@ Definition nat_mapTPCMMix A := union_map_classTPCMMix (nat_mapUMC A). Canonical nat_mapTPCM A := Eval hnf in TPCM (natmap A) (@nat_mapTPCMMix A). Definition nat_map_eqMix (A : eqType) := - @union_map_class_eqMix nat_ordType A _ _ (@natmapMixin A). + @union_map_class_eqMix nat_ordType nat_pred A _ (@natmapMixin A). Canonical nat_map_eqType (A : eqType) := Eval hnf in EqType (natmap A) (@nat_map_eqMix A). +(* installing the Pred structure for writing x \In h *) +Canonical Structure natmap_PredType A : PredType (nat * A) := + Mem_UmMap_PredType (nat_mapUMC A). +Coercion Pred_of_natmap A (f : natmap A) : Pred_Class := [eta Mem_UmMap f]. + Definition nm_pts A (k : nat) (v : A) := - @UMC.pts nat_ordType A (@nat_mapUMC A) k v. + @UMC.pts nat_ordType nat_pred A (@nat_mapUMC A) k v. +(* baking nat_pred into the notation *) Notation "x \-> v" := (@nm_pts _ x v) (at level 30). +Definition nm_foldr V C a c d f := @um_foldr nat_ordType nat_pred V C (@nat_mapUMC V) a c d f. -Lemma nm_dom0 A (h : natmap A) : (h = um_undef \/ h = Unit) <-> (dom h = [::]). +Lemma uniq_dom0 A (h : natmap A) : uniq (0 :: dom h). +Proof. by rewrite /= uniq_dom cond_dom. Qed. + +Hint Resolve uniq_dom0 : core. + +Lemma nm_dom0 A (h : natmap A) : (h = undef \/ h = Unit) <-> (dom h = [::]). Proof. split=>[|E]; first by case=>->; rewrite ?dom_undef ?dom0. case V : (valid h); last by move/invalidE: (negbT V)=>->; left. by rewrite (dom0E V) ?E //; right. Qed. +Definition stamp := nat. + +Lemma In_dom0 A (h : natmap A) k e : (k, e) \In h -> k \in 0::dom h. +Proof. by move=>H; rewrite inE (In_dom H) orbT. Qed. + (***************************************) (***************************************) (* Constructions of last_key and fresh *) (***************************************) (***************************************) +Definition last_key {A} (h : natmap A) := last 0 (dom h). +Definition fresh {A} (h : natmap A) := (last_key h).+1. + +Prenex Implicits last_key fresh. + Section FreshLastKey. Variable A : Type. Implicit Type h : natmap A. -Definition last_key h := last 0 (dom h). -Definition fresh h := (last_key h).+1. +Lemma lastkey_undef : last_key (undef : natmap A) = 0. +Proof. by rewrite /last_key /undef /= !umEX. Qed. -Lemma lastkey_undef : last_key um_undef = 0. -Proof. by rewrite /last_key !umEX. Qed. - -Lemma lastkey0 : last_key Unit = 0. +Lemma lastkey0 : last_key (Unit : natmap A) = 0. Proof. by rewrite /last_key /Unit /= !umEX. Qed. Lemma lastkey_dom h : valid h -> last_key h \notin dom h -> h = Unit. @@ -205,8 +222,8 @@ right; apply: lastkey_dom=>//; rewrite E. by apply: contraT; rewrite negbK; apply: dom_cond. Qed. -Lemma dom_lastkey h : valid h -> ~~ empb h -> last_key h \in dom h. -Proof. by move=>X; apply: contraR; move/(lastkey_dom X)=>->; apply/empbP. Qed. +Lemma dom_lastkey h : valid h -> ~~ unitb h -> last_key h \in dom h. +Proof. by move=>X; apply: contraR; move/(lastkey_dom X)=>->; apply/unitbP. Qed. Lemma lastkeyxx h : valid h -> last_key h = 0 -> h = Unit. Proof. @@ -217,37 +234,59 @@ Lemma dom_lastkeyE h a : a < last_key h -> last_key h \in dom h. Proof. move=>H; case V : (valid h); last first. - by move/invalidE: (negbT V) H=>->; rewrite lastkey_undef. -by apply: dom_lastkey V (contraL _ H)=>/empbE ->; rewrite lastkey0. +by apply: dom_lastkey V (contraL _ H)=>/unitbE ->; rewrite lastkey0. Qed. Lemma lastkey_max h x : x \in dom h -> x <= last_key h. Proof. rewrite /last_key /= !umEX; case: (UMC.from h)=>//; case=>s H _ /=. rewrite /supp /ord /= (leq_eqVlt x) orbC. -by apply: sorted_last_key_max (sorted_oleq H). +by apply: sorted_last_key_maxR (sorted_oleq H)=>//; apply: otrans. Qed. Lemma max_lastkey h x : x \in dom h -> {in dom h, forall y, y <= x} -> last_key h = x. Proof. rewrite /last_key !umEX; case: (UMC.from h)=>//; case=>s H _ /=. -move=>H1 /= H2; apply: sorted_max_key_last (sorted_oleq H) H1 _. +move=>H1 /= H2; apply: sorted_max_key_last (sorted_oleq H) H1 _=>//. +- by apply: otrans. +- by apply: oantisym. by move=>z /(H2 z); rewrite leq_eqVlt orbC. Qed. -Lemma lastkeyPt (x : nat) v : x != 0 -> last_key (x \-> v) = x. +Lemma lastkey_leq h x : + {in dom h, forall y, y <= x} <-> last_key h <= x. +Proof. +split; last by move=>D y Y; apply: leq_trans D; apply: lastkey_max Y. +case V : (valid h); last first. +- by move/negbT/invalidE: V=>->; rewrite lastkey_undef. +case E: (unitb h); first by move/unitbP: E=>->; rewrite lastkey0. +by apply; rewrite dom_lastkey // E. +Qed. + +Lemma lastkey_ltn a h x : + a < last_key h -> {in dom h, forall y, y < x} <-> last_key h < x. +Proof. +move/dom_lastkeyE=>D; split; first by apply. +by move=>Y y X; apply: leq_ltn_trans Y; apply: lastkey_max. +Qed. + +Lemma lastkeyPt (x : nat) (v : A) : x != 0 -> last_key (x \-> v) = x. Proof. by rewrite /last_key domPtK /= => ->. Qed. -Lemma hist_path h : path oleq 0 (dom h). +Lemma hist_path B (h : natmap B) : path oleq 0 (dom h). Proof. -rewrite !umEX; case: (UMC.from h)=>// {h}-h /= _; case: h; case=>//= x s H. +rewrite !umEX; case: (UMC.from h)=>//= {}h _; case: h; case=>//= x s H. rewrite {1}/oleq /ord /= orbC -leq_eqVlt /=. by apply: sub_path H=>z y; rewrite /oleq=>->. Qed. -Lemma lastkey_mono h1 h2 : +Lemma lastkey_mono B (h1 : natmap A) (h2 : natmap B) : {subset dom h1 <= dom h2} -> last_key h1 <= last_key h2. -Proof. by rewrite leq_eqVlt orbC; apply: seq_last_mono; apply: hist_path. Qed. +Proof. +rewrite leq_eqVlt orbC; apply: (seq_last_monoR _ (@otrans nat_ordType))=>//; +by apply: hist_path. +Qed. Lemma lastkeyfUn h1 h2 : valid (h1 \+ h2) -> last_key h1 <= last_key (h1 \+ h2). @@ -271,11 +310,11 @@ Lemma lastkeyUn0 h1 h2 : last_key h1 = last_key h2 -> (h1 = Unit) * (h2 = Unit). Proof. move=>V E. -case E1: (empb h1). -- move/empbE: E1 E V=>->; rewrite unitL lastkey0. +case E1: (unitb h1). +- move/unitbE: E1 E V=>->; rewrite unitL lastkey0. by case/esym/dom_lastkey0P/nm_dom0=>-> //; rewrite valid_undef. -case E2: (empb h2). -- move/empbE: E2 E V=>->; rewrite unitR lastkey0. +case E2: (unitb h2). +- move/unitbE: E2 E V=>->; rewrite unitR lastkey0. by case/dom_lastkey0P/nm_dom0=>-> //; rewrite valid_undef. case: validUn (V)=>// _ _ /(_ _ (dom_lastkey (validL V) (negbT E1))). by rewrite E (dom_lastkey (validR V) (negbT E2)). @@ -285,7 +324,7 @@ Lemma lastkeyUn h1 h2 : last_key (h1 \+ h2) = if valid (h1 \+ h2) then maxn (last_key h1) (last_key h2) else 0. Proof. -have H k1 k2 : valid (k1 \+ k2) -> +have H (k1 k2 : natmap A) : valid (k1 \+ k2) -> last_key k1 < last_key k2 -> last_key (k1 \+ k2) = last_key k2. - move=>V H; apply: max_lastkey=>[|x]. - by rewrite domUn inE V (dom_lastkeyE H) orbT. @@ -300,27 +339,110 @@ by move/esym/(lastkeyUn0 V)=>E; rewrite !E unitL. Qed. Lemma lastkeyPtUn h t u : + last_key (t \-> u \+ h) = + if valid (t \-> u \+ h) then maxn t (last_key h) else 0. +Proof. +by rewrite lastkeyUn; case: ifP=>// /validPtUn_cond /= D; rewrite lastkeyPt. +Qed. + +Lemma lastkeyPtUn_valid h t u : valid h -> last_key h < t -> valid (t \-> u \+ h). Proof. move=>V L; rewrite validPtUn; apply/and3P; split=>//=. -- by rewrite -lt0n; apply: leq_ltn_trans L. +- by rewrite /= -lt0n; apply: leq_ltn_trans L. by apply: contraL L; move/lastkey_max; case: leqP. Qed. +(* monotonicity just gives us <= *) +(* when we remove the last key, we get a strict < *) +Lemma lastkeyF h : + last_key h \in dom h -> + last_key (free h (last_key h)) < last_key h. +Proof. +move=>D; case: (um_eta D)=>v [Ef] Eh. +have N : last_key h \notin dom (free h (last_key h)). +- by rewrite domF inE eq_refl. +have: last_key (free h (last_key h)) <= last_key h. +- by apply: lastkey_mono=>x; rewrite domF inE; case: ifP. +rewrite leq_eqVlt; case/orP=>// /eqP E; rewrite -{1}E in N. +have Dn : last_key h > 0 by move/dom_cond: D; case: (last_key h). +by rewrite -E in Dn; rewrite (dom_lastkeyE Dn) in N. +Qed. + +Lemma lastkeyPtUnE u2 u1 h t : + last_key (t \-> u1 \+ h) = last_key (t \-> u2 \+ h). +Proof. by rewrite !lastkeyPtUn !validPtUn. Qed. + +Lemma lastkeyUnPtE u2 u1 h t : + last_key (h \+ t \-> u1) = last_key (h \+ t \-> u2). +Proof. by rewrite !(joinC h); apply: lastkeyPtUnE. Qed. + +Lemma lastkeyU u h t : + t \in dom h -> last_key (upd t u h) = last_key h. +Proof. by case/um_eta=>v [_ E]; rewrite E updPtUn (lastkeyPtUnE v). Qed. + +Lemma lastkeyUmax h u t : + last_key (upd t u h) = + if t \in dom h then last_key h else + if valid (upd t u h) then maxn t (last_key h) else 0. +Proof. +case: ifP; first by apply: lastkeyU. +move=>H; rewrite validU /=; case: eqP =>[->|/eqP Nt /=]. +- by rewrite upd_condN // lastkey_undef. +case: ifP=>V; last first. +- by move/invalidE: (negbT V)=>->; rewrite upd_undef lastkey_undef. +have W : valid (upd t u h) by rewrite validU /= V Nt. +have E : upd t u h = t \-> u \+ h. +- apply: find_eta=>//; first by rewrite validPtUn /= V H Nt. +- by move=>k; rewrite findU /= Nt V findPtUn2 // validPtUn /= Nt V H. +by rewrite E lastkeyUn -E W lastkeyPt. +Qed. + (* freshness *) +Lemma fresh0N h x : fresh h <= x -> x != 0. +Proof. by case: eqP=>// ->. Qed. + Lemma dom_ordfresh h x : x \in dom h -> x < fresh h. Proof. by move/lastkey_max. Qed. +Lemma dom_freshN h x : x \in dom h -> x != fresh h. +Proof. by move/dom_ordfresh; case: ltngtP. Qed. + Lemma dom_freshn h n : fresh h + n \notin dom h. Proof. by apply: contra (@dom_ordfresh _ _) _; rewrite -leqNgt leq_addr. Qed. +Lemma ordfresh_dom0 h t : fresh h <= t -> t \notin 0::dom h. +Proof. +rewrite leqNgt; apply: contra; rewrite inE. +by case/orP=>[/eqP -> //|]; apply: dom_ordfresh. +Qed. + +Lemma ordfresh_dom h t : fresh h <= t -> t \notin dom h. +Proof. by rewrite leqNgt; apply: contra; apply: dom_ordfresh. Qed. + +Lemma dom_ordfresh0 h x t : fresh h <= t -> x \in 0::dom h -> x < t. +Proof. +by rewrite inE=>D /orP [/eqP ->|/dom_ordfresh H]; apply: leq_ltn_trans D. +Qed. + +Lemma dom_freshN0 h x t : fresh h <= t -> x \in 0::dom h -> x != t. +Proof. by move=>D /(dom_ordfresh0 D); case: ltngtP. Qed. + Lemma dom_freshUn h1 h2 : valid h1 -> [pcm h2 <= h1] -> fresh h1 \notin dom h2. Proof. move=>V [x E]; rewrite {h1}E in V *; apply: contra (@dom_ordfresh _ _) _. by rewrite joinC in V *; rewrite -leqNgt; apply: lastkeyUnf. Qed. +Lemma valid_ge_fresh h v t : + fresh h <= t -> valid (t \-> v \+ h) = valid h. +Proof. by move=>N; rewrite validPtUn /= (fresh0N N) (ordfresh_dom N) andbT. Qed. + +Lemma dom_freshUnN h1 h2 x : + valid h1 -> [pcm h2 <= h1] -> x \in dom h2 -> x != fresh h1. +Proof. by move=>V H; apply: contraL=>/eqP ->; apply: dom_freshUn. Qed. + Lemma dom_fresh h : fresh h \notin dom h. Proof. by move: (dom_freshn h 0); rewrite addn0. Qed. @@ -338,11 +460,9 @@ Lemma lastkey_freshUn h1 h2 v : valid h1 -> [pcm h2 <= h1] -> last_key (fresh h1 \-> v \+ h2) = fresh h1. Proof. -move=>V [x E]; rewrite {h1}E in V *. -apply: max_lastkey=>[|y] /=. -- by rewrite domUn inE valid_freshUn // (validL V) domPt inE eq_refl. -rewrite domUn inE valid_freshUn // (validL V) /= domPt inE /= eq_sym. -rewrite leq_eqVlt; case: eqP=>//= _ D. +move=>V [x E]; rewrite {h1}E in V *; apply: max_lastkey=>[|y] /=; + rewrite domUn inE valid_freshUn // (validE2 V) /= domPt inE ?eq_refl //. +rewrite andbC eq_sym leq_eqVlt; case: eqP=>//= _ D. by apply: lastkey_max; rewrite domUn inE V D. Qed. @@ -354,6 +474,16 @@ rewrite domUn inE /= valid_fresh Vf /= domPt inE /= eq_sym. by rewrite leq_eqVlt; case: eqP=>//= _; apply: dom_ordfresh. Qed. +Lemma dom_freshUnK h1 h2 v : + valid h1 -> [pcm h2 <= h1] -> + dom (fresh h1 \-> v \+ h2) = rcons (dom h2) (fresh h1). +Proof. +move=>V [x E]; subst h1; rewrite joinC domUnPtK //. +- by rewrite joinC valid_freshUn // (validL V). +apply/allP=>k /= D; apply: dom_ordfresh. +by rewrite domUn inE V D. +Qed. + (* pcm induced ordering *) Lemma umpleq_lastkey h1 h2 : @@ -368,18 +498,18 @@ Lemma valid_indb (P : natmap A -> Prop) : P Unit -> (forall u, P (1 \-> u)) -> (forall t u h, P h -> last_key h < t -> - valid (t \-> u \+ h) -> P (t \-> u \+ h)) -> + valid (h \+ t \-> u) -> P (t \-> u \+ h)) -> forall h, valid h -> P h. Proof. move=>H1 H2 H3; elim/um_indb=>//=. - by rewrite -[valid _]negbK; move/negP/invalidE. move=>k v f H V K _. -case E: (empb f); last first. -- apply: H3=>//; first by apply: H (validR V). +case E: (unitb f); last first. +- apply: H3=>//; first by apply: H (validL V). apply: K; apply: contraR (negbT E). - by rewrite -empbE; apply: lastkey_dom (validR V). -move/empbE: {K H} E V=>->; rewrite unitR => V. -move: (H3 k v Unit H1); rewrite unitR lastkey0 lt0n. + by rewrite -unitbE; apply: lastkey_dom (validL V). +move/unitbE: {K H} E V=>->; rewrite unitL unitR => V. +move: (H3 k v Unit H1); rewrite unitL unitR lastkey0 lt0n. by apply=>//; rewrite validPt /= in V. Qed. @@ -402,6 +532,55 @@ Qed. End FreshLastKey. +Lemma range_freshUnK (A : ordType) (h : natmap A) k (v : A) : + valid h -> last_key h < k -> + range (k \-> v \+ h) = rcons (range h) v. +Proof. +move=>V K; rewrite joinC rangeUnPtK //; last first. +- by apply/allP=>x /lastkey_max/leq_ltn_trans; apply. +rewrite validUnPt V /=; case: k K=>//= k K. +by apply: contraL K=>/lastkey_max; rewrite ltnS; case: ltngtP. +Qed. + +Lemma dom_fresh_sub A1 A2 (h1 : natmap A1) (h2 : natmap A2) : + valid h1 -> {subset dom h2 <= dom h1} -> fresh h1 \notin dom h2. +Proof. +move=>V S; apply: contra; first by move/S; apply: dom_ordfresh. +by rewrite ltnn. +Qed. + +Lemma valid_fresh_sub A1 A2 (h1 : natmap A1) (h2 : natmap A2) v : + valid h1 -> {subset dom h2 <= dom h1} -> + valid (fresh h1 \-> v \+ h2) = valid h2. +Proof. by move=>V S; rewrite validPtUn dom_fresh_sub // andbT. Qed. + +(* lemmas with omap *) + +Lemma lastkey_omap V1 V2 (f : nat * V1 -> option V2) (h : natmap V1) : + last_key (omap f h) <= last_key h. +Proof. by apply: lastkey_mono; apply: dom_omap_sub. Qed. + +Lemma lastkey_omap_some V1 V2 (f : nat * V1 -> option V2) (h : natmap V1) : + (forall x, x \In h -> oapp predT false (f x)) -> + last_key (omap f h) = last_key h. +Proof. by rewrite /last_key=>/dom_omap_some ->. Qed. + +(* renaming for freshness *) + +Lemma fresh_omap V1 V2 (f : nat * V1 -> option V2) (h : natmap V1) : + fresh (omap f h) <= fresh h. +Proof. by apply: lastkey_omap. Qed. + +Lemma fresh_omap_some V1 V2 (f : nat * V1 -> option V2) (h : natmap V1) : + (forall x, x \In h -> oapp predT false (f x)) -> + fresh (omap f h) = fresh h. +Proof. by move=>H; rewrite /fresh lastkey_omap_some. Qed. + +(* lemmas with In *) + +Lemma In_lastkey V k v (h : natmap V) : (k, v) \In h -> k <= last_key h. +Proof. by move/In_dom/lastkey_max. Qed. + (*******************) (*******************) @@ -415,7 +594,17 @@ Implicit Type h : natmap A. Definition gapless h := forall k, 0 < k <= last_key h -> k \in dom h. -Lemma gp_undef : gapless um_undef. +Definition gaplessb h := all (fun t => t \in dom h) (iota 1 (last_key h)). + +Lemma gaplessP h : reflect (gapless h) (gaplessb h). +Proof. +case V: (gaplessb h);constructor; last first. +- move=>H;apply/(elimF idP V)/allP=>?. + by rewrite mem_iota add1n ltnS=>?; apply/H. +by move/allP:V=>H t;move: (H t);rewrite mem_iota add1n ltnS. +Qed. + +Lemma gp_undef : gapless undef. Proof. by move=>k; rewrite lastkey_undef; case: k. Qed. Lemma gp0 : gapless Unit. @@ -424,7 +613,7 @@ Proof. by move=>k; rewrite lastkey0; case: k. Qed. Lemma gp_fresh h u : gapless (fresh h \-> u \+ h) <-> gapless h. Proof. case V : (valid h); last first. -- by move/invalidE: (negbT V)=>->; rewrite join_undefR. +- by move/invalidE: (negbT V)=>->; rewrite join_undef. split=>H k; move: (V); rewrite -(valid_fresh _ u)=>V'; last first. - rewrite lastkey_fresh // domPtUn inE V' /= (leq_eqVlt k) eq_sym. by move: (H k); rewrite -(ltnS k); case: (ltngtP k (last_key h).+1). @@ -440,7 +629,7 @@ Proof. move=>V C N; apply/eqP/contraT=>X. have Y : fresh h < k by rewrite leq_eqVlt eq_sym (negbTE X) in N. suff Z : last_key (k \-> u \+ h) = k. -- move: (C (fresh h)); rewrite Z (leq_eqVlt _ k) Y orbT andbT => /(_ erefl). +- move: (C (fresh h)); rewrite Z (leq_eqVlt _ k) Y orbT andbT; move/(_ (erefl _)). rewrite domPtUn inE (negbTE X) /=; case/andP=>_ /dom_ordfresh. by rewrite ltnn. apply: max_lastkey (find_some (findPtUn V)) _ => x. @@ -448,6 +637,23 @@ rewrite domUn inE; case/andP=>_ /orP []. - by rewrite domPt inE; case/andP=>_ /eqP ->. by move/lastkey_max/leq_ltn_trans/(_ N)/ltnW. Qed. + +Lemma gaplessPtUnE u2 u1 h k : + gapless (k \-> u1 \+ h) <-> gapless (k \-> u2 \+ h). +Proof. +rewrite /gapless (lastkeyPtUnE u2); split=>H t /H; +by rewrite !domPtUn !inE !validPtUn. +Qed. + +Lemma gapless_pleq h1 h2 : + gapless h1 -> valid h2 -> [pcm h1 <= h2] -> + exists h, h2 = h1 \+ h /\ forall k, k \in dom h -> last_key h1 < k. +Proof. +move=>G V H; case: H V=>h ->{h2} V; exists h; split=>// k D. +apply: contraR (dom_inNR V D)=>H; apply: G; move/dom_cond: D=>/= D. +by rewrite lt0n D leqNgt. +Qed. + End Gapless. Arguments gp_fresh [A][h] u. @@ -465,14 +671,21 @@ Implicit Type h : natmap (A * A). Definition atval v t h := oapp snd v (find t h). +Lemma atvalUn v t h1 h2 : + valid (h1 \+ h2) -> t <= last_key h1 -> + (forall k : nat, k \in dom h2 -> last_key h1 < k) -> + atval v t (h1 \+ h2) = atval v t h1. +Proof. +move=>V K H; rewrite /atval findUnR //. +by case: ifP=>// /H; rewrite ltnNge K. +Qed. + Lemma umpleq_atval v t h1 h2 : gapless h1 -> valid h2 -> [pcm h1 <= h2] -> t <= last_key h1 -> atval v t h2 = atval v t h1. Proof. -move=>G V H K; rewrite /atval; case E1 : (find t h1)=>[a|]. -- by rewrite (umpleq_some V H E1). -case: t K E1 => [|t] K E1; last by move: (G t.+1 K) (find_none E1)=>->. -by case E2 : (find 0 h2)=>//; move/find_some/dom_cond: E2. +move=>G V H K; case/(gapless_pleq G V): {V} H (V)=>h [->{h2} H] V. +by apply: atvalUn. Qed. Definition last_val v h := atval v (last_key h) h. @@ -514,8 +727,17 @@ rewrite lastkey_freshUn // findPtUn // valid_freshUn //. by case: H V=>h -> /validL. Qed. +Lemma lastval_indom v h : + last_val v h <> v -> last_key h \in dom h. +Proof. by rewrite /last_val /atval; case: dom_find=>// ->. Qed. + End AtVal. +(* in case A = bool *) +Lemma lastval_indomb h : + last_val false h -> last_key h \in dom h. +Proof. by move=>H; apply: (lastval_indom (v:=false)); rewrite H. Qed. + (* Continuous maps with binary range *) Section Continuous. @@ -525,7 +747,7 @@ Implicit Type h : natmap (A * A). Definition continuous v h := forall k x, find k.+1 h = Some x -> oapp snd v (find k h) = x.1. -Lemma cn_undef v : continuous v um_undef. +Lemma cn_undef v : continuous v undef. Proof. by move=>x w; rewrite find_undef. Qed. Lemma cn0 v : continuous v Unit. @@ -605,3 +827,2310 @@ End Complete. Prenex Implicits cm_valid cmPt. + +(************************) +(************************) +(************************) +(* Surgery on histories *) +(* using leq filtering *) +(************************) +(************************) +(************************) + +Notation le t := [pts k _ | k <= t]. +Notation lt t := [pts k _ | k < t]. + +Lemma pts_sub V t1 t2 : t1 <= t2 -> subpred (T:=nat*V) (le t1) (le t2). +Proof. by move=>T [k v] /leq_trans; apply. Qed. + +Lemma pts_sub_lt V t1 t2 : t1 < t2 -> subpred (T:=nat*V) (le t1) (lt t2). +Proof. by move=>T [k v] /leq_ltn_trans; apply. Qed. + +Lemma ptsD V t1 t2 : + t1 <= t2 -> predD (le t2) (le t1) =1 [pts k (v : V) | t1 < k <= t2]. +Proof. by move=>T [k v] /=; rewrite -ltnNge. Qed. + +Lemma ptsD_lt V t1 t2 : + t1 < t2 -> predD (lt t2) (le t1) =1 [pts k (v : V) | t1 < k < t2]. +Proof. by move=>T [k v] /=; rewrite -ltnNge. Qed. + +Lemma lastkey_umfilt_leq A (h : natmap A) t : last_key (um_filter (le t) h) <= t. +Proof. +case V : (valid h); last first. +- by move/negbT/invalidE: V=>->; rewrite umfilt_undef lastkey_undef. +have V' : valid (um_filter (le t) h) by rewrite valid_umfilt. +case E : (unitb (um_filter (le t) h)). +- by move/unitbP: E=>->; rewrite lastkey0. +by case/dom_umfilt: (dom_lastkey V' (negbT E))=>v []. +Qed. + +Lemma lastkey_umfilt_dom A (h : natmap A) t : + last_key (um_filter (le t) h) <= last_key h. +Proof. by apply: lastkey_mono; apply: dom_umfilt_subset. Qed. + +Lemma umfilt_le_last A (h : natmap A) t : + last_key h <= t -> um_filter (le t) h = h. +Proof. +case V : (valid h); last first. +- by move/invalidE: (negbT V)=>->; rewrite umfilt_undef. +move=>N; rewrite -{2}[h]umfilt_predT; apply/eq_in_umfilt. +by case=>k v /In_dom/lastkey_max/leq_trans; apply. +Qed. + +Lemma umfilt_last A (h : natmap A) : um_filter (le (last_key h)) h = h. +Proof. by apply: umfilt_le_last. Qed. + +Lemma umfilt_le_fresh A (h : natmap A) : um_filter (le (fresh h)) h = h. +Proof. by apply: umfilt_le_last; apply: ltnW. Qed. + +Lemma umfilt_le0 A (h : natmap A) t : + valid h -> {in dom h, forall k, t < k} -> um_filter (le t) h = Unit. +Proof. +move=>V D; rewrite -(umfilt_pred0 V). +apply/eq_in_umfilt; case=>k v [/= _][z E]; subst h. +rewrite leqNgt; apply: contraTF (D k _)=>//. +by rewrite domPtUn inE V eq_refl. +Qed. + +Lemma umfilt_le_split A (h : natmap A) t1 t2 : + t1 <= t2 -> + um_filter (le t2) h = + um_filter (le t1) h \+ um_filter [pts x _ | t1 < x <= t2] h. +Proof. +move=>T; rewrite -umfilt_dpredU; last first. +- by case=>x y N /=; rewrite negb_and -leqNgt N. +apply/eq_in_umfilt; case=>k v _ => /=. +by case: (leqP k t1)=>//= /leq_trans; apply. +Qed. + +Lemma umfilt_lt_split A (h : natmap A) t1 t2 k : + t1 <= k <= t2 -> + um_filter [pts x _ | t1 < x <= t2] h = + um_filter [pts x _ | t1 < x <= k] h \+ + um_filter [pts x _ | k < x <= t2] h. +Proof. +move=>/andP [T1 T2]; rewrite -umfilt_dpredU; last first. +- by case=>x y /andP [N1 N2]; rewrite /= negb_and -leqNgt N2. +apply/eq_in_umfilt; case=>k1 v1 _ /=. +case: (leqP k1 k)=>//=; last by move/(leq_ltn_trans T1)=>->. +by move/leq_trans=>-> //; rewrite orbF. +Qed. + +Lemma umfilt_pt_split A (h : natmap A) k v : + (k, v) \In h -> um_filter [pts x _ | k.-1 < x <= k] h = k \-> v. +Proof. +move=>H; have W : valid h by case: H. +have N: 0 < k by move/In_dom/dom_cond: H; case: k. +have subX : forall m n, 0 < n -> (m < n) = (m <= n.-1) by move=>? []. +rewrite (In_eta H) umfiltPtUn -(In_eta H) subX // W !leqnn /=. +rewrite umfilt_mem0L ?unitR ?validF //. +move=>k1 v1 /InF [_ /=]; rewrite andbC; case: (ltngtP k k1) =>//=. +by rewrite subX //; case: (ltngtP k1 k.-1). +Qed. + +Lemma umfilt_leUn A (h1 h2 : natmap A) t : + valid (h1 \+ h2) -> t <= last_key h1 -> + {in dom h2, forall k, k > last_key h1} -> + um_filter (le t) (h1 \+ h2) = um_filter (le t) h1. +Proof. +move=>V K D; rewrite umfiltUn // (umfilt_le0 (validR V)) ?unitR //. +by move=>k /D /(leq_ltn_trans K). +Qed. + +Lemma umfilt_le_gapless A (h1 h2 : natmap A) t : + gapless h1 -> valid h2 -> [pcm h1 <= h2] -> t <= last_key h1 -> + um_filter (le t) h2 = um_filter (le t) h1. +Proof. +move=>G V H K; case: (gapless_pleq G V H)=>h [? D]. +by subst h2; rewrite umfilt_leUn. +Qed. + +Lemma dom_umfilt_le_eq A (h : natmap A) t1 t2 : + t1 \in 0::dom h -> t2 \in 0::dom h -> + um_filter (le t1) h = um_filter (le t2) h -> + t1 = t2. +Proof. +rewrite !inE; case/orP=>[/eqP ->|/In_domX [v1 T1]]. +- case/orP=>[/eqP ->|/In_domX [v2 T2]] //. + move/eq_in_umfilt=>/(_ (t2, v2) T2). + by rewrite leqnn leqn0 eq_sym=>/eqP. +case/orP=>[/eqP ->|/In_domX [v2 T2] L]. +- move/eq_in_umfilt=>/(_ (t1, v1) T1). + by rewrite leqnn leqn0=>/esym/eqP. +move/eq_in_umfilt: (L)=>/(_ (t1, v1) T1). +move/eq_in_umfilt: (L)=>/(_ (t2, v2) T2). +by rewrite !leqnn; case: ltngtP. +Qed. + +Lemma eval_leUn A R a (h1 h2 : [pcm of natmap A]) t (z0 : R) : + valid (h1 \+ h2) -> t <= last_key h1 -> + {in dom h2, forall k, k > last_key h1} -> + eval a (le t) (h1 \+ h2) z0 = eval a (le t) h1 z0. +Proof. by move=>V K D; apply: eq_filt_eval; apply: umfilt_leUn. Qed. + +Lemma eval_le_gapless A R a (h1 h2 : [pcm of natmap A]) t (z0 : R) : + gapless h1 -> valid h2 -> [pcm h1 <= h2] -> t <= last_key h1 -> + eval a (le t) h2 z0 = eval a (le t) h1 z0. +Proof. by move=>G V H K; apply: eq_filt_eval; apply: umfilt_le_gapless. Qed. + +Lemma eval_le0 A R a (h : natmap A) (z0 : R) : eval a (le 0) h z0 = z0. +Proof. +case W : (valid h); last first. +- by move/negbT/invalidE: W=>->; rewrite eval_undef. +rewrite eval_umfilt umfilt_mem0L ?eval0 //. +by move=>k v /In_dom/dom_cond; rewrite /=; case: ltngtP. +Qed. + +Lemma eval_le_split A R a (h : natmap A) t1 t2 (z0 : R) : + t1 <= t2 -> + eval a (le t2) h z0 = + eval a [pts k v | t1 < k <= t2] h (eval a (le t1) h z0). +Proof. +move=>T; case V : (valid h); last first. +- by move/invalidE: (negbT V)=>->; rewrite !eval_undef. +rewrite eval_umfilt (umfilt_predD h (pts_sub T)) evalUn; last first. +- move=>x y /dom_umfilt [vx][X _] /dom_umfilt [wy][/= /andP][]. + by rewrite /= -ltnNge; move/(leq_ltn_trans X). +- by rewrite -(umfilt_predD h (pts_sub T)) valid_umfilt. +by rewrite -!eval_umfilt; apply: eq_in_eval=>kv _; apply: (ptsD T). +Qed. + +Lemma eval_lt_split A R a (h : natmap A) t1 t2 (z0 : R) : + t1 < t2 -> + eval a (lt t2) h z0 = + eval a [pts k v | t1 < k < t2] h (eval a (le t1) h z0). +Proof. +move=>T; case V : (valid h); last first. +- by move/invalidE: (negbT V)=>->; rewrite !eval_undef. +rewrite eval_umfilt (umfilt_predD h (pts_sub_lt T)) evalUn; last first. +- move=>x y /dom_umfilt [vx][X _] /dom_umfilt [wy][/= /andP][]. + by rewrite /= -ltnNge; move/(leq_ltn_trans X). +- by rewrite -(umfilt_predD h (pts_sub_lt T)) valid_umfilt. +by rewrite -!eval_umfilt; apply: eq_in_eval=>kv _; apply: (ptsD_lt T). +Qed. + +Lemma eval_le_lt_split A R a (h : natmap A) t (z0 : R) : + eval a (le t) h z0 = + eval a [pts k v | k == t] h (eval a (lt t) h z0). +Proof. +case V : (valid h); last first. +- by move/invalidE: (negbT V)=>->; rewrite !eval_undef. +have D : subpred (T:=nat*A) (lt t) (le t) by case=>k v /ltnW. +rewrite eval_umfilt (umfilt_predD h D) evalUn; last first. +- move=>x y /dom_umfilt [vx][X _] /dom_umfilt [wy][/= /andP][]. + by rewrite /= -ltnNge; move/(leq_ltn_trans X). +- by rewrite -(umfilt_predD h D) valid_umfilt. +rewrite -!eval_umfilt; apply: eq_in_eval; case=>k v _ /=. +by case: ltngtP. +Qed. + +Lemma eval_eq A R a (h : natmap A) t v (z0 : R) : + (t, v) \In h -> eval a [pts k _ | k == t] h z0 = a z0 t v. +Proof. +move=>H; rewrite eval_umfilt; have N : t != 0 by move/In_dom/dom_cond: H. +suff -> : um_filter [pts k _ | k == t] h = t \-> v by rewrite evalPt /= N. +apply/umem_eq=>[||[k w]]; first by rewrite valid_umfilt; case: H. +- by rewrite validPt. +rewrite In_umfilt. split; last by move/InPt =>[[->->]]. +by move=>[/eqP -> H1]; rewrite (In_fun H H1); apply: In_condPt. +Qed. + +Lemma eval_le_last A R a (h : natmap A) t (z0 : R) : + last_key h <= t -> eval a (le t) h z0 = eval a xpredT h z0. +Proof. +by move=>H; apply: eq_in_eval; case=>k v /In_dom/lastkey_max/leq_trans; apply. +Qed. + +Lemma eval_fresh_le A R a (h : natmap A) t v (z0 : R) : + eval a (le t) (fresh h \-> v \+ h) z0 = + if t <= last_key h then eval a (le t) h z0 else + if valid h then a (eval a predT h z0) (fresh h) v else z0. +Proof. +case V: (valid h); last first. +- by move/invalidE: (negbT V)=>->; rewrite join_undef !eval_undef; case: ifP. +case: ifP=>H. +- by rewrite eval_umfilt umfiltPtUn valid_fresh V ltnNge H -eval_umfilt. +rewrite joinC evalUnPt; last first. +- by apply/allP=>x; apply: lastkey_max. +- by rewrite joinC valid_fresh. +rewrite ltnNge H; congr a; apply: eq_in_eval. +case=>k w /In_dom/lastkey_max /=. +by case: ltngtP H=>// /ltnW H _ /leq_trans; apply. +Qed. + +Lemma eval_fresh_lt A R a (h : natmap A) t v (z0 : R) : + eval a (lt t) (fresh h \-> v \+ h) z0 = + if t <= fresh h then eval a (lt t) h z0 else + if valid h then a (eval a predT h z0) (fresh h) v else z0. +Proof. +case V: (valid h); last first. +- by move/invalidE: (negbT V)=>->; rewrite join_undef !eval_undef; case: ifP. +case: ifPn=>H. +- by rewrite eval_umfilt umfiltPtUn valid_fresh V ltnNge H -eval_umfilt. +rewrite joinC evalUnPt; last first. +- by apply/allP=>x; apply: lastkey_max. +- by rewrite joinC valid_fresh. +rewrite ltnNge H; congr a; apply: eq_in_eval. +case=>k w /In_dom/lastkey_max /=; rewrite /fresh -ltnNge in H. +by case: ltngtP H=>// /ltnW H _ /leq_ltn_trans; apply. +Qed. + +Lemma eval_le_fresh A R a (h : natmap A) t v (z0 : R) : + t <= last_key h -> + eval a (le t) (fresh h \-> v \+ h) z0 = eval a (le t) h z0. +Proof. by move=>H; rewrite eval_fresh_le H. Qed. + +Lemma eval_lt_fresh A R a (h : natmap A) t v (z0 : R) : + t <= fresh h -> + eval a (lt t) (fresh h \-> v \+ h) z0 = eval a (lt t) h z0. +Proof. by move=>H; rewrite eval_fresh_lt H. Qed. + +Lemma eval_le_ind A R a (P : R -> Prop) (h : natmap A) t1 t2 (z0 : R) : + t1 <= t2 -> + P (eval a (le t1) h z0) -> + (forall k v z0, (k, v) \In h -> t1 < k <= t2 -> P z0 -> P (a z0 k v)) -> + P (eval a (le t2) h z0). +Proof. +by move=>T P0 H; rewrite (eval_le_split a h z0 T); apply: eval_ind. +Qed. + +Lemma eval_lt_ind A R a (P : R -> Prop) (h : natmap A) t1 t2 (z0 : R) : + t1 < t2 -> + P (eval a (le t1) h z0) -> + (forall k v z0, (k, v) \In h -> t1 < k < t2 -> P z0 -> P (a z0 k v)) -> + P (eval a (lt t2) h z0). +Proof. +by move=>T P0 H; rewrite (eval_lt_split a h z0 T); apply: eval_ind. +Qed. + +(* common case: functional version of the above le_lemma *) +Lemma eval_leF A R X a (f : R -> X) (h : natmap A) t1 t2 (z0 : R) : + t1 <= t2 -> + (forall k v z0, (k, v) \In h -> t1 < k <= t2 -> f (a z0 k v) = f z0) -> + f (eval a (le t1) h z0) = f (eval a (le t2) h z0). +Proof. +move=>T H. +apply: (eval_le_ind (P := fun x => f (eval a (le t1) h z0) = f x)) T _ _=>//. +by move=>k v z1 H1 /H ->. +Qed. + +(* common case: functional version of the above lt_lemma *) +Lemma eval_ltF A R X a (f : R -> X) (h : natmap A) t1 t2 (z0 : R) : + t1 < t2 -> + (forall k v z0, (k, v) \In h -> t1 < k < t2 -> f (a z0 k v) = f z0) -> + f (eval a (le t1) h z0) = f (eval a (lt t2) h z0). +Proof. +move=>T H. +apply: (eval_lt_ind (P := fun x => f (eval a (le t1) h z0) = f x)) T _ _=>//. +by move=>k v z1 H1 /H ->. +Qed. + +Lemma umcnt_le_split A p (h : natmap A) t1 t2 : + t1 <= t2 -> + um_count (predI p (le t2)) h = + um_count (predI p (le t1)) h + + um_count (predI p [pts k v | t1 < k <= t2]) h. +Proof. +move=>T; rewrite -!umcnt_umfilt !(umcnt_umfiltC p). +by rewrite (umcnt_predD _ (pts_sub T)) (eq_in_umcnt2 _ (ptsD T)). +Qed. + +Lemma umcnt_le0 A p (h : natmap A) t : + valid h -> {in dom h, forall k, t < k} -> + um_count (predI p (le t)) h = 0. +Proof. by rewrite -umcnt_umfilt=>V /(umfilt_le0 V) ->; rewrite umcnt0. Qed. + +Lemma umcnt_leUn A p (h1 h2 : natmap A) t : + valid (h1 \+ h2) -> t <= last_key h1 -> + {in dom h2, forall k, k > last_key h1} -> + um_count (predI p (le t)) (h1 \+ h2) = + um_count (predI p (le t)) h1. +Proof. by move=>V K D; rewrite -!umcnt_umfilt umfilt_leUn. Qed. + +Lemma umcnt_le_gapless A p (h1 h2 : natmap A) t : + gapless h1 -> valid h2 -> [pcm h1 <= h2] -> t <= last_key h1 -> + um_count (predI p (le t)) h2 = um_count (predI p (le t)) h1. +Proof. by move=>G V K D; rewrite -!umcnt_umfilt (umfilt_le_gapless G). Qed. + +Lemma umcnt_le_last A p (h : natmap A) t : + last_key h <= t -> um_count (predI p (le t)) h = um_count p h. +Proof. by move=>K; rewrite -!umcnt_umfilt umfilt_le_last. Qed. + +Lemma umcnt_fresh_le A p (h : natmap A) t v : + um_count (predI p (le t)) (fresh h \-> v \+ h) = + if t <= last_key h then um_count (predI p (le t)) h else + if p (fresh h, v) && valid h then 1 + um_count p h else um_count p h. +Proof. +case V: (valid h); last first. +- move/invalidE: (negbT V)=>->; rewrite join_undef !umcnt_undef. + by rewrite lastkey_undef andbF; case: ifP. +case: ifP=>H. +- by rewrite -!umcnt_umfilt umfiltPtUn valid_fresh V ltnNge H. +rewrite umcntPtUn ?valid_fresh //= ltnNge H /=. +by rewrite umcnt_le_last; [case: ifP | case: ltngtP H]. +Qed. + +Lemma umcnt_le_fresh A p (h : natmap A) t v : + t <= last_key h -> + um_count (predI p (le t)) (fresh h \-> v \+ h) = + um_count (predI p (le t)) h. +Proof. by move=>H; rewrite umcnt_fresh_le H. Qed. + +Definition fresh_le := (umcnt_fresh_le, eval_fresh_le). +Definition le_last := (umcnt_le_last, eval_le_last). +Definition le_fresh := (umcnt_le_fresh, eval_le_fresh). +Definition lt_fresh := (eval_lt_fresh). + +Lemma umcnt_mono A p (h : natmap A) t1 t2 : + t1 <= t2 -> + um_count (predI p (le t1)) h <= um_count (predI p (le t2)) h. +Proof. by move=>T; rewrite (umcnt_le_split _ _ T); apply: leq_addr. Qed. + +Lemma umcnt_leE A p (h : natmap A) t1 t2 : + um_count (predI p (le t1)) h = um_count (predI p (le t2)) h -> + um_count (predI p [pts k v | t1 < k <= t2]) h = 0. +Proof. +case T: (t1 <= t2); last first. +- rewrite -{4}(umcnt_pred0 h) =>_; apply: eq_in_umcnt2=>[[k v]] /=. + by apply: contraFF T=>/and3P [_ /ltnW]; apply: leq_trans. +move/eqP; rewrite (umcnt_le_split _ _ T). +by rewrite -{1}[um_count (predI p (le t1)) h]addn0 eqn_add2l=>/eqP. +Qed. + +Lemma umcnt_umfilt_eq A p (h : natmap A) t1 t2 : + um_count (predI p (le t1)) h = um_count (predI p (le t2)) h -> + um_filter (predI p (le t1)) h = um_filter (predI p (le t2)) h. +Proof. +suff {t1 t2} L t1 t2 : t1 <= t2 -> + um_count (predI p (le t1)) h = um_count (predI p (le t2)) h -> + um_filter (predI p (le t1)) h = um_filter (predI p (le t2)) h. +- by move=>E; case: (leqP t1 t2)=>[|/ltnW] N; [|apply/esym]; apply: L N _. +case V : (valid h); last first. +- by move/negbT/invalidE: V=>->; rewrite !umfilt_undef. +move=>N /umcnt_leE /(umcnt_umfilt0 V). +rewrite !umfilt_predI !(umfiltC p) (umfilt_le_split _ N)=>->. +by rewrite unitR. +Qed. + +(* count and eval leE lemmas put together *) +Lemma evcnt_le_ind A R a (P : R -> Prop) p (h : natmap A) t1 t2 (z0 : R) : + t1 <= t2 -> + P (eval a (le t1) h z0) -> + (forall k v z0, (k, v) \In h -> t1 < k <= t2 -> + ~~ p (k, v) -> P z0 -> P (a z0 k v)) -> + um_count (predI p (le t1)) h = um_count (predI p (le t2)) h -> + P (eval a (le t2) h z0). +Proof. +move=>T P0 H /umcnt_leE/umcnt_mem0=>K; apply: eval_le_ind (T) (P0) _. +move=>k v z1 H' K'; apply: (H k v z1 H' K'). +by move: (K k v H'); rewrite negb_and; case/orP=>//; rewrite K'. +Qed. + +(* common case is when P is an equality of functions *) +Lemma evcnt_leF A R X a (f : R -> X) p (h : natmap A) t1 t2 (z0 : R) : + t1 <= t2 -> + (forall k v z0, (k, v) \In h -> t1 < k <= t2 -> + ~~ p (k, v) -> f (a z0 k v) = f z0) -> + um_count (predI p (le t1)) h = um_count (predI p (le t2)) h -> + f (eval a (le t1) h z0) = f (eval a (le t2) h z0). +Proof. +move=>T H. +apply: (evcnt_le_ind (P := fun x => f (eval a (le t1) h z0) = f x)) T _ _=>//. +by move=>k v z1 H' K' /H ->. +Qed. + +(* The following is a lemma that generalizes evcnt_leF *) +(* from counting to functions that are monotone, in a suitable sense *) +(* that will be defined below *) + +(* The lemma says that if evaluating a growing function preserves *) +(* the result, then all the evaluating steps preserve the result *) + +(*********************************) +(* Some notational abbreviations *) +(*********************************) + +(* exec is evaluating a history up to a timestamp *) +(* run is evaluating a history up to the end *) + +(* In exec and run, the timestamp shouldn't influence *) +(* the val of the operation. So we need a coercion to *) +(* account for the timestamp, which is then ignored *) +Notation exec a t h z0 := (evalv a (le t) h z0). +Notation run a h z0 := (evalv a xpredT h z0). + +Section Exec. +Variables (V R : Type) (aop : V -> R -> R). + +Lemma exec0 a (h : natmap V) (z0 : R) : exec a 0 h z0 = z0. +Proof. +have /(eq_in_eval _) -> : forall kv, kv \In h -> le 0 kv = pred0 kv. +- by case=>k v /In_cond; case: k. +by rewrite eval_pred0. +Qed. + +End Exec. + +Section Growing. +Variables (V R : Type) (X : ordType) (a : R -> V -> R) (f : R -> X). +Implicit Types p : pred (nat*V). + +Definition growing (h : natmap V) := + forall k v z0, (k, v) \In h -> oleq (f z0) (f (a z0 v)). + +Lemma growL h1 h2 : valid (h1 \+ h2) -> growing (h1 \+ h2) -> growing h1. +Proof. by move=>W G k v z0 H; apply: (G k); apply/InL. Qed. + +Lemma growR h1 h2 : valid (h1 \+ h2) -> growing (h1 \+ h2) -> growing h2. +Proof. by rewrite joinC; apply: growL. Qed. + +Lemma helper0 p h z0 : growing h -> oleq (f z0) (f (evalv a p h z0)). +Proof. +elim/um_indf: h z0=>[||k v h IH W /(order_path_min (@trans _)) T] z0 G; +rewrite ?eval_undef ?eval0 // evalPtUn //. +have K: (k, v) \In pts k v \+ h by apply/InPtUnE=>//; left. +have Gh: growing h. +- by move=>k1 v1 z1 X1; apply: (G k1); apply/InPtUnE=>//; right. +case: ifP=>// P; last by apply: IH. +by apply: otrans (IH _ Gh); apply: (G k). +Qed. + +Lemma helper1 p h z0 k v : + growing (k \-> v \+ h) -> + valid (k \-> v \+ h) -> + all (ord k) (dom h) -> + p (k, v) -> + f (evalv a p (k \-> v \+ h) z0) = f z0 -> + f (a z0 v) = f z0. +Proof. +move=>G W D P; move: (growR W G)=>Gh. +have K: (k, v) \In k \-> v \+ h by apply/InPtUnE=>//; left. +rewrite evalPtUn // P => E; apply/eqP; case: ordP=>//. +- by move/(G k v z0): K; rewrite /oleq eq_sym; case: ordP. +by move: (helper0 p (a z0 v) Gh); rewrite -E /oleq eq_sym; case: ordP. +Qed. + +Lemma helper2 p h1 h2 z0 k v : + growing (h1 \+ (k \-> v \+ h2)) -> + valid (h1 \+ (k \-> v \+ h2)) -> + {in dom h1, forall x, x < k} -> + all (ord k) (dom h2) -> + p (k, v) -> + f (evalv a p (h1 \+ (k \-> v \+ h2)) z0) = f z0 -> + f (a (evalv a p h1 z0) v) = f (evalv a p h1 z0). +Proof. +move=>G W D1 D2 P E1; rewrite evalUn ?(validX W) // in E1; last first. +- move=>x y /D1 X1; rewrite domPtUn inE ?(validX W). + by case/orP=>[/eqP <-|/(allP D2)] //; apply: ltn_trans. +suff E2 : f (evalv a p h1 z0) = f z0. +- by apply: helper1 (growR W G) (validR W) D2 P _; rewrite E1 E2. +apply/eqP; case: ordP=>//. +- by move: (helper0 p z0 (growL W G)); rewrite /oleq eq_sym; case: ordP. +move: (helper0 p (evalv a p h1 z0) (growR W G)). +by rewrite -E1 /oleq eq_sym; case: ordP. +Qed. + +(* "introducing" growth *) +Lemma growI h t1 t2 z0 : + growing h -> t1 <= t2 -> + oleq (f (exec a t1 h z0)) (f (exec a t2 h z0)). +Proof. +move=>G N; case W: (valid h); last first. +- by move/negbT/invalidE: W=>->; rewrite !eval_undef. +rewrite eval_umfilt [in X in oleq _ X]eval_umfilt (umfilt_le_split h N). +rewrite evalUn; first by apply: helper0=>x y z /In_umfilt [_ /G]. +- by rewrite -(umfilt_le_split h N) valid_umfilt. +by move=>??/dom_umfilt[?][/leq_ltn_trans Y _]/dom_umfilt[?][/andP[/Y]]. +Qed. + +(* "eliminating" growth *) +Lemma growE h t1 t2 z0 k v : + growing h -> (k, v) \In h -> t1 < k <= t2 -> + f (exec a t2 h z0) = f (exec a t1 h z0) -> + f (a (exec a k.-1 h z0) v) = f (exec a k.-1 h z0). +Proof. +move=>G H /andP [N1 N2] E; have W : valid h by case: H. +pose h0 : natmap V := um_filter (le t1) h. +pose h1 : natmap V := um_filter [pts x _ | t1 < x <= k.-1] h. +pose h2 : natmap V := um_filter [pts x _ | k < x <= t2] h. +have subX : forall k m n, k < n -> (m < n) = (m <= n.-1) by move=>?? []. +have K1 : k.-1 <= k by rewrite ltnW // (subX t1). +have K2 : t1 <= k.-1 by rewrite -(subX t1). +have K3 : k.-1 <= k <= t2 by rewrite K1 N2. +have K4 : t1 <= k.-1 <= t2 by rewrite K2 (leq_trans K1 N2). +have Eh: um_filter (le t2) h = h0 \+ (h1 \+ (k \-> v \+ h2)). +- rewrite (umfilt_le_split h N2) (umfilt_le_split h K1). + by rewrite (umfilt_le_split h K2) (umfilt_pt_split H) -!joinA. +have W1 : valid (h0 \+ (h1 \+ (k \-> v \+ h2))) by rewrite -Eh valid_umfilt. +rewrite eval_umfilt (umfilt_le_split h K2) evalUn ?(validX W1) //; last first. +- by move=>??/dom_umfilt[?][/leq_ltn_trans Y] _ /dom_umfilt[?][] /andP [/Y]. +rewrite -(eval_umfilt (le t1)); apply: helper2 (validR W1) _ _ _ _ =>//. +- by apply: growR W1 _; rewrite -Eh=>k1 v1 z1 /In_umfilt [] _ /G. +- by move=>x /dom_umfilt; rewrite (subX t1 x) //; case=>v0 [] /andP []. +- by apply/allP=>x /dom_umfilt; case=>v0 [] /andP []. +rewrite eval_umfilt Eh evalUn ?(validX W1) -?eval_umfilt // in E. +move=>x y /dom_umfilt; case=>v1 [/leq_ltn_trans Y _]. +rewrite -(umfilt_pt_split H) -(umfilt_lt_split h K3). +by rewrite -(umfilt_lt_split h K4) =>/dom_umfilt; case=>v0 [/andP][/Y]. +Qed. + +End Growing. + +(* The common case of growI and growE is when X = nat. *) + +Section GrowingNat. +Variables (V R : Type) (X : ordType) (a : R -> V -> R) (f : R -> nat). +Implicit Types p : pred (nat*V). + +Lemma growIn h t1 t2 z0 : + growing a f h -> t1 <= t2 -> + f (exec a t1 h z0) <= f (exec a t2 h z0). +Proof. +by move=>G N; move: (growI z0 G N); rewrite leq_eqVlt /oleq/ord orbC. +Qed. + +Lemma growEn h t1 t2 z0 k v : + growing a f h -> (k, v) \In h -> t1 < k <= t2 -> + f (exec a t2 h z0) = f (exec a t1 h z0) -> + f (a (exec a k.-1 h z0) v) = f (exec a k.-1 h z0). +Proof. by apply: growE. Qed. + +End GrowingNat. + + +(**********************************) +(* Morphisms on locking histories *) +(**********************************) + +(* Morphism on locking histories that provides mutual exclusion: *) +(* when one thread has locked, the other can't proceed. *) +(* Because the morphism just looks into the last history entry *) +(* we call it *omega*, or omg for short. *) + +Section OmegaMorph. +Let U := nat_mapPCM (bool * bool). + +Definition omg_s := fun x y : U => + [&& last_val false x ==> (last_key y < last_key x) & + last_val false y ==> (last_key x < last_key y)]. + +Lemma omg_sep_ax : seprel_axiom omg_s. +Proof. +rewrite /omg_s; split=>[|x y|x y|x y z] /=; first by rewrite lastval0. +- by rewrite andbC. +- move=>V /andP [H _]; rewrite lastkey0 lastval0. + by case: (x in x ==> _) H=>// /(leq_trans _) ->. +move=>V /andP [X Y] /andP []. +rewrite !lastkeyUn !lastvalUn !(validLE3 V). +rewrite {1 2}maxnC {1 2}/maxn gtn_max leq_max /=. +case: (ltngtP (last_key x) (last_key y)) X Y=>H X Y Kx Kz; + rewrite ?H ?X ?(negbTE Y) fun_if if_arg ?implybT ?Kx Kz if_same //= ?andbT. +by case: (x in x ==> _) Kz=>// /(ltn_trans H) ->. +Qed. + +Canonical omg_seprel := Eval hnf in seprel omg_s omg_sep_ax. + +Definition omg (x : U) : mtx2 := if last_val false x then own else nown. + +Lemma omg_morph_ax : morph_axiom omg_seprel omg. +Proof. +rewrite /omg; split=>[|x y V /andP [X Y]]; first by rewrite lastval0. +rewrite lastvalUn V; case: ltngtP X Y=>H X Y; +by rewrite ?(negbTE X) ?(negbTE Y) //; case: ifP. +Qed. + +Canonical omg_morph := Morphism' omg omg_morph_ax. + +(* transfer lemmas *) + +Lemma omg_eta (h : natmap (bool * bool)): + valid h -> omg h = own -> + exists h' v, [/\ h' = free h (last_key h), + h = last_key h \-> (v, true) \+ h', + last_key h != 0, + last_key h \in dom h, + last_key h \notin dom h' & + last_key h' < last_key h]. +Proof. +rewrite /omg; case: ifP=>// N V _; set k := last_key h. +have D : k \in dom h. +- rewrite /last_val /atval /oapp in N. + case: dom_find N=>[->//|]. move=>v. simpl. + by []. +have K : k != 0 by apply: dom_cond D. +case: (um_eta D); case=>v1 v2 [Ef Eh]. +set h' := free h k in Eh *; set q := k \-> (v1, true). +have Nd : k \notin dom h'. +- rewrite Eh in V; case: validUn V=>// _ _ X _; apply: X. + by rewrite domPt inE /= K eq_refl. +exists h', v1; split=>//. +- by rewrite /last_val /atval Ef /= in N; rewrite -N. +have: last_key h' <= k. +- by rewrite /k Eh; apply: lastkeyUnf; rewrite -Eh. +rewrite leq_eqVlt; case/orP=>// /eqP E. +rewrite -E in Nd; apply: contraR Nd=>/= _. +by apply: (dom_lastkeyE (a:=0)); rewrite E; case: (k) K. +Qed. + +(* specialize to alternating histories *) +Lemma omg_eta_all (h : natmap (bool * bool)) : + valid h -> omg h = own -> um_all (fun k v => v.2 = ~~ v.1) h -> + exists h', [/\ h' = free h (last_key h), + h = last_key h \-> (false, true) \+ h', + last_key h != 0, + last_key h \in dom h, + last_key h \notin dom h' & + last_key h' < last_key h]. +Proof. +move=>V H A; case: (omg_eta V H)=>h' [v][H1 H2 H3 H4 H5 H6]. +exists h'; split=>//; rewrite H2 in V A; case: (umallPtUnE V A)=>/=. +by case: v {A} V H2. +Qed. + +Lemma omg_lastkey x y : + (omg x = own -> valid (x \+ y) -> omg_s x y -> + last_key (x \+ y) = last_key x) * + (omg y = own -> valid (x \+ y) -> omg_s x y -> + last_key (x \+ y) = last_key y). +Proof. +rewrite /omg_s /omg; split=>L V S; case: ifP L=>L // _; +rewrite L /= in S; rewrite lastkeyUn V; case/andP: S=>S1 S2. + by rewrite maxnC /maxn S1. +by rewrite /maxn S2. +Qed. + +End OmegaMorph. + + +(*******************************) +(* Monotonicity for natmap nat *) +(*******************************) + +Lemma nm_monoPtUn t t' (h : natmap nat) : + last_key h < t -> + (forall k v, (k, v) \In h -> v < t') -> + um_mono (t \-> t' \+ h) = um_mono h. +Proof. +move=>H1 H2; case W: (valid h); last first. +- by move/invalidE: (negbT W)=>->; rewrite join_undef. +by apply: ummonoPtUn (lastkeyPtUn_valid _ W H1) _=>k v X; split; +[move/In_dom/lastkey_max/leq_ltn_trans: X; apply|apply: H2 X]. +Qed. + + +(************************************) +(* Specializing to ordering on nats *) +(* (the distinction is that we need *) +(* a special number, 0, to be added *) +(* to the list) *) +(************************************) + +(* rename index to prevent accidental simplification *) +Definition nat_index := nosimpl (@index nat_eqType). + +Definition seq_le (ks : seq nat) (t1 t2 : nat) := + index t1 (0::ks) <= index t2 (0::ks). + +Definition seq_lt (ks : seq nat) (t1 t2 : nat) := + index t1 (0::ks) < index t2 (0::ks). + +Notation "t1 '<=[' ks ] t2" := (seq_le ks t1 t2) + (at level 10, format "t1 '<=[' ks ] t2"). + +Notation "t1 '<[' ks ] t2" := (seq_lt ks t1 t2) + (at level 10, format "t1 '<[' ks ] t2"). + +Lemma sle_refl ks t : t <=[ks] t. +Proof. by rewrite /seq_le. Qed. + +Hint Resolve sle_refl : core. + +Definition olast_key {A} ks (h : natmap A) := + last 0 (filter (fun x => x \in dom h) ks). +Prenex Implicits olast_key. + +Section OLastKey. +Variable A : Type. +Implicit Type h : natmap A. + +Lemma olastkey_undef ks : olast_key ks (undef : natmap A) = 0. +Proof. by elim: ks=>[|k ks] //=; rewrite /olast_key dom_undef. Qed. + +Lemma olastkey0 ks : olast_key ks (Unit : natmap A) = 0. +Proof. by elim: ks=>[|k ks] //=; rewrite /olast_key dom0. Qed. + +Lemma mem_olastkey ks h : olast_key ks h \in 0 :: ks. +Proof. +move: (mem_last 0 [seq x <- ks | x \in dom h]). +by rewrite !inE mem_filter; case: eqP=>//= _ /andP []. +Qed. + +Hint Resolve mem_olastkey : core. + +(* some interdependence between olastkey components *) + +Lemma olastkey0P ks h : + reflect (forall k, k \in dom h -> k \in ks -> false) + (olast_key ks h == 0). +Proof. +rewrite /olast_key; case: eqP=>H; constructor; last first. +- by move=>K; case/eqP/last_filter_change/andP: H=>H /(K _ H). +move=>k H1 H2; move: (@last_in _ k 0 [seq x <- ks | x \in dom h]). +by rewrite !mem_filter H1 H=>/(_ H2) /andP [] /dom_cond. +Qed. + +Lemma olastkey0_dom ks h : + (olast_key ks h == 0) = (olast_key ks h \notin dom h). +Proof. +apply/idP/idP=>[/eqP ->|H]; first by case E: (0 \in _)=>//; move/dom_cond: E. +by apply: contraR H=>/last_filter_change /andP []. +Qed. + +Lemma olastkeyN_dom ks h : + (olast_key ks h != 0) = (olast_key ks h \in dom h). +Proof. by rewrite olastkey0_dom negbK. Qed. + +Lemma olastkey0_mem ks h : + 0 \notin ks -> (olast_key ks h == 0) = (olast_key ks h \notin ks). +Proof. +move=>N; apply/idP/idP=>[/eqP ->|H] //. +by apply: contraR H=>/last_filter_change /andP []. +Qed. + +Lemma olastkeyN_mem ks h : + 0 \notin ks -> (olast_key ks h != 0) = (olast_key ks h \in ks). +Proof. by move/olastkey0_mem=>->; rewrite negbK. Qed. + +Lemma olastkey0_index ks h : + (olast_key ks h == 0) = (index (olast_key ks h) (0 :: ks) == 0). +Proof. by rewrite /= (eq_sym 0); case: eqP. Qed. + +Lemma olastkey_in k ks h : + k \in dom h -> k \in ks -> + (olast_key ks h \in dom h) && (olast_key ks h \in ks). +Proof. +move=>D K; move: (@last_in _ k 0 [seq x <- ks | x \in dom h]). +by rewrite !mem_filter D K; apply. +Qed. + +Lemma olastkey_indom k ks h : + k \in dom h -> k \in ks -> + olast_key ks h \in dom h. +Proof. by move=>H /(olastkey_in H) /andP []. Qed. + +Lemma olastkey_inmem k ks h : + k \in dom h -> k \in ks -> + olast_key ks h \in ks. +Proof. by move=>H /(olastkey_in H) /andP []. Qed. + +Lemma olastkey_max ks h x : + uniq (0 :: ks) -> x \in dom h -> x \in ks -> x <=[ks] (olast_key ks h). +Proof. by apply: last_filter_last. Qed. + +Lemma max_olastkey ks h x : + uniq (0 :: ks) -> + x \in dom h -> x \in ks -> + (forall y, y \in dom h -> y \in ks -> y <=[ks] x) -> + olast_key ks h = x. +Proof. +move=>/= /andP [N U] Dx Kx H; apply: max_index_last. +- by rewrite filter_uniq. +- by rewrite mem_filter Dx Kx. +move=>z; rewrite mem_filter=>/andP [Dz Kz]. +move: {H}(H z Dz Kz); rewrite /seq_le /nat_index /= !(eq_sym 0). +move: (dom_cond Dx) (dom_cond Dz)=>/= /negbTE -> /negbTE ->. +by apply: index_filter_mono. +Qed. + +Lemma olastkeyPt ks (x : nat) (v : A) : + uniq ks -> x != 0 -> x \in ks -> olast_key ks (x \-> v) = x. +Proof. +rewrite /olast_key domPtK /= =>U -> Kx. +have : [seq x0 <- ks | x0 \in [:: x]] = [seq x0 <- ks | (pred1 x) x0]. +- by apply: eq_in_filter=>/= y; rewrite inE. +by move=>->; rewrite (filter_pred1_uniq U Kx). +Qed. + +Lemma olastkey_mono B ks (h1 : natmap A) (h2 : natmap B) : + {subset dom h1 <= dom h2} -> uniq (0 :: ks) -> + (olast_key ks h1) <=[ks] (olast_key ks h2). +Proof. by apply: index_last_sub. Qed. + +Lemma olastkeyfUn ks h1 h2 : + valid (h1 \+ h2) -> uniq (0 :: ks) -> + (olast_key ks h1) <=[ks] (olast_key ks (h1 \+ h2)). +Proof. by move=>X U; apply: olastkey_mono=>// x; rewrite domUn inE X => ->. Qed. + +Lemma olastkeyUnf ks h1 h2 : + valid (h1 \+ h2) -> uniq (0 :: ks) -> + (olast_key ks h2) <=[ks] (olast_key ks (h1 \+ h2)). +Proof. by rewrite joinC; apply: olastkeyfUn. Qed. + +(* a convenient rephrasing of above lemmas *) + +Lemma olastkeyUn0 ks h1 h2 : + valid (h1 \+ h2) -> uniq (0 :: ks) -> + olast_key ks h1 = olast_key ks h2 -> + ([seq x <- ks | x \in dom h1] = [::]) * + ([seq x <- ks | x \in dom h2] = [::]). +Proof. +move=>V U E. +case E1: ([seq x <- ks | x \in dom h1] == [::]). +- rewrite /olast_key (eqP E1) /= in E *. + by move/esym/eqP/olastkey0P/filter_nilp/eqP: E=>->. +case E2 : ([seq x <- ks | x \in dom h2] == [::]). +- rewrite /olast_key (eqP E2) /= in E *. + by move/eqP/olastkey0P/filter_nilp/eqP: E=>->. +case/filter_nilp: E1=>x Dx /(olastkey_indom Dx). +rewrite E; case: validUn (V)=>// _ _ X _ /X. +by rewrite -olastkey0_dom=>/olastkey0P/filter_nilp; rewrite E2. +Qed. + +Definition omaxn ks m n := if m <[ks] n then n else m. + +Lemma olastkeyUn ks h1 h2 : + uniq (0::ks) -> + olast_key ks (h1 \+ h2) = + if valid (h1 \+ h2) then + omaxn ks (olast_key ks h1) (olast_key ks h2) + else 0. +Proof. +move=>U; have N : 0 \notin ks by case/andP: U. +have H (k1 k2 : natmap A) : valid (k1 \+ k2) -> + (olast_key ks k1) <[ks] (olast_key ks k2) -> + olast_key ks (k1 \+ k2) = olast_key ks k2. +- move=>V H; apply: max_olastkey=>//. + - rewrite domUn inE V -(olastkeyN_dom ks k2). + by rewrite olastkey0_index -lt0n (gt0 H) orbT. + - by rewrite -(olastkeyN_mem k2) // olastkey0_index -lt0n (gt0 H). + move=>y; rewrite domUn inE V; case/orP=>D /(olastkey_max U D) // T. + by apply: leq_trans T (ltnW H). +case V : (valid _); last first. +- by move/invalidE: (negbT V)=>->; rewrite olastkey_undef. +rewrite /omaxn /seq_lt; case: ltngtP. +- by apply: H. +- by rewrite joinC in V *; apply: H V. +move/(index_inj (mem_olastkey _ _)). +case/(olastkeyUn0 V U)=>/eqP E1 /eqP E2. +rewrite /olast_key (eqP E1). +rewrite (_ : [seq x <- ks | _] = [::]) //. +apply/eqP/filter_nilp=>x; rewrite domUn inE V. +by case/orP; apply: filter_nilp. +Qed. + +(* monotonicity just gives us <= *) +(* when we remove the olast key, we get a strict < *) +Lemma olastkeyF ks h : + uniq (0 :: ks) -> + olast_key ks h \in dom h -> + (olast_key ks (free h (olast_key ks h))) <[ks] (olast_key ks h). +Proof. +move=>U D. +have N : olast_key ks h \notin dom (free h (olast_key ks h)). +- by rewrite domF inE eq_refl. +have : index (olast_key ks (free h (olast_key ks h))) (0 :: ks) <= + index (olast_key ks h) (0 :: ks). +- by apply: olastkey_mono=>// x; rewrite domF inE; case: ifP. +rewrite leq_eqVlt; case/orP=>// /eqP /(index_inj (mem_olastkey _ _)) /eqP E. +rewrite -{1}(eqP E) -olastkey0_dom in N. +by rewrite (eqP N) eq_sym olastkey0_dom D in E. +Qed. + +Lemma olastkeyPtUnE ks u2 u1 h t : + olast_key ks (t \-> u1 \+ h) = olast_key ks (t \-> u2 \+ h). +Proof. +have V2 : valid (t \-> u2 \+ h) = valid (t \-> u1 \+ h) by rewrite !validPtUn. +case V1: (valid (t \-> u1 \+ h)); rewrite V1 in V2; last first. +- by move/invalidE: (negbT V1) (negbT V2) => -> /invalidE ->. +by congr last; apply/eq_in_filter=>x K; rewrite !domPtUn !inE V1 V2. +Qed. + +Lemma olastkeyUnPtE ks u2 u1 h t : + olast_key ks (h \+ t \-> u1) = olast_key ks (h \+ t \-> u2). +Proof. by rewrite !(joinC h); apply: olastkeyPtUnE. Qed. + +Lemma olastkeyU ks u h t : + t \in dom h -> olast_key ks (upd t u h) = olast_key ks h. +Proof. by case/um_eta=>v [_ E]; rewrite E updPtUn (olastkeyPtUnE ks v). Qed. + +(* pcm induced ordering *) + +Lemma umpleq_olastkey ks h1 h2 : + uniq (0 :: ks) -> + valid h2 -> [pcm h1 <= h2] -> + (olast_key ks h1) <=[ks] (olast_key ks h2). +Proof. +move=>U V H; case: H V=>z -> V. +by apply: olastkey_mono=>// k D; rewrite domUn inE V D. +Qed. + +End OLastKey. + +(* lemmas with omap *) + +Lemma olastkey_omap (V1 V2 : ordType) (f : nat * V1 -> option V2) (h : natmap V1) ks : + uniq (0 :: ks) -> + (olast_key ks (omap f h)) <=[ks] (olast_key ks h). +Proof. by move=>U; apply: olastkey_mono=>//; apply: dom_omap_sub. Qed. + + +(*************************) +(* Surgery on histories *) +(* using oleq filtering *) +(*************************) + +Notation ole ks t := [pts k _ | k <=[ks] t]. +Notation olt ks t := [pts k _ | k <[ks] t]. + +Lemma pts_osub V ks t1 t2 : + t1 <=[ks] t2 -> subpred (T:=nat*V) (ole ks t1) (ole ks t2). +Proof. by move=>T [k v] /leq_trans; apply. Qed. + +Lemma optsD V ks t1 t2 : + t1 <=[ks] t2 -> + predD (ole ks t2) (ole ks t1) =1 + [pts k (v : V) | t1 <[ks] k && k <=[ks] t2]. +Proof. by move=>T [k v] /=; rewrite -ltnNge. Qed. + +Lemma olastkey_umfilt_leq A ks (h : natmap A) t : + uniq (0 :: ks) -> + (olast_key ks (um_filter (ole ks t) h)) <=[ks] t. +Proof. +move=>U; case: (olast_key ks (um_filter (ole ks t) h) =P 0)=>[->|/eqP] //. +by rewrite {1}olastkeyN_dom=>/dom_umfilt [v][]. +Qed. + +Lemma olastkey_umfilt_dom A ks (h : natmap A) t : + uniq (0::ks) -> + (olast_key ks (um_filter (ole ks t) h)) <=[ks] (olast_key ks h). +Proof. by apply: olastkey_mono; apply: dom_umfilt_subset. Qed. + +(* In the next lemma, we need a new condition *) +(* that the ordering ks is total on dom h. *) +(* In practice, we will always have this, as the ordering *) +(* will in fact be a permutation of dom h. *) + +Lemma umfilt_ole_olast A ks (h : natmap A) t : + uniq (0::ks) -> + {subset dom h <= ks} -> + (olast_key ks h) <=[ks] t -> + um_filter (ole ks t) h = h. +Proof. +move=>U S; case V : (valid h); last first. +- by move/invalidE: (negbT V)=>->; rewrite umfilt_undef. +move=>N; rewrite -{2}[h]umfilt_predT; apply/eq_in_umfilt. +by case=>k v /In_dom H; apply: leq_trans N; apply: olastkey_max (S _ H). +Qed. + +Lemma umfilt_olast A ks (h : natmap A) : + uniq (0::ks) -> {subset dom h <= ks} -> + um_filter (ole ks (olast_key ks h)) h = h. +Proof. by move=>U S; apply: umfilt_ole_olast. Qed. + +Lemma umfilt_ole_fresh A ks (h : natmap A) fresh : + uniq (0::ks) -> + {subset dom h <= ks} -> + fresh \notin (0::ks) -> + um_filter (ole ks fresh) h = h. +Proof. +move=>U S F; apply: umfilt_ole_olast=>//. +move: (index_size fresh (0::ks)). +rewrite /seq_le /nat_index leq_eqVlt index_mem (negbTE F) orbF. +by move/eqP=>->; apply: index_size. +Qed. + +Lemma umfilt_ole0 A ks (h : natmap A) t : + valid h -> {in dom h, forall k, t <[ks] k} -> + um_filter (ole ks t) h = Unit. +Proof. +move=>V D; rewrite -(umfilt_pred0 V). +apply/eq_in_umfilt; case=>k v [_][z E]; subst h. +rewrite /seq_le /nat_index leqNgt; apply: contraTF (D k _)=>//=. +by rewrite domPtUn inE V eq_refl. +Qed. + +Lemma umfilt_ole_split A ks (h : natmap A) t1 t2 : + t1 <=[ks] t2 -> + um_filter (ole ks t2) h = + um_filter (ole ks t1) h \+ + um_filter [pts x _ | t1 <[ks] x && x <=[ks] t2] h. +Proof. +move=>T; rewrite -umfilt_dpredU; last first. +- by case=>x y N /=; rewrite negb_and -leqNgt -/(seq_le ks _ _) N. +apply/eq_in_umfilt; case=>k v _ /=; rewrite /seq_le /seq_lt. +case: (leqP (index k (0::ks)) (index t1 (0::ks)))=>//=. +by move/leq_trans; apply. +Qed. + +Lemma umfilt_olt_split A ks (h : natmap A) t1 t2 k : + t1 <=[ks] k && k <=[ks] t2 -> + um_filter [pts x _ | t1 <[ks] x && x <=[ks] t2] h = + um_filter [pts x _ | t1 <[ks] x && x <=[ks] k] h \+ + um_filter [pts x _ | k <[ks] x && x <=[ks] t2] h. +Proof. +move=>/andP [T1 T2]; rewrite -umfilt_dpredU; last first. +- by case=>x y /andP [N1 N2]; rewrite /= negb_and -leqNgt -/(seq_le _ _ _) N2. +apply/eq_in_umfilt; case=>k1 v1 _ /=; rewrite /seq_le /seq_lt. +case: (leqP (index k1 (0::ks)) (index k (0::ks)))=>//=; last first. +- by move/(leq_ltn_trans T1)=>->. +by move/leq_trans=>-> //; rewrite orbF. +Qed. + +Lemma umfilt_opt_split A ks (h : natmap A) k v : + (k, v) \In h -> k \in ks -> + um_filter [pts x _ | k <=[ks] x && x <=[ks] k] h = k \-> v. +Proof. +move=>H K; have W : valid h by case: H. +have N: 0 < index k (0::ks) by move/In_dom/dom_cond: H; case: k K. clear N. +rewrite (In_eta H) umfiltPtUn -(In_eta H) W sle_refl /=. +rewrite umfilt_mem0L ?unitR ?validF //. +move=>k1 v1 /InF [_ /=]; rewrite andbC -eqn_leq /nat_index. +case: (index k1 _ =P index k _)=>// /esym /index_inj ->; +by rewrite 1?eq_refl // inE K orbT. +Qed. + +Lemma oev_umfilt_le A R ks a (h : natmap A) t (z0 : R) : + oeval a ks (um_filter (ole ks t) h) z0 = + oeval a [seq x <- ks | x <=[ks] t] h z0. +Proof. +rewrite oev_filter; apply: eq_in_oevF=>k v K. +by rewrite !In_umfilt. +Qed. + +Lemma umfilt_oleUn A ks (h1 h2 : natmap A) t : + valid (h1 \+ h2) -> t <=[ks] (olast_key ks h1) -> + {in dom h2, forall k, olast_key ks h1 <[ks] k} -> + um_filter (ole ks t) (h1 \+ h2) = um_filter (ole ks t) h1. +Proof. +move=>V K D; rewrite umfiltUn // (umfilt_ole0 (validR V)) ?unitR //. +by move=>k /D /(leq_ltn_trans K). +Qed. + +(* properties of the ordering *) + +Lemma ole0min x ks : 0 <=[ks] x. +Proof. by []. Qed. + +Lemma olt0min x ks : 0 <[ks] x = (x != 0). +Proof. by rewrite /seq_lt /nat_index /= (eq_sym 0); case: eqP. Qed. + +Lemma ole0 x ks : x <=[ks] 0 = (x == 0). +Proof. by rewrite /seq_le /nat_index /= (eq_sym 0); case: eqP. Qed. + +Lemma olt0 x ks : x <[ks] 0 = false. +Proof. by rewrite /seq_lt /nat_index /= (eq_sym 0); case: eqP. Qed. + +Lemma ole_dom x1 x2 ks : + x1 <=[ks] x2 -> x1 != 0 -> x2 \in ks -> (x1 \in ks). +Proof. +rewrite /seq_le /nat_index /= !(eq_sym 0) -!index_mem. +case: eqP; case: eqP=>//= _ _ N _. +by apply: leq_ltn_trans; rewrite -ltnS. +Qed. + +Lemma oleL k t ks : k <=[k :: ks] t = (k == 0) || (t != 0). +Proof. by rewrite /seq_le /nat_index /= eq_refl !(eq_sym 0); do !case: eqP. Qed. + +Lemma oleR k t ks : k <=[t :: ks] t = (k == 0) || (k == t). +Proof. +rewrite /seq_le /nat_index /= eq_refl !(eq_sym 0) -!(eq_sym t). +by do ![case: eqP=>//]; move=>->->. +Qed. + +Lemma oltL k t ks : k <[k :: ks] t = (k != t) && (t != 0). +Proof. by rewrite /seq_lt ltnNge -/(seq_le _ _ _) oleR negb_or andbC eq_sym. Qed. + +Lemma oltR k t ks : k <[t :: ks] t = (k == 0) && (t != 0). +Proof. by rewrite /seq_lt ltnNge -/(seq_le _ _ _) oleL negb_or negbK andbC. Qed. + +Lemma ole_consL x t k ks : x != k -> x <=[k::ks] t -> x <=[ks] t. +Proof. +rewrite /seq_le /nat_index /= => Kx; do ![case: eqP=>//]=>*; +by subst x; rewrite eq_refl in Kx. +Qed. + +Lemma ole_consR x t k ks : k != t -> x <=[ks] t -> x <=[k::ks] t. +Proof. +rewrite /seq_le /nat_index /= => Kx; do ![case: eqP=>//]=>*. +by subst k; rewrite eq_refl in Kx. +Qed. + +Lemma index_rcons (A : eqType) (a x : A) xs : + index a (rcons xs x) = + if a \in xs then index a xs else + if a == x then size xs else (size xs).+1. +Proof. +rewrite eq_sym; elim: xs=>[|y xs IH] //=. +rewrite inE eq_sym; case: eqP=>//= _. +by rewrite IH; case: ifP=>//; case: eqP. +Qed. + +Lemma index_memN (A : eqType) (x : A) xs : + x \notin xs -> index x xs = size xs. +Proof. +rewrite -index_mem -ltnNge ltnS. +by case: ltngtP (index_size x xs). +Qed. + +Lemma ole_rcons x1 x2 k ks : + x1 \in ks -> k \notin ks -> + x1 <=[rcons ks k] x2 = x1 <=[ks] x2. +Proof. +rewrite /seq_le /nat_index /= !(eq_sym 0). +case: eqP; case: eqP=>// N2 N1; rewrite !ltnS !index_rcons. +case: ifP; case: ifP=>// /negbT/index_memN ->; rewrite index_size. +case: ifP; first by rewrite index_size. +by move/leq_trans: (index_size x1 ks)=>->. +Qed. + +Lemma ole_rcons2 x1 x2 k ks : + x1 <=[ks] x2 -> x2 != k -> x1 <=[rcons ks k] x2. +Proof. +rewrite /seq_le /nat_index /= !(eq_sym 0). +case: (x1 =P 0)=>// N1; case: (x2 =P 0)=>// N2. +rewrite !ltnS !index_rcons. +case X1: (x1 \in ks); last first. +- move/negbT/index_memN: X1=>->. + case X2: (x2 \in ks); first by rewrite -index_mem in X2; case: ltnP X2. + by move/negbT/index_memN: X2=>->; case: eqP; case: eqP. +case X2: (x2 \in ks)=>//. +by move/negbT/index_memN: X2=>->; case: eqP=>// _ /leq_trans ->. +Qed. + +Lemma olt_rcons x t ks : + x \in ks -> t \notin ks -> t != 0 -> x <[rcons ks t] t. +Proof. +rewrite /seq_lt /nat_index /= !(eq_sym 0). +case: (x =P 0)=>[->|_ E1 E2 N]; first by case: eqP. +by rewrite !index_rcons eq_refl E1 (negbTE E2) (negbTE N) ltnS index_mem. +Qed. + +Lemma ole_take ks t : + uniq (0::ks) -> + [seq x <- ks | x <=[ks] t] = take (index t (0::ks)) ks. +Proof. +elim: ks=>//= k ks; rewrite inE negb_or -!(eq_sym k) -!(eq_sym t) -andbA. +move=>IH /and4P [H1 H2 H3 H4]; rewrite oleL (negbTE H1) /=. +case: (t =P 0) IH=>[->|/eqP T] /= IH. + rewrite -(filter_pred0 ks); apply: eq_in_filter=>x Kx /=. + by rewrite ole0; case: eqP Kx H2=>// ->->. +congr cons; case: eqP=>[->|/eqP N]. +- rewrite take0 -(filter_pred0 ks). + apply: eq_in_filter=>x Kx /=; rewrite oleR. + by case: eqP Kx H2 H3=>[->->|] //; case: eqP=>// -> _ ->. +rewrite -IH ?H2 ?H4 //; apply: eq_in_filter=>x Kx; apply/idP/idP. +- by apply: ole_consL; case: eqP Kx H3=>//->->. +by apply: ole_consR; rewrite eq_sym. +Qed. + +Lemma olt_take ks t : + uniq (0::ks) -> + [seq x <- ks | x <[ks] t] = take (index t (0::ks)).-1 ks. +Proof. +elim: ks=>//= k ks; rewrite inE negb_or -!(eq_sym t) -!(eq_sym k) -andbA. +move=>IH /and4P [H1 H2 H3 H4]; rewrite oltL andbC. +case: (t =P 0) IH=>[->|/eqP T] /= IH. +- rewrite -(filter_pred0 ks); apply: eq_in_filter=>x Kx /=. + by rewrite olt0. +case: (k =P t)=>[<-|N] /=. +- rewrite -(filter_pred0 ks); apply/eq_in_filter=>x Kx /=. + by rewrite oltR H1; case: eqP Kx H2=>// ->->. +congr cons; rewrite -IH ?H2 ?H4 //; apply: eq_in_filter=>x Kx. +rewrite /seq_lt !ltnNge; congr negb; apply/idP/idP. +- by apply: ole_consL; case: eqP N=>// ->. +by apply: ole_consR; case: eqP Kx H3=>// ->->. +Qed. + +Lemma ole_drop ks t : + uniq (0::ks) -> + [seq x <- ks | t <=[ks] x] = drop (index t (0::ks)).-1 ks. +Proof. +elim: ks=>//= k ks; rewrite inE negb_or -!(eq_sym k) -!(eq_sym t) -andbA. +move=>IH /and4P [H1 H2 H3 H4]; rewrite oleR. +case: (t =P 0) IH=>[->|/eqP T] /= IH; first by rewrite IH ?H2 ?H4 // drop0. +case: eqP=>E. +- congr cons; rewrite /seq_le -{4}(filter_predT ks). + apply: eq_in_filter=>x Kx. + by rewrite E -/(seq_le _ _ _) oleL orbC; case: eqP Kx H2=>// ->->. +rewrite -IH ?H2 ?H4 //; apply: eq_in_filter=>x Kx. +apply/idP/idP. +- by apply: ole_consL; case: eqP E. +by apply: ole_consR; case: eqP Kx H3=>// ->->. +Qed. + +Lemma olt_drop ks t : + uniq (0::ks) -> + [seq x <- ks | t <[ks] x] = drop (index t (0::ks)) ks. +Proof. +elim: ks=>//= k ks; rewrite inE negb_or -!(eq_sym k) -!(eq_sym t) -andbA. +move=>IH /and4P [H1 H2 H3 H4]; rewrite oltR H1 andbT. +case: (t =P 0) IH=>[->|/eqP T] /= IH. +- congr cons; rewrite drop0 in IH; rewrite /seq_lt -{4}IH ?H2 ?H4 //. + by apply: eq_in_filter=>x Kx; rewrite !ltnNge -/(seq_le _ _ _) ole0 olt0min. +case: (t =P k)=>[->|/eqP N] /=. +- rewrite drop0 /seq_lt -{4}(filter_predT ks). + apply: eq_in_filter=>x Kx; rewrite -/(seq_lt _ _ _) oltL. + case: (k =P x) Kx H1 H2 H3=>[->->|] //=. + by case: (x =P 0)=>// -> _ ->. +rewrite -IH ?H2 ?H4 //; apply: eq_in_filter=>x Kx. +rewrite /seq_lt !ltnNge; congr negb; apply/idP/idP. +- by apply: ole_consL; case: eqP Kx H3=>// ->->. +by apply: ole_consR; rewrite eq_sym. +Qed. + +Lemma olt_consL x t k ks : x != k -> x <[k::ks] t -> x <[ks] t. +Proof. +rewrite /seq_lt /nat_index /= => Kx; do ![case: eqP=>//]=>*. +by subst x; rewrite eq_refl in Kx. +Qed. + +Lemma olt_consR x t k ks : k != t -> x <[ks] t -> x <[k::ks] t. +Proof. +rewrite /seq_lt /nat_index /= => Kx; do ![case: eqP=>//]=>*; +by subst k; rewrite eq_refl in Kx. +Qed. + +Lemma olt_irr x ks : x <[ks] x = false. +Proof. by rewrite /seq_lt ltnn. Qed. + +Lemma index_filter_le (A : eqType) (p : pred A) (ks : seq A) (t1 t2 : A) : + t1 \in ks -> p t1 -> t2 \in ks -> p t2 -> + (index t1 (filter p ks) <= index t2 (filter p ks)) = + (index t1 ks <= index t2 ks). +Proof. +elim: ks t1 t2=>[|k ks IH] t1 t2 //; rewrite !inE /=. +case: ifP=>Pk //=; rewrite !(eq_sym k); +case: eqP Pk=>[->|N1]; case: eqP=>[->|N2] //= Pk Kt1 Pt1 Kt2 Pt2. +- by rewrite !ltnS IH //. +- by rewrite leqnn. +- by rewrite Pk in Pt1. +- by rewrite Pk in Pt2. +by rewrite ltnS IH. +Qed. + +Lemma index_filter_lt (A : eqType) (p : pred A) (ks : seq A) (t1 t2 : A) : + t1 \in ks -> p t1 -> t2 \in ks -> p t2 -> + (index t1 (filter p ks) < index t2 (filter p ks)) = + (index t1 ks < index t2 ks). +Proof. +elim: ks t1 t2=>[|k ks IH] t1 t2 //; rewrite !inE /=. +case: ifP=>Pk //=; rewrite !(eq_sym k); +case: eqP Pk=>[->|N1]; case: eqP=>[->|N2] //= Pk Kt1 Pt1 Kt2 Pt2. +- by rewrite !ltnS IH. +- by rewrite ltnn. +- by rewrite Pk in Pt1. +- by rewrite Pk in Pt2. +by rewrite ltnS IH. +Qed. + +Lemma olt_filter (p : pred nat) (ks : seq nat) (t1 t2 : nat) : + t1 \in filter p ks -> + t2 \in filter p ks -> + t1 <[filter p ks] t2 = t1 <[ks] t2. +Proof. +rewrite !mem_filter; case/andP=>Pt1 Kt1 /andP [Pt2 Kt2]. +rewrite /seq_lt /nat_index /=; case: eqP=>_; case: eqP=>_ //. +by rewrite !ltnS index_filter_lt. +Qed. + +Lemma ole_filter (p : pred nat) (ks : seq nat) (t1 t2 : nat) : + t1 \in filter p ks -> t2 \in filter p ks -> + t1 <=[filter p ks] t2 = t1 <=[ks] t2. +Proof. +rewrite !mem_filter; case/andP=>Pt1 Kt1 /andP [Pt2 Kt2]. +rewrite /seq_le /nat_index /=; case: eqP=>_; case: eqP=>_ //. +by rewrite !ltnS index_filter_le. +Qed. + +Lemma olt_mem x y ks : x <[ks] y -> x \in 0::ks. +Proof. +rewrite /seq_lt /nat_index /= inE (eq_sym 0); case: eqP=>//= _; case: eqP=>// _. +by rewrite ltnS -index_mem; move/leq_trans; apply; apply: index_size. +Qed. + +Lemma olt_catL xs1 xs2 t1 t2 : t1 <[xs1] t2 -> t1 <[xs1++xs2] t2. +Proof. +rewrite /seq_lt /= !index_cat /=. +case: ifP; case: ifP=>// _ _; rewrite !ltnS=>N. +case: ifP=>[|/negbT]; case: ifP=>[|/negbT] //. +- by move=>_ T1; apply: ltn_addr; rewrite index_mem. +- by rewrite -index_mem; move/(ltn_trans N); rewrite index_mem=>->. +rewrite -!index_mem -!leqNgt leq_eqVlt ltnNge index_size orbF. +by move=>/eqP ->; case: ltngtP N. +Qed. + +Lemma olt_catR xs1 xs2 t1 t2 : + uniq (xs1 ++ xs2) -> t2 \in xs2 -> + t1 <[xs2] t2 -> t1 <[xs1++xs2] t2. +Proof. +move=>U T2; rewrite /seq_lt /= !index_cat /=. +case: ifP; case: ifP=>// _ _; rewrite !ltnS=>N. +rewrite cat_uniq in U; case/and3P: U=>U1 U2 U3. +case: ifP=>[|/negbT]; case: ifP=>[|/negbT] //. +- have : t1 \in xs2. + - by rewrite -index_mem; apply: leq_trans N _; rewrite index_size. + by move/hasPn: U2=>U2 /U2 /= /negbTE ->. +- have : t1 \in xs2. + - by rewrite -index_mem; apply: leq_trans N _; rewrite index_size. + by move/hasPn: U2=>U2 /U2 /= /negbTE ->. +- by move/hasPn: U2; move/(_ _ T2)=>/= /negbTE ->. +by rewrite ltn_add2l. +Qed. + +(* various lemmas for splittings of the interval *) +(* categorized wrt. the boundaries of the interval are *) +(* closed/open on left/right *) +(* the 4 versions are: *) +(* ole_le_split: [0, t2] = [0, t1] + (t1, t2] *) +(* ole_lt_split: [0, t2] = [0, t1) + [t1, t2] *) +(* olt_le_split: [0, t2) = [0, t1] + (t1, t2) *) +(* olt_lt_split: [0, t2) = [0, t1) + [t1, t2) *) + +Lemma ole_le_split ks t1 t2 : + uniq (0::ks) -> t1 <=[ks] t2 -> + [seq x <- ks | x <=[ks] t2] = + [seq x <- ks | x <=[ks] t1] ++ + [seq x <- ks | t1 <[ks] x & x <=[ks] t2]. +Proof. +move=>U T. +have E1 : {in ks, forall x, x <=[ks] t1 = + x <=[ks] t2 && x <=[ks] t1}. +- move=>x Kx; case E: (x <=[ks] t1)=>//=; + by rewrite ?andbF // /seq_le (leq_trans E T). +have E2 : {in ks, forall x, t1 <[ks] x && x <=[ks] t2 = + x <=[ks] t2 && t1<[ks] x}. +- by move=>x Kx; rewrite andbC. +rewrite (eq_in_filter E1) (eq_in_filter E2) !filter_predI -filter_cat. +by rewrite [in RHS]ole_take // olt_drop // cat_take_drop. +Qed. + +Lemma ole_lt_split ks t1 t2 : + uniq (0::ks) -> t1 <=[ks] t2 -> + [seq x <- ks | x <=[ks] t2] = + [seq x <- ks | x <[ks] t1] ++ + [seq x <- ks | t1 <=[ks] x & x <=[ks] t2]. +Proof. +move=>U T. +have E1 : {in ks, forall x, x <[ks] t1 = + x <=[ks] t2 && x <[ks] t1}. +- move=>x Kx; case E: (x <[ks] t1)=>//=; + by rewrite ?andbF // /seq_le leq_eqVlt (leq_trans E T) orbT. +have E2 : {in ks, forall x, t1 <=[ks] x && x <=[ks] t2 = + x <=[ks] t2 && t1<=[ks] x}. +- by move=>x Kx; rewrite andbC. +rewrite (eq_in_filter E1) (eq_in_filter E2) !filter_predI -filter_cat. +by rewrite [in RHS]olt_take // ole_drop // cat_take_drop. +Qed. + +Lemma olt_le_split ks t1 t2 : + uniq (0::ks) -> t1 <[ks] t2 -> + [seq x <- ks | x <[ks] t2] = + [seq x <- ks | x <=[ks] t1] ++ + [seq x <- ks | t1 <[ks] x & x <[ks] t2]. +Proof. +move=>U T. +have E1 : {in ks, forall x, x <=[ks] t1 = + x <[ks] t2 && x <=[ks] t1}. +- move=>x Kx; case E: (x <=[ks] t1)=>//=; + by rewrite ?andbF // /seq_lt (leq_ltn_trans E T). +have E2 : {in ks, forall x, t1 <[ks] x && x <[ks] t2 = + x <[ks] t2 && t1<[ks] x}. +- by move=>x Kx; rewrite andbC. +rewrite (eq_in_filter E1) (eq_in_filter E2) !filter_predI -filter_cat. +by rewrite ole_take // [in RHS]olt_drop // cat_take_drop. +Qed. + +Lemma olt_lt_split ks t1 t2 : + uniq (0::ks) -> t1 <=[ks] t2 -> + [seq x <- ks | x <[ks] t2] = + [seq x <- ks | x <[ks] t1] ++ + [seq x <- ks | t1 <=[ks] x & x <[ks] t2]. +Proof. +move=>U T. +have E1 : {in ks, forall x, x <[ks] t1 = + x <[ks] t2 && x <[ks] t1}. +- move=>x Kx; case E: (x <[ks] t1)=>//=; + by rewrite ?andbF // /seq_lt (leq_trans E T). +have E2 : {in ks, forall x, t1 <=[ks] x && x <[ks] t2 = + x <[ks] t2 && t1<=[ks] x}. +- by move=>x Kx; rewrite andbC. +rewrite (eq_in_filter E1) (eq_in_filter E2) !filter_predI -filter_cat. +by rewrite [in RHS]olt_take // ole_drop // cat_take_drop. +Qed. + +(* lemmas about singletons *) +Lemma ole_leS ks k : + uniq ks -> + [seq x <- ks | k <=[ks] x & x <=[ks] k] = + if k \in ks then [:: k] else [::]. +Proof. +move=>U; case: ifP=>K. +- rewrite -(filter_pred1_uniq U K); apply: eq_in_filter=>x Kx. + rewrite -eqn_leq; apply/eqP/eqP=>[|->] //. + by move/esym/index_inj=>-> //; rewrite inE Kx orbT. +rewrite -(filter_pred0 ks); apply: eq_in_filter=>x Kx. +rewrite -eqn_leq; case: eqP=>// /esym/index_inj E. +by rewrite -E ?Kx // in K; rewrite inE Kx orbT. +Qed. + +Lemma ole_ltS ks k : [seq x <- ks | k <=[ks] x & x <[ks] k] = [::]. +Proof. +rewrite -(filter_pred0 ks); apply: eq_in_filter=>x Kx. +by rewrite /seq_le /seq_lt; case: ltngtP. +Qed. + +Lemma oev_ole_last A R a ks (h : natmap A) t (z0 : R) : + uniq (0::ks) -> + olast_key ks h <=[ks] t -> + oeval a [seq x <- ks | x <=[ks] t] h z0 = + oeval a ks h z0. +Proof. +move=>U H; apply: eq_in_oevK; rewrite -filter_predI. +apply: eq_in_filter=>x Kx /=; case D: (x \in dom h)=>//=. +by apply: leq_trans (olastkey_max U D Kx) H. +Qed. + +Lemma oev_fresh_ole A R a ks (h : natmap A) t fresh v (z0 : R) : + uniq (0::ks) -> valid (fresh \-> v \+ h) -> + t <[ks] fresh -> + oeval a [seq x <- ks | x <=[ks] t] (fresh \-> v \+ h) z0 = + oeval a [seq x <- ks | x <=[ks] t] h z0. +Proof. +move=>U V T; rewrite !oev_filter umfiltPtUn V /=. +by rewrite /seq_le leqNgt -/(seq_lt _ _ _) T. +Qed. + +Lemma oev_ole_ind A R a (P : R -> Prop) ks (h : natmap A) t1 t2 (z0 : R) : + uniq (0::ks) -> t1 <=[ks] t2 -> + P (oeval a [seq x <- ks | x <=[ks] t1] h z0) -> + (forall k v z0, (k, v) \In h -> k \in ks -> + t1 <[ks] k -> k <=[ks] t2 -> P z0 -> P (a z0 k v)) -> + P (oeval a [seq x <- ks | x <=[ks] t2] h z0). +Proof. +move=>U T P0 H; rewrite (ole_le_split U T) oev_cat. +apply: oev_ind=>// k v z Kh; rewrite mem_filter -andbA. +by case/and3P=>*; apply: H. +Qed. + +(* common case: functional version of the above lemma *) +Lemma oev_oleF A R X a (f : R -> X) ks (h : natmap A) t1 t2 (z0 : R) : + uniq (0::ks) -> t1 <=[ks] t2 -> + (forall k v z0, (k, v) \In h -> k \in ks -> + t1 <[ks] k -> k <=[ks] t2 -> f (a z0 k v) = f z0) -> + f (oeval a [seq x <- ks | x <=[ks] t1] h z0) = + f (oeval a [seq x <- ks | x <=[ks] t2] h z0). +Proof. +move=>U T H. +apply: (oev_ole_ind (P := fun x => f (oeval a [seq x <- ks | x <=[ks] t1] h z0) = f x)) T _ _=>//. +by move=>k v z1 D K T1 T2; rewrite H. +Qed. + +(* notation for executing up to (and including/excluding) t *) + +Notation oexec_le a ks t h z0 := + (oevalv a (filter (fun x => x <=[ks] t) ks) h z0). +Notation oexec_lt a ks t h z0 := + (oevalv a (filter (fun x => x <[ks] t) ks) h z0). + +(* some lemmas for the new notation *) + +Section OExec. +Variables (V R : Type) (aop : V -> R -> R). + +Lemma oex_le0 a ks (h : natmap V) (z0 : R) : oexec_le a ks 0 h z0 = z0. +Proof. +case W : (valid h); last first. +- by move/negbT/invalidE: W=>->; rewrite oev_undef. +have /eq_in_umfilt E : forall kv, kv \In h -> kv.1 <=[ks] 0 = pred0 kv. +- by move=>kv /In_dom/dom_cond; rewrite ole0=>/= /negbTE ->. +by rewrite oev_filter E umfilt_pred0 // oev0. +Qed. + +Lemma oex_leT a ks t (h : natmap V) (z0 : R) : + t \notin ks -> + oexec_le a ks t h z0 = if t == 0 then z0 else oevalv a ks h z0. +Proof. +move=>K. +case: eqP=>[->|/eqP T]; first by rewrite oex_le0. +suff : [seq x <- ks | x <=[ks] t] = ks by move=>->. +rewrite -[in RHS](filter_predT ks); apply/eq_in_filter=>x D. +rewrite /seq_le /nat_index /= !(eq_sym 0) (negbTE T); case: eqP=>// _. +by rewrite -!index_mem -ltnNge in K D; rewrite (ltn_trans D K). +Qed. + +Lemma oex_leT_last a ks k (h : natmap V) (z0 : R) : + uniq (0::rcons ks k) -> + oexec_le a (rcons ks k) k h z0 = oevalv a (rcons ks k) h z0. +Proof. +move=>/= /andP [U1 U2]. +have Nk : k \notin ks. +- by rewrite rcons_uniq in U2; case/andP: U2. +suff -> : [seq x <- rcons ks k | x <=[rcons ks k] k] = rcons ks k=>//. +rewrite -[in RHS](filter_predT (rcons ks k)); apply/eq_in_filter=>x D. +rewrite /seq_le /nat_index /= !(eq_sym 0); case: eqP=>// _. +case: eqP U1=>[->|_ U1]; first by rewrite /= mem_rcons inE eq_refl. +rewrite (index_rcons k k ks) eq_refl (negbTE Nk). +by rewrite -index_mem size_rcons in D. +Qed. + +Lemma oex_le_cat a ks1 ks2 t (h : natmap V) (z0 : R) : + uniq (0 :: ks1++ks2) -> + oexec_le a (ks1 ++ ks2) t h z0 = + if t \in ks1 then oexec_le a ks1 t h z0 + else oexec_le a ks2 t h (oexec_le a ks1 t h z0). +Proof. +case: (t =P 0)=>[->|/eqP T]; first by rewrite !oex_le0; case: ifP. +rewrite /= mem_cat negb_or -andbA=>/and3P [N1 N2]. +rewrite cat_uniq=>/and3P [U1 /hasPn /= U U2]. +have E1 : [seq x <- ks1++ks2 | x <=[ks1++ks2] t] = + [seq x <- ks1 | x <=[ks1++ks2] t] ++ + [seq x <- ks2 | x <=[ks1++ks2] t]. +- elim: ks1 N1 {U1 U}=>[|k ks1 IH] //=. + rewrite inE negb_or (eq_sym 0)=>/andP [K N1]. + by rewrite !oleL (negbTE K) T /= filter_cat. +have E2 : [seq x <- ks1 | x <=[ks1++ks2] t] = + [seq x <- ks1 | x <=[ks1] t]. +- apply/eq_in_filter=>x K; rewrite /seq_le /nat_index /= !(eq_sym 0) (negbTE T). + case: eqP=>// _; rewrite !index_cat K; case: ifPn=>// T'. + rewrite -!index_mem -ltnNge in K T'. + by rewrite !(ltn_trans K _) // ltnS leq_addr. +have E3 : [seq x <- ks2 | x <=[ks1++ks2] t] = + if t \in ks1 then [::] else [seq x <- ks2 | x <=[ks2] t]. +- case: ifPn=>T'; last first. + - apply/eq_in_filter=>k K; rewrite /seq_le /nat_index /= !(eq_sym 0) (negbTE T). + case: eqP K=>[->|_ K]; first by rewrite (negbTE N2). + by rewrite !index_cat (negbTE T') (negbTE (U _ K)) !ltnS leq_add2l. + rewrite -(filter_pred0 ks2); apply: eq_in_filter=>k K. + rewrite /seq_le /nat_index /= !(eq_sym 0) (negbTE T). + case: eqP K=>[->|_ K]; first by rewrite (negbTE N2). + apply/negbTE; rewrite !index_cat T' (negbTE (U _ K)) -ltnNge ltnS. + by rewrite -!index_mem in K T'; rewrite (leq_trans T' _) // leq_addr. +by rewrite E1 E2 oev_cat E3; case: ifP. +Qed. + +Lemma oex_le_cons a ks k t v (h : natmap V) (z0 : R) : + uniq (0::k::ks) -> + (k, v) \In h -> + oexec_le a (k :: ks) t h z0 = + if t == 0 then z0 else + if t == k then a z0 v else oexec_le a ks t h (a z0 v). +Proof. +move=>U H; case: (t =P 0)=>[->|/eqP T]; first by rewrite oex_le0. +rewrite (_ : k::ks = [::k] ++ ks) in U *; last by rewrite cat1s. +by rewrite oex_le_cat //= inE oleL T orbT /=; move/In_find: H=>->. +Qed. + +Lemma oex_rcons_in a ks k t (h : natmap V) (z0 : R) : + uniq (0::rcons ks k) -> t \in ks -> + oexec_le a (rcons ks k) t h z0 = oexec_le a ks t h z0. +Proof. +move=>U T; have N : t != 0. +- case: eqP T=>// ->; rewrite /= mem_rcons inE negb_or -andbA in U. + by case/and3P: U=>_ /negbTE ->. +rewrite (_ : rcons ks k = ks ++ [:: k]) in U *; last by rewrite cats1. +by rewrite oex_le_cat //= oleL T. +Qed. + +Lemma oex_le_rcons_notin a ks k t v (h : natmap V) (z0 : R) : + uniq (0::rcons ks k) -> (k, v) \In h -> t \notin ks -> + oexec_le a (rcons ks k) t h z0 = + if t == 0 then z0 else a (oexec_le a ks t h z0) v. +Proof. +move=>U H K; case: (t =P 0)=>[->|/eqP T]; first by rewrite oex_le0. +rewrite (_ : rcons ks k = ks ++ [:: k]) in U *; last by rewrite cats1. +by rewrite oex_le_cat //= oleL (negbTE K) T orbT /=; move/In_find: H=>->. +Qed. + +Lemma oex_le_consecutive a x1 t1 t2 x2 v (h : natmap V) (z0 : R) : + uniq (0::x1 ++ [:: t1, t2 & x2]) -> + (t2, v) \In h -> + oexec_le a (x1 ++ [:: t1, t2 & x2]) t2 h z0 = + a (oexec_le a (x1 ++ [:: t1, t2 & x2]) t1 h z0) v. +Proof. +move=>U H; have X : x1 ++ [:: t1, t2 & x2] = rcons x1 t1 ++ [:: t2 & x2]. +- by rewrite -cats1 -catA. +rewrite {}X /= in U *. +move: (U); rewrite !mem_cat cat_uniq /= !inE !negb_or -!andbA => U'. +case/and8P: U'=>U1 U2 U3 U4 U5 /hasPn /= U6 U7 U8. +rewrite !oex_le_cat // (negbTE U5) mem_rcons inE eq_refl. +rewrite (oex_le_cons _ _ _ _ H)=>/=; last by rewrite !inE negb_or U2 U3 U7 U8. +rewrite (eq_sym t2) (negbTE U2) eq_refl. +by rewrite oex_leT // (eq_sym t2) (negbTE U2) oex_leT_last //= U1 U4. +Qed. + +(* Now for oexec_lt, i.e., oexec with lt *) + +Lemma oex_lt0 a ks (h : natmap V) (z0 : R) : oexec_lt a ks 0 h z0 = z0. +Proof. +case W : (valid h); last first. +- by move/negbT/invalidE: W=>->; rewrite oev_undef. +have /eq_in_umfilt E : forall kv, kv \In h -> kv.1 <[ks] 0 = pred0 kv. +- by move=>kv /In_dom/dom_cond; rewrite olt0. +by rewrite oev_filter E umfilt_pred0 // oev0. +Qed. + +Lemma oex_ltT a ks t (h : natmap V) (z0 : R) : + t \notin ks -> + oexec_lt a ks t h z0 = if t == 0 then z0 else oevalv a ks h z0. +Proof. +move=>K. +case: eqP=>[->|/eqP T]; first by rewrite oex_lt0. +suff : [seq x <- ks | x <[ks] t] = ks by move=>->. +rewrite -[in RHS](filter_predT ks); apply/eq_in_filter=>x D. +rewrite /seq_lt /nat_index /= !(eq_sym 0) (negbTE T); case: eqP=>// _. +by rewrite -!index_mem -ltnNge ltnS in K D; rewrite ltnS (leq_trans D K). +Qed. + +Lemma oex_ltT_last a ks k (h : natmap V) (z0 : R) : + uniq (0::rcons ks k) -> + oexec_lt a (rcons ks k) k h z0 = oevalv a ks h z0. +Proof. +move=>/= /andP [U1 U2]. +have Nk : k \notin ks. +- by rewrite rcons_uniq in U2; case/andP: U2. +rewrite mem_rcons inE negb_or (eq_sym 0) in U1; case/andP: U1=>U U1. +suff -> : [seq x <- rcons ks k | x <[rcons ks k] k] = ks =>//. +rewrite /seq_lt filter_rcons ltnn -[in RHS](filter_predT ks); apply/eq_in_filter=>x D. +rewrite /nat_index /= !(eq_sym 0) (negbTE U) ltnS. +case: eqP D U1=>[->->//|_ D U1]. +by rewrite (index_rcons k k ks) eq_refl (negbTE Nk) index_rcons D index_mem. +Qed. + +Lemma oex_lt_cat a ks1 ks2 t (h : natmap V) (z0 : R) : + uniq (0 :: ks1++ks2) -> + oexec_lt a (ks1 ++ ks2) t h z0 = + if t \in ks1 then oexec_lt a ks1 t h z0 + else oexec_lt a ks2 t h (oexec_lt a ks1 t h z0). +Proof. +case: (t =P 0)=>[->|/eqP T]; first by rewrite !oex_lt0; case: ifP. +rewrite /= mem_cat negb_or -andbA=>/and3P [N1 N2]. +rewrite cat_uniq=>/and3P [U1 /hasPn /= U U2]. +have E1 : [seq x <- ks1++ks2 | x <[ks1++ks2] t] = + [seq x <- ks1 | x <[ks1++ks2] t] ++ + [seq x <- ks2 | x <[ks1++ks2] t]. +- elim: ks1 N1 {U1 U}=>[|k ks1 IH] //=. + rewrite inE negb_or (eq_sym 0)=>/andP [K N1]. + by rewrite !oltL T andbT; case: eqP=>//=; rewrite filter_cat. +have E2 : [seq x <- ks1 | x <[ks1++ks2] t] = + [seq x <- ks1 | x <[ks1] t]. +- apply/eq_in_filter=>x K; rewrite /seq_lt /nat_index /= !(eq_sym 0) (negbTE T). + case: eqP=>// _; rewrite !index_cat K; case: ifPn=>// T'. + rewrite -!index_mem -ltnNge in K T'. + by rewrite !ltnS !(leq_trans K _) // leq_addr. +have E3 : [seq x <- ks2 | x <[ks1++ks2] t] = + if t \in ks1 then [::] else [seq x <- ks2 | x <[ks2] t]. +- case: ifPn=>T'; last first. + - apply/eq_in_filter=>k K; rewrite /seq_lt /nat_index /= !(eq_sym 0) (negbTE T). + case: eqP K=>[->|_ K]; first by rewrite (negbTE N2). + by rewrite !index_cat (negbTE T') (negbTE (U _ K)) !ltnS ltn_add2l. + rewrite -(filter_pred0 ks2); apply: eq_in_filter=>k K. + rewrite /seq_lt /nat_index /= !(eq_sym 0) (negbTE T). + case: eqP K=>[->|_ K]; first by rewrite (negbTE N2). + apply/negbTE; rewrite !index_cat T' (negbTE (U _ K)) -ltnNge !ltnS ltnW //. + by rewrite -!index_mem in K T'; rewrite (leq_trans T' _) // leq_addr. +by rewrite E1 E2 oev_cat E3; case: ifP. +Qed. + +Lemma oex_lt_cons_same a ks k (h : natmap V) (z0 : R) : + uniq (0::k::ks) -> + oexec_lt a (k :: ks) k h z0 = z0. +Proof. +move=>U; rewrite (_ : k::ks = [::k] ++ ks) in U *; last by rewrite cat1s. +by rewrite oex_lt_cat //= inE oltL eq_refl. +Qed. + +Lemma oex_lt_cons a ks k t v (h : natmap V) (z0 : R) : + uniq (0::k::ks) -> + (k, v) \In h -> + oexec_lt a (k :: ks) t h z0 = + if t == 0 then z0 else + if t == k then z0 else oexec_lt a ks t h (a z0 v). +Proof. +move=>U H; case: (t =P 0)=>[->|/eqP T]; first by rewrite oex_lt0. +rewrite (_ : k::ks = [::k] ++ ks) in U *; last by rewrite cat1s. +rewrite oex_lt_cat //= inE oltL T andbT (eq_sym t); case: eqP=>//=. +by move/In_find: H=>->. +Qed. + +Lemma oex_lt_rcons_in a ks k t (h : natmap V) (z0 : R) : + uniq (0::rcons ks k) -> t \in ks -> + oexec_lt a (rcons ks k) t h z0 = oexec_lt a ks t h z0. +Proof. +move=>U T; have N : t != 0. +- case: eqP T=>// ->; rewrite /= mem_rcons inE negb_or -andbA in U. + by case/and3P: U=>_ /negbTE ->. +rewrite (_ : rcons ks k = ks ++ [:: k]) in U *; last by rewrite cats1. +by rewrite oex_lt_cat //= oltL T. +Qed. + +Lemma oex_lt_rcons_notin a ks k t v (h : natmap V) (z0 : R) : + uniq (0::rcons ks k) -> (k, v) \In h -> t \notin ks -> + oexec_lt a (rcons ks k) t h z0 = + if t == 0 then z0 else + if t == k then oexec_lt a ks t h z0 else a (oexec_lt a ks t h z0) v. +Proof. +move=>U H K; case: (t =P 0)=>[->|/eqP T]; first by rewrite oex_lt0. +rewrite (_ : rcons ks k = ks ++ [:: k]) in U *; last by rewrite cats1. +rewrite oex_lt_cat //= oltL (negbTE K) T andbT (eq_sym k); case: eqP=>//=. +by move/In_find: H=>->. +Qed. + +Lemma oex_lt_consecutive a x1 t1 t2 x2 (h : natmap V) (z0 : R) : + uniq (0::x1 ++ [:: t1, t2 & x2]) -> + oexec_lt a (x1 ++ [:: t1, t2 & x2]) t2 h z0 = + oexec_le a (x1 ++ [:: t1, t2 & x2]) t1 h z0. +Proof. +move=>U; have X : x1 ++ [:: t1, t2 & x2] = rcons x1 t1 ++ [:: t2 & x2]. +- by rewrite -cats1 -catA. +rewrite {}X /= in U *. +move: (U); rewrite !mem_cat cat_uniq /= !inE !negb_or -!andbA => U'. +case/and8P: U'=>U1 U2 U3 U4 U5 /hasPn /= U6 U7 U8. +rewrite oex_lt_cat // oex_le_cat // (negbTE U5) mem_rcons inE eq_refl. +rewrite (_ : t2 :: x2 = [:: t2] ++ x2) // oex_lt_cat; last first. +- by rewrite /= inE negb_or U2 U3 U7 U8. +rewrite inE eq_refl //= oltL eq_refl /= oex_ltT //. +by rewrite (eq_sym t2) (negbTE U2) oex_leT_last //= U1 U4. +Qed. + +Lemma oex_le_lt_split a x1 x2 (h : natmap V) t v (z0 : R) : + uniq (0::x1 ++ [:: t & x2]) -> + (t, v) \In h -> + oexec_le a (x1 ++ [:: t & x2]) t h z0 = + a (oexec_lt a (x1 ++ [::t & x2]) t h z0) v. +Proof. +move=>U /In_find E; have X : x1 ++ [:: t & x2] = rcons x1 t ++ x2. +- by rewrite -cats1 -catA. +rewrite {}X /= in U *. +move: (U); rewrite !mem_cat cat_uniq /= negb_or -!andbA. +case/and5P=>U1 U2 U3 /hasPn /= U4 U5. +have U' : uniq (0 :: rcons x1 t) by rewrite /= U1 U3. +rewrite oex_le_cat // oex_lt_cat // mem_rcons inE eq_refl /=. +by rewrite oex_leT_last // oex_ltT_last // oev_rconsE E. +Qed. + +End OExec. + +(* induction principle for the new notation *) + +Definition consecutive ks t1 t2 := + [&& t1 <[ks] t2, t2 \in ks & + ~~ has (fun z => t1 <[ks] z && z <[ks] t2) ks]. + +Lemma consec_mem1 ks t1 t2 : consecutive ks t1 t2 -> t1 \in 0::ks. +Proof. +case/and3P=>H1 H2 _; rewrite inE. +move: H1; rewrite /seq_lt /nat_index /= !(eq_sym 0). +case: eqP=>// _; case: eqP=>//= _; rewrite ltnS. +by rewrite -index_mem=>/leq_trans; apply; apply: index_size. +Qed. + +Lemma consec_mem2 ks t1 t2 : consecutive ks t1 t2 -> t2 \in ks. +Proof. by case/and3P. Qed. + +Lemma cons_ind (ks : seq nat) (P : nat -> Prop) : + uniq (0 :: ks) -> + P 0 -> + (forall t1 t2, consecutive ks t1 t2 -> P t1 -> P t2) -> + forall t, t \in 0::ks -> P t. +Proof. +move=>/= /andP [U0 Uk] Bx IHx t; rewrite inE; case/orP=>[/eqP ->//|]. +elim: ks U0 Uk IHx=>[|k ks IH] //=. +rewrite !inE !(eq_sym 0) negb_or=>/andP [Nk K0] /andP [Kk Uk] IHx. +have Nzk z : z \in ks -> z != 0 by case: eqP K0=>// -> /negbTE ->. +have Kzk z : z \in ks -> z != k by case: eqP Kk=>// -> /negbTE ->. +have C0k : consecutive (k::ks) 0 k. +- rewrite /consecutive olt0min Nk inE eq_refl /=. + rewrite negb_or oltL eq_refl andbF /=. + by apply/hasPn=>z /Nzk X; rewrite oltR (negbTE X) andbF. +case/orP=>[/eqP ->|]; first by apply: IHx C0k Bx. +apply: IH=>// t1 t2 /and3P [C1 C2 /hasPn /= C3]. +have N1 : t1 != k. +- by case: eqP C1=>// -> /olt_mem; rewrite inE (negbTE Nk) (negbTE Kk). +have N2 : t2 != k by case: eqP C2 Kk=>// ->->. +have D1 : t1 <[k :: ks] t2 by apply: olt_consR C1; rewrite eq_sym. +have Nt2 : t2 != 0 by case: eqP C1=>// ->; rewrite olt0. +case: (t1 =P 0)=>[E _|/eqP N]. +- suff Ckt2 : consecutive (k::ks) k t2 by apply: (IHx) Ckt2 (IHx _ _ C0k Bx). + rewrite /consecutive /= !oltL eq_sym N2 Nt2 inE C2 orbT eq_refl /=. + apply/hasPn=>z Kz; apply: contra (C3 z Kz). + by rewrite E olt0min (Nzk _ Kz); case/andP=>_ /(olt_consL (Kzk _ Kz)) ->. +suff Ct1t2 : consecutive (k::ks) t1 t2 by apply: IHx. +rewrite /consecutive D1 inE C2 orbT /= negb_or oltR !negb_and N /=. +apply/hasPn=>z Kz; apply: contra (C3 z Kz)=>/andP [X1 X2]. +by rewrite (olt_consL N1 X1) (olt_consL (Kzk _ Kz) X2). +Qed. + +Lemma cons_splitP (ks : seq nat) t1 t2 : + uniq (0 :: ks) -> + reflect (exists x1 x2, 0 :: ks = x1 ++ [:: t1, t2 & x2]) + (consecutive ks t1 t2). +Proof. +move=>U; move: (U)=>/= /andP [K0 Ku]. +case X : (consecutive ks t1 t2); constructor. +- case/and3P: X=>H1 H2 /hasPn H3. + rewrite -[0::ks](cat_take_drop (index t2 (0::ks)).+1). + rewrite /= -(olt_drop t2 U) -(ole_take t2 U). + rewrite (ole_lt_split (t1:=t2)) // ole_leS // H2 -catA /=. + rewrite (olt_le_split U H1). + rewrite (_ : [seq x <- ks | t1 <[ks] x & x <[ks] t2] = [::]); last first. + - rewrite (eq_in_filter (a2:=pred0)) ?filter_pred0 //. + by move=>z /H3 /negbTE ->. + rewrite -catA /=; move: (olt_mem H1); rewrite inE. + case/orP=>[/eqP ->|Kt1]; last first. + - rewrite (ole_lt_split (t1:=t1)) // ole_leS // Kt1 -catA /= -cat1s catA. + by eexists _, _. + exists [::], [seq x <- ks | t2 <[ks] x]. + rewrite (eq_in_filter (a2:=pred0)) ?filter_pred0 //. + by move=>z; rewrite ole0; case: eqP K0=>// -> /negbTE ->. +move=>H; move/negP: X; apply; case: H; case=>[|x xs][ys] /= E. +- case: E U K0 Ku=><- -> U. + rewrite inE negb_or (eq_sym 0)=>/andP [Nt2 Nys] /= /andP [Nty2 Uy]. + rewrite /consecutive /= !oltR (negbTE Nt2) inE !eq_refl /=. + by apply/hasPn=>z _; rewrite olt0min oltR; case: z. +case: E U K0 Ku=>_ -> U K0 Ku; apply/and3P; split. +- apply: olt_catR=>//; first by rewrite !inE eq_refl orbT. + rewrite cat_uniq in Ku; rewrite oltL; case/and3P: Ku=>_ _. + case: eqP=>[->|_ _]; first by rewrite /= inE eq_refl. + by case: eqP K0=>// ->; rewrite mem_cat !inE eq_refl !orbT. +- by rewrite mem_cat !inE eq_refl !orbT. +apply/hasPn=>z Kz; rewrite mem_cat !inE !negb_or in K0. +case/and4P: K0=>X1 X2 X3 X4; rewrite cat_uniq /= 2!negb_or in Ku. +case/and5P: Ku=>Y1 /and3P [Y2 Y3 Y4]. +rewrite inE negb_or=>/andP [Y5 Y6] Y7 Y8. +rewrite /seq_lt /= (negbTE X2) (negbTE X3). +rewrite !index_cat (negbTE Y2) (negbTE Y3) /= (negbTE Y5). +rewrite !eq_refl /= addn0 addn1 ltnS. +case: ifP Kz=>[/eqP <-|Nz]. +- by rewrite mem_cat (negbTE X1) /= !inE (negbTE X2) (negbTE X3) (negbTE X4). +rewrite !ltnS; case: ifP=>_; first by case: ltngtP. +case: ifP=>_; first by rewrite addn0 ltnn. +case: ifP=>[|_] _; first by rewrite addn1 ltnn andbF. +by rewrite negb_and; apply/orP; right; rewrite -ltnNge addnS ltnS leq_addr. +Qed. + +Lemma oex_le_lt V R a t v (h : natmap V) ks (z0 : R) : + uniq (0::ks) -> t \in ks -> (t, v) \In h -> + oexec_le a ks t h z0 = a (oexec_lt a ks t h z0) v. +Proof. by move=>U K H; case/splitPr: K U=>x1 x2 U; apply: oex_le_lt_split. Qed. + +Lemma oex_lt_le_consec V R a t1 t2 (h : natmap V) ks (z0 : R) : + uniq (0::ks) -> + consecutive ks t1 t2 -> + oexec_lt a ks t2 h z0 = oexec_le a ks t1 h z0. +Proof. +move=>U C; case/(cons_splitP _ _ U): C=>x1 [x2] E; rewrite E in U. +case: x1 E U=>[|x x1][??]; last by subst x ks; apply: oex_lt_consecutive. +subst t1 ks; rewrite /= inE negb_or=>/and3P [] /andP [H1 H2] H3 H4. +rewrite /= oltL oleL eq_refl (eq_sym t2) (negbTE H1) /=. +congr oeval; apply: eq_in_filter=>z. +by rewrite oltR ole0 (eq_sym t2) H1 andbT. +Qed. + +Lemma oex_lt_consec V R a t1 t2 v (h : natmap V) ks (z0 : R) : + uniq (0::ks) -> + (t1, v) \In h -> + consecutive ks t1 t2 -> + oexec_lt a ks t2 h z0 = a (oexec_lt a ks t1 h z0) v. +Proof. +move=>U H C; case/(cons_splitP _ _ U): C=>x1 [x2] E; rewrite E in U. +case: x1 E U=>[|x x1] /=; first by case=>?; subst t1; move/In_dom/dom_cond: H. +by case=>? E U; subst x; rewrite E -oex_le_lt_split //; apply: oex_lt_consecutive. +Qed. + +Lemma oex_lt_consec0 V R a t (h : natmap V) ks (z0 : R) : + uniq (0::ks) -> + consecutive ks 0 t -> + oexec_lt a ks t h z0 = z0. +Proof. +move=>U C; case/(cons_splitP _ _ U): C=>x1 [x2]; case: x1=>[|x xs] /=. +- by case=>?; subst ks; rewrite oex_lt_cons_same. +by case=>??; subst x ks; move: U; rewrite /= mem_cat inE eq_refl !orbT. +Qed. + +(* A restatement of the induction principle into a form *) +(* that's more usable in the common cases *) + +Lemma consec0_ind (ks : seq nat) (P : nat -> Prop) : + uniq (0 :: ks) -> + P 0 -> + (forall t, consecutive ks 0 t -> P 0 -> P t) -> + (forall t1 t2, t1 \in ks -> consecutive ks t1 t2 -> P t1 -> P t2) -> + forall t, t \in 0::ks -> P t. +Proof. +move=>U H0 H1 Ht; apply: cons_ind=>// t1 t2 C. +move: {C} (consec_mem1 C) (C); rewrite inE; case/orP. +- by move/eqP=>->; apply: H1. +by apply: Ht. +Qed. + +Lemma consec_ind (ks : seq nat) (P : nat -> Prop) : + uniq (0 :: ks) -> + (forall t, consecutive ks 0 t -> P t) -> + (forall t1 t2, t1 \in ks -> consecutive ks t1 t2 -> P t1 -> P t2) -> + forall t, t \in ks -> P t. +Proof. +move=>U H1 Ht t D. +have D0 : t \in 0::ks by rewrite inE D orbT. +move: t D0 D; apply: consec0_ind=>//. +- by move: U=>/= /andP [/negbTE ->]. +- by move=>t /H1. +by move=>t1 t2 Kt1 C H _; apply: Ht C (H Kt1). +Qed. + +Lemma oleq_eqVlt ks t1 t2 : + t1 \in ks \/ t2 \in ks -> + t1 <=[ks] t2 = (t1 == t2) || (t1 <[ks] t2). +Proof. +move=>H; rewrite /seq_lt /seq_le leq_eqVlt /=. +case: ifP; case: ifP. +- by move/eqP=><- /eqP <-. +- by move=>N /eqP <-; rewrite N. +- by move/eqP=><- N; rewrite /= !ltn0 orbF -(eq_sym 0) N. +move=>N1 N2; rewrite !ltnS eqSS. +case: (t1 =P t2)=>[->|N] /=; first by rewrite eq_refl. +case: eqP=>//=; case: H=>H; first by move/(index_inj H)/N. +by move/esym/(index_inj H)/esym/N. +Qed. + +Lemma olt_sorted ks t1 t2 : + sorted ltn ks -> t2 \in ks -> t1 <[ks] t2 -> t1 < t2. +Proof. +move=>S Kt2; rewrite /seq_lt /=. +case: ifP; case: ifP=>//; first by move=>H2 /eqP <-; case: t2 H2 Kt2. +by move=>_ _; rewrite ltnS; apply: sorted_index_ord ltnn ltn_trans S Kt2. +Qed. + +Lemma ole_sorted ks t1 t2 : + sorted ltn ks -> t2 \in ks -> t1 <=[ks] t2 -> t1 <= t2. +Proof. +move=>S K; rewrite oleq_eqVlt; last by right. +by case/orP=>[/eqP -> //|/(olt_sorted S K)/ltnW]. +Qed. + +Lemma oltnNge ks t1 t2 : t1 <[ks] t2 = ~~ t2 <=[ks] t1. +Proof. +rewrite /seq_lt /seq_le /= !(eq_sym 0). +by case: ifP; case: ifP=>// N2 N1; rewrite !ltnS ltnNge. +Qed. + +Lemma oleqNgt ks t1 t2 : t1 <=[ks] t2 = ~~ t2 <[ks] t1. +Proof. by rewrite oltnNge negbK. Qed. + +Lemma consec_sorted ks t1 t2 : + sorted ltn ks -> consecutive ks t1 t2 -> + forall z, z \in 0::ks -> (z <= t1) = (z < t2). +Proof. +move=>S /and3P [T1 T2 /hasPn T3]; move: (olt_mem T1)=>D1. +move/(olt_sorted S T2): T1=>T1. +move=>z; rewrite inE; case/orP. +- by move/eqP=>->; rewrite leq0n; case: t2 {T2 T3} T1. +move=>Dz; move: (T3 z Dz); rewrite !negb_and -!oleqNgt. +case/orP=>D; last first. +- move: (ole_sorted S Dz D)=>Z2. + by rewrite leqNgt (leq_trans T1 Z2) ltnNge Z2. +rewrite inE in D1; case/orP: D1 D=>[/eqP ->|D1 D]. +- rewrite ole0=>/eqP ->; rewrite leqnn. + by rewrite (leq_ltn_trans _ T1). +move: (ole_sorted S D1 D)=>Z1. +by rewrite Z1 (leq_ltn_trans Z1 T1). +Qed. + +(* we frequently need to talk about consecutivity of stamps *) +(* wrt. some additional property *) + +Definition consecutivep ks p t1 t2 := + [&& t1 <[ks] t2, t2 \in ks & + ~~ has (fun z => [&& t1 <[ks] z, z <[ks] t2 & p z]) ks]. + +Lemma consp_ind (ks : seq nat) (p : pred nat) (P : nat -> Prop) : + uniq (0 :: ks) -> + P 0 -> + (forall t1 t2, t1 <[ks] t2 -> + t2 \in filter p ks -> + ~~ has (fun z => t1 <[ks] z && z <[ks] t2) (filter p ks) -> + P t1 -> P t2) -> + forall t : nat, t \in 0::filter p ks -> P t. +Proof. +move=>/= /andP [U1 U2] H IH; apply: cons_ind=>//=. +- by rewrite mem_filter negb_and filter_uniq // U1 orbT. +move=>t1 t2 /and3P [X H2 /hasPn H3]; move: (olt_mem X); rewrite inE=>H1. +apply: IH=>//. +- case/orP: H1=>[/eqP ->|]; last by move/olt_filter=><-. + by rewrite olt0min; case: eqP X=>// ->; rewrite olt0. +apply/hasPn=>z Kz; apply: contra (H3 _ Kz). +by case/orP: H1=>[/eqP ->|H1]; rewrite ?olt0min !olt_filter. +Qed. + + +(* the generalization of the growing argument *) + +Section OGrowing. +Variables (V R : Type) (X : ordType) (a : R -> V -> R) (f : R -> X). + +Lemma ohelper0 ks h z0 : + growing a f h -> oleq (f z0) (f (oevalv a ks h z0)). +Proof. +move=>G; elim: ks z0=>[|k ks IH] z0 //=; apply: otrans (IH _). +by case E: (find k h)=>[b|] //; move/In_find: E; apply: G. +Qed. + +Lemma ohelper1 h k ks z0 v : + growing a f h -> (k, v) \In h -> + f (oevalv a (k::ks) h z0) = f z0 -> + f (a z0 v) = f z0. +Proof. +move=>G H /=; move/In_find: (H)=>-> E; apply/eqP; case: ordP=>//. +- by move: (G k v z0 H); rewrite /oleq eq_sym; case: ordP. +by move: (ohelper0 ks (a z0 v) G); rewrite E /oleq eq_sym; case: ordP. +Qed. + +Lemma ohelper2 h ks1 k ks2 z0 v : + growing a f h -> (k, v) \In h -> + f (oevalv a (ks1++k::ks2) h z0) = f z0 -> + f (a (oevalv a ks1 h z0) v) = f (oevalv a ks1 h z0). +Proof. +move=>G H E1; move/In_find: (H)=>F. +suff E2 : f (oevalv a ks1 h z0) = f z0. +- by rewrite oev_cat -E2 in E1; apply: ohelper1 G H E1. +apply/eqP; case: ordP=>//. +- by move: (ohelper0 ks1 z0 G); rewrite /oleq eq_sym; case: ordP. +suff: oleq (f (oevalv a ks1 h z0)) (f z0). +- by rewrite /oleq eq_sym; case: ordP. +by rewrite -E1 oev_cat; apply: ohelper0. +Qed. + +(* "introducing" growth *) +Lemma ogrowI ks h t1 t2 z0 : + growing a f h -> uniq (0::ks) -> t1 <=[ks] t2 -> + oleq (f (oexec_le a ks t1 h z0)) (f (oexec_le a ks t2 h z0)). +Proof. by move=>G U N; rewrite (ole_le_split U N) oev_cat; apply: ohelper0. Qed. + +(* "eliminating" growth *) +Lemma ogrowE ks h t1 t2 z0 k v : + growing a f h -> + uniq (0::ks) -> + (k, v) \In h -> k \in ks -> + t1 <[ks] k -> k <=[ks] t2 -> + f (oexec_le a ks t2 h z0) = f (oexec_le a ks t1 h z0) -> + f (a (oevalv a [seq x <- ks | x <[ks] k] h z0) v) = + f (oevalv a [seq x <- ks | x <[ks] k] h z0). +Proof. +move=>G U H K N1 N2 E; have Uk : uniq ks by case/andP: U. +set k0 := [seq x <- ks | x <=[ks] t1] in E. +set k1 := [seq x <- ks | t1 <[ks] x & x <[ks] k]. +set k2 := [seq x <- ks | k <[ks] x & x <=[ks] t2]. +have E1 : [seq x <- ks | x <[ks] k] = k0 ++ k1 by rewrite (olt_le_split U N1). +have E2 : [seq x <- ks | x <=[ks] t2] = (k0 ++ k1) ++ k :: k2. +- rewrite (ole_le_split U N2) -E1 -cat_rcons -cats1. + by rewrite (ole_lt_split U (leqnn _)) ole_leS // K. +by rewrite E1 oev_cat; rewrite E2 -catA oev_cat in E; apply: ohelper2 G H E. +Qed. + +End OGrowing. + +(**********************) +(* umall2 and natmaps *) +(**********************) + +Section Umall2. +Variables (A : Type) (P : A -> A -> Prop). + +Lemma umall2_lastkey (h1 h2 : natmap A) : + um_all2 P h1 h2 -> last_key h1 = last_key h2. +Proof. by move/umall2_dom=>E; rewrite /last_key E. Qed. + +Lemma umall2_umfilt_lastkey (h1 h2 : natmap A) : + um_all2 P h1 (um_filter (le (last_key h1)) h2) <-> + forall k, k <= last_key h1 -> optionR P (find k h1) (find k h2). +Proof. +pose f t (kv : nat*A) := oleq kv.1 t. +have E (h : natmap A) : forall kv : nat * A, kv \In h -> + le (last_key h1) kv = f (last_key h1) kv. +- by case=>k v _; rewrite /f/oleq/ord /=; case: ltngtP. +rewrite -{1}[h1]umfilt_last. +move/eq_in_umfilt: (E h2)=>->. +move/eq_in_umfilt: (E h1)=>->. +rewrite umall2_umfilt_oleq; split=>H k N; apply: H; +by move: N; rewrite /oleq/ord /=; case: ltngtP. +Qed. + +End Umall2. + +(* Transitive closure of relations between stamps *) +(* in a natmap. This is somewhat streamlined compared *) +(* to plan transitive closure rtc. *) +(* First, we don't want reflexive closure. *) +(* Second, we want a different closure condition *) +(* Because this is really about histories, we name the *) +(* lemmas using a prefix "h", rather than "nm" for natmaps *) + +Section HistoryTransitiveClosure. +Variables (A : Type) (h : natmap A) (R : rel nat). +Hypothesis HRclosed : forall x y, R x y -> (x \in 0::dom h) && (y \in 0::dom h). + +Definition hrtc x y := rtc (0::dom h) R x y. +Definition htc x y := tc (0::dom h) R x y. + +Definition acyclic (R : rel nat) := forall x y n, iter' R n.+1 x y -> x != y. + +Lemma hclosed x y : R x y -> (x \in 0::dom h) = (y \in 0::dom h). +Proof. by case/HRclosed/andP=>->->. Qed. + +Lemma hrtc_pathP x y : + reflect (x \in 0::dom h /\ exists2 p, path R x p & y = last x p) + (hrtc x y). +Proof. by apply/(rtc_pathP hclosed). Qed. + +Lemma hrtcP x y : reflect (x \in 0::dom h /\ iter R x y) (hrtc x y). +Proof. by apply/(rtcP hclosed). Qed. + +Lemma hrtc_pathP_last x y : + reflect (y \in 0::dom h /\ exists2 p, path R x p & y = last x p) + (hrtc x y). +Proof. by apply/(rtc_pathP_last hclosed). Qed. + +Lemma hrtcP_last x y : reflect (y \in 0::dom h /\ iter R x y) (hrtc x y). +Proof. by apply/(rtcP_last hclosed). Qed. + +Lemma htcP_closed x y : htc x y -> (x \in 0::dom h) && (y \in 0::dom h). +Proof. +rewrite /htc/tc=>/andP [->] /hasP [z Mz /andP [Rxz]]. +case/(rtcP_last hclosed)=>My; case; case=>[<-|n /iterS [w][_]]; +by [case/HRclosed/andP: Rxz | case/HRclosed/andP]. +Qed. + +Lemma htcP x y : reflect (exists n, iter' R n.+1 x y) (htc x y). +Proof. +rewrite /htc; case: (tcP hclosed)=>H; constructor; first by case: H. +case=>n /= [z][Rxz J]; apply: H; split; last by exists n, z. +by case/HRclosed/andP: Rxz. +Qed. + +Lemma hrtc_acyc_asym : acyclic R -> antisymmetric hrtc. +Proof. +move=>H x y /andP [/hrtcP [_ [n1 H1]] /hrtcP [_ [n2 H2]]]. +case: n1 H1=>[|n1 H1] //; move: (H x x (n1+n2)). +by rewrite -addSn=>/(_ (iter'_add H1 H2)); rewrite eq_refl. +Qed. + +Lemma hrtc_asym_acyc : irreflexive R -> antisymmetric hrtc -> acyclic R. +Proof. +move=>I S x s n /= [y][H1 H2]; case: eqP=>// N; subst s. +case/andP: (HRclosed H1)=>Y1 Y2; move: (S x y) (H1). +have -> : hrtc x y by apply/hrtcP; split=>//; exists 1, y. +have -> : hrtc y x by apply/hrtcP; split=>//; exists n. +by move=>-> //; rewrite I. +Qed. + +Lemma htc_acyc : acyclic R <-> irreflexive htc. +Proof. +split=>[H x|H x y n]; first by case: htcP=>//; case=>n /H; rewrite eq_refl. +case: eqP=>// <-{y} J; have: exists n, iter' R n.+1 x x by exists n. +by move/htcP; rewrite H. +Qed. + +Lemma htc_trans : transitive htc. +Proof. by apply: tc_trans=>x y /HRclosed /andP [->]; rewrite inE orbC=>->. Qed. + +Lemma htc_antisym : acyclic R -> antisymmetric htc. +Proof. by move/htc_acyc=>H x y /andP [X /(htc_trans X)]; rewrite H. Qed. + +Lemma htc1 x y : R x y -> htc x y. +Proof. by move=>Rxy; apply/htcP; exists 0, y. Qed. + +End HistoryTransitiveClosure. + +Lemma htc_idemp A (h : natmap A) (R : rel nat) : + (forall x y, R x y -> (x \in 0::dom h) && (y \in 0::dom h)) -> + htc h (htc h R) =2 htc h R. +Proof. by move=>Rcond x y; rewrite /htc tc_idemp. Qed. + +Lemma htc_sub A (h : natmap A) (R1 R2 : rel nat) x y : + (forall x y, R2 x y -> (x \in 0::dom h) && (y \in 0::dom h)) -> + (forall x y, R1 x y -> R2 x y) -> + htc h R1 x y -> htc h R2 x y. +Proof. by move=>Rcond Rsub; apply: tc_sub. Qed. + +Lemma htc_ind2 A (h : natmap A) (R1 : rel nat) (R2 : nat -> nat -> Prop) x y : + (forall x y, R1 x y -> (x \in 0::dom h) && (y \in 0::dom h)) -> + (forall x y, R1 x y -> R2 x y) -> + Transitive R2 -> + htc h R1 x y -> R2 x y. +Proof. +move=>H S T /(htcP H) [n]. +elim: n x y=>[|n IH] x y /= [s][X]; first by move=><-; apply: S. +by move/IH; apply: T (S _ _ X). +Qed. + +Lemma htc_last A (h : natmap A) (R : rel nat) (P : Pred nat) : + (forall x y, R x y -> (x \in 0::dom h) && (y \in 0::dom h)) -> + (forall x y, R x y -> P y) -> + (forall x y, htc h R x y -> P y). +Proof. +move=>H L x y /htcP [] // n; elim: n x y=>[|n IH] x y /= [s][X]; +by [move=><-; apply: L X |move/IH]. +Qed. + +Lemma hrtc_last A (h : natmap A) (R : rel nat) (P : Pred nat) : + (forall x y, R x y -> (x \in 0::dom h) && (y \in 0::dom h)) -> + (forall x y, R x y -> P y) -> + (forall x y, hrtc h R x y -> P y \/ x = y). +Proof. +move=>H L x y /hrtcP [] // _ [n]; elim: n x y=>[|n IH] x y /=; first by right. +by case=>s [/L X] /IH; case=>[|<-]; left. +Qed. + +(* minimality of m is preserved by htc and hrtc *) + +Lemma htc_min A (h : natmap A) (R : rel nat) m : + (forall x y, R x y -> (x \in 0::dom h) && (y \in 0::dom h)) -> + (forall x, R x m -> x = m) -> + (forall x, htc h R x m -> x = m). +Proof. +move=>H M x /htcP [] // n; elim: n x=>[|n IH] x /=. +- by case=>s [X E]; move: E X=>-> /M. +by case=>s [X] /IH E; move: E X=>-> /M. +Qed. + +Lemma hrtc_min A (h : natmap A) (R : rel nat) m : + (forall x y, R x y -> (x \in 0::dom h) && (y \in 0::dom h)) -> + (forall x, R x m -> x = m) -> + (forall x, hrtc h R x m -> x = m). +Proof. +move=>H M x /hrtcP [] // D [n]; elim: n x D=>[|n IH] x D //= [s][X] /IH. +by case/andP: (H _ _ X)=>_ S /(_ S) E; move: E X=>->; apply: M. +Qed. + +(* when an element is not in range *) + +Lemma htcR A (h : natmap A) (R : rel nat) m : + (forall x y, R x y -> (x \in 0::dom h) && (y \in 0::dom h)) -> + (forall x, ~ R x m) -> + (forall x, ~ htc h R x m). +Proof. +move=>H M x /htcP [] // n; elim: n x=>[|n IH] /= x [s][X]; +by [move=>E; move: E X=>-> /M | move/IH]. +Qed. + +Lemma hrtcR A (h : natmap A) (R : rel nat) m : + (forall x y, R x y -> (x \in 0::dom h) && (y \in 0::dom h)) -> + (forall x, ~ R x m) -> + (forall x, hrtc h R x m -> x = m). +Proof. +move=>H M x /hrtcP [] // _ [n]; elim: n x=>[|n IH] x //=. +by case=>s [X] /IH E; move: E X=>-> /M. +Qed. + +(* there's really not much reason to bother with reflexive closure *) +(* as it's defined simply in terms of transitive closure *) +Lemma hrtcE A (h : natmap A) (R : rel nat) x y : + (forall x y, R x y -> (x \in 0::dom h) && (y \in 0::dom h)) -> + hrtc h R x y = htc h R x y || (x == y) && (x \in 0::dom h). +Proof. +move=>H; apply/(hrtcP H)/idP; last first. +- case/orP; last by case/andP=>/eqP -> Y. + by case/(htcP H)=>n /= [s][X Y]; split; [case/H/andP: X|exists n.+1, s]. +case=>-> [n]; rewrite andbT; elim: n x y=>[|n IH] x y /=. +- by move=>->; rewrite eq_refl orbT. +by case=>s [/(htc1 H) X] /IH /orP [/(htc_trans X) ->|/eqP <-] //; rewrite X. +Qed. diff --git a/pcm/pcm.v b/pcm/pcm.v index aa8f4ab..173f333 100644 --- a/pcm/pcm.v +++ b/pcm/pcm.v @@ -18,10 +18,8 @@ limitations under the License. (******************************************************************************) From Coq Require Import ssreflect ssrbool ssrfun. -From mathcomp Require Import ssrnat bigop. -Set Implicit Arguments. -Unset Strict Implicit. -Unset Printing Implicit Defensive. +From mathcomp Require Import ssrnat eqtype bigop seq. +From fcsl Require Import prelude seqperm pred options. Declare Scope pcm_scope. Delimit Scope pcm_scope with pcm. @@ -60,11 +58,11 @@ Definition clone := fun c & cT -> T & phant_id (pack c) cT => pack c. Definition valid := valid_op class. Definition join := join_op class. -Definition unit := unit_op class. +Definition Unit := unit_op class. End ClassDef. -Arguments unit {cT}. +Arguments Unit {cT}. Module Exports. Coercion sort : type >-> Sortclass. @@ -82,10 +80,11 @@ Notation "[ 'pcm' 'of' T ]" := (@clone T _ _ id id) Notation "x \+ y" := (join x y) (at level 43, left associativity) : pcm_scope. Notation valid := valid. -Notation Unit := unit. +Notation Unit := Unit. -Arguments unit {cT}. -Prenex Implicits valid join unit. +Arguments Unit {cT}. +Arguments valid {cT} !_ / : simpl nomatch. +Prenex Implicits join Unit. (* Restating the laws, with the notation. *) (* Plus some additional derived lemmas. *) @@ -94,10 +93,14 @@ Section Laws. Variable U V : pcm. Lemma joinC (x y : U) : x \+ y = y \+ x. -Proof. by case: U x y=>tp [v j z Cj *]; apply: Cj. Qed. +Proof. +by rewrite /join; case: U x y=>tp [v j z Cj *]; apply: Cj. +Qed. Lemma joinA (x y z : U) : x \+ (y \+ z) = x \+ y \+ z. -Proof. by case: U x y z=>tp [v j z Cj Aj *]; apply: Aj. Qed. +Proof. +by rewrite /join; case: U x y z=>tp [v j z Cj Aj *]; apply: Aj. +Qed. Lemma joinAC (x y z : U) : x \+ y \+ z = x \+ z \+ y. Proof. by rewrite -joinA (joinC y) joinA. Qed. @@ -106,22 +109,29 @@ Lemma joinCA (x y z : U) : x \+ (y \+ z) = y \+ (x \+ z). Proof. by rewrite joinA (joinC x) -joinA. Qed. Lemma validL (x y : U) : valid (x \+ y) -> valid x. -Proof. case: U x y=>tp [v j z Cj Aj Uj /= Mj inv f ?]; apply: Mj. Qed. +Proof. +by rewrite /valid/join; case: U x y=>tp [v j z Cj Aj Uj /= Mj inv f ?]; apply: Mj. +Qed. Lemma validR (x y : U) : valid (x \+ y) -> valid y. Proof. by rewrite joinC; apply: validL. Qed. -Lemma validE (x y : U) : valid (x \+ y) -> valid x * valid y. -Proof. by move=>X; rewrite (validL X) (validR X). Qed. +(* nested pairs are a workaround for https://github.com/coq/coq/issues/8304 *) +Lemma validE2 (x y : U) : valid (x \+ y) -> (valid x * valid y) * (valid (x \+ y) * valid (y \+ x)). +Proof. by move=>X; rewrite (validL X) (validR X) X joinC X. Qed. Lemma unitL (x : U) : Unit \+ x = x. -Proof. by case: U x=>tp [v j z Cj Aj Uj *]; apply: Uj. Qed. +Proof. +by rewrite /Unit/join; case: U x=>tp [v j z Cj Aj Uj *]; apply: Uj. +Qed. Lemma unitR (x : U) : x \+ Unit = x. Proof. by rewrite joinC unitL. Qed. Lemma valid_unit : valid (@Unit U). -Proof. by case: U=>tp [v j z Cj Aj Uj Vm Vu *]. Qed. +Proof. +by rewrite /valid/Unit; case: U=>tp [v j z Cj Aj Uj Vm Vu *]. +Qed. (* some helper lemmas for easier extraction of validity claims *) Lemma validAR (x y z : U) : valid (x \+ y \+ z) -> valid (y \+ z). @@ -130,6 +140,45 @@ Proof. by rewrite -joinA; apply: validR. Qed. Lemma validAL (x y z : U) : valid (x \+ (y \+ z)) -> valid (x \+ y). Proof. by rewrite joinA; apply: validL. Qed. +Lemma validLA (x y z : U) : valid (x \+ y \+ z) -> valid (x \+ (y \+ z)). +Proof. by rewrite joinA. Qed. + +Lemma validRA (x y z : U) : valid (x \+ (y \+ z)) -> valid (x \+ y \+ z). +Proof. by rewrite joinA. Qed. + +(* poor man's automation for a very frequent story of 3 summands *) +(* nested pairs are a workaround for https://github.com/coq/coq/issues/8304 *) +Lemma validLE3 (x y z : U) : valid (x \+ y \+ z) -> + (((valid x * valid y) * (valid z * valid (x \+ y))) * + ((valid (x \+ z) * valid (y \+ x)) * (valid (y \+ z) * valid (z \+ x)))) * + (((valid (z \+ y) * valid (x \+ (y \+ z))) * + (valid (x \+ y \+ z) * valid (x \+ (z \+ y)))) * + ((valid (x \+ z \+ y) * valid (y \+ (x \+ z))) * + (valid (y \+ x \+ z) * valid (y \+ (z \+ x))))) * + (((valid (y \+ z \+ x) * valid (z \+ (x \+ y))) * + (valid (z \+ x \+ y) * valid (z \+ (y \+ x)))) * + valid (z \+ y \+ x)). +Proof. +move=> V3; rewrite !(validE2 V3) joinA V3. +rewrite joinAC in V3; rewrite !(validE2 V3). +rewrite [x \+ z]joinC in V3; rewrite !(validE2 V3). +rewrite joinAC in V3; rewrite !(validE2 V3). +rewrite [z \+ y]joinC in V3; rewrite !(validE2 V3). +by rewrite joinAC in V3; rewrite !(validE2 V3). +Qed. + +Lemma validRE3 (x y z : U) : valid (x \+ (y \+ z)) -> + (((valid x * valid y) * (valid z * valid (x \+ y))) * + ((valid (x \+ z) * valid (y \+ x)) * (valid (y \+ z) * valid (z \+ x)))) * + (((valid (z \+ y) * valid (x \+ (y \+ z))) * + (valid (x \+ y \+ z) * valid (x \+ (z \+ y)))) * + ((valid (x \+ z \+ y) * valid (y \+ (x \+ z))) * + (valid (y \+ x \+ z) * valid (y \+ (z \+ x))))) * + (((valid (y \+ z \+ x) * valid (z \+ (x \+ y))) * + (valid (z \+ x \+ y) * valid (z \+ (y \+ x)))) * + valid (z \+ y \+ x)). +Proof. by rewrite {1}joinA; apply: validLE3. Qed. + End Laws. Hint Resolve valid_unit : core. @@ -138,13 +187,13 @@ Section UnfoldingRules. Variable U : pcm. Lemma pcm_joinE (x y : U) : x \+ y = join_op (class U) x y. -Proof. by []. Qed. +Proof. by rewrite /join. Qed. Lemma pcm_validE (x : U) : valid x = valid_op (class U) x. -Proof. by []. Qed. +Proof. by rewrite /valid. Qed. -Lemma pcm_unitE : unit = unit_op (class U). -Proof. by []. Qed. +Lemma pcm_unitE : Unit = unit_op (class U). +Proof. by rewrite /Unit. Qed. Definition pcmE := (pcm_joinE, pcm_validE, pcm_unitE). @@ -209,7 +258,7 @@ Notation cpcm := type. Notation CPCMMixin := Mixin. Notation CPCM T m := (@pack T _ _ m id). -Notation "[ 'cpcm' 'of' T 'for' cT ]" := (@clone T cT _ id) +Notation "[ 'cpcm' 'of' T 'for' cT ]" := (@clone T cT _ idfun) (at level 0, format "[ 'cpcm' 'of' T 'for' cT ]") : pcm_scope. Notation "[ 'cpcm' 'of' T ]" := (@clone T _ _ id) (at level 0, format "[ 'cpcm' 'of' T ]") : pcm_scope. @@ -262,10 +311,6 @@ Record mixin_of (U : pcm) := Mixin { undef_op : U; unitb_op : U -> bool; _ : forall x, reflect (x = Unit) (unitb_op x); - (* next one doesn't hold of max PCM, so remove eventually *) - _ : forall x y : U, x \+ y = Unit -> x = Unit /\ y = Unit; -(* _ : forall x y z : U, valid (x \+ y \+ z) = - [&& valid (x \+ y), valid (y \+ z) & valid (x \+ z)]; *) _ : ~~ valid undef_op; _ : forall x, undef_op \+ x = undef_op}. @@ -304,14 +349,15 @@ Notation tpcm := type. Notation TPCMMixin := Mixin. Notation TPCM T m := (@pack T _ _ m id). -Notation "[ 'tpcm' 'of' T 'for' cT ]" := (@clone T cT _ id) +Notation "[ 'tpcm' 'of' T 'for' cT ]" := (@clone T cT _ idfun) (at level 0, format "[ 'tpcm' 'of' T 'for' cT ]") : pcm_scope. Notation "[ 'tpcm' 'of' T ]" := (@clone T _ _ id) (at level 0, format "[ 'tpcm' 'of' T ]") : pcm_scope. -Arguments undef {cT}. Notation undef := undef. Notation unitb := unitb. +Arguments undef {cT}. +Prenex Implicits undef. Section Lemmas. Variable U : tpcm. @@ -319,14 +365,11 @@ Variable U : tpcm. Lemma unitbP (x : U) : reflect (x = Unit) (unitb x). Proof. by case: U x=>V [b][u]. Qed. -Lemma unitbE : unitb (Unit : U). +Lemma unitbE (f : U) : f = Unit <-> unitb f. Proof. by case: unitbP. Qed. -Lemma joinE0 (x y : U) : x \+ y = Unit <-> (x = Unit) * (y = Unit). -Proof. -case: U x y=>V [b][u1 u2] H1 H2 H3 T x y; split; first by apply: H2. -by case=>->->; rewrite unitL. -Qed. +Lemma unitb0 : unitb (Unit : U). +Proof. by case: unitbP. Qed. Lemma valid_undefN : ~~ valid (@undef U). Proof. by case: U=>V [b][u]. Qed. @@ -343,7 +386,10 @@ Proof. by rewrite joinC undef_join. Qed. Lemma undef0 : (undef : U) <> (Unit : U). Proof. by move=>E; move: (@valid_unit U); rewrite -E valid_undef. Qed. -Definition undefE := (undef_join, join_undef, valid_undef). +Lemma unitb_undef : unitb (undef : U) = false. +Proof. by case: unitbP =>// /undef0. Qed. + +Definition tpcmE := (undef_join, join_undef, valid_undef, unitb0, unitb_undef). End Lemmas. End Exports. @@ -352,6 +398,60 @@ End TPCM. Export TPCM.Exports. +(********************************) +(* PCMs with decidable equality *) +(********************************) + +Module EQPCM. +Section ClassDef. + +Record mixin_of (U : eqType) := Mixin { + base : PCM.mixin_of U}. + +Record class_of (U : Type) := Class { + eqbase : Equality.mixin_of U; + mixin : mixin_of (EqType U eqbase)}. + +Local Coercion eqbase: class_of >-> Equality.mixin_of. + +Structure type : Type := Pack {sort : Type; _ : class_of sort}. +Local Coercion sort : type >-> Sortclass. + +Variables (T : Type) (cT : type). +Definition class := let: Pack _ c as cT' := cT return class_of cT' in c. +Definition clone c of phant_id class c := @Pack T c. + +(* produce an eqpcm out of the mixin *) +(* equalize m0 and m by means of a phantom *) +Definition pack (eb0 : Equality.mixin_of T) (m0 : mixin_of (EqType T eb0)) := + fun m & phant_id m0 m => Pack (@Class T eb0 m). + +Definition pcm := PCM.Pack (base (mixin class)). +Definition eqtype := Eval hnf in EqType cT class. + +End ClassDef. + +Module Exports. +Coercion eqbase : class_of >-> Equality.mixin_of. +Coercion sort : type >-> Sortclass. +Coercion pcm : type >-> PCM.type. +Coercion eqtype : type >-> eqType. +Canonical Structure pcm. +Canonical Structure eqtype. +Notation eqpcm := type. +Notation EQPCM T m := (@pack T _ _ (Mixin m) id). + +Notation "[ 'eqpcm' 'of' T 'for' cT ]" := (@clone T cT _ idfun) + (at level 0, format "[ 'eqpcm' 'of' T 'for' cT ]") : pcm_scope. +Notation "[ 'eqpcm' 'of' T ]" := (@clone T _ _ id) + (at level 0, format "[ 'eqpcm' 'of' T ]") : pcm_scope. + +End Exports. + +End EQPCM. + +Export EQPCM.Exports. + (***************************************) (* Support for big operators over PCMs *) @@ -363,9 +463,8 @@ Canonical pcm_comoid (U : pcm) := Monoid.ComLaw (@joinC U). Section BigPartialMorph. Variables (R1 : Type) (R2 : pcm) (K : R1 -> R2 -> Type) (f : R2 -> R1). Variables (id1 : R1) (op1 : R1 -> R1 -> R1). -Hypotheses - (f_op : forall x y : R2, valid (x \+ y) -> f (x \+ y) = op1 (f x) (f y)) - (f_id : f Unit = id1). +Hypotheses (f_op : forall x y : R2, valid (x \+ y) -> f (x \+ y) = op1 (f x) (f y)) + (f_id : f Unit = id1). Lemma big_pmorph I r (P : pred I) F : valid (\big[PCM.join/Unit]_(i <- r | P i) F i) -> @@ -388,21 +487,62 @@ End BigPartialMorph. Definition natPCMMix := PCMMixin addnC addnA add0n (fun x y => @id true) (erefl _). Canonical natPCM := Eval hnf in PCM nat natPCMMix. +Canonical natEQPCM := Eval hnf in EQPCM nat natPCMMix. (* also with multiplication, but we don't make that one canonical *) Definition multPCMMix := PCMMixin mulnC mulnA mul1n (fun x y => @id true) (erefl _). Definition multPCM := Eval hnf in PCM nat multPCMMix. +Definition multEQPCM := Eval hnf in EQPCM nat multPCMMix. (* with max too *) Definition maxPCMMix := PCMMixin maxnC maxnA max0n (fun x y => @id true) (erefl _). Definition maxPCM := Eval hnf in PCM nat maxPCMMix. +Definition maxEQPCM := Eval hnf in EQPCM nat maxPCMMix. (* bools with disjunction are a pcm *) Definition bool_orPCMMix := PCMMixin orbC orbA orFb (fun x y => @id true) (erefl _). Definition bool_orPCM := Eval hnf in PCM bool bool_orPCMMix. +Definition bool_orEQPCM := Eval hnf in EQPCM bool bool_orPCMMix. + +(* positive natural numbers under max are a pcm *) +Section PosNatMax. + +Definition posNat := sig (fun x => x > 0). + +Lemma max_pos (x y : posNat) : maxn (val x) (val y) > 0. +Proof. by case: x y=>x pfx [y pfy]; rewrite leq_max pfx pfy. Qed. + +Definition pos_valid (x : posNat) := true. +Definition pos_join x y : posNat := + Sub (maxn (val x) (val y)) (max_pos x y). +Definition pos_unit : posNat := Sub 1 (leqnn 1). + +Lemma pos_joinC x y : pos_join x y = pos_join y x. +Proof. by apply/eqP; rewrite /pos_join -val_eqE /= maxnC. Qed. + +Lemma pos_joinA x y z : pos_join x (pos_join y z) = pos_join (pos_join x y) z. +Proof. by apply/eqP; rewrite /pos_join -val_eqE /= maxnA. Qed. + +Lemma pos_validL x y : pos_valid (pos_join x y) -> pos_valid x. +Proof. by []. Qed. + +Lemma pos_unitL x : pos_join pos_unit x = x. +Proof. +apply/eqP; rewrite /pos_join -val_eqE. +by apply/eqP/maxn_idPr; case: x. +Qed. + +Lemma pos_validU : pos_valid pos_unit. Proof. by []. Qed. + +Definition posnatmaxPCMMix := + PCMMixin pos_joinC pos_joinA pos_unitL + pos_validL pos_validU. +Canonical posnatmaxPCM := Eval hnf in PCM posNat posnatmaxPCMMix. + +End PosNatMax. (* cartesian product of pcms is a pcm *) @@ -416,28 +556,19 @@ Definition pjoin := [fun x1 x2 : tp => (x1.1 \+ x2.1, x1.2 \+ x2.2)]. Definition punit : tp := (Unit, Unit). Lemma joinC x y : pjoin x y = pjoin y x. -Proof. -move: x y => [x1 x2][y1 y2] /=. -by rewrite (joinC x1) (joinC x2). -Qed. +Proof. by rewrite /= (joinC x.1) (joinC x.2). Qed. Lemma joinA x y z : pjoin x (pjoin y z) = pjoin (pjoin x y) z. -Proof. -move: x y z => [x1 x2][y1 y2][z1 z2] /=. -by rewrite (joinA x1) (joinA x2). -Qed. +Proof. by rewrite /= !joinA. Qed. Lemma validL x y : pvalid (pjoin x y) -> pvalid x. -Proof. -move: x y => [x1 x2][y1 y2] /=. -by case/andP=>D1 D2; rewrite (validL D1) (validL D2). -Qed. +Proof. by case/andP => /= /validL-> /validL->. Qed. Lemma unitL x : pjoin punit x = x. -Proof. by case: x=>x1 x2; rewrite /= !unitL. Qed. +Proof. by rewrite /= !unitL -prod_eta. Qed. Lemma validU : pvalid punit. -Proof. by rewrite /pvalid /= !valid_unit. Qed. +Proof. by rewrite /= !valid_unit. Qed. End ProdPCM. End ProdPCM. @@ -447,6 +578,8 @@ Definition prodPCMMixin U V := (@ProdPCM.unitL U V) (@ProdPCM.validL U V) (@ProdPCM.validU U V). Canonical prodPCM U V := Eval hnf in PCM (_ * _) (@prodPCMMixin U V). +Canonical prodEQPCM (U V : eqpcm) := + Eval hnf in EQPCM (_ * _) (prodPCMMixin U V). (* product simplification *) @@ -455,22 +588,25 @@ Variable U V : pcm. Lemma pcmPJ (x1 y1 : U) (x2 y2 : V) : (x1, x2) \+ (y1, y2) = (x1 \+ y1, x2 \+ y2). -Proof. by []. Qed. +Proof. by rewrite pcmE. Qed. -Lemma pcm_peta (x : prodPCM U V) : x = (x.1, x.2). -Proof. by case: x. Qed. +Lemma pcmFJ (x y : U * V) : (x \+ y).1 = x.1 \+ y.1. +Proof. by rewrite pcmE. Qed. + +Lemma pcmSJ (x y : U * V) : (x \+ y).2 = x.2 \+ y.2. +Proof. by rewrite pcmE. Qed. Lemma pcmPV (x : prodPCM U V) : valid x = valid x.1 && valid x.2. +Proof. by []. Qed. + +Lemma pcmPU : Unit = (Unit, Unit) :> prodPCM U V. Proof. by rewrite pcmE. Qed. -Definition pcmPE := (pcmPJ, pcmPV). +Definition pcmPE := (pcmPJ, pcmFJ, pcmSJ, pcmPV, pcmPU). End Simplification. (* product of TPCMs is a TPCM *) -(* With TPCMs we could really remove the pairs *) -(* where one element is valid and the other is not *) -(* But let's not bother now *) Section ProdTPCM. Variables (U V : tpcm). @@ -479,35 +615,23 @@ Lemma prod_unitb (x : prodPCM U V) : reflect (x = Unit) (unitb x.1 && unitb x.2). Proof. case: x=>x1 x2; case: andP=>/= H; constructor. -- by case: H=>/unitbP -> /unitbP ->. -by case=>X1 X2; elim: H; rewrite X1 X2 !unitbE. +- by case: H=>/unitbP -> /unitbP ->; rewrite pcmPE. +by rewrite pcmPE; case=>X1 X2; elim: H; rewrite X1 X2 !tpcmE. Qed. -Lemma prod_join0E (x y : prodPCM U V) : x \+ y = Unit -> x = Unit /\ y = Unit. -Proof. by case: x y=>x1 x2 [y1 y2][] /joinE0 [->->] /joinE0 [->->]. Qed. - -(* -Lemma prod_valid3 (x y z : prodPCM U V) : valid (x \+ y \+ z) = - [&& valid (x \+ y), valid (y \+ z) & valid (x \+ z)]. -Proof. -case: x y z=>x1 x2 [y1 y2][z1 z2]; rewrite pcmE /= !valid3 -!andbA. -by do !bool_congr. -Qed. -*) - Lemma prod_valid_undef : ~~ valid (@undef U, @undef V). -Proof. by rewrite pcmPV negb_and !valid_undef. Qed. +Proof. by rewrite /= !valid_undef. Qed. Lemma prod_undef_join x : (@undef U, @undef V) \+ x = (@undef U, @undef V). -Proof. by case: x=>x1 x2; rewrite pcmPE !undef_join. Qed. +Proof. by rewrite [x]prod_eta /= pcmPJ !undef_join. Qed. -Definition prodTPCMMix := TPCMMixin prod_unitb prod_join0E - prod_valid_undef prod_undef_join. +Definition prodTPCMMix := TPCMMixin prod_unitb prod_valid_undef prod_undef_join. Canonical prodTPCM := Eval hnf in TPCM (U * V) prodTPCMMix. End ProdTPCM. -(* unit is a pcm; just for kicks *) +(* unit is a PCM, but not a TPCM, as there is no undefined element. *) +(* We have to lift for that *) Module UnitPCM. Section UnitPCM. @@ -538,21 +662,14 @@ Definition unitPCMMixin := PCMMixin UnitPCM.ujoinC UnitPCM.ujoinA UnitPCM.uunitL UnitPCM.uvalidL UnitPCM.uvalidU. Canonical unitPCM := Eval hnf in PCM unit unitPCMMixin. - -(* but it's not a TPCM, as there is no undefined element *) -(* we have to lift for that *) +Canonical unitEQPCM := Eval hnf in EQPCM unit unitPCMMixin. (* bools with conjunction are a pcm *) - -Module BoolConjPCM. -Lemma unitL x : true && x = x. Proof. by []. Qed. -End BoolConjPCM. - -Definition boolPCMMixin := PCMMixin andbC andbA BoolConjPCM.unitL +Definition boolPCMMixin := PCMMixin andbC andbA andTb (fun x y => @id true) (erefl _). Canonical boolConjPCM := Eval hnf in PCM bool boolPCMMixin. - +Canonical boolConjEQPCM := Eval hnf in EQPCM bool boolPCMMixin. (*************************) (* PCM-induced pre-order *) @@ -576,7 +693,7 @@ Proof. by exists x; rewrite unitL. Qed. (* We don't have antisymmetry in general, though for common PCMs *) (* e.g., union maps, we do have it; but it is proved separately for them *) -Lemma pleq_refl x : [pcm x <= x]. +Lemma pleq_refl {x} : [pcm x <= x]. Proof. by exists Unit; rewrite unitR. Qed. Lemma pleq_trans x y z : [pcm x <= y] -> [pcm y <= z] -> [pcm x <= z]. @@ -603,7 +720,7 @@ Proof. by case=>x -> /validL. Qed. Lemma pleq_validL (x x1 x2 : U) : [pcm x1 <= x2] -> valid (x \+ x2) -> valid (x \+ x1). -Proof. by case=>a ->; rewrite joinA; apply: validL. Qed. +Proof. by case=>a -> V; rewrite (validRE3 V). Qed. Lemma pleq_validR (x x1 x2 : U) : [pcm x1 <= x2] -> valid (x2 \+ x) -> valid (x1 \+ x). @@ -622,6 +739,7 @@ Proof. by rewrite !(joinC x); apply: pleq_joinxK. Qed. End PleqLemmas. Hint Resolve pleq_unit pleq_refl pleq_joinr pleq_joinl : core. +Prenex Implicits pleq_refl. (*******************) (* Local functions *) @@ -638,3 +756,364 @@ Proof. by move=>H V E; move: (H Unit x y r); rewrite unitL !unitR; case. Qed. Lemma idL (U : pcm) : @local U (fun x y => Some x). Proof. by move=>p x y _ V [<-]; rewrite -joinA. Qed. +(******************) +(* Tuples of PCMs *) +(******************) + +(* Often times we need to iterate PCM-product construction, *) +(* but then, due to the Coq policy of reducing destructors *) +(* the iterated projections are not simplified. Thus, it *) +(* makes (some) sense to have primitive product iterations *) +(* for the few common numbers *) + +(* Also, products are the construction we use by default *) +(* when pairing resources. It makes sense to use different pairing *) +(* constructions for composing resources vs. when we need a tuple *) +(* PCM in one resource. That way, the lemmas and rewrite rules *) +(* that are used to change the views from a larger resource to a smaller one *) +(* can be prevented from descending and unfolding too much. *) +(* Truth be told, the control of this unfolding should really be done *) +(* by using Opaque and Transparent, if they only worked properly in Coq *) + +Inductive Prod3 (U1 U2 U3 : Type) := mk3 {pcm31 : U1; pcm32 : U2; pcm33 : U3}. + +Prenex Implicits pcm31 pcm32 pcm33. + +Section UPCM3. +Variables U1 U2 U3 : pcm. +Notation tp := (Prod3 U1 U2 U3). + +Let uvalid (x : tp) := [&& valid (pcm31 x), valid (pcm32 x) & valid (pcm33 x)]. +Let ujoin (x y : tp) := mk3 (pcm31 x \+ pcm31 y) (pcm32 x \+ pcm32 y) (pcm33 x \+ pcm33 y). +Let uunit : tp := mk3 Unit Unit Unit. + +Lemma ujoinC x y : ujoin x y = ujoin y x. +Proof. by congr mk3; rewrite joinC. Qed. + +Lemma ujoinA x y z : ujoin x (ujoin y z) = ujoin (ujoin x y) z. +Proof. by congr mk3; rewrite joinA. Qed. + +Lemma uvalidL x y : uvalid (ujoin x y) -> uvalid x. +Proof. by rewrite /uvalid; case/and3P; do ![move/validL=>->]. Qed. + +Lemma uunitL x : ujoin uunit x = x. +Proof. by case: x=>*; rewrite /uunit /ujoin /= !unitL. Qed. + +Lemma uvalidU : uvalid uunit. +Proof. by rewrite /uvalid !valid_unit. Qed. + +Definition prod3PCMMixin := PCMMixin ujoinC ujoinA uunitL uvalidL uvalidU. +Canonical prod3_PCM := Eval hnf in PCM (Prod3 U1 U2 U3) prod3PCMMixin. +End UPCM3. + +Section UTPCM3. +Variables U1 U2 U3 : tpcm. +Notation tp := (Prod3 U1 U2 U3). + +Let uunitb (x : tp) := [&& unitb (pcm31 x), unitb (pcm32 x) & unitb (pcm33 x)]. +Let uundef : tp := mk3 undef undef undef. + +Lemma uunitbP x : reflect (x = Unit) (uunitb x). +Proof. +rewrite /uunitb !pcmE; case: x=>x1 x2 x3 /=. +do 3![case: unitbP=>[->|H]; last by constructor; case]. +by constructor. +Qed. + +Lemma uvalid_undef : ~~ valid uundef. +Proof. by rewrite pcmE /valid /= !valid_undef. Qed. + +Lemma uundef_join x : uundef \+ x = uundef. +Proof. by rewrite pcmE /= !undef_join. Qed. + +Definition prod3TPCMMix := TPCMMixin uunitbP uvalid_undef uundef_join. +Canonical prod3_TPCM := Eval hnf in TPCM (Prod3 U1 U2 U3) prod3TPCMMix. +End UTPCM3. + +(* And for 4 *) + +Inductive Prod4 (U1 U2 U3 U4 : Type) := mk4 {pcm41 : U1; pcm42 : U2; pcm43 : U3; pcm44 : U4}. + +Prenex Implicits pcm41 pcm42 pcm43 pcm44. + +Section UPCM4. +Variables U1 U2 U3 U4 : pcm. +Notation tp := (Prod4 U1 U2 U3 U4). + +Let uvalid (x : tp) := [&& valid (pcm41 x), valid (pcm42 x), + valid (pcm43 x) & valid (pcm44 x)]. +Let ujoin (x y : tp) := mk4 (pcm41 x \+ pcm41 y) (pcm42 x \+ pcm42 y) + (pcm43 x \+ pcm43 y) (pcm44 x \+ pcm44 y). +Let uunit : tp := mk4 Unit Unit Unit Unit. + +Lemma ujoinC4 x y : ujoin x y = ujoin y x. +Proof. by congr mk4; rewrite joinC. Qed. + +Lemma ujoinA4 x y z : ujoin x (ujoin y z) = ujoin (ujoin x y) z. +Proof. by congr mk4; rewrite joinA. Qed. + +Lemma uvalidL4 x y : uvalid (ujoin x y) -> uvalid x. +Proof. by rewrite /uvalid; case/and4P; do ![move/validL=>->]. Qed. + +Lemma uunitL4 x : ujoin uunit x = x. +Proof. by case: x=>*; rewrite /uunit /ujoin /= !unitL. Qed. + +Lemma uvalidU4 : uvalid uunit. +Proof. by rewrite /uvalid !valid_unit. Qed. + +Definition prod4PCMMixin := PCMMixin ujoinC4 ujoinA4 uunitL4 uvalidL4 uvalidU4. +Canonical prod4_PCM := Eval hnf in PCM (Prod4 U1 U2 U3 U4) prod4PCMMixin. +End UPCM4. + +Section UTPCM4. +Variables U1 U2 U3 U4 : tpcm. +Notation tp := (Prod4 U1 U2 U3 U4). + +Let uunitb (x : tp) := [&& unitb (pcm41 x), unitb (pcm42 x), + unitb (pcm43 x) & unitb (pcm44 x)]. +Let uundef : tp := mk4 undef undef undef undef. + +Lemma uunitbP4 x : reflect (x = Unit) (uunitb x). +Proof. +rewrite /uunitb !pcmE; case: x=>x1 x2 x3 x4 /=. +do 4![case: unitbP=>[->|H] /=; last by constructor; case]. +by constructor. +Qed. + +Lemma uvalid_undef4 : ~~ valid uundef. +Proof. by rewrite pcmE /valid /= !valid_undef. Qed. + +Lemma uundef_join4 x : uundef \+ x = uundef. +Proof. by rewrite pcmE /= !undef_join. Qed. + +Definition prod4TPCMMix := TPCMMixin uunitbP4 uvalid_undef4 uundef_join4. +Canonical prod4_TPCM := Eval hnf in TPCM (Prod4 U1 U2 U3 U4) prod4TPCMMix. +End UTPCM4. + +(* And for 5 *) + +Inductive Prod5 (U1 U2 U3 U4 U5 : Type) := + mk5 {pcm51 : U1; pcm52 : U2; pcm53 : U3; pcm54 : U4; pcm55 : U5}. + +Prenex Implicits pcm51 pcm52 pcm53 pcm54 pcm55. + +Section UPCM5. +Variables U1 U2 U3 U4 U5 : pcm. +Notation tp := (Prod5 U1 U2 U3 U4 U5). + +Let uvalid (x : tp) := [&& valid (pcm51 x), valid (pcm52 x), + valid (pcm53 x), valid (pcm54 x) & valid (pcm55 x)]. +Let ujoin (x y : tp) := mk5 (pcm51 x \+ pcm51 y) (pcm52 x \+ pcm52 y) + (pcm53 x \+ pcm53 y) (pcm54 x \+ pcm54 y) + (pcm55 x \+ pcm55 y). +Let uunit : tp := mk5 Unit Unit Unit Unit Unit. + +Lemma ujoinC5 x y : ujoin x y = ujoin y x. +Proof. by congr mk5; rewrite joinC. Qed. + +Lemma ujoinA5 x y z : ujoin x (ujoin y z) = ujoin (ujoin x y) z. +Proof. by congr mk5; rewrite joinA. Qed. + +Lemma uvalidL5 x y : uvalid (ujoin x y) -> uvalid x. +Proof. by rewrite /uvalid; case/and5P; do ![move/validL=>->]. Qed. + +Lemma uunitL5 x : ujoin uunit x = x. +Proof. by case: x=>*; rewrite /uunit /ujoin /= !unitL. Qed. + +Lemma uvalidU5 : uvalid uunit. +Proof. by rewrite /uvalid !valid_unit. Qed. + +Definition prod5PCMMixin := PCMMixin ujoinC5 ujoinA5 uunitL5 uvalidL5 uvalidU5. +Canonical prod5_PCM := Eval hnf in PCM (Prod5 U1 U2 U3 U4 U5) prod5PCMMixin. +End UPCM5. + +Section UTPCM5. +Variables U1 U2 U3 U4 U5 : tpcm. +Notation tp := (Prod5 U1 U2 U3 U4 U5). + +Let uunitb (x : tp) := [&& unitb (pcm51 x), unitb (pcm52 x), + unitb (pcm53 x), unitb (pcm54 x) & unitb (pcm55 x)]. +Let uundef : tp := mk5 undef undef undef undef undef. + +Lemma uunitbP5 x : reflect (x = Unit) (uunitb x). +Proof. +rewrite /uunitb !pcmE; case: x=>x1 x2 x3 x4 x5 /=. +do ![case: unitbP=>[->|H] /=; last by constructor; case]. +by constructor. +Qed. + +Lemma uvalid_undef5 : ~~ valid uundef. +Proof. by rewrite pcmE /valid /= !valid_undef. Qed. + +Lemma uundef_join5 x : uundef \+ x = uundef. +Proof. by rewrite pcmE /= !undef_join. Qed. + +Definition prod5TPCMMix := TPCMMixin uunitbP5 uvalid_undef5 uundef_join5. +Canonical prod5_TPCM := Eval hnf in TPCM (Prod5 U1 U2 U3 U4 U5) prod5TPCMMix. +End UTPCM5. + +(* And for 6 *) + +Inductive Prod6 (U1 U2 U3 U4 U5 U6 : Type) := + mk6 {pcm61 : U1; pcm62 : U2; pcm63 : U3; + pcm64 : U4; pcm65 : U5; pcm66 : U6}. + +Prenex Implicits pcm61 pcm62 pcm63 pcm64 pcm65 pcm66. + +Section UPCM6. +Variables U1 U2 U3 U4 U5 U6 : pcm. +Notation tp := (Prod6 U1 U2 U3 U4 U5 U6). + +Let uvalid (x : tp) := [&& valid (pcm61 x), valid (pcm62 x), valid (pcm63 x), + valid (pcm64 x), valid (pcm65 x) & valid (pcm66 x)]. +Let ujoin (x y : tp) := mk6 (pcm61 x \+ pcm61 y) (pcm62 x \+ pcm62 y) + (pcm63 x \+ pcm63 y) (pcm64 x \+ pcm64 y) + (pcm65 x \+ pcm65 y) (pcm66 x \+ pcm66 y). +Let uunit : tp := mk6 Unit Unit Unit Unit Unit Unit. + +Lemma ujoinC6 x y : ujoin x y = ujoin y x. +Proof. by congr mk6; rewrite joinC. Qed. + +Lemma ujoinA6 x y z : ujoin x (ujoin y z) = ujoin (ujoin x y) z. +Proof. by congr mk6; rewrite joinA. Qed. + +Lemma uvalidL6 x y : uvalid (ujoin x y) -> uvalid x. +Proof. by rewrite /uvalid; case/and6P; do ![move/validL=>->]. Qed. + +Lemma uunitL6 x : ujoin uunit x = x. +Proof. by case: x=>*; rewrite /uunit /ujoin /= !unitL. Qed. + +Lemma uvalidU6 : uvalid uunit. +Proof. by rewrite /uvalid !valid_unit. Qed. + +Definition prod6PCMMixin := PCMMixin ujoinC6 ujoinA6 uunitL6 uvalidL6 uvalidU6. +Canonical prod6_PCM := Eval hnf in PCM (Prod6 U1 U2 U3 U4 U5 U6) prod6PCMMixin. +End UPCM6. + +Section UTPCM6. +Variables U1 U2 U3 U4 U5 U6 : tpcm. +Notation tp := (Prod6 U1 U2 U3 U4 U5 U6). + +Let uunitb (x : tp) := [&& unitb (pcm61 x), unitb (pcm62 x), unitb (pcm63 x), + unitb (pcm64 x), unitb (pcm65 x) & unitb (pcm66 x)]. +Let uundef : tp := mk6 undef undef undef undef undef undef. + +Lemma uunitbP6 x : reflect (x = Unit) (uunitb x). +Proof. +rewrite /uunitb !pcmE; case: x=>x1 x2 x3 x4 x5 x6 /=. +do ![case: unitbP=>[->|H] /=; last by constructor; case]. +by constructor. +Qed. + +Lemma uvalid_undef6 : ~~ valid uundef. +Proof. by rewrite pcmE /valid /= !valid_undef. Qed. + +Lemma uundef_join6 x : uundef \+ x = uundef. +Proof. by rewrite pcmE /= !undef_join. Qed. + +Definition prod6TPCMMix := TPCMMixin uunitbP6 uvalid_undef6 uundef_join6. +Canonical prod6_TPCM := Eval hnf in TPCM (Prod6 U1 U2 U3 U4 U5 U6) prod6TPCMMix. +End UTPCM6. + +(* And for 7 *) + +Inductive Prod7 (U1 U2 U3 U4 U5 U6 U7 : Type) := + mk7 {pcm71 : U1; pcm72 : U2; pcm73 : U3; + pcm74 : U4; pcm75 : U5; pcm76 : U6; pcm77 : U7}. + +Prenex Implicits pcm71 pcm72 pcm73 pcm74 pcm75 pcm76 pcm77. + +Section UPCM7. +Variables U1 U2 U3 U4 U5 U6 U7 : pcm. +Notation tp := (Prod7 U1 U2 U3 U4 U5 U6 U7). + +Let uvalid (x : tp) := [&& valid (pcm71 x), valid (pcm72 x), + valid (pcm73 x), valid (pcm74 x), + valid (pcm75 x), valid (pcm76 x) & + valid (pcm77 x)]. +Let ujoin (x y : tp) := mk7 (pcm71 x \+ pcm71 y) (pcm72 x \+ pcm72 y) + (pcm73 x \+ pcm73 y) (pcm74 x \+ pcm74 y) + (pcm75 x \+ pcm75 y) (pcm76 x \+ pcm76 y) + (pcm77 x \+ pcm77 y). +Let uunit : tp := mk7 Unit Unit Unit Unit Unit Unit Unit. + +Lemma ujoinC7 x y : ujoin x y = ujoin y x. +Proof. by congr mk7; rewrite joinC. Qed. + +Lemma ujoinA7 x y z : ujoin x (ujoin y z) = ujoin (ujoin x y) z. +Proof. by congr mk7; rewrite joinA. Qed. + +Lemma uvalidL7 x y : uvalid (ujoin x y) -> uvalid x. +Proof. by rewrite /uvalid; case/and7P; do ![move/validL=>->]. Qed. + +Lemma uunitL7 x : ujoin uunit x = x. +Proof. by case: x=>*; rewrite /uunit /ujoin /= !unitL. Qed. + +Lemma uvalidU7 : uvalid uunit. +Proof. by rewrite /uvalid !valid_unit. Qed. + +Definition prod7PCMMixin := PCMMixin ujoinC7 ujoinA7 uunitL7 uvalidL7 uvalidU7. +Canonical prod7_PCM := Eval hnf in PCM (Prod7 U1 U2 U3 U4 U5 U6 U7) prod7PCMMixin. +End UPCM7. + +Section UTPCM7. +Variables U1 U2 U3 U4 U5 U6 U7 : tpcm. +Notation tp := (Prod7 U1 U2 U3 U4 U5 U6 U7). + +Let uunitb (x : tp) := [&& unitb (pcm71 x), unitb (pcm72 x), unitb (pcm73 x), + unitb (pcm74 x), unitb (pcm75 x), unitb (pcm76 x) & + unitb (pcm77 x)]. +Let uundef : tp := mk7 undef undef undef undef undef undef undef. + +Lemma uunitbP7 x : reflect (x = Unit) (uunitb x). +Proof. +rewrite /uunitb !pcmE; case: x=>x1 x2 x3 x4 x5 x6 x7 /=. +do ![case: unitbP=>[->|H] /=; last by constructor; case]. +by constructor. +Qed. + +Lemma uvalid_undef7 : ~~ valid uundef. +Proof. by rewrite pcmE /valid /= !valid_undef. Qed. + +Lemma uundef_join7 x : uundef \+ x = uundef. +Proof. by rewrite pcmE /= !undef_join. Qed. + +Definition prod7TPCMMix := TPCMMixin uunitbP7 uvalid_undef7 uundef_join7. +Canonical prod7_TPCM := Eval hnf in TPCM (Prod7 U1 U2 U3 U4 U5 U6 U7) prod7TPCMMix. +End UTPCM7. + + +(********************) +(* PCMs and folding *) +(********************) + +(* folding functions that are morphisms in the first argument *) + +Section PCMfold. +Variables (A : Type) (R : pcm) (a : R -> A -> R). +Hypothesis H : (forall x y k, a (x \+ y) k = a x k \+ y). + +(* first a helper lemma *) +Lemma foldl_helper (s1 s2 : seq A) (z0 : R) x : + foldl a z0 (s1 ++ x :: s2) = foldl a z0 (x :: s1 ++ s2). +Proof. +elim: s1 s2 z0=>[|k ks IH] s2' z0 //=. +rewrite IH /=; congr foldl. +rewrite -[a z0 k]unitL H -[z0]unitL !H. +by rewrite -{2}[a Unit x]unitL H joinCA joinA. +Qed. + +Lemma foldl_perm (s1 s2 : seq A) (z0 : R) : + perm s1 s2 -> foldl a z0 s1 = foldl a z0 s2. +Proof. +move=>P; elim: s1 s2 z0 P=>[|k ks IH] s2 z0 P; first by move/pperm_nil: P=>->. +have K: k \In s2 by apply: pperm_in P _; rewrite InE; left. +case: {K} (In_split K) P=>x [y] ->{s2} /= /pperm_cons_cat_cons P. +by rewrite foldl_helper //=; apply: IH. +Qed. + +Lemma foldl_init (s : seq A) (x y : R) : + foldl a (x \+ y) s = foldl a x s \+ y. +Proof. by elim: s x y=>[|k s IH] x y //=; rewrite H IH. Qed. + +End PCMfold. diff --git a/pcm/pred.v b/pcm/pred.v deleted file mode 100644 index 77e37b8..0000000 --- a/pcm/pred.v +++ /dev/null @@ -1,563 +0,0 @@ -(* -Copyright 2010 IMDEA Software Institute -Licensed under the Apache License, Version 2.0 (the "License"); -you may not use this file except in compliance with the License. -You may obtain a copy of the License at - http://www.apache.org/licenses/LICENSE-2.0 -Unless required by applicable law or agreed to in writing, software -distributed under the License is distributed on an "AS IS" BASIS, -WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. -See the License for the specific language governing permissions and -limitations under the License. -*) - -From Coq Require Import ssreflect ssrbool ssrfun Setoid. -From mathcomp Require Import seq. -Set Implicit Arguments. -Unset Strict Implicit. -Unset Printing Implicit Defensive. -Set Warnings "-projection-no-head-constant". - -(******************************************************************************) -(* This file re-implements some of ssrbool's entities in Prop universe *) -(******************************************************************************) - -Lemma andTp p : True /\ p <-> p. Proof. by intuition. Qed. -Lemma andpT p : p /\ True <-> p. Proof. by intuition. Qed. -Lemma andFp p : False /\ p <-> False. Proof. by intuition. Qed. -Lemma andpF p : p /\ False <-> False. Proof. by intuition. Qed. -Lemma orTp p : True \/ p <-> True. Proof. by intuition. Qed. -Lemma orpT p : p \/ True <-> True. Proof. by intuition. Qed. -Lemma orFp p : False \/ p <-> p. Proof. by intuition. Qed. -Lemma orpF p : p \/ False <-> p. Proof. by intuition. Qed. - -Declare Scope rel_scope. -Delimit Scope rel_scope with rel. -Open Scope rel_scope. - -(**************************************************************************) -(* We follow ssrbool, and provide four different types of predicates. *) -(* *) -(* (1) Pred is the type of propositional functions *) -(* (2) Simpl_Pred is the type of predicates that automatically simplify *) -(* when used in an applicative position. *) -(* (3) Mem_Pred is for predicates that support infix notation x \In P *) -(* (4) PredType is the structure for interpreting various types, such as *) -(* lists, tuples, etc. as predicates. *) -(* *) -(* Important point is that custom lemmas over predicates can be stated in *) -(* terms of Pred, while Simpl_Pred, Mem_Pred and PredType are for *) -(* technical developments used in this file only. More on this point *) -(* can be found in ssrbool. *) -(**************************************************************************) - -Definition Pred T := T -> Prop. -Identity Coercion fun_of_Pred : Pred >-> Funclass. - -Notation xPred0 := (fun _ => False). -Notation xPred1 := (fun x y => x = y). -Notation xPredT := (fun _ => True). -Notation xPredI := (fun (p1 p2 : Pred _) x => p1 x /\ p2 x). -Notation xPredU := (fun (p1 p2 : Pred _) x => p1 x \/ p2 x). -Notation xPredC := (fun (p : Pred _) x => ~ p x). -Notation xPredD := (fun (p1 p2 : Pred _) x => ~ p2 x /\ p1 x). -Notation xPreim := (fun f (p : Pred _) x => p (f x)). - -Section Predicates. -Variable T : Type. - -(* simple predicates *) - -Definition Simpl_Pred := simpl_fun T Prop. -Definition SimplPred (p : Pred T) : Simpl_Pred := SimplFun p. -Coercion Pred_of_Simpl (p : Simpl_Pred) : Pred T := p : T -> Prop. - -(* it's useful to declare the operations as simple predicates, so that *) -(* complex expressions automatically reduce when used in applicative *) -(* positions *) - -Definition Pred0 := SimplPred xPred0. -Definition Pred1 x := SimplPred (xPred1 x). -Definition PredT := SimplPred xPredT. -Definition PredI p1 p2 := SimplPred (xPredI p1 p2). -Definition PredU p1 p2 := SimplPred (xPredU p1 p2). -Definition PredC p := SimplPred (xPredC p). -Definition PredD p1 p2 := SimplPred (xPredD p1 p2). -Definition Preim rT f (d : Pred rT) := SimplPred (xPreim f d). - -(* membership predicates *) - -Variant Mem_Pred : Type := MemProp of Pred T. -Definition isMem pT toPred mem := mem = (fun p : pT => MemProp [eta toPred p]). - -(* the general structure for predicates *) - -Structure PredType : Type := PropPredType { - Pred_Sort :> Type; - toPred : Pred_Sort -> Pred T; - _ : {mem | isMem toPred mem}}. - -Definition mkPredType pT toP := - PropPredType (exist (@isMem pT toP) _ (erefl _)). - -(* Pred, SimplPred, Mem_Pred, pred and simpl_pred are PredType's *) -Canonical Structure PredPredType := Eval hnf in @mkPredType (Pred T) id. -Canonical Structure SimplPredPredType := Eval hnf in mkPredType Pred_of_Simpl. -Coercion Pred_of_Mem mp : Pred_Sort PredPredType := - let: MemProp p := mp in [eta p]. -Canonical Structure MemPredType := Eval hnf in mkPredType Pred_of_Mem. -Canonical Structure predPredType := Eval hnf in @mkPredType (pred T) id. -Canonical Structure simplpredPredType := - Eval hnf in @mkPredType (simpl_pred T) (fun p x => p x). - -End Predicates. - -Arguments Pred0 {T}. -Arguments PredT {T}. -Prenex Implicits PredI PredU PredC PredD Preim. - -Notation "r1 +p r2" := (PredU r1 r2 : Pred _) - (at level 55, right associativity) : rel_scope. -Notation "r1 *p r2" := (xPredI r1 r2 : Pred _) - (at level 45, right associativity) : rel_scope. - -Notation "[ 'Pred' : T | E ]" := (SimplPred (fun _ : T => E)) - (at level 0, format "[ 'Pred' : T | E ]") : fun_scope. -Notation "[ 'Pred' x | E ]" := (SimplPred (fun x => E)) - (at level 0, x ident, format "[ 'Pred' x | E ]") : fun_scope. -Notation "[ 'Pred' x : T | E ]" := (SimplPred (fun x : T => E)) - (at level 0, x ident, only parsing) : fun_scope. -Notation "[ 'Pred' x y | E ]" := (SimplPred (fun t => let: (x, y) := t in E)) - (at level 0, x ident, y ident, format "[ 'Pred' x y | E ]") : fun_scope. -Notation "[ 'Pred' x y : T | E ]" := - (SimplPred (fun t : (T*T) => let: (x, y) := t in E)) - (at level 0, x ident, y ident, only parsing) : fun_scope. - -Definition repack_Pred T pT := - let: PropPredType _ a mP := pT return {type of @PropPredType T for pT} -> _ in - fun k => k a mP. - -Notation "[ 'PredType' 'of' T ]" := (repack_Pred (fun a => @PropPredType _ T a)) - (at level 0, format "[ 'PredType' 'of' T ]") : form_scope. - -Notation Pred_Class := (Pred_Sort (PredPredType _)). -Coercion Sort_of_Simpl_Pred T (p : Simpl_Pred T) : Pred_Class := p : Pred T. - -Definition PredArgType := Type. -Coercion Pred_of_argType (T : PredArgType) : Simpl_Pred T := PredT. - -Notation "{ :: T }" := (T%type : PredArgType) - (at level 0, format "{ :: T }") : type_scope. - -(* These must be defined outside a Section because "cooking" kills the *) -(* nosimpl tag. *) -Definition Mem T (pT : PredType T) : pT -> Mem_Pred T := - nosimpl (let: PropPredType _ _ (exist mem _) := pT return pT -> _ in mem). -Definition InMem T x mp := nosimpl Pred_of_Mem T mp x. - -Prenex Implicits Mem. - -(* Membership Predicates can be used as simple ones *) -Coercion Pred_of_Mem_Pred T mp := [Pred x : T | InMem x mp]. - -(* equality and subset *) - -Definition EqPredType T (pT : PredType T) (p1 p2 : pT) := - forall x : T, toPred p1 x <-> toPred p2 x. - -Definition SubPredType T (pT : PredType T) (p1 p2 : pT) := - forall x : T, toPred p1 x -> toPred p2 x. - -Definition EqSimplPred T (p1 p2 : Simpl_Pred T) := EqPredType p1 p2. -Definition SubSimplPred T (p1 p2 : Simpl_Pred T) := SubPredType p1 p2. - -Definition EqPredFun T1 T2 (pT2 : PredType T2) p1 p2 := - forall x : T1, @EqPredType T2 pT2 (p1 x) (p2 x). -Definition SubPredFun T1 T2 (pT2 : PredType T2) p1 p2 := - forall x : T1, @SubPredType T2 pT2 (p1 x) (p2 x). - -Definition EqMem T p1 p2 := forall x : T, InMem x p1 <-> InMem x p2. -Definition SubMem T p1 p2 := forall x : T, InMem x p1 -> InMem x p2. - -Notation "A <~> B" := (@EqPredType _ _ A B) - (at level 70, no associativity) : rel_scope. -Notation "A ~> B" := (@SubPredType _ _ A B) - (at level 70, no associativity) : rel_scope. -Notation "A <~1> B" := (@EqPredFun _ _ _ A B) - (at level 70, no associativity) : rel_scope. -Notation "A ~1> B" := (@SubPredFun _ _ _ A B) - (at level 70, no associativity) : rel_scope. - -Notation "x \In A" := (InMem x (Mem A)) - (at level 70, no associativity) : rel_scope. -Notation "x \Notin A" := (~ (x \In A)) - (at level 70, no associativity) : rel_scope. -Notation "A =p B" := (EqMem (Mem A) (Mem B)) - (at level 70, no associativity) : type_scope. -Notation "A <=p B" := (SubMem (Mem A) (Mem B)) - (at level 70, no associativity) : type_scope. - -(* Some notation for turning PredTypes into Pred or Simple Pred *) -Notation "[ 'Mem' A ]" := (Pred_of_Simpl (Pred_of_Mem_Pred (Mem A))) - (at level 0, only parsing) : fun_scope. -Notation "[ 'PredI' A & B ]" := (PredI [Mem A] [Mem B]) - (at level 0, format "[ 'PredI' A & B ]") : fun_scope. -Notation "[ 'PredU' A & B ]" := (PredU [Mem A] [Mem B]) - (at level 0, format "[ 'PredU' A & B ]") : fun_scope. -Notation "[ 'PredD' A & B ]" := (PredD [Mem A] [Mem B]) - (at level 0, format "[ 'PredD' A & B ]") : fun_scope. -Notation "[ 'PredC' A ]" := (PredC [Mem A]) - (at level 0, format "[ 'PredC' A ]") : fun_scope. -Notation "[ 'Preim' f 'of' A ]" := (Preim f [Mem A]) - (at level 0, format "[ 'Preim' f 'of' A ]") : fun_scope. - -Notation "[ 'Pred' x \In A ]" := [Pred x | x \In A] - (at level 0, x ident, format "[ 'Pred' x \In A ]") : fun_scope. -Notation "[ 'Pred' x \In A | E ]" := [Pred x | (x \In A) /\ E] - (at level 0, x ident, format "[ 'Pred' x \In A | E ]") : fun_scope. -Notation "[ 'Pred' x y \In A & B | E ]" := - [Pred x y | (x \In A) /\ (y \In B) /\ E] - (at level 0, x ident, y ident, - format "[ 'Pred' x y \In A & B | E ]") : fun_scope. -Notation "[ 'Pred' x y \In A & B ]" := [Pred x y | (x \In A) /\ (y \In B)] - (at level 0, x ident, y ident, - format "[ 'Pred' x y \In A & B ]") : fun_scope. -Notation "[ 'Pred' x y \In A | E ]" := [Pred x y \In A & A | E] - (at level 0, x ident, y ident, - format "[ 'Pred' x y \In A | E ]") : fun_scope. -Notation "[ 'Pred' x y \In A ]" := [Pred x y \In A & A] - (at level 0, x ident, y ident, - format "[ 'Pred' x y \In A ]") : fun_scope. - -Section Simplifications. -Variables (T : Type) (pT : PredType T). - -Lemma Mem_toPred : forall (p : pT), Mem (toPred p) = Mem p. -Proof. by rewrite /Mem; case: pT => T1 app1 [mem1 /= ->]. Qed. - -Lemma toPredE x (p : pT) : toPred p x = (x \In p). -Proof. by rewrite -Mem_toPred. Qed. - -Lemma In_Simpl x (p : Simpl_Pred T) : (x \In p) = p x. -Proof. by []. Qed. - -Lemma Simpl_PredE (p : Pred T) : p <~> [Pred x | p x]. -Proof. by []. Qed. - -Lemma Mem_Simpl (p : Simpl_Pred T) : Mem p = p :> Pred T. -Proof. by []. Qed. - -Definition MemE := Mem_Simpl. (* could be extended *) - -Lemma Mem_Mem (p : pT) : (Mem (Mem p) = Mem p) * (Mem [Mem p] = Mem p). -Proof. by rewrite -Mem_toPred. Qed. - -End Simplifications. - -(**************************************) -(* Definitions and lemmas for setoids *) -(**************************************) - -Section RelProperties. -Variables (T : Type) (pT : PredType T). - -Lemma EqPredType_refl (r : pT) : EqPredType r r. Proof. by []. Qed. -Lemma SubPredType_refl (r : pT) : SubPredType r r. Proof. by []. Qed. - -Lemma EqPredType_sym (r1 r2 : pT) : EqPredType r1 r2 -> EqPredType r2 r1. -Proof. by move=>H1 x; split; move/H1. Qed. - -Lemma EqPredType_trans' (r1 r2 r3 : pT) : - EqPredType r1 r2 -> EqPredType r2 r3 -> EqPredType r1 r3. -Proof. by move=>H1 H2 x; split; [move/H1; move/H2 | move/H2; move/H1]. Qed. - -Lemma SubPredType_trans' (r1 r2 r3 : pT) : - SubPredType r1 r2 -> SubPredType r2 r3 -> SubPredType r1 r3. -Proof. by move=>H1 H2 x; move/H1; move/H2. Qed. - -Definition EqPredType_trans r2 r1 r3 := @EqPredType_trans' r1 r2 r3. -Definition SubPredType_trans r2 r1 r3 := @SubPredType_trans' r1 r2 r3. -End RelProperties. - -Hint Resolve EqPredType_refl SubPredType_refl : core. - -(* Declaration of relations *) - -Add Parametric Relation T (pT : PredType T) : pT (@EqPredType _ pT) - reflexivity proved by (@EqPredType_refl _ _) - symmetry proved by (@EqPredType_sym _ _) - transitivity proved by (@EqPredType_trans' _ _) as EqPredType_rel. - -Add Parametric Relation T : (Simpl_Pred T) (@EqSimplPred _) - reflexivity proved by (@EqPredType_refl _ _) - symmetry proved by (@EqPredType_sym _ _) - transitivity proved by (@EqPredType_trans' _ _) as EqSimplPred_rel. - -Add Parametric Relation T : (Simpl_Pred T) (@SubSimplPred _) - reflexivity proved by (@SubPredType_refl _ _) - transitivity proved by (@SubPredType_trans' _ _) as SubSimplPred_rel. - -Add Parametric Relation T : (Mem_Pred T) (@EqMem T) - reflexivity proved by (@EqPredType_refl _ _) - symmetry proved by (@EqPredType_sym _ _) - transitivity proved by (@EqPredType_trans' _ _) as EqMem_rel. - -Add Parametric Relation T : (Mem_Pred T) (@SubMem _) - reflexivity proved by (@SubPredType_refl _ _) - transitivity proved by (@SubPredType_trans' _ _) as SubMem_rel. - -(* Declaring morphisms. *) - -Add Parametric Morphism T : (@Pred_of_Simpl T) with signature - @EqSimplPred _ ==> @EqPredType T (PredPredType T) as Pred_of_Simpl_morph. -Proof. by []. Qed. - -Add Parametric Morphism T (pT : PredType T) : (@EqPredType T pT) with signature - @EqPredType _ _ ==> @EqPredType _ _ ==> iff as EqPredType_morph. -Proof. by move=>r1 s1 H1 r2 s2 H2; rewrite H1 H2. Qed. - -Add Parametric Morphism T (pT : PredType T) : (@SubPredType T pT) with signature - @EqPredType _ _ ==> @EqPredType _ _ ==> iff as SubPred_morph. -Proof. by move=>r1 s1 H1 r2 s2 H2; split=>H x; move/H1; move/H; move/H2. Qed. - -Add Parametric Morphism T : (@InMem T) with signature - @eq _ ==> @EqMem _ ==> iff as InMem_morph. -Proof. by move=>x r s H; split; move/H. Qed. - -Add Parametric Morphism T (pT : PredType T) : (@Mem T pT) with signature - @EqPredType _ _ ==> @EqMem _ as Mem_morhp. -Proof. by move=>x y H p; rewrite /EqPredType -!toPredE in H *; rewrite H. Qed. - -Add Parametric Morphism T : (@PredU T) with signature - @EqPredType _ _ ==> @EqPredType _ _ ==> @EqSimplPred _ as predU_morph. -Proof. -move=>r1 s1 H1 r2 h2 H2 x; split; -by case; [move/H1 | move/H2]=>/=; auto. -Qed. - -Add Parametric Morphism T : (@PredI T) with signature - @EqPredType _ _ ==> @EqPredType _ _ ==> @EqPredType _ _ as predI_morph. -Proof. -move=>r1 s1 H1 r2 s2 H2 x; split; -by case; move/H1=>T1; move/H2=>T2. -Qed. - -Add Parametric Morphism T : (@PredC T) with signature - @EqPredType _ _ ==> @EqPredType _ _ as predC_morph. -Proof. by move=>r s H x; split=>H1; apply/H. Qed. - -Section RelLaws. -Variable (T : Type). - -Lemma orrI (r : Pred nat) : r +p r <~> r. -Proof. by move=>x; split; [case | left]. Qed. - -Lemma orrC (r1 r2 : Pred T) : r1 +p r2 <~> r2 +p r1. -Proof. move=>x; split=>/=; tauto. Qed. - -Lemma orr0 (r : Pred T) : r +p Pred0 <~> r. -Proof. by move=>x; split; [case | left]. Qed. - -Lemma or0r (r : Pred T) : Pred0 +p r <~> r. -Proof. by rewrite orrC orr0. Qed. - -Lemma orrCA (r1 r2 r3 : Pred T) : r1 +p r2 +p r3 <~> r2 +p r1 +p r3. -Proof. by move=>x; split=>/=; intuition. Qed. - -Lemma orrAC (r1 r2 r3 : Pred T) : (r1 +p r2) +p r3 <~> (r1 +p r3) +p r2. -Proof. by move=>?; split=>/=; intuition. Qed. - -Lemma orrA (r1 r2 r3 : Pred T) : (r1 +p r2) +p r3 <~> r1 +p r2 +p r3. -Proof. by rewrite (orrC r2) orrCA orrC. Qed. - -(* absorption *) -Lemma orrAb (r1 a : Pred T) : r1 <~> r1 +p a <-> a ~> r1. -Proof. -split; first by move=>-> x /=; auto. -move=>H x /=; split; first by auto. -by case=>//; move/H. -Qed. - -Lemma sub_orl (r1 r2 : Pred T) : r1 ~> r1 +p r2. Proof. by left. Qed. -Lemma sub_orr (r1 r2 : Pred T) : r2 ~> r1 +p r2. Proof. by right. Qed. - -End RelLaws. - -Section SubMemLaws. -Variable T : Type. - -Lemma subp_refl (p : Pred T) : p <=p p. -Proof. by []. Qed. - -Lemma subp_asym (p1 p2 : Pred T) : p1 <=p p2 -> p2 <=p p1 -> p1 =p p2. -Proof. by move=>H1 H2 x; split; [move/H1 | move/H2]. Qed. - -Lemma subp_trans (p2 p1 p3 : Pred T) : p1 <=p p2 -> p2 <=p p3 -> p1 <=p p3. -Proof. by move=>H1 H2 x; move/H1; move/H2. Qed. - -Lemma subp_or (p1 p2 q : Pred T) : p1 <=p q /\ p2 <=p q <-> p1 +p p2 <=p q. -Proof. -split=>[[H1] H2 x|H1]; first by case; [move/H1 | move/H2]. -by split=>x H2; apply: H1; [left | right]. -Qed. - -Lemma subp_and (p1 p2 q : Pred T) : q <=p p1 /\ q <=p p2 <-> q <=p p1 *p p2. -Proof. -split=>[[H1] H2 x|] H; last by split=>x; case/H. -by split; [apply: H1 | apply: H2]. -Qed. - -Lemma subp_orl (p1 p2 q : Pred T) : p1 <=p p2 -> p1 +p q <=p p2 +p q. -Proof. by move=>H x; case; [move/H; left|right]. Qed. - -Lemma subp_orr (p1 p2 q : Pred T) : p1 <=p p2 -> q +p p1 <=p q +p p2. -Proof. by move=>H x; case; [left | move/H; right]. Qed. - -Lemma subp_andl (p1 p2 q : Pred T) : p1 <=p p2 -> p1 *p q <=p p2 *p q. -Proof. by by move=>H x [H1 H2]; split; [apply: H|]. Qed. - -Lemma subp_andr (p1 p2 q : Pred T) : p1 <=p p2 -> q *p p1 <=p q *p p2. -Proof. by move=>H x [H1 H2]; split; [|apply: H]. Qed. - -End SubMemLaws. - -Hint Resolve subp_refl : core. - -Section ListMembership. -Variable T : Type. - -Fixpoint Mem_Seq (s : seq T) := - if s is y::s' then (fun x => x = y \/ Mem_Seq s' x) else xPred0. - -Definition EqSeq_Class := seq T. -Identity Coercion seq_of_EqSeq : EqSeq_Class >-> seq. - -Coercion Pred_of_Eq_Seq (s : EqSeq_Class) : Pred_Class := [eta Mem_Seq s]. - -Canonical Structure seq_PredType := @mkPredType T (seq T) Pred_of_Eq_Seq. -(* The line below makes Mem_Seq a canonical instance of topred. *) -Canonical Structure Mem_Seq_PredType := mkPredType Mem_Seq. - -Lemma In_cons y s x : (x \In y :: s) <-> (x = y) \/ (x \In s). -Proof. by []. Qed. - -Lemma In_nil x : (x \In [::]) <-> False. -Proof. by []. Qed. - -Lemma Mem_Seq1 x y : (x \In [:: y]) <-> (x = y). -Proof. by rewrite In_cons orpF. Qed. - -Definition InE := (Mem_Seq1, In_cons, In_Simpl). - -Lemma Mem_cat x : forall s1 s2, (x \In s1 ++ s2) <-> x \In s1 \/ x \In s2. -Proof. -elim=>[|y s1 IH] s2 /=; first by split; [right | case]. -rewrite !InE /=. -split. -- case=>[->|/IH]; first by left; left. - by case; [left; right | right]. -case; first by case; [left | move=>H; right; apply/IH; left]. -by move=>H; right; apply/IH; right. -Qed. - -Lemma In_split x s : x \In s -> exists s1 s2, s = s1 ++ x :: s2. -Proof. -elim:s=>[|y s IH] //=; rewrite InE. -case=>[<-|]; first by exists [::], s. -by case/IH=>s1 [s2 ->]; exists (y :: s1), s2. -Qed. - -End ListMembership. - -Lemma Mem_map T T' (f : T -> T') x (s : seq T) : - x \In s -> f x \In (map f s). -Proof. -elim: s=>[|y s IH] //; rewrite InE /=. -by case=>[<-|/IH]; [left | right]. -Qed. - -Lemma Mem_map_inv T T' (f : T -> T') x (s : seq T) : - x \In (map f s) -> exists y, x = f y /\ y \In s. -Proof. -elim: s=>[|y s IH] //=; rewrite InE /=. -case; first by move=>->; exists y; split=>//; left. -by case/IH=>z [->]; exists z; split=>//; right. -Qed. - -Prenex Implicits Mem_map_inv. - -(* Setoids for extensional equality of functions *) - -Lemma eqfun_refl A B (f : A -> B) : f =1 f. Proof. by []. Qed. -Lemma eqfun_sym A B (f1 f2 : A -> B) : f1 =1 f2 -> f2 =1 f1. -Proof. by move=>H x; rewrite H. Qed. -Lemma eqfun_trans A B (f1 f2 f3 : A -> B) : f1 =1 f2 -> f2 =1 f3 -> f1 =1 f3. -Proof. by move=>H1 H2 x; rewrite H1 H2. Qed. - -Add Parametric Relation A B : (A -> B) (@eqfun _ _) - reflexivity proved by (@eqfun_refl A B) - symmetry proved by (@eqfun_sym A B) - transitivity proved by (@eqfun_trans A B) as eqfun_morph. - - -(***********************************) -(* Image of a collective predicate *) -(***********************************) - -Section Image. -Variables (A B : Type) (P : Pred A) (f : A -> B). - -Inductive image_spec b : Prop := Im_mem a of b = f a & a \In P. - -Definition Image' : Pred B := image_spec. - -End Image. - -(* swap to make the notation consider P before E; helps inference *) -Notation Image f P := (Image' P f). - -Notation "[ 'Image' E | i <- s ]" := (Image (fun i => E) s) - (at level 0, E at level 99, i ident, - format "[ '[hv' 'Image' E '/ ' | i <- s ] ']'") : rel_scope. - -Notation "[ 'Image' E | i <- s & C ]" := [Image E | i <- [PredI s & C]] - (at level 0, E at level 99, i ident, - format "[ '[hv' 'Image' E '/ ' | i <- s '/ ' & C ] ']'") : rel_scope. - -Notation "[ 'Image' E | i : T <- s ]" := (Image (fun i : T => E) s) - (at level 0, E at level 99, i ident, only parsing) : rel_scope. - -Notation "[ 'Image' E | i : T <- s & C ]" := - [Image E | i : T <- [PredI s & C]] - (at level 0, E at level 99, i ident, only parsing) : rel_scope. - -Lemma Image_mem A B (f : A -> B) (P : Pred A) x : x \In P -> f x \In Image f P. -Proof. by apply: Im_mem. Qed. - -Lemma Image_inj_sub A B (f : A -> B) (X1 X2 : Pred A) : - injective f -> Image f X1 <=p Image f X2 -> X1 <=p X2. -Proof. by move=>H E x /(Image_mem f) /E [y] /H ->. Qed. - -Lemma Image_inj_eqmem A B (f : A -> B) (X1 X2 : Pred A) : - injective f -> Image f X1 =p Image f X2 -> X1 =p X2. -Proof. by move=>H E; split; apply: Image_inj_sub H _ _; rewrite E. Qed. - -Lemma ImageU A B (f : A -> B) (X1 X2 : Pred A) : - Image f (PredU X1 X2) =p [PredU Image f X1 & Image f X2]. -Proof. -move=>x; split. -- by case=>y -> [H|H]; [left | right]; apply: Image_mem. -by case; case=>y -> H; apply: Image_mem; [left | right]. -Qed. - -Lemma ImageIm A B C (f1 : B -> C) (f2 : A -> B) (X : Pred A) : - Image f1 (Image f2 X) =p Image (f1 \o f2) X. -Proof. -move=>x; split; first by case=>_ -> [x' ->] H; exists x'. -by case=>a -> H; exists (f2 a)=>//; exists a. -Qed. - -Lemma ImageEq A B (f1 f2 : A -> B) (X : Pred A) : - f1 =1 f2 -> Image f1 X =p Image f2 X. -Proof. by move=>H x; split; case=>a ->; exists a. Qed. - diff --git a/pcm/prelude.v b/pcm/prelude.v deleted file mode 100644 index 781a87e..0000000 --- a/pcm/prelude.v +++ /dev/null @@ -1,245 +0,0 @@ -(* -Copyright 2010 IMDEA Software Institute -Licensed under the Apache License, Version 2.0 (the "License"); -you may not use this file except in compliance with the License. -You may obtain a copy of the License at - http://www.apache.org/licenses/LICENSE-2.0 -Unless required by applicable law or agreed to in writing, software -distributed under the License is distributed on an "AS IS" BASIS, -WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. -See the License for the specific language governing permissions and -limitations under the License. -*) - -(******************************************************************************) -(* This file is Prelude -- often used notation definitions and lemmas that *) -(* are not included in the other libraries. *) -(******************************************************************************) - -From Coq Require Import ssreflect ssrbool ssrfun Eqdep. -From mathcomp Require Import ssrnat eqtype. -Set Implicit Arguments. -Unset Strict Implicit. -Unset Printing Implicit Defensive. - -(***********) -(* Prelude *) -(***********) - -(* export inj_pair without exporting the whole Eqdep library *) -Definition inj_pair2 := @inj_pair2. -Arguments inj_pair2 {U P p x y} _. - -(* eta laws for pairs and units *) -Notation prod_eta := surjective_pairing. - -(* eta law often used with injection *) -Lemma prod_inj A B (x y : A * B) : x = y <-> (x.1, x.2) = (y.1, y.2). -Proof. by case: x y=>x1 x2 []. Qed. - -(* This declaration won't be needed after Coq 8.8.0 is out *) -Prenex Implicits Some_inj. - -(***************************) -(* operations on functions *) -(***************************) - -Lemma compA A B C D (h : A -> B) (g : B -> C) (f : C -> D) : - (f \o g) \o h = f \o (g \o h). -Proof. by []. Qed. - -Definition fprod A1 A2 B1 B2 (f1 : A1 -> B1) (f2 : A2 -> B2) := - fun (x : A1 * A2) => (f1 x.1, f2 x.2). - -Notation "f1 \* f2" := (fprod f1 f2) (at level 42, left associativity). - -Lemma ext A (B : A -> Type) (f1 f2 : forall x, B x) : - f1 = f2 -> forall x, f1 x = f2 x. -Proof. by move=>->. Qed. - -(* function splice *) -Definition fsplice A B1 B2 (f1 : A -> B1) (f2 : A -> B2) := - fun x : A => (f1 x, f2 x). - -Notation "[ 'fs' f1 , f2 ]" := (fsplice f1 f2) - (at level 0, format "[ 'fs' f1 , f2 ]"). - -(* composing relation and function *) -(* should i turn relations into functions on products? *) -(* then i won't need the next definition *) -(* it will be described using composition and splicing *) - -Definition relfuncomp A B (D : rel A) (f : B -> A) : rel B := - fun x y => D (f x) (f y). - -Notation "D \\o f" := (@relfuncomp _ _ D f) (at level 42, left associativity). - -(************************) -(* extension to ssrbool *) -(************************) - -Reserved Notation "[ /\ P1 , P2 , P3 , P4 , P5 & P6 ]" (at level 0, format - "'[hv' [ /\ '[' P1 , '/' P2 , '/' P3 , '/' P4 , '/' P5 ']' '/ ' & P6 ] ']'"). - -Reserved Notation "[ /\ P1 , P2 , P3 , P4 , P5 , P6 & P7 ]" (at level 0, format - "'[hv' [ /\ '[' P1 , '/' P2 , '/' P3 , '/' P4 , '/' P5 , '/' P6 ']' '/ ' & P7 ] ']'"). - -Reserved Notation "[ \/ P1 , P2 , P3 , P4 | P5 ]" (at level 0, format - "'[hv' [ \/ '[' P1 , '/' P2 , '/' P3 , '/' P4 ']' '/ ' | P5 ] ']'"). -Reserved Notation "[ \/ P1 , P2 , P3 , P4 , P5 | P6 ]" (at level 0, format - "'[hv' [ \/ '[' P1 , '/' P2 , '/' P3 , '/' P4 , '/' P5 ']' '/ ' | P6 ] ']'"). - -Reserved Notation "[ \/ P1 , P2 , P3 , P4 , P5 , P6 | P7 ]" (at level 0, format - "'[hv' [ \/ '[' P1 , '/' P2 , '/' P3 , '/' P4 , '/' P5 , '/' P6 ']' '/ ' | P7 ] ']'"). - -Inductive and6 (P1 P2 P3 P4 P5 P6 : Prop) : Prop := - And6 of P1 & P2 & P3 & P4 & P5 & P6. -Inductive and7 (P1 P2 P3 P4 P5 P6 P7 : Prop) : Prop := - And7 of P1 & P2 & P3 & P4 & P5 & P6 & P7. - -Inductive or5 (P1 P2 P3 P4 P5 : Prop) : Prop := - Or51 of P1 | Or52 of P2 | Or53 of P3 | Or54 of P4 | Or55 of P5. -Inductive or6 (P1 P2 P3 P4 P5 P6 : Prop) : Prop := - Or61 of P1 | Or62 of P2 | Or63 of P3 | Or64 of P4 | Or65 of P5 | Or66 of P6. -Inductive or7 (P1 P2 P3 P4 P5 P6 P7 : Prop) : Prop := - Or71 of P1 | Or72 of P2 | Or73 of P3 | Or74 of P4 | Or75 of P5 | Or76 of P6 | - Or77 of P7. - -Notation "[ /\ P1 , P2 , P3 , P4 , P5 & P6 ]" := - (and6 P1 P2 P3 P4 P5 P6) : type_scope. -Notation "[ /\ P1 , P2 , P3 , P4 , P5 , P6 & P7 ]" := - (and7 P1 P2 P3 P4 P5 P6 P7) : type_scope. -Notation "[ \/ P1 , P2 , P3 , P4 | P5 ]" := - (or5 P1 P2 P3 P4 P5) : type_scope. -Notation "[ \/ P1 , P2 , P3 , P4 , P5 | P6 ]" := - (or6 P1 P2 P3 P4 P5 P6) : type_scope. -Notation "[ \/ P1 , P2 , P3 , P4 , P5 , P6 | P7 ]" := - (or7 P1 P2 P3 P4 P5 P6 P7) : type_scope. - -Section ReflectConnectives. -Variable b1 b2 b3 b4 b5 b6 b7 : bool. - -Lemma and6P : reflect [/\ b1, b2, b3, b4, b5 & b6] [&& b1, b2, b3, b4, b5 & b6]. -Proof. -by case b1; case b2; case b3; case b4; case b5; case b6; - constructor; try by case. -Qed. - -Lemma and7P : reflect [/\ b1, b2, b3, b4, b5, b6 & b7] - [&& b1, b2, b3, b4, b5, b6 & b7]. -Proof. -by case b1; case b2; case b3; case b4; case b5; case b6; case: b7; - constructor; try by case. -Qed. - -Lemma or5P : reflect [\/ b1, b2, b3, b4 | b5] [|| b1, b2, b3, b4 | b5]. -Proof. -case b1; first by constructor; constructor 1. -case b2; first by constructor; constructor 2. -case b3; first by constructor; constructor 3. -case b4; first by constructor; constructor 4. -case b5; first by constructor; constructor 5. -by constructor; case. -Qed. - -Lemma or6P : reflect [\/ b1, b2, b3, b4, b5 | b6] [|| b1, b2, b3, b4, b5 | b6]. -Proof. -case b1; first by constructor; constructor 1. -case b2; first by constructor; constructor 2. -case b3; first by constructor; constructor 3. -case b4; first by constructor; constructor 4. -case b5; first by constructor; constructor 5. -case b6; first by constructor; constructor 6. -by constructor; case. -Qed. - -Lemma or7P : reflect [\/ b1, b2, b3, b4, b5, b6 | b7] - [|| b1, b2, b3, b4, b5, b6 | b7]. -Proof. -case b1; first by constructor; constructor 1. -case b2; first by constructor; constructor 2. -case b3; first by constructor; constructor 3. -case b4; first by constructor; constructor 4. -case b5; first by constructor; constructor 5. -case b6; first by constructor; constructor 6. -case b7; first by constructor; constructor 7. -by constructor; case. -Qed. - -End ReflectConnectives. - -Arguments and6P {b1 b2 b3 b4 b5 b6}. -Arguments and7P {b1 b2 b3 b4 b5 b6 b7}. -Arguments or5P {b1 b2 b3 b4 b5}. -Arguments or6P {b1 b2 b3 b4 b5 b6}. -Arguments or7P {b1 b2 b3 b4 b5 b6 b7}. - -Lemma andX (a b : bool) : reflect (a * b) (a && b). -Proof. by case: a; case: b; constructor=>//; case. Qed. - -Arguments andX {a b}. - -(**********************************************) -(* Reflexive-transitive closure of a relation *) -(**********************************************) - -Fixpoint iter' T (g : T -> T -> Prop) n s1 s2 := - if n is n'.+1 then exists s, g s1 s /\ iter' g n' s s2 else s1 = s2. -Definition iter T (g : T -> T -> Prop) s1 s2 := exists n, iter' g n s1 s2. -(* curry the iteration *) -Definition iterc A T (g : A -> T -> A -> T -> Prop) a1 s1 a2 s2 := - iter (fun (a b : A * T) => g a.1 a.2 b.1 b.2) (a1, s1) (a2, s2). - -Section IteratedRels. -Variable T : Type. -Implicit Type g : T -> T -> Prop. - -Lemma iter_refl g s : iter g s s. -Proof. by exists 0. Qed. - -Lemma iter_trans g s1 s2 s3 : iter g s1 s2 -> iter g s2 s3 -> iter g s1 s3. -Proof. -case=>n; elim: n s1 =>[|n IH] s1 /=; first by move=>->. -by case=>s [H1 H2] /(IH _ H2) [m]; exists m.+1, s. -Qed. - -Lemma iterS g n s1 s2 : - iter' g n.+1 s1 s2 <-> exists s, iter' g n s1 s /\ g s s2. -Proof. -elim: n s1=>[|n IH] s1. -- by split; [case=>s [H <-]; exists s1 | case=>s [<- H]; exists s2]. -split; first by case=>s [H1] /IH [s'][H G]; exists s'; split=>//; exists s. -by case=>s [[s'][H1 H G]]; exists s'; split=>//; apply/IH; exists s. -Qed. - -Lemma iter'_sub g g' n s1 s2 : - (forall s1 s2, g s1 s2 -> g' s1 s2) -> - iter' g n s1 s2 -> iter' g' n s1 s2. -Proof. by move=>H; elim: n s1=>[|n IH] s1 //= [s][/H G] /IH; exists s. Qed. - -Lemma iter_sub g g' s1 s2 : - (forall s1 s2, g s1 s2 -> g' s1 s2) -> iter g s1 s2 -> iter g' s1 s2. -Proof. by move=>H [n]; exists n; apply: iter'_sub H _. Qed. - -Lemma iter1 g s1 s2 : g s1 s2 -> iter g s1 s2. -Proof. by exists 1, s2. Qed. - -End IteratedRels. - -Lemma iter2 A T (g : A -> T -> A -> T -> Prop) x1 s1 x2 s2 : - g x1 s1 x2 s2 -> iterc g x1 s1 x2 s2. -Proof. by apply: iter1. Qed. - -Prenex Implicits iter1 iter2. -Arguments iter_refl {T g s}. -Hint Resolve iter_refl : core. - -(**************) -(* empty type *) -(**************) - -Lemma emptyset_eqP : Equality.axiom (fun _ _ : Empty_set => true). -Proof. by case. Qed. - -Definition emptysetEqMix := EqMixin emptyset_eqP. -Canonical emptysetEqType := Eval hnf in EqType Empty_set emptysetEqMix. \ No newline at end of file diff --git a/pcm/unionmap.v b/pcm/unionmap.v index 8ade23c..7e9decf 100644 --- a/pcm/unionmap.v +++ b/pcm/unionmap.v @@ -11,18 +11,91 @@ See the License for the specific language governing permissions and limitations under the License. *) -From Coq Require Import ssreflect ssrbool ssrfun. -From mathcomp Require Import ssrnat eqtype seq path. -From fcsl Require Import ordtype finmap pcm. -Set Implicit Arguments. -Unset Strict Implicit. -Unset Printing Implicit Defensive. - (******************************************************************************) (* This file contains the reference implementation of finite maps with keys *) (* satisfying condition p and supporting disjoint union via a top element. *) +(* *) +(* Union maps signature definitions: *) +(* from f == union map f converted to a base finite map. *) +(* to m == finite map m converted to a union map. *) +(* defined f == union map f is valid. *) +(* um_undef == an undefined (i.e., invalid) instance of a union map. *) +(* empty == a valid empty instance of a union map. *) +(* upd k v f == union map f with key-value pair (k,v) inserted. If key k *) +(* already exists in t, its corresponding value is *) +(* overwritten with v. *) +(* dom f == a sequence of keys for union map f. *) +(* dom_eq f1 f2 == the sets of keys are equal for union maps f1 and f2. *) +(* assocs t == a sequence of key-value pairs in union map t. *) +(* free f k == union map f without the key k and its corresponding value. *) +(* find k f == a value corresponding to key k in union map f, wrapped in *) +(* an option type. *) +(* union f1 f2 == a union map formed by taking a disjoint union of maps f1 *) +(* and f2. If the maps are not disjoint, the result is *) +(* undefined. *) +(* empb f == union map f is valid and empty. *) +(* undefb f == union map f is undefined. *) +(* pts k v == a fresh instance of a union map containing a single *) +(* key-value pair (k,v). *) +(* *) +(* Defined notions: *) +(* um_preim p f k == a value corresponding to key k exists in union map f *) +(* and satisfies predicate p. *) +(* range f == a sequence of values for union map f. *) +(* um_mono f == values of union map f are in a monotonically *) +(* increasing order. *) +(* um_foldl a z0 d f == if f is valid, a result of a left fold over its *) +(* key-value pairs using function a and starting *) +(* value z0, d otherwise. *) +(* um_foldr a z0 d f == if f is valid, a result of a right fold over its *) +(* key-value pairs using function a and starting *) +(* value z0, d otherise. *) +(* omap a f == the result of applying a generalized *) +(* mapping/filtering function a to union map t. The *) +(* function may inspect both the key and the value and *) +(* return either a new value wrapped in Some, or None if *) +(* the key-value pair is to be excluded. *) +(* omapv a f == same as omap, but function a may only inspect the *) +(* value. *) +(* mapv a f == same as omapv, but function a may only return the new *) +(* value, i.e. it is a standard functional mapping. *) +(* um_filter p f == the result of applying a filtering function p to *) +(* union map f. Function f may inspect both the key and *) +(* the value. *) +(* um_filterk p f == same as um_filter, but function p may only inspect *) +(* the key. *) +(* um_filterv p f == same as um_filter, but function p may only inspect *) +(* the value. *) +(* oeval a ks f z0 == the result of iteratively applying function a to *) +(* key-value pairs in union map f, in the order *) +(* indicated by sequence of keys ks. *) +(* eval a p f z0 == the result of iteratively applying function a to *) +(* all key-value pairs satisfying predicate p in union *) +(* map f. *) +(* um_count p f == the number of key-values pairs satisfying predicate p *) +(* in union map f. *) +(* dom_map f == a union map f converted to a finite set, i.e., all *) +(* its values are replaced by tt. *) +(* inverse f == a union map f with its keys and values swapped, i.e., *) +(* the values are now serving as keys and vice versa. *) +(* um_comp g f == a union map obtained by composing union maps g and f. *) +(* The keys taken from f, and their corresponding values *) +(* are treated as keys in g for looking up the final *) +(* values. *) +(* um_all P f == type-level predicate P holds for all key-value pairs *) +(* in union map f. *) +(* um_allb p f == predicate p holds for all key-value pairs in union *) +(* map f. *) +(* um_all2 P f1 f2 == binary type-level predicate P holds for all values of *) +(* equal keys in union maps f1 and f2. *) (******************************************************************************) +From Coq Require Import ssreflect ssrbool ssrfun. +From mathcomp Require Import ssrnat eqtype seq path. +From fcsl Require Import prelude finmap seqperm pcm morphism pred. +From fcsl Require Export ordtype. +From fcsl Require Import options. + (* I decided to have union_map_class be a class, rather than a structure. The class packages a condition on keys. Ordinary union_maps have a trivially true condition. Heaps have a condition that the @@ -50,54 +123,61 @@ union_mapUMC), just to keep the universes from jumping. *) Module UM. Section UM. -Variables (K : ordType) (V : Type) (p : pred K). +Variables (K : ordType) (C : pred K) (V : Type). Inductive base := - Undef | Def (f : {finMap K -> V}) of all p (supp f). + Undef | Def (f : finMap K V) of all C (supp f). Section FormationLemmas. -Variable (f g : {finMap K -> V}). +Variable (f g : finMap K V). -Lemma all_supp_insP k v : p k -> all p (supp f) -> all p (supp (ins k v f)). +Lemma all_supp_insP (k : K) v : + C k -> all C (supp f) -> all C (supp (ins k v f)). Proof. move=>H1 H2; apply/allP=>x; rewrite supp_ins inE /=. by case: eqP=>[->|_] //=; move/(allP H2). Qed. -Lemma all_supp_remP k : all p (supp f) -> all p (supp (rem k f)). +Lemma all_supp_remP k : all C (supp f) -> all C (supp (rem k f)). Proof. move=>H; apply/allP=>x; rewrite supp_rem inE /=. by case: eqP=>[->|_] //=; move/(allP H). Qed. Lemma all_supp_fcatP : - all p (supp f) -> all p (supp g) -> all p (supp (fcat f g)). + all C (supp f) -> all C (supp g) -> all C (supp (fcat f g)). Proof. move=>H1 H2; apply/allP=>x; rewrite supp_fcat inE /=. by case/orP; [move/(allP H1) | move/(allP H2)]. Qed. Lemma all_supp_kfilterP q : - all p (supp f) -> all p (supp (kfilter q f)). + all C (supp f) -> all C (supp (kfilter q f)). Proof. move=>H; apply/allP=>x; rewrite supp_kfilt mem_filter. by case/andP=>_ /(allP H). Qed. +Lemma all_supp_mapfP (a : K -> V -> V) : + all C (supp f) -> all C (supp (mapf a f)). +Proof. +by move=>H; apply/allP=>x; rewrite supp_mapf; move/(allP H). +Qed. + End FormationLemmas. Implicit Types (k : K) (v : V) (f g : base). Lemma umapE (f g : base) : f = g <-> match f, g with - Def f' pf, Def g' pg => f' = g' + Def f' pf, Def g' pg => f' = g' | Undef, Undef => true | _, _ => false end. Proof. split; first by move=>->; case: g. case: f; case: g=>// f pf g pg E; rewrite {g}E in pg *. -by congr Def; apply: bool_irrelevance. +by congr Def; apply: eq_irrelevance. Qed. Definition valid f := if f is Def _ _ then true else false. @@ -106,7 +186,7 @@ Definition empty := @Def (finmap.nil K V) is_true_true. Definition upd k v f := if f is Def fs fpf then - if decP (@idP (p k)) is left pf then + if decP (@idP (C k)) is left pf then Def (all_supp_insP v pf fpf) else Undef else Undef. @@ -121,10 +201,13 @@ Definition dom_eq f1 f2 := | _, _ => false end. -Definition free k f := +Definition assocs f : seq (K * V) := + if f is Def fs _ then seq_of fs else [::]. + +Definition free f k := if f is Def fs pf then Def (all_supp_remP k pf) else Undef. -Definition find k f := if f is Def fs _ then fnd k fs else None. +Definition find k f := if f is Def fs _ then fnd (K:=K) k fs else None. Definition union f1 f2 := if (f1, f2) is (Def fs1 pf1, Def fs2 pf2) then @@ -133,9 +216,6 @@ Definition union f1 f2 := else Undef else Undef. -Definition um_filter q f := - if f is Def fs pf then Def (all_supp_kfilterP q pf) else Undef. - Definition empb f := if f is Def fs _ then supp fs == [::] else false. Definition undefb f := if f is Undef then true else false. @@ -146,7 +226,7 @@ Definition pts k v := upd k v empty. Lemma base_indf (P : base -> Prop) : P Undef -> P empty -> (forall k v f, P f -> valid (union (pts k v) f) -> - path ord k (dom f) -> + path ord (k : [ordType of K]) (dom f) -> P (union (pts k v) f)) -> forall f, P f. Proof. @@ -154,7 +234,7 @@ rewrite /empty => H1 H2 H3; apply: base_ind=>//. apply: fmap_ind'=>[|k v f S IH] H. - by set f := Def _ in H2; rewrite (_ : Def H = f) // /f umapE. have N : k \notin supp f by apply: notin_path S. -have [T1 T2] : p k /\ all p (supp f). +have [T1 T2] : C k /\ all C (supp f). - split; first by apply: (allP H k); rewrite supp_ins inE /= eq_refl. - apply/allP=>x T; apply: (allP H x). by rewrite supp_ins inE /= T orbT. @@ -174,8 +254,8 @@ Qed. (* backward induction *) Lemma base_indb (P : base -> Prop) : P Undef -> P empty -> - (forall k v f, P f -> valid (union (pts k v) f) -> - (forall x, x \in dom f -> ord x k) -> + (forall k v f, P f -> valid (union f (pts k v)) -> + (forall x : [ordType of K], x \in dom f -> ord x k) -> P (union (pts k v) f)) -> forall f, P f. Proof. @@ -183,15 +263,15 @@ rewrite /empty => H1 H2 H3; apply: base_ind=>//. apply: fmap_ind''=>[|k v f S IH] H. - by set f := Def _ in H2; rewrite (_ : Def H = f) // /f umapE. have N : k \notin supp f by apply/negP; move/S; rewrite ordtype.irr. -have [T1 T2] : p k /\ all p (supp f). +have [T1 T2] : C k /\ all C (supp f). - split; first by apply: (allP H k); rewrite supp_ins inE /= eq_refl. - apply/allP=>x T; apply: (allP H x). by rewrite supp_ins inE /= T orbT. have E : FinMap (sorted_ins' k v (sorted_nil K V)) = ins k v (@nil K V) by []. -have: valid (union (pts k v) (Def T2)). +have: valid (union (Def T2) (pts k v)). - rewrite /valid /union /pts /upd /=. case: decP; last by rewrite T1. - by move=>T; case: ifP=>//; rewrite E disjC disj_ins N disj_nil. + by move=>T; case: ifP=>//; rewrite E disj_ins N disj_nil. move/(H3 k v _ (IH T2)). rewrite (_ : union (pts k v) (Def T2) = Def H); first by apply; apply: S. rewrite umapE /union /pts /upd /=. @@ -206,7 +286,7 @@ End UM. Module UMC. Section MixinDef. -Variables (K : ordType) (V : Type) (p : pred K). +Variables (K : ordType) (C : pred K) (V : Type). Structure mixin_of (T : Type) := Mixin { defined_op : T -> bool; @@ -215,137 +295,126 @@ Structure mixin_of (T : Type) := Mixin { upd_op : K -> V -> T -> T; dom_op : T -> seq K; dom_eq_op : T -> T -> bool; - free_op : K -> T -> T; + assocs_op : T -> seq (K * V); + free_op : T -> K -> T; find_op : K -> T -> option V; union_op : T -> T -> T; - um_filter_op : pred K -> T -> T; empb_op : T -> bool; undefb_op : T -> bool; pts_op : K -> V -> T; - from_op : T -> UM.base V p; - to_op : UM.base V p -> T; + from_op : T -> UM.base C V; + to_op : UM.base C V -> T; _ : forall b, from_op (to_op b) = b; _ : forall f, to_op (from_op f) = f; _ : forall f, defined_op f = UM.valid (from_op f); - _ : undef_op = to_op (UM.Undef V p); - _ : empty_op = to_op (UM.empty V p); + _ : undef_op = to_op (UM.Undef C V); + _ : empty_op = to_op (UM.empty C V); _ : forall k v f, upd_op k v f = to_op (UM.upd k v (from_op f)); _ : forall f, dom_op f = UM.dom (from_op f); _ : forall f1 f2, dom_eq_op f1 f2 = UM.dom_eq (from_op f1) (from_op f2); - _ : forall k f, free_op k f = to_op (UM.free k (from_op f)); + _ : forall f, assocs_op f = UM.assocs (from_op f); + _ : forall f k, free_op f k = to_op (UM.free (from_op f) k); _ : forall k f, find_op k f = UM.find k (from_op f); _ : forall f1 f2, union_op f1 f2 = to_op (UM.union (from_op f1) (from_op f2)); - _ : forall q f, um_filter_op q f = to_op (UM.um_filter q (from_op f)); _ : forall f, empb_op f = UM.empb (from_op f); _ : forall f, undefb_op f = UM.undefb (from_op f); - _ : forall k v, pts_op k v = to_op (UM.pts p k v)}. + _ : forall k v, pts_op k v = to_op (UM.pts C k v)}. + End MixinDef. Section ClassDef. -Variables (K : ordType) (V : Type). +Variables (K : ordType) (C : pred K) (V : Type). -(* I used to package K and V into the class. I did that to get cond -function over keys to work without supplying the type parameter. I.e., -with K and V out, I have to write cond U k, where U : union_map_class -K V. But the complication wasn't worth it. One has to frequently -reorder arguments to various operations in a "less-ad-hoc" style, to -get the types to be inferred properly. While fun, it's just an -unnecessary hassle, since cond is not used that much. The condition p -can be kept in the structure, however, since no types depend on it. *) +Notation class_of := mixin_of (only parsing). -Structure class_of (T : Type) := Class { - p : pred K; - mixin : mixin_of V p T}. - -Structure type : Type := Pack {sort : Type; _ : class_of sort}. +Structure type : Type := Pack {sort : Type; _ : class_of C V sort}. Local Coercion sort : type >-> Sortclass. Variables (T : Type) (cT : type). -Definition class := let: Pack _ c as cT' := cT return class_of cT' in c. +Definition class := let: Pack _ c as cT' := cT return class_of C V cT' in c. Definition clone c of phant_id class c := @Pack T c. -Definition pack p (m : @mixin_of K V p T) := - @Pack T (@Class T p m). - -(* Definition cond : pred K := @p _ class. *) -Definition cond := [pred x : K | @p _ class x]. -Definition from := from_op (mixin class). -Definition to := to_op (mixin class). -Definition defined := defined_op (mixin class). -Definition um_undef := undef_op (mixin class). -Definition empty := empty_op (mixin class). -Definition upd : K -> V -> cT -> cT := upd_op (mixin class). -Definition dom : cT -> seq K := dom_op (mixin class). -Definition dom_eq := dom_eq_op (mixin class). -Definition free : K -> cT -> cT := free_op (mixin class). -Definition find : K -> cT -> option V := find_op (mixin class). -Definition union := union_op (mixin class). -Definition um_filter : pred K -> cT -> cT := um_filter_op (mixin class). -Definition empb := empb_op (mixin class). -Definition undefb := undefb_op (mixin class). -Definition pts : K -> V -> cT := pts_op (mixin class). +Definition pack (m : @mixin_of K C V T) := + @Pack T m. + +Definition from := from_op class. +Definition to := to_op class. +Definition defined := defined_op class. +Definition um_undef := undef_op class. +Definition empty := empty_op class. +Definition upd : K -> V -> cT -> cT := upd_op class. +Definition dom : cT -> seq K := dom_op class. +Definition dom_eq := dom_eq_op class. +Definition assocs : cT -> seq (K * V) := assocs_op class. +Definition free : cT -> K -> cT := free_op class. +Definition find : K -> cT -> option V := find_op class. +Definition union := union_op class. +Definition empb := empb_op class. +Definition undefb := undefb_op class. +Definition pts : K -> V -> cT := pts_op class. End ClassDef. -Arguments to {K V cT}. -Arguments um_undef {K V cT}. -Arguments empty {K V cT}. -Arguments pts {K V cT} _ _. +Arguments to {K C V cT}. +Arguments um_undef {K C V cT}. +Arguments empty {K C V cT}. +Arguments pts [K C V cT] _ _. +Prenex Implicits pts. Section Lemmas. -Variables (K : ordType) (V : Type) (U : type K V). +Variables (K : ordType) (C : pred K) (V : Type) (U : type C V). Local Coercion sort : type >-> Sortclass. Implicit Type f : U. -Lemma ftE (b : UM.base V (cond U)) : from (to b) = b. -Proof. by case: U b=>x [p][*]. Qed. +Lemma ftE (b : UM.base C V) : from (to b : U) = b. +Proof. by case: U b =>x [*]. Qed. Lemma tfE f : to (from f) = f. -Proof. by case: U f=>x [p][*]. Qed. +Proof. by case: U f=>x [*]. Qed. -Lemma eqE (b1 b2 : UM.base V (cond U)) : - to b1 = to b2 <-> b1 = b2. +Lemma eqE (b1 b2 : UM.base C V) : + to b1 = to b2 :> U <-> b1 = b2. Proof. by split=>[E|-> //]; rewrite -[b1]ftE -[b2]ftE E. Qed. Lemma defE f : defined f = UM.valid (from f). -Proof. by case: U f=>x [p][*]. Qed. +Proof. by case: U f=>x [*]. Qed. -Lemma undefE : um_undef = to (UM.Undef V (cond U)). -Proof. by case: U=>x [p][*]. Qed. +Lemma undefE : um_undef = to (UM.Undef C V) :> U. +Proof. by case: U=>x [*]. Qed. -Lemma emptyE : empty = to (UM.empty V (cond U)). -Proof. by case: U=>x [p][*]. Qed. +Lemma emptyE : empty = to (UM.empty C V) :> U. +Proof. by case: U=>x [*]. Qed. Lemma updE k v f : upd k v f = to (UM.upd k v (from f)). -Proof. by case: U k v f=>x [p][*]. Qed. +Proof. by case: U k v f=>x [*]. Qed. Lemma domE f : dom f = UM.dom (from f). -Proof. by case: U f=>x [p][*]. Qed. +Proof. by case: U f=>x [*]. Qed. Lemma dom_eqE f1 f2 : dom_eq f1 f2 = UM.dom_eq (from f1) (from f2). -Proof. by case: U f1 f2=>x [p][*]. Qed. +Proof. by case: U f1 f2=>x [*]. Qed. + +Lemma assocsE f : assocs f = UM.assocs (from f). +Proof. by case: U f=>x [*]. Qed. -Lemma freeE k f : free k f = to (UM.free k (from f)). -Proof. by case: U k f=>x [p][*]. Qed. +Lemma freeE f k : free f k = to (UM.free (from f) k). +Proof. by case: U k f=>x [*]. Qed. Lemma findE k f : find k f = UM.find k (from f). -Proof. by case: U k f=>x [p][*]. Qed. +Proof. by case: U k f=>x [*]. Qed. Lemma unionE f1 f2 : union f1 f2 = to (UM.union (from f1) (from f2)). -Proof. by case: U f1 f2=>x [p][*]. Qed. - -Lemma um_filterE q f : um_filter q f = to (UM.um_filter q (from f)). -Proof. by case: U q f=>x [p][*]. Qed. +Proof. by case: U f1 f2=>x [*]. Qed. Lemma empbE f : empb f = UM.empb (from f). -Proof. by case: U f=>x [p][*]. Qed. +Proof. by case: U f=>x [*]. Qed. Lemma undefbE f : undefb f = UM.undefb (from f). -Proof. by case: U f=>x [p][*]. Qed. +Proof. by case: U f=>x [*]. Qed. -Lemma ptsE k v : pts k v = to (UM.pts (cond U) k v). -Proof. by case: U k v=>x [p][*]. Qed. +Lemma ptsE k v : pts k v = to (UM.pts C k v) :> U. +Proof. by case: U k v=>x [*]. Qed. End Lemmas. @@ -353,30 +422,29 @@ Module Exports. Coercion sort : type >-> Sortclass. Notation union_map_class := type. Notation UMCMixin := Mixin. -Notation UMC T m := (@pack _ _ T _ m). +Notation UMC T m := (@pack _ _ _ T m). -Notation "[ 'umcMixin' 'of' T ]" := (mixin (class _ _ _ : class_of T)) +Notation "[ 'umcMixin' 'of' T ]" := (class _ _ _ _ : mixin_of T) (at level 0, format "[ 'umcMixin' 'of' T ]") : form_scope. -Notation "[ 'um_class' 'of' T 'for' C ]" := (@clone _ _ T C _ id) +Notation "[ 'um_class' 'of' T 'for' C ]" := (@clone _ _ _ T C _ idfun) (at level 0, format "[ 'um_class' 'of' T 'for' C ]") : form_scope. -Notation "[ 'um_class' 'of' T ]" := (@clone _ _ T _ _ id) +Notation "[ 'um_class' 'of' T ]" := (@clone _ _ _ T _ _ id) (at level 0, format "[ 'um_class' 'of' T ]") : form_scope. -Notation cond := cond. Notation um_undef := um_undef. Notation upd := upd. Notation dom := dom. Notation dom_eq := dom_eq. +Notation assocs := assocs. Notation free := free. Notation find := find. -Notation um_filter := um_filter. Notation empb := empb. Notation undefb := undefb. Notation pts := pts. Definition umEX := - (ftE, tfE, eqE, undefE, defE, emptyE, updE, domE, dom_eqE, - freeE, findE, unionE, um_filterE, empbE, undefbE, ptsE, UM.umapE). + (ftE, tfE, eqE, undefE, defE, emptyE, updE, domE, dom_eqE, assocsE, + freeE, findE, unionE, empbE, undefbE, ptsE, UM.umapE). End Exports. @@ -391,13 +459,13 @@ Export UMC.Exports. Module UnionMapClassPCM. Section UnionMapClassPCM. -Variables (K : ordType) (V : Type) (U : union_map_class K V). +Variables (K : ordType) (C : pred K) (V : Type) (U : union_map_class C V). Implicit Type f : U. -Local Notation "f1 \+ f2" := (@UMC.union _ _ _ f1 f2) +Local Notation "f1 \+ f2" := (@UMC.union _ _ _ _ f1 f2) (at level 43, left associativity). -Local Notation valid := (@UMC.defined _ _ U). -Local Notation unit := (@UMC.empty _ _ U). +Local Notation valid := (@UMC.defined _ _ _ U). +Local Notation unit := (@UMC.empty _ _ _ U). Lemma joinC f1 f2 : f1 \+ f2 = f2 \+ f1. Proof. @@ -436,7 +504,7 @@ End UnionMapClassPCM. Module Exports. Section Exports. -Variables (K : ordType) (V : Type) (U : union_map_class K V). +Variables (K : ordType) (C : pred K) (V : Type) (U : union_map_class C V). (* generic structure for PCM for *all* union maps *) (* I will add specific ones too for individual types *) @@ -444,8 +512,8 @@ Variables (K : ordType) (V : Type) (U : union_map_class K V). (* not against another projection, as is the case *) (* with the generic class here *) Definition union_map_classPCMMix := - PCMMixin (@joinC K V U) (@joinA K V U) (@unitL K V U) - (@validL K V U) (validU U). + PCMMixin (@joinC K C V U) (@joinA K C V U) (@unitL K C V U) + (@validL K C V U) (validU U). Canonical union_map_classPCM := Eval hnf in PCM U union_map_classPCMMix. End Exports. @@ -462,7 +530,7 @@ Export UnionMapClassPCM.Exports. Module UnionMapClassCPCM. Section Cancelativity. -Variables (K : ordType) (V : Type) (U : union_map_class K V). +Variables (K : ordType) (C : pred K) (V : Type) (U : union_map_class C V). Implicit Type f : U. Lemma joinKf f1 f2 f : valid (f1 \+ f) -> f1 \+ f = f2 \+ f -> f1 = f2. @@ -478,9 +546,9 @@ End Cancelativity. Module Exports. Section Exports. -Variables (K : ordType) (V : Type) (U : union_map_class K V). +Variables (K : ordType) (C : pred K) (V : Type) (U : union_map_class C V). -Definition union_map_classCPCMMix := CPCMMixin (@joinKf K V U). +Definition union_map_classCPCMMix := CPCMMixin (@joinKf K C V U). Canonical union_map_classCPCM := Eval hnf in CPCM U union_map_classCPCMMix. End Exports. @@ -490,14 +558,53 @@ End UnionMapClassCPCM. Export UnionMapClassCPCM.Exports. + +(********************) +(* Topped structure *) +(********************) + +Module UnionClassTPCM. +Section UnionClassTPCM. +Variables (K : ordType) (C : pred K) (V : Type) (U : union_map_class C V). +Implicit Type f : U. + +Lemma empbP f : reflect (f = Unit) (empb f). +Proof. +rewrite pcmE /= !umEX /UM.empty /UM.empb -{1}[f]UMC.tfE. +case: (UMC.from f)=>[|f' F]; first by apply: ReflectF; rewrite !umEX. +case: eqP=>E; constructor; rewrite !umEX. +- by move/supp_nilE: E=>->. +by move=>H; rewrite H in E. +Qed. + +Lemma valid_undefN : ~~ valid (um_undef: U). +Proof. by rewrite pcmE /= !umEX. Qed. + +Lemma undef_join f : um_undef \+ f = um_undef. +Proof. by rewrite pcmE /= !umEX. Qed. + +End UnionClassTPCM. + +Module Exports. + +Definition union_map_classTPCMMix K C V (U : union_map_class C V) := + TPCMMixin (@empbP K C V U) (@valid_undefN K C V U) (@undef_join _ _ _ U). +Canonical union_map_classTPCM K C V (U : union_map_class C V) := + Eval hnf in TPCM _ (@union_map_classTPCMMix K C V U). +End Exports. +End UnionClassTPCM. + +Export UnionClassTPCM.Exports. + + (*********************************************************) (* if V is an equality type, then union_map_class is too *) (*********************************************************) Module UnionMapEq. Section UnionMapEq. -Variables (K : ordType) (V : eqType) (p : pred K). -Variables (T : Type) (m : @UMC.mixin_of K V p T). +Variables (K : ordType) (C : pred K) (V : eqType). +Variables (T : Type) (m : @UMC.mixin_of K C V T). Definition unionmap_eq (f1 f2 : UMC T m) := match (UMC.from f1), (UMC.from f2) with @@ -511,9 +618,9 @@ Proof. move=>x y; rewrite -{1}[x]UMC.tfE -{1}[y]UMC.tfE /unionmap_eq. case: (UMC.from x)=>[|f1 H1]; case: (UMC.from y)=>[|f2 H2] /=. - by constructor. -- by constructor; move/(@UMC.eqE _ _ (UMC T m)). -- by constructor; move/(@UMC.eqE _ _ (UMC T m)). -case: eqP=>E; constructor; rewrite (@UMC.eqE _ _ (UMC T m)). +- by constructor; move/(@UMC.eqE _ _ _ (UMC T m)). +- by constructor; move/(@UMC.eqE _ _ _ (UMC T m)). +case: eqP=>E; constructor; rewrite (@UMC.eqE _ _ _ (UMC T m)). - by rewrite UM.umapE. by case. Qed. @@ -522,10 +629,10 @@ End UnionMapEq. Module Exports. Section Exports. -Variables (K : ordType) (V : eqType) (p : pred K). -Variables (T : Type) (m : @UMC.mixin_of K V p T). +Variables (K : ordType) (C : pred K) (V : eqType). +Variables (T : Type) (m : @UMC.mixin_of K C V T). -Definition union_map_class_eqMix := EqMixin (@unionmap_eqP K V p T m). +Definition union_map_class_eqMix := EqMixin (@unionmap_eqP K C V T m). End Exports. End Exports. @@ -540,11 +647,11 @@ Export UnionMapEq.Exports. (*******) Section DomLemmas. -Variables (K : ordType) (V : Type) (U : union_map_class K V). +Variables (K : ordType) (C : pred K) (V : Type) (U : union_map_class C V). Implicit Types (k : K) (v : V) (f g : U). -Lemma dom_undef : dom (um_undef : U) = [::]. -Proof. by rewrite !umEX. Qed. +Lemma dom_undef : dom (undef : U) = [::]. +Proof. by rewrite /undef /= !umEX. Qed. Lemma dom0 : dom (Unit : U) = [::]. Proof. by rewrite pcmE /= !umEX. Qed. @@ -558,9 +665,9 @@ Qed. Lemma domU k v f : dom (upd k v f) =i - [pred x | cond U k & if x == k then valid f else x \in dom f]. + [pred x | C k & if x == k then valid f else x \in dom f]. Proof. -rewrite pcmE /= !umEX /UM.dom /UM.upd /UM.valid /= /cond. +rewrite pcmE /= !umEX /UM.dom /UM.upd /UM.valid /=. case: (UMC.from f)=>[|f' H] /= x. - by rewrite !inE /= andbC; case: ifP. case: decP; first by move=>->; rewrite supp_ins. @@ -568,7 +675,7 @@ by move/(introF idP)=>->. Qed. Lemma domF k f : - dom (free k f) =i + dom (free f k) =i [pred x | if k == x then false else x \in dom f]. Proof. rewrite !umEX; case: (UMC.from f)=>[|f' H] x; rewrite inE /=; @@ -587,9 +694,12 @@ Qed. Lemma dom_valid k f : k \in dom f -> valid f. Proof. by rewrite pcmE /= !umEX; case: (UMC.from f). Qed. -Lemma dom_cond k f : k \in dom f -> cond U k. +Lemma dom_cond k f : k \in dom f -> C k. Proof. by rewrite !umEX; case: (UMC.from f)=>[|f' F] // /(allP F). Qed. +Lemma cond_dom k f : ~~ C k -> k \in dom f = false. +Proof. by apply: contraTF=>/dom_cond ->. Qed. + Lemma dom_inIL k f1 f2 : valid (f1 \+ f2) -> k \in dom f1 -> k \in dom (f1 \+ f2). Proof. by rewrite domUn inE => ->->. Qed. @@ -606,7 +716,7 @@ Lemma dom_NNR k f1 f2 : valid (f1 \+ f2) -> k \notin dom (f1 \+ f2) -> k \notin dom f2. Proof. by move=> D; apply/contra/dom_inIR. Qed. -Lemma dom_free k f : k \notin dom f -> free k f = f. +Lemma dom_free k f : k \notin dom f -> free f k = f. Proof. rewrite -{3}[f]UMC.tfE !umEX /UM.dom /UM.free. by case: (UMC.from f)=>[|f' H] //; apply: rem_supp. @@ -615,7 +725,7 @@ Qed. Variant dom_find_spec k f : bool -> Type := | dom_find_none' : find k f = None -> dom_find_spec k f false | dom_find_some' v : find k f = Some v -> - f = upd k v (free k f) -> dom_find_spec k f true. + f = upd k v (free f k) -> dom_find_spec k f true. Lemma dom_find k f : dom_find_spec k f (k \in dom f). Proof. @@ -631,17 +741,13 @@ by case: ifP=>[/eqP-> //|]; rewrite fnd_rem => ->. Qed. Lemma find_some k v f : find k f = Some v -> k \in dom f. -Proof. by case: dom_find=>// ->. Qed. +Proof. by case: dom_find =>// ->. Qed. -Lemma find_none k f : find k f = None -> k \notin dom f. +Lemma find_none k f : find k f = None <-> k \notin dom f. Proof. by case: dom_find=>// v ->. Qed. -Lemma dom_umfilt p f : dom (um_filter p f) =i [predI p & dom f]. -Proof. -rewrite !umEX /UM.dom /UM.um_filter. -case: (UMC.from f)=>[|f' H] x; first by rewrite !inE /= andbF. -by rewrite supp_kfilt mem_filter. -Qed. +Lemma cond_find k f : ~~ C k -> find k f = None. +Proof. by move/(cond_dom f); case: dom_find. Qed. Lemma dom_prec f1 f2 g1 g2 : valid (f1 \+ g1) -> @@ -677,29 +783,21 @@ case _ : (x \in supp f1') => //= in D1 D2 *. by move/negbTE: D1=>->; move/negbTE: D2=>->. Qed. -Lemma umfilt_dom f1 f2 : - valid (f1 \+ f2) -> um_filter (mem (dom f1)) (f1 \+ f2) = f1. -Proof. -rewrite -{4}[f1]UMC.tfE !pcmE /= !umEX. -rewrite /UM.valid /UM.union /UM.um_filter /UM.dom. -case: (UMC.from f1) (UMC.from f2)=>[|f1' F1][|f2' F2] //. -case: ifP=>// D _; rewrite kfilt_fcat /=. -have E1: {in supp f1', supp f1' =i predT} by []. -have E2: {in supp f2', supp f1' =i pred0}. -- by move=>x; rewrite disjC in D; case: disjP D=>// D _ /D /negbTE ->. -rewrite (eq_in_kfilter E1) (eq_in_kfilter E2). -by rewrite kfilter_predT kfilter_pred0 fcats0. -Qed. - Lemma sorted_dom f : sorted (@ord K) (dom f). Proof. by rewrite !umEX; case: (UMC.from f)=>[|[]]. Qed. +Lemma sorted_dom_oleq f : sorted (@oleq K) (dom f). +Proof. by apply: sorted_oleq (sorted_dom f). Qed. + Lemma uniq_dom f : uniq (dom f). Proof. apply: sorted_uniq (sorted_dom f); by [apply: ordtype.trans | apply: ordtype.irr]. Qed. +Lemma all_dom f : all C (dom f). +Proof. by rewrite !umEX; case: (UMC.from f). Qed. + Lemma perm_domUn f1 f2 : valid (f1 \+ f2) -> perm_eq (dom (f1 \+ f2)) (dom f1 ++ dom f2). Proof. @@ -715,146 +813,64 @@ Qed. Lemma size_domUn f1 f2 : valid (f1 \+ f2) -> size (dom (f1 \+ f2)) = size (dom f1) + size (dom f2). -Proof. by move/perm_domUn/perm_size; rewrite size_cat. Qed. - -End DomLemmas. - -Hint Resolve sorted_dom uniq_dom : core. -Prenex Implicits find_some find_none. - - -(**********) -(* filter *) -(**********) - -Section FilterLemmas. -Variables (K : ordType) (V : Type) (U : union_map_class K V). -Implicit Type f : U. - -Lemma eq_in_umfilt p1 p2 f : - {in dom f, p1 =1 p2} -> um_filter p1 f = um_filter p2 f. -Proof. -rewrite !umEX /UM.dom /UM.um_filter /= /dom. -by case: (UMC.from f)=>[|f' H] //=; apply: eq_in_kfilter. -Qed. - -Lemma umfilt0 q : um_filter q Unit = Unit :> U. -Proof. by rewrite !pcmE /= !umEX /UM.um_filter /UM.empty kfilt_nil. Qed. - -Lemma umfilt_undef q : um_filter q um_undef = um_undef :> U. -Proof. by rewrite !umEX. Qed. - -Lemma umfiltUn q f1 f2 : - valid (f1 \+ f2) -> - um_filter q (f1 \+ f2) = um_filter q f1 \+ um_filter q f2. -Proof. -rewrite !pcmE /= !umEX /UM.valid /UM.union. -case: (UMC.from f1)=>[|f1' H1]; case: (UMC.from f2)=>[|f2' H2] //=. -by case: ifP=>D //= _; rewrite kfilt_fcat disj_kfilt. -Qed. - -Lemma umfilt_pred0 f : valid f -> um_filter pred0 f = Unit. -Proof. -rewrite !pcmE /= !umEX /UM.valid /UM.empty. -case: (UMC.from f)=>[|f' H] //= _; case: f' H=>f' P H. -by rewrite fmapE /= /kfilter' filter_pred0. -Qed. - -Lemma umfilt_predT f : um_filter predT f = f. -Proof. -rewrite -[f]UMC.tfE !umEX /UM.um_filter. -by case: (UMC.from f)=>[|f' H] //; rewrite kfilter_predT. -Qed. - -Lemma umfilt_predI p1 p2 f : - um_filter (predI p1 p2) f = um_filter p1 (um_filter p2 f). -Proof. -rewrite -[f]UMC.tfE !umEX /UM.um_filter. -by case: (UMC.from f)=>[|f' H] //; rewrite kfilter_predI. -Qed. +Proof. by move/perm_domUn/seq.perm_size; rewrite size_cat. Qed. -Lemma umfilt_predU p1 p2 f : - um_filter (predU p1 p2) f = - um_filter p1 f \+ um_filter (predD p2 p1) f. +Lemma size_domF k f : + k \in dom f -> size (dom (free f k)) = (size (dom f)).-1. Proof. -rewrite pcmE /= !umEX /UM.um_filter /UM.union /=. -case: (UMC.from f)=>[|f' H] //=. -rewrite in_disj_kfilt; last by move=>x _; rewrite /= negb_and orbA orbN. -rewrite -kfilter_predU. -by apply: eq_in_kfilter=>x _; rewrite /= orb_andr orbN. +rewrite !umEX; case: (UMC.from f)=>[|f'] //=; case: f'=>f' F /= _. +rewrite /supp /= !size_map size_filter=>H. +move/(sorted_uniq (@trans K) (@irr K))/count_uniq_mem: F=>/(_ k). +rewrite {}H count_map /ssrbool.preim /= => H. +by rewrite -(count_predC [pred x | key x == k]) H addnC addn1. Qed. -Lemma umfilt_dpredU f p q : - {subset p <= predC q} -> - um_filter (predU p q) f = um_filter p f \+ um_filter q f. +Lemma size_domU k v f : + valid (upd k v f) -> + size (dom (upd k v f)) = + if k \in dom f then size (dom f) else (size (dom f)).+1. Proof. -move=>D; rewrite umfilt_predU (eq_in_umfilt (p1:=predD q p) (p2:=q)) //. -by move=>k _ /=; case X : (p k)=>//=; move/D/negbTE: X. +rewrite pcmE /= !umEX /UM.valid /UM.upd /UM.dom. +case: (UMC.from f)=>[|f'] //= H; case: decP=>// P _. +by case: f' H=>f' F H; rewrite /supp /= !size_map size_ins'. Qed. -Lemma umfiltUnK p f1 f2 : - valid (f1 \+ f2) -> - um_filter p (f1 \+ f2) = f1 -> - um_filter p f1 = f1 /\ um_filter p f2 = Unit. -Proof. -move=>V'; rewrite (umfiltUn _ V') => E. -have {V'}-V' : valid (um_filter p f1 \+ um_filter p f2). -- by rewrite E; move/validL: V'. -have F : dom (um_filter p f1) =i dom f1. -- move=>x; rewrite dom_umfilt inE /=. - apply/andP/idP=>[[//]| H1]; split=>//; move: H1. - rewrite -{1}E domUn inE /= !dom_umfilt inE /= inE /=. - by case: (x \in p)=>//=; rewrite andbF. -rewrite -{2}[f1]unitR in E F; move/(dom_prec V' E): F=>X. -by rewrite X in E V' *; rewrite (joinxK V' E). -Qed. +End DomLemmas. -Lemma umfiltU p k v f : - um_filter p (upd k v f) = - if p k then upd k v (um_filter p f) else - if cond U k then um_filter p f else um_undef. -Proof. -rewrite !umEX /UM.um_filter /UM.upd /cond. -case: (UMC.from f)=>[|f' F]; first by case: ifP=>H1 //; case: ifP. -case: ifP=>H1; case: decP=>H2 //=. -- by rewrite !umEX kfilt_ins H1. -- by rewrite H2 !umEX kfilt_ins H1. -by case: ifP H2. -Qed. +Hint Resolve sorted_dom uniq_dom all_dom : core. +Prenex Implicits find_some find_none. -Lemma umfiltF p k f : - um_filter p (free k f) = - if p k then free k (um_filter p f) else um_filter p f. +(* when we compare doms of two differently-typed maps *) +Lemma domE K C V1 V2 (U1 : @union_map_class K C V1) (U2 : @union_map_class K C V2) + (f1 : U1) (f2 : U2) : + dom f1 =i dom f2 <-> dom f1 = dom f2. Proof. -rewrite !umEX /UM.um_filter /UM.free. -by case: (UMC.from f)=>[|f' F]; case: ifP=>// E; rewrite !umEX kfilt_rem E. +split=>[|->] //. +apply: (@eq_sorted_irr _ ord); +by [apply: trans|apply: irr|apply: sorted_dom|apply: sorted_dom]. Qed. -End FilterLemmas. - - (*********) (* valid *) (*********) Section ValidLemmas. -Variables (K : ordType) (V : Type) (U : union_map_class K V). +Variables (K : ordType) (C : pred K) (V : Type) (U : union_map_class C V). Implicit Types (k : K) (v : V) (f g : U). -Lemma invalidE f : ~~ valid f <-> f = um_undef. -Proof. by rewrite !pcmE /= !umEX -2![f]UMC.tfE !umEX; case: (UMC.from f). Qed. - -Lemma valid_undef : valid (um_undef : U) = false. -Proof. by apply/negbTE; apply/invalidE. Qed. +Lemma invalidE f : ~~ valid f <-> f = undef. +Proof. +by rewrite /undef !pcmE /= !umEX -2![f]UMC.tfE !umEX; case: (UMC.from f). +Qed. -Lemma validU k v f : valid (upd k v f) = cond U k && valid f. +Lemma validU k v f : valid (upd k v f) = C k && valid f. Proof. -rewrite !pcmE /= !umEX /UM.valid /UM.upd /cond. +rewrite !pcmE /= !umEX /UM.valid /UM.upd. case: (UMC.from f)=>[|f' F]; first by rewrite andbF. by case: decP=>[|/(introF idP)] ->. Qed. -Lemma validF k f : valid (free k f) = valid f. +Lemma validF k f : valid (free f k) = valid f. Proof. by rewrite !pcmE /= !umEX; case: (UMC.from f). Qed. Variant validUn_spec f1 f2 : bool -> Type := @@ -879,8 +895,36 @@ case: disjP E=>// k T1 T2 _. by apply: (valid_false3 (k:=k)); rewrite !umEX. Qed. +Lemma validUnAE f1 f2 : + valid (f1 \+ f2) = + [&& valid f1, valid f2 & all [predC dom f1] (dom f2)]. +Proof. +apply/idP/idP. +- case: validUn=>// ->-> H _; apply/allP=>k. + by apply: contraL (H _). +case/and3P=>V1 V2 /allP /= D; case: validUn=>//. +- by rewrite V1. +- by rewrite V2. +by move=>k X1 X2; move: (D k X2); rewrite X1. +Qed. + +Lemma validUnAEC f1 f2 : + valid (f1 \+ f2) = + [&& valid f1, valid f2 & all [predC dom f2] (dom f1)]. +Proof. by rewrite validUnAE all_predC_sym. Qed. + +Lemma validUnHE f1 f2 : + valid (f1 \+ f2) = + [&& valid f1, valid f2 & ~~ has [mem dom f1] (dom f2)]. +Proof. by rewrite validUnAE all_predC. Qed. + +Lemma validUnHC f1 f2 : + valid (f1 \+ f2) = + [&& valid f1, valid f2 & ~~ has [mem dom f2] (dom f1)]. +Proof. by rewrite validUnHE has_sym. Qed. + Lemma validFUn k f1 f2 : - valid (f1 \+ f2) -> valid (free k f1 \+ f2). + valid (f1 \+ f2) -> valid (free f1 k \+ f2). Proof. case: validUn=>// V1 V2 H _; case: validUn=>//; last 1 first. - by move=>k'; rewrite domF inE; case: eqP=>// _ /H/negbTE ->. @@ -889,7 +933,7 @@ by rewrite V2. Qed. Lemma validUnF k f1 f2 : - valid (f1 \+ f2) -> valid (f1 \+ free k f2). + valid (f1 \+ f2) -> valid (f1 \+ free f2 k). Proof. by rewrite !(joinC f1); apply: validFUn. Qed. Lemma validUnU k v f1 f2 : @@ -912,9 +956,6 @@ Lemma validUUn k v f1 f2 : k \in dom f1 -> valid (upd k v f1 \+ f2) = valid (f1 \+ f2). Proof. by move=>D; rewrite -!(joinC f2); apply: validUnU D. Qed. -Lemma valid_umfilt p f : valid (um_filter p f) = valid f. -Proof. by rewrite !pcmE /= !umEX; case: (UMC.from f). Qed. - Lemma dom_inNL k f1 f2 : valid (f1 \+ f2) -> k \in dom f1 -> k \notin dom f2. Proof. by case: validUn=>//_ _ H _; apply: H. Qed. @@ -926,12 +967,95 @@ Proof. by rewrite joinC; apply: dom_inNL. Qed. End ValidLemmas. +(************) +(* um_unitb *) +(************) + +(* AKA empb *) + +Section UnitbLemmas. +Variables (K : ordType) (C : pred K) (V : Type) (U : union_map_class C V). +Implicit Types (k : K) (v : V) (f : U). + +Lemma um_unitbU k v f : unitb (upd k v f) = false. +Proof. +rewrite /unitb /= !umEX /UM.empb /UM.upd. +case: (UMC.from f)=>[|f' F] //; case: decP=>// H. +suff: k \in supp (ins k v f') by case: (supp _). +by rewrite supp_ins inE /= eq_refl. +Qed. + +Lemma um_unitbUn f1 f2 : unitb (f1 \+ f2) = unitb f1 && unitb f2. +Proof. +rewrite /unitb !pcmE /= !umEX /UM.empb /UM.union. +case: (UMC.from f1) (UMC.from f2)=>[|f1' F1][|f2' F2] //. +- by rewrite andbF. +case: ifP=>E; case: eqP=>E1; case: eqP=>E2 //; last 2 first. +- by move: E2 E1; move/supp_nilE=>->; rewrite fcat0s; case: eqP. +- by move: E1 E2 E; do 2![move/supp_nilE=>->]; case: disjP. +- by move/supp_nilE: E2 E1=>-> <-; rewrite fcat0s eq_refl. +have [k H3]: exists k, k \in supp f1'. +- case: (supp f1') {E1 E} E2=>[|x s _] //. + by exists x; rewrite inE eq_refl. +suff: k \in supp (fcat f1' f2') by rewrite E1. +by rewrite supp_fcat inE /= H3. +Qed. + +(* some transformation lemmas *) + +(* AKA conicity *) +Lemma join0E f1 f2 : f1 \+ f2 = Unit <-> f1 = Unit /\ f2 = Unit. +Proof. by rewrite !unitbE um_unitbUn; case: andP. Qed. + +Lemma validEb f : reflect (valid f /\ forall k, k \notin dom f) (unitb f). +Proof. +case: unitbP=>E; constructor; first by rewrite E valid_unit dom0. +case=>V' H; apply: E; move: V' H. +rewrite !pcmE /= !umEX /UM.valid /UM.dom /UM.empty -{3}[f]UMC.tfE. +case: (UMC.from f)=>[|f' F] // _ H; rewrite !umEX. +apply/supp_nilE; case: (supp f') H=>// x s /(_ x). +by rewrite inE eq_refl. +Qed. + +Lemma validUnEb f : valid (f \+ f) = unitb f. +Proof. +case E: (unitb f); first by move/unitbE: E=>->; rewrite unitL valid_unit. +case: validUn=>// V' _ L; case: validEb E=>//; case; split=>// k. +by case E: (k \in dom f)=>//; move: (L k E); rewrite E. +Qed. + +End UnitbLemmas. + + +(**********) +(* undefb *) +(**********) + +Section UndefbLemmas. +Variables (K : ordType) (C : pred K) (V : Type) (U : union_map_class C V). +Implicit Types (k : K) (v : V) (f : U). + +Lemma undefb_undef : undefb (undef : U). +Proof. by rewrite /undef /= !umEX. Qed. + +Lemma undefbP f : reflect (f = undef) (undefb f). +Proof. +rewrite /undef /= !umEX /UM.undefb -{1}[f]UMC.tfE. +by case: (UMC.from f)=>[|f' F]; constructor; rewrite !umEX. +Qed. + +Lemma undefbE f : f = undef <-> undefb f. +Proof. by case: undefbP. Qed. + +End UndefbLemmas. + + (**********) (* dom_eq *) (**********) Section DomEqLemmas. -Variables (K : ordType) (V : Type) (U : union_map_class K V). +Variables (K : ordType) (C : pred K) (V : Type) (U : union_map_class C V). Implicit Type f : U. Lemma domeqP f1 f2 : @@ -1022,56 +1146,98 @@ Lemma domeqK f1 f2 f1' f2' : dom_eq f1 f2 -> dom_eq f1' f2'. Proof. move=>V1 /domeqP [E1 H1] /domeqP [E2 H2]; move: (V1); rewrite E1=>V2. -apply/domeqP; split; first by rewrite (validR V1) (validR V2). +apply/domeqP; split; first by rewrite (validE2 V1) (validE2 V2). by apply: domK V1 V2 H1 H2. Qed. +Lemma domeqfU k v f : k \in dom f -> dom_eq f (upd k v f). +Proof. +move=>D; apply/domeqP; rewrite validU (dom_cond D) (dom_valid D). +by split=>// x; rewrite domU inE (dom_cond D) (dom_valid D); case: eqP=>// ->. +Qed. + +Lemma domeqUf k v f : k \in dom f -> dom_eq (upd k v f) f. +Proof. by rewrite domeq_sym; apply: domeqfU. Qed. + End DomEqLemmas. Hint Resolve domeq_refl : core. -(*************************) -(* some precision lemmas *) -(*************************) +(***********) +(* dom_eq2 *) +(***********) -Section Precision. -Variables (K : ordType) (V : Type) (U : union_map_class K V). -Implicit Types (x y : U). +(* for maps of different types *) -Lemma prec_flip x1 x2 y1 y2 : - valid (x1 \+ y1) -> x1 \+ y1 = x2 \+ y2 -> - valid (y2 \+ x2) /\ y2 \+ x2 = y1 \+ x1. -Proof. by move=>X /esym E; rewrite joinC E X joinC. Qed. +Section DomEq2Def. +Variable (K : ordType) (C : pred K) (V1 V2 : Type). +Variable U1 : union_map_class C V1. +Variable U2 : union_map_class C V2. -Lemma prec_domV x1 x2 y1 y2 : - valid (x1 \+ y1) -> x1 \+ y1 = x2 \+ y2 -> - reflect {subset dom x1 <= dom x2} (valid (x1 \+ y2)). +Definition dom_eq2 (f1 : U1) (f2 : U2) := + (valid f1 == valid f2) && (dom f1 == dom f2). + +End DomEq2Def. + +Section DomEq2Lemmas. +Variable (K : ordType) (C : pred K) (V1 V2 V3 : Type). +Variable U1 : union_map_class C V1. +Variable U2 : union_map_class C V2. +Variable U3 : union_map_class C V3. + +Lemma domeq2_undef : dom_eq2 (undef : U1) (undef : U2). +Proof. by rewrite /dom_eq2 !valid_undef !dom_undef. Qed. + +Lemma domeq2P (f1 : U1) (f2 : U2) : + reflect (valid f1 = valid f2 /\ dom f1 = dom f2) (dom_eq2 f1 f2). Proof. -move=>V1 E; case V12 : (valid (x1 \+ y2)); constructor. -- move=>n Hs; have : n \in dom (x1 \+ y1) by rewrite domUn inE V1 Hs. - rewrite E domUn inE -E V1 /= orbC. - by case: validUn V12 Hs=>// _ _ L _ /L /negbTE ->. -move=>Hs; case: validUn V12=>//. -- by rewrite (validL V1). -- by rewrite E in V1; rewrite (validR V1). -by move=>k /Hs; rewrite E in V1; case: validUn V1=>// _ _ L _ /L /negbTE ->. +case E: (dom_eq2 f1 f2); constructor; first by case/andP: E=>/eqP -> /eqP ->. +by case=>V D; move/negbT/negP: E; apply; rewrite /dom_eq2 V D !eq_refl. Qed. -Lemma prec_filtV x1 x2 y1 y2 : - valid (x1 \+ y1) -> x1 \+ y1 = x2 \+ y2 -> - reflect (x1 = um_filter (mem (dom x1)) x2) (valid (x1 \+ y2)). +Lemma domeq2_refl (f : U1) : dom_eq2 f f. +Proof. by rewrite /dom_eq2 !eq_refl. Qed. + +Lemma domeq2_sym (f1 : U1) (f2 : U2) : dom_eq2 f1 f2 -> dom_eq2 f2 f1. +Proof. by case/andP=>V D; apply/andP; rewrite (eqP V) (eqP D). Qed. + +Lemma domeq2_trans (f1 : U1) (f2 : U2) (f3 : U3) : + dom_eq2 f1 f2 -> dom_eq2 f2 f3 -> dom_eq2 f1 f3. +Proof. by rewrite /dom_eq2=>/andP [/eqP -> /eqP ->]. Qed. + +Lemma domeq2VUnE (f1 f1' : U1) (f2 f2' : U2) : + dom_eq2 f1 f2 -> dom_eq2 f1' f2' -> + valid (f1 \+ f1') = valid (f2 \+ f2'). Proof. -move=>V1 E; case X : (valid (x1 \+ y2)); constructor; last first. -- case: (prec_domV V1 E) X=>// St _ H; apply: St. - by move=>n; rewrite H dom_umfilt inE; case/andP. -move: (umfilt_dom V1); rewrite E umfiltUn -?E //. -rewrite (eq_in_umfilt (f:=y2) (p2:=pred0)). -- by rewrite umfilt_pred0 ?unitR //; rewrite E in V1; rewrite (validR V1). -by move=>n; case: validUn X=>// _ _ L _ /(contraL (L _)) /negbTE. +suff L V1' V2' (U1' : union_map_class C V1') + (U2' : union_map_class C V2') (g1 g1' : U1') (g2 g2' : U2') : + dom_eq2 g1 g2 -> dom_eq2 g1' g2' -> + valid (g1 \+ g1') -> valid (g2 \+ g2'). +- by move=>D D'; apply/idP/idP; apply: L=>//; apply: domeq2_sym. +case/andP=>/eqP E /eqP D /andP [/eqP E' /eqP D']; case: validUn=>// V V' G _. +case: validUn=>//; rewrite -?E ?V -?E' ?V' //. +by move=>k; rewrite -D -D'=>/G/negbTE ->. Qed. -End Precision. +Lemma domeq2VUn (f1 f1' : U1) (f2 f2' : U2) : + dom_eq2 f1 f2 -> dom_eq2 f1' f2' -> + valid (f1 \+ f1') -> valid (f2 \+ f2'). +Proof. by move=>D D'; rewrite -(domeq2VUnE D D'). Qed. + +Lemma domeq2Un (f1 f1' : U1) (f2 f2' : U2) : + dom_eq2 f1 f2 -> dom_eq2 f1' f2' -> + dom_eq2 (f1 \+ f1') (f2 \+ f2'). +Proof. +move=>D D'; case/andP: (D)=>V E; case/andP: (D')=>V' E'. +apply/andP; rewrite (domeq2VUnE D D') eq_refl; split=>//. +apply/eqP; apply: eq_sorted_ord; try by apply: sorted_dom. +by move=>k; rewrite !domUn !inE (domeq2VUnE D D') (eqP E) (eqP E'). +Qed. + +End DomEq2Lemmas. + +Hint Resolve domeq2_refl : core. (**********) @@ -1079,16 +1245,23 @@ End Precision. (**********) Section UpdateLemmas. -Variable (K : ordType) (V : Type) (U : union_map_class K V). +Variable (K : ordType) (C : pred K) (V : Type) (U : union_map_class C V). Implicit Types (k : K) (v : V) (f : U). -Lemma upd_invalid k v : upd k v um_undef = um_undef :> U. -Proof. by rewrite !umEX. Qed. +Lemma upd_undef k v : upd k v undef = undef :> U. +Proof. by rewrite /undef /= !umEX. Qed. + +Lemma upd_condN k v f : ~~ C k -> upd k v f = undef. +Proof. +rewrite /undef /= !umEX /UM.upd. +case: (UMC.from f)=>[|f' H']=>//. +by case: decP=>//= ->. +Qed. Lemma upd_inj k v1 v2 f : - valid f -> cond U k -> upd k v1 f = upd k v2 f -> v1 = v2. + valid f -> C k -> upd k v1 f = upd k v2 f -> v1 = v2. Proof. -rewrite !pcmE /= !umEX /UM.valid /UM.upd /cond. +rewrite !pcmE /= !umEX /UM.valid /UM.upd. case: (UMC.from f)=>[|f' F] // _; case: decP=>// H _ E. have: fnd k (ins k v1 f') = fnd k (ins k v2 f') by rewrite E. by rewrite !fnd_ins eq_refl; case. @@ -1107,8 +1280,8 @@ by rewrite ins_ins E. Qed. Lemma updF k1 k2 v f : - upd k1 v (free k2 f) = - if k1 == k2 then upd k1 v f else free k2 (upd k1 v f). + upd k1 v (free f k2) = + if k1 == k2 then upd k1 v f else free (upd k1 v f) k2. Proof. rewrite !umEX /UM.dom /UM.upd /UM.free. case: (UMC.from f)=>[|f' F] //; case: ifP=>// H1; @@ -1135,6 +1308,14 @@ Lemma updUnR k v f1 f2 : if k \in dom f2 then f1 \+ upd k v f2 else upd k v f1 \+ f2. Proof. by rewrite joinC updUnL (joinC f1) (joinC f2). Qed. +Lemma updL k v f1 f2 : + k \in dom f1 -> upd k v (f1 \+ f2) = upd k v f1 \+ f2. +Proof. by move=>D; rewrite updUnL D. Qed. + +Lemma updR k v f1 f2 : + k \in dom f2 -> upd k v (f1 \+ f2) = f1 \+ upd k v f2. +Proof. by move=>D; rewrite updUnR D. Qed. + End UpdateLemmas. @@ -1143,19 +1324,22 @@ End UpdateLemmas. (********) Section FreeLemmas. -Variables (K : ordType) (V : Type) (U : union_map_class K V). +Variables (K : ordType) (C : pred K) (V : Type) (U : union_map_class C V). Implicit Types (k : K) (v : V) (f : U). -Lemma free0 k : free k Unit = Unit :> U. +Lemma free0 k : free Unit k = Unit :> U. Proof. by rewrite !pcmE /= !umEX /UM.free /UM.empty rem_empty. Qed. +Lemma free_undef k : free undef k = undef :> U. +Proof. by rewrite /undef /= !umEX. Qed. + Lemma freeU k1 k2 v f : - free k1 (upd k2 v f) = + free (upd k2 v f) k1 = if k1 == k2 then - if cond U k2 then free k1 f else um_undef - else upd k2 v (free k1 f). + if C k2 then free f k1 else undef + else upd k2 v (free f k1). Proof. -rewrite !umEX /UM.free /UM.upd /cond. +rewrite /undef /= !umEX /UM.free /UM.upd. case: (UMC.from f)=>[|f' F]; first by case: ifP=>H1 //; case: ifP. case: ifP=>H1; case: decP=>H2 //=. - by rewrite H2 !umEX rem_ins H1. @@ -1164,16 +1348,16 @@ by rewrite !umEX rem_ins H1. Qed. Lemma freeF k1 k2 f : - free k1 (free k2 f) = - if k1 == k2 then free k1 f else free k2 (free k1 f). + free (free f k2) k1 = + if k1 == k2 then free f k1 else free (free f k1) k2. Proof. rewrite !umEX /UM.free. by case: (UMC.from f)=>[|f' F]; case: ifP=>// E; rewrite !umEX rem_rem E. Qed. Lemma freeUn k f1 f2 : - free k (f1 \+ f2) = - if k \in dom (f1 \+ f2) then free k f1 \+ free k f2 + free (f1 \+ f2) k = + if k \in dom (f1 \+ f2) then free f1 k \+ free f2 k else f1 \+ f2. Proof. rewrite !pcmE /= !umEX /UM.free /UM.union /UM.dom. @@ -1189,18 +1373,18 @@ by case: disjP E1 E2=>// H _; move/contra: (H k); rewrite negbK. Qed. Lemma freeUnV k f1 f2 : - valid (f1 \+ f2) -> free k (f1 \+ f2) = free k f1 \+ free k f2. + valid (f1 \+ f2) -> free (f1 \+ f2) k = free f1 k \+ free f2 k. Proof. move=>V'; rewrite freeUn domUn V' /=; case: ifP=>//. by move/negbT; rewrite negb_or; case/andP=>H1 H2; rewrite !dom_free. Qed. -Lemma freeUnL k f1 f2 : k \notin dom f1 -> free k (f1 \+ f2) = f1 \+ free k f2. +Lemma freeUnL k f1 f2 : k \notin dom f1 -> free (f1 \+ f2) k = f1 \+ free f2 k. Proof. move=>V1; rewrite freeUn domUn inE (negbTE V1) /=. case: ifP; first by case/andP; rewrite dom_free. move/negbT; rewrite negb_and; case/orP=>V2; last by rewrite dom_free. -suff: ~~ valid (f1 \+ free k f2) by move/invalidE: V2=>-> /invalidE ->. +suff: ~~ valid (f1 \+ free f2 k) by move/invalidE: V2=>-> /invalidE ->. apply: contra V2; case: validUn=>// H1 H2 H _. case: validUn=>//; first by rewrite H1. - by move: H2; rewrite validF => ->. @@ -1208,18 +1392,19 @@ move=>x H3; move: (H _ H3); rewrite domF inE /=. by case: ifP H3 V1=>[|_ _ _]; [move/eqP=><- -> | move/negbTE=>->]. Qed. -Lemma freeUnR k f1 f2 : k \notin dom f2 -> free k (f1 \+ f2) = free k f1 \+ f2. +Lemma freeUnR k f1 f2 : k \notin dom f2 -> free (f1 \+ f2) k = free f1 k \+ f2. Proof. by move=>H; rewrite joinC freeUnL // joinC. Qed. -Lemma free_umfilt p k f : - free k (um_filter p f) = - if p k then um_filter p (free k f) else um_filter p f. -Proof. -rewrite !umEX /UM.free /UM.um_filter. -case: (UMC.from f)=>[|f' F]; case: ifP=>// E; rewrite !umEX. -- by rewrite kfilt_rem E. -by rewrite rem_kfilt E. -Qed. +Lemma freeND k f : k \notin dom f -> free f k = f. +Proof. by move=>W; rewrite -[f]unitR freeUnL // free0. Qed. + +Lemma freeL k f1 f2 : + valid (f1 \+ f2) -> k \in dom f1 -> free (f1 \+ f2) k = free f1 k \+ f2. +Proof. by move=>W /(dom_inNL W) D; rewrite freeUnR. Qed. + +Lemma freeR k f1 f2 : + valid (f1 \+ f2) -> k \in dom f2 -> free (f1 \+ f2) k = f1 \+ free f2 k. +Proof. by rewrite !(joinC f1); apply: freeL. Qed. End FreeLemmas. @@ -1229,16 +1414,16 @@ End FreeLemmas. (********) Section FindLemmas. -Variables (K : ordType) (V : Type) (U : union_map_class K V). +Variables (K : ordType) (C : pred K) (V : Type) (U : union_map_class C V). Implicit Types (k : K) (v : V) (f : U). Lemma find0E k : find k (Unit : U) = None. Proof. by rewrite pcmE /= !umEX /UM.find /= fnd_empty. Qed. -Lemma find_undef k : find k (um_undef : U) = None. -Proof. by rewrite !umEX /UM.find. Qed. +Lemma find_undef k : find k (undef : U) = None. +Proof. by rewrite /undef /= !umEX /UM.find. Qed. -Lemma find_cond k f : ~~ cond U k -> find k f = None. +Lemma find_cond k f : ~~ C k -> find k f = None. Proof. simpl. rewrite !umEX /UM.find; case: (UMC.from f)=>[|f' F] // H. @@ -1249,17 +1434,17 @@ Qed. Lemma findU k1 k2 v f : find k1 (upd k2 v f) = if k1 == k2 then - if cond U k2 && valid f then Some v else None - else if cond U k2 then find k1 f else None. + if C k2 && valid f then Some v else None + else if C k2 then find k1 f else None. Proof. -rewrite pcmE /= !umEX /UM.valid /UM.find /UM.upd /cond. +rewrite pcmE /= !umEX /UM.valid /UM.find /UM.upd. case: (UMC.from f)=>[|f' F]; first by rewrite andbF !if_same. case: decP=>H; first by rewrite H /= fnd_ins. by rewrite (introF idP H) /= if_same. Qed. Lemma findF k1 k2 f : - find k1 (free k2 f) = if k1 == k2 then None else find k1 f. + find k1 (free f k2) = if k1 == k2 then None else find k1 f. Proof. rewrite !umEX /UM.find /UM.free. case: (UMC.from f)=>[|f' F]; first by rewrite if_same. @@ -1293,20 +1478,12 @@ case: (UMC.from f1) (UMC.from f2)=>[|f1' F1][|f2' F2] // _ _ H. by rewrite !umEX; apply/fmapP=>k; move: (H k); rewrite !umEX. Qed. -Lemma find_umfilt p k f : - find k (um_filter p f) = if p k then find k f else None. -Proof. -rewrite !umEX /UM.find /UM.um_filter. -case: (UMC.from f)=>[|f' F]; first by rewrite if_same. -by rewrite fnd_kfilt. -Qed. - End FindLemmas. (* if maps store units, i.e., we keep them just for sets *) (* then we can simplify the find_eta lemma a bit *) -Lemma domeq_eta (K : ordType) (U : union_map_class K unit) (f1 f2 : U) : +Lemma domeq_eta (K : ordType) (C : pred K) (U : union_map_class C unit) (f1 f2 : U) : dom_eq f1 f2 -> f1 = f2. Proof. case/domeqP=>V E; case V1 : (valid f1); last first. @@ -1317,183 +1494,107 @@ move=>k; case D : (k \in dom f1); move: D (D); rewrite {1}E. by do 2![case: dom_find=>// ->]. Qed. -(********) -(* empb *) -(********) -Section EmpbLemmas. -Variables (K : ordType) (V : Type) (U : union_map_class K V). -Implicit Types (k : K) (v : V) (f : U). -Lemma empb_undef : empb (um_undef : U) = false. -Proof. by rewrite !umEX. Qed. +(**********) +(* assocs *) +(**********) -Lemma empbP f : reflect (f = Unit) (empb f). -Proof. -rewrite pcmE /= !umEX /UM.empty /UM.empb -{1}[f]UMC.tfE. -case: (UMC.from f)=>[|f' F]; first by apply: ReflectF; rewrite !umEX. -case: eqP=>E; constructor; rewrite !umEX. -- by move/supp_nilE: E=>->. -by move=>H; rewrite H in E. -Qed. - -Lemma empbU k v f : empb (upd k v f) = false. -Proof. -rewrite !umEX /UM.empb /UM.upd. -case: (UMC.from f)=>[|f' F] //; case: decP=>// H. -suff: k \in supp (ins k v f') by case: (supp _). -by rewrite supp_ins inE /= eq_refl. -Qed. - -Lemma empbUn f1 f2 : empb (f1 \+ f2) = empb f1 && empb f2. -Proof. -rewrite !pcmE /= !umEX /UM.empb /UM.union. -case: (UMC.from f1) (UMC.from f2)=>[|f1' F1][|f2' F2] //. -- by rewrite andbF. -case: ifP=>E; case: eqP=>E1; case: eqP=>E2 //; last 2 first. -- by move: E2 E1; move/supp_nilE=>->; rewrite fcat0s; case: eqP. -- by move: E1 E2 E; do 2![move/supp_nilE=>->]; case: disjP. -- by move/supp_nilE: E2 E1=>-> <-; rewrite fcat0s eq_refl. -have [k H3]: exists k, k \in supp f1'. -- case: (supp f1') {E1 E} E2=>[|x s _] //. - by exists x; rewrite inE eq_refl. -suff: k \in supp (fcat f1' f2') by rewrite E1. -by rewrite supp_fcat inE /= H3. -Qed. - -(* some transformation lemmas *) +Section AssocsLemmas. +Variables (K : ordType) (C : pred K) (V : Type) (U : union_map_class C V). +Implicit Type f : U. -Lemma empbE f : f = Unit <-> empb f. -Proof. by case: empbP. Qed. +Lemma assocs_valid k f : k \In assocs f -> valid f. +Proof. by rewrite !pcmE /= !umEX; case: (UMC.from f). Qed. -Lemma empb0 : empb (Unit : U). -Proof. by apply/empbE. Qed. +Lemma assocs0 : assocs (Unit : U) = [::]. +Proof. by rewrite pcmE /= !umEX. Qed. -Lemma join0E f1 f2 : f1 \+ f2 = Unit <-> f1 = Unit /\ f2 = Unit. -Proof. by rewrite !empbE empbUn; case: andP. Qed. +Lemma assocs_undef : assocs (undef : U) = [::]. +Proof. by rewrite /undef /= !umEX. Qed. -Lemma validEb f : reflect (valid f /\ forall k, k \notin dom f) (empb f). +Lemma assocsPt k v : + assocs (pts k v : U) = + if C k then [:: (k, v)] else [::]. Proof. -case: empbP=>E; constructor; first by rewrite E valid_unit dom0. -case=>V' H; apply: E; move: V' H. -rewrite !pcmE /= !umEX /UM.valid /UM.dom /UM.empty -{3}[f]UMC.tfE. -case: (UMC.from f)=>[|f' F] // _ H; rewrite !umEX. -apply/supp_nilE; case: (supp f') H=>// x s /(_ x). -by rewrite inE eq_refl. +rewrite !umEX /UM.assocs/UM.pts /=. +by case E: decP=>[a|a] /=; [rewrite a | case: ifP]. Qed. -Lemma validUnEb f : valid (f \+ f) = empb f. +Lemma assocsUnPt f k v : + valid (f \+ pts k v) -> all (ord^~ k) (dom f) -> + assocs (f \+ pts k v) = rcons (assocs f) (k, v). Proof. -case E: (empb f); first by move/empbE: E=>->; rewrite unitL valid_unit. -case: validUn=>// V' _ L; case: validEb E=>//; case; split=>// k. -by case E: (k \in dom f)=>//; move: (L k E); rewrite E. +rewrite !pcmE /= !umEX /UM.assocs/UM.union/UM.pts/UM.dom/supp /=. +case: (UMC.from f)=>//=; case: decP=>//= H1 g H2; case: ifP=>//= _ _. +by case: allP=>//; case: g H2=>// s1 H2 H3 H4 _; apply: first_ins' H4. Qed. -Lemma empb_umfilt p f : empb f -> empb (um_filter p f). +Lemma assocsPtUn f k v : + valid (pts k v \+ f) -> all (ord k) (dom f) -> + assocs (pts k v \+ f) = (k, v) :: (assocs f). Proof. -rewrite !umEX /UM.empb /UM.um_filter. -case: (UMC.from f)=>[|f' F] //. -by move/eqP/supp_nilE=>->; rewrite kfilt_nil. +rewrite !pcmE /= !umEX /UM.assocs/UM.union/UM.pts/UM.dom/supp /=. +case: decP=>// H1; case: (UMC.from f)=>//= g H2. +rewrite /disj /= andbT; case: ifP=>//= D _ H3. +rewrite (fcat_inss _ (@nil K V) D) fcat0s. +case: g H2 D H3=>//= g P H2 D H3. +by apply: last_ins'; rewrite path_min_sorted //. Qed. -End EmpbLemmas. - - -(*********) -(* undef *) -(*********) - -Section UndefLemmas. -Variables (K : ordType) (V : Type) (U : union_map_class K V). -Implicit Types (k : K) (v : V) (f : U). - -Lemma undefb_undef : undefb (um_undef : U). -Proof. by rewrite !umEX. Qed. - -Lemma undefbP f : reflect (f = um_undef) (undefb f). +Lemma assocs_perm f1 f2 : + valid (f1 \+ f2) -> perm (assocs (f1 \+ f2)) (assocs f1 ++ assocs f2). Proof. -rewrite !umEX /UM.undefb -{1}[f]UMC.tfE. -by case: (UMC.from f)=>[|f' F]; constructor; rewrite !umEX. +rewrite !pcmE /= !umEX /UM.assocs/UM.union/UM.pts/UM.dom/supp /=. +case: (UMC.from f1)=>//= g1 H1; case: (UMC.from f2)=>//= g2 H2. +by case: ifP=>// D _; apply: seqof_fcat_perm D. Qed. -Lemma undefbE f : f = um_undef <-> undefb f. -Proof. by case: undefbP. Qed. - -Lemma join_undefL f : um_undef \+ f = um_undef. -Proof. by rewrite /PCM.join /= !umEX. Qed. - -Lemma join_undefR f : f \+ um_undef = um_undef. -Proof. by rewrite joinC join_undefL. Qed. - -End UndefLemmas. - - -(**********) -(* um_all *) -(**********) - -(* this is a defined notion; so it does not appear in the signature *) - -Section AllDefLemmas. -Variables (K : ordType) (V : Type) (U : union_map_class K V) (P : V -> Prop). -Implicit Types (k : K) (v : V) (f : U). - -Definition um_all f := forall k v, find k f = Some v -> P v. +Lemma assocs_dom f : dom f = map fst (assocs f). +Proof. by rewrite !umEX; case: (UMC.from f). Qed. -Lemma umall_undef : um_all um_undef. -Proof. by move=>k v; rewrite find_undef. Qed. +Lemma size_assocs f : size (assocs f) = size (dom f). +Proof. by rewrite assocs_dom size_map. Qed. -Hint Resolve umall_undef : core. - -Lemma umall0 : um_all Unit. -Proof. by move=>k v; rewrite find0E. Qed. +End AssocsLemmas. -Lemma umallUn f1 f2 : um_all f1 -> um_all f2 -> um_all (f1 \+ f2). +Lemma uniq_assocs K C (V : eqType) (U : @union_map_class K C V) (f : U) : + uniq (assocs f). Proof. -case W : (valid (f1 \+ f2)); last by move/invalidE: (negbT W)=>->. -by move=>X Y k v; rewrite findUnL //; case: ifP=>_; [apply: X|apply: Y]. +rewrite !umEX /UM.assocs /=; case: (UMC.from f)=>[|[s H _]] //=. +by move/(sorted_uniq (@trans K) (@irr K)): H; apply: map_uniq. Qed. -Lemma umallUnL f1 f2 : valid (f1 \+ f2) -> um_all (f1 \+ f2) -> um_all f1. +Lemma assocs_map K C (V : Type) (U : @union_map_class K C V) (f : U) k v1 v2 : + (k, v1) \In assocs f -> (k, v2) \In assocs f -> v1 = v2. Proof. -by move=>W H k v F; apply: (H k v); rewrite findUnL // (find_some F). +rewrite !umEX; case: (UMC.from f)=>//= g _; case: g=>g S /= H1 H2. +have {S} S' : uniq [seq key i | i <- g]. +- by apply: sorted_uniq S; [apply: trans | apply: irr]. +have X : forall (g : seq (K*V)) k v, (k, v) \In g -> k \in map fst g. +- elim=>[|[ka va] h IH] // kb vb. + - rewrite !InE; case; first by case=>-> _; rewrite inE eq_refl. + by move/IH; rewrite inE=>->; rewrite orbT. +elim: g k v1 v2 S' H1 H2=>[|[ka va] g IH] //= k v1 v2. +case/andP=>U1 U2; rewrite !InE. +case; first by case=>->->; case=>[[]|/X] //; rewrite (negbTE U1). +move=>H1 H2; case: H2 H1. +- by case=>->-> /X; rewrite (negbTE U1). +by move=>H1 H2; apply: IH H2 H1. Qed. -Lemma umallUnR f1 f2 : valid (f1 \+ f2) -> um_all (f1 \+ f2) -> um_all f2. -Proof. by rewrite joinC; apply: umallUnL. Qed. - -End AllDefLemmas. - -Hint Resolve umall_undef umall0 : core. - - -(********************************************) -(* Interaction of subset, dom and um_filter *) -(********************************************) +(*********************************) +(* Interaction of subset and dom *) +(*********************************) Section Interaction. -Variables (K : ordType) (A : Type) (U : union_map_class K A). +Variables (K : ordType) (C : pred K) (A : Type) (U : union_map_class C A). Implicit Types (x y a b : U) (p q : pred K). -(* filter p x is lower bound for p *) -Lemma subdom_umfilt x p : {subset dom (um_filter p x) <= p}. -Proof. by move=>a; rewrite dom_umfilt; case/andP. Qed. - - (* Some equivalent forms for subset with dom *) -Lemma subdom_indomE x p : {subset dom x <= p} = {in dom x, p =1 predT}. -Proof. by []. Qed. - -Lemma subdom_umfiltE x p : {subset dom x <= p} <-> um_filter p x = x. -Proof. -split; last by move=><- a; rewrite dom_umfilt; case/andP. -by move/eq_in_umfilt=>->; rewrite umfilt_predT. -Qed. - -(* specific consequence of subdom_umfiltE *) -Lemma umfilt_memdomE x : um_filter (mem (dom x)) x = x. -Proof. by apply/subdom_umfiltE. Qed. +Lemma subdom_indomE p x : {subset dom x <= p} = {in dom x, p =1 predT}. +Proof. by []. Qed. (* Some equivalent forms for subset expressing disjointness *) @@ -1514,122 +1615,17 @@ Lemma subdom_valid x y : valid x -> valid y -> valid (x \+ y). Proof. by move=>H X Y; case: validUn; rewrite ?X ?Y=>// k /H /negbTE /= ->. Qed. -Lemma subdom_umfilt0 x p : - valid x -> {subset dom x <= predC p} <-> um_filter p x = Unit. +Lemma subdom_validL x y a : + valid a -> + valid (x \+ y) -> {subset dom a <= dom x} -> valid (a \+ y). Proof. -move=>V; split=>H. -- by rewrite (eq_in_umfilt (p2:=pred0)) ?umfilt_pred0 // => a /H /negbTE ->. -move=>k X; case: dom_find X H=>// a _ -> _; rewrite umfiltU !inE. -by case: ifP=>// _ /empbE; rewrite /= empbU. +rewrite !validUnAE=>-> /and3P [_ -> D] S. +by apply: sub_all D=>z /(contra (S z)). Qed. End Interaction. -(************************) -(* PCM-induced ordering *) -(************************) - -(* Union maps and pcm ordering. *) - -Section UmpleqLemmas. -Variables (K : ordType) (A : Type) (U : union_map_class K A). -Implicit Types (x y a b : U) (p : pred K). - -Lemma umpleq_undef x : [pcm x <= um_undef]. -Proof. by exists um_undef; rewrite join_undefR. Qed. - -Hint Resolve umpleq_undef : core. - -(* pcm-induced preorder, is an order in the case of union maps *) - -Lemma umpleq_asym x y : [pcm x <= y] -> [pcm y <= x] -> x = y. -Proof. -case=>a -> [b]; case V : (valid x); last first. -- by move/invalidE: (negbT V)=>->; rewrite join_undefL. -rewrite -{1 2}[x]unitR -joinA in V *. -by case/(joinxK V)/esym/join0E=>->; rewrite unitR. -Qed. - -(* monotonicity lemmas *) - -Lemma umpleq_filt2 x y p : [pcm x <= y] -> [pcm um_filter p x <= um_filter p y]. -Proof. -move=>H; case V : (valid y). -- by case: H V=>a -> V; rewrite umfiltUn //; eexists _. -by move/invalidE: (negbT V)=>->; rewrite umfilt_undef; apply: umpleq_undef. -Qed. - -(* filter p x is lower bound for x *) -Lemma umpleq_filtI x p : [pcm um_filter p x <= x]. -Proof. -exists (um_filter (predD predT p) x); rewrite -umfilt_predU. -by rewrite -{1}[x]umfilt_predT; apply: eq_in_umfilt=>a; rewrite /= orbT. -Qed. - -Hint Resolve umpleq_filtI : core. - -(* in valid case, we can define the order by means of filter *) -Lemma umpleq_filtE a x : - valid x -> [pcm a <= x] <-> um_filter (mem (dom a)) x = a. -Proof. by move=>V; split=>[|<-] // H; case: H V=>b ->; apply: umfilt_dom. Qed. - -(* filter p x is largest lower bound for x and p *) -Lemma umpleq_filt_meet a x p : - {subset dom a <= p} -> [pcm a <= x] -> [pcm a <= um_filter p x]. -Proof. by move=>D /(umpleq_filt2 p); rewrite (eq_in_umfilt D) umfilt_predT. Qed. - -(* join is the least upper bound *) -Lemma umpleq_join x a b : - valid (a \+ b) -> [pcm a <= x] -> [pcm b <= x] -> [pcm a \+ b <= x]. -Proof. -case Vx : (valid x); last by move/invalidE: (negbT Vx)=>->. -move=>V /(umpleq_filtE _ Vx) <- /(umpleq_filtE _ Vx) <-. -by rewrite -umfilt_dpredU //; apply: valid_subdom V. -Qed. - -(* x <= y and subdom *) - -Lemma umpleq_subdom x y : valid y -> [pcm x <= y] -> {subset dom x <= dom y}. -Proof. by move=>V H; case: H V=>a -> V b D; rewrite domUn inE V D. Qed. - -Lemma subdom_umpleq a x y : - valid (x \+ y) -> [pcm a <= x \+ y] -> - {subset dom a <= dom x} -> [pcm a <= x]. -Proof. -move=>V H Sx; move: (umpleq_filt_meet Sx H); rewrite umfiltUn //. -rewrite umfilt_memdomE; move/subset_disjC: (valid_subdom V). -by move/(subdom_umfilt0 _ (validR V))=>->; rewrite unitR. -Qed. - -(* meet is the greatest lower bound *) -Lemma umpleq_meet a x y1 y2 : - valid (x \+ y1 \+ y2) -> - [pcm a <= x \+ y1] -> [pcm a <= x \+ y2] -> [pcm a <= x]. -Proof. -move=>V H1 H2. -have {V} [V1 V V2] : [/\ valid (x \+ y1), valid (y1 \+ y2) & valid (x \+ y2)]. -- rewrite (validL V); rewrite -joinA in V; rewrite (validR V). - by rewrite joinA joinAC in V; rewrite (validL V). -apply: subdom_umpleq (V1) (H1) _ => k D. -move: {D} (umpleq_subdom V1 H1 D) (umpleq_subdom V2 H2 D). -rewrite !domUn !inE V1 V2 /=; case : (k \in dom x)=>//=. -by case: validUn V=>// _ _ L _ /L /negbTE ->. -Qed. - -(* some/none lemmas *) - -Lemma umpleq_some x1 x2 t s : - valid x2 -> [pcm x1 <= x2] -> find t x1 = Some s -> find t x2 = Some s. -Proof. by move=>V H; case: H V=>a -> V H; rewrite findUnL // (find_some H). Qed. - -Lemma umpleq_none x1 x2 t : - valid x2 -> [pcm x1 <= x2] -> find t x2 = None -> find t x1 = None. -Proof. by case E: (find t x1)=>[a|] // V H <-; rewrite (umpleq_some V H E). Qed. - -End UmpleqLemmas. - - (****************************) (* Generic points-to lemmas *) (****************************) @@ -1640,52 +1636,57 @@ lemmas complicated names prefixed by gen, because they are not supposed to be used in applications *) Section PointsToLemmas. -Variables (K : ordType) (V : Type) (U : union_map_class K V). +Variables (K : ordType) (C : pred K) (V : Type) (U : union_map_class C V). Implicit Types (k : K) (v : V) (f : U). Lemma ptsU k v : pts k v = upd k v Unit :> U. Proof. by rewrite !pcmE /= !umEX /UM.pts /UM.upd; case: decP. Qed. -Lemma domPtK k v : dom (pts k v : U) = if cond U k then [:: k] else [::]. +Lemma pts_condN k v : ~~ C k -> pts k v = undef :> U. +Proof. by move=>C'; rewrite ptsU upd_condN. Qed. + +Lemma domPtK k v : dom (pts k v : U) = if C k then [:: k] else [::]. Proof. rewrite !umEX /= /UM.dom /supp /UM.pts /UM.upd /UM.empty /=. by case D : (decP _)=>[a|a] /=; rewrite ?a ?(introF idP a). Qed. -(* a weaker version of gen_domPtK, but legacy *) -Lemma domPt k v : dom (pts k v : U) =i [pred x | cond U k & k == x]. -Proof. by move=>x; rewrite ptsU domU !inE valid_unit dom0; case: eqVneq. Qed. +(* a weaker legacy version of gen_domPtK *) +Lemma domPt k v : dom (pts k v : U) =i [pred x | C k & k == x]. +Proof. +by move=>x; rewrite ptsU domU !inE eq_sym valid_unit dom0; case: eqP. +Qed. -Lemma validPt k v : valid (pts k v : U) = cond U k. +Lemma validPt k v : valid (pts k v : U) = C k. Proof. by rewrite ptsU validU valid_unit andbT. Qed. Lemma domeqPt k v1 v2 : dom_eq (pts k v1 : U) (pts k v2). Proof. by apply/domeqP; rewrite !validPt; split=>// x; rewrite !domPt. Qed. -Lemma findPt k v : find k (pts k v : U) = if cond U k then Some v else None. +Lemma findPt k v : find k (pts k v : U) = if C k then Some v else None. Proof. by rewrite ptsU findU eq_refl /= valid_unit andbT. Qed. Lemma findPt2 k1 k2 v : find k1 (pts k2 v : U) = - if (k1 == k2) && cond U k2 then Some v else None. + if (k1 == k2) && C k2 then Some v else None. Proof. by rewrite ptsU findU valid_unit andbT find0E; case: ifP=>//=; case: ifP. Qed. Lemma findPt_inv k1 k2 v w : - find k1 (pts k2 v : U) = Some w -> [/\ k1 = k2, v = w & cond U k2]. + find k1 (pts k2 v : U) = Some w -> [/\ k1 = k2, v = w & C k2]. Proof. rewrite ptsU findU; case: eqP; last by case: ifP=>//; rewrite find0E. by move=>->; rewrite valid_unit andbT; case: ifP=>// _ [->]. Qed. Lemma freePt2 k1 k2 v : - cond U k2 -> - free k1 (pts k2 v : U) = if k1 == k2 then Unit else pts k2 v. + C k2 -> + free (pts k2 v : U) k1 = if k1 == k2 then Unit else pts k2 v. Proof. by move=>N; rewrite ptsU freeU free0 N. Qed. Lemma freePt k v : - cond U k -> free k (pts k v : U) = Unit. + C k -> free (pts k v : U) k = Unit. Proof. by move=>N; rewrite freePt2 // eq_refl. Qed. Lemma cancelPt k v1 v2 : @@ -1704,13 +1705,13 @@ Qed. Lemma updPt k v1 v2 : upd k v1 (pts k v2) = pts k v1 :> U. Proof. by rewrite !ptsU updU eq_refl. Qed. -Lemma empbPt k v : empb (pts k v : U) = false. -Proof. by rewrite ptsU empbU. Qed. +Lemma um_unitbPt k v : unitb (pts k v : U) = false. +Proof. by rewrite ptsU um_unitbU. Qed. (* valid *) Lemma validPtUn k v f : - valid (pts k v \+ f) = [&& cond U k, valid f & k \notin dom f]. + valid (pts k v \+ f) = [&& C k, valid f & k \notin dom f]. Proof. case: validUn=>//; last 1 first. - rewrite validPt=>H1 -> H2; rewrite H1 (H2 k) //. @@ -1722,22 +1723,33 @@ by rewrite inE=>/eqP ->->; rewrite andbF. Qed. Lemma validUnPt k v f : - valid (f \+ pts k v) = [&& cond U k, valid f & k \notin dom f]. + valid (f \+ pts k v) = [&& C k, valid f & k \notin dom f]. Proof. by rewrite joinC; apply: validPtUn. Qed. +Lemma validPtUnE v2 v1 k f : valid (pts k v1 \+ f) = valid (pts k v2 \+ f). +Proof. by rewrite !validPtUn. Qed. + +Lemma validUnPtE v2 v1 k f : valid (f \+ pts k v1) = valid (f \+ pts k v2). +Proof. by rewrite !validUnPt. Qed. + +Lemma validPtPt v1 v2 k : valid (pts k v1 \+ pts k v2 : U) = false. +Proof. by rewrite (validPtUnE v2) validUnEb um_unitbPt. Qed. + (* the projections from validPtUn are often useful *) -Lemma validPtUn_cond k v f : valid (pts k v \+ f) -> cond U k. -Proof. by rewrite validPtUn; case/and3P. Qed. +Lemma validPtUnI v1 k f : + valid (pts k v1 \+ f) -> [&& C k, valid f & k \notin dom f]. +Proof. by rewrite validPtUn. Qed. -Lemma validUnPt_cond k v f : valid (f \+ pts k v) -> cond U k. -Proof. by rewrite joinC; apply: validPtUn_cond. Qed. +Lemma validUnPtI v1 k f : + valid (f \+ pts k v1) -> [&& C k, valid f & k \notin dom f]. +Proof. by rewrite validUnPt. Qed. -Lemma validPtUnV k v f : valid (pts k v \+ f) -> valid f. +Lemma validPtUn_cond k v f : valid (pts k v \+ f) -> C k. Proof. by rewrite validPtUn; case/and3P. Qed. -Lemma validUnPtV k v f : valid (f \+ pts k v) -> valid f. -Proof. by rewrite joinC; apply: validPtUnV. Qed. +Lemma validUnPt_cond k v f : valid (f \+ pts k v) -> C k. +Proof. by rewrite joinC; apply: validPtUn_cond. Qed. Lemma validPtUnD k v f : valid (pts k v \+ f) -> k \notin dom f. Proof. by rewrite validPtUn; case/and3P. Qed. @@ -1752,7 +1764,7 @@ Lemma domPtUn k v f : [pred x | valid (pts k v \+ f) & (k == x) || (x \in dom f)]. Proof. move=>x; rewrite domUn !inE !validPtUn domPt !inE. -by rewrite -!andbA; case: (UMC.p _ k). +by rewrite -!andbA; case: (C k). Qed. Lemma domUnPt k v f : @@ -1766,9 +1778,40 @@ Proof. by rewrite domPtUn inE eq_refl andbT. Qed. Lemma domUnPtE k v f : k \in dom (f \+ pts k v) = valid (f \+ pts k v). Proof. by rewrite joinC; apply: domPtUnE. Qed. +Lemma domPtUnE2 k v1 v2 f : dom (pts k v1 \+ f) = dom (pts k v2 \+ f). +Proof. by apply/domE=>x; rewrite !domPtUn !validPtUn. Qed. + +Lemma domUnPtE2 k v1 v2 f : dom (f \+ pts k v1) = dom (f \+ pts k v2). +Proof. by rewrite !(joinC f); apply: domPtUnE2. Qed. + Lemma validxx f : valid (f \+ f) -> dom f =i pred0. Proof. by case: validUn=>// _ _ L _ z; case: (_ \in _) (L z)=>//; elim. Qed. +Lemma domPtUnK k v f : + valid (pts k v \+ f) -> + all (ord k) (dom f) -> + dom (pts k v \+ f) = k :: dom f. +Proof. +move=>W H; apply: (@eq_sorted_irr K ord) =>//=. +- by rewrite path_min_sorted //; apply: sorted_dom. +by move=>x; rewrite domPtUn !inE W eq_sym. +Qed. + +Lemma domUnPtK k v f : + valid (f \+ pts k v) -> + all (ord^~k) (dom f) -> + dom (f \+ pts k v) = rcons (dom f) k. +Proof. +move=>W H; apply: (@eq_sorted_irr K ord)=>//=. +- rewrite -(rev_sorted (fun x=>ord^~x)) rev_rcons /=. + by rewrite path_min_sorted ?rev_sorted // all_rev. +by move=>x; rewrite domUnPt inE W mem_rcons inE eq_sym. +Qed. + +Lemma size_domPtUn k v f : + valid (pts k v \+ f) -> size (dom (pts k v \+ f)) = 1 + size (dom f). +Proof. by move=>W; rewrite size_domUn // domPtK (validPtUn_cond W). Qed. + (* dom_eq *) Lemma domeqPtUn k v1 v2 f1 f2 : @@ -1811,8 +1854,8 @@ Proof. by rewrite joinC; apply: findPtUn2. Qed. Lemma freePtUn2 k1 k2 v f : valid (pts k2 v \+ f) -> - free k1 (pts k2 v \+ f) = - if k1 == k2 then f else pts k2 v \+ free k1 f. + free (pts k2 v \+ f) k1 = + if k1 == k2 then f else pts k2 v \+ free f k1. Proof. move=>V'; rewrite freeUn domPtUn inE V' /= eq_sym. rewrite ptsU freeU (validPtUn_cond V') free0. @@ -1822,18 +1865,24 @@ Qed. Lemma freeUnPt2 k1 k2 v f : valid (f \+ pts k2 v) -> - free k1 (f \+ pts k2 v) = - if k1 == k2 then f else free k1 f \+ pts k2 v. + free (f \+ pts k2 v) k1 = + if k1 == k2 then f else free f k1 \+ pts k2 v. Proof. by rewrite !(joinC _ (pts k2 v)); apply: freePtUn2. Qed. Lemma freePtUn k v f : - valid (pts k v \+ f) -> free k (pts k v \+ f) = f. + valid (pts k v \+ f) -> free (pts k v \+ f) k = f. Proof. by move=>V'; rewrite freePtUn2 // eq_refl. Qed. Lemma freeUnPt k v f : - valid (f \+ pts k v) -> free k (f \+ pts k v) = f. + valid (f \+ pts k v) -> free (f \+ pts k v) k = f. Proof. by rewrite joinC; apply: freePtUn. Qed. +(* Frequently we introduce existentials to name *) +(* the "rest" of a map in a representation with a union of a pointer *) +(* This lemma bridges between such representation and one with free *) +Lemma freeX k v f1 f2 : valid f1 -> f1 = pts k v \+ f2 -> f2 = free f1 k. +Proof. by move=>W E; rewrite E freePtUn // -E. Qed. + (* upd *) Lemma updPtUn k v1 v2 f : upd k v1 (pts k v2 \+ f) = pts k v1 \+ f. @@ -1843,76 +1892,58 @@ case V1 : (valid (pts k v2 \+ f)). have V2 : valid (pts k v1 \+ f) = false. - by rewrite !validPtUn in V1 *. move/negbT/invalidE: V1=>->; move/negbT/invalidE: V2=>->. -by apply: upd_invalid. +by apply: upd_undef. Qed. Lemma updUnPt k v1 v2 f : upd k v1 (f \+ pts k v2) = f \+ pts k v1. Proof. by rewrite !(joinC f); apply: updPtUn. Qed. -(* empb *) - -Lemma empbPtUn k v f : empb (pts k v \+ f) = false. -Proof. by rewrite empbUn empbPt. Qed. - -Lemma empbUnPt k v f : empb (f \+ pts k v) = false. -Proof. by rewrite joinC; apply: empbPtUn. Qed. - -(* undef *) - -Lemma pts_undef k v1 v2 : pts k v1 \+ pts k v2 = um_undef :> U. +Lemma upd_eta k v f : upd k v f = pts k v \+ free f k. Proof. -apply/invalidE; rewrite validPtUn validPt domPt !negb_and negb_or eq_refl. -by case: (cond U k). +case W : (valid f); last first. +- by move/negbT/invalidE: W=>->; rewrite upd_undef free_undef join_undef. +case H : (C k); last first. +- by move/negbT: H=>H; rewrite (upd_condN v) // pts_condN // undef_join. +have W' : valid (pts k v \+ free f k). +- by rewrite validPtUn H validF W domF inE eq_refl. +apply: find_eta=>//; first by rewrite validU H W. +by move=>k'; rewrite findU H W findPtUn2 // findF; case: eqP. Qed. -(* um_filter *) - -Lemma umfiltPt p k v : - um_filter p (pts k v : U) = - if p k then pts k v else if cond U k then Unit else um_undef. -Proof. by rewrite ptsU umfiltU umfilt0. Qed. - -Lemma umfiltPtUn p k v f : - um_filter p (pts k v \+ f) = - if valid (pts k v \+ f) then - if p k then pts k v \+ um_filter p f else um_filter p f - else um_undef. -Proof. -case: ifP=>X; last by move/invalidE: (negbT X)=>->; rewrite umfilt_undef. -rewrite umfiltUn // umfiltPt (validPtUn_cond X). -by case: ifP=>//; rewrite unitL. -Qed. +(* um_unitb *) -(* um_all *) +Lemma um_unitbPtUn k v f : unitb (pts k v \+ f) = false. +Proof. by rewrite um_unitbUn um_unitbPt. Qed. -Lemma umallPt (P : V -> Prop) k v : P v -> um_all P (pts k v : U). -Proof. by move=>X u w /findPt_inv [_ <-]. Qed. +Lemma um_unitbUnPt k v f : unitb (f \+ pts k v) = false. +Proof. by rewrite joinC; apply: um_unitbPtUn. Qed. -Lemma umallPtUn (P : V -> Prop) k v f : - P v -> um_all P f -> um_all P (pts k v \+ f). -Proof. by move/(umallPt (k:=k)); apply: umallUn. Qed. +(* undef *) -Lemma umallPtE (P : V -> Prop) k v : cond U k -> um_all P (pts k v : U) -> P v. -Proof. by move=>C /(_ k v); rewrite findPt C; apply. Qed. +Lemma pts_undef k v : pts k v = undef :> U <-> ~~ C k. +Proof. by rewrite -(validPt k v) invalidE. Qed. -Lemma umallPtUnE (P : V -> Prop) k v f : - valid (pts k v \+ f) -> um_all P (pts k v \+ f) -> P v /\ um_all P f. +Lemma pts_undef2 k v1 v2 : pts k v1 \+ pts k v2 = undef :> U. Proof. -move=>W H; move: (umallUnL W H) (umallUnR W H)=>{H} H1 H2. -by split=>//; apply: umallPtE H1; apply: validPtUn_cond W. +apply/invalidE; rewrite validPtUn validPt domPt !negb_and negb_or eq_refl. +by case: (C k). Qed. +(* umpleq *) + +Lemma umpleq_dom k v f : valid f -> [pcm pts k v <= f] -> k \in dom f. +Proof. by move=>W [z E]; subst f; rewrite domPtUn inE W eq_refl. Qed. + (* others *) Lemma um_eta k f : - k \in dom f -> exists v, find k f = Some v /\ f = pts k v \+ free k f. + k \in dom f -> exists v, find k f = Some v /\ f = pts k v \+ free f k. Proof. case: dom_find=>// v E1 E2 _; exists v; split=>//. -by rewrite {1}E2 -{1}[free k f]unitL updUnR domF inE /= eq_refl ptsU. +by rewrite {1}E2 -{1}[free f k]unitL updUnR domF inE /= eq_refl ptsU. Qed. -Lemma um_eta2 k v f : - find k f = Some v -> f = pts k v \+ free k f. +Lemma um_eta2 k v f : find k f = Some v -> f = pts k v \+ free f k. Proof. by move=>E; case/um_eta: (find_some E)=>v' []; rewrite E; case=><-. Qed. Lemma cancel k v1 v2 f1 f2 : @@ -1922,55 +1953,69 @@ Proof. move=>V' E. have: find k (pts k v1 \+ f1) = find k (pts k v2 \+ f2) by rewrite E. rewrite !findPtUn -?E //; case=>X. -by rewrite -{}X in E *; rewrite -(joinxK V' E) (validR V'). +by rewrite -{}X in E *; rewrite -(joinxK V' E) (validE2 V'). Qed. Lemma cancel2 k1 k2 f1 f2 v1 v2 : valid (pts k1 v1 \+ f1) -> pts k1 v1 \+ f1 = pts k2 v2 \+ f2 -> if k1 == k2 then v1 = v2 /\ f1 = f2 - else [/\ free k1 f2 = free k2 f1, - f1 = pts k2 v2 \+ free k1 f2 & - f2 = pts k1 v1 \+ free k2 f1]. + else [/\ free f2 k1 = free f1 k2, + f1 = pts k2 v2 \+ free f2 k1 & + f2 = pts k1 v1 \+ free f1 k2]. Proof. move=>V1 E; case: ifP=>X. - by rewrite -(eqP X) in E; case/(cancel V1): E. move: (V1); rewrite E => V2. -have E' : f2 = pts k1 v1 \+ free k2 f1. -- move: (erefl (free k2 (pts k1 v1 \+ f1))). +have E' : f2 = pts k1 v1 \+ free f1 k2. +- move: (erefl (free (pts k1 v1 \+ f1) k2)). rewrite {2}E freeUn E domPtUn inE V2 eq_refl /=. by rewrite ptsU freeU eq_sym X free0 -ptsU freePtUn. -suff E'' : free k2 f1 = free k1 f2. +suff E'' : free f1 k2 = free f2 k1. - by rewrite -E''; rewrite E' joinCA in E; case/(cancel V1): E. -move: (erefl (free k1 f2)). -by rewrite {1}E' freePtUn // -E'; apply: (validR V2). +move: (erefl (free f2 k1)). +by rewrite {1}E' freePtUn // -E' (validE2 V2). +Qed. + +Lemma um_contra k v f g : f = pts k v \+ g \+ f -> ~~ valid f. +Proof. +move=>E; rewrite E -joinAC -joinA validPtUn. +rewrite {2}E -!joinA domPtUn !inE. +by rewrite eq_refl andbT !joinA -E andbN andbF. Qed. (* induction over union maps, expressed with pts and \+ *) (* forward progressing over keys *) Lemma um_indf (P : U -> Prop) : - P um_undef -> P Unit -> + P undef -> P Unit -> (forall k v f, P f -> valid (pts k v \+ f) -> path ord k (dom f) -> P (pts k v \+ f)) -> forall f, P f. Proof. -rewrite !pcmE /= !umEX => H1 H2 H3 f; rewrite -[f]UMC.tfE. +rewrite !pcmE /undef /= !umEX => H1 H2 H3 f; rewrite -[f]UMC.tfE. apply: (UM.base_indf (P := (fun b => P (UMC.to b))))=>//. move=>k v g H V1; move: (H3 k v _ H); rewrite !pcmE /= !umEX. by apply. Qed. +(* unordered progressing over keys *) +Lemma um_ind' (P : U -> Prop) : + P undef -> P Unit -> + (forall k v f, P f -> valid (pts k v \+ f) -> P (pts k v \+ f)) -> + forall f, P f. +Proof. by move=>H1 H2 H3; apply: um_indf=>// k v f H4 H5 _; apply: H3. Qed. + (* backward progressing over keys *) Lemma um_indb (P : U -> Prop) : - P um_undef -> P Unit -> - (forall k v f, P f -> valid (pts k v \+ f) -> + P undef -> P Unit -> + (forall k v f, P f -> valid (f \+ pts k v) -> (forall x, x \in dom f -> ord x k) -> P (pts k v \+ f)) -> forall f, P f. Proof. -rewrite !pcmE /= !umEX => H1 H2 H3 f; rewrite -[f]UMC.tfE. +rewrite !pcmE /undef /= !umEX => H1 H2 H3 f; rewrite -[f]UMC.tfE. apply: (UM.base_indb (P := (fun b => P (UMC.to b))))=>//. move=>k v g H V1; move: (H3 k v _ H); rewrite !pcmE /= !umEX. by apply. @@ -1981,27 +2026,21 @@ Lemma um_valid3 f1 f2 f3 : valid (f1 \+ f2 \+ f3) = [&& valid (f1 \+ f2), valid (f2 \+ f3) & valid (f1 \+ f3)]. Proof. -apply/idP/idP. -- move=>W; rewrite (validL W). - rewrite -joinA in W; rewrite (validR W). - by rewrite joinCA in W; rewrite (validR W). -case/and3P=>V1 V2 V3; case: validUn=>//. -- by rewrite V1. -- by rewrite (validR V2). -move=>k; rewrite domUn inE V1; case/orP. -- by case: validUn V3=>// _ _ X _ /X /negbTE ->. -by case: validUn V2=>// _ _ X _ /X /negbTE ->. +apply/idP/and3P; first by move=>W; rewrite !(validLE3 W). +case=>V1 V2 V3; case: validUn=>//; rewrite ?V1 ?(validE2 V2) // => k. +by rewrite domUn inE V1; case/orP; [move: V3 | move: V2]; + case: validUn =>// _ _ X _ /X /negbTE ->. Qed. (* points-to is a prime element of the monoid *) Lemma um_prime f1 f2 k v : - cond U k -> + C k -> f1 \+ f2 = pts k v -> f1 = pts k v /\ f2 = Unit \/ f1 = Unit /\ f2 = pts k v. Proof. move: f1 f2; apply: um_indf=>[f1|f2 _|k2 w2 f1 _ _ _ f X E]. -- rewrite join_undefL -(validPt _ v)=>W E. +- rewrite undef_join -(validPt _ v)=>W E. by rewrite -E in W; rewrite valid_undef in W. - by rewrite unitL=>->; right. have W : valid (pts k2 w2 \+ (f1 \+ f)). @@ -2009,140 +2048,3498 @@ have W : valid (pts k2 w2 \+ (f1 \+ f)). rewrite -[pts k v]unitR -joinA in E. move/(cancel2 W): E; case: ifP. - by move/eqP=>-> [->] /join0E [->->]; rewrite unitR; left. -by move=>_ [_ _] /esym/empbP; rewrite empbPtUn. +by move=>_ [_ _] /esym/unitbP; rewrite um_unitbPtUn. Qed. +(* also a simple rearrangment equation *) +Definition pullk (k : K) (v : V) (f : U) := pull (pts k v) f. + End PointsToLemmas. Hint Resolve domeqPt domeqPtUn domeqUnPt : core. -Prenex Implicits validPtUn_cond findPt_inv um_eta2. - -(********************) -(* Topped structure *) -(********************) +Prenex Implicits validPtUn_cond findPt_inv um_eta2 um_contra. +Prenex Implicits validPtUnD validUnPtD. -Module UnionClassTPCM. -Section UnionClassTPCM. -Variables (K : ordType) (V : Type) (U : union_map_class K V). -Implicit Type f : U. -Lemma join0E f1 f2 : f1 \+ f2 = Unit -> f1 = Unit /\ f2 = Unit. -Proof. by rewrite join0E. Qed. +(******************************************************) +(******************************************************) +(* Defined notions that don't appear in the signature *) +(******************************************************) +(******************************************************) -Lemma valid_undefN : ~~ valid (um_undef: U). -Proof. by rewrite valid_undef. Qed. +(************) +(* um_preim *) +(************) -Lemma undef_join f : um_undef \+ f = um_undef. -Proof. by rewrite join_undefL. Qed. -End UnionClassTPCM. +Section PreimDefLemmas. +Variables (K : ordType) (C : pred K) (V : Type) (U : union_map_class C V). +Implicit Types (k : K) (v : V) (f : U) (p : pred V). -Module Exports. -Definition union_map_classTPCMMix K V (U : union_map_class K V) := - TPCMMixin (@empbP K V U) (@join0E K V U) - (@valid_undefN K V U) (@undef_join _ _ U). -Canonical union_map_classTPCM K V (U : union_map_class K V) := - Eval hnf in TPCM _ (@union_map_classTPCMMix K V U). -End Exports. -End UnionClassTPCM. +Definition um_preim p f k := oapp p false (find k f). -Export UnionClassTPCM.Exports. +Lemma umpreim_undef p : um_preim p undef =1 pred0. +Proof. by move=>x; rewrite /um_preim find_undef. Qed. +Lemma umpreim0 p : um_preim p Unit =1 pred0. +Proof. by move=>x; rewrite /um_preim find0E. Qed. -(********************************************************) -(********************************************************) -(* Instance of union maps with trivially true condition *) -(********************************************************) -(********************************************************) +Lemma umpreimUn p f1 f2 : + valid (f1 \+ f2) -> + um_preim p (f1 \+ f2) =1 predU (um_preim p f1) (um_preim p f2). +Proof. +move=>v x; rewrite /um_preim findUnL //=. +case X: (find x f1)=>[a|]; case Y: (find x f2)=>[b|] /=. +- by move: (dom_inNL v (find_some X)); rewrite (find_some Y). +- by rewrite ifT ?(find_some X) // orbC. +- by rewrite ifN //; apply/find_none. +by rewrite ifN //; apply/find_none. +Qed. -(* We build it over the base type with a trival condition on keys. We -seal the definition to avoid any slowdowns (just in case). *) +Lemma umpreim_pred0 f : um_preim pred0 f =1 pred0. +Proof. by move=>x; rewrite /um_preim; by case: (find x f). Qed. -Module Type UMSig. -Parameter tp : ordType -> Type -> Type. +Lemma umpreim_dom f : um_preim predT f =1 mem (dom f). +Proof. +move=>x; rewrite /um_preim /=. +case X: (find x f)=>[a|] /=; first by rewrite (find_some X). +by apply/esym/negbTE/find_none. +Qed. -Section Params. -Variables (K : ordType) (V : Type). -Notation tp := (tp K V). +End PreimDefLemmas. -Parameter um_undef : tp. -Parameter defined : tp -> bool. -Parameter empty : tp. -Parameter upd : K -> V -> tp -> tp. -Parameter dom : tp -> seq K. -Parameter dom_eq : tp -> tp -> bool. -Parameter free : K -> tp -> tp. -Parameter find : K -> tp -> option V. -Parameter union : tp -> tp -> tp. -Parameter um_filter : pred K -> tp -> tp. -Parameter empb : tp -> bool. -Parameter undefb : tp -> bool. -Parameter pts : K -> V -> tp. -Parameter from : tp -> @UM.base K V predT. -Parameter to : @UM.base K V predT -> tp. +(****************************) +(* map membership predicate *) +(****************************) -Axiom ftE : forall b, from (to b) = b. -Axiom tfE : forall f, to (from f) = f. -Axiom undefE : um_undef = to (@UM.Undef K V predT). -Axiom defE : forall f, defined f = UM.valid (from f). -Axiom emptyE : empty = to (@UM.empty K V predT). -Axiom updE : forall k v f, upd k v f = to (UM.upd k v (from f)). -Axiom domE : forall f, dom f = UM.dom (from f). -Axiom dom_eqE : forall f1 f2, dom_eq f1 f2 = UM.dom_eq (from f1) (from f2). -Axiom freeE : forall k f, free k f = to (UM.free k (from f)). -Axiom findE : forall k f, find k f = UM.find k (from f). -Axiom unionE : forall f1 f2, union f1 f2 = to (UM.union (from f1) (from f2)). -Axiom umfiltE : forall q f, um_filter q f = to (UM.um_filter q (from f)). -Axiom empbE : forall f, empb f = UM.empb (from f). -Axiom undefbE : forall f, undefb f = UM.undefb (from f). -Axiom ptsE : forall k v, pts k v = to (@UM.pts K V predT k v). +Section MapMembership. +Variables (K : ordType) (C : pred K) (V : Type) (U : union_map_class C V). +Implicit Type f : U. -End Params. +Definition Mem_UmMap f := + fun x : K * V => valid f /\ [pcm pts x.1 x.2 <= f]. -End UMSig. +Canonical Structure um_PredType := + Eval hnf in @mkPredType (K*V) U Mem_UmMap. +(* This makes Mem_UmMap a canonical instance of topred. *) +Canonical Structure Mem_UmMap_PredType := + Eval hnf in mkPredType Mem_UmMap. -Module UMDef : UMSig. -Section UMDef. -Variables (K : ordType) (V : Type). +Lemma In_undef x : x \In (undef : U) -> False. +Proof. by case; rewrite valid_undef. Qed. -Definition tp : Type := @UM.base K V predT. - -Definition um_undef := @UM.Undef K V predT. -Definition defined f := @UM.valid K V predT f. -Definition empty := @UM.empty K V predT. -Definition upd k v f := @UM.upd K V predT k v f. -Definition dom f := @UM.dom K V predT f. -Definition dom_eq f1 f2 := @UM.dom_eq K V predT f1 f2. -Definition free k f := @UM.free K V predT k f. -Definition find k f := @UM.find K V predT k f. -Definition union f1 f2 := @UM.union K V predT f1 f2. -Definition um_filter q f := @UM.um_filter K V predT q f. -Definition empb f := @UM.empb K V predT f. -Definition undefb f := @UM.undefb K V predT f. -Definition pts k v := @UM.pts K V predT k v. - -Definition from (f : tp) : @UM.base K V predT := f. -Definition to (b : @UM.base K V predT) : tp := b. +Lemma In0 x : x \In (Unit : U) -> False. +Proof. by case=>_; case=>z /esym/unitbP; rewrite um_unitbPtUn. Qed. + +Lemma In_find k v f : (k, v) \In f <-> find k f = Some v. +Proof. +split=>[[W] /= [z E]|E]; first by rewrite E findPtUn // -E. +split; first by move/find_some/dom_valid: E. +by move/um_eta2: E=>->; exists (free f k). +Qed. + +Lemma In_fun k v1 v2 f : (k, v1) \In f -> (k, v2) \In f -> v1 = v2. +Proof. by move/In_find=>H1 /In_find; rewrite H1; case. Qed. + +Lemma InUn x f1 f2 : x \In (f1 \+ f2) -> x \In f1 \/ x \In f2. +Proof. +move/In_find=>F; move/find_some: (F); rewrite domUn inE=>/andP [W D]. +case/orP: D F=>D; first by rewrite findUnL // D=>/In_find; left. +by rewrite findUnR // D=>/In_find; right. +Qed. + +Lemma InL x f1 f2 : valid (f1 \+ f2) -> x \In f1 -> x \In (f1 \+ f2). +Proof. by move=>W /In_find E; apply/In_find; rewrite findUnL // (find_some E). Qed. + +Lemma InR x f1 f2 : valid (f1 \+ f2) -> x \In f2 -> x \In (f1 \+ f2). +Proof. by rewrite joinC; apply: InL. Qed. + +Lemma InPt x k v : x \In pts k v -> x = (k, v) /\ C k. +Proof. +by case: x=>a b /In_find; rewrite findPt2; case: ifP=>//; case/andP=>/eqP ->-> [->]. +Qed. + +Lemma In_condPt k v : C k -> (k, v) \In pts k v. +Proof. by split; [rewrite validPt | exists Unit; rewrite unitR]. Qed. + +Lemma In_dom f x : x \In f -> x.1 \in dom f. +Proof. by case=>W [z E]; subst f; rewrite domPtUn inE W eq_refl. Qed. + +Lemma In_cond f x : x \In f -> C x.1. +Proof. by move/In_dom/dom_cond. Qed. + +Lemma In_domX k f : reflect (exists v, (k, v) \In f) (k \in dom f). +Proof. +case: dom_find=>[E|v /In_find H E]; constructor; last by exists v. +by case=>v /In_find; rewrite E. +Qed. + +Lemma InPtUnE x k v f : + valid (pts k v \+ f) -> + x \In pts k v \+ f <-> x = (k, v) \/ x \In f. +Proof. +move=>W; split; last by case=>[->|/(InR W)]. +by case/InUn; [case/InPt=>->; left|right]. +Qed. + +Lemma InPtUnL k v f : valid (pts k v \+ f) -> (k, v) \In pts k v \+ f. +Proof. by move/InPtUnE=>->; left. Qed. + +Lemma InPtUnEN k v x f : + valid (pts k v \+ f) -> + x \In (pts k v \+ f) <-> x = (k, v) \/ x \In f /\ x.1 != k. +Proof. +move=>W; rewrite InPtUnE //; split; last by case; [left|case; right]. +case=>[->|H]; first by left. +right; split=>//; move/In_dom: H=>H. +case: validUn W=>//; rewrite validPt=>W W' /(_ k). +rewrite domPt inE W eq_refl /= => /(_ (erefl _)). +by case: eqP H=>// ->->. +Qed. + +Lemma InPtUn x k v f : + x \In pts k v \+ f -> + (forall w, valid (pts k w \+ f)) /\ + (x = (k, v) /\ C k \/ x \In f /\ x.1 != k). +Proof. +move=>H; have W w : valid (pts k w \+ f). +- by rewrite (validPtUnE v); apply: dom_valid (In_dom H). +split=>//; move/(InPtUnEN _ (W v)): H (W v). +by case=>[-> /validPtUn_cond|]; [left | right]. +Qed. + +Lemma InPtUnK k v w f : + valid (pts k v \+ f) -> + (k, w) \In (pts k v \+ f) <-> v = w. +Proof. +move=>W; rewrite InPtUnEN //; split=>[|->]; last by left. +by case=>[[->]|[]] //=; rewrite eq_refl. +Qed. + +Lemma In_domY k f : k \in dom f -> sigT (fun v => (k, v) \In f). +Proof. by case: dom_find=>// v /In_find; exists v. Qed. + +Lemma umemE f x : x \In assocs f <-> x \In f. +Proof. +move: f; apply: um_indf=>[||k v f IH W /(order_path_min (@trans K)) P]. +- by rewrite assocs_undef; split=>// /In_undef. +- by rewrite assocs0; split=>// /In0. +by rewrite assocsPtUn ?InE // InPtUnE // IH. +Qed. + +Lemma umem_eq f1 f2 : + valid f1 -> valid f2 -> + (forall x, x \In f1 <-> x \In f2) -> f1 = f2. +Proof. +move=>V1 V2 H; apply: find_eta=>// k. +case K1 : (find k f1)=>[a1|]; case K2 : (find k f2)=>[a2|] //. +- by move/In_find/H/In_find: K1; rewrite K2. +- by move/In_find/H/In_find: K1; rewrite K2. +by move/In_find/H/In_find: K2; rewrite K1. +Qed. + +Lemma InU x k v f : + x \In upd k v f <-> + [/\ valid (upd k v f) & if x.1 == k then x.2 = v else x \In f]. +Proof. +case: x=>x1 v1; split; last first. +- case=>W /= H; split=>//=; exists (free (upd k v f) x1); apply: um_eta2. + move: (W); rewrite validU=>/andP [C' W']; rewrite findU C' W' /=. + by case: eqP H=>[_ ->|_ /In_find]. +move=>H; move: (In_dom H)=>/= D; move: (dom_valid D) (dom_valid D)=>W. +rewrite {1}validU=>/andP [C' W']; split=>//. +move: (D); rewrite domU inE C' /= W'; case: H=>/= _ [z E]. +case: ifP D E=>[/eqP -> D E _|N _ E D]. +- have: find k (upd k v f) = Some v by rewrite findU eq_refl C' W'. + by rewrite E findPtUn -?E //; case=>->. +have: find x1 (pts x1 v1 \+ z) = Some v1 by rewrite findPtUn // -E. +by rewrite -E findU N C'=>/In_find. +Qed. + +Lemma InF x k f : + x \In free f k <-> + [/\ valid (free f k), x.1 != k & x \In f]. +Proof. +case: x=>x1 v1; split; last first. +- by case=>W /= N /In_find E; apply/In_find; rewrite findF (negbTE N). +case=>W /= H; rewrite W; case: H=>z E. +have : find x1 (pts x1 v1 \+ z) = Some v1 by rewrite findPtUn // -E. +by rewrite -E findF; case: eqP=>//= _ /In_find. +Qed. + +Lemma In_eta k v f : (k, v) \In f -> f = pts k v \+ free f k. +Proof. by move=>H; case/In_dom/um_eta: (H)=>w [/In_find/(In_fun H) ->]. Qed. + +End MapMembership. + +Arguments In_condPt {K C V U k}. +Prenex Implicits InPt In_eta InPtUn In_dom. + +(*********) +(* range *) +(*********) + +Section Range. +Variables (K : ordType) (C : pred K) (V : Type) (U : union_map_class C V). +Implicit Types f : U. + +Definition range f := map snd (assocs f). + +Lemma size_range f : size (range f) = size (dom f). +Proof. by rewrite assocs_dom !size_map. Qed. + +Lemma range_undef : range undef = [::]. +Proof. by rewrite /range assocs_undef. Qed. + +Lemma range0 : range Unit = [::]. +Proof. by rewrite /range assocs0. Qed. + +Lemma In_rangeX v f : v \In range f <-> exists k, (k, v) \In f. +Proof. +elim/um_indf: f v=>[v|v|k w f IH W P v]. +- by rewrite range_undef; split=>//; case=>x /In_undef. +- by rewrite range0; split=>//; case=>x /In0. +move: (order_path_min (@trans K) P)=>A. +rewrite /range assocsPtUn //= InE; split. +- case=>[->|/IH [k' H]]; first by exists k; rewrite InPtUnE //; left. + by exists k'; rewrite InPtUnE //; right. +case=>k'; rewrite InPtUnE //; case; first by case; left. +by move=>H; right; apply/IH; exists k'. +Qed. + +Lemma In_range_valid k f : k \In range f -> valid f. +Proof. by case/In_rangeX=>v /In_dom/dom_valid. Qed. + +Lemma In_range k v f : (k, v) \In f -> v \In range f. +Proof. by move=>H; apply/In_rangeX; exists k. Qed. + +Lemma In_rangeUn f1 f2 x : + x \In range (f1 \+ f2) -> + x \In range f1 \/ x \In range f2. +Proof. by rewrite !In_rangeX; case=>k /InUn [] H; [left | right]; exists k. Qed. + +Lemma In_rangeL f1 f2 x : + valid (f1 \+ f2) -> x \In range f1 -> x \In range (f1 \+ f2). +Proof. by move=>W; rewrite !In_rangeX; case=>k H; exists k; apply/InL. Qed. + +Lemma In_rangeR f1 f2 x : + valid (f1 \+ f2) -> x \In range f2 -> x \In range (f1 \+ f2). +Proof. by move=>W; rewrite !In_rangeX; case=>k H; exists k; apply/InR. Qed. + +Lemma In_rangePt k : + C k -> forall v x, x \In range (pts k v) <-> (x = v). +Proof. +move=>C' v x; rewrite In_rangeX. split. +- by case=>w /InPt [[]]. +by move=>->; exists k; apply: In_condPt. +Qed. + +Lemma In_rangePtUn k v f x : + valid (pts k v \+ f) -> + x \In range (pts k v \+ f) <-> x = v \/ x \In range f. +Proof. +move=>W; split. +- case/In_rangeUn; last by right. + by move/(In_rangePt (validPtUn_cond W)); left. +case=>[->|]; last by apply: In_rangeR. +by apply/(In_rangeL W)/(In_rangePt (validPtUn_cond W)). +Qed. + +End Range. + +Prenex Implicits In_range_valid In_range In_rangeUn. + +(* decidable versions, when V is eqtype *) + +Section DecidableRange. +Variables (K : ordType) (C : pred K) (V : eqType) (U : union_map_class C V). +Implicit Types f : U. + +Lemma uniq_rangeP f : + reflect (forall k1 k2 v, (k1, v) \In f -> (k2, v) \In f -> k1 = k2) + (uniq (range f)). +Proof. +case W : (valid f); last first. +- move/negbT/invalidE: W=>->; rewrite range_undef. + by constructor=>k1 k2 v /In_undef. +case H : (uniq (range f)); constructor; last first. +- move=>H'; move/negbT/negP: H; elim. + rewrite map_inj_in_uniq; first by apply: uniq_assocs. + case=>/= k1 v [k2 v'] /mem_seqP/umemE H1 /mem_seqP/umemE H2 /= H3. + by rewrite -H3 in H2 *; rewrite (H' _ _ _ H1 H2). +move/uniqP: H=>H k1 k2 v H1 H2. +set j1 := index k1 (dom f). +set j2 := index k2 (dom f). +have [D1 D2] : k1 \in dom f /\ k2 \in dom f. +- by move/In_dom: H1; move/In_dom: H2. +have [R1 R2] : j1 < size (assocs f) /\ j2 < size (assocs f). +- by rewrite size_assocs !index_mem. +have [M1 M2] : j1 < size (dom f) /\ j2 < size (dom f). +- by rewrite !index_mem. +have [A1 A2] : (k1, v) \in assocs f /\ (k2, v) \in assocs f. +- by move/umemE/mem_seqP: H1=>->; move/umemE/mem_seqP: H2=>->. +have InjF : {in assocs f &, injective fst}. +- case=>a1 v1 [a2 v2] /mem_seqP X1 /mem_seqP X2 /= E. + by move: E X1 X2 => -> X1 /(assocs_map X1) ->. +have /eqP E1 : j1 == index (k1,v) (assocs f). +- rewrite -(nth_uniq (k1,v) R1 _ (uniq_assocs _)); last by rewrite index_mem. + by rewrite /j1 assocs_dom (nth_index_map _ InjF A1) nth_index. +have /eqP E2 : j2 == index (k2,v) (assocs f). +- rewrite -(nth_uniq (k2,v) R2 _ (uniq_assocs _)); last by rewrite index_mem. + by rewrite /j2 assocs_dom (nth_index_map _ InjF A2) nth_index. +have E : nth v (range f) j1 = nth v (range f) j2. +- rewrite /range (nth_map (k1,v) v _ R1) (nth_map (k2,v) v _ R2). + by rewrite E1 E2 !nth_index. +have : j1 = j2 by apply: H E; rewrite inE size_range. +by move/eqP; rewrite -(nth_uniq k1 M1 M2 (uniq_dom _)) !nth_index // =>/eqP. +Qed. + +(* this is just a renaming of mem_seqP for easier finding *) +Lemma mem_rangeP x f : reflect (x \In range f) (x \in range f). +Proof. by apply: mem_seqP. Qed. + +Lemma mem_rangeX v f : v \in range f <-> exists k, (k, v) \In f. +Proof. by split; [move/mem_seqP/In_rangeX|move/In_rangeX/mem_seqP]. Qed. + +Lemma range_valid k f : k \in range f -> valid f. +Proof. by move/mem_seqP/In_range_valid. Qed. + +Lemma mem_range k v f : (k, v) \In f -> v \in range f. +Proof. by move/In_range/mem_seqP. Qed. + +Lemma rangeUn f1 f2 : + range (f1 \+ f2) =i + [pred x | valid (f1 \+ f2) && ((x \in range f1) || (x \in range f2))]. +Proof. +move=>k; apply/idP/idP; last first. +- case/andP=>W /orP [] /mem_seqP H; apply/mem_seqP; + by [apply/(In_rangeL W) | apply/(In_rangeR W)]. +move/mem_seqP=>H; rewrite (In_range_valid H) inE /=. +by case/In_rangeUn: H=>/mem_seqP -> //; rewrite orbT. +Qed. + +Lemma rangePtK k v : C k -> range (U:=U) (pts k v) = [:: v]. +Proof. by move=>C'; rewrite /range assocsPt C'. Qed. + +Lemma rangePt x k v : C k -> x \in range (U:=U) (pts k v) = (x == v). +Proof. by move=>C'; rewrite /range assocsPt C' inE. Qed. + +Lemma rangePtUn k v f : + range (pts k v \+ f) =i + [pred x | valid (pts k v \+ f) & (v == x) || (x \in range f)]. +Proof. +move=>x; rewrite rangeUn !inE; case W : (valid _)=>//=. +by rewrite rangePt ?(validPtUn_cond W) // eq_sym. +Qed. + +Lemma rangePtUnK k v f : + valid (pts k v \+ f) -> + all (ord k) (dom f) -> + range (pts k v \+ f) = v :: range f. +Proof. by move=>W H; rewrite /range assocsPtUn. Qed. + +Lemma rangeUnPtK k v f : + valid (f \+ pts k v) -> + all (ord^~ k) (dom f) -> + range (f \+ pts k v) = rcons (range f) v. +Proof. by move=>W H; rewrite /range assocsUnPt // map_rcons. Qed. + +Lemma uniq_rangeUn f1 f2 : + valid (f1 \+ f2) -> + uniq (range (f1 \+ f2)) = uniq (range f1 ++ range f2). +Proof. +move=>W; apply/esym; case: uniq_rangeP=>H; last first. +- apply/negP; rewrite cat_uniq=>/and3P [H1 /hasP H2 H3]. + elim: H=>k1 k2 v /InUn [] F1 /InUn []; move: F1. + - by move/uniq_rangeP: H1; apply. + - by move/mem_range=>F1 /mem_range F2; elim: H2; exists v. + - by move/mem_range=>F1 /mem_range F2; elim: H2; exists v. + by move/uniq_rangeP: H3; apply. +rewrite cat_uniq; apply/and3P; split; last 1 first. +- by apply/uniq_rangeP=>k1 k2 v F1 F2; apply: (H k1 k2 v); apply/InR. +- by apply/uniq_rangeP=>k1 k2 v F1 F2; apply: (H k1 k2 v); apply/InL. +case: hasP=>//; case=>x /mem_rangeX [k1 H1] /mem_rangeX [k2 H2]. +have [G1 G2] : (k1, x) \In f1 \+ f2 /\ (k2, x) \In f1 \+ f2. +- by split; [apply/InR|apply/InL]. +rewrite -(H k1 k2 x G1 G2) in H2. +by move: (dom_inNR W (In_dom H1)); rewrite (In_dom H2). +Qed. + +Lemma uniq_rangePtUn k v f : + valid (pts k v \+ f) -> + uniq (range (pts k v \+ f)) = uniq (v :: range f). +Proof. by move=>W; rewrite uniq_rangeUn // rangePtK // (validPtUn_cond W). Qed. + +Lemma uniq_rangeL f1 f2 : + valid (f1 \+ f2) -> uniq (range (f1 \+ f2)) -> uniq (range f1). +Proof. by move=>W; rewrite uniq_rangeUn // cat_uniq; case/and3P. Qed. + +Lemma uniq_rangeR f1 f2 : + valid (f1 \+ f2) -> uniq (range (f1 \+ f2)) -> uniq (range f2). +Proof. by move=>W; rewrite uniq_rangeUn // cat_uniq; case/and3P. Qed. + +Lemma uniq_rangeF k f : uniq (range f) -> uniq (range (free f k)). +Proof. +case W: (valid f); last first. +- by move/negbT/invalidE: W=>->; rewrite free_undef !range_undef. +case D : (k \in dom f); last by move/negbT/dom_free: D=>E; rewrite -{1}E. +by case: (um_eta D) W=>x [_] E; rewrite {1 2}E; apply: uniq_rangeR. +Qed. + +End DecidableRange. + + +(********************) +(* Map monotonicity *) +(********************) + +Section MapMonotonicity. +Variables (K V : ordType) (C : pred K) (U : union_map_class C V). +Implicit Types f : U. + +Definition um_mono f := sorted ord (range f). +Definition um_mono_lt f := forall k k' v v', + (k, v) \In f -> (k', v') \In f -> ord k k' -> ord v v'. +Definition um_mono_ltE f := forall k k' v v', + (k, v) \In f -> (k', v') \In f -> ord k k' <-> ord v v'. +Definition um_mono_leE f := forall k k' v v', + (k, v) \In f -> (k', v') \In f -> oleq k k' <-> oleq v v'. + +Lemma ummonoP f : reflect (um_mono_lt f) (um_mono f). +Proof. +rewrite /um_mono/um_mono_lt/range. +apply/(equivP idP); elim/um_indf: f=>[||k v f IH W P]. +- by rewrite assocs_undef; split=>// _ ???? /In_undef. +- by rewrite assocs0; split=>// _ ???? /In0. +rewrite assocsPtUn ?(order_path_min (@trans _) P) //=; split=>H; last first. +- rewrite path_min_sorted; first by apply/IH=>??????; apply: H; apply/InR. + apply/allP=>x /mapP [[y w]] /mem_seqP/umemE X ->. + by apply: H (path_mem (@trans K) P (In_dom X)); [apply/InPtUnL|apply/InR]. +move=>x x' w w'; rewrite !InPtUnE //. +case=>[[->->]|H1]; case=>[[->->]|H2]; first by rewrite irr. +- suff: w' \in [seq i.2 | i <- assocs f] by move/(path_mem (@trans V) H). + by move/umemE/mem_seqP: H2=>H2; apply/mapP; exists (x',w'). +- by move/In_dom/(path_mem (@trans K) P): H1; case: ordP. +by move/path_sorted/IH: H; apply. +Qed. + +Lemma ummono_undef : um_mono undef. +Proof. by apply/ummonoP=>???? /In_undef. Qed. + +Lemma ummono0 : um_mono Unit. +Proof. by apply/ummonoP=>???? /In0. Qed. + +Lemma ummonoUnL f1 f2 : valid (f1 \+ f2) -> um_mono (f1 \+ f2) -> um_mono f1. +Proof. by move=>W /ummonoP H; apply/ummonoP=>??????; apply: H; apply/InL. Qed. + +Lemma ummonoUnR f1 f2 : valid (f1 \+ f2) -> um_mono (f1 \+ f2) -> um_mono f2. +Proof. by rewrite joinC; apply: ummonoUnL. Qed. + +Lemma ummonoPtUn k' v' f : + valid (pts k' v' \+ f) -> + (forall k v, (k, v) \In f -> ord k k' /\ ord v v') -> + um_mono (pts k' v' \+ f) = um_mono f. +Proof. +move=>W H; apply/idP/idP=>/ummonoP X; apply/ummonoP=>x x' w w' Y Y'. +- by apply: X; apply/InR. +case/(InPtUnE _ W): Y'=>[[->->]|]; case/(InPtUnE _ W): Y=>[[->->]|]; +by [rewrite irr|case/H|case/H; case: ordP|apply: X]. +Qed. + +Lemma In_mono_fun k v1 v2 f : + um_mono f -> (v1, k) \In f -> (v2, k) \In f -> v1 = v2. +Proof. +move/ummonoP=>M H1 H2; case: (ordP v1 v2). +- by move/(M _ _ _ _ H1 H2); rewrite irr. +- by move/eqP. +by move/(M _ _ _ _ H2 H1); rewrite irr. +Qed. + +Lemma In_mono_range k f1 f2 : + valid (f1 \+ f2) -> um_mono (f1 \+ f2) -> + k \in range f1 -> k \in range f2 -> false. +Proof. +move=>W M /mem_seqP/In_rangeX [v1 H1] /mem_seqP/In_rangeX [v2 H2]. +have H1' : (v1, k) \In f1 \+ f2 by apply/InL. +have H2' : (v2, k) \In f1 \+ f2 by apply/InR. +rewrite -{v2 H1' H2'}(In_mono_fun M H1' H2') in H2. +move/In_dom: H2; move/In_dom: H1=>H1. +by rewrite (negbTE (dom_inNL W H1)). +Qed. + +Lemma ummono_ltP f : reflect (um_mono_ltE f) (um_mono f). +Proof. +case: ummonoP=>M; constructor; last first. +- by move=>N; elim: M=>x x' y y' X1 X2 /N; apply. +move=>x x' y y' X1 X2. +move: (M _ _ _ _ X1 X2) (M _ _ _ _ X2 X1)=>O1 O2. +split=>// N; case: ordP=>// Y1. +- by move/O2: Y1; case: ordP N. +by move/eqP: Y1=>?; subst x'; rewrite (In_fun X1 X2) irr in N. +Qed. + +Lemma ummono_leP f : um_mono f -> um_mono_leE f. +Proof. +move/ummonoP=>M x x' y y' X1 X2. +move: (M _ _ _ _ X1 X2) (M _ _ _ _ X2 X1)=>O1 O2. +rewrite /oleq (eq_sym x) (eq_sym y). +case: ordP=>Y1; case: ordP=>Y2 //=. +- by move/O2: Y1; rewrite (eqP Y2) irr. +- by move/O2: Y1; case: ordP Y2. +- by rewrite (eqP Y1) in X2; rewrite (In_fun X1 X2) irr in Y2. +by move/O1: Y1; case: ordP Y2. +Qed. + +Lemma index_mem_dom_range f k t : + (k, t) \In f -> uniq (range f) -> index k (dom f) = index t (range f). +Proof. +rewrite /range assocs_dom. +elim/um_indf: f k t=>[||k' t' f IH W /(order_path_min (@trans K)) P] k t. +- by move/In_undef. +- by move/In0. +rewrite assocsPtUn // !map_cons /= InPtUnE //=. +case=>[[->->]|H1 H2]; first by rewrite !eq_refl. +case: eqP (In_dom H1) W=>[-> D|_ _ W]. +- by rewrite validPtUn D !andbF. +case: eqP (In_range H1) H2=>[-> /mem_seqP ->//|_ _ /andP [H2 H3]]. +by rewrite (IH _ _ H1). +Qed. + +Lemma index_dom_range_mem f k t : + index k (dom f) = index t (range f) -> + index k (dom f) != size (dom f) -> (k, t) \In f. +Proof. +rewrite /range assocs_dom. +elim/um_indf: f k t=>[||k' t' f IH W /(order_path_min (@trans K)) P] k t. +- by rewrite assocs_undef. +- by rewrite assocs0. +rewrite assocsPtUn // map_cons /=. +case: eqP=>[|_]; first by case: eqP=>// <-<- _ _; apply/InPtUnE=>//; left. +case: eqP=>// _ [H1]; rewrite eqSS=>H2. +by apply/InPtUnE=>//; right; apply: IH H1 H2. +Qed. + +End MapMonotonicity. + + +(**********************) +(* um_foldl, um_foldr *) +(**********************) + +(* Induction lemmas over PCMs in the proofs *) + +Section FoldDefAndLemmas. +Variables (K : ordType) (C : pred K) (V R : Type) (U : union_map_class C V). +Implicit Type f : U. + +Definition um_foldl (a : R -> K -> V -> R) (z0 d : R) f := + if valid f then foldl (fun z kv => a z kv.1 kv.2) z0 (assocs f) else d. + +Definition um_foldr (a : K -> V -> R -> R) (z0 d : R) f := + if valid f then foldr (fun kv z => a kv.1 kv.2 z) z0 (assocs f) else d. + +Lemma umfoldl_undef a z0 d : um_foldl a z0 d undef = d. +Proof. by rewrite /um_foldl valid_undef. Qed. + +Lemma umfoldr_undef a z0 d : um_foldr a z0 d undef = d. +Proof. by rewrite /um_foldr valid_undef. Qed. + +Lemma umfoldl0 a z0 d : um_foldl a z0 d Unit = z0. +Proof. by rewrite /um_foldl assocs0 valid_unit. Qed. + +Lemma umfoldr0 a z0 d : um_foldr a z0 d Unit = z0. +Proof. by rewrite /um_foldr assocs0 valid_unit. Qed. + +Lemma umfoldlPt a (z0 d : R) k v : + um_foldl a z0 d (pts k v) = + if C k then a z0 k v else d. +Proof. by rewrite /um_foldl validPt assocsPt; case: (C k). Qed. + +Lemma umfoldrPt a (z0 d : R) k v : + um_foldr a z0 d (pts k v) = + if C k then a k v z0 else d. +Proof. by rewrite /um_foldr validPt assocsPt; case: (C k). Qed. + +Lemma umfoldlPtUn a k v z0 d f : + valid (pts k v \+ f) -> all (ord k) (dom f) -> + um_foldl a z0 d (pts k v \+ f) = um_foldl a (a z0 k v) d f. +Proof. by move=>W H; rewrite /um_foldl /= W (validR W) assocsPtUn. Qed. + +Lemma umfoldrPtUn a k v z0 d f : + valid (pts k v \+ f) -> all (ord k) (dom f) -> + um_foldr a z0 d (pts k v \+ f) = a k v (um_foldr a z0 d f). +Proof. by move=>W H; rewrite /um_foldr W (validR W) assocsPtUn. Qed. + +Lemma umfoldlUnPt a k v z0 d f : + valid (f \+ pts k v) -> all (ord^~ k) (dom f) -> + um_foldl a z0 d (f \+ pts k v) = a (um_foldl a z0 d f) k v. +Proof. +by move=>W H; rewrite /um_foldl W (validL W) assocsUnPt // -cats1 foldl_cat. +Qed. + +Lemma umfoldrUnPt a k v z0 d f : + valid (f \+ pts k v) -> all (ord^~ k) (dom f) -> + um_foldr a z0 d (f \+ pts k v) = um_foldr a (a k v z0) d f. +Proof. +by move=>W H; rewrite /um_foldr W (validL W) assocsUnPt // -cats1 foldr_cat. +Qed. + +Lemma umfoldlUn a z0 d f1 f2 : + valid (f1 \+ f2) -> {in dom f1 & dom f2, forall x y, ord x y} -> + um_foldl a z0 d (f1 \+ f2) = um_foldl a (um_foldl a z0 d f1) d f2. +Proof. +move: f2; apply: um_indb=>[W H|W H|k v f2 IH W' P W H]. +- by rewrite join_undef !umfoldl_undef. +- by rewrite unitR umfoldl0. +rewrite -(joinC f2) joinA in W *; rewrite umfoldlUnPt //; last first. +- apply/allP=>x; rewrite domUn inE (validL W). + case/orP=>[/H|]; last by apply: P. + by apply; rewrite domPtUn inE joinC W' eq_refl. +rewrite umfoldlUnPt ?(validAR W) //; last by apply/allP. +rewrite (IH (validL W)) // => k1 k2 D1 D2; apply: H D1 _. +by rewrite domPtUn inE joinC W' D2 orbT. +Qed. + +Lemma umfoldrUn a z0 d f1 f2 : + valid (f1 \+ f2) -> {in dom f1 & dom f2, forall x y, ord x y} -> + um_foldr a z0 d (f1 \+ f2) = um_foldr a (um_foldr a z0 d f2) d f1. +Proof. +move: f1; apply: um_indf=>[W H|W H|k v f1 IH W' P W H]. +- by rewrite undef_join !umfoldr_undef. +- by rewrite unitL umfoldr0. +rewrite -!joinA in W *; rewrite umfoldrPtUn //. +- rewrite umfoldrPtUn ?(order_path_min (@trans K) P) // (IH (validR W)) //. + by move=>k1 k2 D1; apply: H; rewrite domPtUn inE W' D1 orbT. +apply/allP=>x; rewrite domUn inE (validR W) /=. +case/orP; last by apply: H; rewrite domPtUn inE W' eq_refl. +by apply/allP/(order_path_min (@trans K) P). +Qed. + +Lemma umfoldlPtUnE v2 v1 a t (z0 d : R) f : + (forall r, a r t v1 = a r t v2) -> + um_foldl a z0 d (pts t v1 \+ f) = + um_foldl a z0 d (pts t v2 \+ f). +Proof. +move=>H. +case W : (valid (pts t v1 \+ f)); last first. +- move/invalidE: (negbT W)=>->; rewrite umfoldl_undef. + rewrite (validPtUnE v2) in W. + by move/invalidE: (negbT W)=>->; rewrite umfoldl_undef. +elim/um_indf: f z0 W=>[|z0|k v f IH V' P z0 W]; +rewrite ?join_undef ?unitR ?umfoldlPt ?(H z0) //. +case: (ordP k t) W=>E W; last 2 first. +- by move/validAL: W; rewrite (eqP E) (validPtUnE v) validUnEb um_unitbPt. +- have D w : all (ord t) (dom (pts k w \+ f)). + - apply/allP=>x; rewrite domPtUn inE=>/andP [_] /orP [/eqP <-|] //. + by apply: path_mem (@trans K) _; apply: ord_path (@trans K) E P. + by rewrite !(umfoldlPtUn a (k:=t)) ?(validPtUnE v1) // H. +have D w : all (ord k) (dom (pts t w \+ f)). +- apply/allP=>x; rewrite domPtUn inE=>/andP [_] /orP [/eqP <-|] //. + by apply: path_mem (@trans K) P. +rewrite !(joinCA _ (pts k v)) in W *. +rewrite !(umfoldlPtUn a (k:=k)) // ?IH ?(validR W) //. +by rewrite joinCA (validPtUnE v1) joinCA. +Qed. + +Lemma umfoldlUnPtE v2 v1 a t (z0 d : R) f : + (forall r, a r t v1 = a r t v2) -> + um_foldl a z0 d (f \+ pts t v1) = + um_foldl a z0 d (f \+ pts t v2). +Proof. by rewrite !(joinC f); apply: umfoldlPtUnE. Qed. + +Lemma umfoldl_defE a z0 d1 d2 f : + valid f -> um_foldl a z0 d1 f = um_foldl a z0 d2 f. +Proof. +move: f z0; elim/um_indf=>[z0|z0|k v f IH W /(order_path_min (@trans K)) P z0 _]; +by rewrite ?valid_undef ?umfoldl0 // !umfoldlPtUn // IH // (validR W). +Qed. + +Lemma umfoldl_ind (P : R -> Prop) a z0 d f : + valid f -> P z0 -> + (forall z0 k v, (k, v) \In f -> P z0 -> P (a z0 k v)) -> + P (um_foldl a z0 d f). +Proof. +move=>W H1 H2; elim/um_indf: f z0 W H1 H2=>[||k v f IH W O] z0; +rewrite ?valid_undef ?umfoldl0 // => _ H1 H2; rewrite umfoldlPtUn //; +last by apply: order_path_min O; apply: trans. +apply: IH (validR W) _ _; first by apply: H2 (InPtUnL W) H1. +by move=>z1 k0 v0 F; apply: H2 (InR W F). +Qed. + +Lemma umfoldr_ind (P : R -> Prop) a z0 d f : + valid f -> P z0 -> + (forall z0 k v, (k, v) \In f -> P z0 -> P (a k v z0)) -> + P (um_foldr a z0 d f). +Proof. +move=>W H1 H2; elim/um_indb: f z0 W H1 H2=>[||k v f IH W /allP O] z0; +rewrite ?valid_undef ?umfoldr0 // => _ H1 H2. +rewrite joinC umfoldrUnPt //; rewrite joinC in W. +apply: IH (validR W) _ _; first by apply: H2 (InPtUnL W) H1. +by move=>z1 k0 v0 F; apply: H2 (InR W F). +Qed. + +End FoldDefAndLemmas. + +Section PCMFold. +Variables (K : ordType) (C : pred K) (V : Type) (U : union_map_class C V). +Variables (R : pcm) (a : R -> K -> V -> R). +Hypothesis frame : forall x y k v, a (x \+ y) k v = a x k v \+ y. + +Lemma umfoldl0_frame z0 d (f : U) : + valid f -> um_foldl a z0 d f = um_foldl a Unit d f \+ z0. +Proof. +move=>W; elim/um_indf: f W d z0 + =>[||k v f IH _ /(order_path_min (@trans K)) P] W d z0. +- by rewrite valid_undef in W. +- by rewrite !umfoldl0 unitL. +rewrite !umfoldlPtUn // IH 1?[in X in _ = X]IH ?(validR W) //. +by rewrite -joinA -frame unitL. +Qed. + +Lemma umfoldlUn_frame z0 d (f1 f2 : U) : + valid (f1 \+ f2) -> + um_foldl a z0 d (f1 \+ f2) = + um_foldl a Unit d f1 \+ um_foldl a Unit d f2 \+ z0. +Proof. +move=>W; rewrite /um_foldl W (validL W) (validR W). +set b := fun z kv => _. +have X x y kv : b (x \+ y) kv = b x kv \+ y by rewrite /b frame. +rewrite (foldl_perm X _ (assocs_perm W)). +rewrite foldl_cat -{1}[z0]unitL (foldl_init X). +rewrite (foldl_init X) -{1}[foldl b _ (assocs f1)]unitL. +by rewrite (foldl_init X) -!joinA joinCA. +Qed. + +End PCMFold. + +(* Special notation for boolean predicates over K*V *) + +Notation "[ 'pts' k v | E ]" := + (fun kv => let '(k, v) := kv in E%B) + (at level 0, k ident, v ident, format "[ 'pts' k v | E ]"). +Notation "[ 'pts' k ( v : V ) | E ]" := + (fun kv : _*V =>let '(k, v) := kv in E%B) + (at level 0, k ident, v ident, only parsing). +Notation "[ 'pts' ( k : K ) v | E ]" := + (fun kv : K*_ => let '(k, v) := kv in E%B) + (at level 0, k ident, v ident, only parsing). +Notation "[ 'pts' ( k : K ) ( v : V ) | E ]" := + (fun kv : K*V => let '(k, v) := kv in E%B) + (at level 0, k ident, v ident, only parsing). + + +(********) +(* omap *) +(********) + +(* Combines map and filter by having the filtering function *) +(* return an option *) + +Section OMapDefLemmas. +Variables (K : ordType) (C : pred K) (V V' : Type). +Variables (U : union_map_class C V) (U' : union_map_class C V'). +Implicit Types (f : U). + +Definition omap (a : K * V -> option V') f : U' := + if valid f then + foldl (fun z kv => if a kv is Some v' then z \+ pts kv.1 v' else z) + Unit (assocs f) + else undef. + +Lemma omap_undef a : omap a undef = undef. +Proof. by rewrite /omap valid_undef. Qed. + +Lemma omap0 a : omap a Unit = Unit. +Proof. by rewrite /omap valid_unit assocs0. Qed. + +Lemma omapPt a k v : + omap a (pts k v) = + if C k then + if a (k, v) is Some w then pts k w else Unit + else undef. +Proof. +rewrite /omap validPt assocsPt; case C': (C k)=>//=. +by case: (a (k, v))=>// x; rewrite unitL. +Qed. + +Lemma omapUn a f1 f2 : + valid (f1 \+ f2) -> + omap a (f1 \+ f2) = omap a f1 \+ omap a f2. +Proof. +move=>W; rewrite /omap W (validL W) (validR W); set b := fun z kv => _. +have H x y kv : b (x \+ y) kv = b x kv \+ y. +- by rewrite /b; case: (a kv)=>// z; rewrite joinAC. +rewrite (foldl_perm H _ (assocs_perm W)) foldl_cat. +by rewrite joinC -(foldl_init H) unitL. +Qed. + +Lemma omapPtUn a k v f : + omap a (pts k v \+ f) = + if valid (pts k v \+ f) then + if a (k, v) is Some v' then pts k v' \+ omap a f else omap a f + else undef. +Proof. +case: ifP=>X; last first. +- by move/invalidE: (negbT X)=>->; rewrite omap_undef. +rewrite omapUn // omapPt (validPtUn_cond X). +by case: (a _)=>//; rewrite unitL. +Qed. + +Lemma omapUnPt a k v f : + omap a (f \+ pts k v) = + if valid (f \+ pts k v) then + if a (k, v) is Some v' then omap a f \+ pts k v' else omap a f + else undef. +Proof. +rewrite joinC omapPtUn; case: ifP=>// W. +by case: (a _)=>// x; rewrite joinC. +Qed. + +Lemma eq_in_omap a1 a2 f : + (forall kv, kv \In f -> a1 kv = a2 kv) -> + omap a1 f = omap a2 f. +Proof. +elim/um_indf: f=>[||k v f IH W P H]; rewrite ?omap_undef ?omap0 //. +rewrite !omapPtUn W; move/H: (InPtUnL W)=>->. +by case E2: (a2 _)=>[b2|]; (rewrite IH; last by move=>kv /(InR W)/H). +Qed. + +Lemma dom_omap_sub a f : {subset dom (omap a f) <= dom f}. +Proof. +elim/um_indf: f=>[||k v f IH W P] x. +- by rewrite omap_undef dom_undef. +- by rewrite omap0 dom0. +rewrite omapPtUn W domPtUn W !inE /=. +case E : (a _)=>[b|]; last by move/IH=>->; rewrite orbT. +rewrite domPtUn inE; case/andP=>W2. +by case/orP=>[->//|/IH ->]; rewrite orbT. +Qed. + +Lemma valid_omap a f : valid (omap a f) = valid f. +Proof. +elim/um_indf: f=>[||k v f IH W P]. +- by rewrite omap_undef !valid_undef. +- by rewrite omap0 !valid_unit. +rewrite omapPtUn W; case X : (a _)=>[b|]; last by rewrite IH (validR W). +rewrite validPtUn (validPtUn_cond W) IH (validR W). +by apply: contra (notin_path P); apply: dom_omap_sub. +Qed. + +Lemma In_omap_raw a k v f : + (k, v) \In omap a f <-> + exists w, [/\ C k, valid (omap a f), + (k, w) \In f & a (k, w) = Some v]. +Proof. +elim/um_indf: f k v=>[||x w f IH W P] k v. +- by rewrite omap_undef; split=>[/In_undef|[w][]] //; rewrite valid_undef. +- by rewrite omap0; split=>[/In0|[w][??] /In0]. +split=>[H|[w1][C' V1 H E]]. +- move: (dom_valid (In_dom H)); rewrite omapPtUn W in H *. + case E: (a (x, w)) H=>[z|] H W1; last first. + - by case/IH: H=>w1 []; exists w1; split=>//; apply/InR. + rewrite InPtUnE // in H *. + case; first by case=>->->; exists w; rewrite (validPtUn_cond W1). + by case/IH=>w1 []; exists w1; split=>//; apply/InR. +rewrite omapPtUn W in V1 *; move/(InPtUnE _ W): H=>H. +case: H E V1; first by case=>->->-> V1; apply/InPtUnE=>//; left. +case: (a (x, w))=>[b|] H E V1; last by apply/IH; exists w1. +by rewrite InPtUnE //; right; apply/IH; exists w1; rewrite (validR V1). +Qed. + +Lemma In_omapX a k v f : + (k, v) \In omap a f <-> + exists2 w, (k, w) \In f & a (k, w) = Some v. +Proof. +rewrite In_omap_raw; split; first by case=>w []; exists w. +case=>w H X; exists w; split=>//. +- by move/In_dom/dom_cond: H. +by rewrite valid_omap //; move/In_dom/dom_valid: H. +Qed. + +Lemma In_omap a k v w f : + (k, w) \In f -> a (k, w) = Some v -> + (k, v) \In omap a f. +Proof. by move=>H1 H2; apply/In_omapX=>//; exists w. Qed. + +Lemma In_dom_omap a k f : + reflect (exists v w, (k, w) \In f /\ a (k, w) = Some v) + (k \in dom (omap a f)). +Proof. +case: In_domX=>H; constructor. +- by case: H=>v /In_omapX [w H1 H2]; exists v, w. +by case=>v [w][X /(In_omap X) Y]; elim: H; exists v. +Qed. + +(* if the map is total, we have some stronger lemmas *) +Lemma dom_omap_some a f : + (forall x, x \In f -> oapp predT false (a x)) -> dom (omap a f) = dom f. +Proof. +move=>H; apply/domE=>k; apply/idP/idP=>/In_domX [w]. +- by case/In_omapX=>// v /In_dom. +case A : (a (k, w)) {H} (H (k, w))=>[v|] // H1 H2; last by move/H1: H2. +suff : (k, v) \In omap a f by apply: In_dom. +by apply/In_omapX=>//; exists w. +Qed. + +(* if the map is total on f1, f2, don't need validity condition for union *) +Lemma omapUn_some a f1 f2 : + (forall x, x \In f1 -> oapp predT false (a x)) -> + (forall x, x \In f2 -> oapp predT false (a x)) -> + omap a (f1 \+ f2) = omap a f1 \+ omap a f2. +Proof. +move=>H1 H2; have Ev : valid (omap a f1 \+ omap a f2) = valid (f1 \+ f2). +- by rewrite !validUnAE !valid_omap // !dom_omap_some. +case W : (valid (f1 \+ f2)); first by rewrite omapUn. +move: W (W); rewrite -{1}Ev=>/negbT/invalidE -> /negbT/invalidE ->. +by rewrite omap_undef. +Qed. + +Lemma path_omap a f x : path ord x (dom f) -> path ord x (dom (omap a f)). +Proof. +apply: subseq_order_path; first by apply: trans. +apply: (sorted_subset_subseq (leT := ord)); last by apply: dom_omap_sub. +- by apply: irr. +- by apply: trans. +- by apply: sorted_dom. +by apply: sorted_dom. +Qed. + +Lemma omap_predU (a1 a2 : K*V -> option V') f : + (forall kv, a1 kv = None \/ a2 kv = None) -> + omap a1 f \+ omap a2 f = + omap (fun kv => if a1 kv is Some v' then Some v' + else if a2 kv is Some v' then Some v' + else None) f. +Proof. +move=>E; elim/um_indf: f=>[||k v f IH W /(order_path_min (@trans K)) P]. +- by rewrite !omap_undef undef_join. +- by rewrite !omap0 unitL. +rewrite !omapPtUn W /= -IH. +case E1 : (a1 (k, v))=>[b1|]; case E2 : (a2 (k, v))=>[b2|] //. +- by move: (E (k, v)); rewrite E1 E2; case. +- by rewrite joinA. +by rewrite joinCA. +Qed. + +Lemma domeq2_omap (a : K*V -> option V') f : + (forall kv, kv \In f -> isSome (a kv)) -> + dom_eq2 (omap a f) f. +Proof. +move=>H; rewrite /dom_eq2 valid_omap // eq_refl /=. +apply/eqP; apply: eq_sorted_ord=>// k. +apply/idP/idP; first by case/In_dom_omap=>// x [w][] /In_dom. +case/In_domX=>x X; apply/In_dom_omap=>//. +by case F : (a (k, x)) (H _ X)=>[v|] //; exists v, x. +Qed. + +Lemma dom_omapUn (a : K*V -> option V') f1 f2 : + dom (omap a (f1 \+ f2)) =i + [pred x | valid (f1 \+ f2) & + (x \in dom (omap a f1)) || (x \in dom (omap a f2))]. +Proof. +move=>x; case W: (valid (f1 \+ f2)); last first. +- by move/negbT/invalidE: W=>->; rewrite omap_undef dom_undef. +by rewrite omapUn // domUn !inE -omapUn // valid_omap // W. +Qed. + +Lemma In_range_omapUn (a : K*V -> option V') f1 f2 x : + x \In range (omap a (f1 \+ f2)) <-> + valid (f1 \+ f2) /\ (x \In range (omap a f1) \/ x \In range (omap a f2)). +Proof. +split. +- case W : (valid (f1 \+ f2)); first by rewrite omapUn // => /In_rangeUn. + by move/negbT/invalidE: W=>->; rewrite omap_undef range_undef. +case=>W H; rewrite omapUn //; case: H=>H. +- by apply: In_rangeL H; rewrite -omapUn // valid_omap. +by apply: In_rangeR H; rewrite -omapUn // valid_omap. +Qed. + +Lemma omap_morph_ax a : morph_axiom (@sepT _) (omap a). +Proof. +split=>[|x y W _]; first by rewrite omap0. +by rewrite -omapUn // valid_omap. +Qed. + +Canonical omap_morph a := Morphism' (omap a) (omap_morph_ax a). + +Lemma valid_omapPtUn a k v f : + [&& C k, valid f & k \notin dom f] -> + valid (pts k v \+ omap a f). +Proof. +move/and3P => [H1 H2 H3]. +rewrite validPtUn H1 valid_omap // H2 /=. +by apply: contra H3; apply: dom_omap_sub. +Qed. + +Lemma valid_omapUnPt a k v f : + [&& C k, valid f & k \notin dom f] -> + valid (omap a f \+ pts k v). +Proof. by move=>H; rewrite joinC; apply: valid_omapPtUn. Qed. + +Lemma valid_omapUn a1 a2 f1 f2 : + valid (f1 \+ f2) -> valid (omap a1 f1 \+ omap a2 f2). +Proof. +move=>W; rewrite validUnAE !valid_omap // (validL W) (validR W) /=. +apply/allP=>z /In_dom_omap [] //= x1 [w1][/In_dom /= D1 _]. +apply/negP=>/In_dom_omap [] //= x2 [w2][/In_dom /=]. +by move/(dom_inNR W)/negbTE: D1=>->. +Qed. + +Lemma find_omap a k f : + find k (omap a f) = + if find k f is Some v then a (k, v) else None. +Proof. +case E1 : (find k (omap a f))=>[b|]. +- by move/In_find: E1=>/In_omapX [] // x /In_find ->. +move/find_none/negP: E1=>E1. +case E2 : (find k f)=>[c|] //; move/In_find: E2=>E2. +case E3: (a (k, c))=>[d|] //; elim: E1. +by apply/In_dom_omap=>//; exists d, c. +Qed. + +End OMapDefLemmas. + +Arguments omap [K C V V' U U']. +Arguments dom_omap_sub [K C V V' U U' a f] _. +Prenex Implicits dom_omap_sub. + +Section OmapComp. +Variables (K : ordType) (C : pred K) (V1 V2 V3 : Type). +Variables (U1 : union_map_class C V1). +Variables (U2 : union_map_class C V2). +Variables (U3 : union_map_class C V3). +Variables (a1 : K * V1 -> option V2) (a2 : K * V2 -> option V3). + +Lemma omap_comp (f : U1) : + omap a2 (omap a1 f : U2) = + omap (fun x => obind (fun v => a2 (x.1, v)) (a1 x)) f :> U3. +Proof. +elim/um_indf: f=>[||k v f IH W P]; rewrite ?omap_undef ?omap0 //. +rewrite !omapPtUn W; case: (a1 (k, v))=>[w|//]. +rewrite omapPtUn validPtUn (validPtUn_cond W) valid_omap // (validR W) IH. +by rewrite (contra _ (validPtUnD W)) //; apply: dom_omap_sub. +Qed. + +End OmapComp. + +(* special notation for some common variants of omap *) + +(* when we don't supply the key *) +Notation omapv f := (omap (f \o snd)). +(* when the don't supply the key and the map is total *) +Notation mapv f := (omapv (Some \o f)). + +Section OmapvLemmas. +Variables (K : ordType) (C : pred K) (V : Type). +Variables (U : union_map_class C V). + +Variables (V1 V2 V3 : Type). +Variables (U1 : union_map_class C V1). +Variables (U2 : union_map_class C V2). +Variables (U3 : union_map_class C V3). +Variables (a1 : V1 -> option V2) (a2 : V2 -> option V3). + +Lemma omapv_comp (f : U1) : + omapv a2 (omapv a1 f : U2) = omapv (obind a2 \o a1) f :> U3. +Proof. by apply: omap_comp. Qed. + +End OmapvLemmas. + +Section MapvLemmas. +Variables (K : ordType) (C : pred K) (V : Type). +Variables (U : union_map_class C V). + +Lemma mapv_id (f : U) : mapv id f = f. +Proof. +elim/um_indf: f=>[||k v f IH W P]. +- by rewrite omap_undef. +- by rewrite omap0. +by rewrite omapPtUn W /= IH. +Qed. + +(* A bit of a nicer-looking statements can be given for total maps. *) +(* It's the same theorem as omap_comp, but avoids obinds. *) + +Variables (V1 V2 V3 : Type). +Variables (U1 : union_map_class C V1). +Variables (U2 : union_map_class C V2). +Variables (U3 : union_map_class C V3). +Variables (a1 : V1 -> V2) (a2 : V2 -> option V3). + +Lemma mapv_comp (f : U1) : omapv a2 (mapv a1 f : U2) = omapv (a2 \o a1) f :> U3. +Proof. by apply: omap_comp. Qed. + +Lemma In_mapv (f : U1) k x : + injective a1 -> (k, a1 x) \In (mapv a1 f : U2) <-> (k, x) \In f. +Proof. +by move=>G; rewrite In_omapX //=; split; [case=>w H [] /G <-| exists x]. +Qed. + +End MapvLemmas. + + +Section OmapFreeUpd. +Variables (K : ordType) (C : pred K) (V V' : Type). +Variables (U : union_map_class C V) (U' : union_map_class C V'). +Variables (a : K * V -> option V'). + +Lemma omapFE k (f : U) : + free f k = omap (fun kv => if kv.1 == k then None else Some kv.2) f. +Proof. +case W : (valid f); last first. +- by move/negbT/invalidE: W=>->; rewrite free_undef omap_undef. +apply: umem_eq. +- by rewrite validF. +- by rewrite valid_omap. +case=>x w; rewrite InF In_omapX /= validF W; split. +- by case=>_ /negbTE ->; exists w. +by case=>w' H; case: eqP=>// N [<-]. +Qed. + +Lemma omapF k (f : U) : + omap a (free f k) = free (omap a f) k :> U'. +Proof. +case W : (valid f); last first. +- by move/negbT/invalidE: W=>->; rewrite free_undef omap_undef free_undef. +apply: umem_eq. +- by rewrite valid_omap // validF. +- by rewrite validF valid_omap. +case=>x v; split. +- case/In_omapX=>v' /InF /= [_ N M E]. + apply/InF; rewrite validF valid_omap //=; split=>//. + by apply/In_omapX; exists v'. +case/InF=>/= _ N /In_omapX [v'] M E. +by apply/In_omapX; exists v'=>//; apply/InF; rewrite validF. +Qed. + +Lemma omapU k (v : V) (f : U) : + omap a (upd k v f) = + if C k then + if a (k, v) is Some v' then upd k v' (omap a f) + else free (omap a f) k : U' + else undef. +Proof. +case: ifP=>// W; last first. +- have: ~~ valid (upd k v f) by rewrite validU W. + by move/invalidE=>->; apply omap_undef. +case H: (valid f); last first. +- move: H; move/negbT/invalidE=>->; rewrite upd_undef omap_undef free_undef. + by case: (a (k, v)) =>// a0; rewrite upd_undef. +rewrite upd_eta // omapPtUn validPtUn W validF H domF inE eq_refl /=. +case E: (a (k, v))=>[xa|]; last by rewrite omapF. +by rewrite upd_eta omapF. +Qed. + +End OmapFreeUpd. + +(* Other extra lemmas on omap that require different U's *) + +Section OmapExtras. +Variables (K : ordType) (C : pred K) (V V' : Type). +Variables (U : union_map_class C V) (U' : union_map_class C V'). + +Lemma omap_some (a : K * V -> option V) (f : U) : + (forall kv, kv \In f -> a kv = Some kv.2) -> + omap a f = f. +Proof. by move/eq_in_omap=>->; apply: mapv_id. Qed. + +Lemma omap_none (a : K * V -> option V') (f : U) : + valid f -> + (forall kv, kv \In f -> a kv = None) -> + omap a f = Unit :> U'. +Proof. +move=>W /eq_in_omap ->. +elim/um_indf: f W=>[||k v f IH W P]; rewrite ?valid_undef ?omap0 // => _. +by rewrite omapPtUn W IH // (validR W). +Qed. + +Lemma omap_noneR (a : K * V -> option V') (f : U) : + omap a f = Unit :> U' -> forall kv, kv \In f -> a kv = None. +Proof. +elim/um_indf: f=>[H|H|k v f IH W P O] kv. +- by move/In_undef. +- by move/In0. +move: O; rewrite omapPtUn W. +case E: (a (k, v)). +- by move=>/unitbP; rewrite um_unitbPtUn. +move=>O /InUn [/InPt [->]|] //. +by apply: IH =>//; move: W; rewrite validPtUn; case/and3P. +Qed. + +End OmapExtras. + + +(*************) +(* um_filter *) +(*************) + +(* filter that works over both domain and range at once *) + +Section FilterDefLemmas. +Variables (K : ordType) (C : pred K) (V : Type) (U : union_map_class C V). +Implicit Type f : U. +Implicit Type p q : pred (K * V). + +Definition um_filter p f : U := + omap (fun kv => if p kv then Some kv.2 else None) f. + +Lemma umfilt0 p : um_filter p Unit = Unit. +Proof. by rewrite /um_filter omap0. Qed. + +Lemma umfilt_undef p : um_filter p undef = undef. +Proof. by rewrite /um_filter omap_undef. Qed. + +Lemma umfiltPt p k v : + um_filter p (pts k v) = + if C k then + if p (k, v) then pts k v else Unit + else undef. +Proof. by rewrite /um_filter omapPt; case: (ifP (p (k, v))). Qed. + +Lemma umfiltUn p f1 f2 : + valid (f1 \+ f2) -> + um_filter p (f1 \+ f2) = um_filter p f1 \+ um_filter p f2. +Proof. by move=>W; rewrite /um_filter omapUn. Qed. + +Lemma umfiltPtUn p k v f : + um_filter p (pts k v \+ f) = + if valid (pts k v \+ f) then + if p (k, v) then pts k v \+ um_filter p f else um_filter p f + else undef. +Proof. by rewrite /um_filter omapPtUn; case: (ifP (p (k, v))). Qed. + +Lemma umfiltUnPt p k v f : + um_filter p (f \+ pts k v) = + if valid (f \+ pts k v) then + if p (k, v) then um_filter p f \+ pts k v else um_filter p f + else undef. +Proof. by rewrite /um_filter omapUnPt; case: (ifP (p (k, v))). Qed. + +Lemma dom_umfilt_subset p f : + {subset dom (um_filter p f) <= dom f}. +Proof. by rewrite /um_filter; apply: dom_omap_sub. Qed. + +Lemma valid_umfilt p f : valid (um_filter p f) = valid f. +Proof. by rewrite /um_filter valid_omap. Qed. + +Lemma In_umfilt x p f : x \In um_filter p f <-> p x /\ x \In f. +Proof. +case: x => k v; rewrite In_omapX; split=>[[w H]|[I H]]. +- by case E: (p (k, w))=>//=; case=><-. +by exists v=>//; rewrite I. +Qed. + +Lemma dom_umfilt p f k : + reflect (exists v, [/\ p (k, v) & (k, v) \In f]) + (k \in dom (um_filter p f)). +Proof. +case: In_domX=>H; constructor. +- by case: H=>v /In_umfilt []; exists v. +by case=>v [Hp Hf]; elim: H; exists v; apply/In_umfilt. +Qed. + +Lemma valid_umfiltUnL f1 f2 p : + valid (f1 \+ f2) -> valid (um_filter p f1 \+ f2). +Proof. +move=>VH; case: validUn=>//; rewrite ?valid_umfilt ?(validL VH) ?(validR VH) //. +by move=>k /dom_umfilt [v1][_ /In_dom] /(dom_inNL VH) /negbTE ->. +Qed. + +Lemma valid_umfiltUnR f1 f2 p : + valid (f1 \+ f2) -> valid (f1 \+ um_filter p f2). +Proof. by rewrite !(joinC f1); apply: valid_umfiltUnL. Qed. + +Lemma valid_umfiltUn p1 p2 f1 f2 : + valid (f1 \+ f2) -> valid (um_filter p1 f1 \+ um_filter p2 f2). +Proof. by move=>W; apply: valid_umfiltUnL (valid_umfiltUnR _ W). Qed. + +Lemma dom_umfiltE p f : + dom (um_filter p f) = + filter (fun k => if find k f is Some v then p (k, v) else false) + (dom f). +Proof. +apply: (@eq_sorted_irr _ ord); [apply: trans|apply: irr|apply: sorted_dom| | ]. +- by apply: sorted_filter; [apply: trans | apply: sorted_dom]. +move=>k; rewrite mem_filter; apply/idP/idP. +- by case/dom_umfilt=>w [H1 H2]; move/In_find: H2 (H2) H1=>-> /In_dom ->->. +case X: (find k f)=>[v|] // /andP [H1 _]; move/In_find: X=>H2. +by apply/dom_umfilt; exists v. +Qed. + +Lemma umfilt_pred0 f : valid f -> um_filter pred0 f = Unit. +Proof. by move=>W; apply: omap_none. Qed. + +Lemma umfilt_predT f : um_filter predT f = f. +Proof. by apply: omap_some. Qed. + +Lemma umfilt_predI p1 p2 f : + um_filter (predI p1 p2) f = um_filter p1 (um_filter p2 f). +Proof. +rewrite /um_filter omap_comp; apply: eq_in_omap; case=>k v H /=. +by case: (ifP (p2 (k, v))); rewrite ?andbT ?andbF. +Qed. + +Lemma eq_in_umfilt p1 p2 f : + (forall kv, kv \In f -> p1 kv = p2 kv) <-> + um_filter p1 f = um_filter p2 f. +Proof. +split. +- by move=> H; apply /eq_in_omap => kv In; rewrite H. +move=>H x [W][z E]; subst f. +rewrite !umfiltPtUn W -prod_eta in H. +case: ifP H; case: ifP=>// H1 H2 E. +- have: x.1 \in dom (um_filter p2 z). + - by rewrite -E domPtUn inE E eq_refl valid_umfilt (validR W). + by move/dom_umfilt_subset; rewrite (negbTE (validPtUnD W)). +have: x.1 \in dom (um_filter p1 z). +- by rewrite E domPtUn inE -E eq_refl valid_umfilt (validR W). +by move/dom_umfilt_subset; rewrite (negbTE (validPtUnD W)). +Qed. + +Lemma eq_in_umfiltI p1 p2 f : + (forall kv, kv \In f -> p1 kv = p2 kv) -> + um_filter p1 f = um_filter p2 f. +Proof. by move/eq_in_umfilt. Qed. + +(* common use omits the localization part *) +Lemma eq_in_umfiltE p1 p2 f : + p1 =1 p2 -> um_filter p1 f = um_filter p2 f. +Proof. by move=>S; apply/eq_in_umfilt=>kv _; apply: S. Qed. + +Lemma umfiltC p1 p2 f : + um_filter p1 (um_filter p2 f) = um_filter p2 (um_filter p1 f). +Proof. +by rewrite -!umfilt_predI; apply: eq_in_umfiltE=>x /=; rewrite andbC. +Qed. + +Lemma umfilt_predU p1 p2 f : + um_filter (predU p1 p2) f = + um_filter p1 f \+ um_filter (predD p2 p1) f. +Proof. +rewrite /um_filter omap_predU /=. +- apply /eq_in_omap; case=> k v H /=. + by case: (p1 (k, v))=>/=; case: (p2 (k, v))=>/=. +by case=> k v; case: (p1 (k, v))=>/=; [right|left]. +Qed. + +Lemma umfilt_predD p1 p2 f : + subpred p1 p2 -> + um_filter p2 f = um_filter p1 f \+ um_filter (predD p2 p1) f. +Proof. +move=>S; rewrite -umfilt_predU -eq_in_umfilt=>kv _ /=. +by case E: (p1 _)=>//; apply: S E. +Qed. + +Lemma umfilt_dpredU f p q : + subpred p (predC q) -> + um_filter (predU p q) f = um_filter p f \+ um_filter q f. +Proof. +move=>D; rewrite umfilt_predU. +suff : forall kv, kv \In f -> predD q p kv = q kv by move/eq_in_umfilt=>->. +by move=>kv _ /=; case X: (p _)=>//=; move/D/negbTE: X. +Qed. + +Lemma umfiltUnK p f1 f2 : + valid (f1 \+ f2) -> + um_filter p (f1 \+ f2) = f1 -> + um_filter p f1 = f1 /\ um_filter p f2 = Unit. +Proof. +move=>V'; rewrite (umfiltUn _ V') => E. +have W' : valid (um_filter p f1 \+ um_filter p f2). +- by rewrite E; move/validL: V'. +have F : dom (um_filter p f1) =i dom f1. +- move=>x; apply/idP/idP; first by apply: dom_umfilt_subset. + move=>D; move: (D); rewrite -{1}E domUn inE W' /=. + by case/orP=>// /dom_umfilt_subset; case: validUn V'=>// _ _ /(_ _ D) /negbTE ->. +rewrite -{2}[f1]unitR in E F; move/(dom_prec W' E): F=>X. +by rewrite X in E W' *; rewrite (joinxK W' E). +Qed. + +Lemma umfiltU p k v f : + um_filter p (upd k v f) = + if C k then + if p (k, v) then upd k v (um_filter p f) + else free (um_filter p f) k + else undef. +Proof. by rewrite /um_filter omapU //; case: (p (k, v)). Qed. + +Lemma umfiltF p k f : um_filter p (free f k) = free (um_filter p f) k. +Proof. by rewrite /um_filter omapF. Qed. + +Lemma umfilt_morph_ax p : morph_axiom (@sepT _) (um_filter p). +Proof. by rewrite /um_filter; apply: omap_morph_ax. Qed. + +Canonical umfilt_morph p := + Eval hnf in Morphism' (um_filter p) (umfilt_morph_ax p). + +Lemma umfoldl_umfilt R (a : R -> K -> V -> R) (p : pred (K * V)) f z0 d: + um_foldl a z0 d (um_filter p f) = + um_foldl (fun r k v => if p (k, v) then a r k v else r) z0 d f. +Proof. +move: f z0; elim/um_indf=>[||k v f IH W P] z0 /=. +- by rewrite !umfilt_undef !umfoldl_undef. +- by rewrite !umfilt0 !umfoldl0. +have V1 : all (ord k) (dom f) by apply/allP=>x; apply: path_mem (@trans K) P. +have V2 : all (ord k) (dom (um_filter p f)). +- apply/allP=>x /dom_umfilt [w][_] /In_dom. + by apply: path_mem (@trans K) P. +have : valid (um_filter p (pts k v \+ f)) by rewrite valid_umfilt W. +by rewrite !umfiltPtUn W; case: ifP=>E V3; rewrite !umfoldlPtUn // E IH. +Qed. + +Lemma umfilt_mem0 p f : + um_filter p f = Unit -> forall k v, (k, v) \In f -> ~~ p (k, v). +Proof. +rewrite /um_filter => H k v I. +have: ((if p (k, v) then Some v else None) = None). +- by apply: (omap_noneR H). +by case: (p (k, v)). +Qed. + +Lemma umfilt_mem0L p f : + valid f -> (forall k v, (k, v) \In f -> ~~ p (k, v)) -> + um_filter p f = Unit. +Proof. +rewrite /um_filter => VF H. +apply: omap_none => //. +by case=> k v I; rewrite (negbTE (H k v I)). +Qed. + +Lemma umfilt_idemp p f : + um_filter p (um_filter p f) = um_filter p f. +Proof. +rewrite -umfilt_predI; apply/eq_in_umfilt. +by case=>k v H /=; rewrite andbb. +Qed. + +Lemma assocs_umfilt p f : + assocs (um_filter p f) = filter p (assocs f). +Proof. +elim/um_indf: f=>[||k v f IH W /(order_path_min (@trans K)) P]. +- by rewrite umfilt_undef assocs_undef. +- by rewrite umfilt0 assocs0. +rewrite umfiltPtUn W assocsPtUn //=. +case: ifP W=>// H W; rewrite assocsPtUn; first by rewrite IH. +- suff: valid (um_filter p (pts k v \+ f)) by rewrite umfiltPtUn W H. + by rewrite valid_umfilt. +by apply/allP=>x; move/allP: P=>P; move/dom_umfilt_subset/P. +Qed. + +Lemma find_umfilt k p f : + find k (um_filter p f) = + if find k f is Some v then + if p (k, v) then Some v else None + else None. +Proof. +elim/um_indf: f k=>[||k' v f IH W /(order_path_min (@trans K)) P] k. +- by rewrite umfilt_undef find_undef. +- by rewrite umfilt0 find0E. +have W' : valid (pts k' v \+ um_filter p f). +- rewrite validPtUn (validPtUn_cond W)valid_umfilt (validR W). + by apply: contra (validPtUnD W)=>/dom_umfilt [x][_] /In_dom. +rewrite umfiltPtUn W; case: ifP=>E; rewrite !findPtUn2 //; +case: eqP=>// ->; rewrite E // IH. +by move/validPtUnD: W; case: dom_find=>// ->. +Qed. + +Lemma unitb_umfilt p f : unitb f -> unitb (um_filter p f). +Proof. by move/unitbE->; rewrite umfilt0 unitb0. Qed. + +(* filter p x is lower bound for x *) +Lemma umfilt_pleqI x p : [pcm um_filter p x <= x]. +Proof. +exists (um_filter (predD predT p) x); rewrite -umfilt_predU. +by rewrite -{1}[x]umfilt_predT; apply/eq_in_umfilt=>a; rewrite /= orbT. +Qed. + +Hint Resolve umfilt_pleqI : core. + +End FilterDefLemmas. + +Hint Extern 0 [pcm um_filter _ ?X <= ?X] => + apply: umfilt_pleqI : core. + +Notation um_filterk p f := (um_filter (p \o fst) f). +Notation um_filterv p f := (um_filter (p \o snd) f). + +Section FilterKLemmas. +Variables (K : ordType) (C : pred K) (V : Type) (U : union_map_class C V). +Implicit Type f : U. +Implicit Type p q : pred K. + +Lemma dom_umfiltk_filter p f : dom (um_filterk p f) = filter p (dom f). +Proof. +apply: (eq_sorted_irr (leT:=ord)). +- by apply: trans. +- by apply: irr. +- by apply: sorted_dom. +- by apply: sorted_filter; [apply: trans | apply: sorted_dom]. +move=>k; rewrite mem_filter; apply/idP/idP. +- by case/dom_umfilt=>v [/= -> /In_dom]. +by case/andP=>H1 /In_domX [v H2]; apply/dom_umfilt; exists v. +Qed. + +Lemma dom_umfiltk p f : dom (um_filterk p f) =i predI p (mem (dom f)). +Proof. by move=>k; rewrite dom_umfiltk_filter mem_filter. Qed. + +Lemma umfiltk_dom f1 f2 : + valid (f1 \+ f2) -> um_filterk (mem (dom f1)) (f1 \+ f2) = f1. +Proof. +move=>W; apply/umem_eq; first by rewrite valid_umfilt. +- by rewrite (validL W). +case=>k v; rewrite In_umfilt; split=>[|H]. +- by case=>H /InUn [] // /In_dom; rewrite (negbTE (dom_inNL W H)). +by split; [apply: In_dom H | apply: InL]. +Qed. + +Lemma eq_in_umfiltk p1 p2 f : + {in dom f, p1 =1 p2} -> um_filterk p1 f = um_filterk p2 f. +Proof. by move=>H; apply/eq_in_umfilt; case=>k v /In_dom; apply: H. Qed. + +(* filter p x is lower bound for p *) +Lemma subdom_umfiltk p f : {subset dom (um_filterk p f) <= p}. +Proof. by move=>a; rewrite dom_umfiltk; case/andP. Qed. + +Lemma subdom_umfiltkE p f : {subset dom f <= p} <-> um_filterk p f = f. +Proof. +split; last by move=><- a; rewrite dom_umfiltk; case/andP. +by move/eq_in_umfiltk=>->; rewrite umfilt_predT. +Qed. + +(* specific consequence of subdom_umfiltkE *) +Lemma umfiltk_memdomE f : um_filterk (mem (dom f)) f = f. +Proof. by apply/subdom_umfiltkE. Qed. + +Lemma find_umfiltk k (p : pred K) f : + find k (um_filter (p \o fst) f) = if p k then find k f else None. +Proof. by rewrite find_umfilt /=; case: (find _ _)=>[a|]; case: ifP. Qed. + +Lemma subdom_umfiltk0 p f : + valid f -> {subset dom f <= predC p} <-> um_filterk p f = Unit. +Proof. +move=>W; split=>H. +- by rewrite (eq_in_umfiltk (p2:=pred0)) ?umfilt_pred0 // => a /H /negbTE ->. +move=>k X; case: dom_find X H=>// a _ -> _; rewrite umfiltU !inE /=. +case: (ifP (p k)) =>// _ /unitbE. +by case: ifP; [rewrite um_unitbU | rewrite unitb_undef]. +Qed. + +Lemma umfiltkPt p k v : + um_filterk p (pts k v : U) = + if p k then pts k v else if C k then Unit else undef. +Proof. +rewrite ptsU umfiltU umfilt0 free0 /=. +by case: ifP=>//; move/negbT=>N; rewrite upd_condN // if_same. +Qed. + +Lemma umfiltkPtUn p k v f : + um_filterk p (pts k v \+ f) = + if valid (pts k v \+ f) then + if p k then pts k v \+ um_filterk p f else um_filterk p f + else undef. +Proof. +case: ifP=>X; last by move/invalidE: (negbT X)=>->; rewrite umfilt_undef. +rewrite umfiltUn // umfiltPt (validPtUn_cond X) /=. +by case: ifP=>//; rewrite unitL. +Qed. + +Lemma umfiltk_preimUn (q : pred V) f1 f2 : + valid (f1 \+ f2) -> + um_filterk (um_preim q (f1 \+ f2)) f1 = um_filterk (um_preim q f1) f1. +Proof. +move=>v; apply: eq_in_umfiltk; move=>x xf1; rewrite umpreimUn // inE orbC. +have -> : um_preim q f2 x = false=>//. +by move: (dom_inNL v xf1); rewrite /um_preim; case: dom_find=>//->. +Qed. + +Lemma umfiltk_omap V' (U' : union_map_class C V') a p f : + omap a (um_filterk p f) = um_filterk p (omap a f) :> U'. +Proof. +rewrite /um_filter !omap_comp //= /obind/oapp /=. +by apply: eq_in_omap; case=>k v /=; case: ifP; case: a. +Qed. + +Definition omap_umfiltk := umfiltk_omap. + +Lemma umfiltkU p k v f : + um_filterk p (upd k v f) = + if p k then upd k v (um_filterk p f) else + if C k then um_filterk p f else undef. +Proof. +rewrite umfiltU /=; case: ifP; case: ifP=>// Hp Hc. +- by apply: dom_free; rewrite dom_umfiltk inE Hp. +by rewrite upd_condN // Hc. +Qed. + +Lemma umfiltkF p k f : + um_filterk p (free f k) = + if p k then free (um_filterk p f) k else um_filterk p f. +Proof. +rewrite umfiltF; case: ifP=>// N. +by rewrite dom_free // dom_umfiltk inE N. +Qed. + +End FilterKLemmas. + + +Section FilterVLemmas. +Variables (K : ordType) (C : pred K) (V : Type) (U : union_map_class C V). +Implicit Type f : U. +Implicit Type p q : pred V. + +Lemma eq_in_umfiltv (q1 q2 : pred V) f : + (forall v, v \In range f -> q1 v = q2 v) -> + um_filterv q1 f = um_filterv q2 f. +Proof. by move=>H; apply/eq_in_umfilt; case=>k v /In_range; apply: H. Qed. + +Lemma umfiltv_predD (q1 q2 : pred V) f : + subpred q1 q2 -> + um_filterv q2 f = um_filterv q1 f \+ um_filterv (predD q2 q1) f. +Proof. by move=>H; apply: umfilt_predD; case. Qed. + +End FilterVLemmas. + + +Section OmapMembershipLemmas. +Variables aT rT : Type. +Variables (K : ordType) (C : pred K) (Ua : union_map_class C aT) (Ur : union_map_class C rT). +Variables (f : rT -> option aT) (g : aT -> rT). + +Lemma omapv_mapv (h : Ur) : + ocancel f g -> + mapv g (omapv f h : Ua) = um_filterv (isSome \o f) h. +Proof. +move=>O; rewrite /um_filter omap_comp //=. +apply: eq_in_omap; case=>k v H /=. +by case X : (f v)=>[a|] //=; rewrite -(O v) X. +Qed. + +Lemma In_omapv (h : Ur) k x : + ocancel f g -> pcancel g f -> + (k, x) \In (omapv f h : Ua) <-> (k, g x) \In h. +Proof. +move=>O P; rewrite -(In_mapv Ur _ _ _ (pcan_inj P)). +by rewrite omapv_mapv // In_umfilt /= P; split=>//; case. +Qed. + +Lemma In_rangev (h : Ur) (x : aT) : + ocancel f g -> pcancel g f -> + x \In range (omapv f h : Ua) <-> g x \In range h. +Proof. +move=>O P; rewrite !In_rangeX; split; case=>k H; exists k; +by [rewrite -In_omapv | rewrite In_omapv]. +Qed. + +End OmapMembershipLemmas. + + +Section OmapMembershipBool. +Variables aT rT : eqType. +Variables (K : ordType) (C : pred K) (Ua : union_map_class C aT) (Ur : union_map_class C rT). +Variables (f : rT -> option aT) (g : aT -> rT). + +(* decidable version of In_rangev *) + +Lemma mem_rangev (h : Ur) (x : aT) : + ocancel f g -> pcancel g f -> + x \in range (omapv f h : Ua) = (g x \in range h). +Proof. +by move=>O P; apply/idP/idP; move/mem_seqP/(In_rangev _ _ _ O P)/mem_seqP. +Qed. + +End OmapMembershipBool. + + +(************************) +(* PCM-induced ordering *) +(************************) + +(* Union maps and PCM ordering. *) + +Section UmpleqLemmas. +Variables (K : ordType) (C : pred K) (V : Type) (U : union_map_class C V). +Implicit Types (x y a b : U) (p : pred K). + +Lemma umpleq_undef x : [pcm x <= undef]. +Proof. by exists undef; rewrite join_undef. Qed. + +Hint Resolve umpleq_undef : core. + +(* PCM-induced preorder, is an order in the case of union maps *) + +Lemma umpleq_asym x y : [pcm x <= y] -> [pcm y <= x] -> x = y. +Proof. +case=>a -> [b]; case W : (valid x); last first. +- by move/invalidE: (negbT W)=>->; rewrite undef_join. +rewrite -{1 2}[x]unitR -joinA in W *. +by case/(joinxK W)/esym/join0E=>->; rewrite unitR. +Qed. + +(* lemmas on the PCM ordering and um_filter(k) *) + +Lemma umfilt_pleq_mono x y (p : pred (K * V)) : + [pcm x <= y] -> [pcm um_filter p x <= um_filter p y]. +Proof. +move=>H; case W : (valid y). +- by case: H W=>a -> W; rewrite umfiltUn //; eexists _. +by move/invalidE: (negbT W)=>->; rewrite umfilt_undef; apply: umpleq_undef. +Qed. + +(* filter p f is largest lower bound for f and p *) +Lemma umpleq_filtk_meet a p f : + {subset dom a <= p} -> [pcm a <= f] -> [pcm a <= um_filterk p f]. +Proof. by move=>D /(umfilt_pleq_mono (p \o fst)); rewrite (eq_in_umfiltk D) umfilt_predT. Qed. + +(* in valid case, we can define the order by means of filter *) +Lemma umpleqk_pleqE a x : + valid x -> [pcm a <= x] <-> + um_filterk (mem (dom a)) x = a. +Proof. +move=>W; split=>[|<-] // H; case: H W=>b ->. +by apply: umfiltk_dom. +Qed. + +(* join is the least upper bound *) +Lemma umpleq_join a b f : + valid (a \+ b) -> [pcm a <= f] -> [pcm b <= f] -> [pcm a \+ b <= f]. +Proof. +case Vf: (valid f); last by move/invalidE: (negbT Vf)=>->; move: umpleq_undef. +move=>W /(umpleqk_pleqE _ Vf) <- /(umpleqk_pleqE _ Vf) <-. +rewrite -umfilt_dpredU //. +by case=>/= a0 _; move: a0; apply: valid_subdom W. +Qed. + +(* x <= y and subdom *) +Lemma umpleq_subdom x y : valid y -> [pcm x <= y] -> {subset dom x <= dom y}. +Proof. by move=>W H; case: H W=>a -> W b D; rewrite domUn inE W D. Qed. + +Lemma subdom_umpleq a (x y : U) : + valid (x \+ y) -> [pcm a <= x \+ y] -> + {subset dom a <= dom x} -> [pcm a <= x]. +Proof. +move=>W H Sx; move: (umpleq_filtk_meet Sx H); rewrite umfiltUn //. +rewrite umfiltk_memdomE; move/subset_disjC: (valid_subdom W). +by move/(subdom_umfiltk0 _ (validR W))=>->; rewrite unitR. +Qed. + +(* meet is the greatest lower bound *) +Lemma umpleq_meet a (x y1 y2 : U) : + valid (x \+ y1 \+ y2) -> + [pcm a <= x \+ y1] -> [pcm a <= x \+ y2] -> [pcm a <= x]. +Proof. +rewrite um_valid3=> /and3P[V1 V12 V2] H1 H2. +apply: subdom_umpleq (V1) (H1) _ => k D. +move: {D} (umpleq_subdom V1 H1 D) (umpleq_subdom V2 H2 D). +rewrite !domUn !inE V1 V2 /=; case : (k \in dom x)=>//=. +by case: validUn V12=>// _ _ L _ /L /negbTE ->. +Qed. + +(* some/none lemmas *) + +Lemma umpleq_some x1 x2 t s : + valid x2 -> [pcm x1 <= x2] -> find t x1 = Some s -> find t x2 = Some s. +Proof. by move=>W H; case: H W=>a -> W H; rewrite findUnL // (find_some H). Qed. + +Lemma umpleq_none x1 x2 t : + valid x2 -> [pcm x1 <= x2] -> find t x2 = None -> find t x1 = None. +Proof. by case E: (find t x1)=>[a|] // W H <-; rewrite (umpleq_some W H E). Qed. + +End UmpleqLemmas. + + +(********************) +(* Precision lemmas *) +(********************) + +Section Precision. +Variables (K : ordType) (C : pred K) (V : Type) (U : union_map_class C V). +Implicit Types (x y : U). + +Lemma prec_flip x1 x2 y1 y2 : + valid (x1 \+ y1) -> x1 \+ y1 = x2 \+ y2 -> + valid (y2 \+ x2) /\ y2 \+ x2 = y1 \+ x1. +Proof. by move=>X /esym E; rewrite joinC E X joinC. Qed. + +Lemma prec_domV x1 x2 y1 y2 : + valid (x1 \+ y1) -> x1 \+ y1 = x2 \+ y2 -> + reflect {subset dom x1 <= dom x2} (valid (x1 \+ y2)). +Proof. +move=>V1 E; case V12 : (valid (x1 \+ y2)); constructor. +- move=>n Hs; have : n \in dom (x1 \+ y1) by rewrite domUn inE V1 Hs. + rewrite E domUn inE -E V1 /= orbC. + by case: validUn V12 Hs=>// _ _ L _ /L /negbTE ->. +move=>Hs; case: validUn V12=>//. +- by rewrite (validE2 V1). +- by rewrite E in V1; rewrite (validE2 V1). +by move=>k /Hs; rewrite E in V1; case: validUn V1=>// _ _ L _ /L /negbTE ->. +Qed. + +Lemma prec_filtV x1 x2 y1 y2 : + valid (x1 \+ y1) -> x1 \+ y1 = x2 \+ y2 -> + reflect (x1 = um_filterk (mem (dom x1)) x2) (valid (x1 \+ y2)). +Proof. +move=>V1 E; case X : (valid (x1 \+ y2)); constructor; last first. +- case: (prec_domV V1 E) X=>// St _ H; apply: St. + by move=>n; rewrite H dom_umfiltk inE; case/andP. +move: (umfiltk_dom V1); rewrite E umfiltUn -?E //. +rewrite (eq_in_umfiltk (f:=y2) (p2:=pred0)). +- by rewrite umfilt_pred0 ?unitR //; rewrite E in V1; rewrite (validE2 V1). +by move=>n; case: validUn X=>// _ _ L _ /(contraL (L _)) /negbTE. +Qed. + +End Precision. + + +(****************) +(* Ordered eval *) +(****************) + +(* a version of eval, where the user provides a new order of evaluation *) +(* essentially as a list of timestamps, which are then read in-order. *) + +Section OrdEvalDefLemmas. +Variables (K : ordType) (C : pred K) (V R : Type) (U : union_map_class C V). +Implicit Type f : U. +Implicit Type a : R -> K -> V -> R. +Implicit Type p q : pred (K * V). +Implicit Type ks : seq K. + +Fixpoint oeval a ks f z0 := + if ks is k :: ks' then + oeval a ks' f (if find k f is Some v then a z0 k v else z0) + else z0. + +Lemma oev_undef a ks z0 : oeval a ks undef z0 = z0. +Proof. by elim: ks=>[|k ks IH] //=; rewrite find_undef. Qed. + +Lemma oev0 a ks z0 : oeval a ks Unit z0 = z0. +Proof. by elim: ks=>[|k ks IH] //=; rewrite find0E. Qed. + +Lemma oev_dom0 a ks f z0 : dom f =i [::] -> oeval a ks f z0 = z0. +Proof. +case W : (valid f); first by move/(dom0E W)=>->; apply: oev0. +by move/negbT/invalidE: W=>-> _; apply: oev_undef. +Qed. + +(* interaction with constructors that build the ordering mask *) + +Lemma oev_nil a f z0 : oeval a [::] f z0 = z0. +Proof. by []. Qed. + +Lemma oev_consP a k v ks f z0 : + (k, v) \In f -> oeval a (k :: ks) f z0 = oeval a ks f (a z0 k v). +Proof. by move/In_find=>/= ->. Qed. + +Lemma oev_consN a k ks f z0 : + k \notin dom f -> oeval a (k :: ks) f z0 = oeval a ks f z0. +Proof. by case: dom_find=>//= ->. Qed. + +Lemma oev_rconsE a ks k f z0 : + oeval a (rcons ks k) f z0 = + if find k f is Some v then a (oeval a ks f z0) k v + else oeval a ks f z0. +Proof. by elim: ks z0=>[|k' ks IH] z0 /=. Qed. + +Lemma oev_rconsP a k v ks f z0 : + (k, v) \In f -> + oeval a (rcons ks k) f z0 = a (oeval a ks f z0) k v. +Proof. by rewrite oev_rconsE=>/In_find ->. Qed. + +Lemma oev_rconsN a k ks f z0 : + k \notin dom f -> oeval a (rcons ks k) f z0 = oeval a ks f z0. +Proof. by rewrite oev_rconsE=>/find_none ->. Qed. + +Lemma oev_cat a ks1 ks2 f z0 : + oeval a (ks1 ++ ks2) f z0 = oeval a ks2 f (oeval a ks1 f z0). +Proof. by elim: ks1 z0=>/=. Qed. + +(* equalities of oeval wrt. each component *) + +Lemma eq_in_oevA a1 a2 ks f z0 : + (forall r k v, a1 r k v = a2 r k v) -> + oeval a1 ks f z0 = oeval a2 ks f z0. +Proof. +move=>H; elim: ks z0=>[|k ks IH] z0 //=. +by case E : (find k f)=>[b|] //; rewrite IH H. +Qed. + +Lemma eq_in_oevK a ks1 ks2 f z0 : + [seq k <- ks1 | k \in dom f] = [seq k <- ks2 | k \in dom f] -> + oeval a ks1 f z0 = oeval a ks2 f z0. +Proof. +suff oevFK ks : oeval a ks f z0 = + oeval a [seq k <- ks | k \in dom f] f z0. +- by move=>E; rewrite oevFK E -oevFK. +by elim: ks z0=>[|k' ks IH] z0 //=; case: dom_find=>[->|v /= -> _]; rewrite IH. +Qed. + +Lemma eq_in_oevF a ks f1 f2 z0 : + (forall k v, k \in ks -> (k, v) \In f1 <-> (k, v) \In f2) -> + oeval a ks f1 z0 = oeval a ks f2 z0. +Proof. +elim: ks z0=>[|k ks IH] z0 //= H. +case E1: (find k f1)=>[v1|]. +- move/In_find: E1; rewrite H ?(inE,eq_refl) // => /In_find ->. + by apply: IH=>k' v' D; apply: H; rewrite inE D orbT. +case E2: (find k f2)=>[v2|]. +- by move/In_find: E2; rewrite -H ?(inE,eq_refl) // => /In_find; rewrite E1. +by apply: IH=>k' v' D; apply: H; rewrite inE D orbT. +Qed. + +(* restrictions that a, ks, f impose on each other *) + +Lemma oevFK a ks f z0 : + oeval a ks f z0 = oeval a [seq k <- ks | k \in dom f] f z0. +Proof. +by elim: ks z0=>[|k' ks IH] z0 //=; case: dom_find=>[->|v /= -> _]; rewrite IH. +Qed. + +Lemma oevFA a ks f z0 : + oeval a ks f z0 = + oeval (fun r k v => if k \in dom f then a r k v else r) ks f z0. +Proof. +elim: ks a z0=>[|k ks IH] a z0 //=. +by case E: (find k f)=>[b|] //; rewrite (find_some E). +Qed. + +Lemma oevKF a ks f z0 : + oeval a ks f z0 = + oeval a ks (um_filter (fun x => x.1 \in ks) f) z0. +Proof. +elim: ks f z0=>[|k ks IH] f z0 //=. +rewrite find_umfilt /= inE eq_refl [in LHS]IH [in RHS]IH /=. +congr oeval; last by case: (find k f). +by rewrite -umfilt_predI; apply/eq_in_umfilt=>kv /= D; rewrite orKb. +Qed. + +Lemma oevKA a ks f z0 : + oeval a ks f z0 = + oeval (fun r k v => if k \in ks then a r k v else r) ks f z0. +Proof. +elim: ks a z0=>[|k ks IH] a z0 //=; rewrite inE eq_refl [in LHS]IH [in RHS]IH. +by apply: eq_in_oevA=>r k' v; case: ifP=>// D; rewrite inE D orbT. +Qed. + +(* interaction with map constructions *) + +Lemma oev_umfilt a ks p f z0 : + oeval a ks (um_filter p f) z0 = + oeval a [seq k <- ks | if find k f is Some v + then p (k, v) else false] f z0. +Proof. +elim: ks z0=>[|k ks IH] z0 //=; rewrite IH find_umfilt. +by case E: (find k f)=>[b|] //; case: ifP=>//= P; rewrite E. +Qed. + +Lemma oev_filter a ks (p : pred K) f z0 : + oeval a (filter p ks) f z0 = + oeval a ks (um_filterk p f) z0. +Proof. +rewrite oev_umfilt oevFK -filter_predI; congr oeval. +by apply: eq_in_filter=>k D /=; case: dom_find=>[->|v ->]. +Qed. + +Lemma oev_umfiltA a ks p f z0 : + oeval a ks (um_filter p f) z0 = + oeval (fun r k v => if p (k, v) then a r k v else r) ks f z0. +Proof. +elim: ks z0=>[|k ks IH] z0 //=; rewrite IH find_umfilt. +by case E : (find k f)=>[b|] //; case: ifP=>//. +Qed. + +(* convenient derived lemma *) +Lemma oev_dom_umfilt a p f z0 : + oeval a (dom (um_filter p f)) f z0 = + oeval a (dom f) (um_filter p f) z0. +Proof. +rewrite dom_umfiltE oev_filter; apply: eq_in_oevF=>k v _. +by rewrite !In_umfilt /=; split; case=>H Y; move/In_find: Y (Y) H=>->. +Qed. + +Lemma oevF a ks f z0 k : + k \notin ks -> oeval a ks f z0 = oeval a ks (free f k) z0. +Proof. +move=>H; apply: eq_in_oevF=>k' v' D; rewrite InF /= validF. +by case: eqP H D=>[-> /negbTE -> //|???]; split; [move=>H; case: (H) | case]. +Qed. + +Lemma oevUE a k ks v1 v2 f (z0 : R) : + (forall r, a r k v1 = a r k v2) -> + oeval a ks (upd k v1 f) z0 = oeval a ks (upd k v2 f) z0. +Proof. +move=>H; elim: ks z0=>[|k' ks IH] z0 //=. +rewrite !findU; case: eqP=>// ->; rewrite IH. +by congr oeval; case: ifP. +Qed. + +Lemma oevU a k ks v1 v2 f z0 : + (k, v2) \In f -> + (forall r, a r k v1 = a r k v2) -> + oeval a ks (upd k v1 f) z0 = oeval a ks f z0. +Proof. +move=>X H. +have [C' W] : C k /\ valid f by move/In_dom/dom_cond: (X); case: (X). +rewrite [in RHS](_ : f = upd k v2 f); first by apply: oevUE. +apply: umem_eq=>//; first by rewrite validU C' W. +case=>k' v'; rewrite InU validU C' W /=. +case: ifP=>[/eqP ->|_]; last by split=>//; case. +by split=>[/(In_fun X)|[_] ->]. +Qed. + +Lemma oevPtUn a k ks v z0 f : + valid (pts k v \+ f) -> {subset ks <= dom f} -> + oeval a ks (pts k v \+ f) z0 = + oeval a ks f z0. +Proof. +move=>W S; apply: eq_in_oevF=> k0 v0 H. +rewrite InPtUnE //; split; last by move=> H2; right. +by case=>//; case=>E _; move: (S _ H); rewrite E (negbTE (validPtUnD W)). +Qed. + +Lemma oevPtUnE a k ks v1 v2 f z0 : + (forall r, a r k v1 = a r k v2) -> + oeval a ks (pts k v1 \+ f) z0 = oeval a ks (pts k v2 \+ f) z0. +Proof. +move=>H; case W1 : (valid (pts k v1 \+ f)); last first. +- have W2 : valid (pts k v2 \+ f) = false by rewrite !validPtUn in W1 *. + move/invalidE: (negbT W1)=>->; move/invalidE: (negbT W2)=>->. + by rewrite oev_undef. +elim: ks z0=>[|k' ks IH] z0 //=. +have W2 : valid (pts k v2 \+ f) by rewrite !validPtUn in W1 *. +by rewrite !findPtUn2 //; case: eqP=>// ->; rewrite H; apply: IH. +Qed. + +Lemma oev_sub_filter a ks (p : pred K) (h : U) z0 : + {subset dom h <= p} -> + oeval a (filter p ks) h z0 = oeval a ks h z0. +Proof. +move=>S; rewrite oev_filter (eq_in_umfiltI (p2:=predT)) ?umfilt_predT //=. +by case=>k v /In_dom /S. +Qed. + +(* Ideally (- \In f) part would have been folded into p *) +(* but can't because p is decidable and (- \In f) isn't. *) +(* So (- \In f) must be separate. *) +Lemma oev_ind {P : R -> Prop} f ks a z0 : + P z0 -> + (forall k v z0, (k, v) \In f -> k \in ks -> P z0 -> P (a z0 k v)) -> + P (oeval a ks f z0). +Proof. +elim: ks z0=>[|k ks IH] z0 //= Z H; apply: IH; last first. +- by move=>k' v' z' X D; apply: H=>//; rewrite inE D orbT. +case E: (find k f)=>[b|] //; move/In_find: E=>E. +by apply: H=>//; rewrite inE eq_refl. +Qed. + +End OrdEvalDefLemmas. + +Arguments oev_sub_filter [K C V R U a ks p] _. + +Notation oevalv a ks f z0 := (oeval (fun r _ => a r) ks f z0). + +Section OrdEvalRelationalInduction1. +Variables (K : ordType) (C : pred K) (V R1 R2 : Type) (U : union_map_class C V). + +Lemma oev_ind2 {P : R1 -> R2 -> Prop} (f : U) ks a1 a2 z1 z2 : + P z1 z2 -> + (forall k v z1 z2, (k, v) \In f -> k \in ks -> + P z1 z2 -> P (a1 z1 k v) (a2 z2 k v)) -> + P (oeval a1 ks f z1) (oeval a2 ks f z2). +Proof. +elim: ks a1 a2 z1 z2=>[|k ks IH] a1 a2 z1 z2 Z H //=. +apply: IH; last first. +- by move=>k' v' z1' z2' H' K'; apply: H=>//; rewrite inE K' orbT. +case H' : (find k f)=>[b|] //; move/In_find: H'=>H'. +by apply: H=>//; rewrite inE eq_refl. +Qed. + +End OrdEvalRelationalInduction1. + +Section OrdEvalPCM. + +Variables (K : ordType) (C : pred K) (V : Type) (R : pcm) (U : union_map_class C V). +Implicit Type f : U. +Variable (a : R -> K -> V -> R). +Implicit Type p q : pred (K * V). +Implicit Type ks : seq K. +Hypothesis frame : forall x y k v, a (x \+ y) k v = a x k v \+ y. + +Lemma oev1 ks f z0 : oeval a ks f z0 = oeval a ks f Unit \+ z0. +Proof. +apply: (oev_ind2 (P:=fun r1 r2 => r1 = r2 \+ z0)); first by rewrite unitL. +by move=>k v z1 z2 H1 H2 ->; apply: frame. +Qed. + +Lemma oevUn ks (f1 f2 : U) (z0 : R) : + valid (f1 \+ f2) -> + oeval a ks (f1 \+ f2) z0 = oeval a ks f1 (oeval a ks f2 z0). +Proof. +move=>W; elim: ks z0=>[|k ks IH] z0 //=; rewrite findUnL //. +case: dom_find=>[E|v E _]; rewrite E IH //. +move/In_find/In_dom/(dom_inNL W)/find_none: E=>->; congr oeval; apply/esym. +by rewrite oev1 joinC frame joinC -oev1. +Qed. + +End OrdEvalPCM. + + +(********) +(* eval *) +(********) + +(* A special case of oeval with the initial value used *) +(* as default for undefined maps. *) + +Section EvalDefLemmas. +Variables (K : ordType) (C : pred K) (V R : Type) (U : union_map_class C V). +Implicit Type f : U. +Implicit Type a : R -> K -> V -> R. +Implicit Type p q : pred (K * V). + +Definition eval a p f z0 := + oeval a (dom (um_filter p f)) f z0. + +Lemma eval_oevUmfilt a p f z0 : + eval a p f z0 = + oeval a (dom (um_filter p f)) (um_filter p f) z0. +Proof. +apply: eq_in_oevF =>k v H; rewrite In_umfilt. +split; last by case. +split=>//; move: H; move/dom_umfilt => [v' [H H1]]. +by rewrite (In_fun H1 H0) in H. +Qed. + +Lemma eval_undef a p z0 : eval a p undef z0 = z0. +Proof. +by rewrite /eval oev_undef. +Qed. + +Lemma eval0 a p z0 : eval a p Unit z0 = z0. +Proof. +by rewrite /eval oev0. +Qed. + +Lemma eval_dom0 a p f z0 : dom f =i [::] -> eval a p f z0 = z0. +Proof. +by move=> H; rewrite /eval oev_dom0. +Qed. + +Lemma evalPt a p (z0 : R) k v : + eval a p (pts k v) z0 = if C k && p (k, v) then a z0 k v else z0. +Proof. +rewrite /eval umfiltPt. +case E1: (C k); last by rewrite dom_undef oev_nil. +case: (p (k, v)); last by rewrite dom0 oev_nil. +by rewrite domPtK E1 /= findPt E1. +Qed. + +Lemma evalPtUn a p k v z0 f : + valid (pts k v \+ f) -> all (ord k) (dom f) -> + eval a p (pts k v \+ f) z0 = + eval a p f (if p (k, v) then a z0 k v else z0). +Proof. +move=>W /allP H; have: valid (um_filter p (pts k v \+ f)) by rewrite valid_umfilt. +rewrite /eval umfiltPtUn W. +case: (p (k, v))=>W'; last by apply/oevPtUn/dom_umfilt_subset. +rewrite domPtUnK //=; last by apply/allP=>x /dom_umfilt_subset /H. +by rewrite findPtUn // oevPtUn //; apply: dom_umfilt_subset. +Qed. + +Lemma evalUnPt a p k v z0 f : + valid (f \+ pts k v) -> all (ord^~ k) (dom f) -> + eval a p (f \+ pts k v) z0 = + if p (k, v) then a (eval a p f z0) k v else eval a p f z0. +Proof. +move=>W /allP H; have: valid (um_filter p (f \+ pts k v)) by rewrite valid_umfilt. +rewrite /eval umfiltUnPt W. +case: (p (k, v))=>W'; last first. +- by rewrite joinC; apply/oevPtUn/dom_umfilt_subset; rewrite joinC. +rewrite domUnPtK //=; last by apply/allP=>x /dom_umfilt_subset /H. +by rewrite (oev_rconsP _ (v:=v)) // joinC oevPtUn //; + [rewrite joinC | apply: dom_umfilt_subset]. +Qed. + +Lemma evalUn a p z0 f1 f2 : + valid (f1 \+ f2) -> {in dom f1 & dom f2, forall x y, ord x y} -> + eval a p (f1 \+ f2) z0 = eval a p f2 (eval a p f1 z0). +Proof. +elim/um_indb: f2=>[||k v f2 IH W' P W H]. +- by rewrite join_undef valid_undef. +- by rewrite dom0 !unitR eval0. +rewrite -(joinC f2) joinA in W *; rewrite evalUnPt //; last first. +- apply/allP=>x; rewrite domUn inE (validL W). + case/orP=>[/H|]; last by apply: P. + by apply; rewrite domPtUn inE joinC W' eq_refl. +rewrite evalUnPt //; last by apply/allP. +rewrite (IH (validL W)) // => k1 k2 D1 D2; apply: H D1 _. +by rewrite domPtUn inE joinC W' D2 orbT. +Qed. + +Lemma evalPtUnE v2 v1 a p k (z0 : R) f : + (forall r, (if p (k, v1) then a r k v1 else r) = + (if p (k, v2) then a r k v2 else r)) -> + eval a p (pts k v1 \+ f) z0 = eval a p (pts k v2 \+ f) z0. +Proof. +move=>H; rewrite /eval !oev_dom_umfilt !oev_umfiltA. +by rewrite (domPtUnE2 k v1 v2); apply: oevPtUnE. +Qed. + +Lemma evalUEU v2 v1 a p k (z0 : R) f : + (forall r, (if p (k, v1) then a r k v1 else r) = + (if p (k, v2) then a r k v2 else r)) -> + eval a p (upd k v1 f) z0 = eval a p (upd k v2 f) z0. +Proof. +move=>H; rewrite /eval !oev_dom_umfilt !oev_umfiltA. +rewrite (oevUE _ _ _ H); apply: eq_in_oevK; congr filter. +by apply/domE=>x; rewrite !domU. +Qed. + +Lemma evalUE v2 v1 a p k (z0 : R) f : + (k, v2) \In f -> + (forall r, (if p (k, v1) then a r k v1 else r) = + (if p (k, v2) then a r k v2 else r)) -> + eval a p (upd k v1 f) z0 = eval a p f z0. +Proof. +move=>X H; rewrite (evalUEU _ _ H) (_ : upd k v2 f = f) //. +have [C' W] : C k /\ valid f by move/In_dom/dom_cond: (X); case: (X). +apply: umem_eq=>//; first by rewrite validU C' W. +case=>k' v'; rewrite InU validU C' W /=. +by case: ifP=>[/eqP ->|_]; [split=>[[_] ->|/(In_fun X)] | split=>[[]|]]. +Qed. + +Lemma eval_ifP a p z0 f : + eval a p f z0 = + eval (fun r k v => if p (k, v) then a r k v else r) predT f z0. +Proof. by rewrite /eval umfilt_predT oev_dom_umfilt oev_umfiltA. Qed. + +Lemma eq_filt_eval a p1 p2 z0 f1 f2 : + um_filter p1 f1 = um_filter p2 f2 -> + eval a p1 f1 z0 = eval a p2 f2 z0. +Proof. by rewrite !eval_oevUmfilt=>->. Qed. + +Lemma eval_pred0 a z0 f : eval a xpred0 f z0 = z0. +Proof. +case W : (valid f); last first. +- by move/negbT/invalidE: W=>->; rewrite eval_undef. +by rewrite /eval umfilt_pred0 // dom0 /=. +Qed. + +Lemma eval_predI a p q z0 f : + eval a p (um_filter q f) z0 = eval a (predI p q) f z0. +Proof. by rewrite !eval_oevUmfilt -!umfilt_predI. Qed. + +Lemma eval_umfilt p z0 f a : + eval a p f z0 = eval a xpredT (um_filter p f) z0. +Proof. by rewrite eval_predI; apply: eq_filt_eval; apply/eq_in_umfilt. Qed. + +Lemma eq_in_eval p q z0 f a : + (forall kv, kv \In f -> p kv = q kv) -> + eval a p f z0 = eval a q f z0. +Proof. +by rewrite (eval_umfilt p) (eval_umfilt q); move/eq_in_umfilt=>->. +Qed. + +Lemma eval_ind {P : R -> Prop} f p a z0 : + P z0 -> + (forall k v z0, (k, v) \In f -> p (k, v) -> P z0 -> P (a z0 k v)) -> + P (eval a p f z0). +Proof. +move=>Z H; elim/um_indf: f z0 Z H=>[||k v h IH W T] z0 Z H; +rewrite ?eval_undef ?eval0 // evalPtUn // ?(order_path_min (@trans K) T) //. +by apply: IH=>[|k0 v0 z1 /(InR W)]; [case: ifP=>Pk //=; apply: H | apply: H]. +Qed. + +End EvalDefLemmas. + + +Section EvalOmapLemmas. + +Variables (K : ordType) (C : pred K) (V V' R : Type). +Variables (U : union_map_class C V) (U' : union_map_class C V'). + +Lemma eval_omap (b : R -> K -> V' -> R) p a (f : U) z0 : + eval b p (omap a f : U') z0 = + eval (fun r k v => + if a (k, v) is Some v' then b r k v' else r) + (fun kv => + if a kv is Some v' then p (kv.1, v') else false) + f z0. +Proof. +elim/um_indf: f z0=>[||k v f IH W /(order_path_min (@trans K)) P] z0. +- by rewrite omap_undef !eval_undef. +- by rewrite omap0 !eval0. +rewrite omapPtUn W evalPtUn //=; case D : (a (k, v))=>[w|]; last by apply: IH. +rewrite evalPtUn //; last by move/allP: P=>H; apply/allP=>x /dom_omap_sub /H. +rewrite (_ : pts k w \+ omap a f = omap a (pts k v \+ f)) ?valid_omap //. +by rewrite omapPtUn W D. +Qed. + +End EvalOmapLemmas. + + +Section EvalRelationalInduction1. +Variables (K : ordType) (C : pred K) (V R1 R2 : Type) (U : union_map_class C V). + +Lemma eval_ind2 {P : R1 -> R2 -> Prop} (f : U) p0 p1 a0 a1 z0 z1 : + P z0 z1 -> + (forall k v z0 z1, (k, v) \In f -> P z0 z1 -> + P (if p0 (k, v) then a0 z0 k v else z0) + (if p1 (k, v) then a1 z1 k v else z1)) -> + P (eval a0 p0 f z0) (eval a1 p1 f z1). +Proof. +move=>Z H; elim/um_indf: f z0 z1 Z H=>[||k v h IH W T] z0 z1 Z H; +rewrite ?eval_undef ?eval0 // !evalPtUn // ?(order_path_min (@trans K) T) //=. +apply: IH; first by case: ifPn (H k v z0 z1 (InPtUnL W) Z). +by move=>k0 v0 w1 w2 X; apply: H (InR W X). +Qed. + +End EvalRelationalInduction1. + + +Section EvalRelationalInduction2. +Variables (K1 K2 : ordType) (C1 : pred K1) (C2 : pred K2) (V1 V2 R1 R2 : Type). +Variables (U1 : union_map_class C1 V1) (U2 : union_map_class C2 V2). +Variables (U : union_map_class C1 K2) (P : R1 -> R2 -> Prop). + +(* generalization of eval_ind2 to different maps, but related by a bijection *) + +(* f1 and f2 evaluate the same if there exists a monotone bijection *) +(* phi between their timestamps, so that the filtering and *) +(* stepping functions are invariant under the bijection *) + +Lemma eval_ind3 (phi : U) + (f1 : U1) (p1 : pred (K1*V1)) (a1 : R1 -> K1 -> V1 -> R1) (z1 : R1) + (f2 : U2) (p2 : pred (K2*V2)) (a2 : R2 -> K2 -> V2 -> R2) (z2 : R2) : + um_mono phi -> dom phi = dom f1 -> range phi = dom f2 -> + P z1 z2 -> + (forall k1 v1 k2 v2 z1 z2, + (k1, k2) \In phi -> (k1, v1) \In f1 -> (k2, v2) \In f2 -> + P z1 z2 -> + P (if p1 (k1, v1) then a1 z1 k1 v1 else z1) + (if p2 (k2, v2) then a2 z2 k2 v2 else z2)) -> + P (eval a1 p1 f1 z1) (eval a2 p2 f2 z2). +Proof. +elim/um_indf: f1 f2 phi z1 z2 =>[||k1 v1 f1 IH W Po] + f2 phi1 z1 z2 /ummonoP M1 D1 R Hz H. +- rewrite eval_undef eval_dom0 // -R; move/domE: D1; rewrite dom_undef. + case W1 : (valid phi1); first by move/(dom0E W1)=>->; rewrite range0. + by move/negbT/invalidE: W1=>->; rewrite dom_undef range_undef. +- rewrite eval0 eval_dom0 // -R; move/domE: D1; rewrite dom0. + case W1 : (valid phi1); first by move/(dom0E W1)=>->; rewrite range0. + by move/negbT/invalidE: W1=>->; rewrite dom_undef range_undef. +have A1 : all (ord k1) (dom f1) by apply: order_path_min Po; apply: trans. +case W1 : (valid phi1); last first. +- by move/negbT/invalidE: W1 D1 R=>->; rewrite dom_undef domPtUnK. +rewrite domPtUnK // in D1 *; rewrite evalPtUn //. +have [k2 I1] : exists k2, (k1, k2) \In phi1. +- by apply/In_domX; rewrite D1 inE eq_refl. +have I2 : (k1, v1) \In pts k1 v1 \+ f1 by apply/InPtUnE=>//; left. +have [v2 I3] : exists v2, (k2, v2) \In f2. +- by apply/In_domX; rewrite -R; apply/mem_seqP/(In_range I1). +move: (H _ _ _ _ _ _ I1 I2 I3 Hz)=>Hp. +set phi2 := free phi1 k1. +have W2 : valid f2 by move/In_dom/dom_valid: I3. +have E2 : f2 = pts k2 v2 \+ free f2 k2 by apply/In_eta: I3. +have A2 : all (ord k2) (dom (free f2 k2)). +- apply/allP=>x; rewrite domF inE E2 domPtUn inE -E2 W2 /= domF inE. + case: eqP=>//= N; rewrite -R =>R'; move/mem_seqP/In_rangeX: (R'). + case=>k' H1; apply: M1 (I1) (H1) _; move/allP: A1; apply. + move/In_dom: H1 (H1); rewrite D1 inE; case/orP=>//= /eqP ->. + by move/(In_fun I1)/N. +rewrite E2 evalPtUn -?E2 //. +have M2 : um_mono phi2. +- by apply/ummonoP=>???? /InF [_ _ /M1] X /InF [_ _]; apply: X. +have D2 : dom phi2 = dom f1. +- apply/domE=>x; rewrite domF inE D1 inE eq_sym. + by case: eqP=>// ->{x}; rewrite (negbTE (validPtUnD W)). +have R2' : range phi2 = dom (free f2 k2). + move/In_eta: (I1) (R)=>E1; rewrite E1 rangePtUnK. + - by rewrite {1}E2 domPtUnK //; [case | rewrite -E2]. + - by rewrite -E1. + apply/allP=>x; rewrite domF inE D1 inE eq_sym. + by case: eqP=>//= _; apply/allP/A1. +have {}H x1 w1 x2 w2 t1 t2 : (x1, x2) \In phi2 -> (x1, w1) \In f1 -> + (x2, w2) \In free f2 k2 -> P t1 t2 -> + P (if p1 (x1, w1) then a1 t1 x1 w1 else t1) + (if p2 (x2, w2) then a2 t2 x2 w2 else t2). +- case/InF=>_ _ F1 F2 /InF [_ _ F3]. + by apply: H F1 _ F3; apply/(InPtUnE _ W); right. +by case: ifP Hp=>L1; case: ifP=>L2 Hp; apply: IH M2 D2 R2' Hp H. +Qed. + +End EvalRelationalInduction2. + + +(* Evaluating frameable functions *) +Section EvalFrame. +Variables (K : ordType) (C : pred K) (V : Type) (R : pcm) (U : union_map_class C V). +Variables (a : R -> K -> V -> R) (p : pred (K * V)). +Hypothesis frame : (forall x y k v, a (x \+ y) k v = a x k v \+ y). + +Implicit Type f : U. + +(* a lemma on eval update only makes sense if the function a is frameable, *) +(* so that the result is independent of the order of k *) +Lemma evalFPtUn k v z0 f : + valid (pts k v \+ f) -> + eval a p (pts k v \+ f) z0 = + if p (k, v) then a Unit k v \+ eval a p f z0 else eval a p f z0. +Proof. +move=>W; rewrite /eval umfiltPtUn W. +have D : k \notin dom f by apply: (validPtUnD W). +have Ck : C k by apply: (validPtUn_cond W). +case: ifP=>_; last by apply: oevPtUn=>//; apply: dom_umfilt_subset. +rewrite oevUn // -(oev_sub_filter (p:=mem [:: k])) ?(domPtK,Ck) //. +rewrite -dom_umfiltk_filter umfiltPtUn /= valid_umfiltUnR // inE eq_refl. +rewrite umfilt_mem0L ?(inE,valid_umfilt,validR W) //=; first last. +- by move=>?? /In_umfilt [] _ /In_dom Df; rewrite inE; case: eqP Df D=>// ->->. +rewrite unitR domPtK Ck /= findPt Ck -frame unitL. +rewrite -(oev_sub_filter (p:=mem (dom f))) //. +rewrite -dom_umfiltk_filter umfiltPtUn /= valid_umfiltUnR // (negbTE D). +congr (a (oeval a (dom _) f z0) k v). +by rewrite -subdom_umfiltkE; apply: dom_umfilt_subset. +Qed. + +Lemma evalFU k v z0 f : + valid (upd k v f) -> + eval a p (upd k v f) z0 = + if p (k, v) then a Unit k v \+ eval a p (free f k) z0 + else eval a p (free f k) z0. +Proof. +move=>W; move: (W); rewrite validU =>/andP [C1 _]. +have /um_eta2 E : find k (upd k v f) = Some v. +- by rewrite findU eq_refl -(validU k v) W. +by rewrite E evalFPtUn -?E // freeU eq_refl C1. +Qed. + +End EvalFrame. + +Notation evalv a p f z0 := (eval (fun r _ => a r) p f z0). + + +(************) +(* um_count *) +(************) + +Section CountDefLemmas. +Variables (K : ordType) (C : pred K) (V : Type) (U : union_map_class C V). +Implicit Type f : U. +Implicit Type p : pred (K * V). + +Definition um_count p f := size (dom (um_filter p f)). + +Lemma umcnt0 p : um_count p Unit = 0. +Proof. by rewrite /um_count umfilt0 dom0. Qed. + +Lemma umcnt_undef p : um_count p undef = 0. +Proof. by rewrite /um_count umfilt_undef dom_undef. Qed. + +Lemma umcntPt p k v : + um_count p (pts k v) = if C k && p (k, v) then 1 else 0. +Proof. +rewrite /um_count umfiltPt; case C': (C k); last by rewrite dom_undef. +by case: ifP; [rewrite domPtK C' | rewrite dom0]. +Qed. + +Lemma umcntUn p f1 f2 : + valid (f1 \+ f2) -> + um_count p (f1 \+ f2) = um_count p f1 + um_count p f2. +Proof. +move=>W; rewrite /um_count umfiltUn // size_domUn //. +by rewrite -umfiltUn // valid_umfilt. +Qed. + +Lemma umcntPtUn p k v f : + valid (pts k v \+ f) -> + um_count p (pts k v \+ f) = (if p (k, v) then 1 else 0) + um_count p f. +Proof. by move=>W; rewrite umcntUn // umcntPt (validPtUn_cond W). Qed. + +Lemma umcntUnPt p k v f : + valid (f \+ pts k v) -> + um_count p (f \+ pts k v) = um_count p f + if p (k, v) then 1 else 0. +Proof. by rewrite joinC=>W; rewrite umcntPtUn // addnC. Qed. + +Lemma umcntF p k v f : + (k, v) \In f -> + um_count p (free f k) = + if p (k, v) then (um_count p f).-1 else um_count p f. +Proof. +move=>H; move/In_dom: (H)=>/= D; rewrite /um_count. +case E: (k \in dom (um_filter p f)). +- case/dom_umfilt: E=>w [H1 H2]. + rewrite -{w H2}(In_fun H H2) in H1 *. + rewrite H1 umfiltF size_domF ?subn1 //. + by apply/dom_umfilt; exists v. +rewrite umfiltF; case: ifP=>P; last by rewrite dom_free // E. +by move/negP: E; elim; apply/dom_umfilt; exists v. +Qed. + +Lemma umcntU p k v w f : + (k, w) \In f -> + um_count p (upd k v f) = + if p (k, v) then + if p (k, w) then um_count p f else (um_count p f).+1 + else + if p (k, w) then (um_count p f).-1 else um_count p f. +Proof. +move=>H; have E : f = pts k w \+ free f k. +- by apply: um_eta2; apply/In_find. +have W1 : valid f by move/In_dom/dom_valid: H. +have W2 : valid (pts k v \+ free f k). +- by rewrite {1}E !validPtUn in W1 *. +have W3 : valid (pts k w \+ free f k). +- by rewrite {1}E !validPtUn in W1 *. +rewrite {1}E updPtUn umcntPtUn // (umcntF p H). +case: ifP=>H1; case: ifP=>H2; rewrite ?add0n ?add1n //. +have: um_count p f > 0; first by rewrite E umcntPtUn // H2. +by case X: (um_count p f). +Qed. + +Lemma eq_in_umcnt p1 p2 f : + (forall kv, kv \In f -> p1 kv = p2 kv) -> + um_count p1 f = um_count p2 f. +Proof. by rewrite /um_count=>/eq_in_umfilt ->. Qed. + +(* common case *) +Lemma eq_in_umcnt2 p1 p2 f : + p1 =1 p2 -> um_count p1 f = um_count p2 f. +Proof. by move=>S; apply: eq_in_umcnt=>kv _; apply: S. Qed. + +Lemma umcnt_leq p1 p2 f : + subpred p1 p2 -> um_count p1 f <= um_count p2 f. +Proof. +move=>S; case W: (valid f); last first. +- by move/negbT/invalidE: W=>->; rewrite !umcnt_undef. +rewrite /um_count (umfilt_predD _ S) size_domUn ?leq_addr //. +by rewrite -umfilt_predD // valid_umfilt. +Qed. + +Lemma umcnt_count q f : count q (dom f) = um_count (q \o fst) f. +Proof. +rewrite assocs_dom /um_count -size_assocs. +by rewrite assocs_umfilt /= size_filter count_map. +Qed. + +Lemma umcnt_umfilt p q f : + um_count p (um_filter q f) = um_count (predI p q) f. +Proof. by rewrite /um_count umfilt_predI. Qed. + +Lemma umcnt_umfiltC p q f : + um_count p (um_filter q f) = um_count q (um_filter p f). +Proof. by rewrite !umcnt_umfilt; apply: eq_in_umcnt=>x; rewrite /= andbC. Qed. + +Lemma umcnt_pred0 f : um_count pred0 f = 0. +Proof. +case W : (valid f); last first. +- by move/invalidE: (negbT W)=>->; rewrite umcnt_undef. +by rewrite /um_count umfilt_pred0 // dom0. +Qed. + +Lemma umcnt_predT f : um_count predT f = size (dom f). +Proof. by rewrite /um_count umfilt_predT. Qed. + +Lemma umcnt_predU p1 p2 f : + um_count (predU p1 p2) f = + um_count p1 f + um_count (predD p2 p1) f. +Proof. +case W: (valid f); last first. +- by move/invalidE: (negbT W)=>->; rewrite !umcnt_undef. +rewrite /um_count umfilt_predU size_domUn //. +case: validUn=>//; rewrite ?(valid_umfilt,W) //. +move=>k /dom_umfilt [v1 [P1 H1]] /dom_umfilt [v2 [/= P2 H2]]. +by rewrite -(In_fun H1 H2) P1 in P2. +Qed. + +Lemma umcnt_predD p1 p2 f : + subpred p1 p2 -> + um_count p2 f = um_count p1 f + um_count (predD p2 p1) f. +Proof. +move=>S; rewrite -umcnt_predU //; apply: eq_in_umcnt=>kv /= _. +by case E: (p1 kv)=>//; apply: S. +Qed. + +Lemma umcnt_predDE p1 p2 f : + um_count (predD p2 p1) f = + um_count p2 f - um_count (predI p1 p2) f. +Proof. +have S1 : p2 =1 predU (predD p2 p1) (predI p1 p2). +- by move=>x /=; case: (p1 x)=>//; rewrite orbF. +have S2: predD (predI p1 p2) (predD p2 p1) =1 predI p1 p2. +- by move=>x /=; case: (p1 x)=>//; rewrite andbF. +rewrite {1}(eq_in_umcnt2 _ S1) umcnt_predU // (eq_in_umcnt2 _ S2). +by rewrite -addnBA // subnn addn0. +Qed. + +Lemma umcnt_umfiltU p f q1 q2 : + um_count p (um_filter (predU q1 q2) f) = + um_count p (um_filter q1 f) + um_count p (um_filter (predD q2 q1) f). +Proof. +rewrite umcnt_umfiltC umcnt_predU ?valid_umfilt //. +by rewrite !(umcnt_umfiltC p). +Qed. + +Lemma umcnt_umfilt0 f : + valid f -> forall p, um_count p f = 0 <-> um_filter p f = Unit. +Proof. +move=>W; split; last by rewrite /um_count=>->; rewrite dom0. +by move/size0nil=>D; apply: dom0E; rewrite ?valid_umfilt ?D. +Qed. + +Lemma umcnt_eval0 R a p f (z0 : R) : um_count p f = 0 -> eval a p f z0 = z0. +Proof. +case W : (valid f); last first. +- by move/invalidE: (negbT W)=>->; rewrite eval_undef. +by move/(umcnt_umfilt0 W)=>E; rewrite eval_umfilt E eval0. +Qed. + +Lemma umcnt_mem0 p f : + um_count p f = 0 <-> (forall k v, (k, v) \In f -> ~~ p (k, v)). +Proof. +case W : (valid f); last first. +- by move/invalidE: (negbT W)=>->; rewrite umcnt_undef; split=>// _ k v /In_undef. +split; first by move/(umcnt_umfilt0 W)/umfilt_mem0. +by move=>H; apply/(umcnt_umfilt0 W); apply/umfilt_mem0L. +Qed. + +Lemma umcnt_size p f : um_count p f <= size (dom f). +Proof. by rewrite -umcnt_predT; apply: umcnt_leq. Qed. + +Lemma umcnt_memT p f : + um_count p f = size (dom f) <-> + forall k v, (k, v) \In f -> p (k, v). +Proof. +elim/um_indf: f=>[||k v f IH W P]. +- by rewrite umcnt_undef dom_undef; split=>// _ k v /In_undef. +- by rewrite umcnt0 dom0; split=>// _ k v /In0. +rewrite umcntPtUn // size_domPtUn //. +case: ifP=>H; split. +- move/eqP; rewrite !add1n eqSS=>/eqP/IH=>H1 k1 v1. + by case/InPtUnE=>//; [case=>->-> | apply: H1]. +- move=>H1; apply/eqP; rewrite !add1n eqSS; apply/eqP. + by apply/IH=>k1 v1 H2; apply: H1; apply/InR. +- by rewrite add0n=>X; move: (umcnt_size p f); rewrite X add1n ltnn. +by move/(_ k v)/(_ (InPtUnL W)); rewrite H. +Qed. + +Lemma umcnt_filt_eq p f : um_count p f = size (dom f) <-> f = um_filter p f. +Proof. +rewrite umcnt_memT -{2}[f]umfilt_predT -eq_in_umfilt. +by split=>H; [case=>k v /H | move=>k v /H]. +Qed. + +Lemma umcnt_eval p f : um_count p f = eval (fun n _ _ => n.+1) p f 0. +Proof. +elim/um_indf: f=>[||k v f IH W /(order_path_min (@trans K)) P]. +- by rewrite umcnt_undef eval_undef. +- by rewrite umcnt0 eval0. +rewrite umcntPtUn // evalPtUn //=; case: ifP=>// H. +by rewrite /eval oev1 // IH addnC; congr (_ + _). +Qed. + +End CountDefLemmas. + + +(***********) +(* dom_map *) +(***********) + +(* Taking a domain as a finite set, instead of a sequence *) +(* useful when needing to state disjointness of maps that *) +(* have different types *) + +Section DomMap. +Variables (K : ordType) (C : pred K) (V : Type). +Variables (U : union_map_class C V) (U' : union_map_class C unit). + +Definition dom_map (x : U) : U' := omap (fun => Some tt) x. + +Lemma dom_map0 : dom_map (Unit : U) = Unit. +Proof. by rewrite /dom_map omap0. Qed. + +Lemma dom_map_undef : dom_map undef = undef. +Proof. by rewrite /dom_map omap_undef. Qed. + +Lemma dom_mapE x : dom (dom_map x) = dom x. +Proof. by apply: dom_omap_some. Qed. + +Lemma valid_dom_map x : valid (dom_map x) = valid x. +Proof. by rewrite /dom_map valid_omap. Qed. + +Lemma dom_mapUn (x y : U) : + dom_map (x \+ y) = dom_map x \+ dom_map y. +Proof. +case W : (valid (x \+ y)); first by rewrite /dom_map omapUn. +rewrite /dom_map; move/invalidE: (negbT W)=>->; rewrite omap_undef. +case: validUn W=>//. +- by move/invalidE=>->; rewrite omap_undef undef_join. +- by move/invalidE=>->; rewrite omap_undef join_undef. +move=>k D1 D2 _; suff : ~~ valid (dom_map x \+ dom_map y). +- by move/invalidE=>->. +by case: validUn=>// _ _ /(_ k); rewrite !dom_mapE=>/(_ D1); rewrite D2. +Qed. + +Lemma domeq2_dom_mapL x : dom_eq2 (dom_map x) x. +Proof. by rewrite /dom_eq2 valid_dom_map dom_mapE !eq_refl. Qed. + +Lemma domeq2_dom_mapR x : dom_eq2 x (dom_map x). +Proof. by apply: domeq2_sym; apply: domeq2_dom_mapL. Qed. + +End DomMap. + +Arguments dom_map [K C V U U'] _. + +(* composing dom_map *) + +Section DomMapIdempotent. +Variables (K : ordType) (C : pred K) (V : Type) (U : union_map_class C V). +Variables (U' : union_map_class C unit) (U'' : union_map_class C unit). + +Lemma dom_map_idemp (x : U) : dom_map (dom_map x : U') = dom_map x :> U''. +Proof. by rewrite /dom_map omap_comp. Qed. + +End DomMapIdempotent. + + +(*****************) +(* Map inversion *) +(*****************) + +Section MapInverse. +Variables (K V : ordType) (C : pred K) (C' : pred V) (U : union_map_class C V) (U' : union_map_class C' K). + +Definition inverse (f : U) : U' := + um_foldl (fun r k v => r \+ pts v k) Unit undef f. + +Lemma inverse_undef : inverse undef = undef. +Proof. by rewrite /inverse umfoldl_undef. Qed. + +Lemma inverse0 : inverse Unit = Unit. +Proof. by rewrite /inverse umfoldl0. Qed. + +Lemma inverseUn f1 f2 : + valid (f1 \+ f2) -> inverse (f1 \+ f2) = inverse f1 \+ inverse f2. +Proof. +move=>W; rewrite /inverse umfoldlUn_frame ?unitR //. +by move=>*; rewrite joinAC. +Qed. + +Lemma inversePt k v : C k -> inverse (pts k v) = pts v k. +Proof. by move=>C''; rewrite /inverse umfoldlPt unitL C''. Qed. + +Lemma inversePtUn k v f : + valid (pts k v \+ f) -> + inverse (pts k v \+ f) = pts v k \+ inverse f. +Proof. by move=>W; rewrite inverseUn // inversePt // (validPtUn_cond W). Qed. + +Lemma dom_inverse f : valid (inverse f) -> dom (inverse f) =i range f. +Proof. +rewrite /inverse/um_foldl/range; case: ifP=>_; last by rewrite valid_undef. +elim: (assocs f)=>[|x g IH] /= W k; first by rewrite dom0. +rewrite foldl_init in W *; last by move=>*; rewrite joinAC. +by rewrite domUnPt !inE W /= eq_sym IH // (validL W). +Qed. + +Lemma valid_inverse f : + valid (inverse f) = + [&& valid f, uniq (range f) & all C' (range f)]. +Proof. +elim/um_indf: f=>[||k v f IH W /(order_path_min (@trans K)) P]. +- by rewrite inverse_undef !valid_undef. +- by rewrite inverse0 !valid_unit range0. +rewrite inversePtUn W //= rangePtUnK // validPtUn /=. +rewrite IH (validR W) /=; bool_congr; rewrite andbC -!andbA. +rewrite andbC [in X in _ = X]andbC. +case X1 : (uniq (range f)) IH=>//=. +case X2 : (all C' (range f))=>//=. +by rewrite (validR W)=>/dom_inverse ->. +Qed. + +(* The next few lemmas that depend on valid (inverse f) *) +(* can be strenghtened to require just uniq (range f) and *) +(* all C' (range f). *) +(* However, the formulation with the hypothesis valid (inverse f) *) +(* is packaged somewhat better, and is what's encountered in practice. *) + +Lemma range_inverse f : valid (inverse f : U') -> range (inverse f) =i dom f. +Proof. +elim/um_indf: f=>[||k v f IH W P]. +- by rewrite inverse_undef valid_undef. +- by rewrite inverse0 range0 dom0. +rewrite inversePtUn // => W' x. +rewrite rangePtUn inE W' domPtUn inE W /=. +by case: eqP=>//= _; rewrite IH // (validR W'). +Qed. + +Lemma In_inverse k v f : + valid (inverse f) -> (k, v) \In f <-> (v, k) \In inverse f. +Proof. +elim/um_indf: f k v=>[||x w f IH W /(order_path_min (@trans K)) P] k v. +- by rewrite inverse_undef valid_undef. +- by rewrite inverse0; split=>/In0. +move=>W'; rewrite inversePtUn // !InPtUnE //; last by rewrite -inversePtUn. +rewrite IH; first by split; case=>[[->->]|]; auto. +rewrite !valid_inverse rangePtUnK // (validR W) in W' *. +by case/and3P: W'=>_ /= /andP [_ ->] /andP [_ ->]. +Qed. + +Lemma uniq_range_inverse f : uniq (range (inverse f)). +Proof. +case W : (valid (inverse f)); last first. +- by move/invalidE: (negbT W)=>->; rewrite range_undef. +rewrite /range map_inj_in_uniq. +- by apply: (@map_uniq _ _ fst); rewrite -assocs_dom; apply: uniq_dom. +case=>x1 x2 [y1 y] /= H1 H2 E; rewrite {x2}E in H1 *. +move/mem_seqP/umemE/(In_inverse _ _ W): H1=>H1. +move/mem_seqP/umemE/(In_inverse _ _ W): H2=>H2. +by rewrite (In_fun H1 H2). +Qed. + +Lemma all_range_inverse f : all C (range (inverse f)). +Proof. +apply: (@sub_all _ (fun k => k \in dom f)); first by move=>x /dom_cond. +by apply/allP=>x H; rewrite -range_inverse //; apply: range_valid H. +Qed. + +Lemma sorted_range_inverse f : + valid (inverse f) -> + sorted ord (range f) -> sorted ord (range (inverse f)). +Proof. +move=>W /ummonoP X; apply/ummonoP=>v v' k k'. +move/(In_inverse _ _ W)=>H1 /(In_inverse _ _ W) H2. +apply: contraTT; case: ordP=>//= E _; first by case: ordP (X _ _ _ _ H2 H1 E). +by move/eqP: E H2=>-> /(In_fun H1) ->; rewrite irr. +Qed. + +End MapInverse. + +Arguments inverse [K V C C' U U']. + + +Section InverseLaws. +Variables (K V : ordType) (C : pred K) (C' : pred V) (U : union_map_class C V) (U' : union_map_class C' K). + +Lemma valid_inverse_idemp (f : U) : + valid (inverse (inverse f : U') : U) = valid (inverse f : U'). +Proof. +by rewrite valid_inverse uniq_range_inverse all_range_inverse !andbT. +Qed. + +Lemma inverse_idemp (f : U) : + valid (inverse f : U') -> inverse (inverse f : U') = f. +Proof. +move=>W; apply/umem_eq. +- by rewrite valid_inverse_idemp. +- by move: W; rewrite valid_inverse; case/and3P. +move=>x; split=>H. +- by apply/(In_inverse _ _ W); apply/(@In_inverse _ _ _ _ _ U) =>//; case: H. +case: x H=>k v /(In_inverse _ _ W)/In_inverse; apply. +by rewrite valid_inverse_idemp. +Qed. + +End InverseLaws. + + +(***************************) +(* Composition of two maps *) +(***************************) + +Section MapComposition. +Variables (K1 K2 : ordType) (C1 : pred K1) (C2 : pred K2) (V : Type). +Variables (Uf : union_map_class C1 K2). +Variables (Ug : union_map_class C2 V). +Variables (U : union_map_class C1 V). +Implicit Types (f : Uf) (g : Ug). + +Definition um_comp g f : U := + um_foldl (fun r k v => if find v g is Some w + then pts k w \+ r else r) Unit undef f. + +Lemma umcomp_undeff f : valid f -> um_comp undef f = Unit. +Proof. +move=>W; apply: (umfoldl_ind (P:=fun r => r = Unit))=>//. +by move=>*; rewrite find_undef. +Qed. +Lemma umcomp_fundef g : um_comp g undef = undef. +Proof. by rewrite /um_comp umfoldl_undef. Qed. + +Lemma umcomp0f f : valid f -> um_comp Unit f = Unit. +Proof. +move=>W; apply: (umfoldl_ind (P:=fun r => r = Unit))=>//. +by move=>*; rewrite find0E. +Qed. + +Lemma umcompf0 g : um_comp g Unit = Unit. +Proof. by rewrite /um_comp umfoldl0. Qed. + +Lemma dom_umcomp_sub g f : {subset dom (um_comp g f) <= dom f}. +Proof. +rewrite /um_comp; elim/um_indf: f=>[||k v f IH W P] x. +- by rewrite umfoldl_undef dom_undef. +- by rewrite umfoldl0 dom0. +rewrite umfoldlUn_frame //; +last by move=>*; case: (find _ _)=>// a; rewrite joinA. +rewrite unitR umfoldlPt (validPtUn_cond W). +case E : (find v g)=>[b|]; last first. +- by rewrite unitL=>/IH; rewrite domPtUn inE W =>->; rewrite orbT. +rewrite unitR domPtUn inE; case/andP=>_ /orP [/eqP <-|/IH]. +- by rewrite domPtUn inE W eq_refl. +by rewrite domPtUn inE W=>->; rewrite orbT. +Qed. + +Lemma valid_umcomp g f : valid (um_comp g f) = valid f. +Proof. +rewrite /um_comp; elim/um_indf: f=>[||k v f IH W P]. +- by rewrite umfoldl_undef !valid_undef. +- by rewrite umfoldl0 !valid_unit. +rewrite umfoldlUn_frame //; + last by move=>*; case: (find _ _)=>// a; rewrite joinA. +rewrite unitR W umfoldlPt (validPtUn_cond W). +case: (find v g)=>[a|]; last by rewrite unitL IH (validR W). +rewrite unitR validPtUn (validPtUn_cond W) IH (validR W). +by apply: contra (notin_path P); apply: dom_umcomp_sub. +Qed. + +Lemma umcompUnf g1 g2 f : + valid (g1 \+ g2) -> um_comp (g1 \+ g2) f = um_comp g1 f \+ um_comp g2 f. +Proof. +rewrite /um_comp=>Wg; elim/um_indf: f=>[||k v f IH P W]. +- by rewrite !umfoldl_undef undef_join. +- by rewrite !umfoldl0 unitL. +rewrite !umfoldlUn_frame //; +try by move=>*; case: (find _ _)=>// a; rewrite joinA. +rewrite !unitR !umfoldlPt; case: ifP=>C; last by rewrite !undef_join. +rewrite {}IH joinAC [in X in _ = X]joinA [in X in _ = X]joinAC; +rewrite -[in X in _ = X]joinA; congr (_ \+ _). +case: validUn (Wg)=>// W1 W2 X _; rewrite findUnL //. +by case: ifP=>[/X|]; case: dom_find=>// ->; rewrite ?unitL ?unitR. +Qed. + +Lemma umcompfUn g f1 f2 : + valid (f1 \+ f2) -> um_comp g (f1 \+ f2) = um_comp g f1 \+ um_comp g f2. +Proof. +rewrite /um_comp=>W; rewrite umfoldlUn_frame ?unitR //. +by move=>*; case: (find _ _)=>// a; rewrite joinA. +Qed. + +Lemma umcompfPt g k v : + um_comp g (pts k v) = + if C1 k then + if find v g is Some w then pts k w else Unit + else undef. +Proof. +rewrite /um_comp umfoldlPt; case: ifP=>C //. +by case: (find _ _)=>// a; rewrite unitR. +Qed. + +Lemma umcompfPtUn g f k v : + valid (pts k v \+ f) -> + um_comp g (pts k v \+ f) = + if find v g is Some w then pts k w \+ um_comp g f + else um_comp g f. +Proof. +move=>W; rewrite umcompfUn // umcompfPt (validPtUn_cond W). +by case: (find _ _)=>[a|] //; rewrite unitL. +Qed. + +Lemma umcompPtf f k v : + um_comp (pts k v) f = + if C2 k then + omap (fun x => if x.2 == k then Some v else None) f + else if valid f then Unit else undef. +Proof. +elim/um_indf: f=>[||x w f IH W P]. +- rewrite umcomp_fundef omap_undef. + by case: ifP=>//; rewrite valid_undef. +- rewrite umcompf0 omap0. + by case: ifP=>//; rewrite valid_unit. +rewrite umcompfUn // umcompfPt (validPtUn_cond W) omapPtUn. +rewrite W findPt2 andbC. +case C: (C2 k) IH=>/= IH; last by rewrite IH (validR W) unitL. +by case: eqP=>_; rewrite IH // unitL. +Qed. + +Lemma umcompPtUnf g f k v : + valid (pts k v \+ g) -> + um_comp (pts k v \+ g) f = + omap (fun x => if x.2 == k then Some v else None) f \+ um_comp g f. +Proof. +by move=>W; rewrite umcompUnf // umcompPtf (validPtUn_cond W). +Qed. + +Lemma In_umcomp g f k v : + (k, v) \In um_comp g f <-> + valid (um_comp g f) /\ exists k', (k, k') \In f /\ (k', v) \In g. +Proof. +split=>[H|[W][k'][]]. +- split; first by case: H. + elim/um_indf: f H=>[||x w f IH P W]. + - by rewrite umcomp_fundef=>/In_undef. + - by rewrite umcompf0=>/In0. + rewrite /um_comp umfoldlUn_frame //; + last by move=>*; case: (find _ _)=>// a; rewrite joinA. + rewrite unitR !umfoldlPt; case: ifP=>C; last first. + - by rewrite undef_join=>/In_undef. + case/InUn; last by case/IH=>k' [H1 H2]; exists k'; split=>//; apply/InR. + case E: (find _ _)=>[b|]; last by move/In0. + rewrite unitR=>/InPt; case; case=>?? _; subst x b. + by exists w; split=>//; apply/In_find. +elim/um_indf: f W k'=>[||x w f IH P W] W' k'. +- by move/In_undef. +- by move/In0. +rewrite umcompfPtUn // in W'; rewrite InPtUnE //. +case=>[[??] G|H1 H2]. +- subst x w; rewrite umcompfPtUn //. + by move/In_find: (G)=>E; rewrite E in W' *; apply/InPtUnL. +rewrite umcompfPtUn //. +case E: (find _ _) W'=>[b|] W'; last by apply: IH H1 H2. +by apply/InR=>//; apply:IH H1 H2; rewrite (validR W'). +Qed. + +End MapComposition. + + +(**********) +(* um_all *) +(**********) + +Section AllDefLemmas. +Variables (K : ordType) (C : pred K) (V : Type). +Variables (U : union_map_class C V) (P : K -> V -> Prop). +Implicit Types (k : K) (v : V) (f : U). + +Definition um_all f := forall k v, (k, v) \In f -> P k v. + +Lemma umall_undef : um_all undef. +Proof. by move=>k v /In_undef. Qed. + +Lemma umall0 : um_all Unit. +Proof. by move=>k v /In0. Qed. + +Hint Resolve umall_undef umall0 : core. + +Lemma umallUn f1 f2 : um_all f1 -> um_all f2 -> um_all (f1 \+ f2). +Proof. by move=>X Y k v /InUn [/X|/Y]. Qed. + +Lemma umallUnL f1 f2 : valid (f1 \+ f2) -> um_all (f1 \+ f2) -> um_all f1. +Proof. by move=>W H k v F; apply: H; apply/InL. Qed. + +Lemma umallUnR f1 f2 : valid (f1 \+ f2) -> um_all (f1 \+ f2) -> um_all f2. +Proof. by rewrite joinC; apply: umallUnL. Qed. + +Lemma umallPt k v : P k v -> um_all (pts k v : U). +Proof. by move=>X k1 v1 /InPt [[->->]]. Qed. + +Lemma umallPtUn k v f : P k v -> um_all f -> um_all (pts k v \+ f). +Proof. by move/(umallPt (k:=k)); apply: umallUn. Qed. + +Lemma umallUnPt k v f : P k v -> um_all f -> um_all (f \+ pts k v). +Proof. by rewrite joinC; apply: umallPtUn. Qed. + +Lemma umallPtE k v : C k -> um_all (pts k v : U) -> P k v. +Proof. by move/(In_condPt v)=>C'; apply. Qed. + +Lemma umallPtUnE k v f : + valid (pts k v \+ f) -> um_all (pts k v \+ f) -> P k v /\ um_all f. +Proof. +move=>W H; move: (umallUnL W H) (umallUnR W H)=>{H} H1 H2. +by split=>//; apply: umallPtE H1; apply: validPtUn_cond W. +Qed. + +End AllDefLemmas. + +Hint Resolve umall_undef umall0 : core. + + +(***********) +(* um_allb *) +(***********) + +(* decidable um_all *) + +Section MapAllDecidable. +Variables (K : ordType) (C : pred K) (V : Type) (U : union_map_class C V). +Implicit Types (f : U) (p : pred (K*V)). + +Definition um_allb p f := um_count p f == size (dom f). + +Lemma umallb_undef p : um_allb p undef. +Proof. by rewrite /um_allb umcnt_undef dom_undef. Qed. + +Lemma umallb0 p : um_allb p Unit. +Proof. by rewrite /um_allb umcnt0 dom0. Qed. + +Lemma umallbP p f : reflect (forall x, x \In f -> p x) (um_allb p f). +Proof. +apply/(equivP idP); split=>[/eqP | H]. +- by rewrite umcnt_memT=>H [k v]; apply: H. +by apply/eqP; rewrite umcnt_memT=>k v; apply: H. +Qed. + +Lemma umallbUn p f1 f2 : + valid (f1 \+ f2) -> um_allb p (f1 \+ f2) = um_allb p f1 && um_allb p f2. +Proof. +move=>W; apply/idP/idP. +- by move/umallbP=>H; apply/andP; split; apply/umallbP=>x X; + apply: H; [apply: InL | apply: InR]. +clear W. +case/andP=>/umallbP X1 /umallbP X2; apply/umallbP=>x. +by case/InUn; [apply: X1 | apply: X2]. +Qed. + +Lemma umallbPt p k v : C k -> um_allb p (pts k v) = p (k, v). +Proof. by move=>C'; rewrite /um_allb umcntPt domPtK C' /=; case: (p (k, v)). Qed. + +Lemma umallbPtUn p k v f : + valid (pts k v \+ f) -> + um_allb p (pts k v \+ f) = p (k, v) && um_allb p f. +Proof. by move=>W; rewrite umallbUn // umallbPt // (validPtUn_cond W). Qed. + +Lemma umallbU p k v f : + valid (upd k v f) -> + um_allb p (upd k v f) = p (k, v) && um_allb p (free f k). +Proof. +rewrite validU=>/andP [W1 W2]; rewrite upd_eta // umallbPtUn //. +by rewrite validPtUn W1 validF W2 domF inE eq_refl. +Qed. + +Lemma umallbF p k f : um_allb p f -> um_allb p (free f k). +Proof. by move/umallbP=>H; apply/umallbP=>kv /InF [_ _ /H]. Qed. + +End MapAllDecidable. + + +(************************************) +(* Zipping a relation over two maps *) +(************************************) + +(* Binary version of um_all, where comparison is done over values of *) +(* equal keys in both maps. *) + +Section BinaryMapAll. +Variables (K : ordType) (C : pred K) (V : Type) (U : union_map_class C V) (P : V -> V -> Prop). +Implicit Types (k : K) (v : V) (f : U). + +Definition um_all2 (f1 f2 : U) := forall k, optionR P (find k f1) (find k f2). + +Lemma umall2_refl f : Reflexive P -> um_all2 f f. +Proof. by move=>R k; apply: Reflexive_optionR. Qed. + +Lemma umall2_sym : Symmetric P -> Symmetric um_all2. +Proof. by move=>S x y; split=>H k; apply/Symmetric_optionR. Qed. + +Lemma umall2_trans : Transitive P -> Transitive um_all2. +Proof. by move=>T x y z H1 H2 k; apply: Transitive_optionR (H1 k) (H2 k). Qed. + +Lemma umall2_eq : Equivalence_rel P -> Equivalence_rel um_all2. +Proof. +case/Equivalence_relS=>R S T; apply/Equivalence_relS; split. +- by move=>f; apply: umall2_refl. +- by apply: umall2_sym. +by apply: umall2_trans. +Qed. + +Lemma umall20 : um_all2 Unit Unit. +Proof. by move=>k; rewrite /optionR /= find0E. Qed. + +Lemma umall2_undef : um_all2 undef undef. +Proof. by move=>k; rewrite /optionR /= find_undef. Qed. + +Lemma umall2_dom f1 f2 : um_all2 f1 f2 -> dom f1 = dom f2. +Proof. +move=>H; apply/domE=>k; apply/idP/idP; +move: (H k); rewrite /optionR /=; +by case: dom_find =>// v ->; case: dom_find=>// ->. +Qed. + +Lemma umall2_umfilt f1 f2 p : + (forall k v1 v2, + (k, v1) \In f1 -> (k, v2) \In f2 -> P v1 v2 -> + p (k, v1) = p (k, v2)) -> + um_all2 f1 f2 -> um_all2 (um_filter p f1) (um_filter p f2). +Proof. +move=>E H k; move: (H k); rewrite /optionR /= !find_umfilt. +case E1: (find k f1)=>[v1|]; case E2: (find k f2)=>[v2|] // X. +move/In_find: E1=>E1; move/In_find: E2=>E2. +by rewrite -(E k v1 v2 E1 E2) //; case: ifP. +Qed. + +Lemma umall2_umfilt_inv f1 f2 p : + um_all2 (um_filter p f1) (um_filter p f2) -> + forall k, match find k f1, find k f2 with + Some v1, Some v2 => p (k, v1) && p (k, v2) -> P v1 v2 + | Some v1, None => ~~ p (k, v1) + | None, Some v2 => ~~ p (k, v2) + | None, None => True + end. +Proof. +move=>H k; move: (H k); rewrite !find_umfilt. +case: (find k f1)=>[a1|]; case: (find k f2)=>[a2|] //. +- by case: ifP; case: ifP. +- by case: ifP. +by case: ifP. +Qed. + +Lemma umall2_umfilt_ord f1 f2 t : + um_all2 (um_filter (fun kv => ord kv.1 t) f1) + (um_filter (fun kv=>ord kv.1 t) f2) <-> + forall k, ord k t -> optionR P (find k f1) (find k f2). +Proof. +split=>[H k X|H k]; last first. +- move: (H k); rewrite !find_umfilt. + case: (find k f1)=>[a1|]; case: (find k f2)=>[a2|] //=; + by case: ifP=>// X; apply. +move/umall2_umfilt_inv/(_ k): H=>/=. +case: (find k f1)=>[v1|]; case: (find k f2)=>[v2|] //=. +- by rewrite X; apply. +- by rewrite X. +by rewrite X. +Qed. + +Lemma umall2_umfilt_oleq f1 f2 t : + um_all2 (um_filter (fun kv => oleq kv.1 t) f1) + (um_filter (fun kv=>oleq kv.1 t) f2) <-> + forall k, oleq k t -> optionR P (find k f1) (find k f2). +Proof. +split=>[H k X|H k]; last first. +- move: (H k); rewrite !find_umfilt. + case: (find k f1)=>[a1|]; case: (find k f2)=>[a2|] //=; + by case: ifP=>// X; apply. +move/umall2_umfilt_inv/(_ k): H=>/=. +case: (find k f1)=>[v1|]; case: (find k f2)=>[v2|] //=. +- by rewrite X; apply. +- by rewrite X. +by rewrite X. +Qed. + +Lemma umall2_umcnt f1 f2 p : + (forall k v1 v2, + (k, v1) \In f1 -> (k, v2) \In f2 -> P v1 v2 -> + p (k, v1) = p (k, v2)) -> + um_all2 f1 f2 -> um_count p f1 = um_count p f2. +Proof. by move=>*; congr size; apply: umall2_dom; apply: umall2_umfilt. Qed. + +Lemma umall2_find X f1 f2 (f : option V -> X) t : + (forall k v1 v2, + (k, v1) \In f1 -> (k, v2) \In f2 -> P v1 v2 -> + f (Some v1) = f (Some v2)) -> + um_all2 f1 f2 -> f (find t f1) = f (find t f2). +Proof. +move=>E /(_ t). +case E1: (find t f1)=>[v1|]; case E2: (find t f2)=>[v2|] //=. +by move/In_find: E1=>E1; move/In_find: E2=>E2; apply: E E1 E2. +Qed. + +Lemma umall2fUn f f1 f2 : + Reflexive P -> + valid f1 = valid f2 -> um_all2 f1 f2 -> um_all2 (f \+ f1) (f \+ f2). +Proof. +move=>R Ev X; have De : dom_eq f1 f2. +- by apply/domeqP; rewrite (umall2_dom X) Ev. +move=>k; case V1 : (valid (f \+ f1)) (domeqVUnE (domeq_refl f) De)=>/esym V2; +last first. +- by move/negbT/invalidE: V1 V2=>-> /negbT/invalidE ->; rewrite find_undef. +rewrite /optionR /= !findUnL //; case: ifP=>D; last by apply: X. +by case/In_domX: D=>v /In_find ->; apply: R. +Qed. + +Lemma umall2Unf f f1 f2 : + Reflexive P -> + valid f1 = valid f2 -> um_all2 f1 f2 -> um_all2 (f1 \+ f) (f2 \+ f). +Proof. by rewrite -!(joinC f); apply: umall2fUn. Qed. + +Lemma umall2Un f1 f2 g1 g2 : + Reflexive P -> Transitive P -> + valid f1 = valid f2 -> valid g1 = valid g2 -> + um_all2 f1 f2 -> um_all2 g1 g2 -> + um_all2 (f1 \+ g1) (f2 \+ g2). +Proof. +move=>R T Ef Eg Uf Ug; apply: (@umall2_trans T (f2 \+ g1)); +by [apply: umall2Unf | apply: umall2fUn]. +Qed. + +Lemma umall2Pt2 k1 k2 v1 v2 : + um_all2 (pts k1 v1) (pts k2 v2) <-> + if k1 == k2 then C k1 -> P v1 v2 + else ~~ C k1 && ~~ C k2. +Proof. +split; last first. +- case: eqP=>[-> H|N /andP [C1 C2]] k. + - by rewrite /optionR /= !findPt2; case: ifP=>// /andP []. + by rewrite /optionR /= !findPt2 (negbTE C1) (negbTE C2) !andbF. +move=>X; move: (X k1) (X k2). +rewrite /optionR !findPt !findPt2 (eq_sym k2 k1) /=. +case: (k1 =P k2)=>[<-|N] /=; first by case: ifP. +by do 2![case: ifP]. +Qed. + +Lemma umall2Pt k v1 v2 : + C k -> + um_all2 (pts k v1) (pts k v2) <-> P v1 v2. +Proof. by move=>C'; rewrite umall2Pt2 eq_refl; split=>//; apply. Qed. + +Lemma umall2cancel k v1 v2 f1 f2 : + valid (pts k v1 \+ f1) -> valid (pts k v2 \+ f2) -> + um_all2 (pts k v1 \+ f1) (pts k v2 \+ f2) -> + P v1 v2 /\ um_all2 f1 f2. +Proof. +move=>V1 V2 X; split=>[|x]. +- by move: (X k); rewrite !findPtUn. +move: (X x); rewrite !findPtUn2 //; case: ifP=>// /eqP -> _. +by case: dom_find (validPtUnD V1)=>// ->; case: dom_find (validPtUnD V2)=>// ->. +Qed. + +Lemma umall2PtUn k v1 v2 f1 f2 : + Reflexive P -> Transitive P -> + valid f1 = valid f2 -> um_all2 f1 f2 -> + P v1 v2 -> + um_all2 (pts k v1 \+ f1) (pts k v2 \+ f2). +Proof. +move=>R T; case W : (valid (pts k v1 \+ f1)). +- move=>H1 H2 H3; apply: umall2Un=>//. + - by rewrite !validPt (validPtUn_cond W). + by apply/(@umall2Pt _ _ _ (validPtUn_cond W)). +case: validUn W=>//. +- rewrite validPt=>H _ _ _ _. + have L v : pts k v = undef :> U by apply/pts_undef. + by rewrite !L !undef_join; apply: umall2_undef. +- move=>W _; rewrite (negbTE W)=>/esym. + move/invalidE: W=>-> /negbT/invalidE -> _ _; rewrite !join_undef. + by apply: umall2_undef. +move=>x; rewrite domPt inE=>/andP [X /eqP <- D1] _ W /umall2_dom E _. +suff : ~~ valid (pts k v1 \+ f1) /\ ~~ valid (pts k v2 \+ f2). +- by case=>/invalidE -> /invalidE ->; apply: umall2_undef. +by rewrite !validPtUn X -W -E D1 (dom_valid D1). +Qed. + +Lemma umall2fK f f1 f2 : + valid (f \+ f1) -> valid (f \+ f2) -> + um_all2 (f \+ f1) (f \+ f2) -> um_all2 f1 f2. +Proof. +move=>V1 V2 X k; move: (X k); rewrite !findUnL //; case: ifP=>// D _. +by case: dom_find (dom_inNL V1 D)=>// ->; case: dom_find (dom_inNL V2 D)=>// ->. +Qed. + +Lemma umall2KL f1 f2 g1 g2 : + Equivalence_rel P -> + valid (f1 \+ g1) -> valid (f2 \+ g2) -> + um_all2 (f1 \+ g1) (f2 \+ g2) -> + um_all2 f1 f2 -> um_all2 g1 g2. +Proof. +move=>E; case/Equivalence_relS: E=>R S T. +move=>V1 V2 H1 Ha; have /(umall2_sym S) H2: um_all2 (f1 \+ g1) (f2 \+ g1). +- by apply: umall2Unf Ha=>//; rewrite (validL V1) (validL V2). +apply: umall2fK (V2) _; last first. +- by apply: umall2_trans H2 H1. +apply: domeqVUn (V1)=>//; apply/domeqP. +by rewrite (validL V1) (validL V2) (umall2_dom Ha). +Qed. + +Lemma umall2KR f1 f2 g1 g2 : + Equivalence_rel P -> + valid (f1 \+ g1) -> valid (f2 \+ g2) -> + um_all2 (f1 \+ g1) (f2 \+ g2) -> + um_all2 g1 g2 -> um_all2 f1 f2. +Proof. by rewrite (joinC f1) (joinC f2); apply: umall2KL. Qed. + +End BinaryMapAll. + + +(********************************************************) +(********************************************************) +(* Instance of union maps with trivially true condition *) +(********************************************************) +(********************************************************) + +(* We build it over the base type with a trival condition on keys. *) +(* We seal the definition to avoid any slowdowns (just in case). *) + +Module Type UMSig. +Parameter tp : ordType -> Type -> Type. + +Section Params. +Variables (K : ordType) (V : Type). +Notation tp := (tp K V). + +Parameter um_undef : tp. +Parameter defined : tp -> bool. +Parameter empty : tp. +Parameter upd : K -> V -> tp -> tp. +Parameter dom : tp -> seq K. +Parameter dom_eq : tp -> tp -> bool. +Parameter assocs : tp -> seq (K * V). +Parameter free : tp -> K -> tp. +Parameter find : K -> tp -> option V. +Parameter union : tp -> tp -> tp. +Parameter empb : tp -> bool. +Parameter undefb : tp -> bool. +Parameter pts : K -> V -> tp. + +Parameter from : tp -> @UM.base K xpredT V. +Parameter to : @UM.base K xpredT V -> tp. + +Axiom ftE : forall b, from (to b) = b. +Axiom tfE : forall f, to (from f) = f. +Axiom undefE : um_undef = to (@UM.Undef K xpredT V). +Axiom defE : forall f, defined f = UM.valid (from f). +Axiom emptyE : empty = to (@UM.empty K xpredT V). + +Axiom updE : forall k v f, upd k v f = to (UM.upd k v (from f)). +Axiom domE : forall f, dom f = UM.dom (from f). +Axiom dom_eqE : forall f1 f2, dom_eq f1 f2 = UM.dom_eq (from f1) (from f2). +Axiom assocsE : forall f, assocs f = UM.assocs (from f). +Axiom freeE : forall f k, free f k = to (UM.free (from f) k). +Axiom findE : forall k f, find k f = UM.find k (from f). +Axiom unionE : forall f1 f2, union f1 f2 = to (UM.union (from f1) (from f2)). +Axiom empbE : forall f, empb f = UM.empb (from f). +Axiom undefbE : forall f, undefb f = UM.undefb (from f). +Axiom ptsE : forall k v, pts k v = to (@UM.pts K xpredT V k v). + +End Params. + +End UMSig. + + +Module UMDef : UMSig. +Section UMDef. +Variables (K : ordType) (V : Type). + +Definition tp : Type := @UM.base K xpredT V. + +Definition um_undef := @UM.Undef K xpredT V. +Definition defined f := @UM.valid K xpredT V f. +Definition empty := @UM.empty K xpredT V. +Definition upd k v f := @UM.upd K xpredT V k v f. +Definition dom f := @UM.dom K xpredT V f. +Definition dom_eq f1 f2 := @UM.dom_eq K xpredT V f1 f2. +Definition assocs f := @UM.assocs K xpredT V f. +Definition free f k := @UM.free K xpredT V f k. +Definition find k f := @UM.find K xpredT V k f. +Definition union f1 f2 := @UM.union K xpredT V f1 f2. +Definition empb f := @UM.empb K xpredT V f. +Definition undefb f := @UM.undefb K xpredT V f. +Definition pts k v := @UM.pts K xpredT V k v. + +Definition from (f : tp) : @UM.base K xpredT V := f. +Definition to (b : @UM.base K xpredT V) : tp := b. Lemma ftE b : from (to b) = b. Proof. by []. Qed. Lemma tfE f : to (from f) = f. Proof. by []. Qed. -Lemma undefE : um_undef = to (@UM.Undef K V predT). Proof. by []. Qed. +Lemma undefE : um_undef = to (@UM.Undef K xpredT V). Proof. by []. Qed. Lemma defE f : defined f = UM.valid (from f). Proof. by []. Qed. -Lemma emptyE : empty = to (@UM.empty K V predT). Proof. by []. Qed. +Lemma emptyE : empty = to (@UM.empty K xpredT V). Proof. by []. Qed. Lemma updE k v f : upd k v f = to (UM.upd k v (from f)). Proof. by []. Qed. Lemma domE f : dom f = UM.dom (from f). Proof. by []. Qed. Lemma dom_eqE f1 f2 : dom_eq f1 f2 = UM.dom_eq (from f1) (from f2). Proof. by []. Qed. -Lemma freeE k f : free k f = to (UM.free k (from f)). Proof. by []. Qed. +Lemma assocsE f : assocs f = UM.assocs (from f). Proof. by []. Qed. +Lemma freeE f k : free f k = to (UM.free (from f) k). Proof. by []. Qed. Lemma findE k f : find k f = UM.find k (from f). Proof. by []. Qed. Lemma unionE f1 f2 : union f1 f2 = to (UM.union (from f1) (from f2)). Proof. by []. Qed. -Lemma umfiltE q f : um_filter q f = to (UM.um_filter q (from f)). -Proof. by []. Qed. Lemma empbE f : empb f = UM.empb (from f). Proof. by []. Qed. Lemma undefbE f : undefb f = UM.undefb (from f). Proof. by []. Qed. -Lemma ptsE k v : pts k v = to (@UM.pts K V predT k v). Proof. by []. Qed. +Lemma ptsE k v : pts k v = to (@UM.pts K xpredT V k v). Proof. by []. Qed. End UMDef. End UMDef. @@ -2152,14 +5549,14 @@ Notation union_map := UMDef.tp. Definition unionmapMixin K V := UMCMixin (@UMDef.ftE K V) (@UMDef.tfE K V) (@UMDef.defE K V) (@UMDef.undefE K V) (@UMDef.emptyE K V) (@UMDef.updE K V) - (@UMDef.domE K V) (@UMDef.dom_eqE K V) (@UMDef.freeE K V) - (@UMDef.findE K V) (@UMDef.unionE K V) (@UMDef.umfiltE K V) + (@UMDef.domE K V) (@UMDef.dom_eqE K V) (@UMDef.assocsE K V) + (@UMDef.freeE K V) (@UMDef.findE K V) (@UMDef.unionE K V) (@UMDef.empbE K V) (@UMDef.undefbE K V) (@UMDef.ptsE K V). Canonical union_mapUMC K V := Eval hnf in UMC (union_map K V) (@unionmapMixin K V). -(* we add the canonical projections matching against naked type *) +(* We add the canonical projections matching against naked type *) (* thus, there are two ways to get a PCM for a union map: *) (* generic one going through union_map_classPCM, and another by going *) (* through union_mapPCM. Luckily, they produce equal results *) @@ -2182,56 +5579,28 @@ Canonical union_mapTPCM K V := Eval hnf in TPCM (union_map K V) (@union_mapTPCMMix K V). Definition union_map_eqMix K (V : eqType) := - @union_map_class_eqMix K V _ _ (@unionmapMixin K V). + @union_map_class_eqMix K xpredT V _ (@unionmapMixin K V). Canonical union_map_eqType K (V : eqType) := Eval hnf in EqType (union_map K V) (@union_map_eqMix K V). - Definition um_pts (K : ordType) V (k : K) (v : V) := - @UMC.pts K V (@union_mapUMC K V) k v. + @UMC.pts K xpredT V (@union_mapUMC K V) k v. Notation "x \\-> v" := (@um_pts _ _ x v) (at level 30). + (* Finite sets are just union maps with unit range *) Notation fset T := (@union_map T unit). Notation "# x" := (x \\-> tt) (at level 20). -(* Does the notation work? *) -Lemma xx : 1 \\-> true = 1 \\-> false. -Abort. - -(* does the pcm and the equality type work? *) -Lemma xx : ((1 \\-> true) \+ (2 \\-> false)) == (1 \\-> false). -simpl. -rewrite joinC. -Abort. - -(* can we use the base type? *) -Lemma xx (x : union_map nat_ordType nat) : x \+ x == x \+ x. -Abort. - -(* the default dom' lemmas don't work nicely *) -Lemma xx (f : union_map nat_ordType nat) : 3 \in dom (free 2 f). -try by rewrite domF' /=. -rewrite (@domF _ _ (union_mapUMC _ _)). -Abort. - -(* but the dom lemmas do work nicely *) -Lemma xx (f : union_map nat_ordType nat) : 3 \in dom (free 2 f). -rewrite domF /=. -Abort. - -Lemma xx : 1 \\-> (1 \\-> 3) = 2 \\-> (7 \\-> 3). -Abort. - (* Since the union map constructors are opaque, the decidable equality *) (* does not simplify automatically; we need to export explicit equations *) -(* for simplification *) -(* so far, the ones I need are really just for points-to *) +(* for simplification. *) +(* So far, the ones we need are really just for points-to. *) Section UMDecidableEquality. -Variables (K : ordType) (V : eqType) (U : union_map_class K V). +Variables (K : ordType) (V : eqType). Lemma umPtPtE (k1 k2 : K) (v1 v2 : V) : (k1 \\-> v1 == k2 \\-> v2) = (k1 == k2) && (v1 == v2). @@ -2241,23 +5610,23 @@ by rewrite {1}/eq_op /= /feq eqseq_cons andbT. Qed. Lemma umPt0E (k : K) (v : V) : (k \\-> v == Unit) = false. -Proof. by apply: (introF idP)=>/eqP/empbP; rewrite empbPt. Qed. +Proof. by apply: (introF idP)=>/eqP/unitbP; rewrite um_unitbPt. Qed. Lemma um0PtE (k : K) (v : V) : (@Unit [pcm of union_map K V] == k \\-> v) = false. Proof. by rewrite eq_sym umPt0E. Qed. -Lemma umPtUndefE (k : K) (v : V) : (k \\-> v == um_undef) = false. -Proof. by rewrite /eq_op /= /UnionMapEq.unionmap_eq /um_pts !umEX. Qed. +Lemma umPtUndefE (k : K) (v : V) : (k \\-> v == undef) = false. +Proof. by rewrite /eq_op /undef /= /UnionMapEq.unionmap_eq /um_pts !umEX. Qed. Lemma umUndefPtE (k : K) (v : V) : - ((um_undef : union_map_eqType K V) == k \\-> v) = false. + ((undef : union_map_eqType K V) == k \\-> v) = false. Proof. by rewrite eq_sym umPtUndefE. Qed. -Lemma umUndef0E : ((um_undef : union_map_eqType K V) == Unit) = false. -Proof. by apply/(introF idP)=>/eqP/empbP; rewrite empb_undef. Qed. +Lemma umUndef0E : ((undef : union_map_eqType K V) == Unit) = false. +Proof. by apply/eqP/undef0. Qed. -Lemma um0UndefE : ((Unit : union_mapPCM K V) == um_undef) = false. +Lemma um0UndefE : ((Unit : union_mapPCM K V) == undef) = false. Proof. by rewrite eq_sym umUndef0E. Qed. Lemma umPtUE (k : K) (v : V) f : (k \\-> v \+ f == Unit) = false. @@ -2267,35 +5636,33 @@ Lemma umUPtE (k : K) (v : V) f : (f \+ k \\-> v == Unit) = false. Proof. by rewrite joinC umPtUE. Qed. Lemma umPtUPtE (k1 k2 : K) (v1 v2 : V) f : - (k1 \\-> v1 \+ f == k2 \\-> v2) = [&& k1 == k2, v1 == v2 & empb f]. + (k1 \\-> v1 \+ f == k2 \\-> v2) = [&& k1 == k2, v1 == v2 & unitb f]. Proof. apply/idP/idP; last first. -- by case/and3P=>/eqP -> /eqP -> /empbP ->; rewrite unitR. +- by case/and3P=>/eqP -> /eqP -> /unitbP ->; rewrite unitR. move/eqP/(um_prime _); case=>//; case. - move/(cancelPt2 _); rewrite validPt=>/(_ (erefl _)). - by case=>->->->; rewrite !eq_refl empb0. -by move/empbP; rewrite empbPt. + by case=>->->->; rewrite !eq_refl unitb0. +by move/unitbP; rewrite um_unitbPt. Qed. Lemma umPtPtUE (k1 k2 : K) (v1 v2 : V) f : - (k1 \\-> v1 == k2 \\-> v2 \+ f) = [&& k1 == k2, v1 == v2 & empb f]. + (k1 \\-> v1 == k2 \\-> v2 \+ f) = [&& k1 == k2, v1 == v2 & unitb f]. Proof. by rewrite eq_sym umPtUPtE (eq_sym k1) (eq_sym v1). Qed. Lemma umUPtPtE (k1 k2 : K) (v1 v2 : V) f : - (f \+ k1 \\-> v1 == k2 \\-> v2) = [&& k1 == k2, v1 == v2 & empb f]. + (f \+ k1 \\-> v1 == k2 \\-> v2) = [&& k1 == k2, v1 == v2 & unitb f]. Proof. by rewrite joinC umPtUPtE. Qed. Lemma umPtUPt2E (k1 k2 : K) (v1 v2 : V) f : - (k1 \\-> v1 == f \+ k2 \\-> v2) = [&& k1 == k2, v1 == v2 & empb f]. + (k1 \\-> v1 == f \+ k2 \\-> v2) = [&& k1 == k2, v1 == v2 & unitb f]. Proof. by rewrite joinC umPtPtUE. Qed. -Definition umE := (umPtPtE, umPt0E, um0PtE, umPtUndefE, - umUndefPtE, um0UndefE, umUndef0E, - umPtUE, umUPtE, umPtUPtE, umPtPtUE, umUPtPtE, umPtUPt2E, - (* plus a bunch of safe simplifications *) - unitL, unitR, validPt, valid_unit, eq_refl, empb0, empbPt, - join_undefL, join_undefR, empb_undef). +Definition umE := ((((umPtPtE, umPt0E), (um0PtE, umPtUndefE)), + ((umUndefPtE, um0UndefE), (umUndef0E, umPtUE))), + (((umUPtE, umPtUPtE), (umPtPtUE, umUPtPtE, umPtUPt2E)), + (* plus a bunch of safe simplifications *) + ((unitL, unitR), (validPt, valid_unit))), (((eq_refl, unitb0), + (um_unitbPt, undef_join)), (join_undef, unitb_undef))). End UMDecidableEquality. - -