From 9c49b8768f9ec6adbcbb272437177c50bf037a2b Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Thu, 1 Aug 2024 18:25:03 +0200 Subject: [PATCH 01/93] blah --- core/auto.v | 102 +++ core/axioms.v | 135 ++++ core/finmap.v | 1331 ++++++++++++++++++++++++++++++++++++ core/options.v | 3 + core/ordtype.v | 473 +++++++++++++ core/pred.v | 1636 ++++++++++++++++++++++++++++++++++++++++++++ core/prelude.v | 1091 +++++++++++++++++++++++++++++ core/rtc.v | 273 ++++++++ core/seqext.v | 1779 ++++++++++++++++++++++++++++++++++++++++++++++++ core/seqperm.v | 208 ++++++ core/slice.v | 716 +++++++++++++++++++ core/uconsec.v | 953 ++++++++++++++++++++++++++ core/useqord.v | 575 ++++++++++++++++ core/uslice.v | 1007 +++++++++++++++++++++++++++ 14 files changed, 10282 insertions(+) create mode 100644 core/auto.v create mode 100644 core/axioms.v create mode 100644 core/finmap.v create mode 100644 core/options.v create mode 100644 core/ordtype.v create mode 100644 core/pred.v create mode 100644 core/prelude.v create mode 100644 core/rtc.v create mode 100644 core/seqext.v create mode 100644 core/seqperm.v create mode 100644 core/slice.v create mode 100644 core/uconsec.v create mode 100644 core/useqord.v create mode 100644 core/uslice.v diff --git a/core/auto.v b/core/auto.v new file mode 100644 index 0000000..dd77e11 --- /dev/null +++ b/core/auto.v @@ -0,0 +1,102 @@ +(* +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. +*) + +From Coq Require Import ssreflect ssrbool ssrfun. +From mathcomp Require Import ssrnat seq eqtype. +From pcm Require Import options prelude. + +(***************************************************************************) +(* Common infrastructure for syntactifying expressions in automated lemmas *) +(***************************************************************************) + +(* In cancellation algorithms for automated lemmas, we iterate over the *) +(* first expression e1 and remove each of its components from the second *) +(* expression e2. But, typically, we want to remove only one occurrence *) +(* of that component. *) +(* *) +(* First, almost always, only one occurrence will exist, lest e2 be *) +(* undefined. Thus, it's a waste of effort to traverse e2 in full once *) +(* we've found an occurrence. *) +(* *) +(* Second, in some symmetric cancellation problems, e.g., dom_eq e1 e2, *) +(* we *want* to remove only one occurrence from e2 for each component in *) +(* e1. Otherwise, we will not produce a sound reduction. E.g., *) +(* dom (x \+ x) (x \+ x) is valid, since both expressions are undef. *) +(* However, after removing x from the left side, and both x's from the *) +(* right side, we get dom x Unit, which is not valid. *) +(* *) +(* Thus, as a matter of principle, we want a filter that removes just one *) +(* occurrence of a list element. *) +(* *) +(* We write it with p pulled out in a section in order to get it to *) +(* simplify automatically. *) + +Section OneShotFilter. +Variables (A : Type) (p : pred A). + +(* a variant of filter that removes only the first occurence *) + +Fixpoint rfilter {A} (p : pred A) (ts : seq A) : seq A := + if ts is t :: ts' then if p t then ts' else t :: rfilter p ts' else [::]. + +End OneShotFilter. + +(* rfilter can also be thought of as a generalization of rem *) +Lemma rfilter_rem {T : eqType} (ts : seq T) x : + rfilter (pred1 x) ts = rem x ts. +Proof. by elim: ts=> [|t ts IH] //=; case: eqP=>//= _; rewrite IH. Qed. + +(* A canonical structure program for searching and inserting in a list *) + +Section XFind. +Variable A : Type. + +Structure tagged_elem := XTag {xuntag :> A}. + +(* DEVCOMMENT *) +(* remove? *) +(* Local Coercion untag : tagged_elem >-> A. *) +(* /DEVCOMMENT *) + +Definition extend_tag := XTag. +Definition recurse_tag := extend_tag. +Canonical found_tag x := recurse_tag x. + +(* Main structure: *) +(* - xs1 : input sequence *) +(* - xs2 : output sequence; if pivot is found, then xs2 = xs1, else *) +(* xs2 = pivot :: xs1 *) +(* - i : output index of pivot in xs2 *) + +Definition axiom xs1 xs2 i (pivot : tagged_elem) := + onth xs2 i = Some (xuntag pivot) /\ Prefix xs1 xs2. +Structure xfind (xs1 xs2 : seq A) (i : nat) := + Form {pivot :> tagged_elem; _ : axiom xs1 xs2 i pivot}. + +(* found the elements *) +Lemma found_pf x t : axiom (x :: t) (x :: t) 0 (found_tag x). +Proof. by []. Qed. +Canonical found_form x t := Form (@found_pf x t). + +(* recurse *) +Lemma recurse_pf i x xs1 xs2 (f : xfind xs1 xs2 i) : + axiom (x :: xs1) (x :: xs2) i.+1 (recurse_tag (xuntag f)). +Proof. by case: f=>pv [H1 H2]; split=>//; apply/Prefix_cons. Qed. +Canonical recurse_form i x xs1 xs2 f := Form (@recurse_pf i x xs1 xs2 f). + +(* failed to find; attach the element to output *) +Lemma extend_pf x : axiom [::] [:: x] 0 (extend_tag x). +Proof. by []. Qed. +Canonical extend_form x := Form (@extend_pf x). + +End XFind. diff --git a/core/axioms.v b/core/axioms.v new file mode 100644 index 0000000..1154652 --- /dev/null +++ b/core/axioms.v @@ -0,0 +1,135 @@ +(* +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 axioms that are used in some parts of the library. *) +(* The selected set of axioms is known to be consistent with Coq's logic. *) +(* These axioms are: *) +(* - propositional extensionality (pext); *) +(* - functional extensionality (fext). *) +(* This file also defines the dynamic type as an alias for sigT and *) +(* Jonh Major equality via equality cast. *) +(******************************************************************************) + +From Coq Require Import ssreflect ssrfun Eqdep ClassicalFacts. +From mathcomp Require Import eqtype. +From pcm Require Import options. + +(*****************************) +(* Axioms and extensionality *) +(*****************************) + +(* We're additionally using the eq_rect_eq axiom (equivalent to UIP) from + Coq.Logic.Eqdep for its two consequences: inj_pair2 and StreicherK *) + +(* extensionality is needed for domains *) +Axiom pext : forall p1 p2 : Prop, (p1 <-> p2) -> p1 = p2. +Axiom fext : forall A (B : A -> Type) (f1 f2 : forall x, B x), + (forall x, f1 x = f2 x) -> f1 = f2. + +Lemma pf_irr (P : Prop) (p1 p2 : P) : p1 = p2. +Proof. by apply/ext_prop_dep_proof_irrel_cic/@pext. Qed. + +Lemma sval_inj A P : injective (@sval A P). +Proof. +move=>[x Hx][y Hy] /= H; move: Hx Hy; rewrite H=>*. +congr exist; apply: pf_irr. +Qed. + +Lemma svalE A (P : A -> Prop) x H : sval (exist P x H) = x. +Proof. by []. Qed. + +Lemma compf1 A B (f : A -> B) : f = f \o id. +Proof. by apply: fext. Qed. + +Lemma comp1f A B (f : A -> B) : f = id \o f. +Proof. by apply: fext. Qed. + +(*****************************************) +(* Cast and John Major Equality via cast *) +(*****************************************) + +(* depends on StreicherK axiom *) + +Section Cast. +Variable (T : Type) (interp : T -> Type). + +Definition cast A B (pf : A = B) (v : interp B) : interp A := + ecast _ _ (esym pf) v. + +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) := exists pf, v = cast pf w. + +Lemma jm_refl A (v : interp A) : jmeq v v. +Proof. by exists (erefl _); 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=>[[?]|] ->; [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 subst B; rewrite !eqc. Qed. + +End Cast. + +Arguments cast {T} interp [A][B] pf v. +Arguments jmeq {T} interp [A][B] v w. + +#[export] 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). + +(* in case of eqTypes StreicherK not needed *) +Section EqTypeCast. +Variable (T : eqType) (interp : T -> Type). +Lemma eqd a (pf : a = a) (v : interp a) : cast interp pf v = v. +Proof. by rewrite eq_axiomK. Qed. +End EqTypeCast. + + +(* type dynamic is sigT *) + +Section Dynamic. +Variables (A : Type) (P : A -> Type). + +(** eta expand definitions to prevent universe inconsistencies when using + the injectivity of constructors of datatypes depending on [[dynamic]] *) + +Definition dynamic := sigT P. +Definition dyn := existT P. +Definition dyn_tp := @projT1 _ P. +Definition dyn_val := @projT2 _ P. +Definition dyn_eta := @sigT_eta _ P. +Definition dyn_injT := @eq_sigT_fst _ P. +Definition dyn_inj := @inj_pair2 _ P. + +End Dynamic. + +Prenex Implicits dyn_tp dyn_val dyn_injT dyn_inj. +Arguments dyn {T} interp {A} _ : rename. +Notation idyn v := (@dyn _ id _ v). + +Lemma dynE (A B : Type) interp (v : interp A) (w : interp B) : + jmeq interp v w <-> dyn interp v = dyn interp w. +Proof. by split=>[[pf ->]|[pf]]; subst B; [rewrite !eqc | move/dyn_inj=>->]. Qed. diff --git a/core/finmap.v b/core/finmap.v new file mode 100644 index 0000000..86f371f --- /dev/null +++ b/core/finmap.v @@ -0,0 +1,1331 @@ +(* +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 finitely supported maps with keys drawn from *) +(* an ordered type and values from an arbitrary type. *) +(******************************************************************************) + +From HB Require Import structures. +From Coq Require Import ssreflect ssrbool ssrfun. +From Coq Require Setoid. +From mathcomp Require Import ssrnat eqtype seq path. +From pcm Require Export ordtype seqperm. +From pcm Require Import options. + +Section Def. +Variables (K : ordType) (V : Type). + +Definition key (x : K * V) := x.1. +Definition value (x : K * V) := x.2. +Definition predk k := preim key (pred1 k). +Definition predCk k := preim key (predC1 k). + +Record finMap := FinMap { + seq_of : seq (K * V); + _ : sorted ord (map key seq_of)}. + +Definition finMap_for of phant (K -> V) := finMap. + +Identity Coercion finMap_for_finMap : finMap_for >-> finMap. +End Def. + +Notation "{ 'finMap' fT }" := (finMap_for (Phant fT)) + (at level 0, format "{ 'finMap' '[hv' fT ']' }") : type_scope. + +Prenex Implicits key value predk predCk seq_of. + +Section Ops. +Variables (K : ordType) (V : Type). +Notation fmap := (finMap K V). +Notation key := (@key K V). +Notation value := (@value K V). +Notation predk := (@predk K V). +Notation predCk := (@predCk K V). + +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 (eq_irrelevance H1 H2). +Qed. + +Lemma sorted_nil : sorted ord (map key [::]). Proof. by []. Qed. +Definition nil := FinMap sorted_nil. + +Definition fnd k (s : fmap) := + if (filter (predk k) (seq_of s)) is (_, v):: _ + then Some v else None. + +Fixpoint ins' (k : K) (v : V) (s : seq (K * V)) {struct s} : seq (K * V) := + if s is (k1, v1)::s1 then + if ord k k1 then (k, v)::s else + if k == k1 then (k, v)::s1 else (k1, v1)::(ins' k v s1) + else [:: (k, v)]. + +Lemma path_ins' s k1 k2 v : + ord k1 k2 -> path ord k1 (map key s) -> + path ord k1 (map key (ins' k2 v s)). +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: (ord_total k2 k'); rewrite H4 H5. +Qed. + +Lemma sorted_ins' s k v : + sorted ord (map key s) -> sorted ord (map key (ins' k v s)). +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: (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'). + +Lemma sorted_filter' k s : + sorted ord (map key s) -> sorted ord (map key (filter (predCk k) s)). +Proof. by move=>H; rewrite -filter_map sorted_filter //; apply: trans. Qed. + +Definition rem k s := let: FinMap s' p' := s in FinMap (sorted_filter' k p'). + +Lemma sorted_behd s : sorted ord (map key s) -> sorted ord (map key (behead s)). +Proof. by case: s=>//= [[??]] ?; apply: path_sorted. Qed. + +Definition behd s := let: FinMap s' p' := s in FinMap (sorted_behd 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. + +Section Laws. +Variables (K : ordType) (V : Type). +Notation fmap := (finMap K V). +Notation nil := (nil K V). + +(* `path_le` specialized to `transitive ord` *) +Lemma ord_path (x y : K) s : ord x y -> path ord y s -> path ord x s. +Proof. by apply: path_le. Qed. + +Lemma last_ins' (x : K) (v : V) s : + path ord x (map key s) -> ins' x v s = (x, v) :: s. +Proof. by elim: s=>[|[k1 v1] s IH] //=; case: ifP. Qed. + +Lemma first_ins' (k : K) (v : V) s : + (forall x, x \in map key s -> ord x k) -> + ins' k v s = rcons s (k, v). +Proof. +elim: s=>[|[k1 v1] s IH] H //=. +move: (H k1); rewrite inE eq_refl; move/(_ (erefl _)). +case: ordP=>// O _; rewrite IH //. +by move=>x H'; apply: H; rewrite inE /= H' orbT. +Qed. + +Lemma notin_path (x : K) s : path ord x s -> x \notin s. +Proof. +elim: s=>[|k s IH] //=. +rewrite inE negb_or; case/andP=>T1 T2; case: eqP=>H /=. +- by rewrite H irr in T1. +by apply: IH; apply: ord_path T2. +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: 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. + +Lemma notin_filter (x : K) s : + x \notin (map key s) -> filter (predk V x) s = [::]. +Proof. +elim: s=>[|[k v] s IH] //=. +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 /=. +elim: s1 P1 s2 P2 H=>[|[k v] s1 IH] /= P1. +- by case=>[|[k2 v2] s2 P2] //=; move/(_ k2); rewrite eq_refl. +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 H. +move: (notin_path P1) (notin_path P2)=>N1 N2. +case E: (k == k2). +- rewrite -{k2 E}(eqP E) in P2 H N2 *. + move: (H k); rewrite eq_refl=>[[E2]]; rewrite -E2 {v2 E2} in H *. + rewrite IH // => k'. + 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: (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. +Qed. + +Lemma predkN (k1 k2 : K) : predI (predk V k1) (predCk V k2) =1 + if k1 == k2 then pred0 else predk V k1. +Proof. +by move=>x; case: ifP=>H /=; [|case: eqP=>//->]; rewrite ?(eqP H) ?andbN ?H. +Qed. + +Variant supp_spec x (s : fmap) : bool -> Type := +| supp_spec_some v of fnd x s = Some v : supp_spec x s true +| supp_spec_none of fnd x s = None : supp_spec x s false. + +Lemma suppP x (s : fmap) : supp_spec x s (x \in supp s). +Proof. +move E: (x \in supp s)=>b; case: b E; move/idP; last first. +- move=>H; apply: supp_spec_none. + case E: (fnd _ _)=>[v|] //; case: H. + rewrite /supp /fnd in E *; case: s E=>/=. + elim=>[|[y w] s IH] H1 //=. + case: ifP=>H; first by rewrite (eqP H) inE eq_refl. + rewrite -topredE /= eq_sym H; apply: IH. + by apply: path_sorted H1. +case: s; elim=>[|[y w] s IH] /= H1 //; rewrite /supp /= inE in IH *. +case: eqP=>[-> _|H] //=. +- by apply: (@supp_spec_some _ _ w); rewrite /fnd /= eq_refl. +move: (path_sorted H1)=>H1'; move/(IH H1'); case=>[v H2|H2]; +[apply: (@supp_spec_some _ _ v) | apply: supp_spec_none]; +by rewrite /fnd /=; case: eqP H=>// ->. +Qed. + +Lemma suppE (s1 s2 : fmap) : supp s1 =i supp s2 <-> supp s1 = supp s2. +Proof. +split; last by move=>->. +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: (ordP k1 k2)=>/= O H1 H2. +- move/(_ k1); rewrite !inE eq_refl /= eq_sym. + 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. + - by apply: path_sorted H1. + - by apply: path_sorted 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: ordP O=>//= O _. +by move/(ord_path O)/notin_path/negbTE: H1=>->. +Qed. + +Lemma supp_nil : supp nil = [::]. Proof. by []. Qed. + +Lemma supp_nilE (s : fmap) : (supp s = [::]) <-> (s = nil). +Proof. by split=>[|-> //]; case: s; case=>// H; rewrite fmapE. Qed. + +Lemma supp_rem k (s : fmap) : + supp (rem k s) =i predI (predC1 k) (mem (supp s)). +Proof. +case: s => s H1 x; rewrite /supp inE /=. +by case H2: (x == k)=>/=; rewrite -filter_map mem_filter /= H2. +Qed. + +Lemma supp_ins k v (s : fmap) : + supp (ins k v s) =i [predU pred1 k & supp s]. +Proof. +case: s => s H x; rewrite /supp inE /=. +elim: s x k v H=>[|[k1 v1] s IH] //= x k v H. +case: ifP=>H1 /=; first by rewrite inE. +case: ifP=>H2 /=; first by rewrite !inE (eqP H2) orbA orbb. +by rewrite !inE (IH _ _ _ (path_sorted H)) orbCA. +Qed. + +Lemma fnd_empty k : fnd k nil = None. Proof. by []. Qed. + +Lemma fnd_rem k1 k2 (s : fmap) : + fnd k1 (rem k2 s) = if k1 == k2 then None else fnd k1 s. +Proof. +case: s => s H; rewrite /fnd -filter_predI (eq_filter (predkN k1 k2)). +by case: eqP=>//; rewrite filter_pred0. +Qed. + +Lemma fnd_ins k1 k2 v (s : fmap) : + fnd k1 (ins k2 v s) = if k1 == k2 then Some v else fnd k1 s. +Proof. +case: s => s H; rewrite /fnd /=. +elim: s k1 k2 v H=>[|[k' v'] s IH] //= k1 k2 v H. +- by case: ifP=>H1; [rewrite (eqP H1) eq_refl | rewrite eq_sym H1]. +case: ifP=>H1 /=. +- by case: ifP=>H2; [rewrite (eqP H2) eq_refl | rewrite (eq_sym k1) H2]. +case: ifP=>H2 /=. +- rewrite (eqP H2). + by case: ifP=>H3; [rewrite (eqP H3) eq_refl | rewrite eq_sym H3]. +case: ifP=>H3; first by rewrite -(eqP H3) eq_sym H2. +by apply: IH; apply: path_sorted H. +Qed. + +Lemma ins_rem k1 k2 v (s : fmap) : + ins k1 v (rem k2 s) = + if k1 == k2 then ins k1 v s else rem k2 (ins k1 v s). +Proof. +move: k1 k2 v s. +have L3: forall (x : K) s, + path ord x (map key s) -> filter (predCk V x) s = s. +- move=>x t; move/notin_path; elim: t=>[|[k3 v3] t IH] //=. + rewrite inE negb_or; case/andP=>T1 T2. + by rewrite eq_sym T1 IH. +have L5: forall (x : K) (v : V) s, + sorted ord (map key s) -> ins' x v (filter (predCk V x) s) = ins' x v s. +- move=>x v s; elim: s x v=>[|[k' v'] s IH] x v //= H. + case H1: (ord x k'). + - case H2: (k' == x)=>/=; first by rewrite (eqP H2) irr in H1. + by rewrite H1 L3 //; apply: ord_path H1 H. + case H2: (k' == x)=>/=. + - rewrite (eqP H2) eq_refl in H *. + by rewrite L3 //; apply: last_ins' H. + rewrite eq_sym H2 H1 IH //. + by apply: path_sorted H. +move=>k1 k2 v [s H]. +case: ifP=>H1; rewrite /ins /rem fmapE /=. +- rewrite {k1 H1}(eqP H1). + elim: s k2 v H=>[|[k' v'] s IH] //= k2 v H. + case H1: (k' == k2)=>/=. + - rewrite eq_sym H1 (eqP H1) irr in H *. + by rewrite L3 // last_ins'. + rewrite eq_sym H1; case: ifP=>H3. + - by rewrite L3 //; apply: ord_path H3 H. + by rewrite L5 //; apply: path_sorted H. +elim: s k1 k2 H1 H=>[|[k' v'] s IH] //= k1 k2 H1 H; first by rewrite H1. +case H2: (k' == k2)=>/=. +- rewrite (eqP H2) in H *; rewrite H1. + case H3: (ord k1 k2)=>/=. + - by rewrite H1 eq_refl /= last_ins' // L3 //; apply: ord_path H. + by rewrite eq_refl /= IH //; apply: path_sorted H. +case H3: (ord k1 k')=>/=; first by rewrite H1 H2. +case H4: (k1 == k')=>/=; first by rewrite H1. +by rewrite H2 IH //; apply: path_sorted H. +Qed. + +Lemma ins_ins k1 k2 v1 v2 (s : fmap) : + ins k1 v1 (ins k2 v2 s) = if k1 == k2 then ins k1 v1 s + else ins k2 v2 (ins k1 v1 s). +Proof. +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: (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: (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: (ordP k2 k3) H2 H3. + by case: (ordP k1 k3) (trans H3 H2) H4. +- rewrite -(eqP H2) {H2} (H1). + case: (ordP k1 k2) (H1)=>//= H2 _; rewrite ?irr ?eq_refl //. + rewrite eq_sym H1. + by case: (ordP k2 k1) H1 H2. +case: (ordP k1 k3)=>H3 /=. +- rewrite eq_sym H1. + case: (ordP k2 k1) H1 (trans H3 H2)=>//. + by case: (ordP k2 k3) H2=>//=. +- rewrite (eqP H3). + 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. +Qed. + +Lemma rem_empty k : rem k nil = nil. +Proof. by rewrite fmapE. Qed. + +Lemma rem_rem k1 k2 (s : fmap) : + rem k1 (rem k2 s) = if k1 == k2 then rem k1 s else rem k2 (rem k1 s). +Proof. +rewrite /rem; case: s => s H /=. +case H1: (k1 == k2); rewrite fmapE /= -!filter_predI; apply: eq_filter=>x /=. +- by rewrite (eqP H1) andbb. +by rewrite andbC. +Qed. + +Lemma rem_ins k1 k2 v (s : fmap) : + rem k1 (ins k2 v s) = + if k1 == k2 then rem k1 s else ins k2 v (rem k1 s). +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: (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: (ordP k2 k3)=>H2 /=. +- rewrite eq_sym H1 /=. + 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: (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: (ordP k2 k3) H2=>//= H2 _. +by rewrite IH //; apply: path_sorted H. +Qed. + +Lemma rem_supp k (s : fmap) : + k \notin supp s -> rem k s = s. +Proof. +case: s => s H1; rewrite /supp !fmapE /= => H2. +elim: s H1 H2=>[|[k1 v1] s1 IH] //=; move/path_sorted=>H1. +rewrite inE negb_or; case/andP=>H2; move/(IH H1)=>H3. +by rewrite eq_sym H2 H3. +Qed. + +Lemma fnd_supp k (s : fmap) : + k \notin supp s -> fnd k s = None. +Proof. by case: suppP. Qed. + +Lemma fnd_supp_in k (s : fmap) : + k \in supp s -> exists v, fnd k s = Some v. +Proof. by case: suppP=>[v|]; [exists v|]. Qed. + +Lemma cancel_ins k v (s1 s2 : fmap) : + k \notin (supp s1) -> k \notin (supp s2) -> + ins k v s1 = ins k v s2 -> s1 = s2. +Proof. +move: s1 s2=>[s1 p1][s2 p2]; rewrite !fmapE /supp /= {p1 p2}. +elim: s1 k v s2=>[k v s2| [k1 v1] s1 IH1 k v s2] /=. +- case: s2=>[| [k2 v2] s2] //= _. + rewrite inE negb_or; case/andP=>H1 _; case: ifP=>// _. + by rewrite (negbTE H1); case=>E; rewrite E eq_refl in H1. +rewrite inE negb_or; case/andP=>H1 H2 H3. +case: ifP=>H4; case: s2 H3=>[| [k2 v2] s2] //=. +- rewrite inE negb_or; case/andP=>H5 H6. + case: ifP=>H7; first by case=>->->->. + by rewrite (negbTE H5); case=>E; rewrite E eq_refl in H5. +- by rewrite (negbTE H1)=>_; case=>E; rewrite E eq_refl in H1. +rewrite inE negb_or (negbTE H1); case/andP=>H5 H6. +rewrite (negbTE H5); case: ifP=>H7 /=. +- by case=>E; rewrite E eq_refl in H1. +by case=>->-> H; congr (_ :: _); apply: IH1 H. +Qed. + +End Laws. + +Section Append. +Variable (K : ordType) (V : Type). +Notation fmap := (finMap K V). +Notation nil := (nil K V). + +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. +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: (ordP k k2)=>H /=. +- by rewrite H2 H H4. +- by rewrite H2 (eqP H) H4. +rewrite H3 /=. +have H6: sorted ord (map key s) by apply: path_sorted H5. +by move: (IH H6 k2 k v H H4); case: s {IH p H4 H5} H6. +Qed. + +Lemma path_supp_ins_inv k1 k v (s : fmap) : + path ord k (supp s) -> path ord k1 (supp (ins k v s)) -> + ord k1 k && path ord k1 (supp s). +Proof. +case: s=>s p; rewrite /supp /= => H1; rewrite last_ins' //=. +by case/andP=>H2 H3; rewrite H2; apply: ord_path H3. +Qed. + +(* forward induction principle *) +Lemma fmap_ind' (P : fmap -> Prop) : + P nil -> (forall k v s, path ord k (supp s) -> P s -> P (ins k v s)) -> + forall s, P s. +Proof. +move=>H1 H2; case; elim=>[|[k v] s IH] /= H. +- by rewrite (_ : FinMap _ = nil); last by rewrite fmapE. +have S: sorted ord (map key s) by apply: path_sorted H. +rewrite (_ : FinMap _ = ins k v (FinMap S)); last by rewrite fmapE /= last_ins'. +by apply: H2. +Qed. + +(* backward induction principle *) +Lemma fmap_ind'' (P : fmap -> Prop) : + P nil -> (forall k v s, (forall x, x \in supp s -> ord x k) -> + P s -> P (ins k v s)) -> + forall s, P s. +Proof. +move=>H1 H2; case; elim/last_ind=>[|s [k v] IH] /= H. +- by rewrite (_ : FinMap _ = nil); last by rewrite fmapE. +have Sb: subseq (map key s) (map key (rcons s (k, v))). +- by elim: s {IH H}=>[|x s IH] //=; rewrite eq_refl. +have S : sorted ord (map key s). +- by apply: subseq_sorted Sb H; apply: ordtype.trans. +have T : forall x : K, x \in map key s -> ord x k. +- elim: {IH Sb S} s H=>[|[k1 v1] s IH] //= L x. + rewrite inE; case/orP; last by apply: IH; apply: path_sorted L. + move/eqP=>->; elim: s {IH} L=>[|[x1 w1] s IH] /=; first by rewrite andbT. + by case/andP=>O /(ord_path O) /IH. +rewrite (_ : FinMap _ = ins k v (FinMap S)). +- by apply: H2 (IH _)=>x /T. +by rewrite fmapE /= first_ins'. +Qed. + + +Fixpoint fcat' (s1 : fmap) (s2 : seq (K * V)) {struct s2} : fmap := + if s2 is (k, v)::t then fcat' (ins k v s1) t else s1. + +Definition fcat s1 s2 := fcat' s1 (seq_of s2). + +Lemma fcat_ins' k v s1 s2 : + k \notin (map key s2) -> fcat' (ins k v s1) s2 = ins k v (fcat' s1 s2). +Proof. +move=>H; elim: s2 k v s1 H=>[|[k2 v2] s2 IH] k1 v1 s1 //=. +rewrite inE negb_or; case/andP=>H1 H2. +by rewrite -IH // ins_ins eq_sym (negbTE H1). +Qed. + +Lemma fcat_nil' s : fcat' nil (seq_of s) = s. +Proof. +elim/fmap_ind': s=>[|k v s L IH] //=. +by rewrite seqof_ins //= (_ : FinMap _ = ins k v nil) // + fcat_ins' ?notin_path // IH. +Qed. + +Lemma fcat0s s : fcat nil s = s. Proof. by apply: fcat_nil'. Qed. +Lemma fcats0 s : fcat s nil = s. Proof. by []. Qed. + +Lemma fcat_inss k v s1 s2 : + k \notin supp s2 -> fcat (ins k v s1) s2 = ins k v (fcat s1 s2). +Proof. by case: s2=>s2 p2 H /=; apply: fcat_ins'. Qed. + +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: (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: (ordP k2 k1) H2 => // H2 _. +rewrite {1}/fcat seqof_ins //= fcat_ins' ?notin_path // IH ?notin_path //. +rewrite ins_ins; case: (ordP k2 k1) H2 => // H2 _; congr (ins _ _ _). +by rewrite -/(fcat s2 (ins k2 v2 s1)) IH. +Qed. + +Lemma fcat_rems k s1 s2 : + k \notin supp s2 -> fcat (rem k s1) s2 = rem k (fcat s1 s2). +Proof. +elim/fmap_ind': s2 k s1=>[|k2 v2 s2 H IH] k1 v1. +- by rewrite !fcats0. +rewrite supp_ins inE /= negb_or; case/andP=>H1 H2. +by rewrite fcat_sins IH // ins_rem eq_sym (negbTE H1) -fcat_sins. +Qed. + +Lemma fcat_srem k s1 s2 : + k \notin supp s1 -> fcat s1 (rem k s2) = rem k (fcat s1 s2). +Proof. +elim/fmap_ind': s2 k s1=>[|k2 v2 s2 H IH] k1 s1. +- rewrite rem_empty fcats0. + elim/fmap_ind': s1=>[|k3 v3 s3 H1 IH]; first by rewrite rem_empty. + rewrite supp_ins inE /= negb_or. + case/andP=>H2; move/IH=>E; rewrite {1}E . + by rewrite ins_rem eq_sym (negbTE H2). +move=>H1; rewrite fcat_sins rem_ins; case: ifP=>E. +- by rewrite rem_ins E IH. +by rewrite rem_ins E -IH // -fcat_sins. +Qed. + +Lemma fnd_fcat k s1 s2 : + fnd k (fcat s1 s2) = + if k \in supp s2 then fnd k s2 else fnd k s1. +Proof. +elim/fmap_ind': s2 k s1=>[|k2 v2 s2 H IH] k1 s1. +- by rewrite fcats0. +rewrite supp_ins inE /=; case: ifP; last first. +- move/negbT; rewrite negb_or; case/andP=>H1 H2. + by rewrite fcat_sins fnd_ins (negbTE H1) IH (negbTE H2). +case/orP; first by move/eqP=><-; rewrite fcat_sins !fnd_ins eq_refl. +move=>H1; rewrite fcat_sins !fnd_ins. +by case: ifP=>//; rewrite IH H1. +Qed. + +Lemma supp_fcat s1 s2 : supp (fcat s1 s2) =i [predU supp s1 & supp s2]. +Proof. +elim/fmap_ind': s2 s1=>[|k v s L IH] s1. +- by rewrite supp_nil fcats0 => x; rewrite inE /= orbF. +rewrite fcat_sins ?notin_path // => x. +rewrite supp_ins !inE /=. +case E: (x == k)=>/=. +- by rewrite supp_ins inE /= E orbT. +by rewrite IH supp_ins inE /= inE /= E. +Qed. + +End Append. + +(* an induction principle for pairs of finite maps with equal support *) + +Section FMapInd. +Variables (K : ordType) (V : Type). +Notation fmap := (finMap K V). +Notation nil := (@nil K V). + +Lemma supp_eq_ins (s1 s2 : fmap) k1 k2 v1 v2 : + path ord k1 (supp s1) -> path ord k2 (supp s2) -> + supp (ins k1 v1 s1) =i supp (ins k2 v2 s2) -> + k1 = k2 /\ supp s1 =i supp s2. +Proof. +move=>H1 H2 H; move: (H k1) (H k2). +rewrite !supp_ins !inE /= !eq_refl (eq_sym k2). +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). +rewrite !supp_ins !inE /=; case: eqP=>//= -> _. +by rewrite (negbTE (notin_path H1)) (negbTE (notin_path H2)). +Qed. + +Lemma fmap_ind2 (P : fmap -> fmap -> Prop) : + P nil nil -> + (forall k v1 v2 s1 s2, + path ord k (supp s1) -> path ord k (supp s2) -> + P s1 s2 -> P (ins k v1 s1) (ins k v2 s2)) -> + forall s1 s2, supp s1 =i supp s2 -> P s1 s2. +Proof. +move=>H1 H2; elim/fmap_ind'=>[|k1 v1 s1 T1 IH1]; +elim/fmap_ind'=>[|k2 v2 s2 T2 _] //. +- by move/(_ k2); rewrite supp_ins inE /= eq_refl supp_nil. +- by move/(_ k1); rewrite supp_ins inE /= eq_refl supp_nil. +by case/supp_eq_ins=>// E; rewrite -{k2}E in T2 *; move/IH1; apply: H2. +Qed. + +End FMapInd. + + +Section Filtering. +Variables (K : ordType) (V : Type). +Notation fmap := (finMap K V). +Notation nil := (nil K V). + +Definition kfilter' (p : pred K) (s : fmap) := + filter (fun kv => p kv.1) (seq_of s). + +Lemma sorted_kfilter (p : pred K) s : sorted ord (map key (kfilter' p s)). +Proof. +by case: s=>s H; rewrite -filter_map path.sorted_filter //; apply: trans. +Qed. + +Definition kfilter (p : pred K) (s : fmap) := FinMap (sorted_kfilter p s). + +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 H. +by case E: (p k)=>//=; rewrite H. +Qed. + +Lemma kfilt_nil (p : pred K) : kfilter p nil = nil. +Proof. by apply/fmapE. Qed. + +Lemma fnd_kfilt (p : pred K) k (s : fmap) : + fnd k (kfilter p s) = + if p k then fnd k s else None. +Proof. +case: s; rewrite /fnd /kfilter /kfilter' /=. +elim=>[|[k1 v] s IH] /=; first by case: ifP. +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. +Qed. + +Lemma kfilt_ins (p : pred K) k v (s : fmap) : + kfilter p (ins k v s) = + if p k then ins k v (kfilter p s) else kfilter p s. +Proof. +apply/fmapP=>k2; case: ifP=>E1. +- by rewrite fnd_ins !fnd_kfilt fnd_ins; case: eqP=>// ->; rewrite E1. +by rewrite !fnd_kfilt fnd_ins; case: eqP=>// ->; rewrite E1. +Qed. + +Lemma rem_kfilt (p : pred K) k (s : fmap) : + rem k (kfilter p s) = + if p k then kfilter p (rem k s) else kfilter p s. +Proof. +apply/fmapP=>k2; case: ifP=>E1. +- by rewrite fnd_rem !fnd_kfilt fnd_rem; case: eqP=>// ->; rewrite E1. +by rewrite fnd_rem fnd_kfilt; case: eqP=>// ->; rewrite E1. +Qed. + +Lemma kfilt_rem (p : pred K) k (s : fmap) : + kfilter p (rem k s) = + if p k then rem k (kfilter p s) else kfilter p s. +Proof. +apply/fmapP=>k2; case: ifP=>E1. +- by rewrite fnd_kfilt !fnd_rem fnd_kfilt; case: eqP=>// ->; rewrite E1. +by rewrite !fnd_kfilt fnd_rem; case: eqP=>// ->; rewrite E1. +Qed. + +(* filter and fcat *) + +Lemma kfilt_fcat (p : pred K) (s1 s2 : fmap) : + kfilter p (fcat s1 s2) = fcat (kfilter p s1) (kfilter p s2). +Proof. +apply/fmapP=>k; rewrite fnd_kfilt !fnd_fcat !fnd_kfilt supp_kfilt mem_filter. +by case: ifP. +Qed. + +Lemma kfilter_pred0 (s : fmap) : kfilter pred0 s = nil. +Proof. by apply/fmapE; rewrite /= /kfilter' filter_pred0. Qed. + +Lemma kfilter_predT (s : fmap) : kfilter predT s = s. +Proof. by apply/fmapE; rewrite /= /kfilter' filter_predT. Qed. + +Lemma kfilter_predI p1 p2 (s : fmap) : + kfilter (predI p1 p2) s = kfilter p1 (kfilter p2 s). +Proof. by apply/fmapE; rewrite /= /kfilter' filter_predI. Qed. + +Lemma kfilter_predU p1 p2 (s : fmap) : + kfilter (predU p1 p2) s = fcat (kfilter p1 s) (kfilter p2 s). +Proof. +apply/fmapP=>k; rewrite fnd_kfilt fnd_fcat !fnd_kfilt supp_kfilt mem_filter. +rewrite inE /=; case: (ifP (p1 k)); case: (ifP (p2 k))=>//=; +by [case: ifP | case: suppP]. +Qed. + +Lemma eq_in_kfilter p1 p2 s : + {in supp s, p1 =1 p2} -> kfilter p1 s = kfilter p2 s. +Proof. +move=>H; apply/fmapE; rewrite /= /kfilter'. +case: s H; rewrite /supp /=; elim=>[|[k v] s IH] //=. +move/path_sorted/IH=>{IH} H H1. +have ->: p1 k = p2 k by apply: H1; rewrite inE /= eq_refl. +by rewrite H // => x E; apply: H1; rewrite inE /= E orbT. +Qed. + +End Filtering. + +Section DisjointUnion. +Variable (K : ordType) (V : Type). +Notation fmap := (finMap K V). +Notation nil := (nil K V). + +Definition disj (s1 s2 : fmap) := + all (predC (fun x => x \in supp s2)) (supp s1). + +Variant disj_spec (s1 s2 : fmap) : bool -> Type := +| disj_true of (forall x, x \in supp s1 -> x \notin supp s2) : + disj_spec s1 s2 true +| disj_false x of x \in supp s1 & x \in supp s2 : + disj_spec s1 s2 false. + +Lemma disjP s1 s2 : disj_spec s1 s2 (disj s1 s2). +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. +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. +Proof. +case: disjP; case: disjP=>//. +- by move=>x H1 H2; move/(_ x H2); rewrite H1. +by move=>H1 x H2; move/H1; rewrite H2. +Qed. + +Lemma disj_nil (s : fmap) : disj s nil. +Proof. by case: disjP. Qed. + +Lemma disj_ins k v (s1 s2 : fmap) : + disj s1 (ins k v s2) = (k \notin supp s1) && (disj s1 s2). +Proof. +case: disjP=>[H|x H1]. +- case E: (k \in supp s1)=>/=. + - by move: (H _ E); rewrite supp_ins inE /= eq_refl. + case: disjP=>// x H1 H2. + by move: (H _ H1); rewrite supp_ins inE /= H2 orbT. +rewrite supp_ins inE /=; case/orP=>[|H2]. +- by move/eqP=><-; rewrite H1. +rewrite andbC; case: disjP=>[H|y H3 H4] //=. +by move: (H _ H1); rewrite H2. +Qed. + +Lemma disj_rem k (s1 s2 : fmap) : + disj s1 s2 -> disj s1 (rem k s2). +Proof. +case: disjP=>// H _; case: disjP=>// x; move/H. +by rewrite supp_rem inE /= andbC; move/negbTE=>->. +Qed. + +Lemma disj_remE k (s1 s2 : fmap) : + k \notin supp s1 -> disj s1 (rem k s2) = disj s1 s2. +Proof. +move=>H; case: disjP; case: disjP=>//; last first. +- move=>H1 x; move/H1; rewrite supp_rem inE /= => E. + by rewrite (negbTE E) andbF. +move=>x H1 H2 H3; move: (H3 x H1) H. +rewrite supp_rem inE /= negb_and H2 orbF negbK. +by move/eqP=><-; rewrite H1. +Qed. + +Lemma disj_fcat (s s1 s2 : fmap) : + disj s (fcat s1 s2) = disj s s1 && disj s s2. +Proof. +elim/fmap_ind': s s1 s2=>[|k v s L IH] s1 s2. +- by rewrite !(disjC nil) !disj_nil. +rewrite !(disjC (ins _ _ _)) !disj_ins supp_fcat inE /= negb_or. +case: (k \in supp s1)=>//=. +case: (k \in supp s2)=>//=; first by rewrite andbF. +by rewrite -!(disjC s) IH. +Qed. + +Lemma fcatC (s1 s2 : fmap) : disj s1 s2 -> fcat s1 s2 = fcat s2 s1. +Proof. +rewrite /fcat. +elim/fmap_ind': s2 s1=>[|k v s2 L IH] s1 /=; first by rewrite fcat_nil'. +rewrite disj_ins; case/andP=>D1 D2. +by rewrite fcat_ins' // -IH // seqof_ins //= -fcat_ins' ?notin_path. +Qed. + +Lemma fcatA (s1 s2 s3 : fmap) : + disj s2 s3 -> fcat (fcat s1 s2) s3 = fcat s1 (fcat s2 s3). +Proof. +move=>H. +elim/fmap_ind': s3 s1 s2 H=>[|k v s3 L IH] s1 s2 /=; first by rewrite !fcats0. +rewrite disj_ins; case/andP=>H1 H2. +by rewrite fcat_sins ?notin_path // IH // fcat_sins ?notin_path // fcat_sins. +Qed. + +Lemma fcatAC (s1 s2 s3 : fmap) : + [&& disj s1 s2, disj s2 s3 & disj s1 s3] -> + fcat s1 (fcat s2 s3) = fcat s2 (fcat s1 s3). +Proof. by case/and3P=>H1 H2 H3; rewrite -!fcatA // (@fcatC s1 s2). Qed. + +Lemma fcatCA (s1 s2 s3 : fmap) : + [&& disj s1 s2, disj s2 s3 & disj s1 s3] -> + fcat (fcat s1 s2) s3 = fcat (fcat s1 s3) s2. +Proof. +by case/and3P=>H1 H2 H3; rewrite !fcatA // ?(@fcatC s2 s3) ?(disjC s3). +Qed. + +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. +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. +Qed. + +Lemma fcatKs (s s1 s2 : fmap) : + disj s s1 && disj s s2 -> fcat s s1 = fcat s s2 -> s1 = s2. +Proof. +case/andP=>H1 H2. +rewrite (fcatC H1) (fcatC H2); apply: fcatsK. +by rewrite -!(disjC s) H1 H2. +Qed. + +(* DEVCOMMENT *) +(* heh, a theory of submaps would be good here *) +(* but i don't have time for a decent development *) +(* so let's do a quick lemma that's needed for feaps *) +(* /DEVCOMMENT *) + +Lemma disj_kfilt p1 p2 s1 s2 : + disj s1 s2 -> disj (kfilter p1 s1) (kfilter p2 s2). +Proof. +elim/fmap_ind': s2 s1=>[|k v s _ IH] s1 /=. +- by rewrite kfilt_nil => _; case: disjP. +rewrite disj_ins; case/andP=>H1 H2; rewrite kfilt_ins. +case: ifP=>E; last by apply: IH. +rewrite disj_ins supp_kfilt mem_filter negb_and H1 orbT /=. +by apply: IH. +Qed. + +Lemma in_disj_kfilt p1 p2 s : + {in supp s, forall x, ~~ p1 x || ~~ p2 x} -> + disj (kfilter p1 s) (kfilter p2 s). +Proof. +elim/fmap_ind': s=>[|k v s _ IH] //= H. +rewrite !kfilt_ins; case: ifP=>E1; case: ifP=>E2. +- move: (H k); rewrite E1 E2 supp_ins inE /= eq_refl /=. + by move/(_ (erefl _)). +- rewrite disjC disj_ins disjC supp_kfilt mem_filter negb_and E2 /=. + by apply: IH=>x S; apply: H; rewrite supp_ins inE /= S orbT. +- rewrite disj_ins supp_kfilt mem_filter negb_and E1 /=. + by apply: IH=>x S; apply: H; rewrite supp_ins inE /= S orbT. +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. +Variables (K : ordType) (V : eqType). + +Definition feq (s1 s2 : finMap K V) := seq_of s1 == seq_of s2. + +Lemma feqP : Equality.axiom feq. +Proof. +move=>s1 s2; rewrite /feq. +case: eqP; first by move/fmapE=>->; apply: ReflectT. +by move=>H; apply: ReflectF; move/fmapE; move/H. +Qed. + +HB.instance Definition _ := hasDecEq.Build (finMap K V) feqP. +End EqType. + +(* mapping a function over a contents of a finite map *) + +Section Map. +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 (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. +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 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: ordP=>O //=. +by rewrite IH // (path_sorted H). +Qed. + +Lemma mapf_fcat s1 s2 : mapf (fcat s1 s2) = fcat (mapf s1) (mapf s2). +Proof. +elim/fmap_ind': s2 s1=>[|k v s2 H IH] s1 /=. +- rewrite fcats0; set j := FinMap _. + by rewrite (_ : j = nil K V) ?fcat0s //; apply/fmapE. +by rewrite fcat_sins mapf_ins IH -fcat_sins mapf_ins. +Qed. + +Lemma mapf_disjL s1 s2 s : mapf s1 = mapf s2 -> disj s1 s = disj s2 s. +Proof. +case: s1 s2 s=>s1 S1 [s2 S2][s S] /fmapE /=. +elim: s1 S1 s2 S2 s S=>[|[k v] s1 IH] /= S1; case=>//= [[k2 v2]] s2 S2 s S. +case=>E _; rewrite -{k2}E in S2 *. +move/(IH (path_sorted S1) _ (path_sorted S2) _ S). +by rewrite /disj /supp /= => ->. +Qed. + +Lemma mapf_disj s1 s2 s1' s2' : + mapf s1 = mapf s2 -> mapf s1' = mapf s2' -> + disj s1 s1' = disj s2 s2'. +Proof. +move=>X1 X2; apply: eq_trans (mapf_disjL _ X1) _. +by rewrite !(disjC s2); apply: mapf_disjL X2. +Qed. + +End Map. + +Section FoldFMap. +Variables (A B: ordType) (V C: Type). + +Definition foldfmap g (e: C) (s: finMap A V) := + foldr g e (seq_of s). + + +Lemma foldf_nil g e : foldfmap g e (@nil A V) = e. +Proof. by rewrite /foldfmap //=. Qed. + +Lemma foldf_ins g e k v f: + path ord k (supp f) -> + foldfmap g e (ins k v f) = g (k, v) (foldfmap g e f). +Proof. by move=> H; rewrite /foldfmap //= seqof_ins //. Qed. +End FoldFMap. + +Section KeyMap. + +Section MapDef. +Variables (A B: ordType) (V : Type). + +Variable (f: A -> B). +Hypothesis Hf : forall x y, strictly_increasing f x y. + +Definition mapk (m : finMap A V) : finMap B V := + foldfmap (fun p s => ins (f (key p)) (value p) s) (nil B V) m. + +(* mapK preserves sorted *) + +Lemma sorted_mapk m: + sorted ord (supp (mapk m)). +Proof. case: (mapk m)=>[s]I //=. Qed. + + +Lemma path_mapk m k: path ord k (supp m) -> path ord (f k) (supp (mapk m)). +Proof. +elim/fmap_ind': m k =>// k1 v1 s P IH k. +rewrite {1}/supp //= {1}seqof_ins // /= => /andP [H]; move/IH=>H1. +by rewrite /mapk foldf_ins // /supp /= seqof_ins //= H1 andbT (Hf H). +Qed. + +Lemma mapk_nil : mapk (nil A V) = nil B V. +Proof. by rewrite /mapk //=. Qed. + + +Lemma mapk_ins k v s : + path ord k (supp s) -> + mapk (ins k v s) = ins (f k) v (mapk s). +Proof. by move=> H; rewrite /mapk foldf_ins =>//. Qed. +End MapDef. +Arguments mapk {A B V} f m. + +Variables (A B C : ordType)(V : Type)(f : A -> B) (g : B -> C). +Hypothesis Hf : forall x y, strictly_increasing f x y. + +Lemma mapk_id m : @mapk A A V id m = m. +Proof. +by elim/fmap_ind': m=>// k v s L IH; rewrite -{2}IH /mapk foldf_ins //. +Qed. + +Lemma mapk_comp m: + mapk g (@mapk A B V f m) = mapk (comp g f) m. +Proof. +elim/fmap_ind': m =>//= k v s P IH. +rewrite [mapk (g \o f) _]mapk_ins //. +rewrite mapk_ins // mapk_ins //; first by rewrite IH. +exact: (path_mapk Hf P). +Qed. +End KeyMap. +Arguments mapk {A B V} f m. + +(* zipping two finite maps with equal domains and ranges *) +(* zip_p predicate is a test for when contents of the two maps *) +(* at a given key are "combinable". typically zip_p will test *) +(* that the types of contents are equal, in the case the map is *) +(* storing dynamics *) + +Section Zip. +Variables (K : ordType) (V : Type) (zip_f : V -> V -> option 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 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 + [::], [::] => Some [::] + | (k1, v1)::s1', (k2, v2)::s2' => + if k1 == k2 then + if zip_f v1 v2 is Some v then + if zip' s1' s2' is Some s' then Some ((k1, v) :: s') + else None + else None + else None + | _, _ => None + end. + +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] //=; rewrite eq_sym; case: eqP=>// ->{k2}. +by rewrite comm IH. +Qed. + +Lemma zipA' s1 s2 s3 : + obind (zip' s1) (zip' s2 s3) = obind (zip'^~ s3) (zip' s1 s2). +Proof. +elim: s1 s2 s3=>[|[k1 v1] s1 IH]. +- case=>[|[k2 v2] s2]; case=>[|[k3 v3] s3] //=; case: eqP=>// ->{k2}. + by case: (zip_f v2 v3)=>// v; case: (zip' s2 s3). +case=>[|[k2 v2] s2]; case=>[|[k3 v3] s3] //=. +- by case: eqP=>// ->{k1}; case: (zip_f v1 v2)=>// v; case: (zip' s1 s2). +case: (k2 =P k3)=>[->{k2}|E1] /=; last first. +- case: (k1 =P k2)=>E2 //=. + case: (zip_f v1 v2)=>// v. + case: (zip' s1 s2)=>//= s. + by rewrite E2; case: eqP E1. +case: (k1 =P k3)=>[->{k1}|E1] /=; last first. +- by case: (zip_f v2 v3)=>// v; case: (zip' s2 s3)=>//= s; case: eqP E1. +case E1: (zip_f v2 v3)=>[w1|]; last first. +- case E3: (zip_f v1 v2)=>[w3|] //. + case S3: (zip' s1 s2)=>[t3|] //=; rewrite eq_refl. + by move: (assoc v1 v2 v3); rewrite /obind/oapp E1 E3=><-. +case S1: (zip' s2 s3)=>[t1|] /=; last first. +- case E3: (zip_f v1 v2)=>[w3|//]. + case S3: (zip' s1 s2)=>[t3|] //=; rewrite eq_refl. + move: (IH s2 s3); rewrite /obind/oapp S1 S3=><-. + by case: (zip_f w3 v3). +rewrite eq_refl. +case E3: (zip_f v1 v2)=>[w3|]; +move: (assoc v1 v2 v3); rewrite /obind/oapp E1 E3=>-> //=. +case S3: (zip' s1 s2)=>[t3|]; +move: (IH s2 s3); rewrite /obind/oapp S3 S1=>-> /=. +- by rewrite eq_refl. +by case: (zip_f w3 v3). +Qed. + +Lemma zip_unitL' s : zip' (zip_unit' s) s = Some s. +Proof. by elim: s=>[|[k v] s IH] //=; rewrite eq_refl unitL IH. Qed. + +Lemma map_key_zip' s1 s2 s : + zip' s1 s2 = Some s -> map key s = map key s1. +Proof. +elim: s1 s2 s=>[|[k1 v1] s1 IH]; case=>[|[k2 v2] s2] //= s; first by case=><-. +case: eqP=>// ->{k1}; case: (zip_f v1 v2)=>// w; case Z: (zip' s1 s2)=>[t|//]. +by case=><-; rewrite -(IH _ _ Z). +Qed. + +Lemma zip_unitE' s1 s2 : + (exists s, zip' s1 s2 = Some s) <-> zip_unit' s1 = zip_unit' s2. +Proof. +split. +- case; elim: s1 s2 =>[|[k1 v1] s1 IH]; case=>// [[k2 v2] s2] s //=. + case: eqP=>// <-{k2}. + case E1: (zip_f v1 v2)=>[w|//]; case E2: (zip' s1 s2)=>[t|//] _. + by move/IH: E2=>->; congr ((_, _)::_); apply/unitE; exists w. +elim: s1 s2=>[|[k1 v1] s1 IH]; case=>//; first by exists [::]. +case=>k2 v2 s2 //= [<-{k2}]; case/unitE=>s ->; case/IH=>t ->. +by exists ((k1, s)::t); rewrite eq_refl. +Qed. + +Lemma zip_sorted' s1 s2 : + sorted ord (map key s1) -> + forall s, zip' s1 s2 = Some s -> sorted ord (map key s). +Proof. by move=>H s; move/map_key_zip'=>->. Qed. + +Definition zip f1 f2 : option (finMap K V) := + match f1, f2 with + FinMap s1 pf1, FinMap s2 pf2 => + match zip' s1 s2 as z return zip' s1 s2 = z -> _ with + Some s => fun pf => Some (FinMap (zip_sorted' pf1 pf)) + | None => fun pf => None + end (erefl _) + end. + +Lemma zip_unit_sorted' s : + sorted ord (map key s) -> + sorted ord (map key (zip_unit' s)). +Proof. +rewrite (_ : map key s = map key (zip_unit' s)) //. +by apply: (map_key_zip' (s2:=s)); apply: zip_unitL'. +Qed. + +Definition zip_unit f := + let: FinMap s pf := f in FinMap (zip_unit_sorted' pf). + +Lemma zip_unitE f1 f2 : + (exists f, zip f1 f2 = Some f) <-> zip_unit f1 = zip_unit f2. +Proof. +case: f1 f2=>s1 H1 [s2 H2] /=; split. +- case=>s; move: (zip_sorted' _). + case E: (zip' s1 s2)=>[t|//] _ _; apply/fmapE=>/=. + by apply/zip_unitE'; exists t. +case; case/zip_unitE'=>s E; move/(zip_sorted' H1): (E)=>T. +exists (FinMap T); move: (zip_sorted' _); rewrite E=>pf. +by congr Some; apply/fmapE. +Qed. + +(* Now lemmas for interaction of zip with the other primitives *) + +Lemma zip_supp' s1 s2 s : zip' s1 s2 = Some s -> map key s = map key s1. +Proof. +elim: s1 s2 s=>[|[k1 v1] s1 IH] /=; first by case=>// s [<-]. +case=>[|[k2 v2] s2] // s; case: eqP=>// ->{k1}. +case E: (zip_f v1 v2)=>[w|//]; case F: (zip' s1 s2)=>[t|//]. +by move/IH: F=><- [<-]. +Qed. + +Lemma zip_supp f1 f2 f : + zip f1 f2 = Some f -> supp f =i supp f1. +Proof. +case: f1 f2 f=>s1 H1 [s2 H2] [s3 H3] /=; move: (zip_sorted' _). +case E: (zip' s1 s2)=>[t|//]; move=>pf [F x]; rewrite -{s3}F in H3 *. +by rewrite /supp (zip_supp' E). +Qed. + +Lemma zip_filter' s1 s2 s x : + zip' s1 s2 = Some s -> + zip' (filter (predCk V x) s1) (filter (predCk V x) s2) = + Some (filter (predCk V x) s). +Proof. +elim: s1 s2 s=>[|[k1 v1] s1 IH]; case=>[|[k2 v2] s2] //= s; first by case=><-. +case: eqP=>// <-{k2}; case E1: (zip_f v1 v2)=>[a|//]. +case E2: (zip' s1 s2)=>[t|//]; move/IH: E2=>E2 [<-{s}]. +case: eqP=>[->{k1}|] /=; first by rewrite E2 eq_refl. +by rewrite eq_refl E1 E2; case: eqP. +Qed. + +Lemma zip_rem f1 f2 f x : + zip f1 f2 = Some f -> zip (rem x f1) (rem x f2) = Some (rem x f). +Proof. +case: f1 f2 f=>s1 H1 [s2 H2] [s3 H3] /=; do 2![move: (zip_sorted' _)]. +case E1: (zip' s1 s2)=>[t|//]; case E2 : (zip' _ _)=>[q|]; +rewrite (@zip_filter' _ _ _ _ E1) // in E2. +by case: E2=><-{q} pf1 pf2 [E]; congr Some; apply/fmapE; rewrite /= E. +Qed. + +Lemma zip_fnd f1 f2 f x (v : V) : + zip f1 f2 = Some f -> fnd x f = Some v -> + exists v1 v2, + [/\ zip_f v1 v2 = Some v, fnd x f1 = Some v1 & fnd x f2 = Some v2]. +Proof. +case: f1 f2 f=>s1 H1 [s2 H2] [s3 H3] /=; move: (zip_sorted' _). +case E1: (zip' s1 s2)=>[s|//] pf [E]; rewrite /fnd /=. +clear H1 H2 H3 pf; rewrite -{s3}E. +elim: s1 s2 s E1=>[|[k1 v1] s1 IH]; case=>[|[k2 v2] s2] //= s. +- by case=><-. +case: eqP=>// <-{k2}; case E1: (zip_f v1 v2)=>[w|//]. +case E2: (zip' s1 s2)=>[t|//][<-{s}] /=. +case: eqP=>[_ [<-]|_]. + (* DEVCOMMENT: exists v1, v2 errors! *) +- by exists v1; exists v2. +by case: (filter (predk V x) t) (IH _ _ E2). +Qed. + +(* The other direction of the zip_fnd lemma *) + +Lemma fnd_zip f1 f2 f x (v1 v2 : V) : + fnd x f1 = Some v1 -> fnd x f2 = Some v2 -> + zip f1 f2 = Some f -> fnd x f = zip_f v1 v2. +Proof. +case: f1 f2=>s1 H1 [s2 H2] /=; move: (zip_sorted' _). +case E: (zip' s1 s2)=>[s|//] pf; rewrite /fnd /= => F1 F2. +case=><-{f} /= {pf H1 H2}. +elim: s1 s2 s E F1 F2=>[|[k1 w1] s1 IH]; case=>[|[k2 w2] s2] //= s. +case: eqP=>// <-{k2}; case E1: (zip_f w1 w2)=>[w|//]. +case E2: (zip' s1 s2)=>[t|//] [<-{s}]. +case: eqP=>/=; last by case: eqP=>// _ _; apply: IH. +by move=>->{k1}; rewrite eq_refl; case=><- [<-]. +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 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: ordP=>//= O. +by rewrite IH // (path_sorted H). +Qed. + +Lemma zunit_fcat f1 f2 : + zip_unit (fcat f1 f2) = fcat (zip_unit f1) (zip_unit f2). +Proof. +elim/fmap_ind': f2 f1=>[|k v f2 H IH] f1 /=. +- rewrite fcats0; set j := FinMap _. + by rewrite (_ : j = nil K V) ?fcat0s //; apply/fmapE. +by rewrite fcat_sins zunit_ins IH -fcat_sins zunit_ins. +Qed. + +Lemma zunit_supp f : supp (zip_unit f) = supp f. +Proof. +case: f=>s H; rewrite /supp /= {H}. +by elim: s=>[|[k v] s IH] //=; rewrite IH. +Qed. + +Lemma zunit_disj f1 f2 : disj f1 f2 = disj (zip_unit f1) (zip_unit f2). +Proof. +case: disjP; case: disjP=>//; rewrite !zunit_supp. +- by move=>x H1 H2; move/(_ _ H1); rewrite H2. +by move=>H x; move/H/negbTE=>->. +Qed. + +End Zip. diff --git a/core/options.v b/core/options.v new file mode 100644 index 0000000..6269953 --- /dev/null +++ b/core/options.v @@ -0,0 +1,3 @@ +#[export] Set Implicit Arguments. +#[export] Unset Strict Implicit. +#[export] Unset Printing Implicit Defensive. diff --git a/core/ordtype.v b/core/ordtype.v new file mode 100644 index 0000000..13072a1 --- /dev/null +++ b/core/ordtype.v @@ -0,0 +1,473 @@ +(* +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 HB Require Import structures. +From Coq Require Import ssreflect ssrbool ssrfun. +From mathcomp Require Import ssrnat eqtype seq path fintype. +From pcm Require Import options. + +Definition connected (T : eqType) (ord : rel T) := + forall x y, x != y -> ord x y || ord y x. + +Definition ordtype_axiom (T : eqType) (ord : rel T) := + [/\ irreflexive ord, transitive ord & connected ord]. + +HB.mixin Record isOrdered T of Equality T := { + ord : rel T; + ordtype_subproof : ordtype_axiom ord}. + +(* Ideally, we would just generate the structure thus. *) +(* But, we need to insert some Universe Polymorphism declaration *) +(* so we switch to plan B: print the generated code using #[log] *) +(* command, then cut-and-paste it and decorate by hand. *) +(* Unfortunately, this doesn't quite work, HB manuals notwithstanding, *) +(* because custom declarations of log code impoverish the Elpi state *) +(* making subsequent invocations of HB work only partially. *) +(* In this file in particular, the canonicals for each instance *) +(* must be declared by hand as HB.instance does that *) +(* only if Ordered is declared by HB.structure *) +(* +#[log(raw)] +#[short(type="ordType")] +HB.structure Definition Ordered := + {T of Equality T & isOrdered T}. +*) + +(* cut-and-pasted code starts *) +Module Ordered. +Section axioms_. +Local Unset Implicit Arguments. + +Record axioms_ (T : Type) : Type := Class { + eqtype_hasDecEq_mixin : Equality.mixin_of T; + ordtype_isOrdered_mixin : isOrdered.axioms_ T eqtype_hasDecEq_mixin}. +End axioms_. + +Global Arguments axioms_ : clear implicits. +Global Arguments Class [_] [_] _. +Global Arguments eqtype_hasDecEq_mixin [_] _. +Global Arguments ordtype_isOrdered_mixin [_] _. + +Section type. +Local Unset Implicit Arguments. +Polymorphic Cumulative Record + type : Type := Pack { sort : Type; _ : axioms_ sort; }. +End type. + +Definition class (cT : type) := + let: Pack _ c as cT' := cT return axioms_ (sort cT') in c. + +Global Arguments type : clear implicits. +Global Arguments Pack [_] _. +Global Arguments sort : clear implicits. +Global Arguments class : clear implicits. + +(* DEVCOMMENT *) +(* The polymorphism annotations here and below are needed for storing *) +(* ordType instances in finMaps which have an ordType constraint of *) +(* their own. An example of this is KVMap from HTT. *) +(* Ultimately there should be a better solution if we want to switch *) +(* to Mathcomp's orders. *) +(* /DEVCOMMENT *) + +(* Polymorphic Universe annotations added *) +Section PolymorphicClonePack. +Polymorphic Universe ou. +Polymorphic Variables (T : Type@{ou}) (cT : type@{ou}). + +Polymorphic Definition phant_clone : forall (c : axioms_ T), + unify Type Type T (sort cT) nomsg -> + unify type type cT (Pack (sort:=T) c) nomsg -> type := + fun (c : axioms_ T) => + fun=> (fun=> (Pack (sort:=T) c)). + +Polymorphic Definition pack_ := fun (m : Equality.mixin_of T) + (m0 : isOrdered.axioms_ T m) => + Pack (sort:=T) + {|eqtype_hasDecEq_mixin := m; ordtype_isOrdered_mixin := m0 |}. + +End PolymorphicClonePack. + +Local Arguments phant_clone : clear implicits. +Notation clone X2 X1 := (phant_clone X2 X1 _ id_phant id_phant). +Local Arguments pack_ : clear implicits. + +Module Exports. +Notation ordType := Ordered.type. +#[reversible] Coercion sort : Ordered.type >-> Sortclass. + +(* Polymorphic annotation added *) +Polymorphic Definition ordtype_Ordered_class__to__eqtype_Equality_class : + forall T : Type, axioms_ T -> Equality.axioms_ T := + fun (T : Type) (c : axioms_ T) => + {| Equality.eqtype_hasDecEq_mixin := eqtype_hasDecEq_mixin c |}. + +Local Arguments ordtype_Ordered_class__to__eqtype_Equality_class : + clear implicits. + +#[reversible] Coercion ordtype_Ordered_class__to__eqtype_Equality_class : + ordtype.Ordered.axioms_ >-> eqtype.Equality.axioms_. + +Polymorphic Definition ordtype_Ordered__to__eqtype_Equality : + ordType -> eqType := + fun s : ordType => {| Equality.sort := s; Equality.class := class s |}. + +Local Arguments ordtype_Ordered__to__eqtype_Equality : + clear implicits. + +#[reversible] Coercion ordtype_Ordered__to__eqtype_Equality : + ordtype.Ordered.type >-> eqtype.Equality.type. + +Global Canonical ordtype_Ordered__to__eqtype_Equality. + +#[reversible] Coercion eqtype_hasDecEq_mixin : + ordtype.Ordered.axioms_ >-> eqtype.hasDecEq.axioms_. + +#[reversible] Coercion ordtype_isOrdered_mixin : + ordtype.Ordered.axioms_ >-> ordtype.isOrdered.axioms_. + +End Exports. +Import Exports. + +Definition phant_on_ : forall T : ordType, phant T -> axioms_ T := + fun T : ordType => fun=> class T. +Local Arguments phant_on_ : clear implicits. + +Notation on_ X1 := ( phant_on_ _ (Phant X1)). +Notation copy X2 X1 := ( phant_on_ _ (Phant X1) : axioms_ X2). +Notation on X1 := ( phant_on_ _ (Phant _) : axioms_ X1). + +Module EtaAndMixinExports. +Section hb_instance_91. +Variable T : ordType. +Local Arguments T : clear implicits. + +Definition HB_unnamed_factory_92 : axioms_ (eta T) := class T. +Local Arguments HB_unnamed_factory_92 : clear implicits. +Definition ordtype_Ordered__to__eqtype_hasDecEq : + Equality.mixin_of (eta T) := Equality.class T. +Local Arguments ordtype_Ordered__to__eqtype_hasDecEq : + clear implicits. + +Definition ordtype_Ordered__to__ordtype_isOrdered : + isOrdered.axioms_ (eta T) HB_unnamed_factory_92 := + ordtype_isOrdered_mixin HB_unnamed_factory_92. +Local Arguments ordtype_Ordered__to__ordtype_isOrdered : + clear implicits. + +Definition HB_unnamed_mixin_95 := + ordtype_isOrdered_mixin HB_unnamed_factory_92. +Local Arguments HB_unnamed_mixin_95 : clear implicits. + +Definition structures_eta__canonical__ordtype_Ordered : ordType := + Pack (sort := eta T) + {| + eqtype_hasDecEq_mixin := Equality.class T; + ordtype_isOrdered_mixin := HB_unnamed_mixin_95 + |}. + +Local Arguments structures_eta__canonical__ordtype_Ordered : + clear implicits. +Global Canonical structures_eta__canonical__ordtype_Ordered. +End hb_instance_91. +End EtaAndMixinExports. +End Ordered. + +HB.export Ordered.Exports. +HB.export Ordered.EtaAndMixinExports. + +Definition ord : forall s : ordType, rel s := + fun s : ordType => isOrdered.ord (Ordered.class s). +Local Arguments ord : clear implicits. +Global Arguments ord {_}. + +Definition ordtype_subproof : forall s : ordType, ordtype_axiom ord := + fun s : ordType => isOrdered.ordtype_subproof (Ordered.class s). +Local Arguments ordtype_subproof : clear implicits. +Global Arguments ordtype_subproof {_}. + +Notation Ordered X1 := (Ordered.axioms_ X1). +(* end of generated and changed code *) + + +(* Repacking *) + +Lemma irr {T : ordType} : irreflexive (@ord T). +Proof. by case: (@ordtype_subproof T). Qed. + +Lemma trans {T : ordType} : transitive (@ord T). +Proof. by case: (@ordtype_subproof T). Qed. + +Lemma connex {T : ordType} : connected (@ord T). +Proof. by case: (@ordtype_subproof T). Qed. + +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 ord_total x y : [|| ord x y, x == y | ord y x]. +Proof. +case E: (x == y)=>/=; first by rewrite orbT. +by apply: connex; rewrite negbT. +Qed. + +Lemma nsym x y : ord x y -> ~ ord y x. +Proof. by move=>E1 E2; move: (trans E1 E2); rewrite irr. Qed. + +Lemma antisym : antisymmetric (@ord T). +Proof. by move=>x y /andP [H] /(nsym H). 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. 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_path tr (filter_subseq p s)). Qed. + +Lemma ord_sorted_eq (s1 s2 : seq T) : + sorted ord s1 -> sorted ord s2 -> s1 =i s2 -> s1 = s2. +Proof. exact/irr_sorted_eq/irr/(@trans _). Qed. + +Lemma oleq_ord_trans (n m p : T) : + oleq m n -> ord n p -> ord m p. +Proof. by case/orP; [apply: trans | move/eqP=>->]. Qed. + +Lemma ord_oleq_trans (n m p : T) : + ord m n -> oleq n p -> ord m p. +Proof. by move=>H /orP [|/eqP <- //]; apply: trans. Qed. + +End Lemmas. + +Arguments orefl {T}. +Arguments otrans {T}. + +#[export] Hint Resolve orefl irr trans connex +otrans oantisym oleq_ord_trans : 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 ordType *) +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. + +Lemma oleq_eqVord x y : oleq x y = (x == y) || ord x y. +Proof. by rewrite /oleq orbC. 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. + + +(**********************************************) +(* Building hierarchies for some basic orders *) +(**********************************************) + +(* trivial total order for unit type *) +Definition unit_ord (x y : unit) := false. +Lemma unit_is_ordtype : ordtype_axiom unit_ord. Proof. by []. Qed. +HB.instance Definition unit_ord_mix := isOrdered.Build unit unit_is_ordtype. +(* manual canonical declaration, as HB fails to declare it *) +Canonical unit_ordType : ordType := + Ordered.Pack (sort:=unit) (Ordered.Class unit_ord_mix). + +(* trivial total order for nat *) +Lemma nat_is_ordtype : ordtype_axiom ltn. +Proof. by split=>[x||x y] /=; [rewrite ltnn|apply: ltn_trans|case: ltngtP]. Qed. +HB.instance Definition nat_ord_mix := isOrdered.Build nat nat_is_ordtype. +Canonical nat_ordType : ordType := + Ordered.Pack (sort:=nat) (Ordered.Class nat_ord_mix). + +Lemma nat_ord : ord =2 ltn. Proof. by []. Qed. +Lemma nat_oleq : oleq =2 leq. +Proof. by move=>x y; rewrite /oleq nat_ord /=; case: ltngtP. Qed. + +(* product order *) +Section ProdOrdType. +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 prod_is_ordtype : ordtype_axiom lex. +Proof. +split=>[x|[x1 x2][y1 y2][z1 z2]|[x1 x2][y1 y2]]; rewrite /lex /=. +- by rewrite eq_refl irr. +- 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). +by rewrite -pair_eqE /= -(eq_sym x1 y1); case: ifPn=>[_|] /connex. +Qed. + +HB.instance Definition prod_ord_mix := + isOrdered.Build (K * T)%type prod_is_ordtype. +(* manual canonical declaration, as HB fails to declare it *) +Canonical prod_ordType : ordType := + Ordered.Pack (sort:=prod K T) (Ordered.Class prod_ord_mix). +End ProdOrdType. + + +(* Every fintype is ordered *) +Section FinTypeOrd. +Variable T : finType. + +Definition ordf : rel T := + fun x y => index x (enum T) < index y (enum T). + +Lemma fintype_is_ordtype : ordtype_axiom ordf. +Proof. +split=>[x|x y z|x y]; rewrite /ordf /=. +- by rewrite ltnn. +- by apply: ltn_trans. +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. + +(* There isn't canonical projection to latch onto *) +(* so we can't have generic inheritance. *) +(* +HB.instance Definition _ := + isOrdered.Build T%type fintype_is_ordtype. +*) + +End FinTypeOrd. + +(* However, we can get ordtype for any individual fintype *) +(* e.g. ordinals 'I_n *) +HB.instance Definition ordinal_ord_mix n := + isOrdered.Build 'I_n (fintype_is_ordtype 'I_n). +(* manual canonical declaration, as HB fails to declare it *) +Canonical ordinal_ordType n : ordType := + Ordered.Pack (sort:='I_n) (Ordered.Class (ordinal_ord_mix n)). + +Section SeqOrdType. +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 seq_is_ordtype : ordtype_axiom ords. +Proof. +split. +- by elim=>[|x xs IH] //=; rewrite eq_refl. +- elim=>[|y ys /= IH][|x xs][|z zs] //=. + case: (eqVneq x y)=>[->{x}|Hxy]; first by case: ifP=>// _; apply: IH. + case: (eqVneq y z)=>[<-{z}|Hyz]; first by rewrite (negbTE Hxy). + by case: (x =P z)=>[<-{Hyz z} H /(nsym H)//|_]; apply: trans. +elim=>[|x xs IH][|y ys] //=; rewrite (eq_sym y). +case: (x =P y)=>[-> H|/eqP/connex //]. +by apply: IH; apply: contra H=>/eqP ->. +Qed. + +HB.instance Definition seq_ord_mix := isOrdered.Build (seq T) seq_is_ordtype. +(* manual canonical declaration, as HB fails to declare it *) +Canonical seq_ordType : ordType := + Ordered.Pack (sort:=seq T) (Ordered.Class seq_ord_mix). +End SeqOrdType. + +#[deprecated(since="fcsl-pcm 1.4.0", note="Use ord_sorted_eq instead.")] +Notation eq_sorted_ord := ord_sorted_eq (only parsing). + diff --git a/core/pred.v b/core/pred.v new file mode 100644 index 0000000..0d458b0 --- /dev/null +++ b/core/pred.v @@ -0,0 +1,1636 @@ +(* +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 Basics. +From mathcomp Require Import ssrnat seq eqtype. +From pcm Require Import options. + +(* First some basic propositional equalities *) + +(* DEVCOMMENT *) +(* Basically, we need to repeat most of ssrbool.v here but we'll do it as we go. *) +(* /DEVCOMMENT *) + +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 name, format "[ 'Pred' x | E ]") : fun_scope. +Notation "[ 'Pred' x : T | E ]" := (SimplPred (fun x : T => E)) + (at level 0, x name, only parsing) : fun_scope. +Notation "[ 'Pred' x y | E ]" := (SimplPred (fun t => let: (x, y) := t in E)) + (at level 0, x name, y name, 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 name, y name, only parsing) : fun_scope. + +Definition clone_Pred T U := + fun pT & @Pred_Sort T pT -> U => + fun toP (pT' := @PropPredType T U toP) & phant_id pT' pT => pT'. +Notation "[ 'PredType' 'of' T ]" := (@clone_Pred _ T _ id _ id) : form_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. + +(* DEVCOMMENT *) +(* are these needed? *) +(* Definition EqPred T (p1 p2 : Pred T) := EqPredType p1 p2. *) +(* Definition SubPred T (p1 p2 : Pred T) := SubPredType p1 p2. *) +(* +Definition EqMem T (p1 p2 : Mem_Pred T) := EqPredType p1 p2. +Definition SubMem T (p1 p2 : Mem_Pred T) := SubPredType p1 p2. +*) + +(* /DEVCOMMENT *) + +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 name, format "[ 'Pred' x \In A ]") : fun_scope. +Notation "[ 'Pred' x \In A | E ]" := [Pred x | (x \In A) /\ E] + (at level 0, x name, 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 name, y name, + 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 name, y name, + 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 name, y name, + 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 name, y name, + 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. + +(* DEVCOMMENT *) +(* is this needed? *) +(* Definition InE := (In_Simpl, Simpl_PredE). (* to be extended *) *) +(* /DEVCOMMENT *) + +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. + +(* We also add subrelation and relation equality *) +(* for non-collective binary relations *) + +(* Not a definition to avoid universe inconsistencies. *) +(* DEVCOMMENT *) +(* Universe Polymorphism does not quite work yet *) +(* /DEVCOMMENT *) +Local Notation Rel A := (A -> A -> Prop). + +Definition eqrel {A} (R1 R2 : Rel A) := forall x y, R1 x y <-> R2 x y. +Definition subrel {A} (R1 R2 : Rel A) := forall x y, R1 x y -> R2 x y. + +Notation "A <~2> B" := (eqrel A B) + (at level 70, no associativity) : rel_scope. +Notation "A ~2> B" := (subrel A B) + (at level 70, no associativity) : rel_scope. + +Section TransferLemmas. +Variables (A : Type) (R1 R2 : Rel A). +Lemma eqrelI : (forall x y, R1 x y <-> R2 x y) -> R1 <~2> R2. Proof. by []. Qed. +Lemma eqrelE : R1 <~2> R2 -> forall x y, R1 x y <-> R2 x y. Proof. by []. Qed. +Lemma subrelI : (forall x y, R1 x y -> R2 x y) -> R1 ~2> R2. Proof. by []. Qed. +Lemma subrelE : R1 ~2> R2 -> forall x y, R1 x y -> R2 x y. Proof. by []. Qed. +End TransferLemmas. + +(* composing relation and function *) + +Definition Rel_app A B (D : Rel A) (f : B -> A) : Rel B := + fun x y => D (f x) (f y). + +Arguments Rel_app {A B} D f _ _ /. +Notation "D \\o f" := (@Rel_app _ _ D f) + (at level 42, left associativity). + +Lemma rel_app1 T (g : T -> T -> Prop) : g \\o id <~2> g. Proof. by []. Qed. + +(* the same for decidable relations *) + +Definition rel_app U V (S : rel V) f : rel U := + fun x y => S (f x) (f y). +Arguments rel_app {U V} S f _ _ /. + +(**************************************) +(* Definitions and lemmas for setoids *) +(**************************************) + +(* Declaration of relations *) + +Section EqPredType. +Variables (T : Type) (pT : PredType T). +Lemma EqPredType_refl (r : pT) : EqPredType 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. +Definition EqPredType_trans r2 r1 r3 := @EqPredType_trans' r1 r2 r3. +End EqPredType. + +#[export] Hint Resolve EqPredType_refl : core. + +(* declare relations for all implicit coercions and canonical structures *) + +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 : (Mem_Pred T) (@EqMem T) + reflexivity proved by (@EqPredType_refl _ _) + symmetry proved by (@EqPredType_sym _ _) + transitivity proved by (@EqPredType_trans' _ _) as EqMem_rel. + +(* DEVCOMMENT *) +(* are these needed? *) +(* +Add Parametric Relation T : (Pred T) (@EqPred _) + reflexivity proved by (@EqPredType_refl _ _) + symmetry proved by (@EqPredType_sym _ _) + transitivity proved by (@EqPredType_trans' _ _) as EqPred_rel. +*) + +Section SubPredType. +Variables (T : Type) (pT : PredType T). +Lemma SubPredType_refl (r : pT) : SubPredType r r. Proof. by []. 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 SubPredType_trans r2 r1 r3 := @SubPredType_trans' r1 r2 r3. +End SubPredType. + +#[export] Hint Resolve SubPredType_refl : core. + +Add Parametric Relation T (pT : PredType T) : pT (@SubPredType _ pT) + reflexivity proved by (@SubPredType_refl _ _) + transitivity proved by (@SubPredType_trans' _ _) as SubPredType_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) (@SubMem _) + reflexivity proved by (@SubPredType_refl _ _) + transitivity proved by (@SubPredType_trans' _ _) as SubMem_rel. + +(* DEVCOMMENT *) +(* are these needed? *) +(* +Add Parametric Relation T : (Pred T) (@SubPred _) + reflexivity proved by (@SubPredType_refl _ _) + transitivity proved by (@SubPredType_trans' _ _) as SubPred_rel. +*) +(* /DEVCOMMENT *) + +Lemma eqrel_refl {A R} : @eqrel A R R. Proof. by []. Qed. +Lemma eqrel_sym {A R1 R2} : @eqrel A R1 R2 -> eqrel R2 R1. +Proof. by move=>H x y; rewrite H. Qed. +Lemma eqrel_trans {A R1 R2 R3} : + @eqrel A R1 R2 -> eqrel R2 R3 -> eqrel R1 R3. +Proof. by move=>H1 H2 x y; rewrite H1 H2. Qed. + +#[export] Hint Resolve eqrel_refl : core. + +Add Parametric Relation (A : Type) : (Rel A) (@eqrel A) + reflexivity proved by (@eqrel_refl A) + symmetry proved by (@eqrel_sym A) + transitivity proved by (@eqrel_trans A) as eqrel_eq_rel. + +Lemma subrel_refl {A} R : @subrel A R R. Proof. by []. Qed. +Lemma subrel_trans A R1 R2 R3 : + @subrel A R1 R2 -> subrel R2 R3 -> subrel R1 R3. +Proof. by move=>H1 H2 x y /H1/H2. Qed. + +#[export] Hint Resolve subrel_refl : core. + +Add Parametric Relation (A : Type) : (Rel A) (@subrel A) + reflexivity proved by (@subrel_refl A) + transitivity proved by (@subrel_trans A) as subrel_rel. + +(* Declaring morphisms. *) + +(* DEVCOMMENT *) +(* Annoyingly, even the coercions must be declared *) +(* /DEVCOMMENT *) + +Add Parametric Morphism T : (@Pred_of_Simpl T) with signature + @EqSimplPred _ ==> @EqPredType T (PredPredType T) as Pred_of_Simpl_morph. +Proof. by []. Qed. + +(* DEVCOMMENT *) +(* Do we need other coercions? We'll discover as we go *) + +(* Now the other morphisms. Again, not clear which ones are needed. *) +(* However, for all this to work, it seems that morphisms must be *) +(* declared with most specific signatures, or else the system *) +(* complains. For example, we use EqPred _ instead of EqPredType _ _, *) +(* even though the former is an instance of the later. *) + +(* +Add Parametric Morphism T : (@EqPred T) with signature + @EqPred _ ==> @EqPred _ ==> iff as EqPred_morph. +Proof. by move=>r1 s1 H1 r2 s2 H2; rewrite H1 H2. Qed. +*) +(* /DEVCOMMENT *) + +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. + +Add Parametric Morphism A : (@subrel A) with signature + eqrel ==> eqrel ==> iff as subrel_eq. +Proof. by move=>x y H x1 y1 H1; split=>X a b /H/X/H1. Qed. + +Add Parametric Morphism A B : (@Rel_app A B) with signature + eqrel ==> eq ==> eqrel as Relapp_eq. +Proof. by move=>x y H f s s'; split=>/H. Qed. + +Add Parametric Morphism A B : (@Rel_app A B) with signature + subrel ==> eq ==> subrel as Relapp_sub. +Proof. by move=>x y H f s s' /H. Qed. + +(* Can we declare relation application as morphism? *) +(* We can, but it isn't picked up in practice. *) +(* This is so because application is a special term *) +(* that apprently can't be matched on. *) +(* If we want to rewrite under application *) +(* the choices are: *) +(* a) don't use setoids, but prove equalites *) +(* b) move to collective predicates where *) +(* function application is user-level term *) +(* +Add Parametric Morphism A : (fun (R : Rel A) (x y : A) => R x y) + with signature eqrel ==> eq ==> eq ==> iff as app_eq. +Proof. by move=>r1 r2 E x y; split=>/E. 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. + +#[export] 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). + +(* DEVCOMMENT *) +(* I also wanted to add Simpl_PredE, but setoid rewrite returns an error *) +(* and instead of trying the other rules in the tuple, it just stops *) +(* This is ridiculuous *) +(* /DEVCOMMENT *) + +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. + +Prenex Implicits In_split. + +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. + +Lemma MapP T1 T2 (f : T1 -> T2) (s : seq T1) (y : T2) : + y \In map f s <-> exists2 x, x \In s & y = f x. +Proof. +elim: s => [|x s IHs] /=; first by split=>//; case. +rewrite In_cons; split. +- case=>[->|]; first by exists x=>//; apply/In_cons; left. + by case/IHs=>k H ->; exists k=>//; apply/In_cons; right. +case=>k /In_cons [->|H E]; first by left. +by right; apply/IHs; exists k. +Qed. + +Lemma Mem_filter (T : Type) (a : pred T) (x : T) (s : seq T) : + x \In filter a s <-> a x /\ x \In s. +Proof. +elim: s; first by split; case. +move=>y s IHs /=; rewrite 2!fun_if /=. +case E: (a y). +- rewrite InE IHs; split; last first. + - by case=>H1 /In_cons [->|]; [left | right]. + case=>[->|]; first by split=>//; apply/In_cons; left. + by case=>H1 H2; split=>//; apply/In_cons; right. +rewrite IHs; split. +- by case=>H1 H2; split=>//; apply/In_cons; right. +by case=>H1 /In_cons [] // ?; subst y; rewrite E in H1. +Qed. + +Lemma eq_In_filter (T : Type) a1 a2 (s : seq T) : + (forall x, x \In s -> a1 x = a2 x) -> + filter a1 s = filter a2 s. +Proof. +elim: s => //= x s IHs eq_a. +rewrite eq_a; last by rewrite InE; left. +rewrite IHs // => y s_y; apply: eq_a. +by rewrite InE; right. +Qed. + +(* 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. + +(* DEVCOMMENT *) +(* this interferes with the usage of inE, see + https://gitlab.software.imdea.org/mathador/fcsl/-/issues/97 *) +(* +(* 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. +*) +(* /DEVCOMMENT *) + +(* Big \In equivalences for all and has *) + +Section allhasP. +Context {T : Type}. +Variables (p : pred T). + +Lemma allPIn (s : seq T) : + reflect (forall x, x \In s -> p x) (all p s). +Proof. +elim: s=>[|x s IHs] /=; first by constructor. +case: andP=>[[H1 H2]|H]; constructor. +- by move=>z /In_cons [->|/(IHs H2)]. +move=>H1; elim: H; split; first by apply/H1/In_cons; left. +by apply/IHs=>z H; apply/H1/In_cons; right. +Qed. + +Lemma allPnIn (s : seq T) : + reflect (exists2 x, x \In s & ~~ p x) (~~ all p s). +Proof. +elim: s => [|x s IHs] /=; first by constructor=> [[x Hx _]]. +rewrite /= andbC negb_and; case: IHs => IHs /=. +- by constructor; case: IHs => y Hy Hay; exists y=>//; apply/In_cons; right. +apply: (iffP idP) => [|[y]]; first by exists x=>//; apply/In_cons; left. +by case/In_cons=>[->//|H1 H2]; elim: IHs; exists y. +Qed. + +Lemma hasPIn (s : seq T) : + reflect (exists2 x, x \In s & p x) (has p s). +Proof. +elim: s => [|y s IHs] /=; first by right; case. +case Py: (p y); first by left; exists y=>//; apply/In_cons; left. +apply: (iffP IHs)=>[] [x ysx Px]; exists x => //; first by apply/In_cons; right. +by case/In_cons: ysx Px Py=>[->->|]. +Qed. + +Lemma hasPnIn (s : seq T) : + reflect (forall x, x \In s -> ~~ p x) (~~ has p s). +Proof. +apply: (iffP (negPP (hasPIn s)))=>H. +- by move=>x Hx; apply: contra_notN H=>Px; exists x. +by case=>x Hx; apply/negP/H. +Qed. + +(* implication form of hasPIn is frequently used *) +(* and if you don't have it, causes fluff in proofs *) + +Lemma hasPInX x xs : x \In xs -> p x -> has p xs. +Proof. by move=>H1 H2; apply/hasPIn; exists x. Qed. + +End allhasP. + +Arguments hasPInX {T p x xs}. + +Section AllHasP. +Context {T : Type}. +Variables (P : Pred T). + +Fixpoint All xs := if xs is x :: xs then P x /\ All xs else True. + +Lemma AllP xs : All xs <-> forall x, x \In xs -> P x. +Proof. +elim: xs=>[|x xs IH] //=. +split; first by case=>H1 /IH H2 z; rewrite InE; case=>[->//|]; apply: H2. +move=>H; split; first by apply: H; rewrite InE; left. +by apply/IH=>z Z; apply: H; rewrite InE; right. +Qed. + +Lemma All_cat (s1 s2 : seq T) : + All (s1 ++ s2) <-> All s1 /\ All s2. +Proof. +split. +- by move/AllP=>H; split; apply/AllP=>x Hx; apply/H/Mem_cat; [left|right]. +by case=>/AllP H1 /AllP H2; apply/AllP=>x /Mem_cat; case=>Hx; [apply: H1| apply: H2]. +Qed. + +Fixpoint Has xs := if xs is x :: xs then P x \/ Has xs else False. + +Lemma HasP xs : Has xs <-> exists2 x, x \In xs & P x. +Proof. +elim: xs=>[|x xs IH] /=; first by split=>//; case. +split. +- case=>[H|/IH]; first by exists x=>//; rewrite InE; left. + by case=>y H1 H2; exists y=>//; rewrite InE; right. +case=>y; rewrite InE; case=>[->|H1 H2]; first by left. +by right; apply/IH; exists y. +Qed. + +Lemma Has_cat (s1 s2 : seq T) : + Has (s1 ++ s2) <-> Has s1 \/ Has s2. +Proof. +split. +- by move/HasP=>[x] /Mem_cat; case=>Hx Px; [left|right]; apply/HasP; exists x. +by case=>/HasP [x Hx Px]; apply/HasP; exists x=>//; apply/Mem_cat; [left|right]. +Qed. + +End AllHasP. + +Section AllHasT. +Context {T : Type}. +Variables (P : T -> Type). + +Fixpoint AllT xs : Type := if xs is x :: xs then P x * AllT xs else unit. + +Fixpoint HasT xs : Type := if xs is x :: xs then P x + HasT xs else Empty_set. + +End AllHasT. + +(* prop uniqueness *) +Fixpoint Uniq T (xs : seq T) := + if xs is x :: xs then x \Notin xs /\ Uniq xs else True. + +Lemma UniqP (T : eqType) (xs : seq T) : + reflect (Uniq xs) (uniq xs). +Proof. +elim: xs=>[|x xs IH] //=; first by apply: ReflectT. +case: andP=>H; constructor. +- by case: H=>/mem_seqP H /IH. +by case=>/mem_seqP H1 /IH H2; elim: H. +Qed. + +(***********************************) +(* 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 name, + 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 name, + 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 name, 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 name, 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. + +(********************************) +(* Relation preserving function *) +(********************************) + +Definition rel_on_func A B (f : B -> A) (p : Pred B) (g : Rel A) := + forall x y, p x -> g (f x) y -> + exists y', [/\ y = f y', p y' & (g \\o f) x y']. + +Notation "'on-' f" := (rel_on_func f) + (at level 1, format "'on-' f"). + +(* last conjunct can be dropped *) +Lemma onE A B (f : B -> A) p g : + on-f p g <-> + forall x y, p x -> g (f x) y -> exists2 y', y = f y' & p y'. +Proof. +split=>H x y C /[dup] X /(H _ _ C). +- by case=>y' []; exists y'. +by case=>y' E; exists y'; rewrite /Rel_app -E. +Qed. + +Add Parametric Morphism A B : (@rel_on_func A B) with signature + eq ==> @EqPredType _ _ ==> eqrel ==> iff as relon_eq. +Proof. +move=>f p1 p2 H r1 r2 E; split=>X; +by apply/onE=>x y /H /= P /E /(X _ _ P) [y'][-> /H]; eauto. +Qed. + +Add Parametric Morphism A B : (@rel_on_func A B) with signature + eq ==> @EqPredType _ _ ==> subrel ==> flip impl as relon_sub. +Proof. +move=>f p1 p2 H r1 r2 E X. +by apply/onE=>x y /H /= P /E /(X _ _ P) [y'][-> /H]; eauto. +Qed. + +(**********************************) +(**********************************) +(* Theory 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. + +(* DEVCOMMENT *) +(* Using sym_left_Transitive as a view doesn't work. *) +(* see https://github.com/coq/coq/issues/8352 *) +(* /DEVCOMMENT *) + +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). + +(* DEVCOMMENT *) +(* TODO: see if these can be simplified *) +(* /DEVCOMMENT *) + +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. + +#[export] 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. + +Section IteratedRels. +Variable T : Type. +Implicit Type g : T -> T -> Prop. + +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. + +Lemma iterS g n s1 s2 : + iter' g n.+1 s1 s2 <-> exists s, iter' g n s1 s /\ g s s2. +Proof. +rewrite -addn1; split; first by case/iter'_split=>s [H1] [s'][H2 <-]; exists s. +by case=>s [H1 H2]; apply: iter'_add H1 _; exists s2. +Qed. + +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 iter_sub g g' : g ~2> g' -> iter g ~2> iter g'. +Proof. +move=>H s s' [n X]; exists n; elim: n s s' X=>[|n IH] s s' //=. +by case=>x [/H G] /IH; exists x. +Qed. + +Lemma iter_eq g g' : g <~2> g' -> iter g <~2> iter g'. +Proof. by move=>H s s'; split; apply/iter_sub; rewrite H. Qed. + +Lemma iter1 g : g ~2> iter g. +Proof. by move=>s s'; exists 1, s'. Qed. + +Lemma iter_app T' g (f : T' -> T) : iter (g \\o f) ~2> iter g \\o f. +Proof. +move=>s s' [n] H; exists n; elim: n s s' H=>[|n IH] s s' /=; first by move=>->. +by case=>x [H] /IH; exists (f x). +Qed. + +Lemma iter_on T' (f : T' -> T) p g : on-f p g -> on-f p (iter g). +Proof. +move=>O; apply/onE=>x y C [n]. +elim: n x y C=>[|n IH] x y /= C; first by move=><-; exists x. +case=>_ [/(O _ _ C)] [x'][->] X' H /(IH _ _ X') [y' ->] Y'. +by exists y'. +Qed. + +Lemma iter_subf T' g (f : T' -> T) : iter (g \\o f) ~2> iter g \\o f. +Proof. +move=>s s' [n] H; exists n. +elim: n s s' H=>[|n IH] s s'; first by move=>->. +by case=>x [H1] {}/IH; exists (f x). +Qed. + +Lemma iter_fsub T' g (f : T' -> T) : + injective f -> + on-f xpredT g -> + iter g \\o f ~2> iter (g \\o f). +Proof. +move=>H1 H2 s s' [n] H; exists n. +elim: n s s' H =>[|n IH] s s' /=; first by move/H1. +by case=>_ [] /(H2 _ _ erefl) [x][-> _ X] {}/IH; exists x. +Qed. + +Lemma iter_eqf T' g (f : T' -> T) : + injective f -> + on-f xpredT g -> + iter g \\o f <~2> iter (g \\o f). +Proof. by move=>??; split; [apply: iter_fsub|apply: iter_subf]. Qed. + +(* localized variant of iter_fsub *) +Lemma iter_fsubl T' g (f : T' -> T) p s s' : + injective f -> + on-f p g -> + p s -> + (iter g \\o f) s s' -> iter (g \\o f) s s'. +Proof. +move=>H1 H2 P [n] H; exists n. +elim: n s s' P H=>[|n IH] s s' /= Ps; first by move/H1=>->. +by case=>_ [] /(H2 _ _ Ps) [x][-> Px H] /(IH _ _ Px); exists x. +Qed. + +End IteratedRels. + +Lemma iter_pair {T} {g : unit*T -> unit*T -> Prop} : + iter g \\o pair tt <~2> iter (g \\o pair tt). +Proof. by apply: iter_eqf=>[x y []|x [[y]]] //; exists y. Qed. + +Add Parametric Morphism A : (@iter A) with signature + eqrel ==> eqrel as iter_eq'. +Proof. by move=>x y H x1 y1; rewrite (iter_eq H _). Qed. + +Add Parametric Morphism A : (@iter A) with signature + subrel ==> subrel as iter_sub'. +Proof. by move=>x y H x1 y1; apply: iter_sub. 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. + +Arguments iter1 {T}%_type_scope {g _ _}. +Arguments iter_app {T T'}%_type_scope {g f}. +Arguments iter_refl {T g s}. + +#[export] 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 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 _ _ /. + +Lemma union_eq A (x1 y1 x2 y2 : Rel A) : + eqrel x1 y1 -> eqrel x2 y2 -> + eqrel (Union_rel x1 x2) (Union_rel y1 y2). +Proof. by rewrite /eqrel/Union_rel=>E1 E2 x y; rewrite E1 E2. Qed. + +Arguments union_eq {A x1 y1 x2 y2}. +Prenex Implicits union_eq. + +Add Parametric Morphism A : (@Union_rel A) with signature + eqrel ==> eqrel ==> eqrel as union_eq'. +Proof. by move=>*; apply: union_eq. Qed. + +Lemma union_sub A (x1 y1 x2 y2 : Rel A) : + subrel x1 y1 -> subrel x2 y2 -> + subrel (Union_rel x1 x2) (Union_rel y1 y2). +Proof. by move=>E1 E2 x y [/E1|/E2]; [left | right]. Qed. + +Arguments union_sub {A x1 y1 x2 y2}. +Prenex Implicits union_sub. + +Add Parametric Morphism A : (@Union_rel A) with signature + subrel ==> subrel ==> subrel as union_sub1. +Proof. by move=>*; apply: union_sub. Qed. + +Lemma subrelL A (x y : Rel A) : x ~2> Union_rel x y. +Proof. by left. Qed. + +Lemma subrelR A (x y : Rel A) : y ~2> Union_rel x y. +Proof. by right. Qed. + +#[export] +Instance eqrel_subrel A : subrelation (@eqrel A) (@subrel A). +Proof. by move=>x y H ?? /H. Qed. + +Lemma union_app A B D1 D2 (f : B -> A) : + Union_rel D1 D2 \\o f <~2> + Union_rel (D1 \\o f) (D2 \\o f). +Proof. by apply: union_eq. Qed. + +Arguments union_app {A B}%_type_scope {D1 D2 f}%_function_scope. +Prenex Implicits union_app. + +Lemma on_union A B D1 D2 (f : B -> A) p : + on-f p D1 -> + on-f p D2 -> + on-f p (Union_rel D1 D2). +Proof. +move=>H1 H2; apply/onE=>x y P; case. +- by case/(H1 _ _ P)=>y' [->] P'; exists y'. +by case/(H2 _ _ P)=>y' [->] P'; exists y'. +Qed. + +(*************************) +(* 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). + +(* DEVCOMMENT *) +(* (d2 d2' : mem_pred T2) (d3 d3' : mem_pred T3). *) +(* /DEVCOMMENT *) + +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. + +(* we can now state localized version of iter_fsubl *) +Lemma iter_fsubl' T T' (g : Rel T) (f : T' -> T) p : + injective f -> + on-f p g -> + {In p & xPredT, iter g \\o f ~2> iter (g \\o f)}. +Proof. by move=>H1 H2 x y H _; apply: iter_fsubl H1 H2 H. Qed. + + diff --git a/core/prelude.v b/core/prelude.v new file mode 100644 index 0000000..f0fe244 --- /dev/null +++ b/core/prelude.v @@ -0,0 +1,1091 @@ +(* +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 HB Require Import structures. +From Coq Require Import ssreflect ssrbool ssrfun Eqdep. +From mathcomp Require Import ssrnat seq eqtype choice fintype finfun. +From pcm Require Import options axioms. + +(***********) +(* Prelude *) +(***********) + +(* 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. + +Lemma trans_eq A (x y z : A) : x = y -> x = z -> y = z. +Proof. by move/esym; apply: eq_trans. 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. + +(************) +(* Products *) +(************) + +Inductive Prod3 (U1 U2 U3 : Type) := + mk3 {proj31 : U1; proj32 : U2; proj33 : U3}. +Prenex Implicits proj31 proj32 proj33. + +Inductive Prod4 (U1 U2 U3 U4 : Type) := + mk4 {proj41 : U1; proj42 : U2; proj43 : U3; proj44 : U4}. +Prenex Implicits proj41 proj42 proj43 proj44. + +Inductive Prod5 (U1 U2 U3 U4 U5 : Type) := + mk5 {proj51 : U1; proj52 : U2; proj53 : U3; proj54 : U4; proj55 : U5}. +Prenex Implicits proj51 proj52 proj53 proj54 proj55. + +Inductive Prod6 (U1 U2 U3 U4 U5 U6 : Type) := + mk6 {proj61 : U1; proj62 : U2; proj63 : U3; + proj64 : U4; proj65 : U5; proj66 : U6}. +Prenex Implicits proj61 proj62 proj63 proj64 proj65 proj66. + +Inductive Prod7 (U1 U2 U3 U4 U5 U6 U7 : Type) := + mk7 {proj71 : U1; proj72 : U2; proj73 : U3; + proj74 : U4; proj75 : U5; proj76 : U6; proj77 : U7}. +Prenex Implicits proj71 proj72 proj73 proj74 proj75 proj76 proj77. + +Definition eq3 (U1 U2 U3 : eqType) (x y : Prod3 U1 U2 U3) := + [&& proj31 x == proj31 y, proj32 x == proj32 y & proj33 x == proj33 y]. +Definition eq4 (U1 U2 U3 U4 : eqType) (x y : Prod4 U1 U2 U3 U4) := + [&& proj41 x == proj41 y, proj42 x == proj42 y, proj43 x == proj43 y & + proj44 x == proj44 y]. +Definition eq5 (U1 U2 U3 U4 U5 : eqType) (x y : Prod5 U1 U2 U3 U4 U5) := + [&& proj51 x == proj51 y, proj52 x == proj52 y, proj53 x == proj53 y, + proj54 x == proj54 y & proj55 x == proj55 y]. +Definition eq6 (U1 U2 U3 U4 U5 U6 : eqType) (x y : Prod6 U1 U2 U3 U4 U5 U6) := + [&& proj61 x == proj61 y, proj62 x == proj62 y, proj63 x == proj63 y, + proj64 x == proj64 y, proj65 x == proj65 y & proj66 x == proj66 y]. +Definition eq7 (U1 U2 U3 U4 U5 U6 U7 : eqType) (x y : Prod7 U1 U2 U3 U4 U5 U6 U7) := + [&& proj71 x == proj71 y, proj72 x == proj72 y, proj73 x == proj73 y, + proj74 x == proj74 y, proj75 x == proj75 y, proj76 x == proj76 y & + proj77 x == proj77 y]. + +Lemma eq3P U1 U2 U3 : Equality.axiom (@eq3 U1 U2 U3). +Proof. +rewrite /eq3; case=>x1 x2 x3 [y1 y2 y3] /=. +by do ![case: eqP=>[->|]]; constructor=>//; case. +Qed. + +Lemma eq4P U1 U2 U3 U4 : Equality.axiom (@eq4 U1 U2 U3 U4). +Proof. +rewrite /eq4; case=>x1 x2 x3 x4 [y1 y2 y3 y4] /=. +by do ![case: eqP=>[->|]]; constructor=>//; case. +Qed. + +Lemma eq5P U1 U2 U3 U4 U5 : Equality.axiom (@eq5 U1 U2 U3 U4 U5). +Proof. +rewrite /eq5; case=>x1 x2 x3 x4 x5 [y1 y2 y3 y4 y5] /=. +by do ![case: eqP=>[->|]]; constructor=>//; case. +Qed. + +Lemma eq6P U1 U2 U3 U4 U5 U6 : Equality.axiom (@eq6 U1 U2 U3 U4 U5 U6). +Proof. +rewrite /eq6; case=>x1 x2 x3 x4 x5 x6 [y1 y2 y3 y4 y5 y6] /=. +by do ![case: eqP=>[->|]]; constructor=>//; case. +Qed. + +Lemma eq7P U1 U2 U3 U4 U5 U6 U7 : Equality.axiom (@eq7 U1 U2 U3 U4 U5 U6 U7). +Proof. +rewrite /eq7; case=>x1 x2 x3 x4 x5 x6 x7 [y1 y2 y3 y4 y5 y6 y7] /=. +by do ![case: eqP=>[->|]]; constructor=>//; case. +Qed. + +HB.instance Definition _ (U1 U2 U3 : eqType) := + hasDecEq.Build (Prod3 U1 U2 U3) (@eq3P U1 U2 U3). + +HB.instance Definition _ (U1 U2 U3 U4 : eqType) := + hasDecEq.Build (Prod4 U1 U2 U3 U4) (@eq4P U1 U2 U3 U4). + +HB.instance Definition _ (U1 U2 U3 U4 U5 : eqType) := + hasDecEq.Build (Prod5 U1 U2 U3 U4 U5) (@eq5P U1 U2 U3 U4 U5). + +HB.instance Definition _ (U1 U2 U3 U4 U5 U6 : eqType) := + hasDecEq.Build (Prod6 U1 U2 U3 U4 U5 U6) (@eq6P U1 U2 U3 U4 U5 U6). + +HB.instance Definition _ (U1 U2 U3 U4 U5 U6 U7 : eqType) := + hasDecEq.Build (Prod7 U1 U2 U3 U4 U5 U6 U7) (@eq7P U1 U2 U3 U4 U5 U6 U7). + + +(***************************) +(* 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. + +Lemma compE A B C (g : B -> C) (f : A -> B) x : g (f x) = (g \o f) x. +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, aka. fork/fanout/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. + +(************************) +(* 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 , P6 , P7 , P8 , P9 , P10 & P11 ]" (at level 0, format + "'[hv' [ /\ '[' P1 , '/' P2 , '/' P3 , '/' P4 , '/' P5 , '/' P6 , '/' P7 , '/' P8 , '/' P9 , '/' P10 ']' '/ ' & P11 ] ']'"). + +Reserved Notation "[ /\ P1 , P2 , P3 , P4 , P5 , P6 , P7 , P8 , P9 , P10 , P11 & P12 ]" (at level 0, format + "'[hv' [ /\ '[' P1 , '/' P2 , '/' P3 , '/' P4 , '/' P5 , '/' P6 , '/' P7 , '/' P8 , '/' P9 , '/' P10 , '/' P11 ']' '/ ' & P12 ] ']'"). + + +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 and11 (P1 P2 P3 P4 P5 P6 P7 P8 P9 P10 P11 : Prop) : Prop := + And11 of P1 & P2 & P3 & P4 & P5 & P6 & P7 & P8 & P9 & P10 & P11. +Inductive and12 (P1 P2 P3 P4 P5 P6 P7 P8 P9 P10 P11 P12 : Prop) : Prop := + And12 of P1 & P2 & P3 & P4 & P5 & P6 & P7 & P8 & P9 & P10 & P11 & P12. + +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 , P6 , P7 , P8 , P9 , P10 & P11 ]" := + (and11 P1 P2 P3 P4 P5 P6 P7 P8 P9 P10 P11) : type_scope. +Notation "[ /\ P1 , P2 , P3 , P4 , P5 , P6 , P7 , P8 , P9 , P10 , P11 & P12 ]" := + (and12 P1 P2 P3 P4 P5 P6 P7 P8 P9 P10 P11 P12) : 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 *) + +(* DEVCOMMENT *) +(* TODO: we should move some of the following to [ssrbool] in Coq *) +(* /DEVCOMMENT *) + +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 *) +#[export] Program Instance and3_impl_morphism : Proper (iter_respectful impl impl 3) and3 | 1. +#[export] Program Instance and3_iff_morphism : Proper (iter_respectful iff iff 3) and3 | 1. + +#[export] Program Instance and4_impl_morphism : Proper (iter_respectful impl impl 4) and4 | 1. +#[export] Program Instance and4_iff_morphism : Proper (iter_respectful iff iff 4) and4 | 1. + +#[export] Program Instance and5_impl_morphism : Proper (iter_respectful impl impl 5) and5 | 1. +#[export] Program Instance and5_iff_morphism : Proper (iter_respectful iff iff 5) and5 | 1. + +#[export] Program Instance and6_impl_morphism : Proper (iter_respectful impl impl 6) and6 | 1. +#[export] Program Instance and6_iff_morphism : Proper (iter_respectful iff iff 6) and6 | 1. + +#[export] Program Instance and7_impl_morphism : Proper (iter_respectful impl impl 7) and7 | 1. +#[export] Program Instance and7_iff_morphism : Proper (iter_respectful iff iff 7) and7 | 1. + +#[export] Program Instance and8_impl_morphism : Proper (iter_respectful impl impl 8) and8 | 1. +#[export] Program Instance and8_iff_morphism : Proper (iter_respectful iff iff 8) and8 | 1. + +#[export] Program Instance and9_impl_morphism : Proper (iter_respectful impl impl 9) and9 | 1. +#[export] Program Instance and9_iff_morphism : Proper (iter_respectful iff iff 9) and9 | 1. + +#[export] Program Instance and10_impl_morphism : Proper (iter_respectful impl impl 10) and10 | 1. +#[export] Program Instance and10_iff_morphism : Proper (iter_respectful iff iff 10) and10 | 1. + +#[export] Program Instance and11_impl_morphism : Proper (iter_respectful impl impl 11) and11 | 1. +#[export] Program Instance and11_iff_morphism : Proper (iter_respectful iff iff 11) and11 | 1. + +#[export] Program Instance and12_impl_morphism : Proper (iter_respectful impl impl 12) and12 | 1. +#[export] Program Instance and12_iff_morphism : Proper (iter_respectful iff iff 12) and12 | 1. + +(** Logical disjunction *) +#[export] Program Instance or3_impl_morphism : Proper (iter_respectful impl impl 3) or3 | 1. +#[export] Program Instance or3_iff_morphism : Proper (iter_respectful iff iff 3) or3 | 1. + +#[export] Program Instance or4_impl_morphism : Proper (iter_respectful impl impl 4) or4 | 1. +#[export] Program Instance or4_iff_morphism : Proper (iter_respectful iff iff 4) or4 | 1. + +#[export] Program Instance or5_impl_morphism : Proper (iter_respectful impl impl 5) or5 | 1. +#[export] Program Instance or5_iff_morphism : Proper (iter_respectful iff iff 5) or5 | 1. + +#[export] Program Instance or6_impl_morphism : Proper (iter_respectful impl impl 6) or6 | 1. +#[export] Program Instance or6_iff_morphism : Proper (iter_respectful iff iff 6) or6 | 1. + +#[export] Program Instance or7_impl_morphism : Proper (iter_respectful impl impl 7) or7 | 1. +#[export] 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 b11 b12 : bool. + +Lemma and6P : reflect [/\ b1, b2, b3, b4, b5 & b6] [&& b1, b2, b3, b4, b5 & b6]. +Proof. +case: b1=>/=; last by constructor; case. +case: b2=>/=; last by constructor; case. +case: b3=>/=; last by constructor; case. +case: b4=>/=; last by constructor; case. +case: b5=>/=; last by constructor; case. +case: b6=>/=; last by constructor; case. +by constructor. +Qed. + +Lemma and7P : reflect [/\ b1, b2, b3, b4, b5, b6 & b7] [&& b1, b2, b3, b4, b5, b6 & b7]. +Proof. +case: b1=>/=; last by constructor; case. +case: b2=>/=; last by constructor; case. +case: b3=>/=; last by constructor; case. +case: b4=>/=; last by constructor; case. +case: b5=>/=; last by constructor; case. +case: b6=>/=; last by constructor; case. +case: b7=>/=; last by constructor; case. +by constructor. +Qed. + +Lemma and8P : reflect [/\ b1, b2, b3, b4, b5, b6, b7 & b8] [&& b1, b2, b3, b4, b5, b6, b7 & b8]. +Proof. +case: b1=>/=; last by constructor; case. +case: b2=>/=; last by constructor; case. +case: b3=>/=; last by constructor; case. +case: b4=>/=; last by constructor; case. +case: b5=>/=; last by constructor; case. +case: b6=>/=; last by constructor; case. +case: b7=>/=; last by constructor; case. +case: b8=>/=; last by constructor; case. +by constructor. +Qed. + +Lemma and9P : reflect [/\ b1, b2, b3, b4, b5, b6, b7, b8 & b9] [&& b1, b2, b3, b4, b5, b6, b7, b8 & b9]. +Proof. +case: b1=>/=; last by constructor; case. +case: b2=>/=; last by constructor; case. +case: b3=>/=; last by constructor; case. +case: b4=>/=; last by constructor; case. +case: b5=>/=; last by constructor; case. +case: b6=>/=; last by constructor; case. +case: b7=>/=; last by constructor; case. +case: b8=>/=; last by constructor; case. +case: b9=>/=; last by constructor; case. +by constructor. +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. +case: b1=>/=; last by constructor; case. +case: b2=>/=; last by constructor; case. +case: b3=>/=; last by constructor; case. +case: b4=>/=; last by constructor; case. +case: b5=>/=; last by constructor; case. +case: b6=>/=; last by constructor; case. +case: b7=>/=; last by constructor; case. +case: b8=>/=; last by constructor; case. +case: b9=>/=; last by constructor; case. +case: b10=>/=; last by constructor; case. +by constructor. +Qed. + +Lemma and11P : reflect [/\ b1, b2, b3, b4, b5, b6, b7, b8, b9, b10 & b11] + [&& b1, b2, b3, b4, b5, b6, b7, b8, b9, b10 & b11]. +Proof. +case: b1=>/=; last by constructor; case. +case: b2=>/=; last by constructor; case. +case: b3=>/=; last by constructor; case. +case: b4=>/=; last by constructor; case. +case: b5=>/=; last by constructor; case. +case: b6=>/=; last by constructor; case. +case: b7=>/=; last by constructor; case. +case: b8=>/=; last by constructor; case. +case: b9=>/=; last by constructor; case. +case: b10=>/=; last by constructor; case. +case: b11=>/=; last by constructor; case. +by constructor. +Qed. + +Lemma and12P : reflect [/\ b1, b2, b3, b4, b5, b6, b7, b8, b9, b10, b11 & b12] + [&& b1, b2, b3, b4, b5, b6, b7, b8, b9, b10, b11 & b12]. +Proof. +case: b1=>/=; last by constructor; case. +case: b2=>/=; last by constructor; case. +case: b3=>/=; last by constructor; case. +case: b4=>/=; last by constructor; case. +case: b5=>/=; last by constructor; case. +case: b6=>/=; last by constructor; case. +case: b7=>/=; last by constructor; case. +case: b8=>/=; last by constructor; case. +case: b9=>/=; last by constructor; case. +case: b10=>/=; last by constructor; case. +case: b11=>/=; last by constructor; case. +case: b12=>/=; last by constructor; case. +by constructor. +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 and11P {b1 b2 b3 b4 b5 b6 b7 b8 b9 b10 b11}. +Arguments and12P {b1 b2 b3 b4 b5 b6 b7 b8 b9 b10 b11 b12}. + +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}. + +Lemma iffPb (b1 b2 : bool) : reflect (b1 <-> b2) (b1 == b2). +Proof. +case: eqP=>[->|N]; constructor=>//. +case: b1 b2 N; case=>//= _. +- by case=>/(_ erefl). +by case=>_ /(_ erefl). +Qed. + +Lemma iffE (b1 b2 : bool) : b1 = b2 <-> (b1 <-> b2). +Proof. by split=>[->|] //; move/iffPb/eqP. Qed. + +Lemma subsetC T (p q : mem_pred T) : + {subset p <= q} -> {subset predC q <= predC p}. +Proof. by move=>H x; apply: contra (H x). Qed. + +Lemma subsetR T (p q : mem_pred T) : + {subset p <= predC q} -> {subset q <= predC p}. +Proof. by move=>H x; apply: contraL (H x). Qed. + +Lemma subsetL T (p q : mem_pred T) : + {subset predC p <= q} -> {subset predC q <= p}. +Proof. by move=>H x; apply: contraR (H x). Qed. + +Lemma subsetLR T (p q : mem_pred T) : + {subset predC p <= predC q} -> {subset q <= p}. +Proof. by move=>H x; apply: contraLR (H x). Qed. + +Lemma subset_disj T (p q r : mem_pred T) : + {subset p <= q} -> + (forall x, x \in q -> x \in r -> false) -> + (forall x, x \in p -> x \in r -> false). +Proof. by move=>H D x /H/D. Qed. + +Lemma subset_disj2 T (p p1 q q1 : mem_pred T) : + {subset p1 <= p} -> + {subset q1 <= q} -> + (forall x, x \in p -> x \in q -> false) -> + (forall x, x \in p1 -> x \in q1 -> false). +Proof. +move=>H1 H2 D1; apply: subset_disj H1 _ => x H /H2. +by apply: D1 H. +Qed. + +(**************) +(* empty type *) +(**************) + +Lemma False_eqP : Equality.axiom (fun _ _ : False => true). +Proof. by case. Qed. + +HB.instance Definition _ := hasDecEq.Build False False_eqP. + +(*************) +(* 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. + +#[export] Hint Extern 0 (ocancel _ _) => + (apply: lft_inl_ocanc || apply: rgt_inr_ocanc) : core. + +(****************) +(* option types *) +(****************) + +(* Alternative mechanism for manipulating *) +(* properties of the form isSome X. *) +(* To use a lemma of the form isSome X *) +(* one tyoically has to explicitly put the conclusion *) +(* of that lemma onto the context, and then case on X. *) +(* If such lemma is restated using is_some X *) +(* then it suffices to case on the lemma's name. *) +(* This saves typing X explicitly, which *) +(* may be significant if X is large. *) + +Inductive is_some_spec A x : option A -> Prop := +| is_some_case v of x = Some v : is_some_spec x (Some v). + +Hint Resolve is_some_case : core. + +Notation is_some x := (is_some_spec x x). + +Lemma is_someP A (x : option A) : reflect (is_some x) (isSome x). +Proof. by case: x=>[a|]; constructor=>//; case. Qed. + +(* some simplifications *) +Lemma oapp_some A (x : option A) : oapp [eta Some] None x = x. +Proof. by case: x. Qed. + +Lemma obind_some A (x : option A) : obind [eta Some] x = x. +Proof. exact: oapp_some. Qed. + +(********) +(* nats *) +(********) + +Lemma gt0 m n : m < n -> 0 < n. +Proof. by case: n. Qed. + +Lemma neq0 m n : m < n -> n != 0. +Proof. by move/gt0; rewrite lt0n. Qed. + +Lemma neqSn n : n.+1 != n. +Proof. by elim: n. Qed. + +Lemma neqnS n : n != n.+1. +Proof. by elim: n. Qed. + +(**************************************) +(* Inhabited (non-empty) finite types *) +(**************************************) + +Lemma inhabits (T : finType) (t : T) : 0 < #|T|. +Proof. by apply/card_gt0P; exists t. Qed. + +Lemma inhabits_irr (T : finType) (t1 t2 : T) : + inhabits t1 = inhabits t2. +Proof. by apply: bool_irrelevance. Qed. + +Definition inhabited_axiom (T : finType) := 0 < #|T|. + +HB.mixin Record isInhabited T of Finite T := { + card_inhab : inhabited_axiom T}. + +#[short(type="ifinType")] +HB.structure Definition Inhabited := {T of isInhabited T & Finite T}. + +Lemma inhabF (T : ifinType) : ~ (@predT T =1 xpred0). +Proof. by case/card_gt0P: (@card_inhab T)=>x _ /(_ x). Qed. + +Definition inhab {T : ifinType} : T := + match pickP predT with + | Pick x _ => x + | Nopick pf => False_rect T (inhabF pf) + end. + +HB.instance Definition _ := isInhabited.Build unit (inhabits tt). +HB.instance Definition _ := isInhabited.Build bool (inhabits false). +HB.instance Definition _ n := isInhabited.Build 'I_n.+1 (inhabits ord0). +HB.instance Definition _ (T : finType) := + isInhabited.Build (option T) (inhabits None). + +(*************************************) +(* A copy of booleans with mnemonics *) +(* LL and RR for working with sides *) +(*************************************) + +Inductive Side := LL | RR. +Definition Side_eq x y := + match x, y with LL, LL => true | RR, RR => true | _, _ => false end. +Definition nat2Side x := if odd x then LL else RR. +Definition Side2nat x := if x is RR then 0 else 1. + +Lemma ssrcanc : ssrfun.cancel Side2nat nat2Side. Proof. by case. Qed. +HB.instance Definition _ : isCountable Side := CanIsCountable ssrcanc. + +Lemma Side_enumP : Finite.axiom [:: LL; RR]. Proof. by case. Qed. +HB.instance Definition _ : isFinite Side := isFinite.Build Side Side_enumP. +HB.instance Definition _ := isInhabited.Build Side (inhabits LL). + +(*************) +(* sequences *) +(*************) + +(* TODO upstream to mathcomp *) +Section Fold. + +Lemma map_foldr {T1 T2} (f : T1 -> T2) xs : + map f xs = foldr (fun x ys => f x :: ys) [::] xs. +Proof. by []. Qed. + +Lemma fusion_foldr {T R Q} (g : R -> Q) f0 f1 z0 z1 (xs : seq T) : + (forall x y, g (f0 x y) = f1 x (g y)) -> g z0 = z1 -> + g (foldr f0 z0 xs) = foldr f1 z1 xs. +Proof. by move=>Hf Hz; elim: xs=>//= x xs <-. Qed. + +Lemma fusion_foldl {T R Q} (g : R -> Q) f0 f1 z0 z1 (xs : seq T) : + (forall x y, g (f0 x y) = f1 (g x) y) -> g z0 = z1 -> + g (foldl f0 z0 xs) = foldl f1 z1 xs. +Proof. +move=>Hf Hz; elim: xs z0 z1 Hz =>//= x xs IH z0 z1 Hz. +by apply: IH; rewrite Hf Hz. +Qed. + +Lemma foldl_foldr {T R} (f : R -> T -> R) z xs : + foldl f z xs = foldr (fun b g x => g (f x b)) id xs z. +Proof. by elim: xs z=>/=. Qed. + +Lemma foldr_foldl {T R} (f : T -> R -> R) z xs : + foldr f z xs = foldl (fun g b x => g (f b x)) id xs z. +Proof. +elim/last_ind: xs z=>//= xs x IH z. +by rewrite foldl_rcons -IH foldr_rcons. +Qed. + +End Fold. + +(* TODO upstream to mathcomp *) +Lemma pmap_pcomp {S T U} (f : T -> option U) (g : S -> option T) s : + pmap (pcomp f g) s = pmap f (pmap g s). +Proof. by elim: s=>//= x s ->; rewrite /pcomp; case: (g x). Qed. + +(* sequence prefixes *) + +(* Two helper concepts for searching in sequences: *) +(* *) +(* - onth: like nth, but returns None when the element is not found *) +(* - Prefix: a prefix relation on sequences, used for growing *) +(* interpretation contexts *) + + +Fixpoint onth A (s : seq A) n : option A := + if s is x::sx then if n is nx.+1 then onth sx nx else Some x else None. + +Definition Prefix A (s1 s2 : seq A) : Prop := + forall n x, onth s1 n = some x -> onth s2 n = some x. + +(* Lemmas *) + +Section SeqPrefix. +Variable A : Type. +Implicit Type s : seq A. + +Variant onth_spec s n : bool -> Type := +| onth_none : onth s n = None -> onth_spec s n false +| onth_some v : onth s n = Some v -> onth_spec s n true. + +Lemma onth_sizeP s n : onth_spec s n (n < size s). +Proof. +elim: s n=>/= [|a s IH] n; first by rewrite ltn0; constructor. +case: n=>[|n] /=; first by apply: (@onth_some _ _ a). +rewrite ltnS; case: (IH n)=>[|v] H. +- by constructor. +by apply: (@onth_some _ _ v). +Qed. + +Lemma size_onth s n : n < size s -> exists x, onth s n = Some x. +Proof. +by case: onth_sizeP=>// v H _; exists v. +Qed. + +Lemma onth_size s n x : onth s n = Some x -> n < size s. +Proof. +by case: onth_sizeP=>//->. +Qed. + +Lemma size_onthPn s n : reflect (onth s n = None) (size s <= n). +Proof. +by rewrite leqNgt; apply: (iffP idP); case: onth_sizeP=>//= v ->. +Qed. + +Lemma onth_sizeN s : onth s (size s) = None. +Proof. by apply/size_onthPn. Qed. + +Lemma nth_onth x0 n s : nth x0 s n = odflt x0 (onth s n). +Proof. +elim: s n=>/= [|a s IH] n /=; first by apply: nth_nil. +by case: n. +Qed. + +Lemma onth_nth x0 n s : n < size s -> onth s n = Some (nth x0 s n). +Proof. +elim: s n=>//= a s IH n. +by rewrite ltnS; case: n. +Qed. + +Lemma Prefix_refl s : Prefix s s. +Proof. by move=>n x <-. Qed. + +Lemma Prefix_trans s2 s1 s3 : + Prefix s1 s2 -> Prefix s2 s3 -> Prefix s1 s3. +Proof. by move=>H1 H2 n x E; apply: H2; apply: H1. Qed. + +Lemma Prefix_cons x s1 s2 : Prefix (x :: s1) (x :: s2) <-> Prefix s1 s2. +Proof. by split=>E n; [apply: (E n.+1) | case: n]. Qed. + +Lemma Prefix_cons' x y s1 s2 : + Prefix (x :: s1) (y :: s2) -> x = y /\ Prefix s1 s2. +Proof. by move=>H; case: (H 0 x (erefl _)) (H)=>-> /Prefix_cons. Qed. + +Lemma Prefix_rcons x s : Prefix s (rcons s x). +Proof. by elim: s=>//= y ys IH; apply/Prefix_cons; apply: IH. Qed. + +Lemma Prefix_cat s1 s2 : Prefix s1 (s1 ++ s2). +Proof. +elim: s2 s1=>[|x xs IH] s1; first by rewrite cats0. +rewrite -cat_rcons; apply: Prefix_trans (IH _). +by apply: Prefix_rcons. +Qed. + +Lemma Prefix_size s1 s2 : Prefix s1 s2 -> size s1 <= size s2. +Proof. +elim: s1 s2=>[//|a s1 IH] [|b s2] H; first by move: (H 0 a (erefl _)). +by rewrite ltnS; apply: (IH _ (proj2 (Prefix_cons' H))). +Qed. + +Lemma Prefix_onth s t x : + x < size s -> + Prefix s t -> onth s x = onth t x. +Proof. +elim:s t x =>[//|a s IH] [|b t] x H1 H2; first by move: (H2 0 a (erefl _)). +by case/Prefix_cons': H2=><- H2; case: x H1=>[|n] //= H1; apply: IH. +Qed. + +Lemma PrefixE s1 s2 : Prefix s1 s2 <-> exists s3, s2 = s1 ++ s3. +Proof. +split; last by case=>s3 ->; apply: Prefix_cat. +elim: s1 s2=>[|x xs IH] s2; first by exists s2. +case: s2=>[/(_ 0 x erefl)//|y ys /Prefix_cons' [?]]. +by subst y=>/IH [s3 ->]; exists s3. +Qed. + +End SeqPrefix. + +#[export] Hint Resolve Prefix_refl : core. + +Lemma onth_mem (A : eqType) (s : seq A) n x : + onth s n = Some x -> x \in s. +Proof. +by elim: s n=>//= a s IH [[->]|n /IH]; rewrite inE ?eq_refl // orbC =>->. +Qed. + +Lemma onth_index (A : eqType) (s : seq A) x : + onth s (index x s) = if x \in s then Some x else None. +Proof. +by elim: s=>//=h s IH; rewrite inE eq_sym; case: eqP=>//= ->. +Qed. + +Lemma PrefixP (A : eqType) (s1 s2 : seq A) : + reflect (Prefix s1 s2) (prefix s1 s2). +Proof. +apply/(equivP (prefixP (s1:=s1) (s2:=s2))). +by apply: iff_sym; exact: PrefixE. +Qed. + +(******************************) +(* Some commuting conversions *) +(******************************) + +Lemma fun_op A B C (b : option A) (vS : A -> B) (vN : B) (f : B -> C) : + f (if b is Some v then vS v else vN) = + if b is Some v then f (vS v) else f vN. +Proof. by case: b=>//. Qed. + +Lemma op_if A B (b : bool) (vS vN : option A) (vS1 : A -> B) (vN1 : B) : + (if (if b then vS else vN) is Some v then vS1 v else vN1) = + if b then if vS is Some v then vS1 v else vN1 + else if vN is Some v then vS1 v else vN1. +Proof. by case: b. Qed. + +Lemma if_triv b : (if b then true else false) = b. +Proof. by case: b. Qed. + +(*************************************************************) +(* quick ways to extract projections and transitivity proofs *) +(* out of iterated inequalities *) +(*************************************************************) + +Lemma ltn13 a b c : a < b < c -> a < c. +Proof. by case/andP; apply: ltn_trans. Qed. + +Lemma ltn12 a b c : a < b < c -> a < b. +Proof. by case/andP. Qed. + +Lemma ltn23 a b c : a < b < c -> b < c. +Proof. by case/andP. Qed. + +Lemma leq13 a b c : a <= b <= c -> a <= c. +Proof. by case/andP; apply: leq_trans. Qed. + +Lemma leq12 a b c : a <= b <= c -> a <= b. +Proof. by case/andP. Qed. + +Lemma leq23 a b c : a <= b <= c -> b <= c. +Proof. by case/andP. Qed. + +Lemma lqt13 a b c : a <= b < c -> a < c. +Proof. by case/andP; apply: leq_ltn_trans. Qed. + +Lemma lqt12 a b c : a <= b < c -> a <= b. +Proof. by case/andP. Qed. + +Lemma lqt23 a b c : a <= b < c -> b < c. +Proof. by case/andP. Qed. + +(********************) +(* Finite functions *) +(********************) + +Section FinFun. +Variables (T : finType) (Us : T -> Type). +Implicit Type f : {dffun forall t, Us t}. + +(* Explicit name for finite function application. *) +(* Will be used to hang canonical projections onto. *) +(* The function/argument order is reversed to facilitate rewriting. *) +Definition sel tg f : Us tg := f tg. + +Lemma ffinP f1 f2 : (forall t, sel t f1 = sel t f2) <-> f1 = f2. +Proof. by rewrite ffunP. Qed. + +(* beta equality *) +Lemma sel_fin t (f : forall t, Us t) : sel t (finfun f) = f t. +Proof. by rewrite /sel ffunE. Qed. + +(* eta equality *) +Lemma fin_eta f : f = finfun (sel^~ f). +Proof. by apply/ffinP=>t; rewrite sel_fin. Qed. + +(* function *) +Definition splice tg f (v : Us tg) : {dffun _} := + finfun (fun x => + if decP (x =P tg) is left pf then cast Us pf v + else sel x f). + +Lemma sel_splice t f x (v : Us x) : + sel t (splice f v) = + if decP (t =P x) is left pf then cast Us pf v + else sel t f. +Proof. by rewrite sel_fin. Qed. + +Lemma sel_spliceE t f v : sel t (splice f v) = v. +Proof. by rewrite sel_fin; case: eqP=>//= pf; rewrite eqd. Qed. + +Lemma sel_spliceN t x f (w : Us x) : + t <> x -> sel t (splice f w) = sel t f. +Proof. by move=>N; rewrite sel_fin; case: eqP. Qed. + +Lemma splice_eta t f : splice f (sel t f) = f. +Proof. +apply/ffinP=>x; rewrite sel_fin; case: eqP=>//=. +by move/[dup]=>-> pf; rewrite eqd. +Qed. + +End FinFun. + +Arguments sel {T Us} tg f. +Arguments splice {T Us tg} f v. + +(* notation for building finfuns *) +(* DEVCOMMENT *) +(* copied from finfun to fix some bad spacing in formatting *) +(* /DEVCOMMENT *) + +Notation "[ 'ffun' x : aT => E ]" := (finfun (fun x : aT => E)) + (at level 0, x name, format "[ 'ffun' x : aT => E ]") : fun_scope. + +Notation "[ 'ffun' x => E ]" := (@finfun _ (fun=> _) (fun x => E)) + (at level 0, x name, format "[ 'ffun' x => E ]") : fun_scope. + +Notation "['ffun' => E ]" := [ffun _ => E] + (at level 0, format "['ffun' => E ]") : fun_scope. + +Section IteratedNotation. +Variables (T : finType) (Us : T -> Type). + +Variant dfun_delta : Type := DFunDelta t of Us t. + +(* for iteration that starts with function ends with function *) +Definition dapp_fdelta df (f : forall t, Us t) z := + let: DFunDelta t v := df in + if decP (z =P t) is left pf then cast Us pf v + else f z. + +(* for iteration that starts with finfun ends with function *) +Definition splice' df (f : {ffun forall t, Us t}) z := + dapp_fdelta df f z. + +End IteratedNotation. + +Delimit Scope fun_delta_scope with FUN_DELTA. +Arguments dapp_fdelta {T Us} df%_FUN_DELTA f. + +Notation "y \\ x" := (@DFunDelta _ _ x y) (at level 1). + +(* notation for simultaneous update of f with d1,..,dn *) +(* rewrite by sel_fin peels all layers *) +Notation "[ 'ext' f 'with' d1 , .. , dn ]" := + (finfun ( + dapp_fdelta d1%_FUN_DELTA .. (dapp_fdelta dn%_FUN_DELTA f) ..)) + (at level 0, format + "'[hv' [ '[' 'ext' '/ ' f ']' '/' 'with' '[' d1 , '/' .. , '/' dn ']' ] ']'" + ) : fun_scope. + +(* notation for iterated update of f with d1, then d2, ... *) +(* rewrite by sel_fin peels top layer only *) +Notation "[ 'splice' F 'with' d1 , .. , dn ]" := + (finfun (splice' + d1%_FUN_DELTA .. (finfun (splice' dn%_FUN_DELTA (finfun F))) ..)) + (at level 0, format + "'[hv' [ '[' 'splice' '/ ' F ']' '/' 'with' '[' d1 , '/' .. , '/' dn ']' ] ']'" + ) : fun_scope. + +Section TestingNotation. +Variables (T : finType) (Us : T -> Type). +Variables (f : {dffun forall t, Us t}) (t1 t2 : T) (v1 : Us t1) (v2 : Us t2). + +(* we have three different options in displaying splices *) +(* splice, [splice], and [ext] *) +Lemma test : + [/\ sel t2 [splice f with v1 \\ t1, v2 \\ t2] = v2, + sel t2 (splice (splice f v1) v2) = v2, + sel t2 [ext f with v1 \\ t1, v2 \\ t2] = v2 & +(* and we can use underscores to elide some info *) + sel t2 [ext f with v1 \\ _, v2 \\ _] = v2]. +Abort. +End TestingNotation. + +(* mapping over simply-typed finite functions *) + +Section FinFunMap. +Variables (T : finType) (A B : Type). +Implicit Types (f : A -> B) (x : {ffun T -> A}). + +Definition fmap f x := [ffun tg => f (sel tg x)]. + +Lemma sel_fmap f x tg : sel tg (fmap f x) = f (sel tg x). +Proof. exact: sel_fin. Qed. + +Lemma fmap_splice f x tg (v : A) : + fmap f (splice (tg:=tg) x v) = splice (tg:=tg) (fmap f x) (f v). +Proof. +apply/ffinP=>t; rewrite !sel_fin; case: eqP=>//= ?; subst t. +by rewrite !eqc. +Qed. + +End FinFunMap. + + +(* Tagging *) + +Notation Tag := (@existT _ _). + +Lemma Tag_inj T Us (t1 t2 : T) i1 i2 : + Tag t1 i1 = Tag t2 i2 -> + t1 = t2 /\ jmeq Us i1 i2. +Proof. by case=>?; subst t2=>/inj_pair2 ->. Qed. +Arguments Tag_inj {T Us t1 t2 i1 i2}. + +(* tagged union of equality types is equality type *) + +Section TaggedEq. +Variables (T : eqType) (Us : T -> eqType). + +Definition tag_eq : sigT Us -> sigT Us -> bool := + fun '(Tag tx opx) '(Tag ty opy) => + if decP (tx =P ty) is left pf then opx == cast Us pf opy + else false. + +Lemma tag_eqP : Equality.axiom tag_eq. +Proof. +case=>tx opx [ty opy] /=; case: (tx =P ty)=>pf; last first. +- by constructor; case=>/pf. +subst ty; rewrite /= eqc; case: eqP=>pf; constructor; +by [rewrite pf|case=>/inj_pair2/pf]. +Qed. + +HB.instance Definition _ := hasDecEq.Build (sigT Us) tag_eqP. +End TaggedEq. + diff --git a/core/rtc.v b/core/rtc.v new file mode 100644 index 0000000..72e7c72 --- /dev/null +++ b/core/rtc.v @@ -0,0 +1,273 @@ +(* +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 pcm Require Import options axioms pred prelude finmap seqext. + +(* 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 := Finite.clone (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 H; apply: Rsub. +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/seqext.v b/core/seqext.v new file mode 100644 index 0000000..75457cf --- /dev/null +++ b/core/seqext.v @@ -0,0 +1,1779 @@ +(* +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. +*) + +From Coq Require Import ssreflect ssrbool ssrfun. +From mathcomp Require Import ssrnat seq path eqtype choice fintype bigop. +From pcm Require Import options prelude pred. + +(*********************) +(* Extensions to seq *) +(*********************) + +Lemma nilp_hasPn {A} (s : seq A) : nilp s = ~~ has predT s. +Proof. by case: s. Qed. + +Lemma filter_predIC {A} (s : seq A) p1 p2 : + filter (predI p1 p2) s = filter (predI p2 p1) s. +Proof. by apply: eq_filter => z /=; rewrite andbC. Qed. + +Lemma filter_swap {A} (s : seq A) p1 p2 : + filter p1 (filter p2 s) = filter p2 (filter p1 s). +Proof. by rewrite -!filter_predI filter_predIC. Qed. + +(* TODO contribute to mathcomp? *) +Lemma map_nilp {A B} (f : A -> B) (s : seq A) : nilp (map f s) = nilp s. +Proof. by rewrite /nilp; case: s. Qed. + +Lemma filter_nilp {A} (p : pred A) (s : seq A) : nilp (filter p s) = ~~ has p s. +Proof. by rewrite /nilp size_filter -leqn0 leqNgt has_count. Qed. + +Lemma head_map {P Q} (f : P -> Q) z (s : seq P) : + f (head z s) = head (f z) (map f s). +Proof. by case: s. Qed. + +Lemma zip_map2 {P Q R S} (f : P -> R) (g : Q -> S) (s1 : seq P) (s2 : seq Q) : + zip (map f s1) (map g s2) = + map (fun '(x1,x2) => (f x1,g x2)) (zip s1 s2). +Proof. +elim: s1 s2=>/= [|x1 s1 IH] [|x2 s2] //=. +by congr cons. +Qed. + +Corollary zip_mapl {P Q R} (f : P -> R) (s1 : seq P) (s2 : seq Q) : + zip (map f s1) s2 = + map (fun '(x1,x2) => (f x1,x2)) (zip s1 s2). +Proof. by rewrite -{1}(map_id s2) zip_map2. Qed. + +Corollary zip_mapr {P Q S} (g : Q -> S) (s1 : seq P) (s2 : seq Q) : + zip s1 (map g s2) = + map (fun '(x1,x2) => (x1,g x2)) (zip s1 s2). +Proof. by rewrite -{1}(map_id s1) zip_map2. Qed. + +Lemma drop_take_id {A} x (s : seq A) : drop x (take x s) = [::]. +Proof. by rewrite -{2}(add0n x) -take_drop take0. Qed. + +Lemma drop_take_mask {A} (s : seq A) x y : + drop x (take y s) = mask (nseq x false ++ nseq (y-x) true) s. +Proof. +case: (ltnP x (size s))=>Hx; last first. +- rewrite drop_oversize; last by rewrite size_take_min geq_min Hx orbT. + rewrite -{1}(subnKC Hx) nseqD -catA -{3}(cats0 s) mask_cat; last by rewrite size_nseq. + by rewrite mask0 mask_false. +have Hx': size (nseq x false) = size (take x s). +- by rewrite size_nseq size_take_min; symmetry; apply/minn_idPl/ltnW. +rewrite -{2}(cat_take_drop x s) mask_cat // mask_false /= -takeEmask take_drop. +case: (leqP x y)=>[Hxy|/ltnW Hxy]; first by rewrite subnK. +move: (Hxy); rewrite -subn_eq0=>/eqP->; rewrite add0n drop_take_id. +by rewrite drop_oversize // size_take_min geq_min Hxy. +Qed. + +Section LemmasEq. +Variables (A : eqType). +Implicit Type xs : seq A. + +(* With A : Type, we have the In_split lemma. *) +(* With A : eqType, the lemma can be strenghtened to *) +(* not only return the split of xs, but the split of xs *) +(* that uses the first occurrence of x is xs *) +Lemma in_split xs x : + x \in xs -> exists xs1 xs2, xs = xs1 ++ x :: xs2 /\ x \notin xs1. +Proof. +rewrite -has_pred1; case/split_find=>_ s1 s2 /eqP ->. +by rewrite has_pred1=>H; exists s1, s2; rewrite -cats1 -catA. +Qed. + +Lemma subset_nilR xs : + {subset xs <= [::]} -> xs = [::]. +Proof. by case: xs=>// x xs /(_ x); rewrite inE eqxx=>/(_ erefl). Qed. + +Lemma subset_nil xs ys : + {subset xs <= ys} -> ys = [::] -> xs = [::]. +Proof. by move=>X E; move: E X=>->; apply: subset_nilR. Qed. + +Lemma all_mem xs ys : + reflect {subset ys <= xs} (all [mem xs] ys). +Proof. by case: allP=>H; constructor; [move=>x /H | move=>X; apply: H=>x /X]. Qed. + +Lemma all_predC_sym xs ys : + all [predC xs] ys = all [predC ys] xs. +Proof. by rewrite all_predC has_sym -all_predC. Qed. + +Lemma nilp_filter (p : pred A) s : + reflect {in s, forall z, ~~ p z} (nilp (filter p s)). +Proof. +case E : (nilp _); constructor. +- move: E; rewrite nilp_hasPn=>/hasPn H x Kx; apply/negP=>Px. + by move: (H x); rewrite mem_filter Px=>/(_ Kx). +move=>X; move/negP: E; elim. +rewrite nilp_hasPn; apply/negP=>/hasP [x]. +rewrite mem_filter=>/andP [Px Kx] _. +by move: (X x Kx); rewrite Px. +Qed. + +Lemma index_rcons a x 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 x xs : + x \notin xs <-> index x xs = size xs. +Proof. +split; first by exact: memNindex. +by move=>E; rewrite -index_mem E ltnn. +Qed. + +Corollary index_sizeE x xs : + reflect (index x xs = size xs) (x \notin xs). +Proof. by apply/(equivP idP)/index_memN. Qed. + +Lemma size0nilP xs : + reflect (xs = [::]) (size xs == 0). +Proof. +case: eqP=>X; constructor; first by move/size0nil: X. +by move=>N; rewrite N in X. +Qed. + +Lemma has_nilP xs : + reflect (has predT xs) (xs != [::]). +Proof. by case: xs=>[|x xs]; constructor. Qed. + +Lemma map_nilP {B : eqType} (f : B -> A) s : + reflect (exists k, k \in map f s) (map f s != [::]). +Proof. +case: has_nilP=>X; constructor. +- by case/hasP: X=>x; exists x. +by case=>k K; elim: X; apply/hasP; exists k. +Qed. + +Lemma filter_nilP (p : pred A) xs : + reflect (forall x, p x -> x \in xs -> false) + ([seq x <- xs | p x] == [::]). +Proof. +case: eqP=>E; constructor. +- move=>x H1 H2; suff : x \in [seq x <- xs | 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 filter_pred1 x xs : + x \notin xs -> filter (pred1 x) xs = [::]. +Proof. +move=>H; apply/eqP; apply/filter_nilP=>z /eqP ->. +by rewrite (negbTE H). +Qed. + +Lemma filter_predC1 x xs : + x \notin xs -> filter (predC1 x) xs = xs. +Proof. +by move=>H; apply/all_filterP/allP=>y /=; case: eqP=>// ->; apply/contraL. +Qed. + +Lemma filter_mem_sym (s1 s2 : seq A) : + filter (mem s1) s2 =i filter (mem s2) s1. +Proof. by move=>x; rewrite !mem_filter andbC. Qed. + +Lemma index_inj xs x y : + x \in xs -> index x xs = index y xs -> x = y. +Proof. +elim: xs=>[|k xs IH] //=; rewrite inE eq_sym. +case: eqP=>[->{k} _|_ /= S]; case: eqP=>// _ []; apply: IH S. +Qed. + +Lemma cat_cancel (xs1 xs2 ys1 ys2 : seq A) (k : A) : + k \notin xs1 -> k \notin xs2 -> + xs1 ++ k :: ys1 = xs2 ++ k :: ys2 -> + (xs1 = xs2) * (ys1 = ys2). +Proof. +move=>Nk1 Nk2 E. +have Es : size xs1 = size xs2. +- have : index k (xs1++k::ys1) = index k (xs2++k::ys2) by rewrite E. + by rewrite !index_cat /= (negbTE Nk1) (negbTE Nk2) eqxx !addn0. +have Ex : xs1 = take (size xs1) (xs1 ++ k :: ys1). +- by rewrite take_cat ltnn subnn /= cats0. +rewrite E Es take_cat ltnn subnn /= cats0 in Ex. +rewrite {xs1 Nk1 Es}Ex in E *. +have : ys1 = drop (size (xs2++[::k])) (xs2++k::ys1). +- by rewrite drop_cat size_cat /= addn1 ltnNge ltnW //= subSn // subnn /= drop0. +by rewrite E drop_cat size_cat /= addn1 ltnNge ltnW //= subSn // subnn /= drop0. +Qed. + +(* if the list is not empty, the default value in head doesn't matter *) +Lemma head_dflt (x1 x2 x : A) xs : + x \in xs -> head x1 xs = head x2 xs. +Proof. by case: xs. Qed. + +Lemma mem_head (x : A) xs : head x xs \in x :: xs. +Proof. by case: xs=>[|y ys]; rewrite !inE //= eqxx orbT. Qed. + +(* a common pattern of using mem_head that avoids forward reasoning *) +Lemma mem_headI (x : A) xs a : + a = head x xs -> a \in x :: xs. +Proof. by move=>->; apply: mem_head. Qed. + +Lemma head_nilp (x : A) xs : + x \notin xs -> head x xs = x -> nilp xs. +Proof. +elim: xs=>[|y ys IH] //= H1 H2. +by rewrite H2 inE eqxx /= in H1. +Qed. + +Lemma head_notin (x y : A) xs : + y \in xs -> x \notin xs -> head x xs != x. +Proof. +move=>Y X; apply/negP=>/eqP; move/(head_nilp X)/nilP=>E. +by rewrite E in Y. +Qed. + +(* weaker form of in_mask *) +(* TODO contribute to mathcomp *) +Lemma in_mask_count x m xs : + count_mem x xs <= 1 -> + x \in mask m xs = (x \in xs) && nth false m (index x xs). +Proof. +elim: xs m => [|y xs IHs] m /=; first by rewrite mask0 in_nil. +case: m=>/=[|b m]; first by rewrite in_nil nth_nil andbF. +case: b; rewrite !inE eq_sym; case: eqP=>//= _. +- by rewrite add0n; apply: IHs. +- rewrite -{2}(addn0 1%N) leq_add2l leqn0 => /eqP Hc. + rewrite IHs; last by rewrite Hc. + by move/count_memPn/negbTE: Hc=>->. +by rewrite add0n; apply: IHs. +Qed. + +Lemma mem_take_index x xs : + x \notin take (index x xs) xs. +Proof. +elim: xs=>//=h xs; case: eqP=>//= /eqP H IH. +by rewrite inE negb_or eq_sym H. +Qed. + +Lemma prefix_drop_sub (s1 s2 : seq A) : + seq.prefix s1 s2 -> + forall n, {subset (drop n s1) <= drop n s2}. +Proof. +case/seq.prefixP=>s0 {s2}-> n x H. +rewrite drop_cat; case: ltnP=>Hn. +- by rewrite mem_cat H. +by move: H; rewrite drop_oversize. +Qed. + +End LemmasEq. + +(* finding last occurrence *) + +Section FindLast. +Variables (T : Type). +Implicit Types (x : T) (p : pred T) (s : seq T). + +(* helper function for finding the last occurrence in a single pass *) +(* calculates index and size *) +Definition findlast_aux oi0 p s : option nat * nat := + foldl (fun '(o,i) x => (if p x then Some i else o, i.+1)) oi0 s. + +Lemma findlast_auxE oi0 p s : + findlast_aux oi0 p s = + let k := seq.find p (rev s) in + (if k == size s then oi0.1 + else Some (oi0.2 + (size s - k).-1), oi0.2 + size s). +Proof. +rewrite /findlast_aux; elim: s oi0=>/= [|x s IH] [o0 i0] /=; first by rewrite addn0. +rewrite IH /= rev_cons -cats1 find_cat /= has_find. +move: (find_size p (rev s)); rewrite size_rev; case: ltngtP=>// H _. +- case: eqP=>[E|_]; first by rewrite E ltnNge leqnSn in H. + apply: injective_projections=>/=; [congr Some | rewrite addSnnS]=>//. + by rewrite !predn_sub /= -predn_sub addSnnS prednK // subn_gt0. +case: ifP=>_; rewrite addSnnS; last by rewrite addn1 eqxx. +by rewrite addn0 eqn_leq leqnSn /= ltnn subSnn addn0. +Qed. + +Definition findlast p s : nat := + let '(o, i) := findlast_aux (None, 0) p s in odflt i o. + +(* finding last is finding first in reversed list and flipping indices *) +Corollary findlastE p s : + findlast p s = + if has p s then (size s - seq.find p (rev s)).-1 + else size s. +Proof. +rewrite /findlast findlast_auxE /= !add0n -has_rev; case/boolP: (has p (rev s)). +- by rewrite has_find size_rev; case: ltngtP. +by move/hasNfind=>->; rewrite size_rev eqxx. +Qed. + +Lemma findlast_size p s : findlast p s <= size s. +Proof. +rewrite findlastE; case: ifP=>// _. +by rewrite -subnS; apply: leq_subr. +Qed. + +Lemma has_findlast p s : has p s = (findlast p s < size s). +Proof. +symmetry; rewrite findlastE; case: ifP=>H /=; last by rewrite ltnn. +by rewrite -subnS /= ltn_subrL /=; case: s H. +Qed. + +Lemma hasNfindlast p s : ~~ has p s -> findlast p s = size s. +Proof. by rewrite has_findlast; case: ltngtP (findlast_size p s). Qed. + +Lemma findlast0 p x : findlast p [::] = 0. +Proof. by []. Qed. + +Lemma findlast1 p x : findlast p [::x] = ~~ p x. +Proof. by rewrite findlastE /= orbF; case: ifP=>// ->. Qed. + +Lemma findlast_cat p s1 s2 : + findlast p (s1 ++ s2) = + if has p s1 && ~~ has p s2 + then findlast p s1 + else size s1 + findlast p s2. +Proof. +rewrite !findlastE has_cat rev_cat find_cat has_rev size_cat size_rev. +case/boolP: (has p s2)=>H2; last first. +- rewrite orbF; case/boolP: (has p s1)=>//= H1. + by rewrite addnC subnDl. +have H2' : find p (rev s2) < size s2. +- by rewrite -size_rev -has_find has_rev. +rewrite /= orbT andbF -addnBA; last by apply: ltnW. +rewrite -!subn1 -subnDA -addnBA; last by rewrite subn_gt0. +by rewrite subnDA. +Qed. + +Lemma findlast_cons p x s : + findlast p (x::s) = + if p x && ~~ has p s then 0 else (findlast p s).+1. +Proof. +rewrite -cat1s findlast_cat /= !add1n orbF findlast1. +by case: ifP=>// /andP [->]. +Qed. + +Lemma findlast_rcons p x s : + findlast p (rcons s x) = + if p x then size s + else if has p s then findlast p s + else (size s).+1. +Proof. +rewrite -cats1 findlast_cat /= orbF findlast1. +case: (p x)=>/=; first by rewrite andbF addn0. +by rewrite andbT addn1. +Qed. + +Lemma nth_findlast x0 p s : has p s -> p (nth x0 s (findlast p s)). +Proof. +rewrite findlastE=>/[dup] E ->; rewrite -has_rev in E. +rewrite -subnS -nth_rev; last by rewrite -size_rev -has_find. +by apply: nth_find. +Qed. + +Lemma has_drop p s i : has p s -> has p (drop i s) = (i <= findlast p s). +Proof. +rewrite findlastE=>/[dup] E ->; rewrite -has_rev in E. +have Hh: 0 < size s - find p (rev s). +- by rewrite -size_rev subn_gt0 -has_find. +rewrite -size_rev; move/(has_take (size s - i)): (E). +rewrite take_rev -subnS size_rev. +case/boolP: (i < size s)=>[Hi|]. +- rewrite subnA //; last by apply: ltnW. + rewrite subnn add0n has_rev=>->. + rewrite ltn_subRL addnC -ltn_subRL subnS. + by case: (size s - find p (rev s)) Hh. +rewrite -ltnNge ltnS => Hi _. +rewrite drop_oversize //=; symmetry; apply/negbTE. +rewrite -ltnNge subnS prednK //. +by apply/leq_trans/Hi; exact: leq_subr. +Qed. + +Lemma find_geq p s i : has p (drop i s) -> i <= findlast p s. +Proof. +case/boolP: (has p s)=>Hp; first by rewrite (has_drop _ Hp). +suff: ~~ has p (drop i s) by move/negbTE=>->. +move: Hp; apply: contra; rewrite -{2}(cat_take_drop i s) has_cat=>->. +by rewrite orbT. +Qed. + +Lemma find_leq_last p s : find p s <= findlast p s. +Proof. +rewrite findlastE. +case/boolP: (has p s)=>[|_]; last by apply: find_size. +elim: s=>//= h s IH. +rewrite rev_cons -cats1 find_cat has_rev size_rev /=. +case/orP; first by move=>->. +move=>/[dup] H ->; case: ifP=>_ //. +rewrite subSn /=; last first. +- by rewrite -size_rev; apply: find_size. +apply: (leq_ltn_trans (IH H)); rewrite ltn_predL subn_gt0. +by rewrite -size_rev -has_find has_rev. +Qed. + +Variant split_findlast_nth_spec p : seq T -> seq T -> seq T -> T -> Type := + FindLastNth x s1 s2 of p x & ~~ has p s2 : + split_findlast_nth_spec p (rcons s1 x ++ s2) s1 s2 x. + +Lemma split_findlast_nth x0 p s (i := findlast p s) : + has p s -> + split_findlast_nth_spec p s (take i s) (drop i.+1 s) (nth x0 s i). +Proof. +move=> p_s; rewrite -[X in split_findlast_nth_spec _ X](cat_take_drop i s). +rewrite (drop_nth x0 _); last by rewrite -has_findlast. +rewrite -cat_rcons; constructor; first by apply: nth_findlast. +by rewrite has_drop // ltnn. +Qed. + +Variant split_findlast_spec p : seq T -> seq T -> seq T -> Type := + FindLastSplit x s1 s2 of p x & ~~ has p s2 : + split_findlast_spec p (rcons s1 x ++ s2) s1 s2. + +Lemma split_findlast p s (i := findlast p s) : + has p s -> + split_findlast_spec p s (take i s) (drop i.+1 s). +Proof. +by case: s => // x ? in i * =>?; case: split_findlast_nth=>//; constructor. +Qed. + +End FindLast. + +Section FindLastEq. +Variables (T : eqType). +Implicit Type s : seq T. + +Definition indexlast (x : T) : seq T -> nat := findlast (pred1 x). + +Lemma indexlast_size x s : indexlast x s <= size s. +Proof. by rewrite /indexlast; apply: findlast_size. Qed. + +Lemma indexlast_mem x s : (indexlast x s < size s) = (x \in s). +Proof. by rewrite /indexlast -has_findlast has_pred1. Qed. + +Lemma memNindexlast x s : x \notin s -> indexlast x s = size s. +Proof. by rewrite -has_pred1=>/hasNfindlast. Qed. + +Lemma indexlast0 x : indexlast x [::] = 0. +Proof. by []. Qed. + +Lemma indexlast1 x y : indexlast x [::y] = (x != y). +Proof. by rewrite /indexlast findlast1 /= eq_sym. Qed. + +Lemma indexlast_cat x s1 s2 : + indexlast x (s1 ++ s2) = + if (x \in s1) && (x \notin s2) + then indexlast x s1 + else size s1 + indexlast x s2. +Proof. by rewrite /indexlast findlast_cat !has_pred1. Qed. + +Lemma indexlast_cons x y s : + indexlast x (y::s) = + if (y == x) && (x \notin s) then 0 else (indexlast x s).+1. +Proof. by rewrite /indexlast findlast_cons has_pred1. Qed. + +Lemma indexlast_rcons x y s : + indexlast x (rcons s y) = + if y == x then size s + else if x \in s then indexlast x s + else (size s).+1. +Proof. by rewrite /indexlast findlast_rcons has_pred1. Qed. + +Lemma index_geq x s i : x \in drop i s -> i <= indexlast x s. +Proof. by rewrite -has_pred1; apply: find_geq. Qed. + +Lemma index_leq_last x s : index x s <= indexlast x s. +Proof. by apply: find_leq_last. Qed. + +Lemma indexlast_count x s : count_mem x s <= 1 <-> index x s = indexlast x s. +Proof. +elim: s=>//= h t IH; rewrite indexlast_cons. +case: eqP=>/= ?; last first. +- by rewrite add0n IH; split=>[->|[]]. +rewrite add1n ltnS leqn0; split. +- by move/eqP/count_memPn=>->. +by case: ifP=>//= /count_memPn->. +Qed. + +Corollary index_lt_last x s : 1 < count_mem x s -> index x s < indexlast x s. +Proof. +move=>H; move: (index_leq_last x s); rewrite leq_eqVlt. +by case: eqP=>//= /indexlast_count; case: leqP H. +Qed. + +Corollary indexlast_uniq x s : uniq s -> index x s = indexlast x s. +Proof. +move=>H; apply/indexlast_count. +by rewrite count_uniq_mem //; apply: leq_b1. +Qed. + +Lemma indexlast_memN x xs : + x \notin xs <-> indexlast x xs = size xs. +Proof. +split; first by exact: memNindexlast. +by move=>E; rewrite -indexlast_mem E ltnn. +Qed. + +Lemma index_last_inj x y s : + (x \in s) || (y \in s) -> index x s = indexlast y s -> x = y. +Proof. +elim: s=>[|k s IH] //=; rewrite !inE indexlast_cons !(eq_sym k). +case: eqP=>[{k}<- _|_ /= S]; first by case: eqP=>//=. +move: S; case/boolP: (y \in s)=>/=. +- by rewrite andbF=>H _ []; apply: IH; rewrite H orbT. +move=>Ny; rewrite orbF andbT. +by case: eqP=>//; rewrite orbF=>_ H []; apply: IH; rewrite H. +Qed. + +Lemma indexlast_inj x y s : + x \in s -> indexlast x s = indexlast y s -> x = y. +Proof. +elim: s=>[|k s IH] //=; rewrite inE eq_sym !indexlast_cons. +case: eqP=>[->{k} _|_ /= S] /=. +- case: eqP=>//= _. + by case: ifP=>// /negbT; rewrite negbK=>H; case; apply: IH. +by case: ifP=>// _ []; apply: IH. +Qed. + +Lemma mem_drop_indexlast x s : + x \notin drop (indexlast x s).+1 s. +Proof. +elim: s=>//=h s; rewrite indexlast_cons. +case: eqP=>//= _ H. +by case: ifP=>//=; rewrite drop0. +Qed. + +Variant splitLast x : seq T -> seq T -> seq T -> Type := + SplitLast p1 p2 of x \notin p2 : splitLast x (rcons p1 x ++ p2) p1 p2. + +Lemma splitLastP s x (i := indexlast x s) : + x \in s -> + splitLast x s (take i s) (drop i.+1 s). +Proof. +case: s => // y s in i * => H. +case: split_findlast_nth=>//; first by rewrite has_pred1. +move=>_ s1 s2 /= /eqP->; rewrite has_pred1 => H2. +by constructor. +Qed. + +End FindLastEq. + +(* finding all occurrences *) + +Section FindAll. + +Variables (T : Type). +Implicit Types (x : T) (p : pred T) (s : seq T). + +(* helper function for finding all occurrences in a single pass with difference lists *) +Definition findall_aux oi0 p s : (seq nat -> seq nat) * nat := + foldl (fun '(s,i) x => (if p x then s \o cons i else s, i.+1)) oi0 s. + +Lemma findall_auxE oi0 p s : + (forall s1 s2 : seq nat, oi0.1 (s1 ++ s2) = oi0.1 s1 ++ s2) -> + let: (rs, ix) := findall_aux oi0 p s in + (forall s' : seq nat, + rs s' = oi0.1 (unzip1 (filter (p \o snd) (zip (iota oi0.2 (size s)) s))) ++ s') + /\ ix = oi0.2 + size s. +Proof. +move; rewrite /findall_aux; elim: s oi0=>[|x s IH] [o0 i0] /= H0. +- split; last by rewrite addn0. + by move=>s'; rewrite -{1}(cat0s s') H0. +case/boolP: (p x)=>/= Hpx. +- move: (IH (o0 \o cons i0, i0.+1))=>/=. + rewrite addSn addnS; apply=>s1 s2. + by rewrite -cat_cons H0. +by move: (IH (o0, i0.+1))=>/=; rewrite addSn addnS; apply. +Qed. + +Definition findall p s : seq nat := (findall_aux (id, 0) p s).1 [::]. + +Lemma findallE p s : + findall p s = unzip1 (filter (p \o snd) (zip (iota 0 (size s)) s)). +Proof. +rewrite /findall. +move: (@findall_auxE (id, 0) p s)=>/= /(_ (fun s1 s2 : seq nat => erefl)). +case E: (findall_aux (id, 0) p s)=>[rs ix] /=; case=>/(_ [::]) -> _. +by rewrite cats0. +Qed. + +Lemma findall_cat p s1 s2 : + findall p (s1 ++ s2) = + findall p s1 ++ map (fun n => size s1 + n) (findall p s2). +Proof. +rewrite !findallE size_cat iotaD add0n zip_cat; last by rewrite size_iota. +rewrite filter_cat {1}/unzip1 map_cat; congr (_ ++ _). +set n := size s1. +rewrite -{1}(addn0 n) iotaDl zip_mapl filter_map -!map_comp. +rewrite (eq_filter (a2:=(p \o snd))); last by case. +by apply: eq_map; case. +Qed. + +Corollary findall_cons p x s : + findall p (x::s) = + if p x then 0 :: map S (findall p s) else map S (findall p s). +Proof. by rewrite -cat1s findall_cat /= findallE /=; case: (p x). Qed. + +Corollary findall_rcons p x s : + findall p (rcons s x) = + if p x then rcons (findall p s) (size s) else findall p s. +Proof. +rewrite -cats1 findall_cat /= !findallE /=; case: (p x)=>/=. +- by rewrite addn0 cats1. +by rewrite cats0. +Qed. + +Lemma findall_size p s : + size (findall p s) = count p s. +Proof. +by elim: s=>//=x s IH; rewrite findall_cons; case: (p x)=>/=; +rewrite size_map IH. +Qed. + +Lemma findall_nilp p s : + nilp (findall p s) = ~~ has p s. +Proof. by rewrite /nilp findall_size has_count -leqNgt leqn0. Qed. + +Lemma findall_head p s : + find p s = head (size s) (findall p s). +Proof. +elim: s=>//=x s IH; rewrite findall_cons; case: (p x)=>//=. +by rewrite -head_map IH. +Qed. + +Lemma findall_last p s : + findlast p s = last (size s) (findall p s). +Proof. +elim/last_ind: s=>//=s x IH; rewrite findall_rcons findlast_rcons size_rcons. +case: (p x)=>/=; first by rewrite last_rcons. +by rewrite IH -(negbK (has _ _)) -findall_nilp; case: (findall p s). +Qed. + +Lemma findall_count1 p s : + count p s <= 1 -> + findall p s = if has p s then [::find p s] else [::]. +Proof. +rewrite -(findall_size p s); case E: (findall p s)=>[|h t] /=. +- by move/nilP: E; rewrite findall_nilp=>/negbTE ->. +have ->: has p s by rewrite -(negbK (has p s)) -findall_nilp E. +rewrite ltnS leqn0 => /eqP/size0nil Et; move: E; rewrite {t}Et. +by rewrite findall_head=>->. +Qed. + +End FindAll. + +Section FindAllEq. + +Variables (T : eqType). +Implicit Type s : seq T. + +Definition indexall (x : T) : seq T -> seq nat := findall (pred1 x). + +Lemma indexall_size x s : + size (indexall x s) = count_mem x s. +Proof. by rewrite /indexall findall_size. Qed. + +Lemma indexall_mem x s : nilp (indexall x s) = (x \notin s). +Proof. by rewrite /indexall findall_nilp has_pred1. Qed. + +Lemma indexall_head x s : + index x s = head (size s) (indexall x s). +Proof. by rewrite /index /indexall findall_head. Qed. + +Lemma indexall_last x s : + indexlast x s = last (size s) (indexall x s). +Proof. by rewrite /indexlast /indexall findall_last. Qed. + +Lemma indexall_count1 x s : + count_mem x s <= 1 -> + indexall x s = if x \in s then [::index x s] else [::]. +Proof. by rewrite /indexall /index -has_pred1; apply: findall_count1. Qed. + +Corollary indexall_uniq x s : + uniq s -> + indexall x s = if x \in s then [::index x s] else [::]. +Proof. by move=>U; apply: indexall_count1; rewrite (count_uniq_mem _ U) leq_b1. Qed. + +End FindAllEq. + +(* Interaction of filter/last/index *) + +Section FilterLastIndex. +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. + +Lemma last_notin_nilp k (s : seq A) : ~~ nilp s -> k \notin s -> last k s != k. +Proof. +move=>N; apply: (last_notin (x := head k s)). +by case: s N=>//= x s _; rewrite inE eqxx. +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 eqxx. 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 and rcons *) + +Lemma rcons_lastX x y (s : seq A) : + x \in s -> exists s', s = rcons s' (last y s). +Proof. +elim/last_ind: s=>[|ks k IH] //=. +by rewrite last_rcons; exists ks. +Qed. + +Lemma rcons_lastP x (s : seq A) : + reflect (exists s', s = rcons s' (last x s)) (last x s \in s). +Proof. +case X : (last x s \in s); constructor; first by apply: rcons_lastX X. +case=>s' E; move/negP: X; elim. +by rewrite E last_rcons mem_rcons inE eqxx. +Qed. + +Lemma rcons_lastXP x y (s : seq A) : + reflect (exists s', s = rcons s' x) ((x == last y s) && (x \in s)). +Proof. +case: eqP=>[->|N]; first by apply: rcons_lastP. +by constructor; case=>s' E; elim: N; rewrite E last_rcons. +Qed. + +Lemma index_last_size_uniq z (s : seq A) : + uniq s -> + index (last z s) s = (size s).-1. +Proof. +elim: s z=>//= x s IH z. +case/andP=>Nx U; rewrite eq_sym; case: eqVneq=>H. +- by rewrite (last_nochange_nil H Nx). +rewrite {}IH //; apply: prednK. +by case: {Nx U}s H=>//=; rewrite eqxx. +Qed. + +(* last has bigger index than anything in x *) +Lemma index_last_mono x k (s : seq A) : + uniq s -> x \in s -> index x s <= index (last k s) s. +Proof. +elim: s k=>[|k s IH] //= k'; rewrite inE !(eq_sym k). +case/andP=>K U; case: (x =P k)=>//= /eqP N X. +case: (last k s =P k)=>[/last_nochange|/eqP L]. +- by case: eqP X=>[->//|]; rewrite (negbTE K). +by apply: leq_trans (IH k U X) _. +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 : nilp s by move/nilP=>->. +rewrite /nilp 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 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. + +Lemma index_filter_ltL (p : pred A) (ks : seq A) (t1 t2 : A) : + (t1 \notin ks) || p t1 -> + (index t1 ks < index t2 ks) -> + (index t1 (filter p ks) < index t2 (filter p ks)). +Proof. +elim: ks t1 t2=>[|k ks IH] t1 t2 //=; rewrite inE negb_or (eq_sym t1). +case: eqP=>[->{k} /= Pt1|/eqP Nkt1 /= H]. +- by rewrite Pt1 /= eqxx; case: eqP. +case: eqP=>// /eqP Nkt2; case: ifP=>H1 /=. +- by rewrite (negbTE Nkt1) (negbTE Nkt2) !ltnS; apply: IH H. +by rewrite ltnS; apply: IH H. +Qed. + +Lemma index_filter_leL (p : pred A) (ks : seq A) (t1 t2 : A) : + (t1 \notin ks) || p t1 -> + (index t1 ks <= index t2 ks) -> + (index t1 (filter p ks) <= index t2 (filter p ks)). +Proof. +elim: ks t1 t2=>[|k ks IH] t1 t2 //=; rewrite inE negb_or (eq_sym t1). +case: eqP=>[->{k} /= Pt1|/eqP Nkt1 /= H]. +- by rewrite Pt1 /= eqxx; case: eqP. +case: eqP=>// /eqP Nkt2; case: ifP=>H1 /=. +- by rewrite (negbTE Nkt1) (negbTE Nkt2) !ltnS; apply: IH H. +by rewrite ltnS; apply: IH H. +Qed. + +Lemma index_filter_ltR (p : pred A) (ks : seq A) (t1 t2 : A) : + (t2 \notin 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 negb_or /=. +rewrite (eq_sym t2). +case: eqP=>[->{k} /= Pt2|/eqP Nkt2 /=]. +- by rewrite Pt2 /= eqxx; case: eqP. +case: eqP=>[->{t1}//|/eqP Nt1k]. +case: ifP=>H1 H2 /=. +- by rewrite (negbTE Nt1k) (negbTE Nkt2) !ltnS; apply: IH H2. +by rewrite ltnS; apply: IH H2. +Qed. + +Lemma index_filter_leR (p : pred A) (ks : seq A) (t1 t2 : A) : + (t2 \notin 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 negb_or /=. +rewrite (eq_sym t2). +case: eqP=>[->{k} /= Pt2|/eqP Nkt2 /=]. +- by rewrite Pt2 /= eqxx; case: eqP. +case: eqP=>[->{t1}//|/eqP Nt1k]. +case: ifP=>H1 H2 /=. +- by rewrite (negbTE Nt1k) (negbTE Nkt2) !ltnS; apply: IH H2. +by rewrite ltnS; apply: IH H2. +Qed. + +(* we can put the left and right lemmas together *) +Lemma index_filter_lt (p : pred A) (ks : seq A) (t1 t2 : A) : + (t1 \notin ks) || p t1 -> + (t2 \notin ks) || p t2 -> + (index t1 (filter p ks) < index t2 (filter p ks)) = + (index t1 ks < index t2 ks). +Proof. +move=>H1 H2; apply/idP/idP. +- by apply: index_filter_ltR. +by apply: index_filter_ltL. +Qed. + +Lemma index_filter_le (p : pred A) (ks : seq A) (t1 t2 : A) : + (t1 \notin ks) || p t1 -> + (t2 \notin ks) || p t2 -> + (index t1 (filter p ks) <= index t2 (filter p ks)) = + (index t1 ks <= index t2 ks). +Proof. +move=>H1 H2; apply/idP/idP. +- by apply: index_filter_leR. +by apply: index_filter_leL. +Qed. + +(* index and masking *) + +Lemma index_mask (s : seq A) m a b : + uniq s -> + a \in mask m s -> + b \in mask m s -> + index a (mask m s) < index b (mask m s) <-> + index a s < index b s. +Proof. +elim: m s=>[|x m IH][|k s] //= /andP [K U]; case: x=>[|Ma Mb] /=. +- rewrite !inE; case/orP=>[/eqP <-|Ma]. + - by case/orP=>[/eqP ->|]; rewrite eqxx //; case: eqP. + case/orP=>[/eqP ->|Mb]; first by rewrite eqxx. + by case: eqP; case: eqP=>//; rewrite ltnS IH. +case: eqP Ma K=>[-> /mem_mask -> //|Ka]. +case: eqP Mb=>[-> /mem_mask -> //|Kb Mb Ma]. +by rewrite ltnS IH. +Qed. + +Lemma indexlast_mask (s : seq A) m a b : + uniq s -> + a \in mask m s -> + b \in mask m s -> + indexlast a (mask m s) < indexlast b (mask m s) <-> + indexlast a s < indexlast b s. +Proof. +elim: m s=>[|x m IH][|k s] //= /andP [K U]; case: x=>[|Ma Mb] /=; +rewrite !indexlast_cons. +- rewrite !inE; case/orP=>[/eqP Ea|Ma]. + - rewrite -{k}Ea in K *. + have Km: (a \notin mask m s) by apply: contra K; apply: mem_mask. + case/orP=>[/eqP ->|]; rewrite eqxx /=; first by rewrite K ltnn. + case: eqP=>[->|Nab] H /=; first by rewrite !ltnn. + by rewrite K Km. + case/orP=>[/eqP Eb|Mb]. + - rewrite -{k}Eb in K *. + have ->: (b \notin mask m s) by apply: contra K; apply: mem_mask. + by rewrite eqxx /= K. + rewrite Ma Mb (mem_mask Ma) (mem_mask Mb) /= !andbF. + by apply: IH. +case: eqP Ma K=>[-> /mem_mask -> //|Ka] /=. +case: eqP Mb=>[-> /mem_mask -> //|Kb Mb Ma _] /=. +by rewrite ltnS IH. +Qed. + +Lemma index_subseq (s1 s2 : seq A) a b : + subseq s1 s2 -> + uniq s2 -> + a \in s1 -> + b \in s1 -> + index a s1 < index b s1 <-> index a s2 < index b s2. +Proof. by case/subseqP=>m _ ->; apply: index_mask. Qed. + +Lemma indexlast_subseq (s1 s2 : seq A) a b : + subseq s1 s2 -> + uniq s2 -> + a \in s1 -> + b \in s1 -> + indexlast a s1 < indexlast b s1 <-> indexlast a s2 < indexlast b s2. +Proof. by case/subseqP=>m _ ->; apply: indexlast_mask. Qed. + +End FilterLastIndex. + +(* index and mapping *) + +Section IndexPmap. +Variables (A B : eqType). + +Lemma index_pmap_inj (s : seq A) (f : A -> option B) a1 a2 b1 b2 : + injective f -> + f a1 = Some b1 -> + f a2 = Some b2 -> + index b1 (pmap f s) < index b2 (pmap f s) <-> index a1 s < index a2 s. +Proof. +move=>Inj E1 E2; elim: s=>[|k s IH] //=; rewrite /oapp. +case: eqP=>[->{k}|]. +- rewrite E1 /= eqxx. + case: (a1 =P a2) E1 E2=>[-> -> [/eqP ->] //|]. + by case: (b1 =P b2)=>[-> Na <- /Inj /esym/Na|]. +case: eqP=>[->{k} Na|N2 N1]; first by rewrite E2 /= eqxx !ltn0. +case E : (f k)=>[b|] //=. +case: eqP E1 E=>[-><- /Inj/N1 //|_ _]. +by case: eqP E2=>[-><- /Inj/N2 //|_ _ _]; rewrite IH. +Qed. + +Lemma index_pmap_inj_mem (s : seq A) (f : A -> option B) a1 a2 b1 b2 : + {in s &, injective f} -> + a1 \in s -> + a2 \in s -> + f a1 = Some b1 -> + f a2 = Some b2 -> + index b1 (pmap f s) < index b2 (pmap f s) <-> index a1 s < index a2 s. +Proof. +move=>Inj A1 A2 E1 E2. +elim: s Inj A1 A2=>[|k s IH] //= Inj; rewrite /oapp !inE !(eq_sym k). +case: eqP Inj=>[<-{k} /= Inj _|]. +- rewrite E1 /= !eqxx eq_sym. + case: eqP E1 E2=>[->-> [->]|]; first by rewrite eqxx. + case: eqP=>[-> Na <- E /= A2|//]. + by move/Inj: E Na=>-> //; rewrite inE ?(eqxx,A2,orbT). +case eqP=>[<-{k} Na Inj /= A1 _|]; first by rewrite E2 /= eqxx !ltn0. +move=>N2 N1 Inj /= A1 A2. +have Inj1 : {in s &, injective f}. +- by move=>x y X Y; apply: Inj; rewrite inE ?X ?Y ?orbT. +case E : (f k)=>[b|] /=; last by rewrite IH. +case: eqP E1 E=>[-> <- E|_ _]. +- by move/Inj: E N1=>-> //; rewrite inE ?(eqxx,A1,orbT). +case: eqP E2=>[-><- E|_ _ _]; last by rewrite IH. +by move/Inj: E N2=>-> //; rewrite inE ?(eqxx,A2,orbT). +Qed. + +(* we can relax the previous lemma a bit *) +(* the relaxation will be more commonly used than the previous lemma *) +(* because the option type gives us the implication that the second *) +(* element is in the map *) +Lemma index_pmap_inj_in (s : seq A) (f : A -> option B) a1 a2 b1 b2 : + {in s & predT, injective f} -> + f a1 = Some b1 -> + f a2 = Some b2 -> + index b1 (pmap f s) < index b2 (pmap f s) <-> index a1 s < index a2 s. +Proof. +move=>Inj E1 E2. +case A1 : (a1 \in s); last first. +- move/negbT/index_sizeE: (A1)=>->. + suff /index_sizeE -> : b1 \notin pmap f s by rewrite !ltnNge !index_size. + rewrite mem_pmap; apply/mapP; case=>x X /esym; rewrite -E1=>E. + by move/(Inj _ _ X): E A1=><- //; rewrite X. +case A2 : (a2 \in s). +- by apply: index_pmap_inj_mem=>// x y X _; apply: Inj. +move/negbT/index_sizeE: (A2)=>->. +suff /index_sizeE -> : b2 \notin pmap f s. +- by rewrite !index_mem /= A1 mem_pmap; split=>// _; apply/mapP; exists a1. +rewrite mem_pmap; apply/mapP; case=>x X /esym; rewrite -E2=>E. +by move/(Inj _ _ X): E A2=><- //; rewrite X. +Qed. + +End IndexPmap. + +Section SeqRel. +Variable (A : eqType). +Implicit Type (ltT leT : rel A). + +(* ordering with path, seq and last *) + +Lemma eq_last (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 (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 (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 eqxx. +by rewrite (H (last a s)) ?orbT // mem_last. +Qed. + +Lemma path_lastR (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 (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 (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 (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 (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 eqxx. +by move=>z D; apply: H2; rewrite inE D orbT. +Qed. + +Lemma seq_last_mono (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 (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 (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 (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 path_mem_irr (s : seq A) ltT x : + irreflexive ltT -> + transitive ltT -> + path ltT x s -> + x \notin s. +Proof. +move=>I T P; apply: contraFT (I x). +by rewrite negbK; apply: path_mem T P. +Qed. + +Lemma sorted_rcons (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 (s1 s2 : seq A) ltT : + irreflexive ltT -> + transitive ltT -> + sorted ltT s1 -> + sorted ltT 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: irr_sorted_eq S1 _ _=>//; first by rewrite sorted_filter. +by move=>k; rewrite mem_filter; case S : (_ \in _)=>//; rewrite (H _ S). +Qed. + +Lemma sorted_ord_index (s : seq A) ltT x y : + irreflexive ltT -> + transitive ltT -> + sorted ltT s -> + x \in s -> + ltT 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 path_ord_index_leq (s : seq A) leT x y : + transitive leT -> + antisymmetric leT -> + leT x y -> + path leT y s -> + x \in s -> + x = y. +Proof. +move=>T; elim: s x y=>[|a l IH] //= x y As Lxy. +case/andP=>Lya Pal; rewrite inE. +case: eqP Lya Pal As=>[<-{a} Lyx _ As _|Nxa Lya Pal /= As' X]. +- by apply: As=>//; rewrite Lxy Lyx. +by move/Nxa: (IH x a As' (T _ _ _ Lxy Lya) Pal X). +Qed. + +Lemma sorted_ord_index_leq (s : seq A) leT x y : + transitive leT -> + antisymmetric leT -> + sorted leT s -> + x \in s -> + leT x y -> + x != y -> + index x s < index y s. +Proof. +move=>T As S P H N; elim: s S As P=>[|z s IH] //= P As; rewrite inE !(eq_sym z). +case: eqP H P As=>[<-{z} H P As _|Nxz H P As /= X]; first by rewrite eq_sym (negbTE N). +case: eqP Nxz P=>[<-{z} Nxy P|Nyz Nxz P]. +- by move/Nxy: (path_ord_index_leq T As H P X). +by apply: IH X=>//; apply: path_sorted P. +Qed. + +Lemma sorted_index_ord (s : seq A) leT x y : + transitive leT -> + sorted leT s -> + y \in s -> + index x s < index y s -> + leT x y. +Proof. +move=>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. + +(* sorted, uniq, filter *) + +Lemma lt_sorted_uniq_le (s : seq A) ltT : + irreflexive ltT -> + antisymmetric ltT -> + transitive ltT -> + sorted ltT s = uniq s && + (sorted (fun k t => (k == t) || ltT k t) s). +Proof. +move=>I As T; case: s=>// n s; elim: s n=>//= m s IHs n. +rewrite inE negb_or IHs -!andbA /=. +case: (n =P m)=>[->|/eqP Nm /=]; first by rewrite I. +case lTnm : (ltT n m)=>/=; last by rewrite !andbF. +case Ns: (n \in s)=>//=; do !bool_congr. +have T' : transitive (fun k t => (k == t) || ltT k t). +- move=>x y z /orP [/eqP -> //|H]. + case/orP=>[/eqP <-|]; first by rewrite H orbT. + by move/(T _ _ _ H)=>->; rewrite orbT. +apply/negP=>/(order_path_min T')/allP/(_ n Ns). +rewrite eq_sym (negbTE Nm) /= =>lTmn. +by rewrite (As m n) ?eqxx // lTnm lTmn in Nm. +Qed. + +Lemma sort_sorted_in_lt (s : seq A) ltT : + irreflexive ltT -> + antisymmetric ltT -> + transitive ltT -> + uniq s -> + {in s &, total (fun k t => (k == t) || ltT k t)} -> + sorted ltT (sort (fun k t => (k == t) || ltT k t) s). +Proof. +move=>I S T U Tot; rewrite lt_sorted_uniq_le //. +by rewrite sort_uniq U (sort_sorted_in Tot _). +Qed. + +(* filtering and consecutive elements in an order *) +Lemma filterCN (ltT : rel A) f t1 t2 : + t1 \notin f -> + {in f, forall z, ltT z t2 = (z == t1) || ltT z t1} -> + filter (ltT^~ t2) f = filter (ltT^~ t1) f. +Proof. +move=>N C; apply: eq_in_filter=>x T; rewrite C ?inE ?orbT //. +by case: eqP N T=>// -> /negbTE ->. +Qed. + +Lemma filterCE (ltT : rel A) f t1 t2 : + irreflexive ltT -> + transitive ltT -> + sorted ltT f -> + {in f, forall z, ltT z t2 = (z == t1) || ltT z t1} -> + t1 \in f -> + filter (ltT^~ t2) f = filter (ltT^~ t1) f ++ [:: t1]. +Proof. +move=>I T S Z F; have U : uniq f by apply: sorted_uniq T I _ S. +rewrite -(filter_pred1_uniq U F); apply: irr_sorted_eq (T) I _ _ _ _ _. +- by apply: sorted_filter T _ _ S. +- rewrite -[filter (ltT^~ t1) _]revK -[filter (pred1 t1) _]revK -rev_cat. + rewrite rev_sorted -filter_rev filter_pred1_uniq ?(mem_rev,rev_uniq) //. + rewrite /= path_min_sorted ?(rev_sorted, sorted_filter T _ S) //. + by apply/allP=>x; rewrite mem_rev mem_filter=>/andP []. +move=>x; rewrite mem_cat !mem_filter /=. +case X: (x \in f); last by rewrite !andbF. +by rewrite Z // orbC !andbT. +Qed. + +(* frequently we have nested filtering and sorting *) +(* for which the following forms of the lemmas is more effective *) + +Lemma filter2CN (ltT : rel A) p f t1 t2 : + t1 \notin p -> + {in p, forall z, ltT z t2 = (z == t1) || ltT z t1} -> + filter (ltT^~ t2) (filter p f) = filter (ltT^~ t1) (filter p f). +Proof. +move=>N C; apply: filterCN; first by rewrite mem_filter negb_and N. +by move=>z; rewrite mem_filter=>/andP [D _]; apply: C. +Qed. + +Lemma filter2CE (ltT : rel A) (p : pred A) f t1 t2 : + irreflexive ltT -> + antisymmetric ltT -> + transitive ltT -> + {in f &, total (fun k t => (k == t) || ltT k t)} -> + {in p, forall z, ltT z t2 = (z == t1) || ltT z t1} -> + uniq f -> + p t1 -> t1 \in f -> + filter (ltT^~ t2) + (filter p (sort (fun k t => (k == t) || ltT k t) f)) = + filter (ltT^~ t1) + (filter p (sort (fun k t => (k == t) || ltT k t) f)) ++ [:: t1]. +Proof. +move=>I Asym T Tot Z U P F; apply: filterCE (I) (T) _ _ _. +- by rewrite (sorted_filter T _ _) //; apply: sort_sorted_in_lt. +- by move=>z; rewrite mem_filter=>/andP [Pz _]; apply: Z. +by rewrite mem_filter mem_sort P F. +Qed. + +(* nth *) + +Lemma nth_cons (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 (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 (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 eqxx. +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 (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 (s : seq A) x0 k i : nth x0 s i = nth x0 (k::s) i.+1. +Proof. by []. Qed. + +Lemma nth_leT (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 (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 (s : seq A) ltT x0 k i j : + irreflexive ltT -> + transitive ltT -> + i <= size s -> + j <= size s -> + path ltT k s -> + ltT (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 (s : seq A) ltT x0 k z i : + irreflexive ltT -> + transitive ltT -> + path ltT k s -> + ltT (nth x0 (k::s) i) z -> + ltT 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 (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,eqxx,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. + +End SeqRel. + +(* there always exists a nat not in a given list *) +Lemma not_memX (ks : seq nat) : exists k, k \notin ks. +Proof. +have L a xs : foldl addn a xs = a + foldl addn 0 xs. +- elim: xs a=>[|z xs IH] //= a; first by rewrite addn0. + by rewrite add0n // [in LHS]IH [in RHS]IH addnA. +set k := foldl addn 1 ks. +suff K a : a \in ks -> a < k by exists k; apply/negP=>/K; rewrite ltnn. +rewrite {}/k; elim: ks=>[|k ks IH] //=; rewrite inE. +case/orP=>[/eqP ->|/IH]; first by rewrite L add1n addSn ltnS leq_addr. +rewrite L=>N; rewrite L; apply: leq_trans N _. +by rewrite addnAC leq_addr. +Qed. + +Section BigCat. +Context {A B : Type}. +Implicit Types (xs : seq A) (f : A -> seq B). + +Lemma flatten_map_big xs f : + flatten (map f xs) = \big[cat/[::]]_(x <- xs) f x. +Proof. +elim: xs=>/= [|x xs IH]; first by rewrite big_nil. +by rewrite big_cons IH. +Qed. + +Lemma size_big_cat xs f : + size (\big[cat/[::]]_(x <- xs) f x) = + \sum_(x <- xs) (size (f x)). +Proof. +elim: xs=>[|x xs IH] /=; first by rewrite !big_nil. +by rewrite !big_cons size_cat IH. +Qed. + +Lemma has_big_cat (p : pred B) xs f : + has p (\big[cat/[::]]_(x <- xs) f x) = + has (fun x => has p (f x)) xs. +Proof. +elim: xs=>[|x xs IH]; first by rewrite big_nil. +by rewrite big_cons has_cat /= IH. +Qed. + +End BigCat. + +Lemma big_cat_mem_has A (B : eqType) xs (f : A -> seq B) b : + b \in \big[cat/[::]]_(x <- xs) f x = + has (fun x => b \in f x) xs. +Proof. +rewrite -has_pred1 has_big_cat; apply: eq_has=>x. +by rewrite has_pred1. +Qed. + +(* big_cat_mem for A : Type *) +Lemma big_cat_memT A (B : eqType) x xs (f : A -> seq B) : + reflect (exists2 i, i \In xs & x \in f i) + (x \in \big[cat/[::]]_(i <- xs) f i). +Proof. by rewrite big_cat_mem_has; apply/hasPIn. Qed. + +(* big_cat_mem for A : eqType *) +Lemma big_cat_memE (A B : eqType) x xs (f : A -> seq B) : + reflect (exists2 i, i \in xs & x \in f i) + (x \in \big[cat/[::]]_(i <- xs) f i). +Proof. by rewrite big_cat_mem_has; apply: hasP. Qed. + +(* big_cat_mem for A : finType *) +Lemma big_cat_mem (A : finType) (B : eqType) (f : A -> seq B) x : + reflect (exists i, x \in f i) + (x \in \big[cat/[::]]_i f i). +Proof. +case: big_cat_memE=>H; constructor. +- by case: H=>i H1 H2; exists i. +by case=>i X; elim: H; exists i. +Qed. + +(* uniqueness for A : Type *) +Lemma uniq_big_catT A (B : eqType) xs (f : A -> seq B) : + Uniq xs -> + (forall x, x \In xs -> uniq (f x)) -> + (forall x1 x2, x1 \In xs -> x2 \In xs -> + has (mem (f x1)) (f x2) -> x1 = x2) -> + uniq (\big[cat/[::]]_(x <- xs) f x). +Proof. +elim: xs=>[|x xs IH] /=. +- by rewrite big_nil /=; constructor. +case=>Nx U H1 H2; rewrite big_cons cat_uniq. +apply/and3P; split. +- by apply: H1; left. +- rewrite has_big_cat -all_predC; apply/allPIn=>x3 H3 /=. + by apply: contra_notN Nx=>H; rewrite (H2 _ _ _ _ H) //; [left | right]. +apply: IH=>//. +- by move=>z Hz; apply: H1; right. +by move=>z1 z2 Hz1 Hz2 N; apply: H2=>//; right. +Qed. + +Lemma big_cat_uniq_pairwise A (B : eqType) xs (f : A -> seq B) x1 x2 : + uniq (\big[cat/[::]]_(x <- xs) f x) -> + x1 \In xs -> + x2 \In xs -> + has (mem (f x1)) (f x2) -> + x1 = x2. +Proof. +elim: xs=>[|x xs IH] //. +rewrite big_cons cat_uniq; case/and3P=>U N Us. +case/In_cons=>[->|H1]; case/In_cons=>[->|H2] //; last by apply: IH. +- case/hasP=>/= b Hb1 Hb2. + exfalso; move/negP: N; apply; apply/hasP. + by exists b=>//; apply/big_cat_memT; exists x2. +case/hasP=>/= b Hb1 Hb2. +exfalso; move/negP: N; apply. +apply/hasP; exists b=>//; apply/big_cat_memT. +by exists x1. +Qed. + +(* uniqueness for A : eqType *) +Lemma uniq_big_catE (A B : eqType) (f : A -> seq B) (xs : seq A) : + reflect + [/\ forall i, i \in xs -> uniq (f i), + forall i k, i \in xs -> k \in f i -> count_mem i xs = 1 & + forall i j k, i \in xs -> j \in xs -> + k \in f i -> k \in f j -> i = j] + (uniq (\big[cat/[::]]_(i <- xs) f i)). +Proof. +elim: xs=>[|x xs IH] /=; first by rewrite big_nil; constructor. +rewrite big_cons cat_uniq. +case H1 : (uniq (f x))=>/=; last first. +- by constructor; case=>/(_ x); rewrite inE eqxx H1=>/(_ erefl). +case: hasPn=>/= V; last first. +- constructor; case=>H2 H3 H4; elim: V=>z /big_cat_memE [i X Zi]. + apply/negP=>Zx; move/(H3 i z): (Zi); rewrite inE X orbT=>/(_ erefl). + move: (H4 x i z); rewrite !inE eqxx X orbT=>/(_ erefl erefl Zx Zi)=>E. + by rewrite -{i Zi}E eqxx add1n in X *; case=>/count_memPn; rewrite X. +case: IH=>H; constructor; last first. +- case=>H2 H3 H4; apply: H; split; last 1 first. + - by move=>i j k Xi Xj; apply: H4; rewrite inE ?Xi ?Xj orbT. + - by move=>i X; apply: H2; rewrite inE X orbT. + move=>i k X D; move/(H3 i k): (D); rewrite inE X orbT=>/(_ erefl). + case: (x =P i) X D=>[<-{i}|N] X D; last by rewrite add0n. + by rewrite add1n=>[[]] /count_memPn; rewrite X. +case: H=>H2 H3 H4; split; first by move=>i; rewrite inE=>/orP [/eqP ->|/H2]. +- move=>i k; rewrite inE eq_sym; case: (x =P i)=>[<- _|N /= Xi] K; last first. + - by rewrite add0n; apply: H3 Xi K. + rewrite add1n; congr S; apply/count_memPn; apply: contraL (K)=>X. + by apply/V/big_cat_memE; exists x. +move=>i j k; rewrite !inE=>/orP [/eqP ->{i}|Xi] /orP [/eqP ->{j}|Xj] Ki Kj //. +- by suff : k \notin f x; [rewrite Ki | apply/V/big_cat_memE; exists j]. +- by suff : k \notin f x; [rewrite Kj | apply/V/big_cat_memE; exists i]. +by apply: H4 Kj. +Qed. + +(* alternative formulation *) +Lemma uniq_big_catEX (A B : eqType) (f : A -> seq B) (xs : seq A) : + uniq xs -> + reflect + [/\ forall i, i \in xs -> uniq (f i) & + forall i j k, i \in xs -> j \in xs -> + k \in f i -> k \in f j -> i = j] + (uniq (\big[cat/[::]]_(i <- xs) f i)). +Proof. +move=>U; case: uniq_big_catE=>H; constructor; first by case: H. +by case=>H1 H2; elim: H; split=>// i k I K; rewrite count_uniq_mem // I. +Qed. + +(* uniqueness for A : finType *) +Lemma uniq_big_cat (A : finType) (B : eqType) (f : A -> seq B) : + reflect ([/\ forall t, uniq (f t) & + forall t1 t2 k, k \in f t1 -> k \in f t2 -> t1 = t2]) + (uniq (\big[cat/[::]]_tag f tag)). +Proof. +case: uniq_big_catEX=>[|[R1 R2]|R]. +- by rewrite /index_enum -enumT enum_uniq. +- by constructor; split=>[t|t1 t2 k K1 K2]; [apply: R1|apply: R2 K1 K2]. +constructor; case=>R1 R2; elim: R; split=>[t ?|t1 t2 k ??]; +by [apply: R1 | apply: R2]. +Qed. + +Lemma uniq_big_cat_uniqT A (B : eqType) xs (f : A -> seq B) y : + uniq (\big[cat/[::]]_(x <- xs) f x) -> + y \In xs -> + uniq (f y). +Proof. +elim: xs=>[|x xs IH] in y * => //. +rewrite big_cons InE cat_uniq. +case/and3P=>U _ Us; case=>[->|Hy] //. +by apply: IH. +Qed. + +Prenex Implicits uniq_big_cat_uniqT. + +Lemma uniq_big_cat_uniq (A : finType) (f : A -> seq nat) t : + uniq (\big[cat/[::]]_t f t) -> + uniq (f t). +Proof. by move/uniq_big_cat_uniqT; apply; apply/mem_seqP. Qed. + +Lemma uniq_big_cat_uniq0 (A : finType) (f : A -> seq nat) t : + uniq (0 :: \big[cat/[::]]_t f t) -> uniq (0 :: f t). +Proof. +case/andP=>U1 /uniq_big_cat_uniq /= U2. +rewrite U2 andbT (contra _ U1) // => U3. +by apply/big_cat_mem; exists t. +Qed. + +Lemma uniq_big_cat_disj (A : finType) (B : eqType) (f : A -> seq B) t1 t2 x : + uniq (\big[cat/[::]]_t f t) -> + x \in f t1 -> + x \in f t2 -> + t1 = t2. +Proof. by case/uniq_big_cat=>_; apply. Qed. + + diff --git a/core/seqperm.v b/core/seqperm.v new file mode 100644 index 0000000..0f8a4ad --- /dev/null +++ b/core/seqperm.v @@ -0,0 +1,208 @@ +(* +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 pcm Require Import options pred. + +(****************************************************) +(* 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. + +#[export] 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/core/slice.v b/core/slice.v new file mode 100644 index 0000000..6c71b66 --- /dev/null +++ b/core/slice.v @@ -0,0 +1,716 @@ +(* +Copyright 2022 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 path interval order. +From pcm Require Import options prelude seqext. + +Open Scope order_scope. +Import Order.Theory. + +(* DEVCOMMENT *) +(* intervals used to do these simplifications automatically *) +(* but that changed with mathcomp 2.0 *) +(* /DEVCOMMENT *) +Section BSimp_Extension. +Variables (disp : unit) (T : porderType disp). +Implicit Types (x y : T) (b c : bool). + +Lemma binf_inf b c : (BInfty T b == BInfty T c) = (b == c). +Proof. by rewrite eqE /= eqE /= andbT. Qed. + +Lemma bside_inf x b c : + ((BSide b x == BInfty T c) = false) * + ((BInfty T c == BSide b x) = false). +Proof. by rewrite !eqE /= !eqE /= !andbF. Qed. + +Lemma bside_side x y b c : + (BSide b x == BSide c y) = (b == c) && (x == y). +Proof. by rewrite !eqE. Qed. + +End BSimp_Extension. + +Definition bnd_simp := ((@bside_inf, @binf_inf, @bside_side), bnd_simp). + +(* sequence slicing by nat indices *) +(* reuses mathcomp interval infrastructure *) + +(* convert bound to nat *) +(* maps -oo -> 0, +oo -> m *) +Definition bnd (i : itv_bound nat) (m : nat) : nat := + match i with + | BSide b n => n + ~~ b + | BInfty b => if b then 0 else m + end. + +(* slicing is applying an interval to a seq *) +Definition slice {A : Type} (s : seq A) (i : interval nat) := + let sz := size s in + let: Interval l u := i in + drop (bnd l sz) (take (bnd u sz) s). + +Arguments slice {A} s i : simpl never. + +Notation "&: s i" := (slice s i) + (at level 1, i at next level, s at next level, + format "'&:' s i") : fun_scope. + +(* subtract n from bound, convert values past zero to -oo *) +Definition shl_bnd (i : itv_bound nat) (n : nat) := + match i with + | BSide b m => if (m < n)%N then -oo else BSide b (m - n) + | BInfty b => BInfty _ b + end. + +Lemma shl_bnd0 i : shl_bnd i 0%N = i. +Proof. by case: i=>[[] i|[]] //=; rewrite subn0. Qed. + +Lemma shl_bnd_add i n m : shl_bnd (shl_bnd i n) m = shl_bnd i (n+m). +Proof. +case: i=>[[] i|[]] //=. +- case: ltnP=>Hin /=; first by rewrite ltn_addr. + by rewrite ltn_subLR // subnDA. +case: ltnP=>Hin /=; first by rewrite ltn_addr. +by rewrite ltn_subLR // subnDA. +Qed. + +Definition shl_itv (i : interval nat) (n : nat) := + let: Interval l u := i in + Interval (shl_bnd l n) (shl_bnd u n). + +Lemma shl_itv0 i : shl_itv i 0%N = i. +Proof. by case: i=>i j /=; rewrite !shl_bnd0. Qed. + +Lemma shl_itv_add i n m : shl_itv (shl_itv i n) m = shl_itv i (n+m). +Proof. by case: i=>i j /=; rewrite !shl_bnd_add. Qed. + +Lemma mem_shl n m i : + (n \in shl_itv i m) = (m + n \in i). +Proof. +rewrite !in_itv; case: i=>i j /=. +case: i=>[l i|[]]; case: j=>[r j|[]] //=; +rewrite ?andbF ?andbT //. +- case: (ltnP j m)=>/= Hj. + - rewrite andbF; symmetry; apply/negbTE; rewrite negb_and lteifNE negbK -lteifNE. + by apply/orP; right; apply/lteifS/ltn_addr. + have ->: (n < j - m ?<= if ~~ r) = (m + n < j ?<= if ~~ r). + - case: r=>/=; first by rewrite ltEnat /= ltn_subRL. + by rewrite leEnat leq_subRL. + case: (m + n < j ?<= if ~~ r); rewrite ?andbF // !andbT. + case: (ltnP i m)=>/= Hi. + - by symmetry; apply/lteifS/ltn_addr. + case: l=>/=; last by rewrite ltEnat /= ltn_subLR. + by rewrite leEnat leq_subLR. +- case: (ltnP i m)=>/= Hi. + - by symmetry; apply/lteifS/ltn_addr. + case: l=>/=; last by rewrite ltEnat /= ltn_subLR. + by rewrite leEnat leq_subLR. +case: (ltnP j m)=>/= Hj. +- symmetry; apply/negbTE; rewrite lteifNE negbK. + by apply/lteifS/ltn_addr. +case: r=>/=; first by rewrite ltEnat /= ltn_subRL. +by rewrite leEnat leq_subRL. +Qed. + +(* generalize to orderType? *) +Lemma mem_itv_inf (n : nat) : n \in `]-oo,+oo[. +Proof. by rewrite in_itv. Qed. + +Section Lemmas. +Variable (A : Type). +Implicit Type (s : seq A). + +Lemma slice_decompose s l r x y : + &:s (Interval (BSide l x) (BSide r y)) = + &:(&:s (Interval -oo (BSide r y))) (Interval (BSide l x) +oo). +Proof. by rewrite /slice /= drop0 take_size. Qed. + +(* "test" lemmas *) +Lemma slice_ouou s a b : + &:s `]a,b[ = &:(&:s `]-oo,b[) `]a,+oo[. +Proof. by rewrite slice_decompose. Qed. + +(* TODO generalize? *) +Lemma slice_uouo s a b : + &:s `]a,b+a] = &:(&:s `]a,+oo[) `]-oo,b[. +Proof. +by rewrite /slice /= !addn1 !addn0 drop0 take_size take_drop addnS. +Qed. + +(* ... *) + +(* some simplifying equalities on slices *) + +Lemma slice_0L s y : &:s (Interval -oo y) = &:s (Interval (BLeft 0%N) y). +Proof. by rewrite /slice /= addn0. Qed. + +Lemma slice_FR s x : &:s (Interval x +oo) = &:s (Interval x (BLeft (size s))). +Proof. by rewrite /slice /= addn0. Qed. + +Lemma slice_uu s : &:s `]-oo, +oo[ = s. +Proof. by rewrite /slice /= drop0 take_size. Qed. + +Lemma itv_underL s (i j : itv_bound nat) : + bnd i (size s) = 0%N -> + &:s (Interval i j) = &:s (Interval -oo j). +Proof. by move=>Hi; rewrite /slice /= Hi drop0. Qed. + +Lemma itv_underR s (i j : itv_bound nat) : + bnd j (size s) = 0%N -> + &:s (Interval i j) = [::]. +Proof. by move=>Hj; rewrite /slice /= Hj take0. Qed. + +Corollary itv_under0R s (i : itv_bound nat) : + &:s (Interval i (BLeft 0%N)) = [::]. +Proof. by apply: itv_underR. Qed. + +Lemma itv_overL s (i j : itv_bound nat) : + size s <= bnd i (size s) -> + &:s (Interval i j) = [::]. +Proof. +move=>Hi /=; rewrite /slice /=; apply: drop_oversize. +rewrite size_take; case: ifP=>// H. +by apply/ltnW/(leq_trans H). +Qed. + +Lemma itv_overR s (i j : itv_bound nat) : + size s <= bnd j (size s) -> + &:s (Interval i j) = &:s (Interval i +oo). +Proof. by move=>Hj; rewrite /slice /= take_size take_oversize. Qed. + +Lemma itv_minfR s (i : itv_bound nat) : + &:s (Interval i -oo) = [::]. +Proof. by rewrite /slice /= take0. Qed. + +Lemma itv_pinfL s (j : itv_bound nat) : + &:s (Interval +oo j) = [::]. +Proof. +rewrite /slice /=; apply: drop_oversize. +by rewrite size_take_min; apply: geq_minr. +Qed. + +(* TODO unify the next two? *) +Lemma itv_swapped_bnd s (i j : itv_bound nat) : + j <= i -> + &:s (Interval i j) = [::]. +Proof. +move=>H; rewrite /slice; apply: drop_oversize. +move: H; case: i=>[ib ix|[]]; case: j=>[jb jx|[]] //=; +rewrite ?bnd_simp ?take_size ?take0 ?drop_size //= size_take_min. +- rewrite leBSide; case: ib=>/=; case: jb=>/=; rewrite ?addn0 ?addn1=>Hji. + - apply/leq_trans/Hji; exact: geq_minl. + - apply/leq_trans/Hji; exact: geq_minl. + - apply: leq_trans; first by exact: geq_minl. + by apply: ltnW. + by apply: leq_trans; first by exact: geq_minl. +by move=>_; exact: geq_minr. +Qed. + +Lemma itv_swapped s (i j : itv_bound nat) : + bnd j (size s) <= bnd i (size s) -> + &:s (Interval i j) = [::]. +Proof. +move=>H; rewrite /slice; apply: drop_oversize. +rewrite size_take_min; apply/leq_trans/H. +by exact: geq_minl. +Qed. + +Corollary slice_usize s : s = &:s `]-oo, size s[. +Proof. by rewrite itv_overR /=; [rewrite slice_uu | rewrite addn0]. Qed. + +(* slice size *) + +Lemma slice_size s (i : interval nat) : + size (&:s i) = minn (bnd i.2 (size s)) (size s) - bnd i.1 (size s). +Proof. +rewrite /slice; case: i=>[[l i|[]][r j|[]]] //=; rewrite ?take0 ?drop0 ?take_size ?drop_size ?minnn ?min0n ?subn0 //=. +- by rewrite size_drop size_take_min. +- by rewrite size_drop. +- by rewrite size_take_min. +- by rewrite size_drop size_take_min. +by rewrite subnn. +Qed. + +(* splitting slice *) + +Lemma slice_split_bnd s (i : interval nat) (x : itv_bound nat) : + i.1 <= x <= i.2 -> + &:s i = &:s (Interval i.1 x) ++ &:s (Interval x i.2). +Proof. +case: i=>i1 i2 /=. +case: x=>[b x|[]]; rewrite ?bnd_simp /=; first last. +- by move/eqP=>->; rewrite itv_pinfL cats0. +- by rewrite andbT =>/eqP->; rewrite itv_minfR. +case/boolP: (size s <= bnd (BSide b x) (size s)). +- move=>H Hx; rewrite (itv_overL _ H) cats0 (itv_overR _ H). + apply: itv_overR; rewrite /= in H. + case/andP: Hx=>_. + case: i2=>[[] i2|[]] //= Hx; apply: (leq_trans H); case: {H}b Hx; rewrite bnd_simp ?addn0 ?addn1 //. + by move=>Hx; apply/ltnW. +rewrite -ltNge ltEnat /= =>/ltnW/minn_idPl Hx. +case: i1=>[l i|[]]; case: i2=>[r j|[]] //=; +rewrite ?bnd_simp ?andbF ?andbT // ?lteBSide /slice /= ?drop0 ?take_size. +- case/andP=>Hi Hj. + have Hxj : (x + ~~ b <= j + ~~r)%N. + - case: r Hj=>/=; rewrite ?leEnat ?ltEnat; case: {Hx Hi}b=>/=; + rewrite ?addn0 ?addn1 // => Hj; + by apply: ltnW. + have Hxi : (i + ~~ l <= x + ~~ b)%N. + - case: l Hi=>/=; rewrite ?implybT ?implybF /= ?addn0 ?addn1 /= => Hi. + - by apply/(leq_trans Hi)/leq_addr. + by case: {Hx Hj Hxj}b Hi=>/=; rewrite ?leEnat ?ltEnat /= ?addn0 ?addn1. + rewrite -{1}(subnKC Hxj) takeD drop_cat size_take_min Hx take_drop subnK //. + rewrite ltn_neqAle Hxi andbT; case: eqP=>//= ->. + by rewrite subnn drop0 drop_take_id. +- move=>Hi; rewrite -{1}(cat_take_drop (x + ~~ b) s) drop_cat size_take_min Hx. + have Hxi : (i + ~~ l <= x + ~~ b)%N. + - case: l Hi=>/=; rewrite ?implybT ?implybF /= ?addn0 ?addn1 /= => Hi. + - by apply/(leq_trans Hi)/leq_addr. + by case: {Hx}b Hi=>/=; rewrite ?leEnat ?ltEnat /= ?addn0 ?addn1. + rewrite ltn_neqAle Hxi andbT; case: eqP=>//= ->. + by rewrite subnn drop0 drop_take_id. +- move=>Hj. + have Hbj : (x + ~~ b <= j + ~~ r)%N. + - case: r Hj=>/=; rewrite ?leEnat ?ltEnat; case: {Hx}b=>/=; + rewrite ?addn0 ?addn1 // => Hx; + by apply: ltnW. + by rewrite -{1}(subnKC Hbj) takeD take_drop subnK. +by move=>_; rewrite cat_take_drop. +Qed. + +Corollary slice_split s (i : interval nat) b (x : nat) : + x \in i -> + &:s i = &:s (Interval i.1 (BSide b x)) ++ &:s (Interval (BSide b x) i.2). +Proof. +move=>Hx; apply: slice_split_bnd. +case: i Hx=>[[l i|[]][r j|[]]] //=; +rewrite in_itv /= ?andbT ?andbF // ?leBSide; case: b=>/=; +rewrite ?implybF ?implybT //=. +- by case/andP=>->/= /lteifW. +- by case/andP=>/lteifW->. +- by move/lteifW. +by move/lteifW. +Qed. + +(* splitting whole list *) + +Corollary slice_split_full s b (x : nat) : + s = &:s (Interval -oo (BSide b x)) ++ &:s (Interval (BSide b x) +oo). +Proof. by rewrite -[LHS](slice_uu s); apply: slice_split. Qed. + +(* slice extrusion *) + +Lemma slice_extrude s (i : interval nat) : + i.1 < i.2 -> + s = &:s (Interval -oo i.1) ++ &:s i ++ &:s (Interval i.2 +oo). +Proof. +case: i=>[[[] i|[]][[] j|[]]] //=; rewrite bnd_simp=>H; + rewrite ?itv_minfR ?itv_pinfL /= ?cats0. +- by rewrite -{1}(slice_uu s) (slice_split _ true (x:=i)) //= + (slice_split _ true (x:=j) (i:=`[i, +oo[)) //= in_itv /= andbT; apply/ltnW. +- by rewrite -{1}(slice_uu s) (slice_split _ true (x:=i)) //= + (slice_split _ false (x:=j) (i:=`[i, +oo[)) //= in_itv /= andbT. +- by rewrite -{1}(slice_uu s) (slice_split _ true (x:=i)). +- by rewrite -{1}(slice_uu s) (slice_split _ false (x:=i)) //= + (slice_split _ true (x:=j) (i:=`]i, +oo[)) //= in_itv /= andbT. +- by rewrite -{1}(slice_uu s) (slice_split _ false (x:=i)) //= + (slice_split _ false (x:=j) (i:=`]i, +oo[)) //= in_itv /= andbT. +- by rewrite -{1}(slice_uu s) (slice_split _ false (x:=i)). +- by rewrite -{1}(slice_uu s) (slice_split _ true (x:=j)). +- by rewrite -{1}(slice_uu s) (slice_split _ false (x:=j)). +by rewrite slice_uu. +Qed. + +(* "test" lemmas *) +Corollary slice_uxou s i : s = &:s `]-oo, i] ++ &:s `]i, +oo[. +Proof. by rewrite -(slice_split _ _ (i:=`]-oo, +oo[)) // slice_uu. Qed. + +Corollary slice_uoxu s i : s = &:s `]-oo, i[ ++ &:s `[i, +oo[. +Proof. by rewrite -(slice_split _ _ (i:=`]-oo, +oo[)) // slice_uu. Qed. + +Corollary slice_uxxo s a b : + a <= b -> + &:s `]-oo, b] = &:s `]-oo, a] ++ &:s `]a, b]. +Proof. by move=>Hab; rewrite (slice_split _ false (x:=a)). Qed. + +Corollary slice_uxox s a b : + a <= b -> + &:s `]-oo, b] = &:s `]-oo, a[ ++ &:s `[a, b]. +Proof. by move=>Hab; rewrite (slice_split _ true (x:=a)). Qed. + +Corollary slice_uoox (s : seq A) a b : + a < b -> + &:s `]-oo, b[ = &:s `]-oo, a[ ++ &:s `[a, b[. +Proof. by move=>Hab; rewrite (slice_split _ true (x:=a)). Qed. + +(* ... *) + +(* singleton slice *) + +Lemma slice_kk s k l r : + &:s (Interval (BSide l k) (BSide r k)) = + if l && ~~ r then oapp (cons^~ [::]) [::] (onth s k) else [::]. +Proof. +rewrite /slice /=; case: (onth_sizeP s k)=>[|v] H; rewrite H /=. +- move/size_onthPn: H=>H; rewrite if_same. + apply: drop_oversize; rewrite size_take ltnNge. + have ->/= : (size s <= k + ~~ r)%N by apply/leq_trans/leq_addr. + by apply/leq_trans/leq_addr. +move: (onth_size H)=>Hk; case: l; case: r=>/=; +rewrite ?addn0 ?addn1. +- by apply: drop_oversize; rewrite size_take Hk. +- rewrite -addn1 addnC -take_drop. + rewrite (take_nth v); last by rewrite size_drop subn_gt0. + by rewrite take0 /= nth_drop addn0 nth_onth H. +- by apply: drop_oversize; rewrite size_take Hk. +apply: drop_oversize; rewrite size_take; case: ifP=>// /negbT. +by rewrite -leqNgt. +Qed. + +(* "test" lemmas *) +Corollary slice_kkxx s k : + &:s `[k, k] = oapp (cons^~ [::]) [::] (onth s k). +Proof. by rewrite slice_kk. Qed. + +Corollary slice_kkxo s k : &:s `[k,k[ = [::]. +Proof. by rewrite slice_kk. Qed. + +Corollary slice_kkox s k : &:s `]k,k] = [::]. +Proof. by rewrite slice_kk. Qed. + +Corollary slice_kkoo s k : &:s `]k,k[ = [::]. +Proof. by rewrite slice_kk. Qed. + +(* endpoint +1 *) + +Lemma slice_oSL x a s : + &:s (Interval (BLeft x.+1) a) = &:s (Interval (BRight x) a). +Proof. +case/boolP: (x.+1 \in Interval (BRight x) a)=>[H|]. +- rewrite [RHS](slice_split _ true (x:=x.+1)) //=. + suff ->: &:s `]x, x.+1[ = [::] by rewrite cat0s. + by rewrite /slice /= addn0 addn1 drop_take_id. +rewrite in_itv /= negb_and ltEnat /= ltnS leqnn /=. +case: a=>[[] a|[]] //=. +- by rewrite -leNgt leEnat => H; rewrite !itv_swapped //= !addn0 ?addn1 leEnat. +- by rewrite -ltNge ltEnat /= => H; rewrite !itv_swapped //= !addn1 ?addn0 leEnat ltnS. +by move=>_; rewrite !itv_minfR. +Qed. + +Lemma slice_oSR a x s : + &:s (Interval a (BLeft x.+1)) = &:s (Interval a (BRight x)). +Proof. +case/boolP: (x \in Interval a (BLeft x.+1))=>[H|]. +- rewrite (slice_split _ false (x:=x)) //=. + suff ->: &:s `]x, x.+1[ = [::] by rewrite cats0. + by rewrite /slice /= addn0 addn1 drop_take_id. +rewrite in_itv /= negb_and ltEnat /= ltnS leqnn /= orbF. +case: a=>[[] a|[]] //=. +- by rewrite -ltNge ltEnat /= => H; rewrite !itv_swapped //= !addn0 ?addn1 leEnat. +- by rewrite -leNgt leEnat => H; rewrite !itv_swapped //= !addn1 ?addn0 leEnat ltnS. +by move=>_; rewrite !itv_pinfL. +Qed. + +(* endpoint -1 *) + +Corollary slice_oPL a x s : + &:s (Interval (BRight x.-1) a) = if 0 < x + then &:s (Interval (BLeft x) a) + else &:s (Interval (BRight x) a). +Proof. by case: x=>//= n; rewrite slice_oSL. Qed. + +Corollary slice_oPR a x s : + &:s (Interval a (BRight x.-1)) = if 0 < x + then &:s (Interval a (BLeft x)) + else &:s (Interval a (BRight x)). +Proof. by case: x=>//= n; rewrite slice_oSR. Qed. + +(* endpoint split *) + +Lemma slice_xR a x s : + a <= BLeft x -> + &:s (Interval a (BRight x)) = + oapp (rcons (&:s (Interval a (BLeft x)))) + (&:s (Interval a +oo)) + (onth s x). +Proof. +move=>Hax; rewrite (slice_split _ true (x:=x)) /=; last first. +- rewrite in_itv /= lexx andbT. + by case: a Hax=>/=[ax av|ax]; case: ax. +rewrite slice_kk /=; case: (onth_sizeP s x)=>[|v] H; + rewrite H /=; last by rewrite cats1. +rewrite cats0 itv_overR //= addn0. +by apply/size_onthPn. +Qed. + +Lemma slice_xL b x s : + BRight x <= b -> + &:s (Interval (BLeft x) b) = + oapp (cons^~ (&:s (Interval (BRight x) b))) + [::] + (onth s x). +Proof. +move=>Hxb; rewrite (slice_split _ false (x:=x)) /=; last first. +- rewrite in_itv /= lexx /=. + by case: b Hxb=>/=[bx bv|bx]; case: bx. +rewrite slice_kk /=; case: (onth_sizeP s x)=>[|v] H; rewrite H //=. +rewrite itv_overL //= addn1; apply: ltW. +by apply/size_onthPn. +Qed. + +(* slice of empty list *) + +Lemma slice0 i : + &:([::] : seq A) i = [::]. +Proof. +by rewrite /slice /=; case: i=>[[[] av|[]]]; case=>[[] bv|[]]. +Qed. + +(* slice of one-element list *) + +Lemma slice1 (k : A) i : + &:[::k] i = if 0%N \in i then [::k] else [::]. +Proof. +rewrite /slice in_itv /=. +case: i=>[[[] i|[]][[] j|[]]] //=; +rewrite ?drop0 ?addn0 ?addn1 ?andbF ?andbT //=. +- case: j=>[|j]/=; first by rewrite ltxx andbF. + by rewrite andbT; case: i. +- by case: i. +- by case: i. +- by apply: drop_oversize; rewrite size_take /=; case: ifP=>//; case: j. +- by case: j. +by case: j. +Qed. + +(* slice and constructors *) + +Lemma slice_cat s1 s2 i : + &:(s1++s2) i = &:s1 i ++ &:s2 (shl_itv i (size s1)). +Proof. +rewrite /slice; case: i=>[[l i|[]][r j|[]]] /=; +rewrite ?take0 ?drop0 ?take_size ?drop_size // ?size_cat. +- rewrite take_cat; case: ltnP=>Hj. + - have ->/=: (j < size s1)%N by apply/leq_ltn_trans/Hj/leq_addr. + by rewrite take0 /= cats0. + rewrite (take_oversize (s:=s1)) //. + move: Hj; rewrite leq_eqVlt; case/orP=>[/eqP Ej|Hj]. + - rewrite -Ej subnn take0 cats0 /=. + case: r Ej=>/=. + - by rewrite addn0=>->; rewrite ltnn subnn /= take0 /= cats0. + by rewrite addn1=>->; rewrite !ltnS leqnn /= take0 /= cats0. + rewrite (ltnNge j); have H1: (size s1 <= j)%N. + - by case: r Hj=>/=; rewrite ?addn1 // addn0 => /ltnW. + rewrite H1 drop_cat; case: ltnP=>Hi /=. + - have ->/=: (i < size s1)%N by apply/leq_ltn_trans/Hi/leq_addr. + by rewrite drop0 addnBAC. + rewrite (drop_oversize (s:=s1)) //=. + move: Hi; rewrite leq_eqVlt; case/orP=>[/eqP Ei|Hi]. + - rewrite -Ei subnn drop0 /=. + case: l Ei=>/=. + - by rewrite addn0=><-; rewrite ltnn subnn /= drop0 addnBAC. + by rewrite addn1=><-; rewrite leqnn /= drop0 addnBAC. + rewrite (ltnNge i); have H2: (size s1 <= i)%N. + - by case: l Hi=>/=; rewrite ?addn1 // addn0 => /ltnW. + by rewrite H2 /= !addnBAC. +- rewrite drop_cat; case: ltnP=>Hi /=. + - have ->/=: (i < size s1)%N by apply/leq_ltn_trans/Hi/leq_addr. + by rewrite drop0. + rewrite (drop_oversize (s:=s1)) //=. + move: Hi; rewrite leq_eqVlt; case/orP=>[/eqP Ei|Hi]. + - rewrite -Ei subnn drop0 /=. + case: l Ei=>/=. + - by rewrite addn0=><-; rewrite ltnn subnn /= drop0. + by rewrite addn1=><-; rewrite leqnn /= drop0. + rewrite (ltnNge i); have H2: (size s1 <= i)%N. + - by case: l Hi=>/=; rewrite ?addn1 // addn0 => /ltnW. + by rewrite H2 /= addnBAC. +- rewrite take_cat; case: ltnP=>Hj. + - have ->/=: (j < size s1)%N by apply/leq_ltn_trans/Hj/leq_addr. + by rewrite take0 /= cats0. + rewrite (take_oversize (s:=s1)) //. + move: Hj; rewrite leq_eqVlt; case/orP=>[/eqP Ej|Hj]. + - rewrite -Ej subnn take0 cats0 /=. + case: r Ej=>/=. + - by rewrite addn0=>->; rewrite ltnn subnn /= take0 /= cats0. + by rewrite addn1=>->; rewrite !ltnS leqnn /= take0 /= cats0. + rewrite (ltnNge j); have H1: (size s1 <= j)%N. + - by case: r Hj=>/=; rewrite ?addn1 // addn0 => /ltnW. + by rewrite H1 /= addnBAC. +by rewrite !drop_oversize //= size_take_min ?size_cat; apply: geq_minr. +Qed. + +Lemma slice_cat_piecewise s1 s2 i : + &:(s1++s2) i = if size s1 \in i + then &:s1 (Interval i.1 +oo) ++ &:s2 (Interval -oo (shl_bnd i.2 (size s1))) + else if bnd i.2 (size s1 + size s2) <= size s1 + then &:s1 i + else &:s2 (shl_itv i (size s1)). +Proof. +rewrite slice_cat; case: i=>i j /=; case: ifP. +- rewrite in_itv; case/andP=>Hi Hj. + rewrite (itv_overR _ (j:=j)); last first. + - case: j Hj=>[[] j|[]] //=. + - by rewrite addn0; move/ltnW. + by rewrite addn1 leEnat; move/(ltnW (n:=j.+1)). + rewrite (itv_underL (i:=shl_bnd i _)) //. + case: i Hi=>[[] i|[]] //=; last by rewrite ltEnat /=; move=>->. + rewrite leEnat leq_eqVlt; case/orP=>[/eqP->|->] //=. + by rewrite ltnn /= subnn. +rewrite in_itv=>/negbT; rewrite negb_and=>H. +case: ifP=>Hj. +- rewrite (itv_underR (s:=s2)); first by rewrite cats0. + case: {H}j Hj=>[[] j|[]] //=. + - rewrite addn0 leEnat leq_eqVlt; case/orP=>[/eqP->|->] //=. + by rewrite ltnn /= subnn. + - by rewrite addn1 leEnat =>->. + by rewrite leEnat -{2}(addn0 (size s1)) leq_add2l leqn0 => /eqP. +case/orP: H; last first. +- case: j Hj=>[[] j|[]] //=. + - by rewrite addn0 -leNgt=>->. + by rewrite leEnat addn1 -ltnNge =>->. +case: i=>[[] i|[]] //=; last by rewrite !itv_pinfL. +- rewrite leEnat -ltnNge => Hi. + rewrite (ltnNge i) (ltnW Hi) /= itv_overL //= addn0 leEnat. + by apply: ltnW. +rewrite ltEnat /= -leqNgt => Hi. +rewrite ltnNge Hi /= itv_overL //= addn1 leEnat. +by apply: ltnW. +Qed. + +Lemma slice_cons x s i : + &:(x::s) i = if 0%N \in i then x::&:s (shl_itv i 1) else &:s (shl_itv i 1). +Proof. by rewrite -cat1s slice_cat /= slice1; case: ifP. Qed. + +Lemma slice_rcons x s i : + &:(rcons s x) i = if size s \in i then rcons (&:s i) x else &:s i. +Proof. +rewrite -cats1 slice_cat slice1 mem_shl addn0. +by case: ifP=>_; [rewrite cats1 | rewrite cats0]. +Qed. + +(* mask *) + +Lemma slice_mask s i : + &:s i = let x := bnd i.1 (size s) in + let y := bnd i.2 (size s) in + mask (nseq x false ++ nseq (y-x) true) s. +Proof. by rewrite /slice /=; case: i=>i j /=; rewrite drop_take_mask. Qed. + +(* count *) + +Lemma slice_count p s i : + count p (&:s i) = + count (fun j => j \in i) (findall p s). +Proof. +elim: s i=>/= [|x s IH] i; first by rewrite slice0. +rewrite findall_cons slice_cons. +case/boolP: (p x)=>/= Hpx; case/boolP: (0 \in i)=>I0 /=. +- rewrite Hpx !add1n; congr S. + rewrite IH count_map; apply: eq_count=>z /=. + by rewrite mem_shl add1n. +- rewrite IH /= count_map; apply: eq_count=>z /=. + by rewrite mem_shl add1n. +- rewrite (negbTE Hpx) /= IH /= count_map; apply: eq_count=>z /=. + by rewrite mem_shl add1n. +rewrite IH /= count_map; apply: eq_count=>z /=. +by rewrite mem_shl add1n. +Qed. + +(* has *) + +Corollary slice_has p s i : + has p (&:s i) = + has (fun j => j \in i) (findall p s). +Proof. by rewrite !has_count slice_count. Qed. + +End Lemmas. + +(* map *) + +Lemma slice_map {A B} (f : A -> B) i s : + [seq f x | x <- &:s i] = &: [seq f x | x <- s] i. +Proof. by rewrite !slice_mask /= map_mask /= size_map. Qed. + +Section LemmasEq. +Variable (A : eqType). +Implicit Type (s : seq A). + +(* membership *) + +Corollary slice_memE x s i : + x \in &:s i = + has (fun j => j \in i) (indexall x s). +Proof. by rewrite /indexall -has_pred1; apply: slice_has. Qed. + +Corollary slice_memE1 x s i : + (count_mem x s <= 1)%N -> + x \in &:s i = + (x \in s) && (bnd i.1 (size s) <= index x s < bnd i.2 (size s)). +Proof. +move=>H; rewrite slice_memE indexall_count1 //; case: ifP=>//= Hx; rewrite orbF. +by case: i=>/= [[[] i|[]][[] j|[]]]; rewrite in_itv /= ?ltEnat ?leEnat /= + ?ltn0 ?addn0 ?addn1 ?andbT ?andbF // ?(leqNgt (size _)) index_mem Hx //= andbT. +Qed. + +Corollary slice_uniq_memE x s i : + uniq s -> + x \in &:s i = + (x \in s) && (bnd i.1 (size s) <= index x s < bnd i.2 (size s)). +Proof. by move=>U; apply: slice_memE1; rewrite (count_uniq_mem _ U) leq_b1. Qed. + +(* subset *) + +Lemma slice_subset s i1 i2 : + i1 <= i2 -> + {subset (&:s i1) <= &:s i2}. +Proof. +case: i1=>i1 j1; case: i2=>i2 j2. +rewrite subitvE; case/andP=>Hi Hj. +move=>x Hx. +have Hij : i1 <= j1. +- apply: contraLR Hx; rewrite -ltNge=>/ltW Hji. + by rewrite itv_swapped_bnd. +rewrite (@slice_split_bnd _ _ _ i1) /=; last first. +- by rewrite Hi /=; apply/le_trans/Hj. +rewrite mem_cat; apply/orP; right. +rewrite (@slice_split_bnd _ _ _ j1) /=; last first. +- by rewrite Hj andbT. +by rewrite mem_cat Hx. +Qed. + +Corollary slice_subset_full s i : + {subset &:s i <= s}. +Proof. +by rewrite -{2}(slice_uu s); apply/slice_subset/itv_lex1. +Qed. + +(* slicing preserves uniqueness *) + +Lemma slice_uniq s i : + uniq s -> uniq (&:s i). +Proof. +move=>U; rewrite /slice /=; case: i=>l u. +by apply/drop_uniq/take_uniq. +Qed. + +(* slicing preserves sortedness *) + +Lemma slice_sorted r s i : + sorted r s -> sorted r (&:s i). +Proof. +move=>S; rewrite /slice /=; case: i=>l u. +by apply/drop_sorted/take_sorted. +Qed. + +End LemmasEq. diff --git a/core/uconsec.v b/core/uconsec.v new file mode 100644 index 0000000..9da15eb --- /dev/null +++ b/core/uconsec.v @@ -0,0 +1,953 @@ +(* +Copyright 2022 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 path interval order. +From pcm Require Import options prelude ordtype seqext slice useqord uslice. + +Open Scope order_scope. +Import Order.Theory. + +(* We assume the sequences are unique and most lemmas do require this *) +(* condition explicitly. Should it be added to `consec` itself? *) + +(**************************************) +(**************************************) +(* Consecutive elements in a sequence *) +(**************************************) +(**************************************) + +Section ConsecEq. +Variable (A : eqType). +Implicit Type (ks : seq A). + +(* The interval lemmas let us equate evaluations of interval endpoints *) +(* But, when a property we want to prove involves other components *) +(* we need to properly induct over ks. *) +(* We thus first define what it means for elements in ks to be consecutive, *) +(* so that the properties can be used in the proofs of the inductive cases. *) + +(* t2 is consecutive to t1 in ks if it appears after t1 and there are *) +(* no other elements in ks between t1 and t2 *) + +Definition consec ks t1 t2 := + [&& t1 <[ks] t2 & nilp (&=ks `]t1, t2[)]. + +(* several alternative formulations *) +Lemma consecP ks t1 t2 : + reflect (t1 <[ks] t2 /\ ~~ has predT (&=ks `]t1, t2[)) + (consec ks t1 t2). +Proof. +apply: (iffP andP); case=>-> H; split=>//. +- by rewrite -nilp_hasPn. +by rewrite nilp_hasPn. +Qed. + +Lemma consecP_inlt (ks : seq A) t1 t2 : + uniq ks -> + reflect ([/\ t1 \in ks & {in ks, forall z, z <[ks] t2 = z <=[ks] t1}]) + (consec ks t1 t2). +Proof. +move=>U; apply: (iffP (consecP ks t1 t2)); case=>H1 H2; split. +- by apply: slt_memE H1. +- by move=>z Z; apply: (hasNL_oo U H1 H2 Z). +- by rewrite H2 // sle_refl. +by apply: (hasNR_oo U)=>z /H2. +Qed. + +Lemma consecP_ingt (ks : seq A) t1 t2 : + uniq ks -> + reflect ([/\ t1 \in ks & {in ks, forall z, t2 <=[ks] z = t1 <[ks] z}]) + (consec ks t1 t2). +Proof. +move=>U; apply (iffP (consecP_inlt t1 t2 U)); case=>H1 H2; + split=>// z /H2 E. +- by rewrite sleNgt E sltNge. +by rewrite sltNge E sleNgt. +Qed. + +Lemma consec_nil t1 t2 : consec [::] t1 t2 = false. +Proof. by case: consecP_ingt=>//=; case. Qed. + +(* frequent projections *) + +Lemma consec_slt ks t1 t2 : consec ks t1 t2 -> t1 <[ks] t2. +Proof. by case/andP. Qed. + +Lemma consec_sltW ks t1 t2 : consec ks t1 t2 -> t1 <=[ks] t2. +Proof. by move/consec_slt/sltW. Qed. + +Lemma consec_mem ks t1 t2 : consec ks t1 t2 -> t1 \in ks. +Proof. by case/andP=>/slt_memE. Qed. + +Lemma consec_oo ks t1 t2 : consec ks t1 t2 -> &=ks `]t1, t2[ = [::]. +Proof. by case/consecP=>_; rewrite -nilp_hasPn=>/nilP. Qed. + +Lemma consec_in (ks : seq A) t1 t2 : + uniq ks -> + consec ks t1 t2 -> {in ks, forall z, z <[ks] t2 = z <=[ks] t1}. +Proof. by move=>U; case/(consecP_inlt _ _ U). Qed. + +(* and some streamlining *) + +Lemma consec_prev (ks : seq A) x y z : + uniq ks -> + consec ks x y -> z <[ks] y -> z <=[ks] x. +Proof. by move=>U; case/(consecP_inlt _ _ U)=>X E N; rewrite -E // (slt_memE N). Qed. + +Lemma consec_prevN (ks : seq A) x y z : + uniq ks -> + z != x -> consec ks x y -> z <[ks] y -> z <[ks] x. +Proof. +move=>U N C /(consec_prev U C). +by rewrite sle_eqVlt; [rewrite (negbTE N) | rewrite (consec_mem C) orbT]. +Qed. + +Lemma consec_next (ks : seq A) x y z : + uniq ks -> + consec ks x y -> x <[ks] z -> y <=[ks] z. +Proof. +move=>U; case/(consecP_ingt _ _ U)=>X E N; case Dz : (z \in ks); first by rewrite E. +by apply: sle_memI; rewrite Dz. +Qed. + +Lemma consec_nextN (ks : seq A) x y z : + uniq ks -> + y \in ks \/ z \in ks -> + y != z -> consec ks x y -> x <[ks] z -> y <[ks] z. +Proof. +move=>U /orP D N C /(consec_next U C). +by rewrite (sle_eqVlt D) (negbTE N). +Qed. + +(* main splitting properties of consec *) + +(* if t2 \in ks then ks splits *) +Lemma consecP_split (ks : seq A) t1 t2 : + uniq ks -> + t2 \in ks -> + reflect (exists xs1 xs2, ks = xs1++t1::t2::xs2) + (consec ks t1 t2). +Proof. +move=>U T2; apply: (iffP idP). +- case/andP=>+ H; case/slt_splitL=>xs1 [xs2][E] N T1 T2'. + rewrite {ks}E mem_cat /= inE (negbTE T2') eq_sym (negbTE N) /= in U T2 H *. + case/in_split: T2=>ks3 [ks4][E T2'']; rewrite {xs2}E in U H *. + by rewrite eqsl_oo_split // in H; move/nilP: H=>->; exists xs1, ks4. +case=>xs1 [xs2] E; move: U; rewrite {ks}E in T2 *. +rewrite cat_uniq /= inE !negb_or -!andbA (eq_sym t1 t2). +case/and8P=>U1 U2 U3 U4 U5 U6 U7 U8. +by rewrite /consec slt_splitR // eqsl_oo_split_consec=>//; rewrite eq_sym. +Qed. + +(* if t2 \notin ks, then t1 is last *) +Lemma consec_last (ks : seq A) t1 t2 : + uniq ks -> + consec ks t1 t2 -> + t2 \notin ks <-> exists ks', ks = rcons ks' t1. +Proof. +move=>U /andP [T]; rewrite nilp_hasPn => /hasPn H. +case: (lastP ks) U H T=>[|xs x] /= {ks} + H. +- by rewrite slt_nil. +rewrite rcons_uniq slt_rcons mem_rcons inE negb_or !(eq_sym x). +case/andP=>Nx Ux; case: ifP=>X; rewrite ?andbF ?andbT. +- move=>Nt; split=>//; case=>ks' /rcons_inj [??]; subst ks' x. + by rewrite (slt_memE Nt) in Nx. +move/contra: (H x)=>/(_ erefl). +rewrite eqslice_mem_uniq /=; last by rewrite rcons_uniq Nx. +rewrite mem_rcons inE eqxx /= in_itv /= negb_and ltEnat /=. +rewrite -!seqlt_unlock -!sleNgt !sle_rcons (negbTE Nx) X /=. +rewrite eqxx /= orbF andbT (eq_sym x). +case/orP=>[/negbTE->|/eqP->] /=. +- by case/andP=>H1 /eqP->; split=>// _; exists xs. +rewrite eqxx /= orbF => H1; split=>//; case=>ks' /rcons_inj [_ Ex]. +by rewrite Ex H1 in Nx. +Qed. + + +(* restatement using last *) +Lemma consec_lastE ks t1 t2 t3 : + uniq ks -> + consec ks t1 t2 -> + t2 \notin ks <-> t1 = last t3 ks. +Proof. +move=>U C; rewrite (consec_last U C). +split=>[[ks' ->]|E]; first by rewrite last_rcons. +case: (lastP ks) E C=>[-> /andP [] /=|s x]. +- by rewrite slt_nil. +by rewrite last_rcons => ->; exists s. +Qed. + +(* not quite the same lemmas as consec_last, but a useful split *) +Lemma consecP_last ks t1 t2 : + uniq ks -> t2 \notin ks -> + reflect (exists xs, ks = rcons xs t1) + (consec ks t1 t2). +Proof. +move=>U T2; apply: (iffP idP). +- by move/consec_last: T2; apply. +case=>xs E; rewrite E /consec rcons_uniq mem_rcons inE negb_or eq_sym in U T2 *. +case/andP: U T2=>T1 U /andP [N T2]. +rewrite slt_rcons (negbTE T2) (negbTE T1) N eq_refl /= nilp_hasPn. +rewrite -all_predC; apply/allP=>x /=; apply: contraTeq=>_. +rewrite eqslice_mem_uniq; last by rewrite rcons_uniq T1. +rewrite mem_rcons inE in_itv /= ltEnat /= !negb_and negb_or. +rewrite -!seqlt_unlock -!sleNgt !sle_rcons (eq_sym x) eqxx. +rewrite orbF T1 (negbTE T2) (negbTE N) /= andbC orbCA orbb. +case: ifP=>/= _; last by rewrite orbN. +by rewrite orbF sleNgt; apply: contra T1; exact: slt_memE. +Qed. + +Lemma consecP_lastE y ks t1 t2 : + uniq ks -> t2 \notin ks -> + consec ks t1 t2 = (t1 \in ks) && (t1 == last y ks). +Proof. +move=>U T2; case: consecP_last=>//. +- by case=>xs ->; rewrite mem_rcons last_rcons inE eq_refl. +move=>H; apply/esym/negP=>/andP [H1 H2]; elim: H. +by apply/(rcons_lastXP _ y); rewrite H1 H2. +Qed. + +Lemma consecP_nilp_filter ks (p : pred _) t1 t2 : + uniq ks -> + consec (filter p ks) t1 t2 <-> + if p t2 then [/\ t1 <[ks] t2, p t1 & nilp (filter p (&=ks `]t1, t2[))] + else [/\ t1 \in ks, p t1 & nilp (filter p (&=ks `]t1, +oo[))]. +Proof. +move=>U; case: ifP=>P2; split. +- case/consecP=>Cx Nx. + move: (slt_memE Cx); rewrite mem_filter=>/andP [P1 K1]. + rewrite slt_filter ?(P1,P2,orbT) // in Cx. + split=>//; rewrite nilp_hasPn; apply: contra Nx. + case/hasP=>x; rewrite mem_filter; case/andP=>Px. + move/(mem_oo _ _ _ U); case=>N1 N2 Kx _; apply/hasP; exists x=>//. + apply/mem_oo; first by apply: filter_uniq. + by rewrite mem_filter !slt_filter ?(P1,P2,Px,orbT). +- case=>N P1 Nx; apply/consecP; split. + - by rewrite slt_filter ?(P1,P2,orbT). + rewrite nilp_hasPn in Nx; apply: contra Nx. + case/hasP=>x /(mem_oo _ _ _ (filter_uniq _ U)); rewrite mem_filter. + case=>/andP [Px Kx] X1 X2 _. + rewrite !slt_filter ?(P1,P2,Px,orbT) // in X1 X2. + by apply/hasP; exists x=>//; rewrite mem_filter Px /=; apply/mem_oo. +- case/consecP=>Cx Nx. + move: (slt_memE Cx); rewrite mem_filter=>/andP [P1 K1]. + split=>//; rewrite nilp_hasPn; apply: contra Nx. + case/hasP=>x; rewrite mem_filter=>/andP [Px] /(mem_ou _ _ U) [Nx Kx] _. + apply/hasP; exists x=>//; apply/mem_oo; first by apply: filter_uniq. + rewrite mem_filter Px Nx; split=>//. + - by rewrite slt_filter ?(P1,Px,orbT). + by apply: slt_memI; rewrite mem_filter ?Px // P2. +case=>K1 P1 Nx; apply/consecP; split. +- apply: slt_memI; first by rewrite mem_filter P1 K1. + by rewrite mem_filter P2. +rewrite nilp_hasPn in Nx; apply: contra Nx. +case/hasP=>x /(mem_oo _ _ _ (filter_uniq _ U)); rewrite mem_filter. +case=>/andP [Px Kx N1 N2] _. +apply/hasP; exists x=>//; rewrite mem_filter Px /=. +apply/mem_ou=>//; split=>//. +by rewrite slt_filter ?(P1,Px,orbT) in N1. +Qed. + +Lemma consecP_filter ks (p : pred _) t1 t2 : + uniq ks -> + consec (filter p ks) t1 t2 <-> + if p t2 + then [/\ t1 <[ks] t2, p t1 & {in &=ks `]t1, t2[ , forall z, ~~ p z}] + else [/\ t1 \in ks , p t1 & {in &=ks `]t1, +oo[, forall z, ~~ p z}]. +Proof. +move=>U; split=>[|H]. +- by move/(consecP_nilp_filter _ _ _ U); case: ifP=>P [?? /nilp_filter]. +by apply/consecP_nilp_filter; case: ifP H=>P [?? /nilp_filter]. +Qed. + +Lemma olt_consec_prev ks t1 t2 : + uniq ks -> + t1 <[ks] t2 <-> exists t, t1 <=[ks] t /\ consec ks t t2. +Proof. +move=>U; split=>[H|]; last first. +- by case=>t [H1] /consecP [H2 _]; apply: (sle_slt_trans H1 H2). +case/slt_splitL: H U=>ks1 [ks2][-> Nt1t2 N1 N2] /=. +rewrite cat_uniq /= negb_or /= -!andbA. +case/and5P=>Uks1 Nt1ks1 /hasPn Hks2 Nt1ks2 Uks2. +case X : (t2 \in ks2); last first. +- have L : last t1 ks2 \notin ks1. + - move: (mem_last t1 ks2); rewrite inE. + by case/orP=>[/eqP ->//|H]; apply: Hks2. + exists (last t1 ks2); split. + - by rewrite sle_cat (negbTE N1) sleL andbT. + apply/andP; split. + - rewrite slt_cat (negbTE N2) (negbTE L) /= slt_cons (eq_sym t2) Nt1t2 /=. + move: (mem_last t1 ks2); rewrite inE=>/orP [->//|H]. + by rewrite slt_memI ?X ?orbT. + rewrite nilp_hasPn; apply: contra L; case/hasP=>x + _; case/mem_oo. + - rewrite cat_uniq /= negb_or Uks1 Uks2 Nt1ks1 Nt1ks2 /= andbT. + by apply/hasPn. + move=>_; rewrite !slt_cat (negbTE N2) !slt_cons (eq_sym t2) Nt1t2 /=. + case Xks1 : (x \in ks1)=>/=; first by move/slt_memE. + case/orP=>// /andP [Nxt1]; rewrite (negbTE Nxt1) /=. + case/orP=>[/eqP/last_nochange|/[swap] Xp1]. + - by rewrite (negbTE Nt1ks2) /==>/eqP ->; rewrite slt_nil. + move: (@sle_last _ x t1 ks2)=>/(_ Uks2 (slt_memE Xp1)) Z. + by move/(sle_slt_trans Z); rewrite slt_irr. +case/splitP: {ks2} X Hks2 Nt1ks2 Uks2=>p1 p2 H2. +rewrite mem_cat cat_uniq /= negb_or rcons_uniq mem_rcons inE. +rewrite (negbTE Nt1t2) /= -!andbA. +case/andP=>Nt1p1 Nt1p2 /and4P [Nt2p1 Up1 /hasPn Hp2 Up2]. +have L : last t1 p1 \notin ks1. +- move: (mem_last t1 p1); rewrite inE. + case/orP=>[/eqP ->//|H]; apply: H2. + by rewrite mem_cat mem_rcons inE H orbT. +exists (last t1 p1); split. +- by rewrite sle_cat (negbTE Nt1ks1) sleL andbT. +apply/andP; split. +- rewrite slt_cat (negbTE N2) (negbTE L) /= slt_cons (eq_sym t2) Nt1t2 /=. + rewrite slt_cat mem_rcons inE eq_refl /= slt_rcons (negbTE Nt2p1) eq_refl /=. + by move: (mem_last t1 p1); rewrite inE=>/orP [->|->] //=; rewrite orbT. +rewrite nilp_hasPn; apply: contra L; case/hasP=>x. +case/mem_oo. +- rewrite cat_uniq /= cat_uniq mem_cat !negb_or mem_rcons rcons_uniq + negb_or Uks1 N1 Nt1t2 Nt1p1 Nt1p2 Nt2p1 Up1 Up2 /= andbT. + by apply/andP; split; apply/hasPn. +move=>_; rewrite !slt_cat !slt_cons !slt_cat !mem_rcons !inE (negbTE N2) + (eq_sym t2) Nt1t2 /= eqxx /=. +case Xks1 : (x \in ks1)=>/=; first by move/slt_memE. +case/orP=>// /andP [Nxt1]; rewrite (negbTE Nxt1) /=. +case/orP=>[/eqP/last_nochange|/[swap]]. +- rewrite (negbTE Nt1p1)=>/eqP ->/=. + by rewrite slt_cons eqxx. +rewrite slt_rcons (negbTE Nt2p1) eqxx /= orbF => Xp1. +rewrite Xp1 orbT slt_rcons Xp1 sltNge. +by move: (@sle_last _ x t1 p1)=>/(_ Up1 Xp1) ->. +Qed. + +Lemma olt_consec_next ks t1 t2 : + uniq ks -> + t1 <[ks] t2 <-> exists t, consec ks t1 t /\ t <=[ks] t2. +Proof. +move=>U; split=>[H|]; last first. +- by case=>t [/consecP [X _] /(slt_sle_trans X)]. +case/slt_splitL: H U=>ks1 [ks2][-> Nt1t2 N1 N2] /=. +rewrite cat_uniq /= negb_or -!andbA. +case/and5P=>Uks1 _ /hasPn Nks2 Nt1ks2 Uks2. +have H : head t2 ks2 \notin ks1. +- move: (mem_head t2 ks2); rewrite inE. + by case/orP=>[/eqP ->//|]; apply: Nks2. +exists (head t2 ks2); split; last first. +- rewrite sle_cat (negbTE H) N2 /= sle_cons (eq_sym t2) Nt1t2 /=. + by rewrite sle_head orbT. +apply/andP; split. +- rewrite slt_cat (negbTE H) (negbTE N1) /= sltL. + case: eqP Nt1ks2 (mem_head t2 ks2)=>// -> X. + by rewrite inE (negbTE Nt1t2) (negbTE X). +rewrite nilp_hasPn; apply: contra H. +case/hasP=>x; case/mem_oo. +- rewrite cat_uniq /= negb_or Uks1 N1 Nt1ks2 Uks2 /= andbT. + by apply/hasPn. +move=>_; rewrite !slt_cat /= !slt_cons eqxx /= andbT (negbTE N1) /=. +case Xks1 : (x \in ks1)=>/=. +- by move/slt_memE=>E; rewrite E in N1. +move=>Nxt1; rewrite (negbTE Nxt1) /=; case: ifP=>// _. +case/andP=>_ /(sle_slt_trans (@sle_head _ t2 ks2 x))=>+ _. +by rewrite slt_irr. +Qed. + +(* previous element is uniquely determined *) +Lemma consec_prev_inj ks t t1 t2 : + uniq ks -> + consec ks t1 t -> + consec ks t2 t -> + t1 = t2. +Proof. +move=>U /andP [T1 +] /andP [T2]; rewrite !nilp_hasPn => /hasPn H1 /hasPn H2. +move: (@slt_total _ ks t1 t2 (slt_memE T1)). +case/or3P=>[/eqP->|L1|L2] //. +- move: (H1 t2)=>/contraL/(_ erefl); apply: contraNeq=>_. + by apply/mem_oo=>//; split=>//; apply: slt_memE T2. +move: (H2 t1)=>/contraL/(_ erefl); apply: contraNeq=>_. +by apply/mem_oo=>//; split=>//; apply: slt_memE T1. +Qed. + +(* next of a non-last element is uniquely determined *) +Lemma consec_next_inj_nonlast ks t t1 t2 t3 : + uniq ks -> + t != last t3 ks -> + consec ks t t1 -> + consec ks t t2 -> t1 = t2. +Proof. +move=>U N C1 C2. +have K1 : t1 \in ks by apply: contraR N=>/(consec_lastE t3 U C1) ->. +have K2 : t2 \in ks by apply: contraR N=>/(consec_lastE t3 U C2) ->. +case/andP: C1 C2=>T1 + /andP [T2]. +rewrite !nilp_hasPn => /hasPn H1 /hasPn H2. +move: (@slt_total _ ks t1 t2 K1); case/or3P=>[/eqP->|L1|L2] //. +- move: (H2 t1)=>/contraL/(_ erefl); apply: contraNeq=>_. + by apply/mem_oo. +move: (H1 t2)=>/contraL/(_ erefl); apply: contraNeq=>_. +by apply/mem_oo. +Qed. + +(* a restatement in a more useful form *) +Lemma consec_next_inj ks t t1 t2 : + uniq ks -> + t1 \in ks -> + consec ks t t1 -> + consec ks t t2 -> t1 = t2. +Proof. +move=>U T C1 C2; suff N : t != last t1 ks. +- by apply: consec_next_inj_nonlast U N C1 C2. +by apply: contraL T=>/eqP /(consec_lastE t1 U C1). +Qed. + +(* consecutiveness and sortedness under general relation *) + +Lemma consec_sorted_lt ltT ks t1 t2 : + uniq ks -> + irreflexive ltT -> + antisymmetric ltT -> + transitive ltT -> + sorted ltT ks -> + t2 \in ks -> + consec ks t1 t2 -> + {in ks, forall z, ltT z t2 = (z == t1) || ltT z t1}. +Proof. +move=>U I Asym T S T2 C; move: (consec_mem C)=>T1. +have {}Asym : antisymmetric (fun x y => (x == y) || ltT x y). +- move=>x y; rewrite (eq_sym y); case: eqP=>//= _. + by apply: (Asym x y). +have {}T : transitive (fun x y => (x == y) || ltT x y). +- move=>x y z; case: eqP=>[->|/eqP _] //=. + case: eqP=>[->->|/eqP _ /=]; first by rewrite orbT. + by case: eqP=>//= _; apply: T. +have {}S : sorted (fun x y => (x == y) || ltT x y) ks. +- by apply: sub_sorted S=>x y ->; rewrite orbT. +move=>z Z; move/(consec_in U)/(_ z Z): C. +rewrite (slt_sorted_leE Asym T S) //. +rewrite (sle_sorted_leE Asym T S) //. +by rewrite !orbA orbb; case: eqP=>//= ->; rewrite I. +Qed. + +Lemma consec_sorted_le (leT : rel A) ks t1 t2 : + uniq ks -> + {in ks, reflexive leT} -> + antisymmetric leT -> + transitive leT -> + sorted leT ks -> + t2 \in ks -> + consec ks t1 t2 -> + {in ks, forall z, leT z t1 = (z != t2) && leT z t2}. +Proof. +move=>U R Asym T S T2 C; move: (consec_mem C)=>T1. +move=>z Z; move/(consec_in U)/(_ z Z): C. +rewrite (slt_sorted_leE Asym T S) //. +rewrite (sle_sorted_leE Asym T S) //. +by move=>->; case: eqP=>// ->; rewrite R. +Qed. + +(* alternative formulation where we sort ks in consec *) +(* this form is required in some proofs for linearizability *) +Lemma consec_sort_lt ltT ks t1 t2 : + irreflexive ltT -> + antisymmetric ltT -> + transitive ltT -> + {in ks &, total (fun x y => (x == y) || ltT x y)} -> + uniq ks -> + t2 \in ks -> + consec (sort (fun x y => (x == y) || ltT x y) ks) t1 t2 -> + {in ks, forall z, ltT z t2 = (z == t1) || ltT z t1}. +Proof. +set ks' := sort _ _=>I asym T Tot U T2 C z Z. +apply: (@consec_sorted_lt ltT ks')=>//. +- by rewrite sort_uniq. +- by apply: sort_sorted_in_lt. +- by rewrite mem_sort. +by rewrite mem_sort. +Qed. + +Lemma consec_sort_le leT ks t1 t2 : + uniq ks -> + {in ks, reflexive leT} -> + antisymmetric leT -> + transitive leT -> + {in ks &, total leT} -> + t2 \in ks -> + consec (sort leT ks) t1 t2 -> + {in ks, forall z, leT z t1 = (z != t2) && leT z t2}. +Proof. +set ks' := sort _ _=>U R Asym T Tot T2 C z Z. +apply: (@consec_sorted_le leT ks')=>//. +- by rewrite sort_uniq. +- by move=>x; rewrite mem_sort; apply: R. +- by apply: sort_sorted_in Tot _ _. +- by rewrite mem_sort. +by rewrite mem_sort. +Qed. + +(*******************************) +(* consec and cons constructor *) +(*******************************) + +Lemma consec_hdswap k1 k2 ks x : + uniq ks -> + k1 \notin ks -> k2 \notin ks -> + x != k2 -> + consec (k1::ks) k1 x -> consec (k2::ks) k2 x. +Proof. +move=>U K1 K2 N2 C. +have N1 : x != k1 by move/consec_slt: C; rewrite sltL. +move: C; rewrite /consec !sltL N1 N2 /= !nilp_hasPn. +apply: contra=>/hasP [z Z] _; apply/hasP; exists z=>//. +case/mem_oo: Z; first by rewrite /= K2. +rewrite inE sltL slt_cons=>Z1 Nz /andP [_ Z2]. +rewrite (negbTE Nz) /= in Z1 Z2. +apply/mem_oo; first by rewrite /= K1. +rewrite inE sltL slt_cons Z1 N1 Z2 /= orbT; split=>//. +by apply: contra K1 => /eqP <-. +Qed. + +Lemma consec_hd2 k1 k2 ks : + uniq ks -> + k1 \notin ks -> k2 \notin ks -> + k1 != k2 -> consec [:: k1, k2 & ks] k1 k2. +Proof. +move=>U K1 K2 N; rewrite /consec !sltL eq_sym N /= nilp_hasPn. +apply/hasPn=>z; apply: contraL=>/= _. +rewrite eqslice_mem_uniq; last first. +- by rewrite /= inE negb_or N K1 K2. +rewrite negb_and; apply/orP; right. +by rewrite in_itv /= (negbTE N) !eqxx /= negb_and ltEnat /= -!leqNgt leqn0 lt0n orbN. +Qed. + +(* a useful lemma that collects the case analysis *) +(* for consec (k::ks) *) +Lemma consec_consEP' k k1 k2 ks : + uniq ks -> + k \notin ks -> + consec (k::ks) k1 k2 <-> + if k1 == k then + if k2 \in ks then ks = k2 :: behead ks + else k2 != k /\ ks = [::] + else k2 != k /\ consec ks k1 k2. +Proof. +move=>U N; split; last first. +- case: (k1 =P k)=>[->{k1}|/eqP Nk1k [Nk2]]; last first. + - case/consecP=>H1; move=> H2; apply/consecP. + rewrite slt_cons (negbTE Nk1k) /= H1 andbT; split=>//. + apply: contra H2=>/hasP [x H2 _]; apply/hasP; exists x=>//. + move: H2; case/mem_oo; first by rewrite /= N. + rewrite !slt_cons inE (negbTE Nk1k) Nk2 /=; case: eqVneq=>//= _ Hx Hx1 Hx2. + by apply/mem_oo. + case K2 : (k2 \in ks). + - move=>E; rewrite {K2}E /= in N U *. + rewrite inE negb_or in N. + case/andP: U=>U1 U2; case/andP: N=>N1 N2. + by apply: consec_hd2. + case=>Nk2k E; rewrite {ks N U K2}E. + apply/consecP; rewrite sltL; split=>//. + apply/hasPn=>x; apply: contraL=>_; apply: contra Nk2k. + by case/mem_oo=>//; rewrite inE sltL => /eqP->; rewrite eqxx. +case: (k1 =P k)=>[->|/eqP Nk1k]; last first. +- move/consecP; rewrite slt_cons (negbTE Nk1k) /=. + case=>/andP [Nk2k H1 H2]; split=>//; apply/consecP; split=>//. + apply: contra H2=>/hasP [x] H _. + apply/hasP; exists x=>//. + case/mem_oo: H=>// Hx Hx1 Hx2. + apply/mem_oo=>/=; first by rewrite N. + rewrite inE !slt_cons (negbTE Nk1k) (negbTE Nk2k) /=. + by case: eqVneq=>[E|_] //=; rewrite -E Hx in N. +case K2: (k2 \in ks)=>/consecP []; rewrite sltL => Nk2k; last first. +- move=>H; split=>//; apply/nilP; rewrite nilp_hasPn. + apply: contra H=>/hasP [x X _]; apply/hasP; exists x=>//. + apply/mem_oo; first by rewrite /= N. + rewrite inE sltL slt_cons X Nk2k /= orbT. + case: (x =P k)=>//= [E|_]; first by rewrite -E X in N. + by split=>//; apply: slt_memI=>//; rewrite K2. +case: ks U N K2=>//= x ks; rewrite !inE negb_or. +case/andP=>Nkxs U /andP [Nkx Nks] K2 H; congr (_ :: _). +case/orP: K2=>[/eqP->|K2] //. +apply: contraNeq H=>?; apply/hasP; exists x=>//. +apply/mem_oo; first by rewrite /= inE negb_or Nkx Nks Nkxs. +rewrite sltL slt_cons sltL !inE eqxx /= orbT Nk2k /= (eq_sym x) (negbTE Nkx) /=. +by split=>//; rewrite eq_sym. +Qed. + +Lemma consec_consEP k k1 k2 ks : + uniq ks -> + k \notin ks -> + consec (k::ks) k1 k2 <-> + if k1 == k then + if k2 \in ks then k2 = head k ks + else k2 != k /\ ks = [::] + else k2 != k /\ consec ks k1 k2. +Proof. +move=>U /(consec_consEP' _ _ U)=>->. +case: ifP=>// E1; case: ifP=>//. +by case: {U}ks=>//= x ks E2; split=>[[]|->]. +Qed. + +(* let's make them boolean *) +Lemma consec_consE' k k1 k2 ks : + uniq ks -> + k \notin ks -> + consec (k::ks) k1 k2 = + if k1 == k then + if k2 \in ks then ks == k2 :: behead ks + else (k2 != k) && (ks == [::]) + else (k2 != k) && consec ks k1 k2. +Proof. +move=>U K; rewrite iffE consec_consEP' //. +case: ifP=>H1; last by case: andP. +case: ifP=>H2; first by case: eqP. +by split; [case=>->->|case/andP=>-> /eqP]. +Qed. + +Lemma consec_consE k k1 k2 ks : + uniq ks -> + k \notin ks -> + consec (k::ks) k1 k2 = + if k1 == k then + if k2 \in ks then k2 == head k ks + else (k2 != k) && (ks == [::]) + else (k2 != k) && consec ks k1 k2. +Proof. +move=>U K; rewrite iffE consec_consEP //. +case: ifP=>H1; last by case: andP. +case: ifP=>H2; first by case: eqP. +by split; [case=>->->|case/andP=>-> /eqP]. +Qed. + +(* for rcons, we need a uniqueness condition *) +Lemma consec_rconsEP' k k1 k2 ks : + uniq ks -> k \notin ks -> + consec (rcons ks k) k1 k2 <-> + if k1 != k then + if k2 \notin ks then k2 = k /\ exists ks', ks = rcons ks' k1 + else consec ks k1 k2 + else k2 \notin rcons ks k. +Proof. +move=>U K; split; last first. +- case: eqP=>[-> /= K2|/eqP N /=]. + - by apply/consecP_last=>//; [rewrite rcons_uniq K U | exists ks]. + case: ifP=>K2; last first. + - move/negbT: K2; rewrite negbK=>K2 /(consecP_inlt _ _ U) [K1 E]. + apply/consecP_inlt; first by rewrite rcons_uniq K. + split; first by rewrite mem_rcons inE K1 orbT. + move=>z; rewrite mem_rcons inE; case/orP=>[/eqP ->{z}|Z]. + - rewrite slt_rcons K2 sle_rcons (negbTE K) eq_refl andbT K1 /=. + by apply/negP=>/slt_memE; rewrite (negbTE K). + by rewrite slt_rcons K2 sle_rcons Z; apply: E. + case=>->{k2 K2} [ks' E]; rewrite {ks}E in U K *. + apply/consecP_inlt; first by rewrite rcons_uniq K. + split. + - by rewrite mem_rcons inE mem_rcons inE eq_refl orbT. + rewrite rcons_uniq in U; case/andP: U=>N1 U. + rewrite mem_rcons inE negb_or eq_sym N /= in K. + move=>z; rewrite mem_rcons inE mem_rcons inE => Z. + rewrite slt_rcons mem_rcons inE eq_sym (negbTE N) (negbTE K) /=. + rewrite eq_refl orbF !sle_rcons N1 /= eq_refl orbF !mem_rcons !inE eq_refl /=. + case: eqP Z=>[->{z} /= _|/eqP Nz /=]. + - by rewrite eq_sym (negbTE N) /=; case: ifP=>// K'; rewrite K' sle_memI. + rewrite (eq_sym z); case: eqP=>[/= ->|/=]; first by rewrite sle_refl if_same. + by move=>_ Z; rewrite Z sle_memI. +case/consecP_inlt; first by rewrite rcons_uniq K. +move=>+ E; rewrite mem_rcons inE=>K1. +case/orP: K1 E=>[/eqP ->{k1}|K1] E. +- rewrite eq_refl /=. + move: (E k); rewrite mem_rcons inE eq_refl=>/(_ erefl). + rewrite sle_refl; apply: contraL; rewrite mem_rcons inE. + case/orP=>[/eqP ->|K2]; first by rewrite slt_irr. + rewrite slt_rcons K2; apply/negP. + by move/(slt_trans (slt_memI K2 K)); rewrite slt_irr. +case: eqP K1 K=>[->->//|/eqP N K1 K /=]; case: ifP=>K2; last first. +- move/negbT: K2; rewrite negbK=>K2; apply/consecP_inlt=>//. + split=>// z Z; move: (E z); rewrite mem_rcons inE Z orbT=>/(_ erefl). + by rewrite slt_rcons K2 sle_rcons Z. +move: (E k); rewrite mem_rcons inE eq_refl=>/(_ erefl). +rewrite slt_rcons (negbTE K2) (negbTE K) /= eq_refl andbT. +rewrite sle_rcons (negbTE K) K1 /=; case: eqP=>// -> _; split=>//. +suff -> : k1 = last k1 ks. +- move: ks {E N U K K2} k1 K1; apply: last_ind=>[//|ks x IH] k1. + rewrite mem_rcons inE; case/orP=>[/eqP ->|]. + - by rewrite last_rcons; exists ks. + by case/IH=>ks' E; rewrite last_rcons; exists ks. +apply/eqP/contraT; rewrite eq_sym=>M; exfalso. +move: (last_change M)=>L. +move: (E (last k1 ks)); rewrite mem_rcons inE L orbT=>/(_ erefl). +rewrite slt_rcons sle_rcons (negbTE K2) L /=. +move/esym; rewrite sle_eqVlt; last by rewrite L. +rewrite (negbTE M) /=. +by move/(sle_slt_trans (sle_last k1 U K1)); rewrite slt_irr. +Qed. + +Lemma consec_rconsEP k k1 k2 ks : + uniq ks -> k \notin ks -> + consec (rcons ks k) k1 k2 <-> + if k1 != k then + if k2 \notin ks then k2 = k /\ k1 = last k ks + else consec ks k1 k2 + else k2 \notin rcons ks k. +Proof. +move=>U K; rewrite consec_rconsEP' //. +case: eqP=>//= /eqP N1; case: ifP=>// N2. +split; case=>->; first by case=>ks' ->; rewrite last_rcons. +move/esym=>E; split=>//; rewrite -E. +have : last k ks != k by rewrite E. +move/last_change=>{N1 N2}E. +move: ks k U K E; apply: last_ind=>// xs x IH k U K E. +by rewrite last_rcons; exists xs. +Qed. + +(* boolean version *) +Lemma consec_rconsE' y k k1 k2 ks : + uniq ks -> k \notin ks -> + consec (rcons ks k) k1 k2 = + if k1 != k then + if k2 \notin ks then [&& k2 == k, k1 == last y ks & k1 \in ks] + else consec ks k1 k2 + else k2 \notin rcons ks k. +Proof. +move=>U K; rewrite iffE consec_rconsEP' //. +case: ifP=>_ //; case: ifP=>_ //; split. +- by case=>X1 /(rcons_lastXP _ y); rewrite X1 eq_refl. +case/and3P=>/eqP X1 X2 X3; split=>//. +by apply/(rcons_lastXP _ y); rewrite X2 X3. +Qed. + +(* This one is the same, except it drops k1 \in ks condition *) +(* and replaces y by k. The simplifications may be desirable *) +(* depending on the direction of rewriting *) +Lemma consec_rconsE k k1 k2 ks : + uniq ks -> k \notin ks -> + consec (rcons ks k) k1 k2 = + if k1 != k then + if k2 \notin ks then (k2 == k) && (k1 == last k ks) + else consec ks k1 k2 + else k2 \notin rcons ks k. +Proof. +move=>U K; rewrite iffE consec_rconsEP //. +case: ifP=>H1 //; case: ifP=>H2 //. +by split; [case=>->->; rewrite !eq_refl|case/andP=>/eqP -> /eqP ->]. +Qed. + +Lemma consec_behead k ks x y : + uniq ks -> + k \notin ks -> x != k -> + consec (k::ks) x y -> y != k /\ consec ks x y. +Proof. by move=>U K Nx /(consec_consEP _ _ U K); rewrite (negbTE Nx). Qed. + +(* the following isn't a consequence of consec_consE *) +(* as it's independent of k \notin ks *) +(* TODO not anymore *) +Lemma consec_cons k ks x y : + uniq ks -> k \notin ks -> + x != k -> y != k -> consec ks x y -> consec (k::ks) x y. +Proof. +move=>U N Nx Ny; rewrite /consec slt_cons Ny (negbTE Nx) /=. +case/andP=>->/=; rewrite !nilp_hasPn; apply: contra. +case/hasP=>z Z _; apply/hasP; exists z=>//. +case/mem_oo: Z; first by rewrite /= N. +rewrite inE !slt_cons (negbTE Nx) Ny !(eq_sym z) /=. +case: eqVneq=>//= _ Hz Hxz Hzy. +by apply/mem_oo. +Qed. + +(* Pairs of consecutive elements are totally ordered. *) +(* That is: either the two pairs are the same pair, *) +(* or one of the second projections is ordered before the *) +(* first projection of the other pair. *) +Lemma consec2_total ks x1 x2 y1 y2 : + uniq ks -> + y1 \in ks \/ y2 \in ks -> + consec ks x1 y1 -> consec ks x2 y2 -> + [|| (x1 == x2) && (y1 == y2), y2 <=[ks] x1 | y1 <=[ks] x2]. +Proof. +move=>U. +suff L a1 a2 b1 b2 : b1 \in ks -> + consec ks a1 b1 -> consec ks a2 b2 -> + [|| (a1 == a2) && (b1 == b2), b2 <=[ks] a1 | b1 <=[ks] a2]. +- case=>Y Cx1 Cx2; first by apply: L. + case/or3P: (L _ _ _ _ Y Cx2 Cx1)=>[|->|->]; try by rewrite !orbT. + by case/andP=>/eqP/esym -> /eqP/esym ->; rewrite !eq_refl. +move=>Y1 Cx1 Cx2; move: (consec_mem Cx1) (consec_mem Cx2)=>X1 X2. +case/or3P: (slt_total a2 X1) Cx2 X2=>[/eqP <-{a2}|H|H] Cx2 X2. +- by rewrite (consec_next_inj U Y1 Cx1 Cx2) !eq_refl. +- by rewrite (consec_next U Cx1 H) !orbT. +by rewrite (consec_next U Cx2 H) !orbT. +Qed. + +(***************************************) +(* Consecutiveness induction principle *) +(***************************************) + +Lemma consec_ind_raw k ks (P : A -> Prop) : + k \notin ks -> uniq ks -> + P k -> + (forall t1 t2, t2 \in ks -> consec (k::ks) t1 t2 -> P t1 -> P t2) -> + forall t, t \in k::ks -> P t. +Proof. +elim: ks=>[|x ks IH] //= K U H0 Hstep t; first by rewrite inE=>/eqP ->. +rewrite inE negb_or in K U; case/andP: K U=>K1 K2 /andP [U1 U2]. +rewrite !inE; case/or3P; first by move/eqP=>->. +- move/eqP=>->{t}; apply: Hstep H0; first by rewrite inE eq_refl. + by rewrite /consec sltL eq_sym K1 (eqsl_oo_split_consec (s1:=[::])). +have U : uniq (k::x::ks) by rewrite /= inE negb_or K1 K2 U1 U2. +move=>T; apply: IH=>[|||t1 t2 T2 C|] //; last by rewrite inE T orbT. +(* useful sides *) +have Nx2 : t2 != x by case: eqP T2 U1=>// ->->. +have Nk2 : t2 != k by case: eqP T2 K2=>// ->->. +(* main proof *) +case: (t1 =P k) C=>[->{t1} C _|/eqP Nt1k C]. +- (* in this case, we need two induction steps *) + suff [Ckx Cxt2] : consec [:: k, x & ks] k x /\ consec [:: k, x & ks] x t2. + - have : P x by apply: Hstep Ckx H0; rewrite inE eq_refl. + by apply: Hstep Cxt2; rewrite inE T2 orbT. + split; first by apply: consec_hd2. + apply: consec_cons=>//=. + - by rewrite U1. + - by rewrite inE negb_or K1. + - by rewrite eq_sym. + by apply: consec_hdswap C. +(* another useful side which holds only if k != t1 *) +have Nx1 : t1 != x. +- case: eqP U1=>// <-; move: (consec_mem C). + by rewrite inE (negbTE Nt1k) /= =>->. +(* then the proof is straightforward *) +apply: Hstep; first by rewrite inE T2 orbT. +apply: consec_cons=>//=. +- by rewrite U1. +- by rewrite inE negb_or K1. +by apply: consec_cons=>//; case/consec_behead: C. +Qed. + +(* somewhat modified variant of consec_ind_raw *) +(* where we supply the starting k by hand *) +Lemma consec_ind k ks (P : A -> Prop) : + uniq (k :: ks) -> + P k -> + (forall t1 t2, t2 \in ks -> consec (k::ks) t1 t2 -> P t1 -> P t2) -> + forall t, t \in ks -> P t. +Proof. +move=>/= /andP [U1 U2] P0 IH t D; apply: consec_ind_raw IH _ _=>//. +by rewrite inE D orbT. +Qed. + +(* a version that deconstructs ks to find the starting point *) +(* and gives us useful (though derivable) membership predicates t1 \in ks *) +Lemma consec_indX (ks : seq A) (P : A -> Prop) : + uniq ks -> + (forall t1, t1 \in ks -> ks = t1 :: behead ks -> P t1) -> + (forall t1 t2, t1 \in ks -> t2 \in ks -> consec ks t1 t2 -> + P t1 -> P t2) -> + forall t, t \in ks -> P t. +Proof. +case: ks=>//= k ks /andP [K U] H1 H2; apply: consec_ind_raw=>//. +- by apply: (H1 k)=>//; rewrite inE eq_refl. +move=>t1 t2 N Cx; apply: H2=>//. +- by rewrite (consec_mem Cx). +by rewrite inE N orbT. +Qed. + +(* special variants when we induct over an interval *) +(* that is open/closed/unbounded on the right *) +Lemma consec_indo ks k1 k2 (P : A -> Prop) : + uniq ks -> k1 <[ks] k2 -> + P k1 -> + (forall t1 t2, + t2 \in &=ks `]k1, k2[ -> + consec (&=ks `[k1, k2[) t1 t2 -> P t1 -> P t2) -> + forall t, t \in &=ks `]k1,k2[ -> P t. +Proof. +move=>U N H0 IH t; apply: (consec_ind (k:=k1))=>//=. +- by rewrite eqslice_uniq // andbT eqslice_mem_uniq // negb_and in_itv /= ltxx /= orbT. +by move=>t1 t2 H C; apply: IH=>//; rewrite eqsl_xoL N. +Qed. + +Lemma consec_indx ks k1 k2 (P : A -> Prop) : + uniq ks -> k1 <=[ks] k2 -> k1 \in ks -> + P k1 -> + (forall t1 t2, + t2 \in &=ks `]k1, k2] -> + consec (&=ks `[k1, k2]) t1 t2 -> P t1 -> P t2) -> + forall t, t \in &=ks `]k1, k2] -> P t. +Proof. +move=>U N K H0 IH t; apply: (consec_ind (k:=k1))=>//=. + - by rewrite eqslice_uniq // andbT eqslice_mem_uniq //= negb_and in_itv /= ltxx /= orbT. +by move=>t1 t2 H C; apply: IH=>//; rewrite eqsl_xxL N /= K. +Qed. + +Lemma consec_indu ks k (P : A -> Prop) : + uniq ks -> k \in ks -> + P k -> + (forall t1 t2, + t2 \in &=ks `]k, +oo[ -> + consec (&=ks `[k, +oo[) t1 t2 -> P t1 -> P t2) -> + forall t, t \in &=ks `]k, +oo[ -> P t. +Proof. +move=>U K H0 IH t; apply: (consec_ind (k:=k))=>//=. +- by rewrite eqslice_uniq // andbT eqslice_mem_uniq //= negb_and in_itv /= ltxx /= orbT. +by move=>t1 t2 H C; apply: IH=>//; rewrite eqsl_xuL K. +Qed. + +End ConsecEq. + +Section ConsecOrd. +Variable (A : ordType). +Implicit Type (ks : seq A). + +(* consecutiveness and sortedness under A *) + +Lemma consec_sorted ks t1 t2 : + uniq ks -> + sorted ord ks -> t2 \in ks -> consec ks t1 t2 -> + {in ks, forall z, ord z t2 = oleq z t1}. +Proof. +move=>U S T2 /(consecP_inlt _ _ U) [T1 H] z Z. +rewrite -(slt_sortedE S Z T2) -(sle_sortedE S Z T1). +by apply: H Z. +Qed. + +End ConsecOrd. + +Section ConsecNat. + +(* an element is either first, or has a predecessor *) +Lemma consec_prevX (ks : seq nat) x : + uniq ks -> + x \in ks -> + ks = x :: behead ks \/ exists y, consec ks y x. +Proof. +case: (not_memX ks)=>k N U X. +have U' : uniq (k :: ks) by rewrite /= N U. +have : k <[k::ks] x by rewrite sltL; case: eqP X N=>// ->->. +case/(olt_consec_prev _ _ U')=>t [_]; rewrite consec_consEP' //. +by case: eqP X=>[_ ->|_ _ []]; [left|right; exists t]. +Qed. + +(* an element is either last, or has a successor in ks *) +Lemma consec_nextX (ks : seq nat) x : + uniq ks -> + x \in ks -> + (exists ks', ks = rcons ks' x) \/ + exists y, y \in ks /\ consec ks x y. +Proof. +case: (not_memX ks)=>k N U X. +have Ur : uniq (rcons ks k) by rewrite rcons_uniq N U. +have : x <[rcons ks k] k by rewrite slt_rcons (negbTE N) eq_refl orbF. +case/(olt_consec_next _ _ Ur)=>t []. +rewrite consec_rconsEP' //. +case: eqP X N=>[->->//|/eqP Nkx X N /=]. +case T: (t \in ks)=>/=; first by right; exists t. +by case=>-> {t T} [ks' -> _]; left; exists ks'. +Qed. + +End ConsecNat. diff --git a/core/useqord.v b/core/useqord.v new file mode 100644 index 0000000..ec0b84e --- /dev/null +++ b/core/useqord.v @@ -0,0 +1,575 @@ +(* +Copyright 2022 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 path interval order. +From pcm Require Import options prelude ordtype seqext. + +Open Scope order_scope. +Import Order.Theory. + +(* We assume the sequences are unique and use the first index, however most *) +(* lemmas don't require this condition explicitly. The ones that do are *) +(* grouped in a separate section. *) + +(***********************************) +(***********************************) +(* Sequence-induced ordering *) +(* definition and basic properties *) +(***********************************) +(***********************************) + +(* x <[ks] y if first x appears to the left of last y in the sequence ks *) + +(* It turns out it's useful to have 0 <[ks] x, for every x. *) +(* Basically, we use these orderings for reasoning about *) +(* timestamps in histories, and we always keep the null timestamp *) +(* to stand for the initialization step *) +(* That said, the null timestamp is never in any history as *) +(* the initialization step is implicit *) + +Module Type SeqOrdTp. +Parameter seq_le : forall (A : eqType) (ks : seq A), A -> A -> bool. +Parameter seq_lt : forall (A : eqType) (ks : seq A), A -> A -> bool. +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"). +Parameter seqle_unlock : forall (A : eqType) ks (t1 t2 : A), + t1 <=[ks] t2 = (index t1 ks <= index t2 ks)%N. +Parameter seqlt_unlock : forall (A : eqType) ks (t1 t2 : A), + t1 <[ks] t2 = (index t1 ks < index t2 ks)%N. +End SeqOrdTp. + +Module SeqOrd : SeqOrdTp. +Section SeqOrd. +Variables (A : eqType) (ks : seq A) (t1 t2 : A). +Definition seq_le := (index t1 ks <= index t2 ks)%N. +Definition seq_lt := (index t1 ks < index t2 ks)%N. +Definition seqle_unlock := erefl seq_le. +Definition seqlt_unlock := erefl seq_lt. +End SeqOrd. +End SeqOrd. +Export SeqOrd. + +(* alternative rewrites *) +Lemma seqle_unlockE (A : eqType) ks (t1 t2 : A) : + t1 <=[ks] t2 = (index t1 ks <= index t2 ks). +Proof. exact: seqle_unlock. Qed. + +Lemma seqlt_unlockE (A : eqType) ks (t1 t2 : A) : + t1 <[ks] t2 = (index t1 ks < index t2 ks). +Proof. exact: seqlt_unlock. Qed. + +Section SeqLeBase. +Variable (A : eqType). +Implicit Type (ks : seq A). + +(****************** transitivity ****************) + +Lemma sle_trans ks : transitive (seq_le ks). +Proof. by move=>y x z; rewrite !seqle_unlock; apply: leq_trans. Qed. + +Lemma slt_trans ks : transitive (seq_lt ks). +Proof. by move=>y x z; rewrite !seqlt_unlock; apply: ltn_trans. Qed. + +Lemma sle_slt_trans ks t1 t2 t3 : + t1 <=[ks] t2 -> t2 <[ks] t3 -> t1 <[ks] t3. +Proof. by rewrite !seqlt_unlock !seqle_unlock; apply: leq_ltn_trans. Qed. + +Lemma slt_sle_trans ks t1 t2 t3 : + t1 <[ks] t2 -> t2 <=[ks] t3 -> t1 <[ks] t3. +Proof. by rewrite !seqlt_unlock !seqle_unlock; apply: leq_trans. Qed. + + +(****************** reflexivity ****************) + +Lemma sle_refl ks : reflexive (seq_le ks). +Proof. by move=>x; rewrite seqle_unlock. Qed. + +(****************** irreflexivity ***************) + +Lemma slt_irr x ks : x <[ks] x = false. +Proof. by rewrite seqlt_unlock; apply: ltnn. Qed. + +(****************** antisymmetry ****************) + +Lemma sle_antisym ks : {in ks, antisymmetric (seq_le ks)}. +Proof. +move=>x Hx y; rewrite !seqle_unlock. +by rewrite -eqn_leq =>/eqP /index_inj; apply. +Qed. + +(****************** asymmetry ***************) + +Lemma slt_asym x y ks : x <[ks] y -> ~~ y <[ks] x. +Proof. by rewrite !seqlt_unlock; case: ltngtP. Qed. + +(***************** totality ********************) + +Lemma sle_total ks x y : x <=[ks] y || y <=[ks] x. +Proof. by rewrite !seqle_unlock; case: ltngtP. Qed. + +Lemma slt_total ks x y : + x \in ks -> + [|| x == y, x <[ks] y | y <[ks] x]. +Proof. +rewrite !seqlt_unlock=>H; case: ltngtP; rewrite ?orbT ?orbF //. +by move/index_inj=>->. +Qed. + +(* transfer properties of sequence ordering *) + +(****************** sle_eqVlt ***************) + +Lemma sle_eqVlt ks t1 t2 : + (t1 \in ks) || (t2 \in ks) -> + t1 <=[ks] t2 = (t1 == t2) || (t1 <[ks] t2). +Proof. +move=>H; rewrite seqlt_unlock seqle_unlock leq_eqVlt /=. +case: (t1 =P t2)=>[->|N] /=; first by rewrite eq_refl. +case: eqP=>//=; case/orP: H=>H; first by move/(index_inj H)/N. +by move/esym/(index_inj H)/esym/N. +Qed. + +(****************** slt_neqAle ***************) + +Lemma slt_neqAle ks t1 t2 : + (t1 \in ks) || (t2 \in ks) -> + t1 <[ks] t2 = (t1 != t2) && (t1 <=[ks] t2). +Proof. +move=>H. +rewrite seqlt_unlock seqle_unlock ltn_neqAle. +case: (t1 =P t2)=>[->|N] /=; first by rewrite eq_refl. +case: eqP=>//=; case/orP: H=>H; first by move/(index_inj H)/N. +by move/esym/(index_inj H)/esym/N. +Qed. + +(****************** sltNge ***************) + +Lemma sltNge ks t1 t2 : t1 <[ks] t2 = ~~ t2 <=[ks] t1. +Proof. by rewrite seqlt_unlock seqle_unlock ltnNge. Qed. + +(****************** sleNgt ***************) + +Lemma sleNgt ks t1 t2 : t1 <=[ks] t2 = ~~ t2 <[ks] t1. +Proof. by rewrite sltNge negbK. Qed. + +(* order properties of the sequence orderings *) + +(****************** slt_neq ***************) + +Corollary slt_neq x y ks : x <[ks] y -> x != y. +Proof. by apply/contraL=>/eqP->; rewrite slt_irr. Qed. + +End SeqLeBase. + +#[export] Hint Resolve sle_refl : core. + +Section SeqLeProp. +Variable (A : eqType). +Implicit Type (ks : seq A). + +Lemma sltW ks t1 t2 : t1 <[ks] t2 -> t1 <=[ks] t2. +Proof. by rewrite seqlt_unlock seqle_unlock; apply: ltnW. Qed. + + +(* membership properties of the sequence orderings *) + +Lemma slt_memI x y ks : x \in ks -> y \notin ks -> x <[ks] y. +Proof. by move=>X /index_memN E; rewrite seqlt_unlock E index_mem. Qed. + +Lemma sle_memI x y ks : y \notin ks -> x <=[ks] y. +Proof. by move/index_memN=>E; rewrite seqle_unlock E index_size. Qed. + +Lemma slt_memE x y ks : x <[ks] y -> x \in ks. +Proof. +rewrite seqlt_unlock -index_mem=>/leq_trans. +by apply; rewrite index_size. +Qed. + +Lemma sle_memE x y ks : x <=[ks] y -> y \in ks -> x \in ks. +Proof. by rewrite seqle_unlock -!index_mem; apply: leq_ltn_trans. Qed. + +(* sequence orderings and constructors *) + +Lemma slt_nil x y : x <[Nil A] y = false. +Proof. by rewrite seqlt_unlock. Qed. +Lemma sle_nil x y : x <=[Nil A] y. +Proof. by rewrite seqle_unlock. Qed. + +(* cons *) + +Lemma slt_cons x y k ks : + x <[k :: ks] y = (y != k) && ((x == k) || (x <[ks] y)). +Proof. by rewrite !seqlt_unlock /= !(eq_sym k); case: eqP; case: eqP. Qed. + +Lemma sle_cons x y k ks : + x <=[k :: ks] y = (x == k) || (y != k) && x <=[ks] y. +Proof. by rewrite sleNgt slt_cons negb_and negbK negb_or -sleNgt. Qed. + +Lemma sltL x y ks : x <[x :: ks] y = (y != x). +Proof. by rewrite slt_cons eq_refl andbT. Qed. + +Lemma sleL x y ks : x <=[x :: ks] y. +Proof. by rewrite sle_cons eq_refl. Qed. + +Lemma sltR x y ks : x <[y :: ks] y = false. +Proof. by rewrite sltNge sleL. Qed. + +Lemma sleR x y ks : x <=[y :: ks] y = (y == x). +Proof. by rewrite sleNgt sltL negbK. Qed. + +(* sequence ordering and head *) + +Lemma sle_head x ks y : (head x ks) <=[ks] y. +Proof. +case: ks=>[|k ks] /=; first by rewrite sle_nil. +by rewrite sleL. +Qed. + +(* sequence orderings and rcons *) + +Lemma slt_rcons x y k ks : + x <[rcons ks k] y = if y \in ks then x <[ks] y + else (x \in ks) || (k != y) && (k == x). +Proof. +rewrite !seqlt_unlock !index_rcons. +case X: (x \in ks); case Y: (y \in ks)=>//=. +- by case: eqP; [rewrite index_mem | rewrite ltnS index_size]. +- move/negbT/index_memN: X=>X; rewrite [LHS]ltnNge [RHS]ltnNge. + rewrite X index_size /=. + case: eqP=>//; first by rewrite index_size. + by rewrite ltnW // ltnS index_size. +rewrite !(eq_sym k). +case: eqP=>_; case: eqP=>_ /=. +- by rewrite ltnn. +- by rewrite ltnS leqnn. +- by rewrite ltnNge leqnSn. +by rewrite ltnn. +Qed. + +Lemma sle_rcons x y k ks : + x <=[rcons ks k] y = if x \in ks then x <=[ks] y + else (y \notin ks) && ((k == x) || (k != y)). +Proof. +by rewrite !sleNgt slt_rcons; case: ifP=>//; rewrite negb_or negb_and negbK. +Qed. + +(* some shortcuts for slt/sle_rcons *) + +Lemma slt_rconsI x y k ks : x <[ks] y -> x <[rcons ks k] y. +Proof. by move=>H; rewrite slt_rcons H (slt_memE H) if_same. Qed. + +Lemma sle_rconsI x y k ks : k != y -> x <=[ks] y -> x <=[rcons ks k] y. +Proof. +move=>N H; rewrite sle_rcons H N orbT andbT. +case: ifP=>// K; apply: contraFT K. +by rewrite negbK; apply: sle_memE H. +Qed. + +Lemma slt_rcons_in x y k ks : + x \in ks -> x <[rcons ks k] y = x <[ks] y. +Proof. +move=>H; rewrite slt_rcons H /=; case: ifP=>// K. +by apply/esym; apply: slt_memI=>//; rewrite K. +Qed. + +Lemma sle_rcons_in x y k ks : + x \in ks -> x <=[rcons ks k] y = x <=[ks] y. +Proof. by move=>X; rewrite sle_rcons X. Qed. + +Lemma slt_rcons_inE ks x y k1 k2 : + (x \in ks) || (y \in ks) -> + x <[rcons ks k1] y = x <[rcons ks k2] y. +Proof. by rewrite !slt_rcons=>/orP [] ->. Qed. + +Lemma sle_rcons_inE ks x y k1 k2 : + (x \in ks) || (y \in ks) -> + x <=[rcons ks k1] y = x <=[rcons ks k2] y. +Proof. by rewrite !sle_rcons=>/orP [] ->. Qed. + +Lemma slt_rconsR ks x k : x <[rcons ks k] k -> x \in ks. +Proof. by rewrite slt_rcons eq_refl orbF; case: ifP=>[_ /slt_memE|]. Qed. + +Lemma sle_rconsR ks x k : x <=[rcons ks k] k -> x \in rcons ks k. +Proof. +rewrite sle_rcons eq_refl orbF mem_rcons inE. +case X: (x \in ks); first by rewrite orbT. +by rewrite orbF eq_sym; case/andP. +Qed. + +(* sequence orderings and concatenation *) +Lemma sle_cat ks1 ks2 x y : + x <=[ks1++ks2] y = if x \in ks1 then x <=[ks1] y + else (y \notin ks1) && x <=[ks2] y. +Proof. +rewrite !seqle_unlock !index_cat. +case X: (x \in ks1); case Y: (y \in ks1)=>//=. +- move/negbT/index_memN: Y=>Y. + by rewrite Y index_size ltnW // ltn_addr // index_mem. +- rewrite -index_mem in Y. + apply/negP=>H; move/(leq_ltn_trans H): Y. + by rewrite ltnNge leq_addr. +by rewrite leq_add2l. +Qed. + +Lemma slt_cat ks1 ks2 x y : + x <[ks1++ks2] y = if y \in ks1 then x <[ks1] y + else (x \in ks1) || x <[ks2] y. +Proof. by rewrite !sltNge sle_cat; case: ifP=>//; rewrite negb_and negbK. Qed. + +(* shortcuts *) + +Lemma slt_catL ks1 ks2 x y : x <[ks1] y -> x <[ks1++ks2] y. +Proof. by move=>H; rewrite slt_cat H (slt_memE H) if_same. Qed. + +Lemma slt_splitR x y ks1 ks2 : y != x -> y \notin ks1 -> x <[ks1++x::ks2] y. +Proof. +by move=>N Y; rewrite slt_cat slt_cons eq_refl andbT (negbTE Y) N orbT. +Qed. + +Lemma sle_splitR x y ks1 ks2 : y \notin ks1 -> x <=[ks1++x::ks2] y. +Proof. +move=>Y; rewrite sle_eqVlt; last first. +- by apply/orP; left; rewrite mem_cat inE eq_refl orbT. +by case: eqP=>[|/eqP N] //=; rewrite (slt_splitR _ _ Y) // eq_sym. +Qed. + +(* the other direction of slt_splitR, further strenghtened *) +(* with an additional fact that x \notin ks1 *) +(* by picking a split with the first occurrence of x *) +(* in fact, we can have both directions here, so we prove a reflect lemma *) +(* but it should really only be used in the direction x <[ks] y -> .. *) +(* because in the other direction slt_splitR is already stronger. *) +Lemma slt_splitL ks x y : + reflect (exists ks1 ks2, [/\ ks = ks1++x::ks2, x != y, + x \notin ks1 & y \notin ks1]) + (x <[ks] y). +Proof. +case H : (x <[ks] y); constructor; last first. +- apply/contraFnot: H; case=>ks1 [ks2][-> N _]. + by apply: slt_splitR; rewrite eq_sym. +rewrite seqlt_unlock in H. +have : x \in ks by rewrite -index_mem (leq_trans H _) // index_size. +case/in_split=>ks1 [ks2][E X]; exists ks1, ks2. +rewrite /seq_lt {ks}E !index_cat /= eq_refl in H *. +case: eqP H=>[->|/eqP N] /=; first by rewrite ltnn. +rewrite (negbTE X) addn0; case: ifP=>//= _. +by rewrite ltnNge index_size. +Qed. + +(* ditto for ole_split *) +Lemma sle_splitL ks x y : + x \in ks -> + reflect (exists ks1 ks2, [/\ ks = ks1++x::ks2, + x \notin ks1 & y \notin ks1]) + (x <=[ks] y). +Proof. +move=>X; case H : (x <=[ks] y); constructor; last first. +- apply/contraFnot: H; case=>ks1 [ks2][-> _ N]. + by apply: sle_splitR. +case/in_split: X=>ks1 [ks2][E X]; exists ks1, ks2; split=>//. +rewrite seqle_unlock {ks}E !index_cat /= eq_refl (negbTE X) addn0 in H. +by case: ifP H=>//; rewrite -index_mem; case: ltngtP. +Qed. + +(* sequence orderings and filter *) + +Lemma slt_filterL (p : pred A) ks x y : + (x \notin ks) || p x -> + x <[ks] y -> x <[filter p ks] y. +Proof. by rewrite !seqlt_unlock; apply: index_filter_ltL. Qed. + +Lemma sle_filterL (p : pred A) ks x y : + (x \notin ks) || p x -> + x <=[ks] y -> x <=[filter p ks] y. +Proof. by rewrite !seqle_unlock; apply: index_filter_leL. Qed. + +Lemma slt_filterR (p : pred A) ks x y : + (y \notin ks) || p y -> + x <[filter p ks] y -> x <[ks] y. +Proof. by rewrite !seqlt_unlock; apply: index_filter_ltR. Qed. + +Lemma sle_filterR (p : pred A) ks x y : + (y \notin ks) || p y -> + x <=[filter p ks] y -> x <=[ks] y. +Proof. by rewrite !seqle_unlock; apply: index_filter_leR. Qed. + +Lemma slt_filter (p : pred A) ks x y : + (x \notin ks) || p x -> (y \notin ks) || p y -> + x <[filter p ks] y = x <[ks] y. +Proof. +by move=>H1 H2; apply/idP/idP; [apply: slt_filterR | apply: slt_filterL]. +Qed. + +Lemma sle_filter (p : pred A) ks x y : + (x \notin ks) || p x -> (y \notin ks) || p y -> + x <=[filter p ks] y = x <=[ks] y. +Proof. +by move=>H1 H2; apply/idP/idP; [apply: sle_filterR | apply: sle_filterL]. +Qed. + +(* sequence orderings and sortedness *) + +(* slt/sle under general sorted relations *) +Lemma slt_sorted_lt ltT ks x y : + transitive ltT -> + sorted ltT ks -> + y \in ks -> x <[ks] y -> ltT x y. +Proof. by rewrite seqlt_unlock; apply: sorted_index_ord. Qed. + +Lemma sle_sorted_lt ltT ks x y : + transitive ltT -> + sorted ltT ks -> + y \in ks -> x <=[ks] y -> (x == y) || ltT x y. +Proof. +move=>T S Y; rewrite sle_eqVlt; last by rewrite Y orbT. +by case/orP=>[->//|/(slt_sorted_lt T S Y) ->]; rewrite orbT. +Qed. + +(* we can get the other direction as well *) +(* if we add antisymmetry *) +(* and the condition that x \in ks *) +Lemma slt_sorted_leE leT ks x y : + antisymmetric leT -> + transitive leT -> + sorted leT ks -> + x \in ks -> y \in ks -> + x <[ks] y = (x != y) && leT x y. +Proof. +move=>As T S X Y; apply/idP/idP. +- by case: eqP=>[->|/eqP N] /=; [apply: contraLR; rewrite slt_irr | apply: slt_sorted_lt]. +by rewrite seqlt_unlock; case/andP=>H K; apply: sorted_ord_index_leq K H. +Qed. + +(* if we add antisymmetry and t1 \in ks *) +Lemma sle_sorted_leE leT ks x y : + antisymmetric leT -> + transitive leT -> + sorted leT ks -> + x \in ks -> y \in ks -> + x <=[ks] y = (x == y) || leT x y. +Proof. +move=>As T S X Y; rewrite sle_eqVlt; last by rewrite X. +by rewrite (slt_sorted_leE As T S X Y); case: eqP. +Qed. + +End SeqLeProp. + +#[export] Hint Resolve slt_nil sle_nil : core. + + +Section SeqLeUniq. +Variable (A : eqType). +Implicit Type (ks : seq A). + +Lemma slt_subseq ks1 ks2 k t : + subseq ks1 ks2 -> uniq ks2 -> k \in ks1 -> t \in ks1 -> + k <[ks1] t = k <[ks2] t. +Proof. +rewrite !seqlt_unlock=>S U T K. +by apply/idP/idP=>/(index_subseq S U T K). +Qed. + +Lemma sle_subseq ks1 ks2 k t : + subseq ks1 ks2 -> uniq ks2 -> k \in ks1 -> t \in ks1 -> + k <=[ks1] t = k <=[ks2] t. +Proof. by move=>S U T K; rewrite !sleNgt (slt_subseq S U K T). Qed. + +(* sequence orderings and last *) + +Lemma sle_last x k ks : + uniq ks -> x \in ks -> x <=[ks] (last k ks). +Proof. by rewrite seqle_unlock; apply: index_last_mono. Qed. + +Lemma sle_last_cons x k ks : + uniq (k :: ks) -> x \in k::ks -> x <=[k::ks] (last k ks). +Proof. +move=>/= /andP [U1 U2]. +rewrite inE sle_cons; case: eqP=>//= /eqP Nxk K. +by rewrite (last_notin K) //=; apply: sle_last. +Qed. + +Lemma slt_last x k ks : + uniq ks -> x \in ks -> last k ks != x -> x <[ks] (last k ks). +Proof. +move=>U X N; move: (sle_last k U X); rewrite sle_eqVlt; last by rewrite X. +by rewrite eq_sym (negbTE N). +Qed. + +Lemma slt_last_cons x k ks : + uniq (k :: ks) -> x \in k::ks -> + last k ks != x -> x <[k::ks] (last k ks). +Proof. +move=>U X N; rewrite slt_neqAle; last by rewrite X. +by rewrite eq_sym N sle_last_cons. +Qed. + +(* every list is sorted by its slt relation, assuming uniqueness *) +Lemma sorted_slt ks : uniq ks -> sorted (seq_lt ks) ks. +Proof. +case: ks=>//= k ks; elim: ks k=>[|k1 ks IH] k2 //=. +rewrite inE negb_or -andbA=>/and4P [N1 N2 N3 N4]. +rewrite sltL eq_sym N1 /=. +have : path (seq_lt [:: k1 & ks]) k1 ks by apply: IH; rewrite N3 N4. +apply: (@sub_in_path _ (mem (k1::ks))); last by apply/allP. +move=>x y /=; rewrite !inE !slt_cons. +case/orP=>[/eqP ->{x}|X]. +- rewrite (eq_sym k1 k2) (negbTE N1) /= eq_refl andbT. + case/orP=>[/eqP ->|Y ->]; first by rewrite eq_refl. + by case: eqP Y N2=>// ->->. +case/orP=>[/eqP ->|Y]; first by rewrite eq_refl. +case: eqP Y N3=>[->|/eqP N Y N3] //=. +case: eqP X N3=>[->->|/eqP M X K1] //= H. +by rewrite H orbT andbT; case: eqP Y N2=>// ->->. +Qed. + +Lemma sorted_sle ks : uniq ks -> sorted (seq_le ks) ks. +Proof. +move=>U; apply: sub_sorted (sorted_slt U). +by move=>x y /sltW. +Qed. + +End SeqLeUniq. + +Section SeqLeOrd. +Variable (A : ordType). +Implicit Type (ks : seq A). + +(* olt/ole and sortedness under ordering on A *) + +Lemma slt_sorted ks x y : + sorted ord ks -> y \in ks -> x <[ks] y -> ord x y. +Proof. by apply/slt_sorted_lt/trans. Qed. + +Lemma sle_sorted ks x y : + sorted ord ks -> y \in ks -> x <=[ks] y -> oleq x y. +Proof. by rewrite oleq_eqVord; apply/sle_sorted_lt/trans. Qed. + +Lemma slt_sortedE ks x y : + sorted ord ks -> + x \in ks -> y \in ks -> + x <[ks] y = ord x y. +Proof. +move=>S X Y; apply/idP/idP; first by apply: slt_sorted S Y. +by rewrite seqlt_unlock; apply: (sorted_ord_index (@irr _) (@trans _)) S X. +Qed. + +Lemma sle_sortedE ks x y : + sorted ord ks -> + x \in ks -> y \in ks -> + x <=[ks] y = oleq x y. +Proof. by move=>S X Y; rewrite oleqNord sleNgt (slt_sortedE S Y X). Qed. + +End SeqLeOrd. diff --git a/core/uslice.v b/core/uslice.v new file mode 100644 index 0000000..3624796 --- /dev/null +++ b/core/uslice.v @@ -0,0 +1,1007 @@ +(* +Copyright 2022 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 path interval order. +From pcm Require Import options prelude ordtype seqext slice useqord. + +(* We assume the sequences are unique and use the first index, *) +(* however most lemmas don't require this condition explicitly. *) +(* The ones that do are grouped in separate sections. *) + +Open Scope order_scope. +Import Order.Theory. + +(* slicing by element index *) + +Definition ix_bnd {A : eqType} (s : seq A) (i : itv_bound A) : itv_bound nat := + match i with + | BSide b x => BSide b (index x s) + | BInfty b => BInfty _ b + end. + +Definition ix_itv {A : eqType} (s : seq A) (i : interval A) : interval nat := + let: Interval l u := i in + Interval (ix_bnd s l) (ix_bnd s u). + +Definition eq_slice {A : eqType} (s : seq A) (i : interval A) := + slice s (ix_itv s i). + +Notation "&= s i" := (eq_slice s i) + (at level 1, i at next level, s at next level, + format "&= s i") : fun_scope. + +Section Lemmas. +Variable (A : eqType). +Implicit Type (s : seq A). + +Lemma eqsl_uu s : + &=s `]-oo,+oo[ = s. +Proof. by apply: slice_uu. Qed. + +Lemma eqsl_underL s (i j : itv_bound A) : + bnd (ix_bnd s i) (size s) = 0%N -> + &=s (Interval i j) = &=s (Interval -oo j). +Proof. by move=>H; apply: itv_underL. Qed. + +Lemma eqsl_underR s (i j : itv_bound A) : + bnd (ix_bnd s j) (size s) = 0%N -> + &=s (Interval i j) = [::]. +Proof. by move=>H; apply: itv_underR. Qed. + +Lemma eqsl_overL s (i j : itv_bound A) : + size s <= bnd (ix_bnd s i) (size s) -> + &=s (Interval i j) = [::]. +Proof. by move=>H; apply: itv_overL. Qed. + +Lemma eqsl_overR s (i j : itv_bound A) : + size s <= bnd (ix_bnd s j) (size s) -> + &=s (Interval i j) = &=s (Interval i +oo). +Proof. by move=>Hj; apply: itv_overR. Qed. + +Lemma eqsl_swapped s (i j : itv_bound A) : + bnd (ix_bnd s j) (size s) <= bnd (ix_bnd s i) (size s) -> + &=s (Interval i j) = [::]. +Proof. +by move=>H; rewrite /eq_slice /=; apply: itv_swapped. +Qed. + +Corollary eqsl_notinL y b t s : + t \notin s -> + &=s (Interval (BSide b t) y) = [::]. +Proof. +move=>T; apply: eqsl_overL=>/=. +by rewrite (memNindex T); apply: leq_addr. +Qed. + +Corollary eqsl_notinR x b t s : + t \notin s -> + &=s (Interval x (BSide b t)) = &=s (Interval x +oo). +Proof. +move=>T; apply: eqsl_overR=>/=. +by rewrite (memNindex T); apply: leq_addr. +Qed. + +Lemma eqsl_minfR s (i : itv_bound A) : + &=s (Interval i -oo) = [::]. +Proof. by rewrite /eq_slice /=; apply: itv_minfR. Qed. + +Lemma eqsl_pinfL s (j : itv_bound A) : + &=s (Interval +oo j) = [::]. +Proof. by rewrite /eq_slice /=; apply: itv_pinfL. Qed. + +Lemma eqsliceRO_notin s x a : + x \notin &=s (Interval a (BLeft x)). +Proof. +rewrite /eq_slice /slice /= addn0. +apply/negP=>/mem_drop; apply/negP. +by apply: mem_take_index. +Qed. + +Corollary eqslice_subset_full s i : + {subset &=s i <= s}. +Proof. by exact: slice_subset_full. Qed. + +(* splitting *) + +Lemma eqslice_split (s : seq A) (i : interval A) b (x : A) : + ix_bnd s i.1 <= BSide b (index x s) <= ix_bnd s i.2 -> + &=s i = &=s (Interval i.1 (BSide b x)) ++ &=s (Interval (BSide b x) i.2). +Proof. by case: i=>i j H; rewrite /eq_slice; apply: slice_split_bnd. Qed. + +Corollary eqslice_split_full s b (x : A) : + s = &=s (Interval -oo (BSide b x)) ++ &=s (Interval (BSide b x) +oo). +Proof. by rewrite -[LHS](eqsl_uu s); apply: eqslice_split. Qed. + +(* test lemmas / instantiations *) + +Corollary eqsl_uoxx (ks : seq A) t1 t2 : + t1 <=[ks] t2 -> + &=ks `]-oo, t2] = &=ks `]-oo, t1[ ++ &=ks `[t1, t2]. +Proof. by rewrite seqle_unlock=>H; apply: eqslice_split. Qed. + +Corollary eqsl_uxoo (ks : seq A) t1 t2 : + t1 <[ks] t2 -> + &=ks `]-oo, t2[ = &=ks `]-oo, t1] ++ &=ks `]t1, t2[. +Proof. by rewrite seqlt_unlock=>H; apply: eqslice_split. Qed. + +Corollary eqsl_uoxo (ks : seq A) t1 t2 : + t1 <=[ks] t2 -> + &=ks `]-oo, t2[ = &=ks `]-oo,t1[ ++ &=ks `[t1, t2[. +Proof. by rewrite seqle_unlock=>H; apply: eqslice_split. Qed. + +Corollary eqsl_uxox (ks : seq A) t1 t2 : + t1 <=[ks] t2 -> + &=ks `]-oo, t2] = &=ks `]-oo, t1] ++ &=ks `]t1, t2]. +Proof. by rewrite seqle_unlock=>H; apply: eqslice_split. Qed. + +Corollary eqsl_xxou (ks : seq A) t1 t2 : + t1 <=[ks] t2 -> + &=ks `[t1, +oo[ = &=ks `[t1, t2] ++ &=ks `]t2, +oo[. +Proof. +rewrite seqle_unlock=>H. +by apply: eqslice_split=>/=; rewrite andbT. +Qed. + +Corollary eqsl_uxou (ks : seq A) t : ks = &=ks `]-oo, t] ++ &=ks `]t, +oo[. +Proof. exact: eqslice_split_full. Qed. + +Corollary eqsl_uoxu (ks : seq A) t : ks = &=ks `]-oo, t[ ++ &=ks `[t, +oo[. +Proof. exact: eqslice_split_full. Qed. + +(* +Lemma eqsl_kk_out s l r k : + ~~ l || r -> + &=s (Interval (BSide l k) (BSide r k)) = [::]. +Proof. +move=>H; apply: eqsl_swapped=>/=. +case/orP: H. +- by move/negbTE=>-> /=; case: r=>//=; rewrite addn1 addn0. +move=>-> /=; rewrite addn0; case: l=>/=; first by rewrite addn0. +by rewrite addn1. +Qed. +*) + +Lemma eqsl_kk1 (s : seq A) l r k : + &=s (Interval (BSide l k) (BSide r k)) = + if [&& l, ~~ r & k \in s] then [:: k] else [::]. +Proof. +rewrite /eq_slice /= slice_kk onth_index. +apply/esym; case: ifP; first by case/and3P=>->->->. +move/negbT; rewrite !negb_and negbK. +case/or3P=>[/negbTE->|->|/negbTE->] //=. +- by rewrite andbF. +by rewrite if_same. +Qed. + +(* endpoint split of single-bounded interval *) + +Lemma eqsl_uxR t s : + &=s `]-oo, t] = if t \in s + then rcons (&=s `]-oo, t[) t + else &=s `]-oo, t[. +Proof. +rewrite /eq_slice /= (@slice_split _ _ _ true (index t s)) /=; last first. +- by rewrite in_itv /=. +rewrite slice_kk /= onth_index; case: ifP=>/= H. +- by rewrite cats1. +by rewrite cats0. +Qed. + +Lemma eqsl_xuL t s : + &=s `[t, +oo[ = if t \in s + then t :: &=s `]t, +oo[ + else &=s `]t, +oo[. +Proof. +rewrite /eq_slice /= (@slice_split _ _ _ false (index t s)) //=; last first. +- by rewrite in_itv /= andbT. +by rewrite slice_kk /= onth_index; case: ifP. +Qed. + +Lemma eqsl_xxL t1 t2 s : + &=s `[t1, t2] = if (t1 <=[s] t2) && (t1 \in s) + then t1 :: &=s `]t1, t2] + else [::]. +Proof. +rewrite /eq_slice seqle_unlock /=. +case: leqP=>I /=; last by rewrite itv_swapped_bnd. +rewrite (@slice_split _ _ _ false (index t1 s)) /=; last first. +- by rewrite in_itv /= lexx. +rewrite slice_kk /= onth_index; case: ifP=>//= /negbT N1. +by rewrite (memNindex N1) itv_overL //= addn1. +Qed. + +Lemma eqsl_xxR t1 t2 s : + &=s `[t1, t2] = if t1 <=[s] t2 then + if t2 \in s + then rcons (&=s `[t1, t2[) t2 + else &=s `[t1, +oo[ + else [::]. +Proof. +rewrite /eq_slice seqle_unlock /=. +case: leqP=>I /=; last by rewrite itv_swapped_bnd //. +rewrite (@slice_split _ _ _ true (index t2 s)) /=; last first. +- by rewrite in_itv /= lexx andbT. +rewrite slice_kk /= onth_index /=; case: ifP=>/=; first by rewrite cats1. +rewrite cats0 => /negbT/memNindex->. +by apply: itv_overR=>/=; rewrite addn0. +Qed. + +Lemma eqsl_xoL t1 t2 s : + &=s `[t1, t2[ = if t1 <[s] t2 + then t1 :: &=s `]t1, t2[ + else [::]. +Proof. +rewrite /eq_slice seqlt_unlock /=. +case: ltnP=>I; last by rewrite itv_swapped_bnd. +rewrite (@slice_split _ _ _ false (index t1 s)) /=; last first. +- by rewrite in_itv /= lexx. +rewrite slice_kk /= onth_index; case: ifP=>//= /negbT/memNindex E. +by move: I; rewrite E ltnNge index_size. +Qed. + +Lemma eqsl_oxR t1 t2 s : + &=s `]t1, t2] = if t1 <[s] t2 then + if t2 \in s + then rcons (&=s `]t1, t2[) t2 + else &=s `]t1, +oo[ + else [::]. +Proof. +rewrite /eq_slice seqlt_unlock /=. +case: ltnP=>I; last by rewrite itv_swapped_bnd. +rewrite (@slice_split _ _ _ true (index t2 s)) /=; last first. +- by rewrite in_itv /= lexx andbT. +rewrite slice_kk /= onth_index /=; case: ifP=>/=; first by rewrite cats1. +rewrite cats0 =>/negbT/memNindex->. +by apply: itv_overR=>/=; rewrite addn0. +Qed. + +(* some simplifying equalities on intervals *) + +Lemma eqsl_uL_notinE s b t : + t \notin s -> + &=s `(Interval -oo (BSide b t)) = s. +Proof. +move=>N; rewrite /eq_slice /= itv_overR /=; first by exact: slice_uu. +by rewrite (memNindex N); exact: leq_addr. +Qed. + +Lemma eqsl_uR_notinE s b t : + t \notin s -> + &=s `(Interval (BSide b t) +oo) = [::]. +Proof. +move=>N; rewrite /eq_slice /= itv_overL //=. +by rewrite (memNindex N); exact: leq_addr. +Qed. + +(* eqslice of nil *) +Lemma eqslice0 i : + &=([::] : seq A) i = [::]. +Proof. by rewrite /eq_slice slice0. Qed. + +(* concat *) +(* TODO unify *) + +Lemma eqsl_uL_catE s1 s2 b t : + &= (s1 ++ s2) (Interval -oo (BSide b t)) = + if t \in s1 then &= s1 (Interval -oo (BSide b t)) + else s1 ++ &= s2 (Interval -oo (BSide b t)). +Proof. +rewrite /eq_slice slice_cat /= index_cat; case: ifP=>H1. +- by rewrite index_mem H1 itv_minfR cats0. +rewrite ltnNge leq_addr /= addnC addnK itv_overR /=; first by rewrite slice_uu. +by rewrite -addnA addnCA; exact: leq_addr. +Qed. + +Lemma eqsl_uR_catE s1 s2 b t : + &= (s1 ++ s2) (Interval (BSide b t) +oo) = + if t \in s1 then &= s1 (Interval (BSide b t) +oo) ++ s2 + else &= s2 (Interval (BSide b t) +oo). +Proof. +rewrite /eq_slice slice_cat /= index_cat; case: ifP=>H1. +- by rewrite index_mem H1 slice_uu. +rewrite ltnNge leq_addr /= addnC addnK itv_overL //=. +by rewrite -addnA addnCA; exact: leq_addr. +Qed. + +Lemma eqsl_catE s1 s2 b1 t1 b2 t2 : + &= (s1 ++ s2) (Interval (BSide b1 t1) (BSide b2 t2)) = + if t1 \in s1 + then if t2 \in s1 + then &= s1 (Interval (BSide b1 t1) (BSide b2 t2)) + else &= s1 (Interval (BSide b1 t1) +oo) ++ &= s2 (Interval -oo (BSide b2 t2)) + else if t2 \in s1 + then [::] + else &= s2 (Interval (BSide b1 t1) (BSide b2 t2)). +Proof. +rewrite /eq_slice slice_cat /= !index_cat. +case/boolP: (t1 \in s1)=>H1; case/boolP: (t2 \in s1)=>H2. +- by rewrite !index_mem H1 H2 itv_minfR cats0. +- rewrite index_mem H1 ltnNge leq_addr /= itv_overR /=; last by rewrite -addnA; exact: leq_addr. + by congr (_ ++ _); rewrite addnC addnK. +- by rewrite ltnNge leq_addr /= index_mem H2 itv_minfR cats0 itv_overL //= -addnA; exact: leq_addr. +rewrite !ltnNge !leq_addr /= itv_overL /=; last by rewrite -addnA; exact: leq_addr. +by do 2!rewrite addnC addnK. +Qed. + +(* cons lemmas *) +(* TODO unify *) + +Lemma eqsl_uL_consE s b k t : + &=(k :: s) (Interval -oo (BSide b t)) = + if t == k then + if b + then [::] + else [::k] + else k :: &=s (Interval -oo (BSide b t)). +Proof. +rewrite -cat1s eqsl_uL_catE /= inE; case: eqVneq=>// {t}->. +by rewrite /eq_slice /= eqxx; case: b. +Qed. + +Corollary eqsl_uL_consL s b t : + &=(t :: s) (Interval -oo (BSide b t)) = + if b then [::] else [::t]. +Proof. by rewrite eqsl_uL_consE eqxx. Qed. + +Lemma eqsl1_uL (k : A) b t : + &=[::k] (Interval -oo (BSide b t)) = + if ~~ b || (t!=k) then [::k] else [::]. +Proof. +rewrite eqsl_uL_consE; case: eqVneq=>_. +- by rewrite orbF; case: b. +by rewrite orbT eqslice0. +Qed. + +Lemma eqsl_uR_consE s b k t : + &=(k :: s) (Interval (BSide b t) +oo) = + if t == k then + if b + then k::s + else s + else &=s (Interval (BSide b t) +oo). +Proof. +rewrite -cat1s eqsl_uR_catE /= inE; case: eqVneq=>// {t}->. +by rewrite /eq_slice /= eqxx; case: b. +Qed. + +Corollary eqsl_uR_consL s b t : + &=(t :: s) (Interval (BSide b t) +oo) = if b then t::s else s. +Proof. by rewrite eqsl_uR_consE eqxx. Qed. + +Lemma eqsl1_uR (k : A) b t : + &=[::k] (Interval (BSide b t) +oo) = + if b && (t==k) then [::k] else [::]. +Proof. +rewrite eqsl_uR_consE; case: eqVneq=>_; first by rewrite andbT. +by rewrite eqslice0 andbF. +Qed. + +Lemma eqsl_consE s k b1 t1 b2 t2 : + &=(k :: s) (Interval (BSide b1 t1) (BSide b2 t2)) = + if t1 == k then + if t2 == k + then if b1 && ~~ b2 then [::k] else [::] + else if b1 then k :: &=s (Interval -oo (BSide b2 t2)) + else &=s (Interval -oo (BSide b2 t2)) + else if t2 == k + then [::] + else &=s (Interval (BSide b1 t1) (BSide b2 t2)). +Proof. +rewrite -cat1s eqsl_catE /= !inE. +case: eqVneq=>[{t1}->|N1]; case: eqVneq=>[E2|N2] //=. +- by rewrite /eq_slice /= E2 eqxx; case: b1; case: b2. +by rewrite /eq_slice /= eqxx; case: b1. +Qed. + +(* rcons lemmas *) + +Lemma eqsl_uL_rconsE s b k t : + &=(rcons s k) (Interval -oo (BSide b t)) = + if t \in s + then &=s (Interval -oo (BSide b t)) + else if ~~ b || (t!=k) then rcons s k else s. +Proof. +rewrite -cats1 eqsl_uL_catE /=; case: ifP=>//= _. +by rewrite eqsl1_uL; case: ifP=>_ //; rewrite cats0. +Qed. + +Lemma eqsl_uR_rconsE s b k t : + &=(rcons s k) (Interval (BSide b t) +oo) = + if t \in s + then rcons (&=s (Interval (BSide b t) +oo)) k + else if b && (t == k) then [:: k] else [::]. +Proof. +rewrite -cats1 eqsl_uR_catE /=; case: ifP=>//= _. +- by rewrite cats1. +by rewrite eqsl1_uR. +Qed. + +Lemma eqsl_rconsE s k b1 t1 b2 t2 : + &=(rcons s k) (Interval (BSide b1 t1) (BSide b2 t2)) = + if t1 \in s then + if t2 \in s + then &=s (Interval (BSide b1 t1) (BSide b2 t2)) + else if b2 && (k==t2) + then &=s (Interval (BSide b1 t1) +oo) + else rcons (&=s (Interval (BSide b1 t1) +oo)) k + else if t2 \in s + then [::] + else if [&& b1, k==t1 & (k==t2) ==> ~~b2] + then [::k] else [::]. +Proof. +rewrite -cats1 eqsl_catE. +case/boolP: (t1 \in s)=>H1; case/boolP: (t2 \in s)=>H2 //. +- rewrite /eq_slice /=; case: eqVneq; case B2: b2=>/= _; try by rewrite cats1. + by rewrite cats0. +by rewrite /eq_slice /=; case: eqVneq=>_; case: eqVneq=>_; case: b1; case: b2. +Qed. + +(* test surgery lemmas *) + +Lemma eqsl_uo_split t (s1 s2 : seq A) : + t \notin s1 -> + &=(s1++t::s2) `]-oo, t[ = s1. +Proof. +by move=>X1; rewrite eqsl_uL_catE (negbTE X1) eqsl_uL_consE eqxx cats0. +Qed. + +Lemma eqsl_ou_split t (s1 s2 : seq A) : + t \notin s1 -> + &=(s1++t::s2) `]t, +oo[ = s2. +Proof. +by move=>X1; rewrite eqsl_uR_catE (negbTE X1) eqsl_uR_consE eqxx. +Qed. + +Lemma eqsl_oo_split t1 t2 (s1 s2 s : seq A) : + t1 != t2 -> + t1 \notin s1 -> + t2 \notin s1 -> + t2 \notin s -> + &=(s1++t1::s++t2::s2) `]t1, t2[ = s. +Proof. +move=>N X1 X21 X2. +rewrite eqsl_catE -cat_cons eqsl_uL_catE eqsl_catE /= + eqsl_consE !eqsl_uL_consE eqsl_uR_consE /= !inE !eqxx /= !cats0. +by rewrite (negbTE X1) (negbTE X21) eq_sym (negbTE N) /= (negbTE X2). +Qed. + +Corollary eqsl_oo_nil_split t1 t2 (s s2 : seq A) : + t1 != t2 -> + t2 \notin s -> + &=(t1::s++t2::s2) `]t1, t2[ = s. +Proof. +move=>N X2. +by rewrite -(cat0s (_ :: _ ++ _)); apply: eqsl_oo_split. +Qed. + +Corollary eqsl_oo_split_consec t1 t2 (s1 s2 : seq A) : + t1 != t2 -> + t1 \notin s1 -> t2 \notin s1 -> + &=(s1++t1::t2::s2) `]t1, t2[ = [::]. +Proof. +move=>N X1 X2. +by rewrite -(cat1s t1); apply: (@eqsl_oo_split _ _ _ _ [::]). +Qed. + +Corollary eqsl_oo_nil_split_consec t1 t2 s : + t1 != t2 -> + &=(t1::t2::s) `]t1, t2[ = [::]. +Proof. +move=>N. +by rewrite -(cat0s (_ :: _ :: _)); apply: eqsl_oo_split_consec. +Qed. + +(* intervals and filter *) + +(* TODO unify / better theory *) +Lemma eqsl_filterL (p : pred A) b (y : A) s : + (y \notin s) || p y -> + &= (filter p s) (Interval -oo (BSide b y)) = filter p (&= s (Interval -oo (BSide b y))). +Proof. +case/orP=>Hy. +- rewrite !eqsl_notinR //=; first by rewrite !eqsl_uu. + by apply: contra Hy; rewrite mem_filter; case/andP. +elim: s=>//= h s IH. +case/boolP: (p h)=>/= Hp; last first. +- rewrite {}IH eqsl_uL_consE; case: eqVneq=>[E|_] /=. + - by rewrite -E Hy in Hp. + by rewrite (negbTE Hp). +rewrite !eqsl_uL_consE; case: eqVneq=>_ /=. +- by case: {IH}b=>//=; rewrite Hp. +by rewrite Hp IH. +Qed. + +Lemma eqsl_filterR (p : pred A) b (x : A) s : + (x \notin s) || p x -> + &= (filter p s) (Interval (BSide b x) +oo) = filter p (&= s (Interval (BSide b x) +oo)). +Proof. +case/orP=>Hx. +- by rewrite !eqsl_notinL //= mem_filter negb_and Hx orbT. +elim: s=>//=h s IH. +case/boolP: (p h)=>/= Hp; last first. +- rewrite {}IH eqsl_uR_consE; case: eqVneq=>[E|_] //=. + by rewrite -E Hx in Hp. +rewrite !eqsl_uR_consE; case: eqVneq=>//= _. +by case: {IH}b=>//=; rewrite Hp. +Qed. + +Lemma eqsl_filter (p : pred A) b1 t1 b2 t2 s : + (t1 \notin s) || (p t1 && ((t2 \notin s) || p t2)) -> + &= (filter p s) (Interval (BSide b1 t1) (BSide b2 t2)) = + filter p (&= s (Interval (BSide b1 t1) (BSide b2 t2))). +Proof. +case/orP=>[N1|/andP [H1]]. +- by rewrite !eqsl_notinL //= mem_filter negb_and N1 orbT. +case/orP=>H2. +- rewrite !eqsl_notinR //=. + - by rewrite eqsl_filterR // H1 orbT. + by rewrite mem_filter negb_and H2 orbT. +elim: s=>//= h s IH. +case/boolP: (p h)=>/= Hp; last first. +- rewrite {}IH eqsl_consE; case: eqVneq=>[E1|_]. + - by rewrite -E1 H1 in Hp. + case: eqVneq=>[E2|_] //. + by rewrite -E2 H2 in Hp. +rewrite !eqsl_consE; case: eqVneq=>/=_; case: eqVneq=>//=_. +- by case: ifP=>//= _; rewrite Hp. +rewrite eqsl_filterL; last by rewrite H2 orbT. +by case: ifP=>//= _; rewrite Hp. +Qed. + +Lemma eqslice_mem (i : interval A) (ks : seq A) (k : A) : + k \in &=ks i = + has (fun j => j \in ix_itv ks i) (indexall k ks). +Proof. by rewrite /eq_slice slice_memE. Qed. + +Lemma eqslice_sorted r s i : + sorted r s -> sorted r (&=s i). +Proof. by exact: slice_sorted. Qed. + +End Lemmas. + +Section LemmasUniq. +Variable (A : eqType). +Implicit Type (s : seq A). + +Lemma eqsliceLO_notin (s : seq A) x b : + uniq s -> + x \notin &=s (Interval (BRight x) b). +Proof. +rewrite /eq_slice /slice /= addn1 => U. +move: (mem_drop_indexlast x s); rewrite indexlast_uniq //. +by apply/contra/prefix_drop_sub/prefix_take. +Qed. + +Lemma eqsl_lastR_uniq s x : + uniq s -> + s = &=s `]-oo, (last x s)]. +Proof. +move=>U; rewrite /eq_slice /= [LHS]slice_usize index_last_size_uniq // slice_oPR. +by case: ifP=>// /negbT; rewrite -ltnNge /= ltnS leqn0 => /eqP/size0nil->. +Qed. + +Lemma eqslice_uniq s i : + uniq s -> uniq (&=s i). +Proof. +rewrite /eq_slice; apply: slice_uniq. +Qed. + +Lemma eqslice_mem_uniq (i : interval A) s (x : A) : + uniq s -> + x \in &=s i = + (x \in s) && (index x s \in ix_itv s i). +Proof. +move=>U; rewrite eqslice_mem indexall_uniq //. +by case: ifP=>//= _; rewrite orbF. +Qed. + +End LemmasUniq. + +(* interaction of slicing and sequence ordering under uniqueness assumptions *) + +Section SliceSeqOrd. +Variable (A : eqType). +Implicit Type (s : seq A). + +(* TODO generalize / refactor to useq? *) +Lemma uniq_ux_filter s i : + uniq s -> + &=s `]-oo, i] = [seq x <- s | x <=[s] i]. +Proof. +move=>U; rewrite /eq_slice /= /slice /= drop0 addn1. +elim: s U=>//=h s IH. +case/andP=>Nh U; rewrite sle_cons eqxx /=. +congr (cons _); case: (eqVneq i h)=>[E|N]. +- rewrite -{h}E in Nh *; rewrite take0 -(filter_pred0 s). + apply: eq_in_filter=>z Hz /=; rewrite sle_cons eqxx /= orbF. + by apply/esym/negbTE; apply: contraNneq Nh=><-. +rewrite IH //; apply: eq_in_filter=>z Hz; rewrite sle_cons N /=. +suff: (z != h) by move/negbTE=>->. +by apply: contraNneq Nh=><-. +Qed. + +Lemma uniq_uo_filter s i : + uniq s -> + &=s `]-oo, i[ = [seq x <- s | x <[s] i]. +Proof. +move=>U; rewrite /eq_slice /= /slice /= drop0 addn0. +elim: s U=>//=h s IH. +case/andP=>Nh U; rewrite slt_cons eqxx /= andbT. +case: (eqVneq i h)=>[E|N] /=. +- rewrite -{h}E in Nh *; rewrite -(filter_pred0 s). + by apply: eq_in_filter=>z Hz /=; rewrite slt_cons eqxx. +congr (cons _); rewrite IH //; apply: eq_in_filter=>z Hz. +rewrite slt_cons N /=. +suff: (z != h) by move/negbTE=>->. +by apply: contraNneq Nh=><-. +Qed. + +(* sequence ordering, intervals, and last *) + +Lemma olt_ole_last k s x t : + uniq (k::s) -> t != k -> + x <[k::s] t = x <=[k::s] (last k (&=s `]-oo, t[)). +Proof. +elim: s k=>/= [|y s IH] k //=. +- rewrite slt_cons slt_nil sle_cons sle_nil andbT eqxx orbF. + by move=>_ ->. +rewrite inE negb_or -andbA; case/and4P=>Nky K Y U Nkt. +rewrite slt_cons sle_cons Nkt /=; case: (eqVneq x k)=>//= Nkx. +rewrite eqsl_uL_consE; case: (eqVneq t y)=>/= [E|Nty]. +- by rewrite eqxx /= E sltR. +suff -> /= : last y (&=s `]-oo, t[) != k by rewrite IH //= Y. +apply: contraTneq (mem_last y (&=s `]-oo, t[))=> E. +rewrite inE E negb_or Nky /=; apply: contra K. +by exact: slice_subset_full. +Qed. + +(* a slight variation to add the cons to last *) +Corollary olt_ole_last' k s x t : + uniq (k::s) -> t != k -> + x <[k::s] t = x <=[k::s] (last k (&=(k::s) `]-oo, t[)). +Proof. by move=>U K; rewrite olt_ole_last // eqsl_uL_consE (negbTE K). Qed. + +Corollary eqsl_uox_last k s t : + uniq (k::s) -> t != k -> + &=(k :: s) `]-oo, last k (&=s `]-oo, t[) ] = &=(k :: s) `]-oo,t[. +Proof. +move=>U K. +rewrite (uniq_ux_filter _ U) (uniq_uo_filter _ U). +by apply: eq_in_filter=>x _; rewrite -olt_ole_last. +Qed. + +(* membership *) + +Lemma mem_oo t1 t2 (ks : seq A) (k : A) : + uniq ks -> + reflect ([/\ k \in ks, t1 <[ks] k & k <[ks] t2]) + (k \in &=ks `]t1, t2[). +Proof. +rewrite !seqlt_unlock=>U. +by rewrite eqslice_mem_uniq // in_itv; apply: and3P. +Qed. + +Lemma mem_xo t1 t2 (ks : seq A) k : + uniq ks -> + reflect ([/\ k \in ks, t1 <=[ks] k & k <[ks] t2]) + (k \in &=ks `[t1, t2[). +Proof. +rewrite seqlt_unlock seqle_unlock=>U. +by rewrite eqslice_mem_uniq // in_itv; apply: and3P. +Qed. + +Lemma mem_ox t1 t2 (ks : seq A) k : + uniq ks -> + reflect ([/\ k \in ks, t1 <[ks] k & k <=[ks] t2]) + (k \in &=ks `]t1, t2]). +Proof. +rewrite seqlt_unlock seqle_unlock=>U. +by rewrite eqslice_mem_uniq // in_itv; apply: and3P. +Qed. + +Lemma mem_xx t1 t2 (ks : seq A) k : + uniq ks -> + reflect ([/\ k \in ks, t1 <=[ks] k & k <=[ks] t2]) + (k \in &=ks `[t1, t2]). +Proof. +rewrite !seqle_unlock=>U. +by rewrite eqslice_mem_uniq // in_itv; apply: and3P. +Qed. + +Lemma mem_uo t (ks : seq A) k : + uniq ks -> + reflect ([/\ k \in ks & k <[ks] t])(k \in &=ks `]-oo, t[). +Proof. +rewrite seqlt_unlock=>U. +by rewrite eqslice_mem_uniq // in_itv; apply: andP. +Qed. + +Lemma mem_ux t (ks : seq A) k : + uniq ks -> + reflect ([/\ k \in ks & k <=[ks] t])(k \in &=ks `]-oo, t]). +Proof. +rewrite seqle_unlock=>U. +by rewrite eqslice_mem_uniq // in_itv; apply: andP. +Qed. + +Lemma mem_ou t (ks : seq A) k : + uniq ks -> + reflect ([/\ k \in ks & t <[ks] k])(k \in &=ks `]t, +oo[). +Proof. +rewrite seqlt_unlock=>U. +by rewrite eqslice_mem_uniq // in_itv /= andbT; apply: andP. +Qed. + +Lemma mem_xu t (ks : seq A) k : + uniq ks -> + reflect ([/\ k \in ks & t <=[ks] k])(k \in &=ks `[t, +oo[). +Proof. +rewrite seqle_unlock=>U. +by rewrite eqslice_mem_uniq // in_itv /= andbT; apply: andP. +Qed. + +(* has predT lemmas *) + +Lemma has_predT_uslice (ks : seq A) i : + uniq ks -> + has predT (&=ks i) = has (fun z => index z ks \in ix_itv ks i) ks. +Proof. +move=>U; apply/hasP/hasP; case=>x. +- by rewrite eqslice_mem_uniq // =>/andP [Hx Hi] _; exists x. +by move=>Hx Hi; exists x=>//; rewrite eqslice_mem_uniq // Hx. +Qed. + +(* corollaries *) +Lemma has_oo t1 t2 (ks : seq A) : + uniq ks -> + has predT (&=ks `]t1, t2[) = has (fun z => t1 <[ks] z && z <[ks] t2) ks. +Proof. +move/has_predT_uslice=>->; apply: eq_has=>z. +by rewrite !seqlt_unlock in_itv. +Qed. + +Lemma has_ox t1 t2 (ks : seq A) : + uniq ks -> + has predT (&=ks `]t1, t2]) = has (fun z => t1 <[ks] z && z <=[ks] t2) ks. +Proof. +move/has_predT_uslice=>->; apply: eq_has=>z. +by rewrite seqlt_unlock seqle_unlock in_itv. +Qed. + +Lemma has_xo t1 t2 (ks : seq A) : + uniq ks -> + has predT (&=ks `[t1, t2[) = has (fun z => t1 <=[ks] z && z <[ks] t2) ks. +Proof. +move/has_predT_uslice=>->; apply: eq_has=>z. +by rewrite seqle_unlock seqlt_unlock in_itv. +Qed. + +Lemma has_xx t1 t2 (ks : seq A) : + uniq ks -> + has predT (&=ks `[t1, t2]) = has (fun z => t1 <=[ks] z && z <=[ks] t2) ks. +Proof. +move/has_predT_uslice=>->; apply: eq_has=>z. +by rewrite !seqle_unlock in_itv. +Qed. + +Lemma has_ou t (ks : seq A) : + uniq ks -> + has predT (&=ks `]t, +oo[) = has (fun z => t <[ks] z) ks. +Proof. +move/has_predT_uslice=>->; apply: eq_has=>z. +by rewrite seqlt_unlock in_itv /= andbT. +Qed. + +Lemma has_xu t (ks : seq A) : + uniq ks -> + has predT (&=ks `[t, +oo[) = has (fun z => t <=[ks] z) ks. +Proof. +move/has_predT_uslice=>->; apply: eq_has=>z. +by rewrite seqle_unlock in_itv /= andbT. +Qed. + +Lemma has_uo t (ks : seq A) : + uniq ks -> + has predT (&=ks `]-oo, t[) = has (fun z => z <[ks] t) ks. +Proof. +move/has_predT_uslice=>->; apply: eq_has=>z. +by rewrite seqlt_unlock in_itv. +Qed. + +Lemma has_ux t (ks : seq A) : + uniq ks -> + has predT (&=ks `]-oo, t]) = has (fun z => z <=[ks] t) ks. +Proof. +move/has_predT_uslice=>->; apply: eq_has=>z. +by rewrite seqle_unlock in_itv. +Qed. + +(* negation of has on the left side *) + +Lemma hasNL_oo (ks : seq A) t1 t2 (p : pred A) : + uniq ks -> t1 <[ks] t2 -> + ~~ has p (&=ks `]t1, t2[) -> + {in ks, forall z, p z -> z <[ks] t2 = z <=[ks] t1}. +Proof. +move=>U T /hasPn H z K P. +move: (H z); rewrite eqslice_mem_uniq // in_itv K /= =>/contraL/(_ P). +rewrite negb_and -!seqlt_unlockE -!sleNgt; case/orP=>Hz. +- by rewrite Hz; apply: (sle_slt_trans Hz). +rewrite sltNge Hz sleNgt; congr negb; apply/esym. +by apply: (slt_sle_trans T). +Qed. + +Lemma hasNL_ox (ks : seq A) t1 t2 (p : pred A) : + uniq ks -> t1 <=[ks] t2 -> + ~~ has p (&=ks `]t1, t2]) -> + {in ks, forall z, p z -> z <=[ks] t2 = z <=[ks] t1}. +Proof. +move=>U T /hasPn H z K P. +move: (H z); rewrite eqslice_mem_uniq // in_itv K /= =>/contraL/(_ P). +rewrite negb_and -seqlt_unlockE -seqle_unlockE -sleNgt -sltNge. +case/orP=>Hz. +- by rewrite Hz; apply: (sle_trans Hz). +rewrite !sleNgt Hz; congr negb; apply/esym. +by apply: (sle_slt_trans T). +Qed. + +Lemma hasNL_xo (ks : seq A) t1 t2 (p : pred A) : + uniq ks -> t1 <=[ks] t2 -> + ~~ has p (&=ks `[t1, t2[) -> + {in ks, forall z, p z -> z <[ks] t2 = z <[ks] t1}. +Proof. +move=>U T /hasPn H z K P. +move: (H z); rewrite eqslice_mem_uniq // in_itv K /= =>/contraL/(_ P). +rewrite negb_and -seqle_unlockE -seqlt_unlockE -sltNge -sleNgt. +case/orP=>Hz. +- by rewrite Hz; apply: (slt_sle_trans Hz). +rewrite !sltNge Hz; congr negb; apply/esym. +by apply: (sle_trans T). +Qed. + +Lemma hasNL_xx (ks : seq A) t1 t2 (p : pred A) : + uniq ks -> t1 <=[ks] t2 -> + ~~ has p (&=ks `[t1, t2]) -> + {in ks, forall z, p z -> z <=[ks] t2 = z <[ks] t1}. +Proof. +move=>U T /hasPn H z K P. +move: (H z); rewrite eqslice_mem_uniq // in_itv K /= =>/contraL/(_ P). +rewrite negb_and -!seqle_unlockE -!sltNge; case/orP=>Hz. +- by rewrite Hz; apply/sltW/(slt_sle_trans Hz). +rewrite sltNge sleNgt Hz; congr negb; apply/esym. +by apply/sltW/(sle_slt_trans T). +Qed. + +Lemma hasNL_ou (ks : seq A) t (p : pred A) : + uniq ks -> + ~~ has p (&=ks `]t, +oo[) -> {in ks, forall z, p z -> z <=[ks] t}. +Proof. +move=>U /hasPn H z K P. +move: (H z); rewrite eqslice_mem_uniq // in_itv K /= andbT =>/contraL/(_ P). +by rewrite -seqlt_unlockE -sleNgt. +Qed. + +Lemma hasNL_xu (ks : seq A) t (p : pred A) : + uniq ks -> + ~~ has p (&=ks `[t, +oo[) -> {in ks, forall z, p z -> z <[ks] t}. +Proof. +move=>U /hasPn H z K P. +move: (H z); rewrite eqslice_mem_uniq // in_itv K /= andbT =>/contraL/(_ P). +by rewrite -seqle_unlockE -sltNge. +Qed. + +Lemma hasNL_uo (ks : seq A) t (p : pred A) : + uniq ks -> + ~~ has p (&=ks `]-oo, t[) -> {in ks, forall z, p z -> t <=[ks] z}. +Proof. +move=>U /hasPn H z K P. +move: (H z); rewrite eqslice_mem_uniq // in_itv K /= =>/contraL/(_ P). +by rewrite -seqlt_unlockE -sleNgt. +Qed. + +Lemma hasNL_ux (ks : seq A) t (p : pred A) : + uniq ks -> + ~~ has p (&=ks `]-oo, t]) -> {in ks, forall z, p z -> t <[ks] z}. +Proof. +move=>U /hasPn H z K P. +move: (H z); rewrite eqslice_mem_uniq // in_itv K /= =>/contraL/(_ P). +by rewrite -seqle_unlockE -sltNge. +Qed. + +(* negation of has on the right side *) + +Lemma hasNR_oo (ks : seq A) t1 t2 (p : pred A) : + uniq ks -> + {in ks, forall z, p z -> z <[ks] t2 = z <=[ks] t1} -> + ~~ has p (&=ks `]t1, t2[). +Proof. +move=>U T; apply/hasPn=>z; rewrite eqslice_mem_uniq // in_itv /=. +rewrite -!seqlt_unlockE; case/and3P=>H1 H2 H3; apply: contraL H2=>P. +by rewrite -sleNgt -(T _ H1 P). +Qed. + +Lemma hasNR_ox (ks : seq A) t1 t2 (p : pred A) : + uniq ks -> + {in ks, forall z, p z -> z <=[ks] t2 = z <=[ks] t1} -> + ~~ has p (&=ks `]t1, t2]). +Proof. +move=>U T; apply/hasPn=>z; rewrite eqslice_mem_uniq // in_itv /=. +rewrite -seqlt_unlockE -seqle_unlockE. +case/and3P=>H1 H2 H3; apply: contraL H2=>P. +by rewrite -sleNgt -(T _ H1 P). +Qed. + +Lemma hasNR_xo (ks : seq A) t1 t2 (p : pred A) : + uniq ks -> + {in ks, forall z, p z -> z <[ks] t2 = z <[ks] t1} -> + ~~ has p (&=ks `[t1, t2[). +Proof. +move=>U T; apply/hasPn=>z; rewrite eqslice_mem_uniq // in_itv /=. +rewrite -seqlt_unlockE -seqle_unlockE. +case/and3P=>H1 H2 H3; apply: contraL H2=>P. +by rewrite -sltNge -(T _ H1 P). +Qed. + +Lemma hasNR_xx (ks : seq A) t1 t2 (p : pred A) : + uniq ks -> + {in ks, forall z, p z -> z <=[ks] t2 = z <[ks] t1} -> + ~~ has p (&=ks `[t1, t2]). +Proof. +move=>U T; apply/hasPn=>z; rewrite eqslice_mem_uniq // in_itv /=. +rewrite -!seqle_unlockE. +case/and3P=>H1 H2 H3; apply: contraL H2=>P. +by rewrite -sltNge -(T _ H1 P). +Qed. + +Lemma hasNR_ou (ks : seq A) t (p : pred A) : + uniq ks -> + {in ks, forall z, p z -> z <=[ks] t} -> + ~~ has p (&=ks `]t, +oo[). +Proof. +move=>U T; apply/hasPn=>z; rewrite eqslice_mem_uniq // in_itv /= andbT. +rewrite -seqlt_unlockE; case/andP=>H1 H2; apply: contraL H2=>P. +by rewrite -sleNgt; apply: T. +Qed. + +Lemma hasNR_xu (ks : seq A) t (p : pred A) : + uniq ks -> + {in ks, forall z, p z -> z <[ks] t} -> + ~~ has p (&=ks `[t, +oo[). +Proof. +move=>U T; apply/hasPn=>z; rewrite eqslice_mem_uniq // in_itv /= andbT. +rewrite -seqle_unlockE; case/andP=>H1 H2; apply: contraL H2=>P. +by rewrite -sltNge; apply: T. +Qed. + +Lemma hasNR_uo (ks : seq A) t (p : pred A) : + uniq ks -> + {in ks, forall z, p z -> t <=[ks] z} -> + ~~ has p (&=ks `]-oo, t[). +Proof. +move=>U T; apply/hasPn=>z; rewrite eqslice_mem_uniq // in_itv /=. +rewrite -seqlt_unlockE; case/andP=>H1 H2; apply: contraL H2=>P. +by rewrite -sleNgt; apply: T. +Qed. + +Lemma hasNR_ux (ks : seq A) t (p : pred A) : + uniq ks -> + {in ks, forall z, p z -> t <[ks] z} -> + ~~ has p (&=ks `]-oo, t]). +Proof. +move=>U T; apply/hasPn=>z; rewrite eqslice_mem_uniq // in_itv /=. +rewrite -seqle_unlockE; case/andP=>H1 H2; apply: contraL H2=>P. +by rewrite -sltNge; apply: T. +Qed. + +End SliceSeqOrd. From b94f334beb4a0ff538e502d5f51009d6bec97927 Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Thu, 1 Aug 2024 18:25:25 +0200 Subject: [PATCH 02/93] adding files --- pcm/automap.v | 764 ++++++ pcm/autopcm.v | 357 +++ pcm/heap.v | 573 ++++ pcm/invertible.v | 124 + pcm/morphism.v | 3324 ++++++++++++++++++++++ pcm/mutex.v | 583 ++++ pcm/natmap.v | 2712 ++++++++++++++++++ pcm/pcm.v | 1497 ++++++++++ pcm/unionmap.v | 6853 ++++++++++++++++++++++++++++++++++++++++++++++ 9 files changed, 16787 insertions(+) create mode 100644 pcm/automap.v create mode 100644 pcm/autopcm.v create mode 100644 pcm/heap.v create mode 100644 pcm/invertible.v create mode 100644 pcm/morphism.v create mode 100644 pcm/mutex.v create mode 100644 pcm/natmap.v create mode 100644 pcm/pcm.v create mode 100644 pcm/unionmap.v diff --git a/pcm/automap.v b/pcm/automap.v new file mode 100644 index 0000000..2666050 --- /dev/null +++ b/pcm/automap.v @@ -0,0 +1,764 @@ +(* +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 eqtype seq. +From pcm Require Import options pred prelude. +From pcm Require Export auto. +From pcm Require Import pcm unionmap natmap. + +(**************************************************************************) +(**************************************************************************) +(* 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 *) +(* *) +(**************************************************************************) +(**************************************************************************) + +(* DEVCOMMENT *) +(* 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. *) +(* /DEVCOMMENT *) + +(* 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. *) +(* *) +(* The context of keys is thus seq K. The context of vars is seq U. *) + +Section ReflectionContexts. +Variables (K : ordType) (C : pred K) (T : Type) (U : union_map C T). + +Structure ctx := Context {keyx : seq K; varx : seq U}. + +Definition empx := Context [::] [::]. + +(* because contexts grow during computation, *) +(* we need a notion of sub-context *) + +Definition sub_ctx (i j : ctx) := + Prefix (keyx i) (keyx j) /\ Prefix (varx i) (varx j). + +Lemma sc_refl i : sub_ctx i i. +Proof. by []. Qed. + +Lemma sc_trans i j k : sub_ctx i j -> sub_ctx j k -> sub_ctx i k. +Proof. +by case=>K1 V1 [K2 V2]; split; [move: K2 | move: V2]; apply: Prefix_trans. +Qed. + +End ReflectionContexts. + +(* Keys and map variables are syntactified as De Bruijn indices in the context. *) +(* Disjoint union is syntactified as concatenation of lists. *) +(* *) +(* Pts n v : syntax for "key indexed the context under number n" \-> v *) +(* Var n : syntax for "expression indexed in the context under number n" *) + +(* now for reflection *) + +Section Reflection. +Variables (K : ordType) (C : pred K) (T : Type) (U : union_map C T). +Implicit Type i : ctx U. + +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 undef + | Var n => if onth (varx i) n is Some f then f else undef + end. + +(* main interpretation function *) +Notation fx i := (fun t f => interp' i t \+ f). +Definition interp i ts := foldr (fx i) Unit ts. + +Lemma fE i ts x : foldr (fx i) x ts = x \+ interp i ts. +Proof. by elim: ts x=>[|t ts IH] x; rewrite /= ?unitR // IH joinCA. Qed. + +Lemma interp_rev i ts : interp i (rev ts) = interp i ts. +Proof. +elim: ts=>[|t ts IH] //=; rewrite rev_cons -cats1. +by rewrite /interp foldr_cat fE /= unitR IH. +Qed. + +(* we also need a pretty-printer, which works the same as interp *) +(* but removes trailing Unit's *) + +Fixpoint pprint i ts := + if ts is t :: ts' then + if ts' is [::] then interp' i t else interp' i t \+ pprint i ts' + else Unit. + +Lemma pp_interp i ts : pprint i ts = interp i ts. +Proof. by elim: ts=>[|t ts /= <-] //; case: ts=>//; rewrite unitR. Qed. + +Definition key n t := if t is Pts m _ then n == m else false. +Definition var n t := if t is Var m then n == m else false. + +Definition kfree n t := rfilter (key n) t. +Definition vfree n t := rfilter (var n) t. +Arguments kfree /. +Arguments vfree /. + +Lemma keyN i n ts : ~~ has (key n) ts -> interp i (kfree n ts) = interp i ts. +Proof. by elim: ts=>[|t ts IH] //=; case: ifP=>//= _ /IH ->. Qed. + +Lemma varN i n ts : ~~ has (var n) ts -> interp i (vfree n ts) = interp i ts. +Proof. by elim: ts=>[|t ts IH] //=; case: ifP=>//= _ /IH ->. Qed. + +Lemma keyP i n k ts : + has (key n) ts -> onth (keyx i) n = Some k -> + exists v, interp i ts = pts k v \+ interp i (kfree n ts). +Proof. +elim: ts=>[|t ts IH] //=; case: ifP=>[|_ H]. +- by case: t=>//= _ v /eqP <- _ ->; exists v. +by case/(IH H)=>v ->; exists v; rewrite joinCA. +Qed. + +Lemma varP i n u ts : + has (var n) ts -> onth (varx i) n = Some u -> + interp i ts = u \+ interp i (vfree n ts). +Proof. +elim: ts=>[|t ts IH] //=; case: ifP=>[|_ H]. +- by case: t=>//= _ /eqP <- _ ->. +by move/(IH H)=>->; rewrite joinCA. +Qed. + +(* interpretation is invariant under context weakening *) +(* under assumption that the interpreted term is well-formed *) + +Definition wf i t := + match t with + Pts n _ => n < size (keyx i) + | Var n => n < size (varx i) + end. + +Lemma sc_wf i j ts : sub_ctx i j -> all (wf i) ts -> all (wf j) ts. +Proof. +case=>/Prefix_size H1 /Prefix_size H2; elim: ts=>[|t ts IH] //=. +case/andP=>H /IH ->; rewrite andbT. +by case: t H=>[n v|v] H; apply: leq_trans H _. +Qed. + +Lemma sc_interp i j ts : + sub_ctx i j -> all (wf i) ts -> interp i ts = interp j ts. +Proof. +case=>H1 H2; elim: ts=>[|t ts IH] //= /andP [H] /IH ->. +by case: t H=>[n v|n] /= /Prefix_onth <-. +Qed. + +Lemma valid_wf i ts : valid (interp i ts) -> all (wf i) ts. +Proof. +elim: ts=>[|t ts IH] //= V; rewrite (IH (validR V)) andbT. +case: t {V IH} (validL V)=>[n v|n] /=; +by case X : (onth _ _)=>[a|]; rewrite ?(onth_size X) // valid_undef. +Qed. + +Lemma wf_kfree i n ts : all (wf i) ts -> all (wf i) (kfree n ts). +Proof. by elim: ts=>//= t ts IH; case: ifP=>_ /andP [] //= -> /IH ->. Qed. + +Lemma wf_vfree i n ts : all (wf i) ts -> all (wf i) (vfree n ts). +Proof. by elim: ts=>//= t ts IH; case: ifP=>_ /andP [] //= -> /IH ->. Qed. + +(* sometimes we want to get keys in a list, not in a predicate *) + +Definition getkeys := + foldr (fun t ks => if t is Pts k _ then k :: ks else ks) [::]. + +Lemma has_getkeys ts n : n \in getkeys ts = has (key n) ts. +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 validX *) +(* and domeqX and invalidX respectively. *) +(**********************************************************************) +(**********************************************************************) + +(* subterm is purely functional version of validX *) + +Section Subterm. +Variables (K : ordType) (C : pred K) (T : Type) (U : union_map C T). +Implicit Types (i : ctx U) (ts : seq (term T)). + +Fixpoint subterm ts1 ts2 := + match ts1 with + Pts n _ :: tsx1 => + if has (key n) ts2 then subterm tsx1 (kfree n ts2) else false + | Var n :: tsx1 => + if has (var n) ts2 then subterm tsx1 (vfree n ts2) else false + | [::] => true + end. + +(* the procedure is sound for deciding wf subterms *) +Lemma subterm_sound i ts1 ts2 : + all (wf i) ts1 -> all (wf i) ts2 -> subterm ts1 ts2 -> + exists u, dom_eq (interp i ts1 \+ u) (interp i ts2). +Proof. +elim: ts1 ts2=>[|t ts1 IH] ts2 /= A1 A2. +- by exists (interp i ts2); rewrite unitL. +case/andP: A1; case: t=>[n v|n] /= /size_onth [k] X A1; +rewrite X; case: ifP=>Y //. +- case: (keyP Y X)=>w -> /(IH _ A1 (wf_kfree n A2)) [xs D]. + by exists xs; rewrite -joinA; apply: domeqUn D; apply: domeqPt. +move: (varP Y X)=>-> /(IH _ A1 (wf_vfree n A2)) [xs D]. +by exists xs; rewrite -joinA; apply: domeqUn D. +Qed. + +End Subterm. + +(* subtract is purely functional version of domeqX *) + +Section Subtract. +Variables (K : ordType) (C : pred K) (T : Type) (U : union_map C T). +Implicit Types (i : ctx U) (ts : seq (term T)). + +(* We need a subterm lemma that returns the uncancelled stuff from *) +(* both sides. xs accumulates uncancelled part of ts1, starting as nil *) + +Fixpoint subtract ts1 ts2 xs := + match ts1 with + Pts n v :: tsx1 => + if has (key n) ts2 then subtract tsx1 (kfree n ts2) xs + else subtract tsx1 ts2 (Pts n v :: xs) + | Var n :: tsx1 => + if has (var n) ts2 then subtract tsx1 (vfree n ts2) xs + else subtract tsx1 ts2 (Var T n :: xs) + | [::] => (xs, ts2) + end. + +(* below, the existentially quantified u is the cancelled part *) +Lemma subtract_sound i ts1 ts2 rs1 rs2 xs : + all (wf i) ts1 -> all (wf i) ts2 -> + subtract ts1 ts2 xs = (rs1, rs2) -> + exists u, dom_eq (interp i ts1 \+ interp i xs) (interp i rs1 \+ u) /\ + dom_eq (interp i ts2) (interp i rs2 \+ u). +Proof. +elim: ts1 ts2 xs=>[|t ts1 IH] ts2 xs /= A1 A2. +- by case=><-<-; exists Unit; rewrite unitL !unitR. +case/andP: A1; case: t=>[n v|n] /= /size_onth [x X] A1; rewrite X; case: ifP=>Y. +- case: (keyP Y X)=>w -> /(IH _ _ A1 (wf_kfree n A2)) [u][H1 H2]. + exists (pts x v \+ u); rewrite -joinA !(pull (pts x _)). + by split=>//; apply: domeqUn=>//; apply: domeqPt. +- by case/(IH _ _ A1 A2)=>u [/= H1 H2]; rewrite X joinCA joinA in H1; exists u. +- move: (varP Y X)=>-> /(IH _ _ A1 (wf_vfree n A2)) [u][H1 H2]. + by exists (x \+ u); rewrite -joinA !(pull x); split=>//; apply: domeqUn. +by case/(IH _ _ A1 A2)=>u [/= H1 H2]; rewrite X joinCA joinA in H1; exists u. +Qed. + +End Subtract. + +(* invalid is a purely functional test of invalidX *) + +Section Invalid. +Variables (K : ordType) (C : pred K) (T : Type) (U : union_map C T). +Implicit Types (i : ctx U) (t : term T) (ts : seq (term T)). + +Definition undefx i t := + if t is Var n then + if onth (varx i) n is Some x then undefb x else false + else false. + +Definition isundef i ts := ~~ uniq (getkeys ts) || has (undefx i) ts. + +Lemma isundef_sound i ts : + all (wf i) ts -> isundef i ts -> ~~ valid (interp i ts). +Proof. +elim: ts=>[|t ts IH] //= /andP [W A]. +case: t W=>[n v|n] /= /size_onth [k] X; rewrite /isundef /= X; last first. +- rewrite orbCA=>H; case: validUn=>// V; rewrite (negbTE (IH A _)) //. + 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_undef2 undef_join valid_undef. +Qed. + +End Invalid. + + +(********************************) +(********************************) +(* Canonical structure programs *) +(********************************) +(********************************) + +(* 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) (C : pred K) (T : Type) (U : union_map C T). +Implicit Types (i : ctx U) (ts : seq (term T)). + +(* a tagging structure to control the flow of computation *) +Structure tagged_map := Tag {untag : U}. + +Local Coercion untag : tagged_map >-> UMC.sort. + +(* in reversed order; first test for unions, then empty, pts and vars *) +Definition var_tag := Tag. +Definition key_tag := var_tag. +Definition empty_tag := key_tag. +Canonical Structure union_tag hc := empty_tag hc. + +(* Main structure *) +(* - i : input context *) +(* - j : output context *) +(* - ts : syntactification of map_of using context j *) + +Definition axiom i j ts (pivot : tagged_map) := + [/\ interp j ts = pivot :> U, sub_ctx i j & all (wf j) ts]. +Structure form i j ts := Form {pivot : tagged_map; _ : axiom i j ts pivot}. + +Local Coercion pivot : form >-> tagged_map. + +(* check for union *) + +Lemma union_pf i j k ts1 ts2 (f1 : form i j ts1) (f2 : form j k ts2) : + axiom i k (ts1 ++ ts2) (union_tag (untag f1 \+ untag f2)). +Proof. +case: f1 f2=>_ [<- S1 W1][_][<- S2 W2]; split. +- by rewrite /interp foldr_cat fE joinC -(sc_interp S2 W1). +- by apply: sc_trans S1 S2. +by rewrite all_cat (sc_wf S2 W1) W2. +Qed. + +Canonical union_form i j k ts1 ts2 f1 f2 := + Form (@union_pf i j k ts1 ts2 f1 f2). + +(* check if reached empty *) + +Lemma empty_pf i : axiom i i [::] (empty_tag Unit). +Proof. by []. Qed. + +Canonical empty_form i := Form (@empty_pf i). + +(* check for pts k v *) + +Lemma pts_pf vars keys1 keys2 k v (f : xfind keys1 keys2 k): + axiom (Context keys1 vars) (Context keys2 vars) + [:: Pts k v] (key_tag (pts (xuntag f) v)). +Proof. by case: f=>p [E H]; split=>//=; rewrite ?E ?unitR // (onth_size E). Qed. + +Canonical pts_form vars keys1 keys2 k v f := + Form (@pts_pf vars keys1 keys2 k v f). + +(* check for var *) + +Lemma var_pf keys vars1 vars2 n (f : xfind vars1 vars2 n) : + axiom (Context keys vars1) (Context keys vars2) + [:: Var T n] (var_tag (xuntag f)). +Proof. by case: f=>p [E H]; split=>//=; rewrite ?E ?unitR // (onth_size E). Qed. + +Canonical var_form keys vars1 vars2 v f := Form (@var_pf keys vars1 vars2 v f). + +End Syntactify. + +Module Exports. +Coercion untag : tagged_map >-> UMC.sort. +Coercion pivot : form >-> tagged_map. +Canonical union_tag. +Canonical union_form. +Canonical empty_form. +Canonical pts_form. +Canonical var_form. +End Exports. +End Syntactify. + +Export Syntactify.Exports. + +(*********************) +(* Automating validX *) +(*********************) + +(* validX is a refined lemma for subterm checking which *) +(* automatically discharges the spurious argument from above *) + +Module ValidX. +Section ValidX. +Variables (K : ordType) (C : pred K) (T : Type) (U : union_map C T). +Implicit Types (j : ctx U) (ts : seq (term T)). +Notation form := Syntactify.form. +Notation untag := Syntactify.untag. + +(* The rform structure has two important components: *) +(* *) +(* -- a packed/hoisted map m, which will be reified into the ts2 argument *) +(* of subterm ts2 ts1 *) +(* *) +(* -- a boolean b, which will be instantiated with true in the validX *) +(* lemma, and will be unified with subterm ts2 ts1 in the start *) +(* instance *) +(* *) +(* The other components of rform are j ts1 and pivot, which are forced by *) +(* needing to compose the proofs, but behave plainly in unification. *) + +Structure packed_map (m : U) := Pack {unpack : U}. +Canonical equate (m : U) := Pack m m. + +Definition raxiom j ts1 m (b : bool) (pivot : packed_map m) := + all (wf j) ts1 -> valid (interp j ts1) -> b -> valid (unpack pivot). + +Structure rform j ts1 m b := + RForm {pivot :> packed_map m; _ : raxiom j ts1 b pivot}. + +(* start instance: note how subterm ts2 ts1 is unified with *) +(* the boolean component of rform *) + +Lemma start_pf j k ts1 ts2 (f2 : form j k ts2) : + @raxiom j ts1 (untag f2) (subterm ts2 ts1) (equate f2). +Proof. +case: f2=>f2 [<- S A2] A1; rewrite (sc_interp S A1)=>V. +case/(subterm_sound A2 (sc_wf S A1))=>xs. +by case/domeqP; rewrite V=>/validL ->. +Qed. + +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=>/[swap]; exact: joinKx. Qed. + +Lemma joinxK' (x x1 x2 : U) : x \+ x1 = x \+ x2 -> valid (x \+ x1) -> x1 = x2. +Proof. by move=>/[swap]; exact: joinxK. 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) (C : pred K) (T : Type) (U : union_map C T). +Implicit Types (j : ctx U) (ts : seq (term T)). +Notation form := Syntactify.form. +Notation untag := Syntactify.untag. + +(* the main lemma; note how the boolean component of rform is set to true *) + +Lemma validX m j ts1 (f1 : form (empx U) j ts1) (g: rform j ts1 m true) : + valid (untag f1) -> valid (unpack (pivot g)). +Proof. by case: g f1; case=>pivot H [f1][<- Sc A] /(H A); apply. Qed. + +End Exports. + +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 domeqX *) +(*********************) + +Module DomeqX. +Section DomeqX. +Variables (K : ordType) (C : pred K) (T : Type) (U : union_map C T). +Implicit Types (j : ctx U) (ts : seq (term T)). +Notation form := Syntactify.form. +Notation untag := Syntactify.untag. + +Structure packed_map (m : U) := Pack {unpack : U}. +Canonical equate (m : U) := Pack m m. + +(* b is the pair of residual terms *) +Definition raxiom j k ts1 m b (pivot : packed_map m) := + all (wf j) ts1 -> sub_ctx j k /\ + (dom_eq (interp k b.1) (interp k b.2) -> + dom_eq (interp k ts1) (unpack pivot)). + +Structure rform j k ts1 m b := + RForm {pivot :> packed_map m; _ : raxiom j k ts1 b pivot}. + +(* start instance: note how subtract ts1 ts2 [::] is unified with *) +(* the b component of rform thus passing the residual terms *) + +Lemma start_pf j k ts1 ts2 (f2 : form j k ts2) : + @raxiom j k ts1 (untag f2) (subtract ts1 ts2 [::]) (equate f2). +Proof. +case: f2=>f2 [<- S A2]; case E : (subtract _ _ _)=>[rs1 rs2] A1; split=>//. +case/(subtract_sound (sc_wf S A1) A2): E=>ys [/= D1 D2 D]. +rewrite unitR in D1; apply: domeq_trans D1 _. +rewrite domeq_symE; apply: domeq_trans D2 _. +by rewrite domeq_symE; apply: domeqUn. +Qed. + +Canonical start j k ts1 ts2 f2 := RForm (@start_pf j k ts1 ts2 f2). + +End DomeqX. + +Module Exports. +Canonical equate. +Canonical start. + +Section Exports. +Variables (K : ordType) (C : pred K) (T : Type) (U : union_map C T). +Implicit Types (j : ctx U) (ts : seq (term T)). +Notation form := Syntactify.form. +Notation untag := Syntactify.untag. + +(* the main lemma; notice how residuals rs1, rs2 are passed to g to compute *) + +Lemma domeqX m j k rs1 rs2 ts1 (f1 : form (empx U) j ts1) + (g : rform j k ts1 m (rs1, rs2)) : + dom_eq (pprint k (rev rs1)) (pprint k rs2) -> + dom_eq (untag f1) (unpack (pivot g)). +Proof. +case: g f1; case=>pivot R [f1][<- _ A1] /=; case/(_ A1): R=>S D. +by rewrite !pp_interp interp_rev (sc_interp S A1). +Qed. + +End Exports. + +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). +Proof. apply: domeqX=>/=. Abort. + +End Exports. +End DomeqX. + +Export DomeqX.Exports. + +(***********************) +(* Automating invalidX *) +(***********************) + +Module InvalidX. +Section InvalidX. +Variables (K : ordType) (C : pred K) (T : Type) (U : union_map C T). +Implicit Types (i : ctx U) (ts : seq (term T)). +Notation form := Syntactify.form. +Notation untag := Syntactify.untag. + +Structure packed_map (m : U) := Pack {unpack : U}. +Canonical equate (m : U) := Pack m m. + +(* b is the boolean that we want to get out of isundef *) +Definition raxiom i ts m b (pivot : packed_map m) := + b -> valid (unpack pivot) = false. + +Structure rform i ts m b := + RForm {pivot :> packed_map m; _ : raxiom i ts b pivot}. + +(* start instance: note how isundef i ts is unified with *) +(* the b component of rform, which will be set to true by lemma statements *) + +Lemma start_pf i ts (f : form (empx U) i ts) : + @raxiom i ts (untag f) (isundef i ts) (equate f). +Proof. by case: f=>f [<- S A] /(isundef_sound A) /negbTE. Qed. + +Canonical start i ts f := RForm (@start_pf i ts f). + +End InvalidX. + +Module Exports. +Canonical equate. +Canonical start. + +Section Exports. +Variables (K : ordType) (C : pred K) (T : Type) (U : union_map 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) = undef. +Proof. by case: g; case=>pivot /= /(_ (erefl _))/negbT/invalidE. Qed. + +Lemma invalidX m i ts (g : rform i ts m true) : + valid (unpack (pivot g)) = false. +Proof. by rewrite undefX valid_undef. Qed. + +End Exports. + +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). +Proof. rewrite invalidX. Abort. + +End Exports. +End InvalidX. + +Export InvalidX.Exports. + +(************************************************) +(* Pushing an omap-like function inside a union *) +(************************************************) + +Module OmfX. +Section OmfX. + +Inductive syntx A (U : natmap A) := + | UnitOmf + | PtOmf of nat & A + | UnOmf of syntx U & syntx U + | OtherOmf of U. + +Fixpoint omf_interp A B (U : natmap A) (V : natmap B) (f : omap_fun U V) e : V := + match e with + UnitOmf => Unit + | PtOmf k v => if omf f (k, v) is Some w then pts k w else Unit + | UnOmf h1 h2 => omf_interp f h1 \+ omf_interp f h2 + | OtherOmf h => f h + end. + +Structure tagged_natmap A (U : natmap A) := Tag {untag : U}. +Definition unit_tag := Tag. +Definition pt_tag := unit_tag. +Definition union_tag := pt_tag. +Definition recurse_unguard_tag := union_tag. +Canonical recurse_guard_tag A (U : natmap A) f := @recurse_unguard_tag A U f. + +Definition guard_ax A B (U : natmap A) (V : natmap B) ts e (f : omap_fun U V) := + untag e = f (omf_interp idfun ts). + +Structure guarded_form A B (U : natmap A) (V : natmap B) ts := GForm { + gpivot :> tagged_natmap V; + guard_of : omap_fun U V; + _ : guard_ax ts gpivot guard_of}. + +Definition unguard_ax A (U : natmap A) (ts : syntx U) e := + untag e = omf_interp idfun ts. + +Structure unguarded_form A (U : natmap A) (ts : syntx U) := UForm { + upivot :> tagged_natmap U; + _ : unguard_ax ts upivot}. + +(* we first try to see if there's more function guards *) +Lemma recurse_guard_pf A B C (U : natmap A) (V : natmap B) (W : natmap C) + ts (f : omap_fun V W) + (g : @guarded_form A B U V ts) : + guard_ax ts (recurse_guard_tag (f (untag g))) + (f \o guard_of g). +Proof. by case: g=>e g pf /=; rewrite pf. Qed. +Canonical recurse_guard A B C U V W ts f g := GForm (@recurse_guard_pf A B C U V W ts f g). + +(* if not, we descend to unguarded form to syntactify the underlying expression *) +Lemma recurse_unguard_pf A (U : natmap A) ts (u : @unguarded_form A U ts) : + guard_ax ts (recurse_unguard_tag (untag u)) idfun. +Proof. by case: u. Qed. +Canonical recurse_unguard A U ts u := GForm (@recurse_unguard_pf A U ts u). + +(* syntactifying union recursively descends to both sides *) +Lemma unguard_union_pf A (U : natmap A) ts1 ts2 + (u1 : @unguarded_form A U ts1) (u2 : @unguarded_form A U ts2) : + unguard_ax (UnOmf ts1 ts2) (union_tag (untag u1 \+ untag u2)). +Proof. by case: u1 u2=>u1 pf1 [u2 pf2]; rewrite /= pf1 pf2. Qed. +Canonical unguard_union A U ts1 ts2 u1 u2 := UForm (@unguard_union_pf A U ts1 ts2 u1 u2). + +(* base case for syntactifying points-to *) +Lemma unguard_pts_pf A (U : natmap A) k (v : A) : + unguard_ax (PtOmf U k v) (pt_tag (pts k v)). +Proof. by []. Qed. +Canonical unguard_pts A U k v := UForm (@unguard_pts_pf A U k v). + +(* base case for syntactifying empty map Unit *) +Lemma unguard_unit_pf A (U : natmap A) : unguard_ax (UnitOmf U) (unit_tag Unit). +Proof. by []. Qed. +Canonical unguard_unit A U := UForm (@unguard_unit_pf A U). + +(* base case for syntactifying all other expressions *) +Lemma unguard_other_pf A (U : natmap A) (e : U) : unguard_ax (OtherOmf e) (Tag e). +Proof. by []. Qed. +Canonical unguard_other A U e := UForm (@unguard_other_pf A U e). + +End OmfX. + +Module Exports. +Canonical recurse_guard_tag. +Canonical recurse_guard. +Canonical recurse_unguard. +Canonical unguard_union. +Canonical unguard_pts. +Canonical unguard_unit. +Canonical unguard_other. + +(* main lemma *) + +Lemma omfX A B C (U : natmap A) (V : natmap B) (W : natmap C) ts + (f : omap_fun V W) + (g : @guarded_form A B U V ts) : + valid (omf_interp idfun ts) -> + f (untag g) = + omf_interp (f \o guard_of g) ts. +Proof. +case: g=>eq g /=; elim: ts eq=>[|s a|ts1 IH1 ts2 IH2|t] /= eq pf X. +- by rewrite pf !pfunit. +- rewrite validPt in X; rewrite pf omfPt // omf_comp /=. + by case: (omf g _)=>[x|]; rewrite ?omfPt ?pfunit. +- rewrite pf /= !omfUn //; last by rewrite -omfUn ?pfVE. + by rewrite IH1 ?(validL X) // IH2 ?(validR X). +by rewrite pf. +Qed. +End Exports. +End OmfX. + +Export OmfX.Exports. + diff --git a/pcm/autopcm.v b/pcm/autopcm.v new file mode 100644 index 0000000..8372161 --- /dev/null +++ b/pcm/autopcm.v @@ -0,0 +1,357 @@ +(* +Copyright 2022 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. +From pcm Require Import options pred prelude auto pcm. + +(**************************************************************************) +(**************************************************************************) +(* Canonical structure lemmas for automating the following task: *) +(* *) +(* Splitting a PCM expression e into a combination of a input *) +(* subexpression e1 (modulo AC) and a residual expression obtained by *) +(* dropping all subterms of e1 from e. *) +(* *) +(* This allows us to formulate a generalized "pull"-style transformation, *) +(* which can factor out a given subexpression with one invocation. Note, *) +(* however, that this rewrites the goal into a right-biased form, losing *) +(* the associativity information. *) +(* *) +(**************************************************************************) +(**************************************************************************) + +(* Context structure for reflection of PCM expressions. An expression is *) +(* everything that are not recognized as a disjoint union, i.e., it can *) +(* be a variable or a literal. *) + +Section ReflectionContexts. +Variables (U : pcm). + +Structure ctx := Context {expx : seq U}. + +Definition empx := Context [::]. + +(* because contexts grow during computation, we need a notion of sub-context *) + +Definition sub_ctx (i j : ctx) := + Prefix (expx i) (expx j). + +Lemma sc_refl i : sub_ctx i i. +Proof. by []. Qed. + +Lemma sc_trans i j k : sub_ctx i j -> sub_ctx j k -> sub_ctx i k. +Proof. by apply: Prefix_trans. Qed. + +End ReflectionContexts. + +(* Expressions are syntactified as De Bruijn indices in the context. *) +(* Disjoint union is syntactified as concatenation of lists. *) + +(* now for reflection *) + +Section Reflection. +Variables (U : pcm). +Implicit Type i : ctx U. + +Variant term := Expr of nat. + +Definition nat_of (t : term) : nat := + let: Expr n := t in n. + +(* interpretation function for elements *) +Definition interp' (i : ctx U) (t : term) : option U := + let: Expr n := t in onth (expx i) n. + +(* main interpretation function *) +Notation fx i := (fun t f => interp' i t \+ f). +Definition interp (i : ctx U) (ts : seq term) : option U := + foldr (fx i) Unit ts. + +Lemma fE i ts x : foldr (fx i) x ts = x \+ interp i ts. +Proof. by elim: ts x=>[|t ts IH] x /=; rewrite ?unitR // IH /= joinCA. Qed. + +Lemma interp_rev i ts : interp i (rev ts) = interp i ts. +Proof. +elim: ts=>[|t ts IH] //=; rewrite rev_cons -cats1. +by rewrite /interp foldr_cat fE /= unitR IH. +Qed. + +(* we also need a pretty-printer, which works the same as interp *) +(* but removes trailing Unit's *) + +Fixpoint pprint i ts := + if ts is t :: ts' then + if ts' is [::] then interp' i t else interp' i t \+ pprint i ts' + else Unit. + +Lemma pp_interp i ts : pprint i ts = interp i ts. +Proof. by elim: ts=>[|t ts /= <-] //; case: ts=>//; rewrite unitR. Qed. + +Definition exp n t := let: Expr m := t in n == m. + +Definition efree n t := rfilter (exp n) t. +Arguments efree /. + +Lemma expP i n ts : + has (exp n) ts -> + interp i ts = onth (expx i) n \+ interp i (efree n ts). +Proof. +elim: ts=>[|t ts IH] //=; case: ifP=>/= [|_ H]. +- by case: t=>/= m /eqP->. +by rewrite (IH H) joinCA. +Qed. + +(* interpretation is invariant under context weakening *) +(* under assumption that the interpreted term is well-formed *) + +Definition wf i t := + let: Expr n := t in n < size (expx i). + +Lemma sc_wf i j ts : sub_ctx i j -> all (wf i) ts -> all (wf j) ts. +Proof. +move/Prefix_size=>H; elim: ts=>[|t ts IH] //=. +case/andP=>Hi /IH ->; rewrite andbT. +by case: t Hi=>v /= Hi; apply: leq_trans H. +Qed. + +Lemma sc_interp i j ts : + sub_ctx i j -> all (wf i) ts -> interp i ts = interp j ts. +Proof. +move=>H; elim: ts=>[|t ts IH] //= /andP [H0] /IH ->. +by case: t H0=>n /= /Prefix_onth <-. +Qed. + +Lemma wf_efree i n ts : all (wf i) ts -> all (wf i) (efree n ts). +Proof. by elim: ts=>//= t ts IH; case: ifP=>_ /andP [] //= -> /IH ->. Qed. + +End Reflection. + + +(************************************************************************) +(************************************************************************) +(* Purely functional decision procedure for the subtraction task. *) +(* Further below, it will be embedded into the canonical program pullX. *) +(************************************************************************) +(************************************************************************) + +(* subtract is purely functional version of pullX *) + +Section Subtract. +Variables (U : pcm). +Implicit Types (i : ctx U) (ts : seq term). + +(* We need a subterm lemma that returns the residiual of ts1. *) +(* xs accumulates uncancelled part of ts1, starting as nil *) + +Fixpoint subtract ts1 ts2 xs : option (seq term) := + match ts1 with + Expr n :: tsx1 => + if has (exp n) ts2 then subtract tsx1 (efree n ts2) xs + else subtract tsx1 ts2 (Expr n :: xs) + | [::] => if ts2 is [::] then Some xs else None + end. + +Lemma subtract_sound i ts1 ts2 rs1 xs : + all (wf i) ts1 -> all (wf i) ts2 -> + subtract ts1 ts2 xs = Some rs1 -> + interp i ts1 \+ interp i xs = interp i rs1 \+ interp i ts2. +Proof. +elim: ts1 ts2 xs=>[|t ts1 IH] ts2 xs /= =>[_|/andP [At A1]]. +- by case: ts2=>//=_; case=>->; rewrite unitR unitL. +case: t At=>n /= /size_onth [x X] A2; case: ifP=>Y. +- rewrite -joinA; move: (expP i Y)=>-> /(IH _ _ A1 (wf_efree n A2))->. + by rewrite joinCA. +by move/(IH _ _ A1 A2)=><-/=; rewrite joinCA joinA. +Qed. + +End Subtract. + + +(********************************) +(********************************) +(* Canonical structure programs *) +(********************************) +(********************************) + +Module Syntactify. +Section Syntactify. +Variables (U : pcm). +Implicit Types (i : ctx U) (ts : seq term). + +(* a tagging structure to control the flow of computation *) +Structure tagged_pcm := Tag {untag : U}. + +Local Coercion untag : tagged_pcm >-> PCM.sort. + +(* in reversed order; first test for unions, then empty, and exprs *) +Definition expr_tag := Tag. +Definition empty_tag := expr_tag. +Canonical Structure union_tag hc := empty_tag hc. + +(* Main structure *) +(* - i : input context *) +(* - j : output context *) +(* - ts : syntactification of pcm using context j *) + +Definition axiom i j ts (pivot : tagged_pcm) := + [/\ interp j ts = Some (untag pivot), sub_ctx i j & all (wf j) ts]. +Structure form i j ts := Form {pivot : tagged_pcm; _ : axiom i j ts pivot}. + +Local Coercion pivot : form >-> tagged_pcm. + +(* check for union *) + +Lemma union_pf i j k ts1 ts2 (f1 : form i j ts1) (f2 : form j k ts2) : + axiom i k (ts1 ++ ts2) (union_tag (untag f1 \+ untag f2)). +Proof. +case: f1 f2 =>[[u1]] /= [E1 S1 W1][[u2]][E2 S2 W2]; split=>/=. +- by rewrite /interp foldr_cat !fE unitL joinC -(sc_interp S2 W1) E1 E2. +- by apply: sc_trans S1 S2. +by rewrite all_cat (sc_wf S2 W1) W2. +Qed. + +Canonical union_form i j k ts1 ts2 (f1 : form i j ts1) (f2 : form j k ts2) := + Form (@union_pf i j k ts1 ts2 f1 f2). + +(* check if reached empty *) + +Lemma empty_pf i : axiom i i [::] (empty_tag Unit). +Proof. by split. Qed. + +Canonical empty_form i := + Form (@empty_pf i). + +(* check for expr *) + +Lemma expr_pf exprs1 exprs2 n (f : xfind exprs1 exprs2 n) : + axiom (Context exprs1) (Context exprs2) + [:: Expr n] (expr_tag (xuntag f)). +Proof. +case: f=>p [E H]; split=>//=; first by rewrite unitR. +by rewrite andbT (onth_size E). +Qed. + +Canonical expr_form exprs1 exprs2 n (f : xfind exprs1 exprs2 n) := + Form (@expr_pf exprs1 exprs2 n f). + +End Syntactify. + +Module Exports. +Coercion untag : tagged_pcm >-> PCM.sort. +Coercion pivot : form >-> tagged_pcm. +Canonical union_tag. +Canonical union_form. +Canonical empty_form. +Canonical expr_form. + +End Exports. +End Syntactify. + +Export Syntactify.Exports. + +(********************) +(* Automating pullX *) +(********************) + +Module PullX. +Section PullX. +Variables (U : pcm). +Implicit Types (j k : ctx U) (ts : seq term). +Notation form := Syntactify.form. +Notation untag := Syntactify.untag. + +Structure packed_pcm (m : U) := Pack {unpack : U}. + +Local Coercion unpack : packed_pcm >-> PCM.sort. + +Canonical equate (m : U) := Pack m m. + +(* j : input context *) +(* k : output context *) +(* ts : (input) the syntactification of subtractee in j *) +(* g : (aux) copy of the goal *) +(* r : (output) the optional residual term *) + +(* We have a choice of leaving the residual in a syntactified form or printing *) +(* it back into a PCM. Here we choose the second option since in HTT we need *) +(* to pass the residual to a partitioning autolemma which operates on heaps. *) + +Definition raxiom j k ts g (r : option U) (pivot : packed_pcm g) := + all (wf j) ts -> r -> sub_ctx j k /\ + Some (unpack pivot) = interp k ts \+ r. + +Structure rform j k ts g r := + RForm {pivot : packed_pcm g; _ : @raxiom j k ts g r pivot}. + +Local Coercion pivot : rform >-> packed_pcm. + +(* start instance: syntactify the goal (by triggering fg), *) +(* run the subtraction on quoted terms and print the residual. *) + +Lemma start_pf j k ts tg (fg : form j k tg) : + @raxiom j k ts (untag fg) + (obind (pprint k \o rev) (subtract tg ts [::])) + (equate fg). +Proof. +case: fg=>fg [Eg S Ag]; case E: (subtract _ _ _)=>[rs|] //= Am _; split=>//. +move/(subtract_sound Ag (sc_wf S Am)): E=>/=. +by rewrite unitR joinC Eg pp_interp interp_rev. +Qed. + +Canonical start j k ts tg (fg : form j k tg) := + RForm (@start_pf j k ts tg fg). + +End PullX. + +Module Exports. +Coercion unpack : packed_pcm >-> PCM.sort. +Coercion pivot : rform >-> packed_pcm. +Canonical equate. +Canonical start. + +Section Exports. +Variables (U : pcm). +Implicit Types (j : ctx U) (ts : seq term). +Notation form := Syntactify.form. +Notation untag := Syntactify.untag. + +(* we need to syntactify first the subtractee (fm), then the goal (fg) *) + +Lemma pullX' s j k ts g r (fs : form (empx U) j ts) + (fg : rform j k ts g (Some r)) : + untag fs = s -> + unpack (pivot fg) = s \+ r. +Proof. +move=><-; case: fg fs; case=>pivot R [fm][E _ A1] /=. +case/(_ A1 erefl): R=>S /=; rewrite -(sc_interp S A1) E. +by case. +Qed. + +End Exports. + +Arguments pullX' [U] s [j k ts g r fs fg] _. +Notation pullX s := (pullX' s erefl). + +Example ex0 (x y z : nat) : + 1 \+ x \+ 2 \+ y \+ 3 \+ z = 0. +Proof. +rewrite [LHS](pullX (3 \+ 2)) /=. +Abort. + +End Exports. +End PullX. + +Export PullX.Exports. + + diff --git a/pcm/heap.v b/pcm/heap.v new file mode 100644 index 0000000..3477902 --- /dev/null +++ b/pcm/heap.v @@ -0,0 +1,573 @@ +(* +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 heaps as possibly undefined finite maps from pointer *) +(* type to dynamic type. *) +(* Heaps are a special case of Partial Commutative Monoids (pcm) *) +(******************************************************************************) + +From HB Require Import structures. +From Coq Require Import ssreflect ssrbool ssrfun Eqdep. +From mathcomp Require Import ssrnat eqtype seq path. +From pcm Require Import options axioms finmap. +From pcm Require Import pcm unionmap natmap. + +(************) +(* Pointers *) +(************) + +(* ptr/null is alternative name for nat/0 *) +Definition ptr := nat. +Definition null : ptr := 0. + +(* some pointer arithmetic: offsetting from a base *) +Definition ptr_offset (x : ptr) i := x + i. + +Notation "x .+ i" := (ptr_offset x i) + (at level 5, format "x .+ i"). + +Lemma ptr0 x : x.+0 = x. +Proof. by rewrite /ptr_offset addn0. Qed. + +Lemma ptr1 x : x .+ 1 = x.+1. +Proof. by rewrite /ptr_offset addn1. Qed. + +Lemma ptrA x i j : x.+i.+j = x.+(i+j). +Proof. by rewrite /ptr_offset addnA. Qed. + +Lemma ptrK x i j : (x.+i == x.+j) = (i == j). +Proof. by rewrite /ptr_offset eqn_add2l. Qed. + +Lemma ptr_null x m : (x.+m == null) = (x == null) && (m == 0). +Proof. by rewrite /ptr_offset addn_eq0. Qed. + +Lemma ptrT x y : {m : nat | (x == y.+m) || (y == x.+m)}. +Proof. +exists (if x <= y then (y - x) else (x - y)). +rewrite /ptr_offset leq_eqVlt /=. +by case: (ltngtP x y)=>/= E; rewrite subnKC ?(ltnW E) ?eq_refl ?orbT // E. +Qed. + +(*********) +(* Heaps *) +(*********) + +Inductive heap := + Undef | Def (finmap : {finMap ptr -> dynamic id}) of + null \notin supp finmap. + +Module Heap. +Section NullLemmas. +Variables (f g : {finMap ptr -> dynamic id}) (x : ptr) (v : dynamic id). + +Lemma upd_nullP : + x != null -> null \notin supp f -> null \notin supp (ins x v f). +Proof. by move=>H1 H2; rewrite supp_ins negb_or /= inE /= eq_sym H1. Qed. + +Lemma free_nullP : null \notin supp f -> null \notin supp (rem x f). +Proof. by move=>H; rewrite supp_rem negb_and /= H orbT. Qed. + +Lemma un_nullP : + null \notin supp f -> null \notin supp g -> + null \notin supp (fcat f g). +Proof. by move=>H1 H2; rewrite supp_fcat negb_or H1 H2. Qed. + +Lemma filt_nullP (q : pred ptr) : + null \notin supp f -> null \notin supp (kfilter q f). +Proof. by move=>H; rewrite supp_kfilt mem_filter negb_and H orbT. Qed. + +Lemma heap_base : null \notin supp f -> all (fun k => k != null) (supp f). +Proof. by move=>H; apply/allP=>k; case: eqP H=>// -> /negbTE ->. Qed. + +Lemma base_heap : all (fun k => k != null) (supp f) -> null \notin supp f. +Proof. by move/allP=>H; apply: (introN idP); move/H. Qed. + +Lemma heapE (h1 h2 : heap) : + h1 = h2 <-> + match h1, h2 with + Def f' pf, Def g' pg => f' = g' + | Undef, Undef => true + | _, _ => false + end. +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: eq_irrelevance. +Qed. + +End NullLemmas. + +(* methods *) + +Notation base := (@UM.base ptr (fun k => k != null) (dynamic id)). + +Definition def h := if h is Def _ _ then true else false. +Definition empty := @Def (finmap.nil _ _) is_true_true. +Definition upd k v h := + if h is Def hs ns then + if decP (@idP (k != null)) is left pf then + Def (@upd_nullP _ _ v pf ns) + else Undef + else Undef. +Definition dom h : seq ptr := if h is Def f _ then supp f else [::]. +Definition assocs h : seq (ptr * 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 := + if h is Def hs _ then fnd x hs else None. +Definition union h1 h2 := + if (h1, h2) is (Def hs1 ns1, Def hs2 ns2) then + if disj hs1 hs2 then + Def (@un_nullP _ _ ns1 ns2) + else Undef + else Undef. +Definition pts (x : ptr) v := upd x v empty. +Definition empb h := if h is Def hs _ then supp hs == [::] else false. +Definition undefb h := if h is Undef then true else false. +Definition keys_of h : seq ptr := + if h is Def f _ then supp f else [::]. +Definition from (f : heap) : base := + if f is Def hs ns then UM.Def (heap_base ns) else UM.Undef _ _. +Definition to (b : base) : heap := + if b is UM.Def hs ns then Def (base_heap ns) else Undef. + +Module Exports. +(* heap has union map structure *) +Lemma heap_is_umc : union_map_axiom def empty Undef upd dom + assocs free find union empb + undefb pts from to. +Proof. +split; first by split=>[[]|[]] // f H; rewrite ?UM.umapE ?heapE. +split. +- split; first by split; [split=>[[]|]|rewrite heapE]. + split; last by case=>[|f H k] //; rewrite heapE. + split=>[|[]] //; split=>[k v [|h H]|[]] //=. + by case: decP=>// H1; rewrite heapE. +split=>[|k v]; last first. +- by rewrite /pts /UM.pts /UM.upd /=; case: decP=>// H; rewrite heapE. +split=>[|[]] //; split=>[|[]] //; split=>[k []|[|f1 H1][|f2 H2]] //. +by rewrite /union /UM.union /=; case: ifP=>D //; rewrite heapE. +Qed. + +HB.instance Definition _ := + isUnion_map.Build ptr (fun k => k != null) (dynamic id) heap heap_is_umc. +HB.instance Definition _ := isNatMap.Build (dynamic id) heap. +End Exports. +End Heap. +Export Heap.Exports. + +Notation "x :-> v" := (ptsT heap x (idyn v)) (at level 30). + +(********************) +(* points-to lemmas *) +(********************) + +(* union_map pts lemmas combined with dyn_inj *) + +Section HeapPointsToLemmas. +Implicit Types (x : ptr) (h : heap). + +Lemma hcancelPtT A1 A2 x (v1 : A1) (v2 : A2) : + valid (x :-> v1) -> x :-> v1 = x :-> v2 -> A1 = A2. +Proof. by move=>V /(cancelPt V)/dyn_injT. Qed. + +Lemma hcancelPtT2 A1 A2 x1 x2 (v1 : A1) (v2 : A2) : + valid (x1 :-> v1) -> x1 :-> v1 = x2 :-> v2 -> (x1, A1) = (x2, A2). +Proof. by move=>V; case/(cancelPt2 V)=>-> E _; rewrite E. Qed. + +Lemma hcancelPtV A x (v1 v2 : A) : + valid (x :-> v1) -> x :-> v1 = x :-> v2 -> v1 = v2. +Proof. by move=>V; move/(cancelPt V)/dyn_inj. Qed. + +Lemma hcancelPtV2 A x1 x2 (v1 v2 : A) : + valid (x1 :-> v1) -> x1 :-> v1 = x2 :-> v2 -> (x1, v1) = (x2, v2). +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 h x. +Proof. by case/um_eta; case=>A v H; exists A, v. Qed. + +Lemma hcancelT A1 A2 x (v1 : A1) (v2 : A2) h1 h2 : + valid (x :-> v1 \+ h1) -> + x :-> v1 \+ h1 = x :-> v2 \+ h2 -> A1 = A2. +Proof. by move=>V; case/(cancel V); move/dyn_injT. Qed. + +Lemma hcancelV A x (v1 v2 : A) h1 h2 : + valid (x :-> v1 \+ h1) -> + x :-> v1 \+ h1 = x :-> v2 \+ h2 -> [/\ v1 = v2, valid h1 & h1 = h2]. +Proof. by move=>V; case/(cancel V); move/dyn_inj. Qed. + +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 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. + +End HeapPointsToLemmas. + + +(*******************************************) +(* Block update for reasoning about arrays *) +(*******************************************) + +Section BlockUpdate. +Variable (A : Type). + +Fixpoint updi (x : ptr) (vs : seq A) {struct vs} : heap := + if vs is v'::vs' then x :-> v' \+ updi x.+1 vs' else Unit. + +Lemma updiS x v vs : updi x (v :: vs) = x :-> v \+ updi x.+1 vs. +Proof. by []. Qed. + +Lemma updi_last x v vs : + updi x (rcons vs v) = updi x vs \+ x.+(size vs) :-> v. +Proof. +elim: vs x v=>[|w vs IH] x v /=. +- by rewrite ptr0 unitR unitL. +by rewrite -(addn1 (size vs)) addnC -ptrA IH joinA ptr1. +Qed. + +Lemma updi_cat x vs1 vs2 : + updi x (vs1 ++ vs2) = updi x vs1 \+ updi x.+(size vs1) vs2. +Proof. +elim: vs1 x vs2=>[|v vs1 IH] x vs2 /=. +- by rewrite ptr0 unitL. +by rewrite -(addn1 (size vs1)) addnC -ptrA IH joinA ptr1. +Qed. + +Lemma updi_catI x y vs1 vs2 : + y = x + size vs1 -> updi x vs1 \+ updi y vs2 = updi x (vs1 ++ vs2). +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; first by rewrite dom0. +rewrite /= -addnS domPtUn inE /= negb_and negb_or -{4}(addn0 x). +by rewrite eqn_add2l -lt0n H IH // andbT orbT. +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 valid_unit orbC. +by rewrite /= validPtUn -addn1 updiVm' // orbF IH addn1 /= andbT. +Qed. + +Lemma updiVm x m xs : + x \in dom (updi x.+m xs) = [&& x != null, m == 0 & size xs > 0]. +Proof. +case: m=>[|m] /=; last first. +- by rewrite andbF; apply/negbTE/updiVm'. +case: xs=>[|v xs]; rewrite ptr0 ?andbF ?andbT ?dom0 //=. +by rewrite domPtUn inE /= eq_refl -updiS updiD orbF andbT /=. +Qed. + +Lemma updimV x m xs : + x.+m \in dom (updi x xs) = (x != null) && (m < size xs). +Proof. +case H: (x == null)=>/=. +- 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 ptr1 IH //. +by rewrite -updiS updiD H /= -{1}(ptr0 x) -ptr1 ptrA ptrK. +Qed. + +Lemma updiP x y xs : + reflect (y != null /\ exists m, x = y.+m /\ m < size xs) + (x \in dom (updi y xs)). +Proof. +case H: (y == null)=>/=. +- 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=>->. +- by rewrite updimV H /= => H1; split=>//; exists m. +rewrite updiVm; case/and3P=>H1; move/eqP=>-> H2. +by split=>//; exists 0; rewrite ptrA addn0 ptr0. +Qed. + +(* Invertibility *) +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. +- by case/esym/umap0E=>/unitbP; rewrite um_unitbPt. +- by case/umap0E=>/unitbP; rewrite um_unitbPt. +by case/(hcancelV D)=><- {}D /(IH _ _ D) <-. +Qed. + +Lemma updi_iinv x xs1 xs2 h1 h2 : + size xs1 = size xs2 -> valid (updi x xs1 \+ h1) -> + updi x xs1 \+ h1 = updi x xs2 \+ h2 -> 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; case/(hcancelV D)=><-{}D. +by case/(IH _ _ _ _ E D)=>->->. +Qed. + +End BlockUpdate. + +Lemma domeqUP A1 A2 x (xs1 : seq A1) (xs2 : seq A2) : + size xs1 = size xs2 -> dom_eq (updi x xs1) (updi x xs2). +Proof. +move=>E; apply/domeqP; split; first by rewrite !updiD E. +apply/domE=>z; case: updiP=>[[H][m][->]|X]; first by rewrite updimV H E. +by case: updiP=>// [[H]][m][Ez S]; elim: X; split=>//; exists m; rewrite Ez E. +Qed. + +(*****************************) +(*****************************) +(* Automation of PtUn lemmas *) +(*****************************) +(*****************************) + +(* First, the mechanism for search-and-replace for the overloaded lemas, *) +(* pattern-matching on heap expressions. *) + +Structure tagged_heap := HeapTag {untag :> heap}. + +Definition right_tag := HeapTag. +Definition left_tag := right_tag. +Canonical found_tag i := left_tag i. + +Definition partition_axiom k r (h : tagged_heap) := untag h = k \+ r. + +Structure partition (k r : heap) := + Partition {heap_of :> tagged_heap; + _ : partition_axiom k r heap_of}. + +Lemma partitionE r k (f : partition k r) : untag f = k \+ r. +Proof. by case: f=>[[j]] /=; rewrite /partition_axiom /= => ->. Qed. + +Lemma found_pf k : partition_axiom k Unit (found_tag k). +Proof. by rewrite /partition_axiom unitR. Qed. + +Canonical found_struct k := Partition (found_pf k). + +Lemma left_pf h r (f : forall k, partition k r) k : + partition_axiom k (r \+ h) (left_tag (untag (f k) \+ h)). +Proof. by rewrite partitionE /partition_axiom /= joinA. Qed. + +Canonical left_struct h r (f : forall k, partition k r) k := + Partition (left_pf h f k). + +Lemma right_pf h r (f : forall k, partition k r) k : + partition_axiom k (h \+ r) (right_tag (h \+ f k)). +Proof. by rewrite partitionE /partition_axiom /= joinCA. Qed. + +Canonical right_struct h r (f : forall k, partition k r) k := + Partition (right_pf h f k). + +(* now the actual lemmas *) + +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}. + +Lemma defPt_nullO A h x (v : A) (f : partition (x :-> v) h) : + valid (untag f) -> x != null. +Proof. by rewrite partitionE; apply: validPtUn_cond. Qed. + +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 => /validR. Qed. + +Arguments defPt_defO [A][h] x [v][f] _. + +Lemma defPt_domO A h x (v : A) (f : partition (x :-> v) h) : + valid (untag f) -> x \notin dom h. +Proof. by rewrite partitionE; apply: validPtUnD. Qed. + +Arguments defPt_domO [A][h] x [v][f] _. + +Lemma domPtUnO A h x (v : A) (f : partition (x :-> v) h) : + dom (untag f) =i + [pred y | valid (untag f) & (x == y) || (y \in dom h)]. +Proof. by rewrite partitionE; apply: domPtUn. Qed. + +Arguments domPtUnO [A][h] x [v][f] _. + +Lemma lookPtUnO A h x (v : A) (f : partition (x :-> v) h) : + valid (untag f) -> find x (untag f) = Some (idyn v). +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 (untag f) x = h. +Proof. by rewrite partitionE; apply: freePtUn. Qed. + +Arguments freePtUnO [A h x v f] _. + +Lemma updPtUnO A1 A2 x i (v1 : A1) (v2 : A2) + (f : forall k, partition k i) : + upd x (idyn v2) (untag (f (x :-> v1))) = f (x :-> v2). +Proof. by rewrite !partitionE; apply: updPtUn. Qed. + +Arguments updPtUnO [A1 A2 x i v1 v2] f. + +Lemma cancelTO A1 A2 h1 h2 x (v1 : A1) (v2 : A2) + (f1 : partition (x :-> v1) h1) + (f2 : partition (x :-> v2) h2) : + valid (untag f1) -> f1 = f2 :> heap -> A1 = A2. +Proof. by rewrite !partitionE; apply: hcancelT. Qed. + +Arguments cancelTO [A1 A2 h1 h2] x [v1 v2 f1 f2] _ _. + +Lemma cancelO A h1 h2 x (v1 v2 : A) + (f1 : partition (x :-> v1) h1) + (f2 : partition (x :-> v2) h2) : + valid (untag f1) -> f1 = f2 :> heap -> + [/\ v1 = v2, valid h1 & h1 = h2]. +Proof. by rewrite !partitionE; apply: hcancelV. Qed. + +Arguments cancelO [A h1 h2] x [v1 v2 f1 f2] _ _. + +Lemma domPtUnXO A (v : A) x i (f : partition (x :-> v) i) : + valid (untag f) -> x \in dom (untag f). +Proof. by rewrite partitionE domPtUnE. Qed. + +(*******************************************************) +(*******************************************************) +(* Custom lemmas about singly-linked lists in the heap *) +(*******************************************************) +(*******************************************************) + +Fixpoint llist A p (xs : seq A) {struct xs} := + if xs is x::xt then + fun h => exists r h', h = p :-> (x, r) \+ h' /\ llist r xt h' + else fun h : heap => p = null /\ h = Unit. + +Lemma llist_prec A p (l1 l2 : seq A) h1 h2 g1 g2 : + valid (h1 \+ g1) -> + llist p l1 h1 -> llist p l2 h2 -> + h1 \+ g1 = h2 \+ g2 -> + [/\ l1 = l2, h1 = h2 & g1 = g2]. +Proof. +move=>V H1 H2 E; elim: l1 l2 h1 h2 p H1 H2 E V=>[|a1 l1 IH]. +- case=>[|a2 l2] h1 h2 p /= [->->]. + - by case=>_ ->; rewrite !unitL. + by case=>r [h'][-> _ ->] /validL /validPtUn_cond. +case=>[|a2 l2] h1 h2 p /= [p1][k1][-> H1]. +- by case=>->-> _ /validL /validPtUn_cond. +case=>p2 [k2][-> H2]; rewrite -!joinA => E V. +case: {E V} (hcancelV V E) H1 H2; case=>->-> V E H1 H2. +by case: (IH _ _ _ _ H1 H2 E V)=>->->->. +Qed. + +Lemma first_exists A p h (ls : seq A) : + p != null -> llist p ls h -> + exists x r h', + [/\ ls = x :: behead ls, h = p :-> (x, r) \+ h' & + llist r (behead ls) h']. +Proof. +case: ls=>[|a ls] /= H []; first by case: eqP H. +by move=>r [h'][-> H1]; eexists a, r, h'. +Qed. + +Lemma llist_null A (xs : seq A) h : + valid h -> llist null xs h -> xs = [::] /\ h = Unit. +Proof. +case: xs=>[|x xs] /= V H; first by case: H. +by case: H V=>p [h'][-> _] /validPtUn_cond. +Qed. + +(******************************) +(******************************) +(* Custom lemmas about arrays *) +(******************************) +(******************************) + +From mathcomp Require Import fintype tuple finfun. + +Definition indx {I : finType} (x : I) := index x (enum I). + +Prenex Implicits indx. + +(***********************************) +(* Arrays indexed by a finite type *) +(***********************************) + +Section Array. +Variables (p : ptr) (I : finType). + +(* helper lemmas *) + +Lemma enum_split k : + enum I = take (indx k) (enum I) ++ k :: drop (indx k).+1 (enum I). +Proof. +rewrite -{2}(@nth_index I k k (enum I)) ?mem_enum //. +by rewrite -drop_nth ?index_mem ?mem_enum // cat_take_drop. +Qed. + +Lemma updi_split T k (f : {ffun I -> T}) : + updi p (fgraph f) = updi p (take (indx k) (fgraph f)) \+ + p.+(indx k) :-> f k \+ + updi (p.+(indx k).+1) (drop (indx k).+1 (fgraph f)). +Proof. +rewrite fgraph_codom /= codomE {1}(enum_split k) map_cat updi_cat /=. +rewrite map_take map_drop size_takel ?joinA; first by rewrite -ptr1 ptrA addn1. +by rewrite size_map index_size. +Qed. + +Lemma takeord T k x (f : {ffun I -> T}) : + take (indx k) (fgraph [ffun y => [eta f with k |-> x] y]) = + take (indx k) (fgraph f). +Proof. +set f' := (finfun _). +suff E: {in take (indx k) (enum I), f =1 f'}. +- by rewrite !fgraph_codom /= !codomE -2!map_take; move/eq_in_map: E. +move: (enum_uniq I); rewrite {1}(enum_split k) cat_uniq /= =>H4. +move=>y H5; rewrite /f' /= !ffunE /=; case: eqP H5 H4=>// -> ->. +by rewrite andbF. +Qed. + +Lemma dropord T k x (f : {ffun I -> T}) : + drop (indx k).+1 (fgraph [ffun y => [eta f with k |->x] y]) = + drop (indx k).+1 (fgraph f). +Proof. +set f' := (finfun _). +suff E: {in drop (indx k).+1 (enum I), f =1 f'}. +- by rewrite !fgraph_codom /= !codomE -2!map_drop; move/eq_in_map: E. +move: (enum_uniq I); rewrite {1}(enum_split k) cat_uniq /= => H4. +move=>y H5; rewrite /f' /= !ffunE /=; case: eqP H5 H4=>// -> ->. +by rewrite !andbF. +Qed. + +Lemma size_fgraph T1 T2 (r1 : {ffun I -> T1}) (r2 : {ffun I -> T2}) : + size (fgraph r1) = size (fgraph r2). +Proof. by rewrite !fgraph_codom /= !codomE !size_map. Qed. + +Lemma fgraphE T (r1 r2 : {ffun I -> T}) : + fgraph r1 = fgraph r2 -> r1 = r2. +Proof. +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..89e6677 --- /dev/null +++ b/pcm/invertible.v @@ -0,0 +1,124 @@ +(* +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 pcm Require Import options axioms prelude. +From pcm Require Import pcm morphism. + +(* in the POPL21 paper, we use the notation + x D y D z =def= x D y /\ (x \+ y) D z +*) +Definition inv_rel (U : pcm) (D : rel U) := + 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 (U V : pcm) (f : pcm_morph U V) := + forall (a : U) (b1 b2 : V), + valid a -> sep f a Unit -> f a = b1 \+ b2 -> + exists a1 a2, [/\ a = a1 \+ a2, valid (a1 \+ a2), + sep f a1 a2, f a1 = b1 & f a2 = b2 ]. + +Definition inv_morph (U V : pcm) (f : pcm_morph U V) := + inv_rel (sep f) /\ inv_morph_axiom f. + + +(* trivial seprel in invertible *) +Lemma inv_sepT (U : pcm) : inv_rel (@relT U). Proof. by []. Qed. + +Lemma inv_idfun (U : pcm) : inv_morph (@idfun U). +Proof. by split=>// a b1 b2 V _ E; exists b1, b2; rewrite -E. Qed. + +(* composition of invertible morphisms is invertible *) + +Lemma inv_comp (U V W : pcm) + (f : pcm_morph U V) (g : pcm_morph V W) : + inv_morph f -> inv_morph g -> inv_morph (g \o f). +Proof. +case=>Sf Ff [Sg Fg]; split=>[a1 a2 a' V1|a b1 b2 V1] /andP [] /=. +- move/(Sf _ _ _ V1)=>H1 H2 /andP [/H1] /andP [D1 D2] /=. + rewrite /sepx/= !pfjoin ?(validLE3 V1,sepAxx V1 D1 D2) //= in H2 *. + by apply: Sg H2; apply: pfV3. +move=>D1; rewrite pfunit=>C1. +case/(Fg _ _ _ (pfV _ _ V1 D1) C1)=>c1 [c2][Ef Vc Xc <-<-]. +case/(Ff _ _ _ V1 D1): Ef Xc=>a1 [a2][-> Va Xc <-<- Xd]. +by exists a1, a2; rewrite /sepx/= Xc Xd. +Qed. + +(* morphism product of invertibles is invertible *) +Lemma inv_cprod (U1 U2 V1 V2 : pcm) + (f : pcm_morph U1 V1) (g : pcm_morph U2 V2) : + inv_morph f -> inv_morph g -> inv_morph (f \* g). +Proof. +case=>Sf Ff [Sg Fg]; split=>[a1 a2 a'|] /=. +- rewrite /rel_prod=>/andP [Va Vb] /andP [/(Sf _ _ _ Va) H1 /(Sg _ _ _ Vb) H2]. + by rewrite /sepx/=; case/andP=>/H1/andP [->->] /H2/andP [->->]. +move=>/= a1 b1 b2 /andP [Va Vb] /andP [H1 H2][]. +case/(Ff _ _ _ Va H1)=>c1 [c2][E1 Vc1 Y1 Fc1 Fc2]. +case/(Fg _ _ _ Vb H2)=>d1 [d2][E2 Vd1 Y2 Fd1 Fd2]. +exists (c1, d1), (c2, d2); rewrite /valid /= Vc1 Vd1. +rewrite /sepx/=/rel_prod/fprod /= Fc1 Fc2 Fd1 Fd2 Y1 Y2. +by rewrite pcmPE -E1 -E2 -!prod_eta. +Qed. + +(* ker requires only that sep f is invertible (not that f itself is) *) +Lemma inv_ker (U V : pcm) (f : pcm_morph U V) : + inv_rel (sep f) -> inv_rel (ker f). +Proof. +move=>Sf a1 a2 a' W /and3P [D1 /unitbP Eq1 _] /and3P [D2 /unitbP Eq2 /unitbP]. +case/andP: (Sf _ _ _ W D1 D2)=>{}D1 {}D2. +rewrite /kerx !pfjoin ?(validLE3 W) //; last by rewrite (sepAxx W). +by rewrite /kerx Eq1 Eq2 D1 D2 !unitL=>->; rewrite /sepU pcmE. +Qed. + +(* equalizer is invertible only for cancellative ranges *) +Lemma inv_eqlz (U : pcm) (V : eqcpcm) + (f : pcm_morph U V) (g : pcm_morph U V) : + inv_rel (sep f) -> inv_rel (sep g) -> inv_rel (eqlz f g). +Proof. +move=>Sf Sg x y z W; rewrite /eqlzx. +case/and4P=>X1 X2 Ex _ /and4P [Y1 Y2] Ey /eqP. +case/andP: (Sf _ _ _ W X1 Y1)=>E1 E1'. +case/andP: (Sg _ _ _ 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 => /joinxK ->; first by rewrite eqxx. +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 V : pcm) (S : subpcm_struct U V) + (W : pcm) (f : pcm_morph V W) : + subsep S =2 sep f -> + inv_morph f -> inv_morph (f \o pval S). +Proof. +move=>E [Sf Hf]; split=>[a1 a2 a' V1|a b1 b2 V1] /=; +rewrite /sepx/= !pfT /=. +- by rewrite !pfjoin ?(validLE3 V1) //; apply: Sf; apply: pfV3. +rewrite pfunit /= => H1 /(Hf _ _ _ (pfVI V1) H1). +case=>c1 [c2][Eq Vc H2 F1 F2]; exists (psub S c1), (psub S c2). +rewrite pfT /= -!pfjoin ?E //= -Eq psub_pval //=. +by rewrite !pval_psub ?(validL Vc,validR Vc) // E (sep0E Vc). +Qed. + + diff --git a/pcm/morphism.v b/pcm/morphism.v new file mode 100644 index 0000000..b6a4a8b --- /dev/null +++ b/pcm/morphism.v @@ -0,0 +1,3324 @@ +(* +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 HB Require Import structures. +From Coq Require Import ssreflect ssrbool ssrfun. +From mathcomp Require Import ssrnat eqtype fintype finfun. +From pcm Require Import options pred axioms prelude. +From pcm Require Import pcm. + +(*****************) +(*****************) +(* PCM Morphisms *) +(*****************) +(*****************) + +(* Orthogonality relations (in the POPL21 paper called separating relations) *) +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)]. + +(* Here, by separating relation we take a somewhat different definition *) +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)]. + +(* The two definitions differ in that seprel elide the validity requirement. *) +(* This is pragmatic distinction, because in practice we often have *) +(* the validity in context so we don't have to bother proving it. *) +(* 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 relation of a morphism f *) +(* we will write x \orth_f y iff valid (x \+ y) and sep f x y *) +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 sperels 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. + +HB.mixin Record isSeprel (U : pcm) (sep : rel U) := { + seprel_subproof : seprel_axiom sep}. + +#[short(type="seprel")] +HB.structure Definition Seprel (U : pcm) := {sep of isSeprel U sep}. + +Section Repack. +Variables (U : pcm) (sep : seprel U). + +Lemma sep00 : sep Unit Unit. +Proof. by case: (@seprel_subproof U sep). Qed. + +Lemma sepC x y : + valid (x \+ y) -> + sep x y = sep y x. +Proof. by case: (@seprel_subproof U sep)=>_ H _ _; apply: H. Qed. + +Lemma sepx0 x y : + valid (x \+ y) -> + sep x y -> + sep x Unit. +Proof. by case: (@seprel_subproof U sep)=>_ _ 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. +case: (@seprel_subproof U sep)=>_ _ _ H V R1 R2. +by 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; rewrite -!(@sepC y) ?unitR ?(validE2 V) //. +by 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 Repack. + + +(***********************************) +(* Sep relations and constructions *) +(***********************************) + +(* always-true relation *) + +Definition relT {U} : rel U := fun x y => true. + +Lemma relT_is_seprel {U : pcm} : seprel_axiom (@relT U). Proof. by []. Qed. +HB.instance Definition _ U := isSeprel.Build U relT relT_is_seprel. + + +(* always-unit relation *) +(* always-false relation is not seprel, because we need sep0 Unit Unit *) +(* i.e., sepU is really the smallest seprel we can have *) + +Definition sepU {U : pcm} : rel U := fun x y => unitb x && unitb y. + +Section SepU. +Variable U : pcm. + +Lemma sepU_is_seprel : seprel_axiom (@sepU U). +Proof. +rewrite /sepU; split=>[|x y|x y|x y z]. +- by rewrite pcmE. +- by rewrite andbC. +- by move=>V /andP [->]; rewrite pcmE. +move=>V /andP [/unitbP -> /unitbP ->] /andP [_ /unitbP ->]. +by rewrite unitL !pcmE. +Qed. + +HB.instance Definition _ := isSeprel.Build U sepU sepU_is_seprel. +End SepU. + + +(************************************) +(* Conjunction of seprels is seprel *) +(************************************) + +(* first conjunction of plain rels *) +Definition relI U (X Y : rel U) (x y : U) := X x y && Y x y. +(* and iterated variants *) +Definition rel3I U (X Y Z : rel U) (x y : U) := + [&& X x y, Y x y & Z x y]. +Definition rel4I U (X Y Z W : rel U) (x y : U) := + [&& X x y, Y x y, Z x y & W x y]. + +Arguments relI {U} X Y x y /. +Arguments rel3I {U} X Y Z x y /. +Arguments rel4I {U} X Y Z W x y /. + +Section SepI. +Variables (U : pcm) (X Y : seprel U). + +Lemma relI_is_seprel : seprel_axiom (relI X Y). +Proof. +split=>[|x y|x y|x y z] /=. +- by rewrite !sep00. +- by move=>V; rewrite !(sepC _ V). +- by move=>V /andP []; do ![move/(sepx0 V) ->]. +move=>V /andP [X1 Y1] /andP [X2 Y2]. +by rewrite !(sepAxx V X1 X2, sepAxx V Y1 Y2). +Qed. + +HB.instance Definition _ := isSeprel.Build U (relI X Y) relI_is_seprel. +End SepI. + +(* 3-way conjunction *) +Section Sep3I. +Variables (U : pcm) (X Y Z : seprel U). + +Lemma rel3I_is_seprel : seprel_axiom (rel3I X Y Z). +Proof. +split=>[|x y|x y|x y z] /=. +- by rewrite !sep00. +- by move=>V; rewrite !(sepC _ V). +- by move=>V /and3P []; do ![move/(sepx0 V) ->]. +move=>V /and3P [X1 Y1 Z1] /and3P [X2 Y2 Z2]. +by rewrite !(sepAxx V X1 X2, sepAxx V Y1 Y2, sepAxx V Z1 Z2). +Qed. + +HB.instance Definition _ := isSeprel.Build U (rel3I X Y Z) rel3I_is_seprel. +End Sep3I. + +(* 4-way conjunction *) +Section Sep4I. +Variables (U : pcm) (X Y Z W : seprel U). + +Lemma rel4I_is_seprel : seprel_axiom (rel4I X Y Z W). +Proof. +split=>[|x y|x y|x y z] /=. +- by rewrite !sep00. +- by move=>V; rewrite !(sepC _ V). +- by move=>V /and4P []; do ![move/(sepx0 V) ->]. +move=>V /and4P [X1 Y1 Z1 W1] /and4P [X2 Y2 Z2 W2]. +by rewrite !(sepAxx V X1 X2, sepAxx V Y1 Y2, sepAxx V Z1 Z2, sepAxx V W1 W2). +Qed. + +HB.instance Definition _ := isSeprel.Build U (rel4I X Y Z W) rel4I_is_seprel. +End Sep4I. + +(************************************) +(* projections and pairwise product *) +(************************************) + +Definition rel_fst U V (X : rel U) (x y : U * V) := X x.1 y.1. +Arguments rel_fst {U V} X x y /. +Definition rel_snd U V (Y : rel V) (x y : U * V) := Y x.2 y.2. +Arguments rel_snd {U V} Y x y /. + +Section SepFst. +Variables (U V : pcm) (X : seprel U). +Lemma relfst_is_seprel : seprel_axiom (@rel_fst U V X). +Proof. +split=>[|x y|x y|x y z] //=. +- by rewrite !sep00. +- by case/andP=>V1 _; rewrite sepC. +- by case/andP=>V1 _ /(sepx0 V1). +by case/andP=>V1 V2 X1 Y1; rewrite !(sepAxx V1 X1 Y1). +Qed. +HB.instance Definition _ := + isSeprel.Build (U * V)%type (rel_fst X) relfst_is_seprel. +End SepFst. + +Section SepSnd. +Variables (U V : pcm) (Y : seprel V). +Lemma relsnd_is_seprel : seprel_axiom (@rel_snd U V Y). +Proof. +split=>[|x y|x y|x y z] //=. +- by rewrite !sep00. +- by case/andP=>_ V2; rewrite sepC. +- by case/andP=>_ V2 /(sepx0 V2). +by case/andP=>V1 V2 X2 Y2; rewrite !(sepAxx V2 X2 Y2). +Qed. +HB.instance Definition _ := + isSeprel.Build (U * V)%type (rel_snd Y) relsnd_is_seprel. +End SepSnd. + +Definition rel_prod U V (X : rel U) (Y : rel V) := + relI (rel_fst X) (rel_snd Y). +Arguments rel_prod {U V} X Y /. + +HB.instance Definition _ (U V : pcm) (X : seprel U) (Y : seprel V) := + Seprel.copy (rel_prod X Y) (rel_prod X Y). + +(**********************) +(**********************) +(* Classes of seprels *) +(**********************) +(**********************) + +(* Non-symmetric seprels *) + +(* Often times, we build a seprel out of a non-symmetric relation *) +(* by taking the symmetric closure of the relation explicitly. *) +(* This is such a common approach, that its useful *) +(* to introduce a class of primitive non-symmetric seprels that *) +(* are made full seprels by closing up. *) + +Definition nseprel_axiom (U : pcm) (nsep : rel U) := + [/\ (* unit is separated from unit *) + nsep Unit Unit, + (* 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) -> nsep x y -> nsep x Unit, + (* if the first arguments is Unit, then nsep trivially holds *) + forall y, valid y -> nsep Unit y, + (* associativity 1 *) + forall x y z, valid (x \+ y \+ z) -> + nsep x y -> nsep (x \+ y) z -> nsep x (y \+ z) & + (* associativity 2 *) + forall x y z, valid (x \+ y \+ z) -> + nsep x (y \+ z) -> nsep y (x \+ z) -> nsep x y /\ nsep (x \+ y) z]. + +Lemma orth_nsep (U : pcm) (nsep : rel U) : + nseprel_axiom nsep -> + seprel_axiom (fun x y => nsep x y && nsep y x). +Proof. +case=>H1 H2 H3 H4 H5; split=>[|x y V|x y V|x y z V]. +- by rewrite H1. +- by rewrite andbC. +- by case/andP=>/(H2 _ _ V) -> _; rewrite (H3 _ (validL V)). +case/andP=>X1 X2 /andP [X3 X4]. +move: (H4 x y z V X1 X3)=>X5; rewrite X5 /=. +rewrite joinC in X3. +move: (H4 y x z); rewrite (validLE3 V)=>/(_ erefl X2 X3) X6. +rewrite joinC in X4; rewrite joinC in X6. +move: (H5 y z x); rewrite (validLE3 V)=>/(_ erefl X6 X4) [->->] /=. +by move: (H5 z y x); rewrite (validLE3 V)=>/(_ erefl X4 X6) [] ->. +Qed. + +(* Guarded non-symmetric seprels *) + +(* Further optimization, if the nsep has the form forall k, P k x -> Q k x y *) +(* for some P and Q *) + +Section GuardedSeprels. +Variable (U : pcm) (X : Type) (P : X -> U -> Prop) (Q : X -> U -> U -> Prop). +Variable (nsep : rel U). + +Hypothesis pf1 : forall x y, reflect (forall k, P k x -> Q k x y) (nsep x y). + +Definition gseprel_axiom := + [/\ (* P is disjunctive *) + forall k x y, valid (x \+ y) -> P k (x \+ y) <-> P k x \/ P k y, + (* unit is separated from unit *) + forall k, P k Unit -> Q k Unit Unit, + (* is x is in the domain (i.e., x is considered) *) + (* then it is separate from at least Unit *) + forall k x y, valid (x \+ y) -> P k x -> Q k x y -> Q k x Unit, + (* if the first argument is Unit, then nsep trivially holds *) + forall k y, valid y -> P k Unit -> Q k Unit y, + (* associativity 1 *) + forall k x y z, valid (x \+ y \+ z) -> + P k x -> Q k x y -> Q k (x \+ y) z -> Q k x (y \+ z) & + (* associativity 2 *) + forall k x y z, valid (x \+ y \+ z) -> + P k x -> Q k x (y \+ z) -> Q k x y /\ Q k (x \+ y) z]. + +Lemma orth_gsep : + gseprel_axiom -> + seprel_axiom (fun x y => nsep x y && nsep y x). +Proof. +case=>H1 H2 H3 H4 H5 H6; apply: orth_nsep. +split=>[|x y V|y V|x y z V|x y z V]. +- by apply/pf1. +- by move/pf1=>H; apply/pf1=>k /[dup] Y /H; apply: H3 V Y. +- by apply/pf1=>x; apply: H4. +- move/pf1=>X1 /pf1 X2; apply/pf1=>k H. + apply: H5=>//; first by apply: X1 H. + by apply: X2; apply/H1; [rewrite (validLE3 V)|left]. +move/pf1=>X1 /pf1 X2; split; apply/pf1=>k H. +- by case: (H6 k x y z V H (X1 _ H)). +case/H1: H; first by rewrite (validLE3 V). +- by move=>H; case: (H6 k x y z V H (X1 _ H)). +move=>H; rewrite joinC. +case: (H6 k y x z)=>//; first by rewrite (validLE3 V). +by apply: X2 H. +Qed. + +End GuardedSeprels. + +(* Further optimization *) +(* Most of the time, the guard is not only disjunctive, but *) +(* exclusively disjunctive. This allows to weaken the conditions. *) +(* We have to prove a couple of additional side-conditions however *) +(* so we'll see how useful this is. *) +(* I wasn't able to find a simpler formulation. *) + +Section ExclusiveGuardedSeprels. +Variable (U : pcm) (X : Type) (P : X -> U -> Prop) (Q : X -> U -> U -> Prop). +Variable (nsep : rel U). + +Hypothesis pf1 : forall x y, reflect (forall k, P k x -> Q k x y) (nsep x y). + +Definition xgseprel_axiom := + [/\ (* P is disjunctive *) + forall k x y, valid (x \+ y) -> P k (x \+ y) <-> P k x \/ P k y, + (* P is exclusive *) + forall k1 k2 x y, valid (x \+ y) -> P k1 x -> P k2 y -> k1 <> k2, + forall k1 k2 x y z, valid (x \+ y \+ z) -> + P k1 x -> Q k1 x y -> Q k2 (x \+ y) z -> k1 = k2, + forall k1 k2 x y z, valid (x \+ y \+ z) -> + P k1 x -> Q k1 x (y \+ z) -> P k2 y-> Q k2 y (x \+ z) -> k1 = k2, + (* unit is separated from unit *) + forall k, P k Unit -> Q k Unit Unit, + (* is x is in the domain (i.e., x is considered) *) + (* then it is separate from at least Unit *) + forall k x y, valid (x \+ y) -> P k x -> Q k x y -> Q k x Unit, + (* if the first argument is Unit, then nsep trivially holds *) + forall k y, valid y -> P k Unit -> Q k Unit y, + (* associativity 1 *) + forall k x y z, valid (x \+ y \+ z) -> + P k x -> (forall k', ~ P k' y) -> + Q k x y -> Q k (x \+ y) z -> Q k x (y \+ z) & + (* associativity 2 *) + forall k x y z, valid (x \+ y \+ z) -> + P k x -> (forall k', ~ P k' y) -> + Q k x (y \+ z) -> Q k x y /\ Q k (x \+ y) z]. + +Lemma xorth_gsep : + xgseprel_axiom -> + seprel_axiom (fun x y => nsep x y && nsep y x). +Proof. +case=>H1 Xp Xq1 Xq2 H2 H3 H4 H5 H6; apply: orth_nsep. +split=>[|x y V|y V|x y z V|x y z V]. +- by apply/pf1. +- by move/pf1=>H; apply/pf1=>k /[dup] Y /H; apply: H3 V Y. +- by apply/pf1=>x; apply: H4. +- have V' : valid (x \+ y) by rewrite (validLE3 V). + move/pf1=>X1 /pf1 X2; apply/pf1=>k Hx; apply: (H5)=>//; last first. + - by apply: X2; apply/H1=>//; left. + - by apply: X1. + move=>k' Hy; apply: Xp (V') (Hx) (Hy) _. + have Hxy : P k' (x \+ y) by rewrite H1 //; right. + by apply: Xq1 V (Hx) (X1 _ Hx) (X2 _ Hxy). +have V' : valid (x \+ y) by rewrite (validLE3 V). +move/pf1=>X1 /pf1 X2; split; apply/pf1=>k Hx. +- case: (H6 k x y z V Hx)=>//; last by apply: X1. + move=>k' Hy; apply: Xp V' (Hx) (Hy) _. + by apply: Xq2 (X1 _ Hx) _ (X2 _ Hy). +case/H1: Hx; first by rewrite (validLE3 V). +- move=>Hx; case: (H6 k x y z V Hx)=>//; last by apply: X1. + by move=>k' Hy; apply: Xp V' (Hx) (Hy) _; apply: Xq2 (X1 _ Hx) _ (X2 _ Hy). +move=>Hy; rewrite joinC. +case: (H6 k y x z)=>//; first by rewrite (validLE3 V). +- by move=>k' Hx; apply: Xp V' (Hx) (Hy) _; apply: Xq2 (X1 _ Hx) _ (X2 _ Hy). +by apply: X2. +Qed. + +End ExclusiveGuardedSeprels. + + +(*****************) +(*****************) +(* PCM morphisms *) +(*****************) +(*****************) + +(* morphism comes with a seprel on which it is valid *) + +Definition pcm_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]. + +HB.mixin Record isPCM_morphism (U V : pcm) (f : U -> V) := { + sep_op : seprel U; + pcm_morph_subproof : pcm_morph_axiom sep_op f}. + +#[short(type=pcm_morph)] +HB.structure Definition PCM_morphism (U V : pcm) := + {f of isPCM_morphism U V f}. + +Arguments sep_op {U V} _ /. +Arguments pcm_morph_subproof {U V}. + +(* When sep_op is applied to function f, the system infers *) +(* the pcm_morph structure F of f, as intended. *) +(* However, it then proceeds to display sep_op F instead of sep_op f, *) +(* which is often unreadable, as F is typically large. *) +(* The following arranges that f is printed instead of F. *) + +(* There are two options that differ in issues orthogonal to f/F. *) +(* 1. sepx is declared seprel *) +(* 2. sepx is declared rel, with canonical seprel attached *) +(* Option 1 automatically simplifies nested sepx's of composed functions. *) +(* Option 2 makes it easier to attach canonical structures to sepx, *) +(* but simplification requires explicit and often iterated unfolding *) +(* e.g. rewrite /sepx/=/sepx/=/sepx/=. *) +(* Here, choosing (2); if iterated unfolding of some seprel is *) +(* frequently needed, it should be exported as a lemma. *) + +(* +(* option 1 *) +Definition sepx U V (f : pcm_morph U V) + of phantom (U -> V) f := sep_op f. +Notation sep f := (sepx (Phantom (_ -> _) f)). +*) + +(* option 2 *) +Definition sepx U V (f : pcm_morph U V) + of phantom (U -> V) f : rel U := sep_op f. +Notation sep f := (sepx (Phantom (_ -> _) f)). +HB.instance Definition _ U V (f : pcm_morph U V) := + Seprel.on (sep f). + +(* we can similarly get uniform name for the structure itself *) +(* but we won't use this *) +Definition morphx (U V : pcm) (f : pcm_morph U V) + of phantom (U -> V) f := f. +Notation morph f := (morphx (Phantom (_ -> _) f)). + + +Section Laws. +Variables (U V : pcm) (f : pcm_morph U V). + +Lemma pfunit : f Unit = Unit. +Proof. by case: (pcm_morph_subproof f). Qed. + +Lemma pfjoinV (x y : U) : + valid (x \+ y) -> + sep f x y -> + valid (f x \+ f y) /\ f (x \+ y) = f x \+ f y. +Proof. by case: (pcm_morph_subproof f)=>_; apply. Qed. + +Lemma pfV2 x y : + valid (x \+ y) -> + sep f x y -> + valid (f x \+ f y). +Proof. by move=>H S; case: (pfjoinV H S). Qed. + +Lemma pfjoin x y : + valid (x \+ y) -> + sep f 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) -> + sep f x y -> + f x \+ f y = f (x \+ y). +Proof. by move=>*; rewrite pfjoin. Qed. + +Lemma pfV x : + valid x -> + sep f x Unit -> + valid (f x). +Proof. by rewrite -{1 3}[x]unitR=>W S; rewrite pfjoin // pfV2. Qed. + +Lemma pfV3 x y z : + valid (x \+ y \+ z) -> + sep f x y -> + sep f (x \+ y) z -> + valid (f x \+ f y \+ f z). +Proof. by move=>W D1 D2; rewrite -pfjoin ?(validL W) // pfV2. Qed. + +Lemma pfL (x y z : U) : + valid (x \+ z) -> + valid (y \+ z) -> + sep f x z -> + sep f y z -> + f y = f x -> + f (y \+ z) = f (x \+ z). +Proof. by move=>V1 V2 S1 S2 E; rewrite !pfjoin // E. Qed. + +Lemma pfR (x y z : U) : + valid (z \+ x) -> + valid (z \+ y) -> + sep f z x -> + sep f z y -> + f y = f x -> + f (z \+ y) = f (z \+ x). +Proof. by move=>V1 V2 S1 S2 E; rewrite !pfjoin // E. Qed. + +Lemma pfUnE (x1 y1 x2 y2 : U) : + valid (x1 \+ y1) -> + valid (x2 \+ y2) -> + sep f x1 y1 -> + sep f x2 y2 -> + f x2 = f x1 -> + f y2 = f y1 -> + f (x2 \+ y2) = f (x1 \+ y1). +Proof. +move=>V1 V2 S1 S2 E1 E2. +by rewrite !pfjoin // E1 E2. +Qed. + +Lemma pfunitL x : f Unit \+ x = x. +Proof. by rewrite pfunit unitL. Qed. + +Lemma pfunitR x : x \+ f Unit = x. +Proof. by rewrite pfunit unitR. Qed. + +End Laws. + +Arguments pfunit {U V}. +Arguments pfjoinV {U V}. +Arguments pfV2 {U V}. +Arguments pfjoin {U V}. +Arguments joinpf {U V}. +Arguments pfV {U V}. +Arguments pfV3 {U V}. +Arguments pfL {U V} f {x y}. +Arguments pfR {U V} f {x y}. +Arguments pfunitL {U V}. +Arguments pfunitR {U V}. + +(*********************) +(* Morphism equality *) +(*********************) + +(* morphisms are equal iff equal as functions and have equal seprels *) + +Definition pcm_morph_eq (U V : pcm) (f g : pcm_morph U V) := + f =1 g /\ sep f =s sep g. + +Notation "f =m g" := (pcm_morph_eq f g) (at level 55). + +Lemma pcm_morph_eq_refl {U V} : Reflexive (@pcm_morph_eq U V). +Proof. by []. Qed. + +Lemma pcm_morph_eq_sym {U V} : Symmetric (@pcm_morph_eq U V). +Proof. +by apply/sym_iff_pre_sym=>f1 f2 [H1 H2]; split=>// x y W; rewrite H2. +Qed. + +Lemma pcm_morph_eq_trans {U V} : Transitive (@pcm_morph_eq U V). +Proof. +move=>f2 f1 f3 [H12 S12][H23 S23]; split=>[x|x y W]. +- by rewrite H12. +by rewrite S12 // S23. +Qed. + +(***************************************) +(* Operations on morphisms and seprels *) +(***************************************) + +(* morphism preimage of seprel is seprel *) + +Definition preimx (U V : pcm) (f : pcm_morph U V) + of phantom (U -> V) f : rel V -> rel U := + fun R x y => sep f x y && R (f x) (f y). +Notation preim f := (preimx (Phantom (_ -> _) f)). +Arguments preimx {U V f} _ _ _ _ /. + +Section Preim. +Variables (U V : pcm) (f : pcm_morph U V) (R : seprel V). + +Lemma preim_is_seprel : seprel_axiom (preim f R). +Proof. +split=>[|x y|x y|x y z] /=. +- by rewrite !pfunit !sep00. +- move=>W; rewrite (sepC (sep f))=>//=; case H : (sep f y x)=>//. + by rewrite sepC // pfV2 // sepC. +- move=>H /andP [H1 H2]. + by rewrite !pfunit (sep0E H H1) (sep0E _ H2) // pfV2. +move=>W /andP [G1 F1] /andP [G2 F2]. +rewrite pfjoin ?(validLE3 W) // in F2. +rewrite pfjoin ?(validLE3 W) ?(sepAxx W G1 G2) //=. +move: (pfV2 _ _ _ W G2); rewrite pfjoin ?(validLE3 W) // => W2. +by rewrite !(sepAxx W2 F1 F2). +Qed. + +HB.instance Definition _ := isSeprel.Build U (preim f R) preim_is_seprel. +End Preim. + +(* kernel of morphism is seprel *) + +Definition kerx (U V : pcm) (f : pcm_morph U V) + of phantom (U -> V) f : rel U := + fun x y => sep f x y && sepU (f x) (f y). +Notation ker f := (kerx (Phantom (_ -> _) f)). +HB.instance Definition _ U V (f : pcm_morph U V) := + isSeprel.Build U (ker f) (preim_is_seprel f _). + +(* restriction of morphism by seprel is morphism *) +(* restriction has tpcm range because it returns undef results *) +Section Restriction. +Variables (U : pcm) (V : tpcm). + +Definition res (f : U -> V) (S : rel U) : U -> V := + fun x : U => if S x Unit then f x else undef. + +Variables (f : pcm_morph U V) (S : seprel U). + +Lemma res_is_morphism : pcm_morph_axiom (relI S (sep f)) (res f S). +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. + +HB.instance Definition _ := isPCM_morphism.Build U V (res f S) res_is_morphism. +End Restriction. + +(* equalizer of morphism is seprel *) + +Definition eqlzx (U : pcm) (V : eqpcm) (f1 f2 : pcm_morph U V) + of phantom (U -> V) f1 & phantom (U -> V) f2 : rel U := + fun x y => [&& sep f1 x y, sep f2 x y, f1 x == f2 x & f1 y == f2 y]. +Notation eqlz f1 f2 := (eqlzx (Phantom (_ -> _) f1) (Phantom (_ -> _) f2)). + +Section Equalizer. +Variables (U : pcm) (V : eqpcm) (f1 f2 : pcm_morph U V). + +Lemma eqlz_is_seprel : seprel_axiom (eqlz f1 f2). +Proof. +rewrite /eqlzx; split=>[|x y W|x y W|x y z W]. +- by rewrite !sep00 !pfunit eq_refl. +- by rewrite !andbA andbAC (sepC (sep f1)) // (sepC (sep f2)). +- 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. + +HB.instance Definition _ := isSeprel.Build U (eqlz f1 f2) eqlz_is_seprel. +End Equalizer. + +(* join of two morphism is a morphism with an appropriate seprel *) + +Definition join_relx (U V : pcm) (f1 f2 : pcm_morph U V) + of phantom (U -> V) f1 & phantom (U -> V) f2 : rel U := + fun x y => [&& sep f1 x y, sep f2 x y & + valid ((f1 x \+ f2 x) \+ (f1 y \+ f2 y))]. +Notation join_rel f1 f2 := + (join_relx (Phantom (_ -> _) f1) (Phantom (_ -> _) f2)). + +Definition join_fun (U V : pcm) (f1 f2 : U -> V) : U -> V := + fun x => f1 x \+ f2 x. + +Section JoinMorph. +Variables (U V : pcm) (f1 f2 : pcm_morph U V). + +Lemma join_rel_is_seprel : seprel_axiom (join_rel f1 f2). +Proof. +rewrite /join_relx; split=>[|x y W|x y W|x y z W]. +- by rewrite !sep00 ?unitL !pfunit ?unitL valid_unit. +- by rewrite (sepC (sep f1)) // (sepC (sep f2)) // (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. + +HB.instance Definition _ := + isSeprel.Build U (join_rel f1 f2) join_rel_is_seprel. + +Lemma join_fun_is_morph : + pcm_morph_axiom (join_rel f1 f2) (join_fun f1 f2). +Proof. +rewrite /join_fun; 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. + +HB.instance Definition _ := + isPCM_morphism.Build U V (join_fun f1 f2) join_fun_is_morph. +End JoinMorph. + +(*************************) +(* Categorical morphisms *) +(*************************) + +(* identity *) + +Lemma id_is_morphism (U : pcm) : pcm_morph_axiom relT (@idfun U). +Proof. by []. Qed. +HB.instance Definition _ (U : pcm) := + isPCM_morphism.Build U U idfun (id_is_morphism U). + +(* composition *) + +Section CompMorphism. +Variables U V W : pcm. +Variables (f : pcm_morph U V) (g : pcm_morph V W). +Lemma comp_is_morphism : pcm_morph_axiom (preim f (sep g)) (g \o f). +Proof. +split=>[|x y] /=; first by rewrite !pfunit. +by move=>D /andP [H1 H2]; rewrite !pfjoin // !pfV2. +Qed. +HB.instance Definition _ := + isPCM_morphism.Build U W (g \o f) comp_is_morphism. +End CompMorphism. + + +Section CategoricalLaws. +Variables (U V W X : pcm). +Variables (f : pcm_morph V U) (g : pcm_morph W V) (h : pcm_morph X W). + +Lemma morph0L : f \o idfun =m f. +Proof. by []. Qed. + +Lemma morph0R : idfun \o f =m f. +Proof. by split=>// x y; rewrite {1}/sepx /= andbT. Qed. + +Lemma morphA : f \o (g \o h) =m (f \o g) \o h. +Proof. by split=>//= x y H; rewrite /sepx/= andbA. Qed. +End CategoricalLaws. + +(***********************) +(* Cartesian morphisms *) +(***********************) + +(* always-unit function on valid elements *) +Section UnitFun. +Context {U : pcm} {V : tpcm}. + +Definition unit_fun (x : U) : V := + if valid x then Unit else undef. + +Lemma unitfun_is_pcm_morph : pcm_morph_axiom relT unit_fun. +Proof. +rewrite /unit_fun; split=>[|x y W _]. +- by rewrite valid_unit. +by rewrite W (validL W) (validR W) unitL. +Qed. + +HB.instance Definition _ := + isPCM_morphism.Build U V + unit_fun unitfun_is_pcm_morph. +End UnitFun. + +(* pairwise product of morphisms is a morphism *) +Section FProdMorph. +Variables U1 U2 V1 V2 : pcm. +Variables (f1 : pcm_morph U1 V1) (f2 : pcm_morph U2 V2). + +Lemma fprod_is_morph : + pcm_morph_axiom (rel_prod (sep f1) (sep f2)) (f1 \* f2). +Proof. +rewrite /fprod; split=>[|x y]; first by rewrite !pfunit. +by case/andP=>/= W1 W2 /andP [H1 H2]; rewrite pcmE /= !pfV2 //= !pfjoin. +Qed. + +HB.instance Definition _ := + isPCM_morphism.Build (U1 * U2)%type (V1 * V2)%type (f1 \* f2) fprod_is_morph. +End FProdMorph. + +(* projections are morphisms *) +Section ProjMorph. +Variables U1 U2 : pcm. + +Lemma fst_is_morph : pcm_morph_axiom relT (@fst U1 U2). +Proof. by split=>[|x y] // /andP []. Qed. + +HB.instance Definition _ := + isPCM_morphism.Build (U1 * U2)%type U1 fst fst_is_morph. + +Lemma snd_is_morph : pcm_morph_axiom relT (@snd U1 U2). +Proof. by split=>[|x y] // /andP []. Qed. + +HB.instance Definition _ := + isPCM_morphism.Build (U1 * U2)%type U2 snd snd_is_morph. +End ProjMorph. + +Section Proj3Morph. +Variables U1 U2 U3 : pcm. +Notation tp := (Prod3 U1 U2 U3). + +Lemma proj31_morph_ax : pcm_morph_axiom relT (proj31 : tp -> _). +Proof. by split=>[|x y] // /and3P []. Qed. +Lemma proj32_morph_ax : pcm_morph_axiom relT (proj32 : tp -> _). +Proof. by split=>[|x y] // /and3P []. Qed. +Lemma proj33_morph_ax : pcm_morph_axiom relT (proj33 : tp -> _). +Proof. by split=>[|x y] // /and3P []. Qed. + +HB.instance Definition _ := isPCM_morphism.Build tp U1 _ proj31_morph_ax. +HB.instance Definition _ := isPCM_morphism.Build tp U2 _ proj32_morph_ax. +HB.instance Definition _ := isPCM_morphism.Build tp U3 _ proj33_morph_ax. +End Proj3Morph. + +Section Proj4Morph. +Variables U1 U2 U3 U4 : pcm. +Notation tp := (Prod4 U1 U2 U3 U4). + +Lemma proj41_morph_ax : pcm_morph_axiom relT (proj41 : tp -> _). +Proof. by split=>[|x y] // /and4P []. Qed. +Lemma proj42_morph_ax : pcm_morph_axiom relT (proj42 : tp -> _). +Proof. by split=>[|x y] // /and4P []. Qed. +Lemma proj43_morph_ax : pcm_morph_axiom relT (proj43 : tp -> _). +Proof. by split=>[|x y] // /and4P []. Qed. +Lemma proj44_morph_ax : pcm_morph_axiom relT (proj44 : tp -> _). +Proof. by split=>[|x y] // /and4P []. Qed. + +HB.instance Definition _ := isPCM_morphism.Build tp U1 _ proj41_morph_ax. +HB.instance Definition _ := isPCM_morphism.Build tp U2 _ proj42_morph_ax. +HB.instance Definition _ := isPCM_morphism.Build tp U3 _ proj43_morph_ax. +HB.instance Definition _ := isPCM_morphism.Build tp U4 _ proj44_morph_ax. +End Proj4Morph. + +Section Proj5Morph. +Variables U1 U2 U3 U4 U5 : pcm. +Notation tp := (Prod5 U1 U2 U3 U4 U5). + +Lemma proj51_morph_ax : pcm_morph_axiom relT (proj51 : tp -> _). +Proof. by split=>[|x y] // /and5P []. Qed. +Lemma proj52_morph_ax : pcm_morph_axiom relT (proj52 : tp -> _). +Proof. by split=>[|x y] // /and5P []. Qed. +Lemma proj53_morph_ax : pcm_morph_axiom relT (proj53 : tp -> _). +Proof. by split=>[|x y] // /and5P []. Qed. +Lemma proj54_morph_ax : pcm_morph_axiom relT (proj54 : tp -> _). +Proof. by split=>[|x y] // /and5P []. Qed. +Lemma proj55_morph_ax : pcm_morph_axiom relT (proj55 : tp -> _). +Proof. by split=>[|x y] // /and5P []. Qed. + +HB.instance Definition _ := isPCM_morphism.Build tp U1 _ proj51_morph_ax. +HB.instance Definition _ := isPCM_morphism.Build tp U2 _ proj52_morph_ax. +HB.instance Definition _ := isPCM_morphism.Build tp U3 _ proj53_morph_ax. +HB.instance Definition _ := isPCM_morphism.Build tp U4 _ proj54_morph_ax. +HB.instance Definition _ := isPCM_morphism.Build tp U5 _ proj55_morph_ax. +End Proj5Morph. + +Section Proj6Morph. +Variables U1 U2 U3 U4 U5 U6 : pcm. +Notation tp := (Prod6 U1 U2 U3 U4 U5 U6). + +Lemma proj61_morph_ax : pcm_morph_axiom relT (proj61 : tp -> _). +Proof. by split=>[|x y] // /and6P []. Qed. +Lemma proj62_morph_ax : pcm_morph_axiom relT (proj62 : tp -> _). +Proof. by split=>[|x y] // /and6P []. Qed. +Lemma proj63_morph_ax : pcm_morph_axiom relT (proj63 : tp -> _). +Proof. by split=>[|x y] // /and6P []. Qed. +Lemma proj64_morph_ax : pcm_morph_axiom relT (proj64 : tp -> _). +Proof. by split=>[|x y] // /and6P []. Qed. +Lemma proj65_morph_ax : pcm_morph_axiom relT (proj65 : tp -> _). +Proof. by split=>[|x y] // /and6P []. Qed. +Lemma proj66_morph_ax : pcm_morph_axiom relT (proj66 : tp -> _). +Proof. by split=>[|x y] // /and6P []. Qed. + +HB.instance Definition _ := isPCM_morphism.Build tp U1 _ proj61_morph_ax. +HB.instance Definition _ := isPCM_morphism.Build tp U2 _ proj62_morph_ax. +HB.instance Definition _ := isPCM_morphism.Build tp U3 _ proj63_morph_ax. +HB.instance Definition _ := isPCM_morphism.Build tp U4 _ proj64_morph_ax. +HB.instance Definition _ := isPCM_morphism.Build tp U5 _ proj65_morph_ax. +HB.instance Definition _ := isPCM_morphism.Build tp U6 _ proj66_morph_ax. +End Proj6Morph. + +Section Proj7Morph. +Variables U1 U2 U3 U4 U5 U6 U7 : pcm. +Notation tp := (Prod7 U1 U2 U3 U4 U5 U6 U7). + +Lemma proj71_morph_ax : pcm_morph_axiom relT (proj71 : tp -> _). +Proof. by split=>[|x y] // /and7P []. Qed. +Lemma proj72_morph_ax : pcm_morph_axiom relT (proj72 : tp -> _). +Proof. by split=>[|x y] // /and7P []. Qed. +Lemma proj73_morph_ax : pcm_morph_axiom relT (proj73 : tp -> _). +Proof. by split=>[|x y] // /and7P []. Qed. +Lemma proj74_morph_ax : pcm_morph_axiom relT (proj74 : tp -> _). +Proof. by split=>[|x y] // /and7P []. Qed. +Lemma proj75_morph_ax : pcm_morph_axiom relT (proj75 : tp -> _). +Proof. by split=>[|x y] // /and7P []. Qed. +Lemma proj76_morph_ax : pcm_morph_axiom relT (proj76 : tp -> _). +Proof. by split=>[|x y] // /and7P []. Qed. +Lemma proj77_morph_ax : pcm_morph_axiom relT (proj77 : tp -> _). +Proof. by split=>[|x y] // /and7P []. Qed. + +HB.instance Definition _ := isPCM_morphism.Build tp U1 _ proj71_morph_ax. +HB.instance Definition _ := isPCM_morphism.Build tp U2 _ proj72_morph_ax. +HB.instance Definition _ := isPCM_morphism.Build tp U3 _ proj73_morph_ax. +HB.instance Definition _ := isPCM_morphism.Build tp U4 _ proj74_morph_ax. +HB.instance Definition _ := isPCM_morphism.Build tp U5 _ proj75_morph_ax. +HB.instance Definition _ := isPCM_morphism.Build tp U6 _ proj76_morph_ax. +HB.instance Definition _ := isPCM_morphism.Build tp U7 _ proj77_morph_ax. +End Proj7Morph. + +(* product morphism of morphisms is a morphism *) +Section PMorphMorph. +Variables U V1 V2 : pcm. +Variables (f1 : pcm_morph U V1) (f2 : pcm_morph U V2). + +Lemma pmorph_is_morph : pcm_morph_axiom (relI (sep f1) (sep f2)) (f1 \** f2). +Proof. +rewrite /pmorphism; split=>[|x y] /=; first by rewrite !pcmPE !pfunit. +by move=>V /andP [S1 S2]; rewrite pcmE /= !pfV2 // !pfjoin. +Qed. + +HB.instance Definition _ := + isPCM_morphism.Build U (V1 * V2)%type (f1 \** f2) pmorph_is_morph. +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 *) + +(* morphisms for finite products *) + +Section FinFunMorph. +Variables (T : finType) (Us : T -> pcm). +Implicit Types (f : {dffun forall t, Us t}) (t : T). + +Definition app f := [eta f]. + +Lemma app_is_morph : pcm_morph_axiom relT app. +Proof. +split=>[|x y]; first by apply: fext=>t; rewrite /app ffunE. +move/forallP=>V _; split; last by apply: fext=>t; rewrite /app ffunE. +by apply/forallP=>t; move: (V t); rewrite sel_fin. +Qed. + +(* +HB.instance Definition _ := + isPCM_morphism.Build {ffun forall t, Us t} (forall t, Us t) app app_is_morph. +*) + +(* above lemma isn't so useful in practice because can't rewrite *) +(* by pfjoin when sel is fully applied (which is always). *) +(* Thus, we need separate lemma for the case of applied sel. *) +(* This doesn't require function extensionality *) + +Lemma sel_is_morph t : pcm_morph_axiom relT (@sel T Us t). +Proof. by split=>[|x y]; [|move/forallP/(_ t)]; rewrite sel_fin. Qed. + +HB.instance Definition _ t := + isPCM_morphism.Build {dffun forall t, Us t} (Us t) (sel t) (sel_is_morph t). + +Lemma finfun_is_morph : pcm_morph_axiom relT (@finfun T Us : _ -> {dffun _}). +Proof. +split=>[|f1 f2]; first by apply/ffinP=>t; rewrite sel_fin. +move/forallP=>V _; split. +- by apply/forallP=>t; rewrite !sel_fin; apply: V. +by apply/ffunP=>t; rewrite !ffunE !sel_fin. +Qed. + +HB.instance Definition _ := + isPCM_morphism.Build (forall t, Us t) {dffun forall t, Us t} + (@finfun T Us : _ -> {dffun _}) + finfun_is_morph. + +(* Now we can prove pcm-like lemmas for splice *) + +Lemma valid_spliceUn t (x y : {dffun forall t, Us t}) (v : Us t) : + valid (x \+ y) -> + valid (v \+ sel t y) -> + valid (splice x v \+ y). +Proof. +move=>V V'; apply/forallP=>z; rewrite !sel_fin. +by case: eqP=>/= ?; [subst z; rewrite eqc | rewrite pfV2]. +Qed. + +Lemma spliceUn t (x y : {dffun forall t, Us t}) (v w : Us t) : + splice (x \+ y) (v \+ w) = splice x v \+ splice y w. +Proof. +apply/ffinP=>a; rewrite !sel_fin. +by case: eqP=>//= ?; subst a; rewrite !eqc. +Qed. + +End FinFunMorph. + +(****************************) +(****************************) +(* Classes of PCM morphisms *) +(****************************) +(****************************) + +(*********************************) +(* Normal and binormal morphisms *) +(*********************************) + +(* Normal morphisms: map invalids to invalids *) +(* contrapositively: valid images -> valid originals *) +(* basic healthiness property ensuring *) +(* that no invalid element is "rehabilitated" *) +(* by having the morphism map it to something valid *) + +(* Binormal morphism: map overlapping to overlapping *) +(* contrapositively: separated images -> separated originals *) +(* binary variant of normal -- no overlapping elements *) +(* are rehabilitated by having the morphism map them to *) +(* separated ones *) + +(* binormal -> normal but not vice-versa *) + +(* Both properties are needed, as there are useful morphisms *) +(* that are normal but not binormal. *) +(* Example is injection xval : U -> V in xsep construction below, *) +(* where new PCM U is created out of old one V by moding out V by *) +(* seprel D. This creates new notion of separatenss *) +(* by adding D to the separateness of V *) +(* But then, if two elements are separate in V (old notion) *) +(* it doesn't mean they should be separate in U (new notion) *) +(* as the whole point is that new separateness in U should be more *) +(* demaning. In other words, xval isn't binormal *) + +Definition norm_pcm_morph_axiom (U V : pcm) (f : pcm_morph U V) := + forall x : U, valid (f x) -> valid x /\ sep f x Unit. + +Definition binorm_pcm_morph_axiom (U V : pcm) (f : pcm_morph U V) := + forall x y, valid (f x \+ f y) -> valid (x \+ y) /\ sep f x y. + +(* structure definitions *) +HB.mixin Record isNorm_PCM_morphism (U V : pcm) (f : U -> V) + of @PCM_morphism U V f := { + norm_pcm_morph_subproof : norm_pcm_morph_axiom f}. + +#[short(type=norm_pcm_morph)] +HB.structure Definition Norm_PCM_morphism (U V : pcm) := + {f of isNorm_PCM_morphism U V f & @PCM_morphism U V f}. + +(* helper mixin to define binormal structure *) +HB.mixin Record isBinorm_PCM_morphism U V + f of @PCM_morphism U V f := { + binorm_subproof : binorm_pcm_morph_axiom f}. + +(* require both binorm_axiom and norm_axiom for sake of inheritance *) +(* we prove below that binorm_axiom implies norm_axiom *) +(* so we can actually build binorm_pcm_morph just out of binorm_axiom *) +#[short(type=binorm_pcm_morph)] +HB.structure Definition Binorm_PCM_morphism U V := + {f of isBinorm_PCM_morphism U V f & @Norm_PCM_morphism U V f}. + +(* Norm mixin follows from Binorm mixin *) +HB.builders Context U V f of isBinorm_PCM_morphism U V f. + +Lemma binorm_morph_is_norm_morph : norm_pcm_morph_axiom f. +Proof. +move=>x W; rewrite -{1}[x]unitR; apply: binorm_subproof. +by rewrite pfunit unitR. +Qed. +(* default instance of norm_pcm_morph *) +HB.instance Definition _ := + isNorm_PCM_morphism.Build U V f binorm_morph_is_norm_morph. +HB.end. + +(* Normality Lemmas *) +(* We use names that start with f to indicate that these lemmas *) +(* all have validity preconditions over f x *) + +Lemma fpVI (U V : pcm) (f : norm_pcm_morph U V) (x : U) : + valid (f x) -> valid x /\ sep f x Unit. +Proof. exact: norm_pcm_morph_subproof. Qed. + +Lemma fpVE (U V : pcm) (f : norm_pcm_morph U V) (x : U) : + valid (f x) = valid x && sep f x Unit. +Proof. by apply/idP/idP; [move/fpVI/andP|case/andP; apply: pfV]. Qed. + +Lemma fpV (U V : pcm) (f : norm_pcm_morph U V) (x : U) : + valid (f x) -> valid x. +Proof. by case/fpVI. Qed. + +Lemma fpVS (U V : pcm) (f : norm_pcm_morph U V) (x : U) : + valid (f x) -> sep f x Unit. +Proof. by case/fpVI. Qed. + +(* Binormality Lemmas *) + +Lemma fpV2I (U V : pcm) (f : binorm_pcm_morph U V) (x y : U) : + valid (f x \+ f y) -> valid (x \+ y) /\ sep f x y. +Proof. by apply: binorm_subproof. Qed. + +Lemma fpV2E (U V : pcm) (f : binorm_pcm_morph U V) (x y : U) : + valid (f x \+ f y) = valid (x \+ y) && sep f x y. +Proof. by apply/idP/idP; [move/fpV2I/andP|case/andP; apply: pfV2]. Qed. + +Lemma fpV2 (U V : pcm) (f : binorm_pcm_morph U V) (x y : U) : + valid (f x \+ f y) -> valid (x \+ y). +Proof. by case/fpV2I. Qed. + +Lemma fpV2S (U V : pcm) (f : binorm_pcm_morph U V) (x y : U) : + valid (f x \+ f y) -> sep f x y. +Proof. by case/fpV2I. Qed. + +Arguments fpVI {U V}. +Arguments fpVE {U V}. +Arguments fpV {U V}. +Arguments fpVS {U V}. +Arguments fpV2I {U V}. +Arguments fpV2E {U V}. +Arguments fpV2 {U V}. +Arguments fpV2S {U V}. + +(*******************************************) +(* Categorical morphisms and (bi)normality *) +(*******************************************) + +(* id is binormal *) +Lemma id_is_binorm_morphism {U : pcm} : binorm_pcm_morph_axiom (@idfun U). +Proof. by []. Qed. +HB.instance Definition _ (U : pcm) := + isBinorm_PCM_morphism.Build U U idfun id_is_binorm_morphism. + +(* composition preserves (bi)normality *) +Section CompMorphism. +Variables U V W : pcm. + +Section NormCompIsNormal. +Variables (f : norm_pcm_morph U V) (g : norm_pcm_morph V W). + +Lemma comp_is_norm_morphism : norm_pcm_morph_axiom (g \o f). +Proof. +move=>x /fpVI [] /fpVI [H1 H2] H3 /=. +by rewrite /sepx/= H1 H2 pfunit H3. +Qed. + +HB.instance Definition _ := + isNorm_PCM_morphism.Build U W (g \o f) comp_is_norm_morphism. + +End NormCompIsNormal. + +Section BinormCompIsBinorm. +Variables (f : binorm_pcm_morph U V) (g : binorm_pcm_morph V W). + +Lemma comp_is_binorm_morphism : binorm_pcm_morph_axiom (g \o f). +Proof. +move=>x y /fpV2I [/fpV2I][H1 H2 H3]. +by rewrite /sepx/= H1 H2 H3. +Qed. + +HB.instance Definition _ := + isBinorm_PCM_morphism.Build U W (g \o f) comp_is_binorm_morphism. +End BinormCompIsBinorm. +End CompMorphism. + +(***************************************) +(* Cartesian morphisms and binormality *) +(***************************************) + +(* unit_fun is normal, but not binormal *) +Section UnitFun. +Variables (U : pcm) (V : tpcm). + +Lemma unitfun_is_normal : norm_pcm_morph_axiom (@unit_fun U V). +Proof. +move=>x; case W : (valid x)=>//. +by move: (valid_undef V); rewrite /valid/=/unit_fun/= W =>->. +Qed. + +HB.instance Definition _ := + isNorm_PCM_morphism.Build U V + (@unit_fun U V) unitfun_is_normal. +End UnitFun. + +(* projections aren't (bi)normal, but product morphisms are *) + +Section FProdMorph. +Variables U1 U2 V1 V2 : pcm. + +Section Normal. +Variables (f1 : norm_pcm_morph U1 V1) (f2 : norm_pcm_morph U2 V2). + +Lemma fprod_is_norm_pcm_morph : norm_pcm_morph_axiom (f1 \* f2). +Proof. +move=>x /=; rewrite !pcmPV /sepx /=. +by case/andP=>/fpVI [->->] /fpVI [->->]. +Qed. + +HB.instance Definition _ := + isNorm_PCM_morphism.Build (U1 * U2)%type (V1 * V2)%type (f1 \* f2) + fprod_is_norm_pcm_morph. +End Normal. + +Section Binormal. +Variables (f1 : binorm_pcm_morph U1 V1) (f2 : binorm_pcm_morph U2 V2). + +Lemma fprod_is_binorm_pcm_morph : binorm_pcm_morph_axiom (f1 \* f2). +Proof. +move=>x y /=; rewrite !pcmPV /sepx/=. +by case/andP=>/fpV2I [->->] /fpV2I. +Qed. + +HB.instance Definition _ := +isBinorm_PCM_morphism.Build (U1 * U2)%type (V1 * V2)%type (f1 \* f2) + fprod_is_binorm_pcm_morph. +End Binormal. +End FProdMorph. + +Section PMorphMorph. +Variables U V1 V2 : pcm. + +Section Normal. +Variables (f1 : norm_pcm_morph U V1) (f2 : norm_pcm_morph U V2). + +Lemma pmorph_is_norm_pcm_morph : norm_pcm_morph_axiom (f1 \** f2). +Proof. +move=>x; rewrite /sepx /=. +by case/andP=>/fpVI [->->] /fpVI []. +Qed. + +HB.instance Definition _ := + isNorm_PCM_morphism.Build U (V1 * V2)%type (f1 \** f2) + pmorph_is_norm_pcm_morph. +End Normal. + +Section Binormal. +Variables (f1 : binorm_pcm_morph U V1) (f2 : binorm_pcm_morph U V2). + +Lemma pmorph_is_binorm_pcm_morph : binorm_pcm_morph_axiom (f1 \** f2). +Proof. +move=>x y /= /andP [] /fpV2I [-> H1] /fpV2I [_ H2]. +by rewrite /sepx/= H1 H2. +Qed. + +HB.instance Definition _ := + isBinorm_PCM_morphism.Build U (V1 * V2)%type (f1 \** f2) + pmorph_is_binorm_pcm_morph. +End Binormal. +End PMorphMorph. + +Section FinFunMorph. +Variables (T : finType) (Us : T -> pcm). + +Lemma finfun_is_binorm_pcm_morph : + binorm_pcm_morph_axiom (@finfun T Us : _ -> {dffun _}). +Proof. +move=>f1 f2 /forallP /= V; split=>//; apply/forallP=>t. +by move: (V t); rewrite !sel_fin. +Qed. + +HB.instance Definition _ := + isBinorm_PCM_morphism.Build (forall t, Us t) {dffun forall t, Us t} + (@finfun T Us : _ -> {dffun _}) finfun_is_binorm_pcm_morph. +End FinFunMorph. + + +(*****************) +(* Full morphism *) +(*****************) + +(* morphism with seprel relT (largest possible) *) +(* an alternative name would be "total" morphism *) +(* as this morphism is defined for all valid elements *) +(* but we keep "total" for functions that are definied on *) +(* *all* elements, valid or not *) + +Definition full_pcm_morph_axiom (U V : pcm) (f : pcm_morph U V) := + sep f =2 relT. + +HB.mixin Record isFull_PCM_morphism (U V : pcm) (f : U -> V) + of @PCM_morphism U V f := { + full_pcm_morphism_subproof : full_pcm_morph_axiom f}. + +#[short(type=full_pcm_morph)] +HB.structure Definition Full_PCM_morphism (U V : pcm) := + {f of isFull_PCM_morphism U V f & @PCM_morphism U V f}. + +#[short(type=full_norm_pcm_morph)] +HB.structure Definition Full_Norm_PCM_morphism (U V : pcm) := + {f of @Norm_PCM_morphism U V f & @Full_PCM_morphism U V f}. + +#[short(type=full_binorm_pcm_morph)] +HB.structure Definition Full_Binorm_PCM_morphism (U V : pcm) := + {f of @Binorm_PCM_morphism U V f & @Full_PCM_morphism U V f}. + +(* fullness lemmas *) + +Lemma pfSE (U V : pcm) (f : full_pcm_morph U V) : sep f =2 relT. +Proof. by apply: full_pcm_morphism_subproof. Qed. + +Lemma pfT (U V : pcm) (f : full_pcm_morph U V) x y : sep f x y. +Proof. by rewrite pfSE. Qed. + +#[export] Hint Resolve pfT : core. + +Lemma pfVI (U V : pcm) (f : full_pcm_morph U V) (x : U) : + valid x -> + valid (f x). +Proof. by move=>W; apply: pfV W _. Qed. + +Lemma pfV2I (U V : pcm) (f : full_pcm_morph U V) (x y : U) : + valid (x \+ y) -> + valid (f x \+ f y). +Proof. by move=>W; apply: pfV2 W _. Qed. + +Lemma pfV2IC (U V : pcm) (f : full_pcm_morph U V) (x y : U) : + valid (x \+ y) -> + valid (f y \+ f x). +Proof. by rewrite joinC=>/pfV2I. Qed. + +Lemma pfV3I (U V : pcm) (f : full_pcm_morph U V) (x y z : U) : + valid (x \+ y \+ z) -> + valid (f x \+ f y \+ f z). +Proof. by move=>W; rewrite -pfjoin ?(validL W) // pfV2. Qed. + +Lemma pfVE (U V : pcm) (f : full_norm_pcm_morph U V) (x : U) : + valid (f x) = valid x. +Proof. by rewrite fpVE pfT andbT. Qed. + +Lemma pfV2E (U V : pcm) (f : full_binorm_pcm_morph U V) (x y : U) : + valid (f x \+ f y) = valid (x \+ y). +Proof. by rewrite fpV2E pfT andbT. Qed. + +Lemma pfjoinT (U V : pcm) (f : full_pcm_morph U V) (x y : U) : + valid (x \+ y) -> + f (x \+ y) = f x \+ f y. +Proof. by move=>D; rewrite pfjoin. Qed. + +Lemma pfLT (U V : pcm) (f : full_pcm_morph U V) (x y z : U) : + valid (x \+ z) -> + valid (y \+ z) -> + f y = f x -> + f (y \+ z) = f (x \+ z). +Proof. by move=>V1 V2 E; rewrite !pfjoinT // E. Qed. + +Lemma pfRT (U V : pcm) (f : full_pcm_morph U V) (x y z : U) : + valid (z \+ x) -> + valid (z \+ y) -> + f y = f x -> + f (z \+ y) = f (z \+ x). +Proof. by move=>V1 V2 E; rewrite !pfjoinT // E. Qed. + +Lemma pfUnTE (U V : pcm) (f : full_pcm_morph U V) (x1 y1 x2 y2 : U) : + valid (x1 \+ y1) -> + valid (x2 \+ y2) -> + f x2 = f x1 -> + f y2 = f y1 -> + f (x2 \+ y2) = f (x1 \+ y1). +Proof. by move=>V1 V2 E1 E2; rewrite !pfjoinT // E1 E2. Qed. + +Arguments pfVI {U V f x}. +Arguments pfV2I {U V f x y}. +Arguments pfVE {U V f x}. +Arguments pfV2E {U V f x y}. +Arguments pfjoinT {U V f x}. +Arguments pfLT {U V} f {x y}. +Arguments pfRT {U V} f {x y}. + +(* Categorical morphisms and fullness *) + +(* id is full *) +Lemma id_is_full_morphism {U : pcm} : full_pcm_morph_axiom (@idfun U). +Proof. by []. Qed. +HB.instance Definition _ (U : pcm) := + isFull_PCM_morphism.Build U U idfun id_is_full_morphism. + +(* composition preserves fullness *) +Section CompMorphism. +Variables U V W : pcm. + +Section FullCompIsFull. +Variables (f : full_pcm_morph U V) (g : full_pcm_morph V W). + +Lemma comp_is_full_morphism : full_pcm_morph_axiom (g \o f). +Proof. by move=>x y /=; rewrite /sepx/= !pfT. Qed. + +HB.instance Definition _ := + isFull_PCM_morphism.Build U W (g \o f) comp_is_full_morphism. +End FullCompIsFull. + +(* instances for combinations must be declared explicitly *) +(* DEVCOMMENT *) +(* this seems unintended, should check with mathcomp people *) +(* /DEVCOMMENT *) + +HB.instance Definition _ (f : full_norm_pcm_morph U V) + (g : full_norm_pcm_morph V W) := + Full_Norm_PCM_morphism.copy (g \o f) + (Full_Norm_PCM_morphism.pack_ + (Norm_PCM_morphism.class (g \o f)) + (Full_PCM_morphism.class (g \o f))). + +HB.instance Definition _ (f : full_binorm_pcm_morph U V) + (g : full_binorm_pcm_morph V W) := + Full_Binorm_PCM_morphism.copy (g \o f) + (Full_Binorm_PCM_morphism.pack_ + (Binorm_PCM_morphism.class (g \o f)) + (Norm_PCM_morphism.class (g \o f)) + (Full_PCM_morphism.class (g \o f))). + +End CompMorphism. + +(* Cartesian morphisms and fullness *) + +Section UnitFun. +Variables (U : pcm) (V : tpcm). + +Lemma unitfun_is_full_pcm_morph : full_pcm_morph_axiom (@unit_fun U V). +Proof. by []. Qed. + +HB.instance Definition _ := + isFull_PCM_morphism.Build U V + (@unit_fun U V) unitfun_is_full_pcm_morph. +End UnitFun. + + +Section Cartesians. +Notation pf := (fun _ _ => erefl _). + +Section FProdMorph. +Variables U1 U2 V1 V2 : pcm. +Variables (f1 : full_pcm_morph U1 V1) (f2 : full_pcm_morph U2 V2). + +Lemma fprod_is_full_pcm_morph : full_pcm_morph_axiom (f1 \* f2). +Proof. by move=>x y /=; rewrite /sepx/=!pfT. Qed. + +HB.instance Definition _ := + isFull_PCM_morphism.Build (U1 * U2)%type (V1 * V2)%type (f1 \* f2) + fprod_is_full_pcm_morph. +End FProdMorph. + +Section ProjMorph. +Variables U1 U2 : pcm. +HB.instance Definition _ := isFull_PCM_morphism.Build (U1 * U2)%type U1 fst pf. +HB.instance Definition _ := isFull_PCM_morphism.Build (U1 * U2)%type U2 snd pf. +End ProjMorph. + +Section Proj3Morph. +Variables U1 U2 U3 : pcm. +Notation tp := (Prod3 U1 U2 U3). +HB.instance Definition _ := isFull_PCM_morphism.Build tp U1 proj31 pf. +HB.instance Definition _ := isFull_PCM_morphism.Build tp U2 proj32 pf. +HB.instance Definition _ := isFull_PCM_morphism.Build tp U3 proj33 pf. +End Proj3Morph. + +Section Proj4Morph. +Variables U1 U2 U3 U4 : pcm. +Notation tp := (Prod4 U1 U2 U3 U4). +HB.instance Definition _ := isFull_PCM_morphism.Build tp U1 proj41 pf. +HB.instance Definition _ := isFull_PCM_morphism.Build tp U2 proj42 pf. +HB.instance Definition _ := isFull_PCM_morphism.Build tp U3 proj43 pf. +HB.instance Definition _ := isFull_PCM_morphism.Build tp U4 proj44 pf. +End Proj4Morph. + +Section Proj5Morph. +Variables U1 U2 U3 U4 U5 : pcm. +Notation tp := (Prod5 U1 U2 U3 U4 U5). +HB.instance Definition _ := isFull_PCM_morphism.Build tp U1 proj51 pf. +HB.instance Definition _ := isFull_PCM_morphism.Build tp U2 proj52 pf. +HB.instance Definition _ := isFull_PCM_morphism.Build tp U3 proj53 pf. +HB.instance Definition _ := isFull_PCM_morphism.Build tp U4 proj54 pf. +HB.instance Definition _ := isFull_PCM_morphism.Build tp U5 proj55 pf. +End Proj5Morph. + +Section Proj6Morph. +Variables U1 U2 U3 U4 U5 U6 : pcm. +Notation tp := (Prod6 U1 U2 U3 U4 U5 U6). +HB.instance Definition _ := isFull_PCM_morphism.Build tp U1 proj61 pf. +HB.instance Definition _ := isFull_PCM_morphism.Build tp U2 proj62 pf. +HB.instance Definition _ := isFull_PCM_morphism.Build tp U3 proj63 pf. +HB.instance Definition _ := isFull_PCM_morphism.Build tp U4 proj64 pf. +HB.instance Definition _ := isFull_PCM_morphism.Build tp U5 proj65 pf. +HB.instance Definition _ := isFull_PCM_morphism.Build tp U6 proj66 pf. +End Proj6Morph. + +Section Proj7Morph. +Variables U1 U2 U3 U4 U5 U6 U7 : pcm. +Notation tp := (Prod7 U1 U2 U3 U4 U5 U6 U7). +HB.instance Definition _ := isFull_PCM_morphism.Build tp U1 proj71 pf. +HB.instance Definition _ := isFull_PCM_morphism.Build tp U2 proj72 pf. +HB.instance Definition _ := isFull_PCM_morphism.Build tp U3 proj73 pf. +HB.instance Definition _ := isFull_PCM_morphism.Build tp U4 proj74 pf. +HB.instance Definition _ := isFull_PCM_morphism.Build tp U5 proj75 pf. +HB.instance Definition _ := isFull_PCM_morphism.Build tp U6 proj76 pf. +HB.instance Definition _ := isFull_PCM_morphism.Build tp U7 proj77 pf. +End Proj7Morph. + +Section PMorphMorph. +Variables U V1 V2 : pcm. +Variables (f1 : full_pcm_morph U V1) (f2 : full_pcm_morph U V2). + +Lemma pmorph_is_full_pcm_morph : full_pcm_morph_axiom (f1 \** f2). +Proof. by move=>x y /=; rewrite /sepx/= !pfT. Qed. + +HB.instance Definition _ := + isFull_PCM_morphism.Build U (V1 * V2)%type (f1 \** f2) + pmorph_is_full_pcm_morph. +End PMorphMorph. + +Section FinFunMorph. +HB.instance Definition _ (T : finType) (Us : T -> pcm) (t : T) := + isFull_PCM_morphism.Build {dffun forall t, Us t} (Us t) (sel t) pf. + +HB.instance Definition _ (T : finType) (Us : T -> pcm) := + isFull_PCM_morphism.Build (forall t, Us t) {dffun forall t, Us t} + (@finfun T Us : _ -> {dffun _}) pf. +End FinFunMorph. + +End Cartesians. + +(* fullness and other constructions *) + +Section RelApp. +Variables (U V : pcm) (S : seprel V) (f : full_pcm_morph U V). + +Lemma relapp_is_seprel : seprel_axiom (rel_app S f). +Proof. +split=>[|x y|x y|x y z] /=. +- by rewrite pfunit sep00. +- by move=>W; rewrite sepC // pfV2I. +- by move=>W; rewrite pfunit; move/sep0E=>-> //; rewrite pfV2I. +move=>W; rewrite !pfjoinT ?(validLE3 W) // => X Y. +by rewrite !(sepAxx _ X Y) // -!pfjoinT ?(validLE3 W) ?pfVI. +Qed. + +HB.instance Definition _ := + isSeprel.Build U (rel_app S f) relapp_is_seprel. +End RelApp. + +Section Pmorph. +Variables (U V1 V2 : pcm). +Variables (f1 : full_pcm_morph U V1) (f2 : full_pcm_morph U V2). + +Lemma pmorphism_is_full : full_pcm_morph_axiom (f1 \** f2). +Proof. by move=>x y /=; rewrite !pfT. Qed. + +HB.instance Definition _ := + isFull_PCM_morphism.Build U (V1 * V2)%type + (f1 \** f2) pmorphism_is_full. +End Pmorph. + +(******************) +(******************) +(* TPCM morphisms *) +(******************) +(******************) + +(* tpcm morphisms further preserve undef *) + +Definition tpcm_morph_axiom (U V : tpcm) (f : U -> V) := + f undef = undef. + +HB.mixin Record isTPCM_morphism (U V : tpcm) (f : U -> V) := { + tpcm_morphism_subproof : tpcm_morph_axiom f}. + +#[short(type=tpcm_morph)] +HB.structure Definition TPCM_morphism (U V : tpcm) := + {f of isTPCM_morphism U V f & @PCM_morphism U V f}. + +(* introduce names for the combinations of sub-properties *) + +#[short(type=norm_tpcm_morph)] +HB.structure Definition Norm_TPCM_morphism (U V : tpcm) := + {f of @Norm_PCM_morphism U V f & @TPCM_morphism U V f}. + +#[short(type=binorm_tpcm_morph)] +HB.structure Definition Binorm_TPCM_morphism (U V : tpcm) := + {f of @Binorm_PCM_morphism U V f & @TPCM_morphism U V f}. + +#[short(type=full_tpcm_morph)] +HB.structure Definition Full_TPCM_morphism (U V : tpcm) := + {f of @Full_PCM_morphism U V f & @TPCM_morphism U V f}. + +#[short(type=full_norm_tpcm_morph)] +HB.structure Definition Full_Norm_TPCM_morphism (U V : tpcm) := + {f of @Full_PCM_morphism U V f & @Norm_TPCM_morphism U V f}. + +#[short(type=full_binorm_tpcm_morph)] +HB.structure Definition Full_Binorm_TPCM_morphism (U V : tpcm) := + {f of @Full_PCM_morphism U V f & @Binorm_TPCM_morphism U V f}. + +(* TPCM lemmas *) + +Lemma pfundef {U V : tpcm} (f : tpcm_morph U V) : f undef = undef. +Proof. by apply: tpcm_morphism_subproof. Qed. + +(* full morphism on normal tpcm is normal *) + +Lemma fullmorph_is_norm (U : normal_tpcm) V (f : full_tpcm_morph U V) : + norm_pcm_morph_axiom f. +Proof. +move=>x; rewrite pfT; case: (normalP x)=>[->|] //=. +by rewrite pfundef valid_undef. +Qed. +Arguments fullmorph_is_norm {U V f}. + +(* Following lemma is a crutch whose use is discouraged. *) +(* Instead, declare f as normal, and use pfVE. *) +Lemma pfnVE (U : normal_tpcm) V (f : full_tpcm_morph U V) x : + valid (f x) = valid x. +Proof. by apply/idP/idP; [case/fullmorph_is_norm|apply/pfVI]. Qed. + +(* Categoricals are tpcm morphism *) + +HB.instance Definition _ (U : tpcm) := + isTPCM_morphism.Build U U (@idfun U) (erefl _). + +Section Comp. +Variables U V W : tpcm. + +Section CompTPCM. +Variables (f : tpcm_morph U V) (g : tpcm_morph V W). + +Lemma comp_is_tpcm_morphism : tpcm_morph_axiom (g \o f). +Proof. by rewrite /tpcm_morph_axiom /= !pfundef. Qed. + +HB.instance Definition _ := + isTPCM_morphism.Build U W (g \o f) comp_is_tpcm_morphism. +End CompTPCM. + +(* combinations declared explicitly *) +(* DEVCOMMENT: boilerplate -- could it be automated? *) + +HB.instance Definition _ + (f : norm_tpcm_morph U V) (g : norm_tpcm_morph V W) := + Norm_TPCM_morphism.copy (g \o f) + (Norm_TPCM_morphism.pack_ + (TPCM_morphism.class (g \o f)) + (Norm_PCM_morphism.class (g \o f))). + +HB.instance Definition _ + (f : binorm_tpcm_morph U V) (g : binorm_tpcm_morph V W) := + Binorm_TPCM_morphism.copy (g \o f) + (Binorm_TPCM_morphism.pack_ + (TPCM_morphism.class (g \o f)) + (Binorm_PCM_morphism.class (g \o f)) + (Norm_PCM_morphism.class (g \o f))). + +HB.instance Definition _ + (f : full_tpcm_morph U V) (g : full_tpcm_morph V W) := + Full_TPCM_morphism.copy (g \o f) + (Full_TPCM_morphism.pack_ + (TPCM_morphism.class (g \o f)) + (Full_PCM_morphism.class (g \o f))). + +HB.instance Definition _ + (f : full_norm_tpcm_morph U V) (g : full_norm_tpcm_morph V W) := + Full_Norm_TPCM_morphism.copy (g \o f) + (Full_Norm_TPCM_morphism.pack_ + (TPCM_morphism.class (g \o f)) + (Full_PCM_morphism.class (g \o f)) + (Norm_PCM_morphism.class (g \o f))). + +HB.instance Definition _ + (f : full_binorm_tpcm_morph U V) (g : full_binorm_tpcm_morph V W) := + Full_Binorm_TPCM_morphism.copy (g \o f) + (Full_Binorm_TPCM_morphism.pack_ + (TPCM_morphism.class (g \o f)) + (Full_PCM_morphism.class (g \o f)) + (Binorm_PCM_morphism.class (g \o f)) + (Norm_PCM_morphism.class (g \o f))). + +End Comp. + +(* Cartesians *) + +Section UnitFun. +Variables (U V : tpcm). + +Lemma unitfun_is_tpcm_morph : tpcm_morph_axiom (@unit_fun U V). +Proof. by rewrite /tpcm_morph_axiom/unit_fun valid_undef. Qed. + +HB.instance Definition _ := + isTPCM_morphism.Build U V + (@unit_fun U V) unitfun_is_tpcm_morph. +End UnitFun. + + +Section FProdMorph. +Variables U1 U2 V1 V2 : tpcm. +Variables (f1 : tpcm_morph U1 V1) (f2 : tpcm_morph U2 V2). + +Lemma fprod_is_tpcm_morph : tpcm_morph_axiom (f1 \* f2). +Proof. by rewrite /tpcm_morph_axiom /fprod !pfundef. Qed. + +HB.instance Definition _ := + isTPCM_morphism.Build (U1 * U2)%type (V1 * V2)%type (f1 \* f2) + fprod_is_tpcm_morph. +End FProdMorph. + +(* projections are tpcm morphisms *) +Section ProjMorph. +Variables U1 U2 : tpcm. + +Lemma fst_is_tpcm_morph : tpcm_morph_axiom (@fst U1 U2). +Proof. by rewrite /tpcm_morph_axiom /undef. Qed. + +HB.instance Definition _ := + isTPCM_morphism.Build (U1 * U2)%type U1 fst fst_is_tpcm_morph. + +Lemma snd_is_tpcm_morph : tpcm_morph_axiom (@snd U1 U2). +Proof. by rewrite /tpcm_morph_axiom /undef. Qed. + +HB.instance Definition _ := + isTPCM_morphism.Build (U1 * U2)%type U2 snd snd_is_tpcm_morph. +End ProjMorph. + +(* projections for iterated products are morphisms *) + +Section Proj3Morph. +Variables U1 U2 U3 : tpcm. +Notation tp := (Prod3 U1 U2 U3). + +Lemma proj31_is_tpcm_morph : tpcm_morph_axiom (proj31 : tp -> _). +Proof. by rewrite /tpcm_morph_axiom /undef. Qed. +Lemma proj32_is_tpcm_morph : tpcm_morph_axiom (proj32 : tp -> _). +Proof. by rewrite /tpcm_morph_axiom /undef. Qed. +Lemma proj33_is_tpcm_morph : tpcm_morph_axiom (proj33 : tp -> _). +Proof. by rewrite /tpcm_morph_axiom /undef. Qed. + +HB.instance Definition _ := isTPCM_morphism.Build tp U1 _ proj31_is_tpcm_morph. +HB.instance Definition _ := isTPCM_morphism.Build tp U2 _ proj32_is_tpcm_morph. +HB.instance Definition _ := isTPCM_morphism.Build tp U3 _ proj33_is_tpcm_morph. +End Proj3Morph. + +Section Proj4Morph. +Variables U1 U2 U3 U4 : tpcm. +Notation tp := (Prod4 U1 U2 U3 U4). + +Lemma proj41_is_tpcm_morph : tpcm_morph_axiom (proj41 : tp -> _). +Proof. by rewrite /tpcm_morph_axiom /undef. Qed. +Lemma proj42_is_tpcm_morph : tpcm_morph_axiom (proj42 : tp -> _). +Proof. by rewrite /tpcm_morph_axiom /undef. Qed. +Lemma proj43_is_tpcm_morph : tpcm_morph_axiom (proj43 : tp -> _). +Proof. by rewrite /tpcm_morph_axiom /undef. Qed. +Lemma proj44_is_tpcm_morph : tpcm_morph_axiom (proj44 : tp -> _). +Proof. by rewrite /tpcm_morph_axiom /undef. Qed. + +HB.instance Definition _ := isTPCM_morphism.Build tp U1 _ proj41_is_tpcm_morph. +HB.instance Definition _ := isTPCM_morphism.Build tp U2 _ proj42_is_tpcm_morph. +HB.instance Definition _ := isTPCM_morphism.Build tp U3 _ proj43_is_tpcm_morph. +HB.instance Definition _ := isTPCM_morphism.Build tp U4 _ proj44_is_tpcm_morph. +End Proj4Morph. + +Section Proj5Morph. +Variables U1 U2 U3 U4 U5 : tpcm. +Notation tp := (Prod5 U1 U2 U3 U4 U5). + +Lemma proj51_is_tpcm_morph : tpcm_morph_axiom (proj51 : tp -> _). +Proof. by rewrite /tpcm_morph_axiom /undef. Qed. +Lemma proj52_is_tpcm_morph : tpcm_morph_axiom (proj52 : tp -> _). +Proof. by rewrite /tpcm_morph_axiom /undef. Qed. +Lemma proj53_is_tpcm_morph : tpcm_morph_axiom (proj53 : tp -> _). +Proof. by rewrite /tpcm_morph_axiom /undef. Qed. +Lemma proj54_is_tpcm_morph : tpcm_morph_axiom (proj54 : tp -> _). +Proof. by rewrite /tpcm_morph_axiom /undef. Qed. +Lemma proj55_is_tpcm_morph : tpcm_morph_axiom (proj55 : tp -> _). +Proof. by rewrite /tpcm_morph_axiom /undef. Qed. + +HB.instance Definition _ := isTPCM_morphism.Build tp U1 _ proj51_is_tpcm_morph. +HB.instance Definition _ := isTPCM_morphism.Build tp U2 _ proj52_is_tpcm_morph. +HB.instance Definition _ := isTPCM_morphism.Build tp U3 _ proj53_is_tpcm_morph. +HB.instance Definition _ := isTPCM_morphism.Build tp U4 _ proj54_is_tpcm_morph. +HB.instance Definition _ := isTPCM_morphism.Build tp U5 _ proj55_is_tpcm_morph. +End Proj5Morph. + +Section Proj6Morph. +Variables U1 U2 U3 U4 U5 U6 : tpcm. +Notation tp := (Prod6 U1 U2 U3 U4 U5 U6). + +Lemma proj61_is_tpcm_morph : tpcm_morph_axiom (proj61 : tp -> _). +Proof. by rewrite /tpcm_morph_axiom /undef. Qed. +Lemma proj62_is_tpcm_morph : tpcm_morph_axiom (proj62 : tp -> _). +Proof. by rewrite /tpcm_morph_axiom /undef. Qed. +Lemma proj63_is_tpcm_morph : tpcm_morph_axiom (proj63 : tp -> _). +Proof. by rewrite /tpcm_morph_axiom /undef. Qed. +Lemma proj64_is_tpcm_morph : tpcm_morph_axiom (proj64 : tp -> _). +Proof. by rewrite /tpcm_morph_axiom /undef. Qed. +Lemma proj65_is_tpcm_morph : tpcm_morph_axiom (proj65 : tp -> _). +Proof. by rewrite /tpcm_morph_axiom /undef. Qed. +Lemma proj66_is_tpcm_morph : tpcm_morph_axiom (proj66 : tp -> _). +Proof. by rewrite /tpcm_morph_axiom /undef. Qed. + +HB.instance Definition _ := isTPCM_morphism.Build tp U1 _ proj61_is_tpcm_morph. +HB.instance Definition _ := isTPCM_morphism.Build tp U2 _ proj62_is_tpcm_morph. +HB.instance Definition _ := isTPCM_morphism.Build tp U3 _ proj63_is_tpcm_morph. +HB.instance Definition _ := isTPCM_morphism.Build tp U4 _ proj64_is_tpcm_morph. +HB.instance Definition _ := isTPCM_morphism.Build tp U5 _ proj65_is_tpcm_morph. +HB.instance Definition _ := isTPCM_morphism.Build tp U6 _ proj66_is_tpcm_morph. +End Proj6Morph. + +Section Proj7Morph. +Variables U1 U2 U3 U4 U5 U6 U7 : tpcm. +Notation tp := (Prod7 U1 U2 U3 U4 U5 U6 U7). + +Lemma proj71_is_tpcm_morph : tpcm_morph_axiom (proj71 : tp -> _). +Proof. by rewrite /tpcm_morph_axiom /undef. Qed. +Lemma proj72_is_tpcm_morph : tpcm_morph_axiom (proj72 : tp -> _). +Proof. by rewrite /tpcm_morph_axiom /undef. Qed. +Lemma proj73_is_tpcm_morph : tpcm_morph_axiom (proj73 : tp -> _). +Proof. by rewrite /tpcm_morph_axiom /undef. Qed. +Lemma proj74_is_tpcm_morph : tpcm_morph_axiom (proj74 : tp -> _). +Proof. by rewrite /tpcm_morph_axiom /undef. Qed. +Lemma proj75_is_tpcm_morph : tpcm_morph_axiom (proj75 : tp -> _). +Proof. by rewrite /tpcm_morph_axiom /undef. Qed. +Lemma proj76_is_tpcm_morph : tpcm_morph_axiom (proj76 : tp -> _). +Proof. by rewrite /tpcm_morph_axiom /undef. Qed. +Lemma proj77_is_tpcm_morph : tpcm_morph_axiom (proj77 : tp -> _). +Proof. by rewrite /tpcm_morph_axiom /undef. Qed. + +HB.instance Definition _ := isTPCM_morphism.Build tp U1 _ proj71_is_tpcm_morph. +HB.instance Definition _ := isTPCM_morphism.Build tp U2 _ proj72_is_tpcm_morph. +HB.instance Definition _ := isTPCM_morphism.Build tp U3 _ proj73_is_tpcm_morph. +HB.instance Definition _ := isTPCM_morphism.Build tp U4 _ proj74_is_tpcm_morph. +HB.instance Definition _ := isTPCM_morphism.Build tp U5 _ proj75_is_tpcm_morph. +HB.instance Definition _ := isTPCM_morphism.Build tp U6 _ proj76_is_tpcm_morph. +HB.instance Definition _ := isTPCM_morphism.Build tp U7 _ proj77_is_tpcm_morph. +End Proj7Morph. + +(* product morphism of tpcm morphisms is a tpcm morphism *) +Section PMorphMorph. +Variables U V1 V2 : tpcm. +Variables (f1 : tpcm_morph U V1) (f2 : tpcm_morph U V2). + +Lemma pmorph_is_tpcm_morph : tpcm_morph_axiom (f1 \** f2). +Proof. by rewrite /tpcm_morph_axiom/pmorphism !pfundef. Qed. + +HB.instance Definition _ := + isTPCM_morphism.Build U (V1 * V2)%type (f1 \** f2) pmorph_is_tpcm_morph. + +End PMorphMorph. + +(* combination of full and tpcm declared by hand *) +HB.instance Definition _ (U V1 V2 : tpcm) + (f1 : full_tpcm_morph U V1) (f2 : full_tpcm_morph U V2) := + Full_TPCM_morphism.copy (f1 \** f2) + (Full_TPCM_morphism.pack_ + (TPCM_morphism.class (f1 \** f2)) + (Full_PCM_morphism.class (f1 \** f2))). + +(* finite products *) +Section FinFunMorph. +Variables (T : ifinType) (Us : T -> tpcm). + +Lemma sel_is_tpcm_morph t : tpcm_morph_axiom (sel (Us:=Us) t). +Proof. by rewrite /tpcm_morph_axiom sel_fin. Qed. + +HB.instance Definition _ (t : T) := + isTPCM_morphism.Build {dffun forall t, Us t} (Us t) (sel t) + (sel_is_tpcm_morph t). + +Lemma finfun_is_tpcm_morph : + tpcm_morph_axiom (V:={dffun forall t, Us t}) finfun. +Proof. by []. Qed. + +HB.instance Definition _ := + isTPCM_morphism.Build (forall t, Us t) {dffun forall t, Us t} + (@finfun T Us : _ -> {dffun _}) finfun_is_tpcm_morph. +End FinFunMorph. + +(* restriction *) +Section RestrictionTPCM. +Variables (U V : tpcm) (f : tpcm_morph U V) (S : seprel U). + +Lemma res_is_tpcm_morph : tpcm_morph_axiom (res f S). +Proof. by rewrite /tpcm_morph_axiom /res pfundef if_same. Qed. + +HB.instance Definition _ := + isTPCM_morphism.Build U V (res f S) res_is_tpcm_morph. +End RestrictionTPCM. + +(* join *) +Section JoinFunTPCM. +Variables (U V : tpcm) (f1 f2 : tpcm_morph U V). + +Lemma joinfun_is_tpcm_morph : tpcm_morph_axiom (join_fun f1 f2). +Proof. by rewrite /tpcm_morph_axiom/join_fun !pfundef join_undef. Qed. + +HB.instance Definition _ := + isTPCM_morphism.Build U V (join_fun f1 f2) joinfun_is_tpcm_morph. +End JoinFunTPCM. + +(********************) +(* Option morphisms *) +(********************) + +Definition odfltu {A : pcm} (x : option A) := odflt Unit x. + +(* Some is full binormal morphism *) +Lemma some_is_morph {A : pcm} : pcm_morph_axiom relT (@Some A). +Proof. by []. Qed. +HB.instance Definition _ A := + isPCM_morphism.Build A (option A) Some some_is_morph. +Lemma some_is_binorm {A : pcm} : binorm_pcm_morph_axiom (@Some A). +Proof. by []. Qed. +HB.instance Definition _ A := + isBinorm_PCM_morphism.Build A (option A) Some some_is_binorm. +HB.instance Definition _ A := + isFull_PCM_morphism.Build A (option A) Some (fun _ _ => erefl _). + +(* odflt is full morphism, but doesn't preserve undef (hence, not normal) *) +Lemma odfltu_is_morph {A : pcm} : pcm_morph_axiom relT (@odfltu A). +Proof. by split=>//; case=>[x|][y|]. Qed. +HB.instance Definition _ (A : pcm) := + isPCM_morphism.Build (option A) A odfltu odfltu_is_morph. +HB.instance Definition _ (A : pcm) := + isFull_PCM_morphism.Build (option A) A odfltu (fun _ _ => erefl _). + +(* strenghtening of pfjoin for Some morphism (elides validity) *) +Lemma pfsome (A : pcm) (x y : A) : Some x \+ Some y = Some (x \+ y). +Proof. by []. Qed. + +Lemma odflt_some (A : pcm) (x : A) : odfltu (Some x) = x. +Proof. by []. Qed. + +Lemma odflt_someUn (A : pcm) (x y : A) : odfltu (Some x \+ Some y) = x \+ y. +Proof. by []. Qed. + +Lemma some_odflt (A : pcm) (x : option A) : valid x -> Some (odfltu x) = x. +Proof. by case: x. Qed. + +Lemma some_odfltUn (A : pcm) (x y : option A) : + valid (x \+ y) -> Some (odfltu x \+ odfltu y) = x \+ y. +Proof. by case: x; case: y. Qed. + +(* constructions for option nat, specifically *) + +Lemma option_nat_is_normal : normal_tpcm_axiom (option nat). +Proof. by apply: option_is_normal. Qed. +HB.instance Definition _ := + isNormal_TPCM.Build (option nat) option_nat_is_normal. + +Lemma onat0E (x y : option nat) : x \+ y = Unit -> (x = Unit) * (y = Unit). +Proof. +case: x=>[x|]; case: y=>[y|] // [] /eqP. +by rewrite addn_eq0=>/andP [/eqP -> /eqP ->]. +Qed. + + +(**********) +(**********) +(* SubPCM *) +(**********) +(**********) + +(* Subpcm structure between U and V consists of two morphisms *) +(* pval : U -> V and psub : V -> U with special properties. *) +(* The definition thus resembles Galois connections and adjunctions *) +(* which also postulate existence of morphisms in the opposite directions. *) +(* This formulation allows reasoning about composition *) +(* of subpcm structures *) + +(* We thus define the structure collecting the two morphisms *) +(* (with identity and composition) *) + +Record sub_struct (U V : Type) := SubStruct { + pval : U -> V; + psub : V -> U}. + +Arguments pval : simpl never. +Arguments psub : simpl never. + +Definition id_sub {U} := @SubStruct U U idfun idfun. + +Definition comp_sub U V W (X : sub_struct U V) (Y : sub_struct V W) := + SubStruct (pval Y \o pval X) (psub X \o psub Y). + +(* explicit rewrite rules *) +Lemma pval_id U : pval (@id_sub U) = idfun. Proof. by []. Qed. +Lemma psub_id U : psub (@id_sub U) = idfun. Proof. by []. Qed. +Lemma pval_comp U V W X Y : + pval (@comp_sub U V W X Y) = pval Y \o pval X. +Proof. by []. Qed. +Lemma psub_comp U V W X Y : + psub (@comp_sub U V W X Y) = psub X \o psub Y. +Proof. by []. Qed. + +(* subpcm axoms *) +(* pval - normality needed so that transitions of subresources *) +(* don't require side condition on validity *) +(* pval - fullness is required for simplicity *) +(* psub - binormality needed so that xsep (defined below) is pcm *) +(* remaining two axioms are expected for injection/retraction pair *) + +Definition subpcm_struct_axiom' (U V : pcm) + (pval : full_norm_pcm_morph U V) + (psub : binorm_pcm_morph V U) := + [/\ (* inject then retract is id *) + forall u, psub (pval u) = u & + (* retract then inject is id on valid elements *) + forall v, valid v -> sep psub v Unit -> pval (psub v) = v]. + +Notation subpcm_struct_axiom S := + (subpcm_struct_axiom' (pval S) (psub S)). + +HB.mixin Record isSubPCM_struct (U V : pcm) (S : sub_struct U V) := { + pval_submix : @Full_Norm_PCM_morphism U V (pval S); + psub_submix : @Binorm_PCM_morphism V U (psub S); + subpcm_struct_subproof : + subpcm_struct_axiom' + (Full_Norm_PCM_morphism.Pack pval_submix) + (Binorm_PCM_morphism.Pack psub_submix)}. + +#[short(type=subpcm_struct)] +HB.structure Definition SubPCM_struct U V := {S of isSubPCM_struct U V S}. + +Arguments subpcm_struct_subproof {U V}. + +(* declare pval/psub to be pcm morphisms *) +HB.instance Definition _ (U V : pcm) (S : subpcm_struct U V) := + Full_Norm_PCM_morphism.copy (pval S) + (Full_Norm_PCM_morphism.Pack pval_submix). +HB.instance Definition _ (U V : pcm) (S : subpcm_struct U V) := + Binorm_PCM_morphism.copy (psub S) + (Binorm_PCM_morphism.Pack psub_submix). + +Notation subsep S := (sep (psub S)). + +Section Repack. +Variables (U V : pcm) (S : subpcm_struct U V). + +Lemma psub_pval (x : U) : psub S (pval S x) = x. +Proof. by case: (subpcm_struct_subproof S)=>H ?; apply: H. Qed. + +Lemma pval_psub (x : V) : + valid x -> subsep S x Unit -> pval S (psub S x) = x. +Proof. by case: (subpcm_struct_subproof S)=>?; apply. Qed. + +End Repack. + +Arguments psub_pval {U V}. +Arguments pval_psub {U V}. + +Lemma sep_pval (U V : pcm) (S : subpcm_struct U V) x y : sep (pval S) x y. +Proof. exact: pfT. Qed. + +#[export] Hint Resolve sep_pval : core. + +Section DerivedLemmas. +Variables (U V : pcm) (S : subpcm_struct U V). + +(* unary lemmas *) + +Lemma valid_sepE (x : U) : + valid x = valid (pval S x) && subsep S (pval S x) Unit. +Proof. by rewrite -fpVE /= psub_pval. Qed. + +Lemma valid_sep (x : U) : + valid x -> valid (pval S x) /\ subsep S (pval S x) Unit. +Proof. by rewrite valid_sepE=>/andP. Qed. + +Lemma valid_pval (x : U) : valid x -> valid (pval S x). +Proof. by case/valid_sep. Qed. + +Lemma valid_sepS (x : U) : valid x -> subsep S (pval S x) Unit. +Proof. by case/valid_sep. Qed. + +(* pvalE is full and normal *) +Lemma valid_pvalE (x : U) : valid (pval S x) = valid x. +Proof. by rewrite pfVE. Qed. + +(* iff variant for convenience *) +Lemma valid_pvalEP (x : U) : valid (pval S x) <-> valid x. +Proof. by rewrite valid_pvalE. Qed. + +Lemma valid_pvalS (x : U) : valid (pval S x) -> subsep S (pval S x) Unit. +Proof. by move/valid_pvalEP/valid_sepS. Qed. + +(* we can collect the two conditions of pval_psub into one *) +Lemma valid_psubS (x : V) : valid (psub S x) -> pval S (psub S x) = x. +Proof. by case/fpVI; apply: pval_psub. Qed. + + +(* binary lemmas *) + +Lemma valid_sepUnE (x y : U) : + valid (x \+ y) = + valid (pval S x \+ pval S y) && subsep S (pval S x) (pval S y). +Proof. by rewrite -fpV2E /= !psub_pval. Qed. + +Lemma valid_sepUn (x y : U) : + valid (x \+ y) -> + valid (pval S x \+ pval S y) /\ subsep S (pval S x) (pval S y). +Proof. by rewrite valid_sepUnE=>/andP. Qed. + +Lemma valid_pvalUn (x y : U) : + valid (x \+ y) -> valid (pval S x \+ pval S y). +Proof. by case/valid_sepUn. Qed. + +Lemma valid_sepUnS (x y : U) : valid (x \+ y) -> subsep S (pval S x) (pval S y). +Proof. by rewrite valid_sepUnE=>/andP []. Qed. + +(* binary versions of valid_pvalE, valid_pvalS only hold if *) +(* subsep operates on arguments independently *) +Lemma valid_pvalUnE (x y : U) : + subsep S (pval S x) (pval S y) = + subsep S (pval S x) Unit && subsep S (pval S y) Unit -> + valid (pval S x \+ pval S y) = valid (x \+ y). +Proof. +rewrite valid_sepUnE=>->; case W : (valid _)=>//=. +by rewrite !valid_pvalS ?(validL W, validR W). +Qed. + +Lemma valid_pvalUnS (x y : U) : + subsep S (pval S x) (pval S y) = + subsep S (pval S x) Unit && subsep S (pval S y) Unit -> + valid (pval S x \+ pval S y) -> subsep S (pval S x) (pval S y). +Proof. by move/valid_pvalUnE=>-> /valid_sepUnS. Qed. + + +(* ternary lemmas (occasionally useful) *) + +Lemma valid_sep3E (x y z : U) : + valid (x \+ y \+ z) = + [&& valid (pval S x \+ pval S y \+ pval S z), + subsep S (pval S x) (pval S y) & + subsep S (pval S x \+ pval S y) (pval S z)]. +Proof. +apply/idP/idP=>[/[dup] W /validL D|/and3P [W S1 S2]]. +- by rewrite -pfjoin // andbCA -valid_sepUnE W andbT valid_sepUnS. +by rewrite valid_sepUnE pfjoin ?W // valid_sepUnE (validL W) S1. +Qed. + +Lemma valid_sep3 (x y z : U) : + valid (x \+ y \+ z) -> + [/\ valid (pval S x \+ pval S y \+ pval S z), + subsep S (pval S x) (pval S y) & + subsep S (pval S x \+ pval S y) (pval S z)]. +Proof. by rewrite valid_sep3E=>/and3P. Qed. + + +(* lemmas for zig-zag interaction of psub and pval *) + +Lemma valid_psubUnX (x : V) (y : U) : + valid (psub S x \+ y) = valid (x \+ pval S y) && subsep S x (pval S y). +Proof. by rewrite -{1}(psub_pval S y) fpV2E. Qed. + +Lemma valid_psubXUn (x : U) (y : V) : + valid (x \+ psub S y) = valid (pval S x \+ y) && subsep S (pval S x) y. +Proof. by rewrite -{1}(psub_pval S x) fpV2E. Qed. + +Lemma psubUnX (x : V) (y : U) : + valid (psub S x \+ y) -> psub S x \+ y = psub S (x \+ pval S y). +Proof. by rewrite valid_psubUnX=>/andP [W H]; rewrite pfjoin //= psub_pval. Qed. + +Lemma psubXUn (x : U) (y : V) : + valid (x \+ psub S y) -> x \+ psub S y = psub S (pval S x \+ y). +Proof. by rewrite valid_psubXUn=>/andP [W H]; rewrite pfjoin //= psub_pval. Qed. + +Lemma pvalXUn (x : V) (y : U) : + valid (psub S x \+ y) -> x \+ pval S y = pval S (psub S x \+ y). +Proof. +by move/[dup]=>W /validL/fpVI [/= V1 V2]; rewrite pfjoin //= pval_psub. +Qed. + +Lemma pvalUnX (x : U) (y : V) : + valid (x \+ psub S y) -> pval S x \+ y = pval S (x \+ psub S y). +Proof. by rewrite joinC=>/pvalXUn <-; rewrite joinC. Qed. + + +(* injectivity *) + +Lemma pval_inj : injective (pval S). +Proof. by move=>x y E; rewrite -(psub_pval S x) E psub_pval. Qed. + +Lemma psub_inj (x y : V) : valid (psub S x) -> psub S x = psub S y -> x = y. +Proof. +move/[swap]=>E /[dup]; rewrite {2}E. +case/fpVI=>/= W1 H1 /fpVI [/= W2 H2]. +by rewrite -(pval_psub S _ W1 H1) -(pval_psub S _ W2 H2) E. +Qed. + +End DerivedLemmas. + +Prenex Implicits valid_sepE valid_pvalE valid_pvalEP valid_pvalS valid_psubS +valid_sepUnE valid_pvalUnE valid_pvalUnS valid_sep3E valid_psubUnX valid_psubXUn +psubUnX psubXUn pvalXUn pvalUnX pval_inj psub_inj. + + +(* properties of V propagate to U *) +Section SubPCMPropagation. +Variable U : pcm. + +(* if V is tpcm, psub undef isn't valid *) +Lemma psub_undef (V : tpcm) (S : subpcm_struct U V) : ~~ valid (psub S undef). +Proof. by rewrite fpVE /= valid_undef. Qed. + +(* if V is cancellative, so is U *) +Lemma subpcm_is_cpcm (V : cpcm) (S : subpcm_struct U V) : cpcm_axiom U. +Proof. +move=>x1 x2 x W E; move: (W) (W). +rewrite {1}E !(valid_sepUnE S)=>/andP [W2 D2] /andP [W1 D1]. +move: E; rewrite -(psub_pval S x1) -(psub_pval S x2) -(psub_pval S x). +rewrite -pfjoin // -[R in _ = R]pfjoin //; move/psub_inj. +by rewrite fpVE W1 sepU0 // => /(_ (erefl _)) /(joinKx W1) ->. +Qed. + +(* no canonical projection to latch onto, so no generic inheritance *) +(* but subpcm_is_cpcm can be used for any individual type U *) +(* +HB.instance Definition _ (V : cpcm) (S : subpcm_struct U V) := + isCPCM.Build (PCM.pack_ (PCM.class U)) (subpcm_is_cpcm S). +*) +End SubPCMPropagation. + + +(* identity subpcm structure *) +Lemma id_sub_is_subpcm_struct {U : pcm} : subpcm_struct_axiom (@id_sub U). +Proof. by []. Qed. +HB.instance Definition _ U := + isSubPCM_struct.Build U U id_sub id_sub_is_subpcm_struct. +Lemma subsep_id (U : pcm) : subsep (@id_sub U) = relT. Proof. by []. Qed. + +(* Composition of subpcm structures *) +Section SubPCMStructCompose. +Variables U V W : pcm. +Variables (X : subpcm_struct U V) (Y : subpcm_struct V W). + +Lemma comp_is_subpcm_struct : subpcm_struct_axiom (comp_sub X Y). +Proof. +split=>x/=; first by rewrite !psub_pval. +rewrite /sepx/= pfunit => D /andP [/= H1 H2]. +by rewrite !pval_psub // fpVE /= D H1. +Qed. + +HB.instance Definition _ := + isSubPCM_struct.Build U W (comp_sub X Y) comp_is_subpcm_struct. + +Lemma subsep_comp : + subsep (comp_sub X Y) = + fun y1 y2 => subsep Y y1 y2 && subsep X (psub Y y1) (psub Y y2). +Proof. by []. Qed. + +End SubPCMStructCompose. + +(***********) +(* SubTPCM *) +(***********) + +(* subtpcm structure is a subpcm struct on tpcms *) +(* where pval is further required to preserve undef *) +(* (i.e., pval is tpcm morphism) *) +(* it then follows that psub preserves undef as well *) + +Definition subtpcm_struct_axiom (U V : tpcm) (S : subpcm_struct U V) := + pval S undef = undef. + +HB.mixin Record isSubTPCM_struct (U V : tpcm) + S of @SubPCM_struct U V S := { + subtpcm_struct_subproof : subtpcm_struct_axiom S}. + +#[short(type=subtpcm_struct)] +HB.structure Definition SubTPCM_struct (U V : tpcm) := + {S of isSubTPCM_struct U V S & @SubPCM_struct U V S}. + +Arguments subtpcm_struct_subproof {U V}. + +(* psub preserves undef *) +Lemma psub_is_tpcm_morph (U V : tpcm) (S : subtpcm_struct U V) : + tpcm_morph_axiom (psub S). +Proof. +by rewrite /tpcm_morph_axiom -(subtpcm_struct_subproof S) psub_pval. +Qed. + +(* declare pval/psub to be tpcm morphisms *) +HB.instance Definition _ (U V : tpcm) (S : subtpcm_struct U V) := + isTPCM_morphism.Build U V (pval S) (subtpcm_struct_subproof S). +HB.instance Definition _ (U V : tpcm) (S : subtpcm_struct U V) := + isTPCM_morphism.Build V U (psub S) (psub_is_tpcm_morph S). + +(* identity subtpcm structure *) +HB.instance Definition _ (U : tpcm) := + isSubTPCM_struct.Build U U id_sub (erefl undef). + +(* composition of subtpcm structures *) +Section SubTPCMStructCompose. +Variables U V W : tpcm. +Variables (X : subtpcm_struct U V) (Y : subtpcm_struct V W). +Lemma comp_is_subtpcm_struct : subtpcm_struct_axiom (comp_sub X Y). +Proof. by rewrite /subtpcm_struct_axiom/pval /= !pfundef. Qed. +HB.instance Definition _ := + isSubTPCM_struct.Build U W (comp_sub X Y) comp_is_subtpcm_struct. +End SubTPCMStructCompose. + + +(*****************************) +(*****************************) +(* SubPCM/TPCM constructions *) +(*****************************) +(*****************************) + + +(**********************************) +(* Modding out TPCM V by seprel D *) +(**********************************) + +(* requires TPCM, because morphisms need undef *) +(* to return when D not satisfied *) + +(* xsep is the type of the subpcm *) +(* xsub is the subpcm structure *) +Definition xsepP (V : tpcm) (D : seprel V) (x : V) := + valid x && D x Unit \/ x = undef. +Inductive xsep (V : tpcm) (D : seprel V) := + ex_sep x of xsepP D x. + +(* makes defs opaque to avoid performance penalty *) + +Module Type XSepSig. +Parameter valx : forall (V : tpcm) (D : seprel V), xsep D -> V. +Parameter subx : forall (V : tpcm) (D : seprel V), V -> xsep D. +Definition xsub V D := SubStruct (@valx V D) (@subx V D). + +(* pcm operation *) +Parameter xsep_valid : forall (V : tpcm) (D : seprel V), xsep D -> bool. +Parameter xsep_unit : forall (V : tpcm) (D : seprel V), xsep D. +Parameter xsep_unitb : forall (V : tpcm) (D : seprel V), xsep D -> bool. +Parameter xsep_join' : forall (V : tpcm) (D : seprel V), + xsep D -> xsep D -> V. +Parameter xsep_join : forall (V : tpcm) (D : seprel V), + xsep D -> xsep D -> xsep D. +(* tpcm operations *) +Parameter xsep_undef : forall (V : tpcm) (D : seprel V), xsep D. +Parameter xsep_undefb : forall (V : tpcm) (D : seprel V), xsep D -> bool. + +Parameter valxE : forall (V : tpcm) (D : seprel V), + @valx V D = fun (x : xsep D) => let: ex_sep v _ := x in v. +Parameter subxE : forall (V : tpcm) (D : seprel V), + @subx V D = fun x => + if decP (valid x && D x Unit =P true) is left pf + then ex_sep (or_introl pf) + else ex_sep (or_intror (erefl undef)). +Parameter xsep_validE : forall (V : tpcm) (D : seprel V), + @xsep_valid V D = fun x => valid (valx x) && D (valx x) Unit. +Parameter xsep_unitP : forall (V : tpcm) (D : seprel V), + xsepP D (Unit : V). +Parameter xsep_unitE : forall (V : tpcm) (D : seprel V), + @xsep_unit V D = ex_sep (@xsep_unitP V D). +Parameter xsep_unitbE : forall (V : tpcm) (D : seprel V), + @xsep_unitb V D = fun x => unitb (valx x). +Parameter xsep_joinE' : forall (V : tpcm) (D : seprel V), + @xsep_join' V D = fun x y => + if valid (valx x \+ valx y) && D (valx x) (valx y) + then valx x \+ valx y else undef. +Parameter xsep_joinP : forall (V : tpcm) (D : seprel V) x y, + xsepP D (@xsep_join' V D x y). +Parameter xsep_joinE : forall (V : tpcm) (D : seprel V), + @xsep_join V D = fun x y => ex_sep (xsep_joinP x y). +Parameter xsep_undefP : forall (V : tpcm) (D : seprel V), + xsepP D undef. +Parameter xsep_undefE : forall (V : tpcm) (D : seprel V), + @xsep_undef V D = ex_sep (@xsep_undefP V D). +Parameter xsep_undefbE : forall (V : tpcm) (D : seprel V), + @xsep_undefb V D = fun x => undefb (valx x). +End XSepSig. + +Module XSepSubPCMDef : XSepSig. +Section XSepSubPCMDef. +Variables (V : tpcm) (D : seprel V). + +Definition valx (x : xsep D) := let: ex_sep v _ := x in v. +Definition subx (x : V) := + if decP (valid x && D x Unit =P true) is left pf + then ex_sep (or_introl pf) + else ex_sep (or_intror (erefl undef)). +Definition xsep_valid x := valid (valx x) && D (valx x) Unit. +Lemma xsep_unitP : xsepP D (Unit : V). +Proof. by rewrite /xsepP valid_unit sep00; left. Qed. +Definition xsep_unit := ex_sep xsep_unitP. +Definition xsep_unitb x := unitb (valx x). +Definition xsep_join' x y := + if valid (valx x \+ valx y) && D (valx x) (valx y) + then valx x \+ valx y else undef. +Lemma xsep_joinP x y : xsepP D (xsep_join' x y). +Proof. +rewrite /xsepP /xsep_join'; case: ifP; last by right. +by case/andP=>W /(sepU0 W) ->; rewrite W; left. +Qed. +Definition xsep_join x y := ex_sep (xsep_joinP x y). +Lemma xsep_undefP : xsepP D undef. Proof. by right. Qed. +Definition xsep_undef := ex_sep xsep_undefP. +Definition xsep_undefb (x : xsep D) := undefb (valx x). + +Definition valxE := erefl valx. +Definition subxE := erefl subx. +Definition xsub := SubStruct valx subx. +Definition xsep_validE := erefl xsep_valid. +Definition xsep_unitE := erefl xsep_unit. +Definition xsep_unitbE := erefl xsep_unitb. +Definition xsep_joinE' := erefl xsep_join'. +Definition xsep_joinE := erefl xsep_join. +Definition xsep_undefE := erefl xsep_undef. +Definition xsep_undefbE := erefl xsep_undefb. +End XSepSubPCMDef. +End XSepSubPCMDef. + +Export XSepSubPCMDef. + +Section XSepSubPCM. +Variables (V : tpcm) (D : seprel V). + +(* helper lemma *) +Lemma valx_inj (x y : xsep D) : + valx x = valx y -> + x = y. +Proof. +case: x y=>x Hx [y Hy]; rewrite !valxE => E. +by subst y; rewrite (pf_irr Hx). +Qed. + +(* unary and binary orthogonality relations *) +Notation orth1 x := (valid x && D x Unit). +Notation orth2 x y := (valid (x \+ y) && D x y). + +Notation xsep_valid := (@xsep_valid V D). +Notation xsep_join := (@xsep_join V D). +Notation xsep_unit := (@xsep_unit V D). +Notation xsep_unitb := (@xsep_unitb V D). +Notation xsep_undef := (@xsep_undef V D). +Notation xsep_undefb := (@xsep_undefb V D). + +(* xsep is pcm *) +Lemma xsep_is_pcm : pcm_axiom xsep_valid xsep_join xsep_unit xsep_unitb. +Proof. +have joinC : commutative xsep_join. +- case=>x Hx [y Hy]; apply: valx_inj; rewrite valxE xsep_joinE xsep_joinE'. + by rewrite joinC; case W: (valid _)=>//=; rewrite -sepC. +split=>[//||[x Hx]|x y||x]. +- suff joinAC : right_commutative xsep_join. + - by move=>a b c; rewrite !(joinC a) joinAC. + case=>a Ha [b Hb][c Hc]; apply: valx_inj; rewrite valxE. + rewrite xsep_joinE; do ![rewrite {1}xsep_joinE' !valxE /=]. + case Sab: (orth2 a b); case Sac: (orth2 a c); rewrite ?tpcmE //=; last first. + - case/andP: Sac=>_ Sac; case: andP=>//; case=>W Sacb. + by rewrite (sepAxx W Sac Sacb) andbT (validLE3 W) in Sab. + - case/andP: Sab=>_ Sab; case: andP=>//; case=>W Sabc. + by rewrite (sepAxx W Sab Sabc) andbT (validLE3 W) in Sac. + case/andP: Sab=>_ Sab; case/andP: Sac=>_ Sac. + case Sabc: (orth2 (a \+ b) c). + - case/andP: Sabc=>W Sabc; rewrite sepC (joinAC a c b) W //=. + by rewrite (sepAxx W Sab Sabc). + case Sacb: (orth2 (a \+ c) b)=>//. + case/andP: Sacb=>W Sacb; rewrite sepC (joinAC a b c) W // in Sabc. + by rewrite (sepAxx W Sac Sacb) in Sabc. +- apply: valx_inj; rewrite !valxE /=. + rewrite xsep_joinE xsep_unitE xsep_joinE' !{1}valxE. + rewrite unitL; case: Hx=>[|->]; last by rewrite tpcmE. + by case/andP=>W E; rewrite sepC ?unitL // W E. +- rewrite xsep_validE xsep_joinE valxE xsep_joinE' !{1}valxE /=. + case: ifP; last by rewrite tpcmE. + by case/andP=>W E; rewrite !(validE2 W) (sepx0 W E). +- by rewrite xsep_validE xsep_unitE valxE valid_unit sep00. +rewrite xsep_unitbE xsep_unitE; case: x=>x H /=; rewrite valxE. +case: unitbP=>X; constructor; last by case=>/X. +by rewrite X in H *; rewrite (pf_irr H (@xsep_unitP V D)). +Qed. + +HB.instance Definition _ : isPCM (xsep D) := + isPCM.Build (xsep D) xsep_is_pcm. + +(* xsep is tpcm *) +Lemma xsep_is_tpcm : tpcm_axiom xsep_undef xsep_undefb. +Proof. +split=>[/= x||/= x]. +- rewrite xsep_undefbE xsep_undefE valxE; case: x=>x H /=. + case: undefbP=>X; constructor; last by case=>/X. + by rewrite X in H *; rewrite (pf_irr H (xsep_undefP D)). +- by rewrite pcmE /= xsep_validE xsep_undefE valxE tpcmE. +apply: valx_inj; rewrite xsep_undefE !valxE. +by rewrite /join/= xsep_joinE xsep_joinE' valxE /= !tpcmE. +Qed. + +HB.instance Definition _ : isTPCM (xsep D) := + isTPCM.Build (xsep D) xsep_is_tpcm. + +(* xsep is normal tpcm *) +Lemma xsep_is_normal : normal_tpcm_axiom (xsep D). +Proof. +case=>x [] H; [left|right]. +- by rewrite /valid/= xsep_validE valxE. +by apply/valx_inj; rewrite !valxE /undef /= xsep_undefE. +Qed. + +HB.instance Definition _ : isNormal_TPCM (xsep D) := + isNormal_TPCM.Build (xsep D) xsep_is_normal. + +(* Next, morphisms *) + +(* valx is morphism *) +Lemma valx_is_morph : pcm_morph_axiom relT (@valx V D). +Proof. +split=>[|x y]; first by rewrite valxE /Unit /= xsep_unitE. +rewrite {1}/valid /= xsep_validE !valxE /=. +rewrite pcm_joinE /= xsep_joinE xsep_joinE' !valxE. +by case: ifP; rewrite ?tpcmE //; case/andP. +Qed. + +HB.instance Definition _ := + isPCM_morphism.Build (xsep D) V (@valx V D) valx_is_morph. + +(* valx is full morphism *) +Lemma valx_is_full_morph : full_pcm_morph_axiom (@valx V D). +Proof. by []. Qed. + +HB.instance Definition _ := + isFull_PCM_morphism.Build (xsep D) V (@valx V D) valx_is_full_morph. + +(* valx is normal morphism *) +Lemma valx_is_norm_morph : norm_pcm_morph_axiom (@valx V D). +Proof. +move=>x; rewrite {2}/valid/= xsep_validE. +rewrite !{1}valxE /= => W; split=>//. +by case: x W=>x /= [/andP [->->]|] // ->; rewrite tpcmE. +Qed. + +HB.instance Definition _ := + isNorm_PCM_morphism.Build (xsep D) V (@valx V D) valx_is_norm_morph. + +(* NOTE: valx is *not* binormal morphism as it doesn't preserve separation. *) +(* The subpcm gives new notion of separation that is more stringent *) +(* compared to super-pcm. If valx were binormal, that would imply that *) +(* the new notion isn't more stringent, and thus defeat the purpose *) +(* of the construction. *) +Lemma valx_is_binorm_morph : binorm_pcm_morph_axiom (@valx V D). +Proof. +case=>x Hx [y Hy] /=; rewrite !{1}valxE => W. +rewrite /valid/=/sepx/= xsep_validE /= !{1}valxE /=. +rewrite pcm_joinE /= xsep_joinE /= xsep_joinE' !{1}valxE. +case H : (D x y); first by rewrite !W (sepU0 W H). +rewrite W /=. +Abort. + +(* subx is morphism *) +Lemma subx_is_morph : pcm_morph_axiom D (@subx V D). +Proof. +rewrite subxE; split=>[|x y W E]. +- apply: valx_inj; rewrite !valxE /Unit /= xsep_unitE; case: eqP=>//=. + by rewrite valid_unit /= sep00. +case: eqP=>Hx /=; last by rewrite (sep0E W E) (validE2 W) in Hx. +case: eqP=>Hy /=; last by rewrite (sep0E W E) (validE2 W) in Hy. +case: eqP=>H /=; last by rewrite W (sepU0 W E) in H. +rewrite /valid/= xsep_validE pcm_joinE valxE /= xsep_joinE /=. +do ![rewrite {1}xsep_joinE' valxE]. +rewrite {1 2}W {1 2}E {1}W {1}(sepU0 W E) /=. +split=>//; apply: valx_inj; rewrite valxE /=. +by rewrite xsep_joinE' valxE W E. +Qed. + +HB.instance Definition _ := + isPCM_morphism.Build V (xsep D) (subx D) subx_is_morph. + +(* subx is binormal morphism *) +Lemma subx_is_binorm_morph : binorm_pcm_morph_axiom (subx D). +Proof. +move=>x y /=. +rewrite /sepx/= subxE {1}/valid/= xsep_validE valxE !pcm_joinE /= xsep_joinE. +case: eqP=>Hx; case: eqP=>Hy; rewrite xsep_joinE' valxE; +case H: (valid (x \+ y) && D x y) Hx Hy; +by rewrite /= ?tpcmE //=; case/andP: H. +Qed. + +HB.instance Definition _ := + isBinorm_PCM_morphism.Build V (xsep D) (subx D) subx_is_binorm_morph. + +(* Next, substructures *) + +(* xsub is subpcm *) +Definition xsub := SubStruct (@valx V D) (subx D). + +Lemma xsub_is_subpcm_struct : subpcm_struct_axiom xsub. +Proof. +split=>[x|x] //=; last first. +- rewrite /sepx/= subxE valxE => W H. + by case: eqP=>//=; rewrite W H. +apply: valx_inj; rewrite valxE subxE /=. +by case: eqP; case: x=>// x []. +Qed. + +HB.instance Definition _ := + isSubPCM_struct.Build (xsep D) V xsub xsub_is_subpcm_struct. + +Lemma xsub_is_subtpcm_struct : subtpcm_struct_axiom xsub. +Proof. +rewrite /subtpcm_struct_axiom. +by rewrite /pval/= /undef/= xsep_undefE valxE. +Qed. + +HB.instance Definition _ := + isSubTPCM_struct.Build (xsep D) V xsub xsub_is_subtpcm_struct. +End XSepSubPCM. + +Lemma psub_undefN (V : tpcm) (D : seprel V) (x : V) : + ~~ D x Unit -> + psub (xsub D) x = undef. +Proof. +move=>X; apply: valx_inj. +rewrite /undef/= xsep_undefE valxE /psub/= subxE /=. +by case: decP=>//; rewrite (negbTE X) andbF. +Qed. + +(*****************************************) +(* Normalize = mod out trivially by relT *) +(*****************************************) + +(* removes non-valid elements != undef *) + +Definition normalize (U : tpcm) := xsep (@relT U). +Definition norm_sub {U : tpcm} := xsub (@relT U). + +(* redeclare structures for normalize, to save on unfolding *) +HB.instance Definition _ (U : tpcm) := + Normal_TPCM.on (normalize U). +HB.instance Definition _ (U : tpcm) := + SubTPCM_struct.on (@norm_sub U). +(* psub also becomes full morphism *) +HB.instance Definition _ (U : tpcm) := + isFull_PCM_morphism.Build U (normalize U) (psub norm_sub) + (fun _ _ => erefl _). + +(***************************) +(* Normal product of TPCMs *) +(***************************) + +(* product which is immediately normalized *) +(* to remove spurious invalid elements *) +(* that arose out of pairing invalids of one TPCM *) +(* with valids of the other. *) + +Section NormalProd. +Variables U V : tpcm. + +Definition nprod := normalize (prod U V). +Definition nprod_sub := @norm_sub (prod U V). + +Definition nfst : nprod -> U := fst \o pval nprod_sub. +Definition nsnd : nprod -> V := snd \o pval nprod_sub. +Definition npair : U * V -> nprod := psub nprod_sub. + +End NormalProd. + +Arguments nprod_sub {U V}. +Prenex Implicits nfst nsnd npair nprod_sub. + +(* redeclare again *) +HB.instance Definition _ (U V : tpcm) := + Normal_TPCM.on (nprod U V). +HB.instance Definition _ (U V : tpcm) := + SubTPCM_struct.on (@nprod_sub U V). + +(* redeclare morphisms *) +HB.instance Definition _ (U V : tpcm) := + Full_Binorm_TPCM_morphism.on (@npair U V). + +(* nfst and nsnd are full by inheritance *) +HB.instance Definition _ (U V : tpcm) := + Full_TPCM_morphism.on (@nfst U V). +HB.instance Definition _ (U V : tpcm) := + Full_TPCM_morphism.on (@snd U V). + +(* but now they are also normal *) +(* though not binormal (as they shouldn't) *) +Lemma nfst_is_norm_pcm_morphism U V : norm_pcm_morph_axiom (@nfst U V). +Proof. +move=>/= x W; split=>//; case: (normalP x) W=>// ->. +by rewrite pfundef valid_undef. +Qed. + +HB.instance Definition _ (U V : tpcm) := + isNorm_PCM_morphism.Build (nprod U V) U (@nfst U V) + (@nfst_is_norm_pcm_morphism U V). + +Lemma nsnd_is_norm_pcm_morphism U V : norm_pcm_morph_axiom (@nsnd U V). +Proof. +move=>/= x W; split=>//; case: (normalP x) W=>// ->. +by rewrite pfundef valid_undef. +Qed. + +HB.instance Definition _ (U V : tpcm) := + isNorm_PCM_morphism.Build (nprod U V) V (@nsnd U V) + (@nsnd_is_norm_pcm_morphism U V). + +(* lemmas for normal products *) + +Lemma nprod_eta (U V : tpcm) (x : nprod U V) : + x = npair (nfst x, nsnd x). +Proof. by rewrite -prod_eta /npair psub_pval. Qed. + +Lemma nfst_pair (U V : tpcm) (x : U * V) : + valid x -> nfst (npair x) = x.1. +Proof. by move=>W; rewrite /nfst /= pval_psub. Qed. + +Lemma nsnd_pair (U V : tpcm) (x : U * V) : + valid x -> nsnd (npair x) = x.2. +Proof. by move=>W; rewrite /nsnd /= pval_psub. Qed. + +Lemma nprodV (U V : tpcm) (x : nprod U V) : + valid x = valid (nfst x) && valid (nsnd x). +Proof. by rewrite {1}[x]nprod_eta pfVE. Qed. + +Lemma nprodUnV (U V : tpcm) (x y : nprod U V) : + valid (x \+ y) = valid (nfst x \+ nfst y) && + valid (nsnd x \+ @nsnd U V y). +Proof. by rewrite {1}[x]nprod_eta {1}[y]nprod_eta pfV2E. Qed. + + +(***********************) +(* Relativized seprels *) +(***********************) + +(* Relation S is seprel-relative-to-morphism-f *) +(* if it becomes seprel once composed with f and paired with sep f. *) +(* Useful when sequencing subpcm constructions *) + +Definition seprel_on_axiom U V (f : pcm_morph U V) (S : rel V) := + [/\ S Unit Unit, + forall x y, valid (x \+ y) -> sep f x y -> + S (f x) (f y) = S (f y) (f x), + forall x y, valid (x \+ y) -> sep f x y -> + S (f x) (f y) -> S (f x) Unit & + forall x y z, valid (x \+ y \+ z) -> + sep f x y -> sep f (x \+ y) z -> + S (f x) (f y) -> S (f x \+ f y) (f z) -> + S (f y) (f z) && S (f x) (f y \+ f z)]. + +(* S is seprel-on-f iff *) +(* relI (sep f) (rel_app S f) is seprel on U *) +Lemma seprel_on_app U V (f : pcm_morph U V) S : + seprel_on_axiom f S <-> + seprel_axiom (relI (sep f) (rel_app S f)). +Proof. +split=>[[H1 H2 H3 H4]|[/andP [H1 H2 H3 H4 H5]]]. +- split=>[|x y W|x y W|x y z W] /=. + - by rewrite pfunit sep00 H1. + - rewrite (sepC _ W) /=. + by case X: (sep f y x)=>//; rewrite H2 // sepC. + - by case/andP=>X1 X2; rewrite pfunit (sep0E W) // (H3 _ _ W). + case/andP=>X1 X2 /andP [X3 X4]. + rewrite pfjoin ?(sepAxx W X1 X3) ?(validLE3 W) //=. + by rewrite H4 // -pfjoin // (validL W). +rewrite /= !pfunit in H2 H4; split=>[//|x y|x y|x y z] W. +- rewrite (sepC (sep f) W) => X. + by move: (H3 x y W); rewrite /= (sepC (sep f) W) /= X. +- by move=>X; move: (H4 x y W); rewrite X (sep0E W). +move=>X1 X2 X3 X4; case: (sepAx W X1 X2)=>/= X5 X6. +move: (H5 x y z W); rewrite /= X1 X2 X5 X6 !pfjoin ?(validLE3 W) //=. +by apply. +Qed. + +HB.mixin Record isSeprelOn U V (f : pcm_morph U V) (S : rel V) := { + seprel_on_subproof : seprel_on_axiom f S}. + +#[short(type=seprel_on)] +HB.structure Definition SeprelOn U V f := {S of @isSeprelOn U V f S}. + +Section Repack. +Variables (U V : pcm) (f : pcm_morph U V) (S : seprel_on f). + +Lemma sepon00 : S Unit Unit. +Proof. by case: (@seprel_on_subproof U V f S). Qed. + +Lemma seponC x y : + valid (x \+ y) -> + sep f x y -> + S (f x) (f y) = S (f y) (f x). +Proof. by case: (@seprel_on_subproof U V f S)=>_ H _ _; apply: H. Qed. + +Lemma seponx0 x y : + valid (x \+ y) -> + sep f x y -> + S (f x) (f y) -> + S (f x) Unit. +Proof. by case: (@seprel_on_subproof U V f S)=>_ _ H _; apply: H. Qed. + +Lemma seponAx x y z : + valid (x \+ y \+ z) -> + sep f x y -> + sep f (x \+ y) z -> + S (f x) (f y) -> + S (f x \+ f y) (f z) -> + S (f y) (f z) * S (f x) (f y \+ f z). +Proof. +case: (@seprel_on_subproof U V f S)=>_ _ _ H W R1 R2 R3 R4. +by rewrite !(andX (H _ _ _ W R1 R2 R3 R4)). +Qed. + +(* derived lemmas *) + +Lemma sepon0x x y : + valid (x \+ y) -> + sep f x y -> + S (f x) (f y) -> + S Unit (f y). +Proof. +move=>W sf; rewrite seponC // => Sf. +rewrite sepC //= in sf; rewrite joinC in W. +rewrite -(pfunit f) -(@seponC y) ?pfunit ?unitR ?(validE2 W) //. +- by apply: seponx0 W sf Sf. +by apply: sepx0 W sf. +Qed. + +Lemma sepon0E x y : + valid (x \+ y) -> + sep f x y -> + S (f x) (f y) -> + S (f x) Unit * S (f y) Unit. +Proof. +move=>W sf Sf; rewrite (seponx0 W sf Sf). +rewrite -(pfunit f) seponC ?pfunit ?unitR ?(validE2 W) //. +- by rewrite (sepon0x W sf Sf). +by rewrite (sep0E W sf). +Qed. + +Lemma seponAx23_helper x y z : + valid (x \+ y \+ z) -> + sep f x y -> + sep f (x \+ y) z -> + S (f x) (f y) -> + S (f x \+ f y) (f z) -> + ((S (f x) (f z) * S (f z) (f x)) * (S (f y) (f z) * S (f z) (f y))) * + ((S (f x) (f y \+ f z) * S (f x) (f z \+ f y)) * + (S (f y) (f x \+ f z) * S (f y) (f z \+ f x))). +Proof. +move=>W s1 s2 S1 S2. +rewrite !(@seponC z) ?(validLE3 W) ?(sepAxx W s1 s2) ?(joinC (f z)) //. +rewrite !(seponAx W s1 s2 S1 S2). +rewrite seponC ?(validLE3 W) // in S1. +rewrite sepC ?(validLE3 W) //= in s1. +rewrite (joinC x) in W s2; rewrite (joinC (f x)) in S2. +by rewrite !(seponAx W s1 s2 S1 S2). +Qed. + +Lemma seponxA x y z : + valid (x \+ (y \+ z)) -> + sep f y z -> + sep f x (y \+ z) -> + S (f y) (f z) -> + S (f x) (f y \+ f z) -> + S (f x \+ f y) (f z) * S (f x) (f y). +Proof. +move=>W s1 s2 S1 S2; have /= R := (sepxxA W s1 s2, validRE3 W). +rewrite -pfjoin ?R // seponC ?R // pfjoin ?R // in S2. +rewrite sepC ?R //= in s2. +rewrite -pfjoin ?R // seponC ?R // pfjoin ?R //. +by rewrite !(seponAx23_helper _ s1 s2 S1 S2) ?R. +Qed. + +Lemma seponAxx x y z : + valid (x \+ y \+ z) -> + sep f x y -> + sep f (x \+ y) z -> + S (f x) (f y) -> + S (f x \+ f y) (f z) -> + (((S (f x) (f y \+ f z) * S (f x) (f z \+ f y)) * + (S (f y) (f x \+ f z) * S (f y) (f z \+ f x))) * + ((S (f z) (f x \+ f y) * S (f z) (f y \+ f x)) * + (S (f y \+ f z) (f x) * S (f z \+ f y) (f x)))) * + (((S (f x \+ f z) (f y) * S (f z \+ f x) (f y)) * + (S (f x \+ f y) (f z) * S (f y \+ f x) (f z))) * + (((S (f x) (f y) * S (f y) (f x)) * + (S (f x) (f z) * S (f z) (f x))))) * + (S (f y) (f z) * S (f z) (f y)). +Proof. +move=>W s1 s2 S1 S2; have /= R := (sepAxx W s1 s2, validLE3 W). +move: (seponAx23_helper W s1 s2 S1 S2)=>E; rewrite S1 S2 !E. +rewrite -!pfjoin ?R // -!(@seponC x) ?R // !pfjoin ?R // S1 !E. +rewrite -!pfjoin ?R // !(@seponC z) ?R // !pfjoin ?R //. +rewrite (joinC (f y)) S2. +rewrite -!pfjoin ?R // -!(@seponC y) ?R // !pfjoin ?R //. +by rewrite (joinC (f z)) E. +Qed. + +Lemma seponxxA x y z : + valid (x \+ (y \+ z)) -> + sep f y z -> + sep f x (y \+ z) -> + S (f y) (f z) -> + S (f x) (f y \+ f z) -> + (((S (f x) (f y \+ f z) * S (f x) (f z \+ f y)) * + (S (f y) (f x \+ f z) * S (f y) (f z \+ f x))) * + ((S (f z) (f x \+ f y) * S (f z) (f y \+ f x)) * + (S (f y \+ f z) (f x) * S (f z \+ f y) (f x)))) * + (((S (f x \+ f z) (f y) * S (f z \+ f x) (f y)) * + (S (f x \+ f y) (f z) * S (f y \+ f x) (f z))) * + (((S (f x) (f y) * S (f y) (f x)) * + (S (f x) (f z) * S (f z) (f x))))) * + (S (f y) (f z) * S (f z) (f y)). +Proof. +move=>W s1 s2 S1 S2; have /= R := (sepxxA W s1 s2, validRE3 W). +rewrite -pfjoin ?R // seponC // pfjoin ?R // in S2. +rewrite sepC //= in s2; rewrite joinC in W. +rewrite !(seponAx23_helper W s1 s2 S1 S2). +by rewrite !(seponAxx W s1 s2 S1 S2). +Qed. + +Lemma seponU0 x y : + valid (x \+ y) -> + sep f x y -> + S (f x) (f y) -> + S (f x \+ f y) Unit. +Proof. +move=>W s1 S1; rewrite -(pfunit f). +rewrite seponAxx ?pfunit ?unitL //. +- by rewrite sepC ?unitL ?(validL W) ?(sepx0 W s1). +rewrite seponC // in S1; rewrite sepC //= in s1; rewrite joinC in W. +by apply: (@sepon0x y). +Qed. + +Lemma sepon0U x y : + valid (x \+ y) -> + sep f x y -> + S (f x) (f y) -> + S Unit (f x \+ f y). +Proof. +move=>W s1 S1. +rewrite -(pfunit f) -pfjoin // seponC ?unitL ?(sep0U W s1) //. +by rewrite pfjoin ?pfunit ?seponU0. +Qed. + +End Repack. + + +(* if f is full tpcm morphism *) +(* S is seprel-on-f iff (S o f) is seprel on U *) +Lemma seprel_on_full U V (f : full_pcm_morph U V) S : + seprel_on_axiom f S <-> + seprel_axiom (rel_app S f). +Proof. +rewrite seprel_on_app; apply/sepXEQ=>x y W. +by rewrite /relI /= pfSE. +Qed. + +Lemma seprel_on_id (U : pcm) (S : rel U) : + seprel_on_axiom idfun S <-> + seprel_axiom S. +Proof. by rewrite seprel_on_full; apply/sepXEQ. Qed. + +(* S is seprel-on-f iff (S o f) is seprel on subpcm U/(sep f) *) +Lemma seprel_on_pval (U : tpcm) V (f : pcm_morph U V) S : + seprel_on_axiom f S <-> + seprel_axiom (rel_app S (f \o pval (xsub (sep f)))). +Proof. +set xf := (xsub (sep f)). +split; case=>/= H1 H2 H3 H4. +- split=>[|x y W|x y W|x y z W /= X1 X2] /=. + - by rewrite !pfunit. + - by case/(valid_sepUn xf): W=>W /(H2 _ _ W). + - by rewrite !pfunit; case/(valid_sepUn xf): W=>W /(H3 _ _ W). + case: (valid_sepUn xf W)=>Wxyz Sxyz. + case: (valid_sepUn xf (validL W))=>Wxy Sxy. + case: (valid_sepUn xf (validAR W))=>Wyz Syz. + rewrite !pfjoin ?(validLE3 W) //=. + rewrite H4 ?pfV3 // -!pfjoin ?(validLE3 W) //=. + by rewrite pfjoin ?(validLE3 W). +rewrite pfunit in H1 H3. +split=>[|x y W X|x y W X|x y z W X1 X2]. +- by rewrite pfunit in H1. +- rewrite -(pval_psub xf x) ?(sep0E W X,validL W) //. + rewrite -(pval_psub xf y) ?(sep0E W X,validR W) //. + by apply/H2/pfV2. +- rewrite -(pval_psub xf x) ?(sep0E W X,validL W) //. + rewrite -(pval_psub xf y) ?(sep0E W X,validR W) //. + by move/H3; rewrite pfunit; apply; apply: (pfV2 (psub xf)). +rewrite -!pfjoin ?(sepAxx W X1 X2,validLE3 W) //. +rewrite -(pval_psub xf x) ?(sep0E _ X1,validLE3 W) //. +rewrite -(pval_psub xf y) ?(sep0E _ X1,validLE3 W) //. +rewrite -(pval_psub xf z) ?(sep0E _ X2,validLE3 W) //. +rewrite -!(pfjoin (pval xf)) ?(pfV2,sepAxx W X1 X2,validLE3 W) //. +by apply/H4/pfV3. +Qed. + +(* when the morphism is restriction of idfun with seprel K *) +(* then S is seprel alongside K *) +Lemma seprel_with_on (U : tpcm) (K : seprel U) (S : rel U) : + seprel_on_axiom (res idfun K) S <-> + seprel_axiom (relI K S). +Proof. +rewrite seprel_on_app; apply/sepXEQ=>x y V. +rewrite /sepx/= /res/= andbT. +by case X: (K x y)=>//=; rewrite !(sep0E _ X). +Qed. + +Notation "'seprel_with' K" := (seprel_on (res idfun K)) + (at level 1, format "'seprel_with' K"). + +(* Every seprel is trivially seprel_on idfun morphism *) +Section SeprelOnIdCoercion. +Variables (U : pcm) (S : seprel U). +(* dummy name to enable inheritance *) +Definition eta_rel : rel U := fun x y => S x y. +Lemma seprel_is_seprelonid : seprel_on_axiom idfun eta_rel. +Proof. by apply/seprel_on_id/seprel_subproof. Qed. +HB.instance Definition _ := + isSeprelOn.Build U U idfun eta_rel seprel_is_seprelonid. +Definition seprel_on_idfun : seprel_on idfun := eta_rel. +End SeprelOnIdCoercion. + +Coercion seprel_on_idfun : seprel >-> seprel_on. + +(* naming the completion seprel for specific case *) +(* of seprel_on relations where the morphism is full *) +Section SeprelOnFull. +Variables (U V : pcm) (f : full_pcm_morph U V) (S : seprel_on f). +Definition onsep := rel_app S f. +Lemma onsep_is_seprel : seprel_axiom onsep. +Proof. by apply/seprel_on_full/seprel_on_subproof. Qed. +HB.instance Definition _ := + isSeprel.Build U onsep onsep_is_seprel. +End SeprelOnFull. + +(* trivially true seprel is seprel_on for any morphism *) + +Lemma relT_is_on U V (f : pcm_morph U V) : seprel_on_axiom f relT. +Proof. by []. Qed. + +HB.instance Definition _ U V (f : pcm_morph U V) := + isSeprelOn.Build U V f relT (relT_is_on f). + +(* conjunction of seprel-on's is seprel-on *) +Section SepOnI. +Variables (U V : pcm) (f : pcm_morph U V) (X Y : seprel_on f). + +Lemma relI_is_seprelon : seprel_on_axiom f (relI X Y). +Proof. +split=>[|x y W s|x y W s|x y z W s1 s2] /=. +- by rewrite !sepon00. +- by rewrite !(seponC _ W s). +- by case/andP; do ![move/(seponx0 W s) ->]. +case/andP=>X1 Y1 /andP [X2 Y2]. +by rewrite !(seponAxx W s1 s2 X1 X2) !(seponAxx W s1 s2 Y1 Y2). +Qed. + +HB.instance Definition _ := + isSeprelOn.Build U V f (relI X Y) relI_is_seprelon. +End SepOnI. + +(* 3-way conjunction *) +Section SepOn3I. +Variables (U V : pcm) (f : pcm_morph U V) (X Y Z : seprel_on f). + +Lemma rel3I_is_seprelon : seprel_on_axiom f (rel3I X Y Z). +Proof. +split=>[|x y W s|x y W s|x y z W s1 s2] /=. +- by rewrite !sepon00. +- by rewrite !(seponC _ W s). +- by case/and3P; do ![move/(seponx0 W s) ->]. +case/and3P=>X1 Y1 Z1 /and3P [X2 Y2 Z2]. +by rewrite !(seponAxx W s1 s2 X1 X2, seponAxx W s1 s2 Y1 Y2, + seponAxx W s1 s2 Z1 Z2). +Qed. + +HB.instance Definition _ := + isSeprelOn.Build U V f (rel3I X Y Z) rel3I_is_seprelon. +End SepOn3I. + +(* 4-way conjunction *) +Section SepOn4I. +Variables (U V : pcm) (f : pcm_morph U V) (X Y Z W : seprel_on f). + +Lemma rel4I_is_seprelon : seprel_on_axiom f (rel4I X Y Z W). +Proof. +split=>[|x y VV s|x y VV s|x y z VV s1 s2] /=. +- by rewrite !sepon00. +- by rewrite !(seponC _ VV s). +- by case/and4P; do ![move/(seponx0 VV s) ->]. +case/and4P=>X1 Y1 Z1 W1 /and4P [X2 Y2 Z2 W2]. +by rewrite !(seponAxx VV s1 s2 X1 X2, seponAxx VV s1 s2 Y1 Y2, + seponAxx VV s1 s2 Z1 Z2, seponAxx VV s1 s2 W1 W2). +Qed. + +HB.instance Definition _ := + isSeprelOn.Build U V f (rel4I X Y Z W) rel4I_is_seprelon. +End SepOn4I. + +(*******************) +(* Local functions *) +(*******************) + +(* Unlike morphisms which operate over a single PCM element *) +(* local functions operate over a pair (self, other) *) +(* and then have to support a form of framing, i.e., moving from other *) +(* to self. I haven't quite yet found a need to develop their theory *) +(* but, they could be something we can call "bimorphisms" *) + +(* Transitions should fit into this category of local functions *) + +Definition local_fun (U : pcm) (f : U -> U -> option U) := + forall p x y r, valid (x \+ (p \+ y)) -> f x (p \+ y) = Some r -> + valid (r \+ p \+ y) /\ f (x \+ p) y = Some (r \+ p). + +Lemma localV U f x y r : + @local_fun U f -> valid (x \+ y) -> f x y = Some r -> valid (r \+ y). +Proof. by move=>H V E; move: (H Unit x y r); rewrite unitL !unitR; case. Qed. + +Lemma idL (U : pcm) : @local_fun U (fun x y => Some x). +Proof. by move=>p x y _ V [<-]; rewrite -joinA. Qed. + +(********************) +(* Global functions *) +(********************) + +(* given function over two arguments, make it global *) +(* by making it operate over the join *) + +Definition glob (U : pcm) R (f : U -> U -> R) := + fun x y => f (x \+ y) Unit. +Arguments glob {U R} f x y /. + +(*******************) +(* Local relations *) +(*******************) + +(* Local relations are needed at some places *) +(* but are weaker than separating relations *) +(* For example, separating relation would allow moving p from y to x *) +(* only if R p y; this is the associativity property *) +(* of seprels, and is essential for the subPCM construction *) +(* But here we don't require that property, because we won't be *) +(* modding out U by a local rel to obtain a subPCM *) +(* Also, we don't require any special behavior wrt unit. *) +(* And no commutativity (for now) *) +(* Also, its sometimes useful to have a condition under *) +(* which the relation is local *) +(* In practice, it frequently happens that some relation is a seprel *) +(* only if some other seprel already holds. Thus, it makes sense to *) +(* consider locality under a binary condition too. *) + +(* Local rel is the main concept *) +Definition local_rel (U : pcm) (R : U -> U -> Prop) (cond : U -> Prop) := + forall p x y, cond (x \+ p \+ y) -> R x (p \+ y) -> R (x \+ p) y. + + diff --git a/pcm/mutex.v b/pcm/mutex.v new file mode 100644 index 0000000..d7af5e1 --- /dev/null +++ b/pcm/mutex.v @@ -0,0 +1,583 @@ +(* +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 defines the generalized mutex type. *) +(* We need a notion of mutex where the options are not just nown/own but *) +(* an arbitrary number of elements in-between. This reflects the possible *) +(* stages of acquiring a lock. Once a thread embarks on the first stage, it *) +(* excludes others, but it may require more steps to fully acquire the lock. *) +(******************************************************************************) + +From HB Require Import structures. +From Coq Require Import ssreflect ssrbool ssrfun. +From Coq Require Setoid. +From mathcomp Require Import ssrnat eqtype seq. +From pcm Require Import options prelude pred. +From pcm Require Import pcm morphism natmap. + +(***********************) +(* Generalized mutexes *) +(***********************) + +(* T encodes mutex stages, excluding undef and unit *) + +Inductive mutex T := mx_undef | nown | mx of T. + +Arguments mx_undef {T}. +Arguments nown {T}. +Prenex Implicits mx_undef nown. + +Section GeneralizedMutex. +Variable T : Type. +Definition def_mx_valid := [fun x : mutex T => + if x is mx_undef then false else true]. +Definition def_mx_join := [fun x y : mutex T => + match x, y with + mx_undef, _ => mx_undef + | _, mx_undef => mx_undef + | nown, x => x + | x, nown => x + | mx _, mx _ => mx_undef + end]. +Definition def_mx_unitb := [fun x : mutex T => + if x is nown then true else false]. + +Lemma mutex_is_pcm : pcm_axiom def_mx_valid def_mx_join nown def_mx_unitb. +Proof. +by split=>[[||x][||y]|[||x][||y][||z]|[]|[]||[||x]] //; constructor. +Qed. +HB.instance Definition _ : isPCM (mutex T) := + isPCM.Build (mutex T) mutex_is_pcm. + +(* cancellativity *) +Lemma mutex_is_cpcm : cpcm_axiom (mutex T). +Proof. by case=>[||m][||m1][||m2]; rewrite !pcmE. Qed. +HB.instance Definition _ : isCPCM (mutex T) := + isCPCM.Build (mutex T) mutex_is_cpcm. + +(* topped structure *) +Definition def_mx_undefb (x : mutex T) : bool := + if x is mx_undef then true else false. + +Lemma mutex_is_tpcm : tpcm_axiom mx_undef def_mx_undefb. +Proof. by split=>//; case=>[||x]; constructor. Qed. +HB.instance Definition _ : isTPCM (mutex T) := + isTPCM.Build (mutex T) mutex_is_tpcm. + +(* normality *) +Lemma mutex_is_normal : normal_tpcm_axiom (mutex T). +Proof. by case; [right|rewrite valid_unit; left|move=>t; left]. Qed. +HB.instance Definition _ : isNormal_TPCM (mutex T) := + isNormal_TPCM.Build (mutex T) mutex_is_normal. +End GeneralizedMutex. + +(* if T is eqType, so is mutex T *) +Section Equality. +Variable T : eqType. + +Definition mutex_eq (x y : mutex T) := + match x, y with + mx_undef, mx_undef => true + | nown, nown => true + | mx x', mx y' => x' == y' + | _, _ => false + end. + +Lemma mutex_eqP : Equality.axiom mutex_eq. +Proof. +case=>[||x]; case=>[||y] /=; try by constructor. +by case: eqP=>[->|H]; constructor=>//; case=>/H. +Qed. + +HB.instance Definition _ := hasDecEq.Build (mutex T) mutex_eqP. +End Equality. + +(* 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)). + +(* some lemmas for generalized mutexes *) + +Section MutexLemmas. +Variable T : Type. +Implicit Types (t : T) (x y : mutex T). + +Variant mutex_spec x : mutex T -> Type := +| mutex_undef of x = undef : mutex_spec x undef +| mutex_nown of x = Unit : mutex_spec x Unit +| 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. + +Lemma mxE0 x y : x \+ y = Unit -> (x = Unit) * (y = Unit). +Proof. by case: x; case: y. 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. + +Lemma cancelxM t1 t2 x : (x \+ mx t1 = mx t2) <-> (t1 = t2) * (x = Unit). +Proof. by rewrite joinC cancelMx. Qed. + +(* DEVCOMMENT *) +(* the next batch of lemmas about validity should be automated *) +(* /DEVCOMMENT *) +Lemma mxMx t x : valid (mx t \+ x) -> x = Unit. +Proof. by case: x. Qed. + +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 joinA=>/mxxM. Qed. + +Lemma mxMxy t x y : valid (mx t \+ x \+ y) -> x \+ y = Unit. +Proof. by rewrite -joinA=>/mxMx. Qed. + +Lemma mxxMy t x y : valid (x \+ (mx t \+ y)) -> x \+ y = Unit. +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. + +(* 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 = Unit & y = Unit : mxjoin_spec x y Unit Unit Unit +| leftmx t of x = mx t & y = Unit : mxjoin_spec x y (mx t) (mx t) Unit +| rightmx t of x = Unit & y = mx t : mxjoin_spec x y (mx t) Unit (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. + +(* DEVCOMMENT *) +(* and the same for own; do we need to repeat? *) +(* /DEVCOMMENT *) +Section OwnMutex. +Variables T : Type. +Implicit Types x y : mtx T. + +Lemma mxOx x : valid (own \+ x) -> x = Unit. +Proof. by apply: mxMx. Qed. + +Lemma mxxO x : valid (x \+ own) -> x = Unit. +Proof. by apply: mxxM. Qed. + +Lemma mxxyO x y : valid (x \+ (y \+ own)) -> x \+ y = Unit. +Proof. by apply: mxxyM. Qed. + +Lemma mxOxy x y : valid (own \+ x \+ y) -> x \+ y = Unit. +Proof. by apply: mxMxy. Qed. + +Lemma mxxOy x y : valid (x \+ (own \+ y)) -> x \+ y = Unit. +Proof. by apply: mxxMy. Qed. + +Lemma mxyOx x y : valid (y \+ own \+ x) -> y \+ x = Unit. +Proof. by apply: mxyMx. Qed. +End OwnMutex. + +Prenex Implicits mxOx mxxO mxxyO mxOxy mxxOy mxyOx. + +(* specific lemmas for binary mutexes *) + +(* DEVCOMMENT *) +(* these also are a bit of a featuritis *) +(* /DEVCOMMENT *) +Lemma mxON (x : mtx2) : valid x -> x != own -> x = Unit. +Proof. by case: x=>//; case. Qed. + +Lemma mxNN (x : mtx2) : valid x -> x != Unit -> x = own. +Proof. by case: x=>//; case=>//; case. Qed. + +(* the next batch of lemmas is for automatic simplification *) + +Section MutexRewriting. +Variable T : eqType. +Implicit Types (t : T) (x : mutex T). + +Lemma mxE1 : (((undef == Unit :> mutex T) = false) * + ((Unit == undef :> mutex T) = false)). +Proof. by []. Qed. + +Lemma mxE2 t : (((mx t == Unit) = false) * + ((Unit == mx t) = false)) * + (((mx t == undef) = false) * + ((undef == mx t) = false)). +Proof. by []. Qed. + +Lemma mxE3 t x : ((((mx t \+ x == Unit) = false) * + ((x \+ mx t == Unit) = false)) * + (((Unit == mx t \+ x) = false) * + ((Unit == 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))). +Proof. +have L : forall t1 t2 x, (mx t1 \+ x == mx t2) = (t1 == t2) && (x == Unit). +- 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_injE t1 t2 : (mx t1 == mx t2) = (t1 == t2). +Proof. by []. Qed. + +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))). + +End MutexRewriting. + +(* function mapping all mx's to own *) +Definition mxown {T} (x : mutex T) : mtx2 := + match x with + mx_undef => undef + | nown => Unit + | _ => own + end. + +(* this is a morphism under full domain *) +Lemma mxown_is_pcm_morph {T} : pcm_morph_axiom relT (@mxown T). +Proof. by split=>[|x y] //; case: x; case: y. Qed. +HB.instance Definition _ T := + isPCM_morphism.Build (mutex T) mtx2 mxown mxown_is_pcm_morph. +HB.instance Definition _ T := + isFull_PCM_morphism.Build (mutex T) mtx2 (@mxown T) (fun _ _ => erefl _). +HB.instance Definition _ T := + isTPCM_morphism.Build (mutex T) mtx2 mxown (erefl undef). +Lemma mxown_is_binormal {T} : binorm_pcm_morph_axiom (@mxown T). +Proof. by case=>[||x][||y]. Qed. +HB.instance Definition _ T := + isBinorm_PCM_morphism.Build (mutex T) mtx2 mxown mxown_is_binormal. + +(* 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 = Unit -> x = Unit. +Proof. by case: x. Qed. + +Definition mxundef {T} (x : mtx T) : mtx2 := + match x with + | nown => Unit + | 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 sep_mxundef_is_seprel {T} : seprel_axiom (@sep_mxundef T). +Proof. by split=>//; case=>[||x] // [||y] // [||[]] //=; case: y. Qed. +HB.instance Definition _ T := + isSeprel.Build (mtx T) sep_mxundef sep_mxundef_is_seprel. + +Lemma mxundef_is_pcm_morph {T} : pcm_morph_axiom (@sep_mxundef T) mxundef. +Proof. by split=>//; case=>[||x] // [||[]] //; case: x. Qed. +HB.instance Definition _ T := + isPCM_morphism.Build (mtx T) mtx2 mxundef mxundef_is_pcm_morph. +HB.instance Definition _ T := + isTPCM_morphism.Build (mtx T) mtx2 mxundef (erefl undef). +Lemma mxundef_is_binormal {T} : binorm_pcm_morph_axiom (@mxundef T). +Proof. by case=>[||[x|]][||[y|]]. Qed. +HB.instance Definition _ T := + isBinorm_PCM_morphism.Build (mtx T) mtx2 mxundef mxundef_is_binormal. + +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 Unit Unit +| 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 Unit else own. + +Variant nxown_spec n : mtx2 -> Type := +| nxZ of n = 0 : nxown_spec n Unit +| 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 Unit +| 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. + +(**********************************) +(* 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. +Variable U : natmap (bool * bool). + +Definition omg_sep := fun x y : U => + [&& last_atval false x ==> (last_key y < last_key x) & + last_atval false y ==> (last_key x < last_key y)]. + +Lemma omg_is_seprel : seprel_axiom omg_sep. +Proof. +rewrite /omg_sep; split=>[|x y|x y|x y z] /=. +- by rewrite lastatval0. +- by rewrite andbC. +- move=>V /andP [H _]; rewrite lastkey0 lastatval0 /=. + by case: (x in x ==> _) H=>// /(leq_trans _) ->. +move=>V /andP [X Y] /andP []. +rewrite !lastkeyUn !lastatvalUn !(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. + +HB.instance Definition _ := isSeprel.Build U omg_sep omg_is_seprel. + +Definition omg (x : U) : mtx2 := + if undefb x then undef else + if last_atval false x then own else nown. + +(* DEVCOMMENT *) +(* omg isn't tpcm morphism because it doesn't preserve undef *) +(* this makes it less useful than it might be *) +(* but we wait with fixing the definition until it becomes necessary *) +(* (the definition should branch on x being undef *) +(* /DEVCOMMENT *) + +Lemma omg_is_pcm_morph : pcm_morph_axiom omg_sep omg. +Proof. +rewrite /omg; split=>[|x y V /andP [X Y]] /=. +- by rewrite undefb0 lastatval0. +rewrite !undefNV !(validE2 V) /= lastatvalUn V; case: ltngtP X Y=>H X Y; +by rewrite ?(negbTE X) ?(negbTE Y) //; case: ifP. +Qed. +HB.instance Definition _ := isPCM_morphism.Build U mtx2 omg omg_is_pcm_morph. + +Lemma omg_is_tpcm_morph : tpcm_morph_axiom omg. +Proof. by rewrite /tpcm_morph_axiom/omg undefb_undef. Qed. +HB.instance Definition _ := isTPCM_morphism.Build U mtx2 omg omg_is_tpcm_morph. + +Lemma omg_is_norm : norm_pcm_morph_axiom omg. +Proof. +move=>/= x; rewrite /sepx/=/omg/omg_sep lastkey0 lastatval0. +case: (normalP x)=>// Vx; case: ifP=>H _ //=. +by case: lastkeyP (lastatval_indomb H). +Qed. +HB.instance Definition _ := isNorm_PCM_morphism.Build U mtx2 omg omg_is_norm. + +(* transfer lemmas *) + +(* omg isn't full, but admits valid h -> valid (omg h) *) +Lemma valid_omg h : valid (omg h) = valid h. +Proof. +apply/idP/idP=>[/fpV//|V]. +rewrite pfV //= /sepx/= /omg_sep /= lastatval0 andbT lastkey0. +case N: (last_key h); last by apply/implyP. +by rewrite /last_atval /atval /= cond_find // N. +Qed. + +Lemma omgPos (V : pcm) (v : V) (ht : V -> U) : + last_atval false (ht v) = (omg (ht v) == own). +Proof. +rewrite /omg; case: normalP=>[->//|]; last by case: ifP. +by rewrite lastatval_undef. +Qed. + +Lemma omgPosMorph (V : pcm) (v1 v2 : V) (ht : pcm_morph V U): + valid (v1 \+ v2) -> + preim ht omg_sep v1 v2 -> + last_atval false (ht v1 \+ ht v2) = + (omg (ht v1) \+ omg (ht v2) == own). +Proof. +move=>W /andP [G] /andP []; rewrite /omg /= in G *. +rewrite !undefNV !(validE2 (pfV2 _ _ _ W G)) lastatvalUn pfV2 //=. +by case: ltngtP=>H H1 H2; rewrite ?(negbTE H1) ?(negbTE H2) //; case: ifP. +Qed. + +Lemma omgNeg' (V : pcm) (v : V) (ht : V -> U) : + ~~ last_atval false (ht v) = + (omg (ht v) == nown) || undefb (omg (ht v)). +Proof. by rewrite omgPos; case: (omg _)=>//; case. Qed. + +Lemma omgNeg (V : pcm) (v : V) (ht : V -> U) : + valid (ht v) -> + ~~ last_atval false (ht v) = (omg (ht v) == nown). +Proof. by move=>W; rewrite omgNeg' undefNV valid_omg W; case: (omg _). Qed. + +Lemma omgNegMorph (V : pcm) (v1 v2 : V) (ht : pcm_morph V U) : + valid (v1 \+ v2) -> preim ht omg_sep v1 v2 -> + ~~ last_atval false (ht v1 \+ ht v2) = + (omg (ht v1) \+ omg (ht v2) == nown). +Proof. +move=>W /andP [H1 H2]. +have W' : valid (ht (v1 \+ v2)) by rewrite pfV // (sepU0 W H1). +by rewrite -!pfjoin //= omgNeg. +Qed. + +Lemma omgidPos (v : U) : last_atval false v = (omg v == own). +Proof. by rewrite (omgPos _ id). Qed. + +Lemma omgidPosMorph (v1 v2 : U) : + valid (v1 \+ v2) -> omg_sep v1 v2 -> + last_atval false (v1 \+ v2) = (omg v1 \+ omg v2 == own). +Proof. by move=>W S; rewrite (omgPosMorph (ht:=idfun)). Qed. + +Lemma omgidNeg (v : U) : + valid v -> + ~~ last_atval false v = (omg v == Unit). +Proof. exact: (@omgNeg _ _ id). Qed. + +Lemma omgidNegMorph (v1 v2 : U) : + valid (v1 \+ v2) -> omg_sep v1 v2 -> + ~~ last_atval false (v1 \+ v2) = + (omg v1 \+ omg v2 == Unit). +Proof. by move=>W S; rewrite (omgNegMorph (ht:=idfun)). Qed. + +Definition omgP := + (valid_omg, + (omgidNegMorph, omgidPosMorph, omgPosMorph, omgNegMorph), + (omgidPos, omgidNeg, omgPos, omgNeg)). + +Lemma omg_fresh_val (V : pcm) (v1 v2 : V) (ht : pcm_morph V U) : + valid (v1 \+ v2) -> + sep ht v1 v2 -> + (omg (pts (fresh (ht v1 \+ ht v2)) (false, true) \+ ht v1) = own) * + (omg (pts (fresh (ht v1 \+ ht v2)) (true, false) \+ ht v1) = nown) * + (omg (pts (fresh (ht v1 \+ ht v2)) (false, true) \+ ht v2) = own) * + (omg (pts (fresh (ht v1 \+ ht v2)) (true, false) \+ ht v2) = nown). +Proof. +move=>A O; have Vh : valid (ht v1 \+ ht v2) by rewrite pfV2. +rewrite /omg !undefNV !validPtUn /= !(validE2 Vh) !negbK /=. +by rewrite !lastatval_freshUn // !(negbTE (fresh_dom _)) ?(freshUnL,freshUnR). +Qed. + +Lemma omg_fresh_sep (V : pcm) (v1 v2 : V) (ht : pcm_morph V U) op : + valid (v1 \+ v2) -> + sep ht v1 v2 -> + (omg (ht v2) = Unit -> + omg_sep (pts (fresh (ht v1 \+ ht v2)) op \+ ht v1) (ht v2)) * + (omg (ht v1) = Unit -> + omg_sep (ht v1) (pts (fresh (ht v1 \+ ht v2)) op \+ ht v2)). +Proof. +move=>A O; have Vh : valid (ht v1 \+ ht v2) by rewrite pfV2. +rewrite /omg_sep lastatval_freshUn //. +rewrite !lastkeyPtUn ?(validE2 Vh,freshUnL,freshUnR) //. +by split=>N; rewrite implybT omgP N. +Qed. + +Definition omg_fresh := (omg_fresh_val, omg_fresh_sep). + +(* DEVCOMMENT *) +(* some extra properties the need for which appeared later *) +(* TODO eventually organize into their proper places *) +(* /DEVCOMMENT *) +Lemma omg_eta (h : U): + valid h -> omg h = own -> + exists h' v, [/\ h' = free h (last_key h), + h = pts (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. +move=>V; rewrite /omg /=; rewrite undefNV V /=. +case: ifP=>// N _; set k := last_key h. +have D : k \in dom h. +- rewrite /last_atval /atval /oapp in N. + by case: dom_find N. +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_atval /atval Ef /= in N; rewrite -N. +have: last_key h' <= k. +- by rewrite /k Eh lastkeyPtUnE -Eh V leq_maxr. +rewrite leq_eqVlt; case/orP=>// /eqP E. +by rewrite -E in Nd K; case: lastkeyP K Nd. +Qed. + +(* specialize to alternating histories *) +Lemma omg_eta_all (h : U) : + valid h -> omg h = own -> um_all (fun k v => v.2 = ~~ v.1) h -> + exists h', [/\ h' = free h (last_key h), + h = pts (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_sep x y -> + last_key (x \+ y) = last_key x) * + (omg y = own -> valid (x \+ y) -> omg_sep x y -> + last_key (x \+ y) = last_key y). +Proof. +rewrite /omg_sep /omg !undefNV. +split=>L V S; rewrite (validE2 V) /= in L; 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. + +Arguments omg {U}. +Arguments omg_sep {U}. + diff --git a/pcm/natmap.v b/pcm/natmap.v new file mode 100644 index 0000000..7676532 --- /dev/null +++ b/pcm/natmap.v @@ -0,0 +1,2712 @@ +(* +Copyright 2015 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 maps over non-zero natural *) +(* numbers, pcm instance for natmap, gapless natmaps, natmaps with binary *) +(* range, several sorts of continuous natmaps. *) +(* Histories are a special case of natmaps. *) +(******************************************************************************) + +From HB Require Import structures. +From Coq Require Import ssreflect ssrbool ssrfun. +From mathcomp Require Import ssrnat eqtype seq path interval. +From pcm Require Import axioms options prelude pred finmap rtc seqext. +From pcm Require Export useqord uslice uconsec pcm morphism unionmap. +Import order.Order.NatOrder. (* listed last to avoid notation clash *) + +(************************) +(* Maps over non-0 nats *) +(************************) + +Definition null := 0. +Notation nat_pred := (fun x => x != 0). + +(* natmap is union map of non-0 nat keys *) +HB.mixin Record isNatMap V U of UMC nat (fun x => x != 0) V U. +#[short(type="natmap")] +HB.structure Definition NatMap V := { + U of UMC nat (fun x => x != 0) V U & isNatMap V U}. + +(* pred structure (to write x \In h) is copied from union_map *) +Canonical natmap_PredType A (U : natmap A) : PredType (nat * A) := + Mem_UmMap_PredType U. +Coercion Pred_of_natmap A (U : natmap A) (x : U) : Pred_Class := + [eta Mem_UmMap x]. + +(* canonical natmap is histories *) + +Record history A := Hist {hist_base : @UM.base nat nat_pred A}. + +Section HistoryUMC. +Variables A : Type. +Implicit Type f : history A. +Local Coercion hist_base : history >-> UM.base. +Let ht_valid f := @UM.valid nat nat_pred A f. +Let ht_empty := Hist (@UM.empty nat nat_pred A). +Let ht_undef := Hist (@UM.Undef nat nat_pred A). +Let ht_upd k v f := Hist (@UM.upd nat nat_pred A k v f). +Let ht_dom f := @UM.dom nat nat_pred A f. +Let ht_assocs f := @UM.assocs nat nat_pred A f. +Let ht_free f k := Hist (@UM.free nat nat_pred A f k). +Let ht_find k f := @UM.find nat nat_pred A k f. +Let ht_union f1 f2 := Hist (@UM.union nat nat_pred A f1 f2). +Let ht_empb f := @UM.empb nat nat_pred A f. +Let ht_undefb f := @UM.undefb nat nat_pred A f. +Let ht_from (f : history A) : UM.base _ _ := f. +Let ht_to (b : @UM.base nat nat_pred A) : history A := Hist b. +Let ht_pts k v := Hist (@UM.pts nat nat_pred A k v). + +Lemma history_is_umc : + union_map_axiom ht_valid ht_empty ht_undef ht_upd ht_dom + ht_assocs ht_free ht_find ht_union ht_empb + ht_undefb ht_pts ht_from ht_to. +Proof. by split=>//; split=>[|[]]. Qed. + +HB.instance Definition _ := + isUnion_map.Build nat nat_pred A (history A) history_is_umc. +End HistoryUMC. + +HB.instance Definition _ A := isNatMap.Build A (history A). +HB.instance Definition _ (A : eqType) := + hasDecEq.Build (history A) (@union_map_eqP nat _ A (history A)). +Canonical history_PredType A : PredType (nat * A) := + Mem_UmMap_PredType (history A). +Coercion Pred_of_history A (x : history A) : Pred_Class := + [eta Mem_UmMap x]. + +Notation "x \-> v" := (ptsT (history _) x v) (at level 30). + +(* DEVCOMMENT *) +(* tests *) +Lemma xx : 1 \-> true = null \-> false. +Abort. + +Lemma xx : ((1 \-> true) \+ (2 \-> false)) == (1 \-> false). +rewrite joinC. +Abort. + +Lemma xx (x : history nat) : x \+ x == x \+ x. +Abort. + +Lemma xx : 1 \-> (1 \-> 3) = 2 \-> (7 \-> 3). +Abort. + +Lemma xx : (1, 3) \In (1 \-> 3). +Abort. +(* /DEVCOMMENT *) + +(*************) +(* Main defs *) +(*************) + +Definition last_key {A} {U : natmap A} (h : U) := last 0 (dom h). +Definition fresh {A} {U : natmap A} (h : U) := (last_key h).+1. +Definition last_val {A} {U : natmap A} (h : U) := find (last_key h) h. + +(**********************) +(* generic properties *) +(**********************) + +(* alternative to In_cond *) +Lemma In_neq A (U : natmap A) (h : U) k : + k \In h -> + k.1 > 0. +Proof. by move/In_cond; rewrite lt0n. Qed. + +Lemma mem_dom0 A (U : natmap A) {h : U} : 0 \in dom h = false. +Proof. by rewrite cond_dom. Qed. + +Lemma uniq_dom0 A (U : natmap A) {h : U} : uniq (0 :: dom h). +Proof. by rewrite /= uniq_dom cond_dom. Qed. + +Lemma all_leq0 A (U : natmap A) {h : U} : all (leq 0) (dom h). +Proof. by apply/allP=>x /dom_cond; case: ltngtP. Qed. + +Lemma all_ltn0 A (U : natmap A) {h : U} : all (ltn 0) (dom h). +Proof. by apply/allP=>x /dom_cond; case: ltngtP. Qed. + +#[export] Hint Resolve uniq_dom0 all_leq0 all_ltn0 : core. + +Lemma In_dom0 {A} {U : natmap A} (h : U) k e : + (k, e) \In h -> k \in 0 :: dom h. +Proof. by move=>H; rewrite inE (In_dom H) orbT. Qed. + +Lemma sorted_leq_dom A {U : natmap A} {h : U} : sorted leq (dom h). +Proof. by rewrite -(eq_sorted nat_oleq) sorted_dom_oleq. Qed. + +Lemma sorted_ltn_dom A {U : natmap A} {h : U} : sorted ltn (dom h). +Proof. by rewrite sorted_dom. Qed. + +Lemma path_leq_dom A {U : natmap A} {h : U} : path leq 0 (dom h). +Proof. by rewrite path_min_sorted // sorted_leq_dom. Qed. + +Lemma path_ltn_dom A {U : natmap A} {h : U} : path ltn 0 (dom h). +Proof. by rewrite path_min_sorted. Qed. + +#[export] Hint Resolve sorted_leq_dom sorted_ltn_dom + path_leq_dom path_ltn_dom : core. + +(* form of totality of key order *) +Lemma umfiltT A (U : natmap A) k1 k2 (h : U) : + k1 \in dom h -> k1 \notin dom (um_filterk (ltn^~ k2) h) -> + k2 \in dom h -> k2 \notin dom (um_filterk (ltn^~ k1) h) -> + k1 = k2. +Proof. +case/In_domX=>v1 D1 S1 /In_domX [v2 D2] S2. +case: (ltngtP k1 k2)=>// N. +- by move/(In_umfiltk (ltn^~ k2))/(_ N)/In_dom: D1; rewrite (negbTE S1). +by move/(In_umfiltk (ltn^~ k1))/(_ N)/In_dom: D2; rewrite (negbTE S2). +Qed. + +Lemma umfiltT0 A (U : natmap A) k1 k2 (h : U) : + k1 \in dom h -> um_filterk (ltn^~ k1) h = Unit -> + k2 \in dom h -> um_filterk (ltn^~ k2) h = Unit -> + k1 = k2. +Proof. +by move=>D1 U1 D2 U2; apply: umfiltT D1 _ D2 _; rewrite ?U1 ?U2 dom0. +Qed. + +Lemma omf_subdom0 {A1 A2} {U1 : natmap A1} {U2 : natmap A2} + {f : omap_fun U1 U2} {h : U1} : + {subset 0::dom (f h) <= 0::dom h}. +Proof. +move=>x; rewrite !inE; case/orP=>[|/omf_subdom] -> //. +by rewrite orbT. +Qed. + +(**************) +(* Base cases *) +(**************) + +Lemma lastkey_undef A (U : natmap A) : last_key (undef : U) = 0. +Proof. by rewrite /last_key dom_undef. Qed. + +Lemma lastkey0 A (U : natmap A) : last_key (Unit : U) = 0. +Proof. by rewrite /last_key dom0. Qed. + +Lemma lastkeyPt A (U : natmap A) (x : nat) (v : A) : + last_key (pts x v : U) = x. +Proof. by rewrite /last_key domPtK; case: (x =P 0). Qed. + +Lemma lastval0 A (U : natmap A) : last_val (Unit : U) = None. +Proof. by rewrite /last_val lastkey0 find0E. Qed. + +Lemma lastval_undef A (U : natmap A) : last_val (undef : U) = None. +Proof. by rewrite /last_val lastkey_undef find_undef. Qed. + +Lemma lastvalPt A (U : natmap A) x v : + last_val (pts x v : U) = if x != 0 then Some v else None. +Proof. by rewrite /last_val lastkeyPt findPt. Qed. + +(* basic transfer lemma *) +Lemma In_last A (U : natmap A) (h : U) v : + (last_key h, v) \In h <-> last_val h = Some v. +Proof. exact: In_find. Qed. + +(* for inlined rewriting *) +Lemma In_lastE A (U : natmap A) (h : U) v : + (last_key h, v) \In h -> last_val h = Some v. +Proof. exact: In_findE. Qed. + +(* fresh and 0 *) + +Lemma fresh_undef A (U : natmap A) : fresh (undef : U) = 1. +Proof. by rewrite /fresh lastkey_undef. Qed. + +Lemma fresh0 A (U : natmap A) : fresh (Unit : U) = 1. +Proof. by rewrite /fresh lastkey0. Qed. + +Lemma fresh_lt0 A (U : natmap A) (h : U) x : fresh h <= x -> 0 < x. +Proof. by case: x. Qed. + +Lemma freshN0 A (U : natmap A) (h : U) x : fresh h <= x -> x != 0. +Proof. by case: x. Qed. + +(*************) +(* Main view *) +(*************) + +(* captures that following properties are equivalent: *) +(* (undefb && unitb) = (last_key != 0) = (last_key > 0) = ... *) +(* ... = (last_key \in dom) = (dom <> [::]) = lastval = Some *) + +Variant lastkey_dom_spec A (U : natmap A) (h : U) : + bool -> bool -> bool -> nat -> seq nat -> + bool -> bool -> bool -> option A -> Type := +| lastkey_dom_undef of h = undef : + lastkey_dom_spec h false true false 0 [::] true false false None +| lastkey_dom_unit of h = Unit : + lastkey_dom_spec h true false true 0 [::] true false false None +| lastkey_dom_valid v of (last_key h, v) \In h : + lastkey_dom_spec h true false false (last_key h) + (dom h) false true true (Some v). + +Lemma lastkeyP A (U : natmap A) (h : U) : + lastkey_dom_spec h (valid h) (undefb h) (unitb h) + (last_key h) (dom h) (last_key h == 0) + (last_key h > 0) (last_key h \in dom h) + (last_val h). +Proof. +have L (x : U) : valid x -> last_key x \notin dom x -> x = Unit. +- rewrite /last_key !umEX /UM.valid/UM.dom/UM.empty -{4}[x]tfE. + case: (UMC_from x)=>//=; case=>s H H1 _ /seq_last_in. + by rewrite eqE UM.umapE /supp fmapE /= {H H1}; elim: s. +case: (normalP0 h)=>[->|->|]. +- by rewrite lastkey_undef dom_undef lastval_undef; apply: lastkey_dom_undef. +- by rewrite lastkey0 dom0 lastval0; apply: lastkey_dom_unit. +move=>V H. +have H1 : last_key h \in dom h by apply: contraT=>/(L _ V)/H. +have H2 : last_key h != 0 by apply: dom_cond H1. +have [v H4] : {v | last_val h = Some v}. +- by rewrite /last_val; case: dom_find H1=>// v; exists v. +by rewrite H1 lt0n (negbTE H2) H4; apply/lastkey_dom_valid/In_last. +Qed. + +(* main membership invariant of last_key *) +Lemma lastkey_mem0 A (U : natmap A) (h : U) : last_key h \in 0 :: dom h. +Proof. by rewrite inE; case: lastkeyP. Qed. + +(* DEVCOMMENT: *) +(* Weird variation. Worth preserving? *) +(* /DEVCOMMENT *) +Lemma lastkey_memk A (U : natmap A) (h : U) k : + (k < last_key h -> last_key h \notin dom h) -> + last_key h <= k. +Proof. by case: lastkeyP=>//= v _; case: ltngtP=>// _; apply. Qed. + +(************) +(* last_key *) +(************) + +Section LastkeyLemmas. +Variables (A : Type) (U : natmap A). +Implicit Type h : U. + +(* last_key is upper bound *) +Lemma dom_lastkey h k : k \in dom h -> k <= last_key h. +Proof. +rewrite /last_key !umEX /UM.dom; case: (UMC_from h)=>//; case=>s H _ /=. +rewrite /supp/ord /= (leq_eqVlt k) orbC. +by apply: sorted_last_key_maxR otrans (sorted_oleq H). +Qed. + +(* useful variants with 0 :: dom and contrapositives *) +Lemma dom0_lastkey h k : k \in 0 :: dom h -> k <= last_key h. +Proof. by rewrite inE; case/orP=>[/eqP -> //|]; apply: dom_lastkey. Qed. +Lemma lastkey_dom h k : last_key h < k -> k \notin dom h. +Proof. by apply: contraL; rewrite -leqNgt; apply: dom_lastkey. Qed. +Lemma lastkey_dom0 h k : last_key h < k -> k \notin 0 :: dom h. +Proof. by apply: contraL; rewrite -leqNgt; apply: dom0_lastkey. Qed. + +(* DEVCOMMENT: *) +(* Weird variation. Worth preserving? *) +(* /DEVCOMMENT *) +Lemma dom_lastkey_eq h k : + k \in dom h -> + (k < last_key h -> last_key h \notin dom h) -> + k = last_key h. +Proof. +move/dom_lastkey=>D /lastkey_memk H. +by rewrite (@anti_leq (last_key _) k) // H D. +Qed. + +(* last_key is least upper bound *) +Lemma lastkey_lub h k : {in dom h, forall x, x <= k} -> last_key h <= k. +Proof. by case lastkeyP=>// v /In_dom H; apply. Qed. + +(* useful variant with equality *) +Lemma lub_lastkey h k : + k \in dom h -> + (forall x, k < x -> x \notin dom h) -> + k = last_key h. +Proof. +case: lastkeyP=>// v H /dom_lastkey Dx X. +rewrite (@anti_leq (last_key _) k) // Dx andbT. +apply: lastkey_lub=>x; apply: contraTT. +by rewrite -ltnNge; apply: X. +Qed. + +(* equivalences *) + +Lemma lastkey_leq h k : (last_key h <= k) = all (fun x => x <= k) (dom h). +Proof. +apply/idP/allP=>[H x /dom_lastkey D|]; +by [apply: leq_trans D H | apply: lastkey_lub]. +Qed. + +Lemma leq_lastkey h k : (k <= last_key h) = has (fun x => k <= x) (0 :: dom h). +Proof. +apply/idP/hasP=>[|[x] D H]; last by apply: leq_trans H (dom0_lastkey D). +by exists (last_key h)=>//; apply: lastkey_mem0. +Qed. + +Lemma ltn_lastkey h k : (k < last_key h) = has (fun x => k < x) (dom h). +Proof. +apply/idP/hasP=>[|[x] D H]; last by apply: leq_trans H (dom_lastkey D). +by exists (last_key h)=>//; case: lastkeyP H. +Qed. + +Lemma lastkey_ltn h k : (last_key h < k) = all (fun x => x < k) (0 :: dom h). +Proof. +apply/idP/allP=>[H x D|]; last by apply; apply: lastkey_mem0. +by apply: leq_ltn_trans (dom0_lastkey D) H. +Qed. + +(* unfolding equivalences into implications (forward) *) +Lemma lastkey_leq_dom h x k : last_key h <= k -> x \in dom h -> x <= k. +Proof. by rewrite lastkey_leq=>/allP; apply. Qed. +Lemma lastkey_leq_dom0 h x k : last_key h <= k -> x \in 0 :: dom h -> x <= k. +Proof. by move=>H; case/orP=>[/eqP ->//|]; apply: lastkey_leq_dom H. Qed. +Lemma lastkey_ltn_dom0 h x k : last_key h < k -> x \in 0 :: dom h -> x < k. +Proof. by rewrite lastkey_ltn=>/allP; apply. Qed. +Lemma lastkey_ltn_dom h x k : last_key h < k -> x \in dom h -> x < k. +Proof. by move=>H D; apply: lastkey_ltn_dom0 H _; rewrite inE D orbT. Qed. + +(* unfolding equivalences into implications (backward) *) +Lemma dom0_leq_lastkey h x k : x \in 0 :: dom h -> k <= x -> k <= last_key h. +Proof. by move=>D H; rewrite leq_lastkey; apply/hasP; exists x. Qed. +Lemma dom_leq_lastkey h x k : x \in dom h -> k <= x -> k <= last_key h. +Proof. by move=>D; apply: dom0_leq_lastkey; rewrite inE D orbT. Qed. +Lemma dom_ltn_lastkey h x k : x \in dom h -> k < x -> k < last_key h. +Proof. by move=>D H; rewrite ltn_lastkey; apply/hasP; exists x. Qed. +Lemma dom0_ltn_lastkey h x k : x \in 0 :: dom h -> k < x -> k < last_key h. +Proof. by case/orP=>[/eqP ->//|]; apply: dom_ltn_lastkey. Qed. +End LastkeyLemmas. + +(* interaction of two lastkeys *) + +Lemma lastkey_mono A1 A2 (U1 : natmap A1) (U2 : natmap A2) (h1 : U1) (h2 : U2) : + {subset dom h1 <= dom h2} -> + last_key h1 <= last_key h2. +Proof. +rewrite leq_eqVlt orbC; apply: (seq_last_monoR orefl otrans); +by rewrite (eq_path nat_oleq). +Qed. + +Lemma lastkey_subdom A1 A2 (U1 : natmap A1) (U2 : natmap A2) (h1 : U1) (h2 : U2) : + {subset dom h1 <= dom h2} -> + last_key h2 \in dom h1 -> + last_key h1 = last_key h2. +Proof. +by move=>S D; apply/eqP; rewrite eqn_leq (dom_lastkey D) (lastkey_mono S). +Qed. + +Lemma lastkeyUnL A (U : natmap A) (h1 h2 : U) : + valid (h1 \+ h2) -> + last_key h1 <= last_key (h1 \+ h2). +Proof. by move=>V; apply: lastkey_mono=>x; rewrite domUn inE V => ->. Qed. + +Lemma lastkeyUnLT A (U : natmap A) (h1 h2 : U) k : + valid (h1 \+ h2) -> + last_key (h1 \+ h2) < k -> + last_key h1 < k. +Proof. by move=>V; apply/leq_ltn_trans/lastkeyUnL/V. Qed. + +Lemma lastkeyUnR A (U : natmap A) (h1 h2 : U) : + valid (h1 \+ h2) -> + last_key h2 <= last_key (h1 \+ h2). +Proof. by rewrite joinC; apply: lastkeyUnL. Qed. + +Lemma lastkeyUnRT A (U : natmap A) (h1 h2 : U) k : + valid (h1 \+ h2) -> + last_key (h1 \+ h2) < k -> + last_key h2 < k. +Proof. by rewrite joinC; apply: lastkeyUnLT. Qed. + +(* disjoint maps with equal last keys are both empty *) +Lemma lastkeyUn0T A1 A2 (U1 : natmap A1) (U2 : natmap A2) (h1 : U1) (h2 : U2) : + valid h1 -> valid h2 -> + (forall x, x \in dom h1 -> x \in dom h2 -> false) -> + last_key h1 = last_key h2 -> + (h1 = Unit) * (h2 = Unit). +Proof. +case: (lastkeyP h1)=>[//|-> _ V _ /esym/eqP|]; first by case: (lastkeyP h2) V. +move=>v1 H1 _ V2 /(_ _ (In_dom H1)) /[swap]/eqP /=. +case: (lastkeyP h2) V2=>//; first by rewrite (negbTE (In_cond H1)). +by move=>v2 H2 _ /eqP -> /(_ (In_dom H2)). +Qed. + +(* specialization to equally-typed maps *) +Lemma lastkeyUn0 A (U : natmap A) (h1 h2 : U) : + valid (h1 \+ h2) -> + last_key h1 = last_key h2 -> + (h1 = Unit) * (h2 = Unit). +Proof. by case: validUn=>// ?? D _; apply: lastkeyUn0T=>// x /D/negbTE ->. Qed. + +(* variation to avoid validity *) +Lemma lastkey00 A1 A2 (U1 : natmap A1) (U2 : natmap A2) (h1 : U1) (h2 : U2) : + (forall x, x \in dom h1 -> x \in dom h2 -> false) -> + last_key h1 = last_key h2 -> + (last_key h1 = 0) * (last_key h2 = 0). +Proof. +case: (normalP h1)=>[->|V1] D; first by rewrite lastkey_undef. +case: (normalP h2) D=>[->|V2] D; first by rewrite lastkey_undef. +by move=>H; rewrite !(lastkeyUn0T V1 V2 D H) !lastkey0. +Qed. + +(* variation with subdoms *) +Lemma lastkey_subdom00 A1 A2 (U1 : natmap A1) (U2 : natmap A2) + (h1 : U1) (h2 : U2) (s1 s2 : pred nat) : + (forall x, x \in s1 -> x \in s2 -> false) -> + {subset dom h1 <= s1} -> + {subset dom h2 <= s2} -> + last_key h1 = last_key h2 -> + (last_key h1 = 0) * (last_key h2 = 0). +Proof. by move=>D S1 S2; apply: lastkey00=>x /S1 H /S2; apply: D H. Qed. + +(* interaction with constructors *) + +Section LastkeyConstructors. +Variables (A : Type) (U : natmap A). +Implicit Type h : U. + +Lemma lastkey_find h k : last_key h < k -> find k h = None. +Proof. by move/lastkey_dom; case: dom_find. Qed. + +(* monotonicity gives <=; removing last_key gives strict < *) +Lemma lastkeyF h : + last_key h != 0 -> + last_key (free h (last_key h)) < last_key h. +Proof. +move=>D; have {}D : last_key h \in dom h by case: lastkeyP D. +case: (um_eta D)=>v [Ef Eh]. +have N : last_key h \notin dom (free h (last_key h)). +- by rewrite domF inE eqxx. +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 : last_key h > 0 by move/dom_cond: D; case: (last_key h). +by case: lastkeyP N. +Qed. + +(* lastkey of union is max of lastkeys *) +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 : U) : valid (k1 \+ k2) -> + last_key k1 < last_key k2 -> last_key (k1 \+ k2) = last_key k2. +- move=>V H; apply/eqP; rewrite eqn_leq lastkeyUnR // andbT lastkey_leq. + apply/allP=>/= x; rewrite domUn inE V /=. + by case/orP=>/dom_lastkey // T; apply: leq_trans T (ltnW H). +case: (normalP (h1 \+ h2))=>[->|V]; first by rewrite lastkey_undef. +case: (ltngtP (last_key h2) (last_key h1)). +- by rewrite joinC in V *; apply: H. +- by apply: H. +by move/esym/(lastkeyUn0 V)=>E; rewrite !E unitL. +Qed. + +Lemma lastkeyPtUnV h k v : last_key h < k -> valid (pts k v \+ h) = valid h. +Proof. +move=>N; rewrite validPtUn -lt0n (leq_ltn_trans _ N) //= andbC. +by case D: (k \in dom h)=>//; move/dom_lastkey: D; case: leqP N. +Qed. + +Lemma lastkey_domPtUn h k v : + valid h -> last_key h < k -> dom (pts k v \+ h) = rcons (dom h) k. +Proof. +move=>V N; rewrite joinC domUnPtK //. +- by rewrite joinC lastkeyPtUnV. +by apply/allP=>x /dom_lastkey D; apply: leq_ltn_trans D N. +Qed. + +Lemma lastkey_rangePtUn h k v : + valid h -> last_key h < k -> range (pts k v \+ h) = rcons (range h) v. +Proof. +move=>V K; rewrite joinC rangeUnPtK //; last first. +- by apply/allP=>x /dom_lastkey H; apply: leq_ltn_trans H K. +by rewrite joinC lastkeyPtUnV. +Qed. + +Lemma lastkeyPtUnE h k v : + last_key (pts k v \+ h) = + if valid (pts k v \+ h) then maxn k (last_key h) else 0. +Proof. by rewrite lastkeyUn lastkeyPt. Qed. + +Lemma lastkeyPtUn h k v : + valid h -> last_key h < k -> last_key (pts k v \+ h) = k. +Proof. +move=>V N; rewrite lastkeyPtUnE (lastkeyPtUnV _ N) V. +by apply: maxn_idPl (ltnW N). +Qed. + +Lemma lastkeyPtUnEX v2 v1 h k : + last_key (pts k v1 \+ h) = last_key (pts k v2 \+ h). +Proof. by rewrite !lastkeyPtUnE !validPtUn. Qed. + +Lemma lastkeyUnPtEX v2 v1 h k : + last_key (h \+ pts k v1) = last_key (h \+ pts k v2). +Proof. by rewrite !(joinC h) (lastkeyPtUnEX v2). Qed. + +Lemma lastkeyU h k v : k \in dom h -> last_key (upd k v h) = last_key h. +Proof. by case/um_eta=>u [_ ->]; rewrite updPtUn (lastkeyPtUnEX u). Qed. + +Lemma lastkeyUE h k v : + last_key (upd k v h) = + if k \in dom h then last_key h else + if valid (upd k v h) then maxn k (last_key h) else 0. +Proof. +case: ifP=>[|H]; first by apply: lastkeyU. +rewrite validU; case: eqVneq=>[->|N] /=. +- by rewrite upd_condN // lastkey_undef. +case: (normalP h)=>[->|V]; first by rewrite upd_undef lastkey_undef. +by rewrite upd_eta2 ?H // lastkeyUn -upd_eta2 ?H // validU N V lastkeyPt. +Qed. + +(* backward induction on valid natmaps *) +Lemma valid_indb (P : U -> Prop) : + P Unit -> + (forall v, P (pts 1 v)) -> + (forall k v h, + P h -> last_key h < k -> + valid (pts k v \+ h) -> P (pts k v \+ h)) -> + forall h, valid h -> P h. +Proof. +move=>H1 H2 H3; elim/um_indb=>[||k v f H V K _]; rewrite ?valid_undef //. +rewrite joinC in V; move: (validR V)=>Vf; case E: (unitb f); last first. +- by apply: H3 (H Vf) (K _ _) V; case: lastkeyP Vf E. +move/unitbP: {K H} E V (H3 k v Unit H1)=>->. +by rewrite unitR validPt lastkey0 lt0n=>V; apply. +Qed. + +(* forward induction on valid natmaps *) +Lemma valid_indf (P : U -> Prop) : + P Unit -> + (forall k v h, P h -> + {in dom h, forall x, k < x} -> + valid (pts k v \+ h) -> P (pts k v \+ h)) -> + forall h, valid h -> P h. +Proof. +move=>H1 H2; elim/um_indf=>[||k v f H V K _]; rewrite ?valid_undef //. +by apply: H2 (H (validR V)) _ V; apply/allP/(order_path_min _ K). +Qed. + +(* membership *) + +Lemma In_lastkey h k v : (k, v) \In h -> k <= last_key h. +Proof. by move/In_dom/dom_lastkey. Qed. + +Lemma In_lastkeyPtUn h x k v w : + last_key h < k -> (x, w) \In h -> (x, w) \In pts k v \+ h. +Proof. by move=>N H; apply: InR=>//; rewrite lastkeyPtUnV ?(In_valid H). Qed. + +Lemma In_lastkeyPtUn_inv h x k v w : + x <= last_key h -> last_key h < k -> + (x, w) \In pts k v \+ h -> (x, w) \In h. +Proof. +move=>N1 N2 H; case/(InPtUnE _ (In_valid H)): H N2=>//. +by case=><- _; case: ltngtP N1. +Qed. + +End LastkeyConstructors. + +(* last_key and omap_fun -- compositions with omf_subdom/omf_subdom0 *) + +Section LastKeyOmapFun. +Variables (A1 A2 : Type) (U1 : natmap A1) (U2 : natmap A2). +Implicit Type f : omap_fun U1 U2. + +(* upper bound lemmas *) +Lemma odom_lastkey f h k : k \in dom (f h) -> k <= last_key h. +Proof. by move/omf_subdom/dom_lastkey. Qed. +Lemma odom0_lastkey f h k : k \in 0 :: dom (f h) -> k <= last_key h. +Proof. by move/omf_subdom0/dom0_lastkey. Qed. +Lemma lastkey_odom f h k : last_key h < k -> k \notin dom (f h). +Proof. by move/lastkey_dom; apply/(subsetC omf_subdom). Qed. +Lemma lastkey_odom0 f h k : last_key h < k -> k \notin 0 :: dom (f h). +Proof. move/lastkey_dom0; apply/(subsetC omf_subdom0). Qed. + +(* DEVCOMMENT: *) +(* Weird variation. Worth preserving? *) +(* composes dom_lastkey_eq with In_dom_omfX *) +(* /DEVCOMMENT *) +Lemma odom_lastkey_eq f h k v : + (k, v) \In h -> omf f (k, v) -> + (forall w, k < last_key (f h) -> + (last_key (f h), w) \In h -> + omf f (last_key (f h), w) -> False) -> + k = last_key (f h). +Proof. +move=>D1 D2 H; apply: dom_lastkey_eq=>[|X]. +- by apply/In_dom_omfX; exists v. +by apply/In_dom_omfX; case=>w []; apply: H X. +Qed. + +(* lub lemmas: combine lastkey_lub/lub_lastkey with In_dom_omfX *) +Lemma lastkey_olub f h k : + (forall x v, (x, v) \In h -> omf f (x, v) -> x <= k) -> + last_key (f h) <= k. +Proof. +move=>H; apply: lastkey_lub=>x /In_dom_omfX [v][X E]. +by apply: H X _; rewrite E. +Qed. + +Lemma olub_lastkey f h k : + k \in dom (f h) -> + (forall x v, k < x -> (x, v) \In h -> omf f (x, v) -> False) -> + k = last_key (f h). +Proof. +move=>D H; apply: lub_lastkey=>// x K; apply/In_dom_omfX; case=>y [X E]. +by apply: H K X _; rewrite E. +Qed. + +(* equivalence lemmas (forward) *) +Lemma lastkey_leq_odom f h x k : last_key h <= k -> x \in dom (f h) -> x <= k. +Proof. by move/lastkey_leq_dom=>H /omf_subdom /H. Qed. +Lemma lastkey_leq_odom0 f h x k : last_key h <= k -> x \in 0 :: dom (f h) -> x <= k. +Proof. by move/lastkey_leq_dom0=>H /omf_subdom0 /H. Qed. +Lemma lastkey_ltn_odom0 f h x k : last_key h < k -> x \in 0 :: dom (f h) -> x < k. +Proof. by move/lastkey_ltn_dom0=>H /omf_subdom0 /H. Qed. +Lemma lastkey_ltn_odom f h x k : last_key h < k -> x \in dom (f h) -> x < k. +Proof. by move/lastkey_ltn_dom=>H /omf_subdom /H. Qed. + +(* equivalence lemms (backward) *) +Lemma odom0_leq_lastkey f h x k : x \in 0 :: dom (f h) -> k <= x -> k <= last_key h. +Proof. by move/omf_subdom0; apply/dom0_leq_lastkey. Qed. +Lemma odom_leq_lastkey f h x k : x \in dom (f h) -> k <= x -> k <= last_key h. +Proof. by move/omf_subdom; apply/dom_leq_lastkey. Qed. +Lemma odom_ltn_lastkey f h x k : x \in dom (f h) -> k < x -> k < last_key h. +Proof. by move/omf_subdom; apply: dom_ltn_lastkey. Qed. +Lemma odom0_ltn_lastkey f h x k : x \in 0 :: dom (f h) -> k < x -> k < last_key h. +Proof. by move/omf_subdom0; apply: dom0_ltn_lastkey. Qed. + +(* other lemmas *) + +Lemma lastkey_omf f h : last_key (f h) <= last_key h. +Proof. by apply: lastkey_mono; apply: omf_subdom. Qed. + +Lemma lastkey_omfT f h k : last_key h <= k -> last_key (f h) <= k. +Proof. by apply/leq_trans/lastkey_omf. Qed. + +Lemma lastkey_omfT' f h k : last_key h < k -> last_key (f h) < k. +Proof. by apply/leq_ltn_trans/lastkey_omf. Qed. + +Lemma lastkey_omfUn f h1 h2 : + valid (h1 \+ h2) -> last_key (f h1 \+ f h2) <= last_key (h1 \+ h2). +Proof. by move=>V; rewrite -omfUn // lastkey_omf. Qed. + +Lemma lastkey_omf_some f h : + (forall x, x \In h -> oapp predT false (omf f x)) -> + last_key (f h) = last_key h. +Proof. by rewrite /last_key=>/dom_omf_some ->. Qed. + +Lemma lastkey_omfE f h : last_key h \in dom (f h) -> last_key (f h) = last_key h. +Proof. by apply: lastkey_subdom omf_subdom. Qed. + +Lemma lastkey_omfPtUn f k v h : + last_key h < k -> + f (pts k v \+ h) = + if omf f (k, v) is Some w then pts k w \+ f h else f h. +Proof. +case: (normalP h)=>[->|V N]; last by rewrite omfPtUn lastkeyPtUnV // V. +by case: (omf _ _)=>[a|]; rewrite pfundef !join_undef pfundef. +Qed. + +Lemma omf_sublastkey f k v h : + last_key h < k -> + {subset dom (f h) <= dom (f (pts k v \+ h))}. +Proof. +move=>N x H; rewrite lastkey_omfPtUn //; case: (omf _ _)=>[a|//]. +by rewrite domPtUn inE H orbT lastkeyPtUnV ?(dom_valid H,lastkey_omfT'). +Qed. + +End LastKeyOmapFun. + +(* last_key and filter *) + +Lemma lastkey_umfilt_le A (U : natmap A) (h : U) k : + last_key h <= k -> um_filterk [pred x | x <= k] h = h. +Proof. +move=>N; rewrite -[RHS]umfilt_predT -eq_in_umfilt. +by case=>x v /In_dom/dom_lastkey/leq_trans/(_ N). +Qed. + +Lemma lastkey_umfilt_lt A (U : natmap A) (h : U) k : + last_key h < k -> um_filterk [pred x | x < k] h = h. +Proof. +move=>N; rewrite -[RHS]umfilt_predT -eq_in_umfilt. +by case=>x v /In_dom/dom_lastkey/leq_ltn_trans/(_ N). +Qed. + +Lemma lastkey_umfiltPtUn A (U : natmap A) (p : pred (nat * A)) (h : U) k v : + last_key h < k -> + um_filter p (pts k v \+ h) = + if p (k, v) then pts k v \+ um_filter p h else um_filter p h. +Proof. by move=>V; rewrite lastkey_omfPtUn ?omf_omap //=; case: ifP. Qed. + +(* last_key and side *) + +Lemma lastkey_sidePtUn (T : eqType) (Us : T -> Type) + (U : natmap (sigT Us)) (Ut : forall t, natmap (Us t)) + (h : U) t k v : + fresh h <= k -> + side_map Ut t (pts k v \+ h) = + if decP (t =P tag v) is left pf then + pts k (cast Us pf (tagged v)) \+ side_map Ut t h + else side_map Ut t h. +Proof. by case: v=>tx vx N; rewrite lastkey_omfPtUn //= /omfx/=; case: eqP. Qed. + +Lemma lastkey_dom_sidePtUn (T : eqType) (Us : T -> Type) + (U : natmap (sigT Us)) (Ut : forall t, natmap (Us t)) + (h : U) t k v : + last_key h < k -> + dom (side_map Ut t (pts k v \+ h)) = + if valid h then + if t == tag v then rcons (dom (side_map Ut t h)) k + else dom (side_map Ut t h) + else [::]. +Proof. +case: (normalP h)=>[->|V N]; first by rewrite join_undef pfundef dom_undef. +rewrite lastkey_sidePtUn //; case: eqP=>//= ?; subst t. +by rewrite eqc lastkey_domPtUn ?(pfV (side_map Ut _)) ?lastkey_omfT'. +Qed. + +(* last_key and non-omap morphisms *) +Lemma plastkey (A : Type) (U1 : pcm) (U2 : natmap A) + (f : pcm_morph U1 U2) k (x y : U1) : + valid (x \+ y) -> + sep f x y -> + last_key (f (x \+ y)) < k -> + last_key (f x) < k. +Proof. +move=>V D; apply: leq_ltn_trans. +by rewrite pfjoin // lastkeyUnL // pfV2. +Qed. + +Lemma plastkeyT A (U1 : pcm) (U2 : natmap A) + (f : full_pcm_morph U1 U2) k (x y : U1) : + valid (x \+ y) -> + last_key (f (x \+ y)) < k -> + last_key (f x) < k. +Proof. by move=>V; apply/(plastkey V)/pfT. Qed. + +(* monotonicity for natmap nat *) +Lemma lastkey_monoPtUn (U : natmap nat) k w (h : U) : + last_key h < k -> + (forall k v, (k, v) \In h -> v < w) -> + um_mono (pts k w \+ h) = um_mono h. +Proof. +move=>N H; case: (normalP h)=>[->|V]; first by rewrite join_undef. +apply: ummonoPtUn; first by rewrite lastkeyPtUnV. +move=>x v X; split; last by apply: H X. +by apply: leq_ltn_trans (dom_lastkey (In_dom X)) N. +Qed. + +(* last_key and pcm ordering *) +Lemma lastkey_fun A (U1 : pcm) (U2 : natmap A) + (f : full_pcm_morph U1 U2) (x y : U1) : + [pcm y <= x] -> + valid x -> + last_key (f y) <= last_key (f x). +Proof. by case=>z -> V; rewrite pfjoinT // lastkeyUnL // pfV2I. Qed. + +(* special case of lastkey_fun when f = id *) +Lemma lastkey_umpleq A (U : natmap A) (h1 h2 : U) : + [pcm h1 <= h2] -> + valid h2 -> + last_key h1 <= last_key h2. +Proof. by apply: (lastkey_fun idfun). Qed. + +Lemma lastkeyL A (U1 : pcm) (U2 : natmap A) + (f : full_pcm_morph U1 U2) (x y : U1) : + valid (x \+ y) -> + last_key (f x) <= last_key (f (x \+ y)). +Proof. by apply: lastkey_fun. Qed. + +Lemma lastkeyR A (U1 : pcm) (U2 : natmap A) + (f : full_pcm_morph U1 U2) (x y : U1) : + valid (x \+ y) -> + last_key (f y) <= last_key (f (x \+ y)). +Proof. by apply: lastkey_fun. Qed. + +(*********) +(* fresh *) +(*********) + +(* mostly new names for lastkey lemmas *) +(* but there are several new lemmas also *) +(* to be systematic, we list all *) + +Section FreshLemmas. +Variables (A : Type) (U : natmap A). +Implicit Type h : U. + +(* fresh is strict upper bound *) +Lemma dom_fresh h k : k \in dom h -> k < fresh h. +Proof. exact: dom_lastkey. Qed. + +(* variants with 0 :: dom and contrapositives *) +Lemma dom0_fresh h k : k \in 0 :: dom h -> k < fresh h. +Proof. exact: dom0_lastkey. Qed. +Lemma fresh_dom h k : fresh h <= k -> k \notin dom h. +Proof. exact: lastkey_dom. Qed. +Lemma fresh_dom0 h k : fresh h <= k -> k \notin 0 :: dom h. +Proof. exact: lastkey_dom0. Qed. + +(* equivalences (similar to lastkey, but with x.+1) *) +Lemma fresh_leq h k : (fresh h <= k) = all (fun x => x.+1 <= k) (0 :: dom h). +Proof. exact: lastkey_ltn. Qed. + +Lemma leq_fresh h k : (k <= fresh h) = has (fun x => k <= x.+1) (0 :: dom h). +Proof. +apply/idP/hasP=>[|[x] D H]; last by apply: leq_trans H (dom0_fresh D). +by exists (last_key h)=>//; apply: lastkey_mem0. +Qed. + +Lemma ltn_fresh h k : (k < fresh h) = has (fun x => k < x.+1) (0 :: dom h). +Proof. by rewrite ltnS leq_lastkey. Qed. + +Lemma fresh_ltn h k : (fresh h < k) = all (fun x => x.+1 < k) (0 :: dom h). +Proof. +apply/idP/allP=>[H x D|]; last by apply; apply: lastkey_mem0. +by apply: leq_ltn_trans (dom0_fresh D) H. +Qed. + +(* unfolding equivalences into implications (forward) *) +Lemma fresh_leq_dom0 h x k : fresh h <= k -> x \in 0 :: dom h -> x < k. +Proof. exact: lastkey_ltn_dom0. Qed. +Lemma fresh_leq_dom h x k : fresh h <= k -> x \in dom h -> x < k. +Proof. exact: lastkey_ltn_dom. Qed. +Lemma fresh_ltn_dom0 h x k : fresh h < k -> x \in 0 :: dom h -> x.+1 < k. +Proof. by rewrite fresh_ltn=>/allP; apply. Qed. +Lemma fresh_ltn_dom h x k : fresh h < k -> x \in dom h -> x.+1 < k. +Proof. by move=>H D; apply: fresh_ltn_dom0 H _; rewrite inE D orbT. Qed. + +(* unfolding equivalences into implications (backward) *) +Lemma dom0_leq_fresh h x k : x \in 0 :: dom h -> k <= x.+1 -> k <= fresh h. +Proof. by move=>D H; rewrite leq_fresh; apply/hasP; exists x. Qed. +Lemma dom_leq_fresh h x k : x \in dom h -> k <= x.+1 -> k <= fresh h. +Proof. by move=>D; apply: dom0_leq_fresh; rewrite inE D orbT. Qed. +Lemma dom0_ltn_fresh h x k : x \in 0 :: dom h -> k <= x -> k < fresh h. +Proof. exact: dom0_leq_lastkey. Qed. +Lemma dom_ltn_fresh h x k : x \in dom h -> k <= x -> k < fresh h. +Proof. exact: dom_leq_lastkey. Qed. +End FreshLemmas. + +(* interaction of two fresh's *) + +Lemma fresh_mono A1 A2 (U1 : natmap A1) (U2 : natmap A2) + (h1 : U1) (h2 : U2) : + {subset dom h1 <= dom h2} -> + fresh h1 <= fresh h2. +Proof. exact: lastkey_mono. Qed. + +Lemma fresh_subdom A1 A2 (U1 : natmap A1) (U2 : natmap A2) + (h1 : U1) (h2 : U2) : + {subset dom h1 <= dom h2} -> + fresh h2 \notin dom h1. +Proof. by move/fresh_mono/fresh_dom. Qed. + +Lemma fresh_subdom_lt A1 A2 (U1 : natmap A1) (U2 : natmap A2) + (h1 : U1) (h2 : U2) x : + {subset dom h1 <= dom h2} -> + x \in dom h1 -> + x < fresh h2. +Proof. by move=>H /H /dom_fresh. Qed. + +Lemma fresh_subdomN A1 A2 (U1 : natmap A1) (U2 : natmap A2) + (h1 : U1) (h2 : U2) x : + {subset dom h1 <= dom h2} -> + x \in dom h1 -> + x != fresh h2. +Proof. by move=>S /(fresh_subdom_lt S); case: ltngtP. Qed. + +Lemma freshUnL A (U : natmap A) (h1 h2 : U) : + valid (h1 \+ h2) -> + fresh h1 <= fresh (h1 \+ h2). +Proof. exact: lastkeyUnL. Qed. + +Lemma freshUnLT A (U : natmap A) (h1 h2 : U) k : + valid (h1 \+ h2) -> + fresh (h1 \+ h2) <= k -> + fresh h1 <= k. +Proof. exact: lastkeyUnLT. Qed. + +Lemma freshUnR A (U : natmap A) (h1 h2 : U) : + valid (h1 \+ h2) -> + fresh h2 <= fresh (h1 \+ h2). +Proof. exact: lastkeyUnR. Qed. + +Lemma freshUnRT A (U : natmap A) (h1 h2 : U) k : + valid (h1 \+ h2) -> + fresh (h1 \+ h2) <= k -> + fresh h2 <= k. +Proof. exact: lastkeyUnRT. Qed. + +(* interaction with constructors *) + +Section FreshConstructors. +Variables (A : Type) (U : natmap A). +Implicit Type h : U. + +Lemma fresh_find h k : + fresh h <= k -> find k h = None. +Proof. exact: lastkey_find. Qed. + +Lemma freshF h : + last_key h != 0 -> + fresh (free h (last_key h)) < fresh h. +Proof. exact: lastkeyF. Qed. + +Lemma freshUn h1 h2 : + fresh (h1 \+ h2) = + if valid (h1 \+ h2) then maxn (fresh h1) (fresh h2) else 1. +Proof. by rewrite /fresh lastkeyUn maxnSS; case: ifP. Qed. + +Lemma freshPtUnV h k v : + fresh h <= k -> + valid (pts k v \+ h) = valid h. +Proof. exact: lastkeyPtUnV. Qed. + +(* easier name to remember *) +Definition valid_fresh := freshPtUnV. + +Lemma fresh_domPtUn h k v : + valid h -> + fresh h <= k -> + dom (pts k v \+ h) = rcons (dom h) k. +Proof. exact: lastkey_domPtUn. Qed. + +Lemma fresh_rangePtUn h k v : + valid h -> + fresh h <= k -> + range (pts k v \+ h) = rcons (range h) v. +Proof. exact: lastkey_rangePtUn. Qed. + +Lemma freshPtUnE h k v : + fresh (pts k v \+ h) = + if valid (pts k v \+ h) then maxn k.+1 (fresh h) else 1. +Proof. by rewrite /fresh lastkeyPtUnE maxnSS; case: ifP. Qed. + +Lemma freshPtUn h k v : + valid h -> fresh h <= k -> fresh (pts k v \+ h) = k.+1. +Proof. by move=>V N; rewrite /fresh lastkeyPtUn. Qed. + +Lemma freshPtUnEX v2 v1 h k : fresh (pts k v1 \+ h) = fresh (pts k v2 \+ h). +Proof. by congr (_.+1); apply: lastkeyPtUnEX. Qed. + +Lemma freshUnPtEX v2 v1 h k : fresh (h \+ pts k v1) = fresh (h \+ pts k v2). +Proof. by rewrite !(joinC h) (freshPtUnEX v2). Qed. + +Lemma freshU h k v : k \in dom h -> fresh (upd k v h) = fresh h. +Proof. by move=>D; congr _.+1; apply: lastkeyU D. Qed. + +Lemma freshUE h k v : + fresh (upd k v h) = + if k \in dom h then fresh h else + if valid (upd k v h) then maxn k.+1 (fresh h) else 1. +Proof. by rewrite /fresh lastkeyUE maxnSS; case: ifP=>//; case: ifP. Qed. + +(* membership *) + +Lemma In_fresh h k v : (k, v) \In h -> k < fresh h. +Proof. exact: In_lastkey. Qed. + +Lemma In_freshPtUn h x k v w : + fresh h <= k -> (x, w) \In h -> (x, w) \In pts k v \+ h. +Proof. exact: In_lastkeyPtUn. Qed. + +Lemma In_freshPtUn_inv h x k v w : + x < fresh h -> fresh h <= k -> + (x, w) \In pts k v \+ h -> (x, w) \In h. +Proof. exact: In_lastkeyPtUn_inv. Qed. + +End FreshConstructors. + +(* fresh and omap_fun -- compositions with omf_subdom/omf_subdom0 *) +Section FreshOmapFun. +Variables (A1 A2 : Type) (U1 : natmap A1) (U2 : natmap A2). +Implicit Types (f : omap_fun U1 U2) (h : U1). + +(* strict upper bound lemmas *) +Lemma odom_fresh f h k : k \in dom (f h) -> k < fresh h. +Proof. exact: odom_lastkey. Qed. +Lemma odom0_fresh f h k : k \in 0 :: dom (f h) -> k < fresh h. +Proof. exact: odom0_lastkey. Qed. +Lemma fresh_odom f h k : fresh h <= k -> k \notin dom (f h). +Proof. exact: lastkey_odom. Qed. +Lemma fresh_odom0 f h k : fresh h <= k -> k \notin 0 :: dom (f h). +Proof. exact: lastkey_odom0. Qed. + +(* equivalence lemmas (forward) *) +Lemma fresh_leq_odom0 f h x k : fresh h <= k -> x \in 0 :: dom (f h) -> x < k. +Proof. exact: lastkey_ltn_odom0. Qed. +Lemma fresh_leq_odom f h x k : fresh h <= k -> x \in dom (f h) -> x < k. +Proof. exact: lastkey_ltn_odom. Qed. +Lemma fresh_ltn_odom0 f h x k : fresh h < k -> x \in 0 :: dom (f h) -> x.+1 < k. +Proof. by move/fresh_ltn_dom0=>H /omf_subdom0 /H. Qed. +Lemma fresh_ltn_odom f h x k : fresh h < k -> x \in dom (f h) -> x.+1 < k. +Proof. by move/fresh_ltn_dom=>H /omf_subdom /H. Qed. + +(* equivalence lemmas (backward) *) +Lemma odom0_leq_fresh f h x k : x \in 0 :: dom (f h) -> k <= x.+1 -> k <= fresh h. +Proof. by move/omf_subdom0; apply/dom0_leq_fresh. Qed. +Lemma odom_leq_fresh f h x k : x \in dom (f h) -> k <= x.+1 -> k <= fresh h. +Proof. by move/omf_subdom; apply/dom_leq_fresh. Qed. +Lemma odom0_ltn_fresh f h x k : x \in 0 :: dom (f h) -> k <= x -> k < fresh h. +Proof. exact: odom0_leq_lastkey. Qed. +Lemma odom_ltn_fresh f h x k : x \in dom (f h) -> k <= x -> k < fresh h. +Proof. exact: odom_leq_lastkey. Qed. + +(* other lemmas *) + +Lemma fresh_omf f h : fresh (f h) <= fresh h. +Proof. exact: lastkey_omf. Qed. + +Lemma fresh_omfT f k h : + fresh h <= k -> + fresh (f h) <= k. +Proof. exact: lastkey_omfT'. Qed. + +Lemma fresh_omfUn f h1 h2 : + valid (h1 \+ h2) -> + fresh (f h1 \+ f h2) <= fresh (h1 \+ h2). +Proof. exact: lastkey_omfUn. Qed. + +Lemma fresh_omf_some f h : + (forall x, x \In h -> oapp predT false (omf f x)) -> + fresh (f h) = fresh h. +Proof. by move=>H; congr _.+1; apply: lastkey_omf_some H. Qed. + +Lemma fresh_omfE f h : + last_key h \in dom (f h) -> + fresh (f h) = fresh h. +Proof. by move=>H; congr _.+1; apply: lastkey_omfE. Qed. + +Lemma fresh_omfPtUn f k v h : + fresh h <= k -> + f (pts k v \+ h) = + if omf f (k, v) is Some w then pts k w \+ f h else f h. +Proof. exact: lastkey_omfPtUn. Qed. + +Lemma omf_subfresh f k v h : + fresh h <= k -> + {subset dom (f h) <= dom (f (pts k v \+ h))}. +Proof. exact: omf_sublastkey. Qed. + +End FreshOmapFun. + +(* fresh and filter *) + +Lemma fresh_umfilt_lt A (U : natmap A) (h : U) k : + fresh h <= k -> + um_filterk [pred x | x < k] h = h. +Proof. exact: lastkey_umfilt_lt. Qed. + +Lemma fresh_umfiltPtUn A (U : natmap A) (p : pred (nat * A)) (h : U) k v : + fresh h <= k -> + um_filter p (pts k v \+ h) = + if p (k, v) then pts k v \+ um_filter p h else um_filter p h. +Proof. exact: lastkey_umfiltPtUn. Qed. + +(* fresh and side *) + +Lemma fresh_sidePtUn (T : eqType) (Us : T -> Type) + (U : natmap (sigT Us)) (Ut : forall t, natmap (Us t)) + (h : U) t k v : + fresh h <= k -> + side_map Ut t (pts k v \+ h) = + if decP (t =P tag v) is left pf then + pts k (cast Us pf (tagged v)) \+ side_map Ut t h + else side_map Ut t h. +Proof. exact: lastkey_sidePtUn. Qed. + +Lemma fresh_dom_sidePtUn (T : eqType) (Us : T -> Type) + (U : natmap (sigT Us)) (Ut : forall t, natmap (Us t)) + (h : U) t k v : + fresh h <= k -> + dom (side_map Ut t (pts k v \+ h)) = + if valid h then + if t == tag v then rcons (dom (side_map Ut t h)) k + else dom (side_map Ut t h) + else [::]. +Proof. exact: lastkey_dom_sidePtUn. Qed. + +(* fresh and non-omap morphisms *) +Lemma pfresh (A : Type) (U1 : pcm) (U2 : natmap A) + (f : pcm_morph U1 U2) k (x y : U1) : + valid (x \+ y) -> + sep f x y -> + fresh (f (x \+ y)) <= k -> + fresh (f x) <= k. +Proof. exact: plastkey. Qed. + +Lemma pfreshT (A : Type) (U1 : pcm) (U2 : natmap A) + (f : full_pcm_morph U1 U2) k (x y : U1) : + valid (x \+ y) -> + fresh (f (x \+ y)) <= k -> + fresh (f x) <= k. +Proof. exact: plastkeyT. Qed. + +(* monotonicity for natmap nat *) +Lemma fresh_monoPtUn (U : natmap nat) k w (h : U) : + fresh h <= k -> + (forall k v, (k, v) \In h -> v < w) -> + um_mono (pts k w \+ h) = um_mono h. +Proof. exact: lastkey_monoPtUn. Qed. + +(* fresh and pcm ordering *) +Lemma fresh_fun A (U1 : pcm) (U2 : natmap A) + (f : full_pcm_morph U1 U2) (x y : U1) : + [pcm y <= x] -> + valid x -> + fresh (f y) <= fresh (f x). +Proof. exact: lastkey_fun. Qed. + +(* special case of fresh_fun when f = id *) +Lemma fresh_umpleq A (U : natmap A) (h1 h2 : U) : + [pcm h1 <= h2] -> + valid h2 -> + fresh h1 <= fresh h2. +Proof. exact: lastkey_umpleq. Qed. + +Lemma freshL A (U1 : pcm) (U2 : natmap A) + (f : full_pcm_morph U1 U2) (x y : U1) : + valid (x \+ y) -> + fresh (f x) <= fresh (f (x \+ y)). +Proof. exact: lastkeyL. Qed. + +Lemma freshR A (U1 : pcm) (U2 : natmap A) + (f : full_pcm_morph U1 U2) (x y : U1) : + valid (x \+ y) -> + fresh (f y) <= fresh (f (x \+ y)). +Proof. by apply: lastkeyR. Qed. + +(***********************************************) +(***********************************************) +(* Executing up to (and including/excluding) t *) +(***********************************************) +(***********************************************) + +Definition oexec_le V (U : natmap V) R (a : R -> V -> R) ks t (h : U) z0 := + oevalv a (&=ks `]-oo, t]) h z0. +Definition oexec_lt V (U : natmap V) R (a : R -> V -> R) ks t (h : U) z0 := + oevalv a (&=ks `]-oo, t[) h z0. + +(* explicit rewrites to facilitate folding *) +Lemma oexleP V U R a ks t h z0 : + oevalv a (&=ks `]-oo, t]) h z0 = @oexec_le V U R a ks t h z0. +Proof. by []. Qed. +Lemma oexltP V U R a ks t h z0 : + oevalv a (&=ks `]-oo, t[) h z0 = @oexec_lt V U R a ks t h z0. +Proof. by []. Qed. + +(* Main transfer lemmas *) + +Lemma oexleE V (U : natmap V) R a t v (h : U) ks (z0 : R) : + 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=>K H; rewrite /oexec_le/oexec_lt eqsl_uxR K (oev_rconsP _ _ _ H). +Qed. + +Arguments oexleE [V U R a t v h ks z0]. + +Lemma oexleNE V (U : natmap V) R a t (h : U) ks (z0 : R) : + (t \notin ks) || (t \notin dom h) -> + oexec_le a ks t h z0 = oexec_lt a ks t h z0. +Proof. +rewrite /oexec_le/oexec_lt; case K: (t \in ks)=>/= H; last first. +- rewrite (eqsl_uoxx (t1:=t) (t2:=t)); last by exact: sle_refl. + by rewrite eqsl_kk1 /= K cats0. +rewrite [LHS]oevFK [RHS]oevFK; congr oeval. +by rewrite eqsl_uxR K filter_rcons (negbTE H). +Qed. + +Arguments oexleNE [V U R a t h ks z0]. + +(**********************************************************) +(* Interaction of oexec_lt and oexec_le with constructors *) +(**********************************************************) + +(* DEVCOMMENT *) +(* +Lemma oexlt0 V (U : natmap V) R a ks (h : U) (z0 : R) : oexec_lt a ks 0 h z0 = z0. +Proof. by rewrite /oexec_lt squo0. Qed. + +Lemma oexle0 V R a ks (h : natmap V) (z0 : R) : oexec_le a ks 0 h z0 = z0. +Proof. +rewrite /oexec_le squx0; case: ifP=>//= _. +set xs := filter _ _; rewrite oevFK; set ys := filter _ _. +rewrite (_ : ys = [::]) //. +rewrite -[RHS](filter_pred0 xs); apply: eq_in_filter. +by move=>x; rewrite mem_filter=>/andP [/eqP ->]; rewrite cond_dom. +Qed. +*) +(* /DEVCOMMENT *) + +Lemma oexlt_notin V (U : natmap V) R a ks t (h : U) (z0 : R) : + t \notin ks -> + oexec_lt a ks t h z0 = oevalv a ks h z0. +Proof. by move=>K; rewrite /oexec_lt eqsl_uL_notinE. Qed. + +Arguments oexlt_notin [V U R a ks t h z0]. + +Lemma oexle_notin V (U : natmap V) R a ks t (h : U) (z0 : R) : + t \notin ks -> + oexec_le a ks t h z0 = oevalv a ks h z0. +Proof. by move=>K; rewrite /oexec_le eqsl_uL_notinE. Qed. + +Arguments oexle_notin [V U R a ks t h z0]. + +Lemma oexlt_cat V (U : natmap V) R a ks1 ks2 t (h : U) (z0 : R) : + t \notin ks1 -> + 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. +move=>T; rewrite /oexec_lt eqsl_uL_catE (negbTE T). +by rewrite oev_cat (eqsl_uL_notinE (s:=ks1)). +Qed. + +Arguments oexlt_cat [V U R a ks1 ks2 t h z0]. + +Lemma oexle_cat V (U : natmap V) R a ks1 ks2 t (h : U) (z0 : R) : + t \notin ks1 -> + 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. +move=>T; rewrite /oexec_le eqsl_uL_catE (negbTE T). +by rewrite oev_cat (eqsl_uL_notinE (s:=ks1)). +Qed. + +Arguments oexle_cat [V U R a ks1 ks2 t h z0]. + +Lemma oexlt_cons V (U : natmap V) R a ks k t v (h : U) (z0 : R) : + (k, v) \In h -> t != k -> + oexec_lt a (k :: ks) t h z0 = oexec_lt a ks t h (a z0 v). +Proof. +move=>H Ntk; rewrite /oexec_lt eqsl_uL_consE (negbTE Ntk) /=. +by move/In_find: H=>->. +Qed. + +Arguments oexlt_cons [V U R a ks k t v h z0]. + +Lemma oexle_cons V (U : natmap V) R a ks k t v (h : U) (z0 : R) : + (k, v) \In h -> + oexec_le a (k :: ks) t h z0 = + if t == k then a z0 v else oexec_le a ks t h (a z0 v). +Proof. +move=>H; rewrite /oexec_le eqsl_uL_consE. +by case: eqP=>/=; move/In_find: H=>->. +Qed. + +Arguments oexle_cons [V U R a ks k t v h z0]. + +(* for oexlt, we need to cover the case when k == t *) +Lemma oexlt_cons_same V (U : natmap V) R a ks k (h : U) (z0 : R) : + oexec_lt a (k :: ks) k h z0 = z0. +Proof. by rewrite /oexec_lt eqsl_uL_consL. Qed. + +Arguments oexlt_cons_same {V U R a ks k h z0}. + +Lemma oexlt_cons_notin V (U : natmap V) R a ks k t (h : U) (z0 : R) : + k \notin dom h -> t != k -> + oexec_lt a (k :: ks) t h z0 = oexec_lt a ks t h z0. +Proof. +move=>H Ntk; rewrite /oexec_lt eqsl_uL_consE (negbTE Ntk) /=. +by case: dom_find H=>//= -> _. +Qed. + +Arguments oexlt_cons_notin [V U R a ks k t h z0]. + +(* for oexle, we have two "notin" lemmas *) +Lemma oexle_cons_same_notin V (U : natmap V) R a ks k (h : U) (z0 : R) : + k \notin dom h -> oexec_le a (k :: ks) k h z0 = z0. +Proof. +move=>H. +by rewrite /oexec_le [in LHS]oevFK eqsl_uL_consL /= (negbTE H). +Qed. + +Arguments oexle_cons_same_notin [V U R a ks k h z0]. + +Lemma oexle_cons_notin V (U : natmap V) R a ks k t (h : U) (z0 : R) : + k \notin dom h -> t != k -> + oexec_le a (k :: ks) t h z0 = oexec_le a ks t h z0. +Proof. +move=>H Ntk; rewrite /oexec_le [in LHS]oevFK [in RHS]oevFK. +by rewrite eqsl_uL_consE (negbTE Ntk) /= (negbTE H). +Qed. + +Arguments oexle_cons_notin [V U R a ks k t h z0]. + +Lemma oexlt_rcons V (U : natmap V) R a ks k t (h : U) (z0 : R) : + t \in ks -> + oexec_lt a (rcons ks k) t h z0 = oexec_lt a ks t h z0. +Proof. by move=>T; rewrite /oexec_lt eqsl_uL_rconsE T. Qed. + +Arguments oexlt_rcons [V U R a ks k t h z0]. + +Lemma oexle_rcons V (U : natmap V) R a ks k t (h : U) (z0 : R) : + t \in ks -> + oexec_le a (rcons ks k) t h z0 = oexec_le a ks t h z0. +Proof. by move=>T; rewrite /oexec_le eqsl_uL_rconsE T. Qed. + +Arguments oexle_rcons [V U R a ks k t h z0]. + +Lemma oexlt_rcons_notin V (U : natmap V) R a ks k t v (h : U) (z0 : R) : + (k, v) \In h -> t \notin ks -> t != k -> + oexec_lt a (rcons ks k) t h z0 = + a (oexec_lt a ks t h z0) v. +Proof. +move=>H T Ntk; rewrite /oexec_lt eqsl_uL_rconsE /= (negbTE T) Ntk. +by rewrite oev_rconsE eqsl_uL_notinE //; move/In_find: H=>->. +Qed. + +Arguments oexlt_rcons_notin [V U R a ks k t v h z0]. + +Lemma oexle_rcons_notin V (U : natmap V) R a ks k t (h : U) (z0 : R) : + t \notin ks -> + oexec_le a (rcons ks k) t h z0 = oevalv a (rcons ks k) h z0. +Proof. by move=>T; rewrite /oexec_le eqsl_uL_rconsE (negbTE T). Qed. + +Arguments oexle_rcons_notin [V U R a ks k t h z0]. + +(* in case of oexlt we must cover the case when t == k *) +Lemma oexlt_rcons_same V (U : natmap V) R a ks k (h : U) (z0 : R) : + k \notin ks -> + oexec_lt a (rcons ks k) k h z0 = oevalv a ks h z0. +Proof. by move=>N; rewrite /oexec_lt eqsl_uL_rconsE eqxx /= (negbTE N). Qed. + +Arguments oexlt_rcons_same [V U R a ks k h z0]. + +(* in case of oexle, case t == k can be optimized *) +(* DEVCOMMENT *) +(* TODO doesn't simplify anything now, remove? *) +(* /DEVCOMMENT *) +Lemma oexle_rcons_same V (U : natmap V) R a ks k (h : U) (z0 : R) : + k \notin ks -> + oexec_le a (rcons ks k) k h z0 = oevalv a (rcons ks k) h z0. +Proof. exact: oexle_rcons_notin. Qed. + +Arguments oexle_rcons_same [V U R a ks k h z0]. + +(* in case of oexle, when we know the value at k *) +(* we can further expand the right-hand side *) +Lemma oexle_rcons_notin_in V (U : natmap V) R a ks k t v (h : U) (z0 : R) : + (k, v) \In h -> t \notin ks -> + oexec_le a (rcons ks k) t h z0 = a (oexec_le a ks t h z0) v. +Proof. +move=>H T; rewrite oexle_rcons_notin // oev_rconsE. +by move/In_find: (H)=>->; rewrite oexleNE ?T // /oexec_lt eqsl_uL_notinE. +Qed. + +Arguments oexle_rcons_notin_in [V U R a ks k t v h z0]. + +(* oexlt/oexle filter lemmas *) + +Lemma oexlt_umfilt V (U : natmap V) R a ks p t (h : U) (z0 : R) : + oexec_lt a ks t (um_filterk p h) z0 = + oevalv a (filter p (&=ks `]-oo, t[)) h z0. +Proof. by rewrite /oexec_lt oev_filter. Qed. + +Arguments oexlt_umfilt {V U R a ks p t h z0}. + +Lemma oexle_umfilt V (U : natmap V) R a ks p t (h : U) (z0 : R) : + oexec_le a ks t (um_filterk p h) z0 = + oevalv a (filter p (&=ks `]-oo, t])) h z0. +Proof. by rewrite /oexec_le oev_filter. Qed. + +Arguments oexle_umfilt {V U R a ks p t h z0}. + +(* two somewhat simpler rewrites under a condition *) +Lemma oexlt_umfiltN V (U : natmap V) R a ks p t (h : U) (z0 : R) : + (t \notin ks) || (p t) -> + oexec_lt a ks t (um_filterk p h) z0 = + oexec_lt a (filter p ks) t h z0. +Proof. +move=>H; rewrite /oexec_lt oev_umfilt eqsl_filterL //=. +apply: eq_in_oevK; rewrite -!filter_predI; apply: eq_in_filter=>k Ks /=. +by case D : (k \in dom h)=>//=; case/In_domX: D=>v /In_find ->. +Qed. + +Arguments oexlt_umfiltN [V U R a ks p t h z0]. + +Lemma oexle_umfiltN V (U : natmap V) R a ks p t (h : U) (z0 : R) : + (t \notin ks) || (p t) -> + oexec_le a ks t (um_filterk p h) z0 = + oexec_le a (filter p ks) t h z0. +Proof. +move=>H; rewrite /oexec_le oev_umfilt eqsl_filterL //=. +apply: eq_in_oevK; rewrite -!filter_predI; apply: eq_in_filter=>k Ks /=. +by case D : (k \in dom h)=>//=; case/In_domX: D=>v /In_find ->. +Qed. + +Arguments oexle_umfiltN [V U R a ks p t h z0]. + +(* restating the last two lemmas for the other direction *) +(* DEVCOMMENT *) +(* TODO why not just state them like this initially? *) +(* /DEVCOMMENT *) +Lemma oexlt_filter V (U : natmap V) R a ks p t (h : U) (z0 : R) : + (t \notin ks) || (p t) -> + oexec_lt a (filter p ks) t h z0 = + oexec_lt a ks t (um_filterk p h) z0. +Proof. by move=>H; rewrite (oexlt_umfiltN H). Qed. + +Lemma oexle_filter V (U : natmap V) R a ks p t (h : U) (z0 : R) : + (t \notin ks) || (p t) -> + oexec_le a (filter p ks) t h z0 = + oexec_le a ks t (um_filterk p h) z0. +Proof. by move=>H; rewrite (oexle_umfiltN H). Qed. + +(* interaction with the map *) + +Lemma oexlt_filter_dom V (U : natmap V) R a ks t (h : U) (z0 : R) : + (t \notin ks) || (t \in dom h) -> + oexec_lt a ks t h z0 = + oexec_lt a (filter (mem (dom h)) ks) t h z0. +Proof. by move=>H; rewrite oexlt_filter // umfiltk_dom'. Qed. + +Lemma oexle_filter_dom V (U : natmap V) R a ks t (h : U) (z0 : R) : + (t \notin ks) || (t \in dom h) -> + oexec_le a ks t h z0 = + oexec_le a (filter (mem (dom h)) ks) t h z0. +Proof. by move=>H; rewrite oexle_filter // umfiltk_dom'. Qed. + +(* interaction of oexlt, oexle and last *) + +Lemma oexlt_oexle_last V (U : natmap V) R a k ks t (h : U) (z0 : R) : + uniq (k::ks) -> t != k -> + oexec_lt a (k::ks) t h z0 = + oexec_le a (k::ks) (last k (&=ks `]-oo, t[)) h z0. +Proof. +move=>Uq Ntk; rewrite /oexec_lt/oexec_le [LHS]oevFK [RHS]oevFK. +by rewrite eqsl_uox_last. +Qed. + +Lemma oev_oexle_last V (U : natmap V) R a k ks (h : U) (z0 : R) : + uniq (k::ks) -> + oevalv a (k::ks) h z0 = + oexec_le a (k::ks) (last k ks) h z0. +Proof. +case/andP=>U1 U2; rewrite /oexec_le eqsl_uL_consE. +case: eqP=>/= [/last_nochange|/eqP/last_change H]. +- by rewrite (negbTE U1)=>/eqP->. +by congr oeval; apply/eqsl_lastR_uniq. +Qed. + +(*******************************************************) +(* oexlt/oexle on endpoints of intervals of invariance *) +(*******************************************************) + +(* these can be called "induction" lemmas over the interval *) + +Definition oex_inv V (U : natmap V) R (P : R -> Prop) a ks (h : U) (z0 : R) := + forall k v, (k, v) \In h -> + let z := oexec_lt a ks k h z0 in P z -> P (a z v). + +(* The naming schema in these lemmas is based on backward reasoning *) +(* E.g., the suffix ox means we're proving an lt_fact (o) from le_fact (x) *) + +(* two-sided intervals *) + +Lemma oex_ox V (U : natmap V) R (P : R -> Prop) a ks (h : U) t1 t2 (z0 : R) : + uniq ks -> t1 <[ks] t2 -> + {in &=ks `]t1, t2[, oex_inv P a ks h z0} -> + P (oexec_le a ks t1 h z0) -> P (oexec_lt a ks t2 h z0). +Proof. +move=>Uq T IH P0; rewrite /oexec_lt (eqsl_uxoo T) oev_cat. +apply: oev_indX=>// k ks1 ks2 v _ Kh Eh ->; rewrite -oev_cat. +have K : k \in &=ks `]t1, t2[ by rewrite Eh mem_cat inE eqxx orbT. +have Nk : k \notin ks1. +- move/(eqslice_uniq `]t1, t2[): Uq. + by rewrite Eh cat_uniq /= negb_or -andbA; case/and5P. +case/mem_oo: (K)=>// [_ T1K T2K]. +suff {IH Uq K}-> : ks1 = &=ks `]t1, k[ by rewrite -eqsl_uxoo //; apply: IH. +move: Eh; rewrite (eqslice_split (b:=true) (x:=k)) /=; last first. +- rewrite !lteBSide /= leEnat -seqlt_unlockE -seqle_unlock. + by rewrite T1K (sltW T2K). +rewrite eqsl_xoL T2K /= => Eh; rewrite (cat_cancel _ _ Eh) //. +by apply: eqsliceRO_notin. +Qed. + +(* one can try to derive the next lemma from the previous *) +(* but we need to reason about successor of t1 in ks *) +(* which we don't have yet. Hence, we prove the lemma directly *) +(* using the same approach as in the previous lemma. *) + +Lemma oex_oo V (U : natmap V) R (P : R -> Prop) a ks (h : U) t1 t2 (z0 : R) : + uniq ks -> t1 <=[ks] t2 -> + {in &=ks `[t1, t2[, oex_inv P a ks h z0} -> + P (oexec_lt a ks t1 h z0) -> P (oexec_lt a ks t2 h z0). +Proof. +move=>Uq T IH P0; rewrite /oexec_lt (eqsl_uoxo T) oev_cat. +apply: oev_indX=>// k ks1 ks2 v _ Kh Eh ->; rewrite -oev_cat. +have K : k \in &=ks `[t1, t2[ by rewrite Eh mem_cat inE eqxx orbT. +have Nk : k \notin ks1. +- move/(eqslice_uniq `[t1, t2[): Uq. + by rewrite Eh cat_uniq /= negb_or -andbA; case/and5P. +case/mem_xo: (K)=>// [_ T1K T2K]. +suff {IH Uq K}-> : ks1 = &=ks `[t1, k[ by rewrite -eqsl_uoxo //; apply: IH. +move: Eh; rewrite (eqslice_split (b:=true) (x:=k)) /=; last first. +- by rewrite !lteBSide /= !leEnat -!seqle_unlock T1K (sltW T2K). +rewrite (eqsl_xoL k) T2K /= => Eh; rewrite (cat_cancel _ _ Eh) //. +by apply: eqsliceRO_notin. +Qed. + +Lemma oex_xo V (U : natmap V) R (P : R -> Prop) a ks (h : U) t1 t2 (z0 : R) : + uniq ks -> t1 <=[ks] t2 -> + {in &=ks `[t1, t2], oex_inv P a ks h z0} -> + P (oexec_lt a ks t1 h z0) -> P (oexec_le a ks t2 h z0). +Proof. +move=>Uq T IH P0; rewrite /oexec_le (eqsl_uoxx T) oev_cat. +apply: oev_indX=>// k ks1 ks2 v _ Kh Eh ->; rewrite -oev_cat. +have K : k \in &=ks `[t1, t2] by rewrite Eh mem_cat inE eqxx orbT. +have Nk : k \notin ks1. +- move/(eqslice_uniq `[t1, t2]): Uq. + by rewrite Eh cat_uniq /= negb_or -andbA; case/and5P. +case/mem_xx: (K)=>// [Ks T1K T2K]. +suff {IH Uq K}-> : ks1 = &=ks `[t1, k[ by rewrite -eqsl_uoxo //; apply: IH. +move: Eh; rewrite (eqslice_split (b:=true) (x:=k)) /=; last first. +- by rewrite /order.Order.le/=/order.Order.le/= -!seqle_unlock T1K T2K. +rewrite eqsl_xxL T2K Ks /= => Eh; rewrite (cat_cancel _ _ Eh) //. +by apply: eqsliceRO_notin. +Qed. + +Lemma oex_xx V (U : natmap V) R (P : R -> Prop) a ks (h : U) t1 t2 (z0 : R) : + uniq ks -> t1 <=[ks] t2 -> + {in &=ks `]t1, t2], oex_inv P a ks h z0} -> + P (oexec_le a ks t1 h z0) -> P (oexec_le a ks t2 h z0). +Proof. +move=>Uq T IH P0; rewrite /oexec_le (eqsl_uxox T) oev_cat. +apply: oev_indX=>// k ks1 ks2 v _ Kh Eh ->; rewrite -oev_cat. +have K : k \in &=ks `]t1, t2] by rewrite Eh mem_cat inE eqxx orbT. +have Nk : k \notin ks1. +- move/(eqslice_uniq `]t1, t2]): Uq. + by rewrite Eh cat_uniq /= negb_or -andbA; case/and5P. +case/mem_ox: (K)=>// [Ks T1K T2K]. +suff {IH Uq K}-> : ks1 = &=ks `]t1, k[ by rewrite -eqsl_uxoo //; apply: IH. +move: Eh; rewrite (eqslice_split (b:=true) (x:=k)) /=; last first. +- rewrite /order.Order.le/=/order.Order.le/=/order.Order.lt /=. + by rewrite -seqlt_unlock -seqle_unlock T1K T2K. +rewrite eqsl_xxL T2K Ks /= => Eh; rewrite (cat_cancel _ _ Eh) //. +by apply: eqsliceRO_notin. +Qed. + +Arguments oex_ox [V U R] P [a ks h t1 t2 z0]. +Arguments oex_oo [V U R] P [a ks h t1 t2 z0]. +Arguments oex_xo [V U R] P [a ks h t1 t2 z0]. +Arguments oex_xx [V U R] P [a ks h t1 t2 z0]. + +(* one-sided intervals *) + +Lemma oex_ux V (U : natmap V) R (P : R -> Prop) a ks (h : U) t (z0 : R) : + uniq ks -> + {in &=ks `]t,+oo[, oex_inv P a ks h z0} -> + P (oexec_le a ks t h z0) -> P (oevalv a ks h z0). +Proof. +move=>Uq IH P0; rewrite (eqsl_uxou ks t) oev_cat. +apply: oev_indX=>// k ks1 ks2 v _ Kh Eh ->; rewrite -oev_cat. +have K : k \in &=ks `]t, +oo[ by rewrite Eh mem_cat inE eqxx orbT. +have Nk : k \notin ks1. +- move/(eqslice_uniq `]t, +oo[): Uq. + by rewrite Eh cat_uniq /= negb_or -andbA; case/and5P. +case/mem_ou: (K)=>// [Ks TK]. +suff {IH Uq K}-> : ks1 = &=ks `]t, k[ by rewrite -eqsl_uxoo //; apply: IH. +move: Eh; rewrite (eqslice_split (b:=true) (x:=k)) /=; last first. +- by rewrite /order.Order.le/=/order.Order.lt/= -seqlt_unlock TK. +rewrite eqsl_xuL Ks /= => Eh; rewrite (cat_cancel _ _ Eh) //. +by apply: eqsliceRO_notin. +Qed. + +Lemma oex_uo V (U : natmap V) R (P : R -> Prop) a ks (h : U) t (z0 : R) : + uniq ks -> + {in &=ks `[t,+oo[, oex_inv P a ks h z0} -> + P (oexec_lt a ks t h z0) -> P (oevalv a ks h z0). +Proof. +move=>Uq IH P0; rewrite (eqsl_uoxu ks t) oev_cat. +apply: oev_indX=>// k ks1 ks2 v _ Kh Eh ->; rewrite -oev_cat. +have K : k \in &=ks `[t, +oo[ by rewrite Eh mem_cat inE eqxx orbT. +have Nk : k \notin ks1. +- move/(eqslice_uniq `[t, +oo[): Uq. + by rewrite Eh cat_uniq /= negb_or -andbA; case/and5P. +case/mem_xu: (K)=>// [Ks TK]. +suff {IH Uq K}-> : ks1 = &=ks `[t, k[ by rewrite -eqsl_uoxo //; apply: IH. +move: Eh; rewrite (eqslice_split (b:=true) (x:=k)) /=; last first. +- by rewrite /order.Order.le/=/order.Order.le/= -seqle_unlock TK. +rewrite eqsl_xuL Ks => Eh; rewrite (cat_cancel _ _ Eh) //. +by apply: eqsliceRO_notin. +Qed. + +Lemma oex_xu V (U : natmap V) R (P : R -> Prop) a ks (h : U) t (z0 : R) : + uniq ks -> + {in &=ks `]-oo, t], oex_inv P a ks h z0} -> + P z0 -> P (oexec_le a ks t h z0). +Proof. +move=>Uq IH P0; apply: oev_indX=>// k ks1 ks2 v _ Kh Eh ->. +have K : k \in &=ks `]-oo, t] by rewrite Eh mem_cat inE eqxx orbT. +have Nk : k \notin ks1. +- move/(eqslice_uniq `]-oo, t]): Uq. + by rewrite Eh cat_uniq /= negb_or -andbA; case/and5P. +case/mem_ux: (K)=>// [Ks TK]. +suff {IH Uq K}-> : ks1 = &=ks `]-oo, k[ by apply: IH. +move: Eh; rewrite (eqslice_split (b:=true) (x:=k)) //=; last first. +- by rewrite /order.Order.le/=/order.Order.le/= -seqle_unlock. +rewrite eqsl_xxL TK Ks /= => Eh; rewrite (cat_cancel _ _ Eh) //. +by apply: eqsliceRO_notin. +Qed. + +Lemma oex_ou V (U : natmap V) R (P : R -> Prop) a ks (h : U) t (z0 : R) : + uniq ks -> + {in &=ks `]-oo, t[, oex_inv P a ks h z0} -> + P z0 -> P (oexec_lt a ks t h z0). +Proof. +move=>Uq IH P0; apply: oev_indX=>// k ks1 ks2 v _ Kh Eh ->. +have K : k \in &=ks `]-oo, t[ by rewrite Eh mem_cat inE eqxx orbT. +have Nk : k \notin ks1. +- move/(eqslice_uniq `]-oo, t[): Uq. + by rewrite Eh cat_uniq /= negb_or -andbA; case/and5P. +case/mem_uo: (K)=>// [Ks TK]. +suff {IH Uq K}-> : ks1 = &=ks `]-oo, k[ by apply: IH. +move: Eh; rewrite (eqslice_split (b:=true) (x:=k)) /=; last first. +- by rewrite lteBSide /= leEnat -seqle_unlock (sltW TK). +rewrite eqsl_xoL TK => Eh; rewrite (cat_cancel _ _ Eh) //. +by apply: eqsliceRO_notin. +Qed. + +Arguments oex_ux [V U R] P [a ks h t z0]. +Arguments oex_uo [V U R] P [a ks h t z0]. +Arguments oex_xu [V U R] P [a ks h t z0]. +Arguments oex_ou [V U R] P [a ks h t z0]. + +(* whole list *) + +Lemma oex_uu V (U : natmap V) R (P : R -> Prop) a ks (h : U) (z0 : R) : + uniq ks -> + {in ks, oex_inv P a ks h z0} -> + P z0 -> P (oevalv a ks h z0). +Proof. +move=>Uq IH P0; apply: oev_indX=>// k ks1 ks2 v _ Kh Eh ->. +have K : k \in ks by rewrite Eh mem_cat inE eqxx orbT. +have Nk : k \notin ks1. +- move: Uq. + by rewrite Eh cat_uniq /= negb_or -andbA; case/and5P. +suff -> : ks1 = &=ks `]-oo, k[ by apply: IH. +move: Eh; rewrite {1}(eqsl_uxou ks k) eqsl_uxR K -cats1 -catA /= => Eh. +by rewrite (cat_cancel _ _ Eh) //=; apply: eqsliceRO_notin. +Qed. + +Arguments oex_uu [V U R] P [a ks h z0]. + +(**********************************************) +(* functional versions of the interval lemmas *) +(**********************************************) + +Definition oexF_inv V (U : natmap V) R X (f : R -> X) a ks (h : U) (z0 : R) := + forall k v, (k, v) \In h -> + let z := oexec_lt a ks k h z0 in f (a z v) = f z. + +(* two-sided intervals *) + +Lemma oexF_ox V (U : natmap V) R X (f : R -> X) a ks (h : U) t1 t2 (z0 : R) : + uniq ks -> t1 <[ks] t2 -> + {in &=ks `]t1, t2[, oexF_inv f a ks h z0} -> + f (oexec_lt a ks t2 h z0) = f (oexec_le a ks t1 h z0). +Proof. +move=>Uq T H; pose P := fun x => f x = f (oexec_le a ks t1 h z0). +by apply: (oex_ox P) T _ _=>// x S v K E ; rewrite /P -E; apply: H. +Qed. + +Lemma oexF_oo V (U : natmap V) R X (f : R -> X) a ks (h : U) t1 t2 (z0 : R) : + uniq ks -> t1 <=[ks] t2 -> + {in &=ks `[t1, t2[, oexF_inv f a ks h z0} -> + f (oexec_lt a ks t2 h z0) = f (oexec_lt a ks t1 h z0). +Proof. +move=>Uq T H; pose P := fun x => f x = f (oexec_lt a ks t1 h z0). +by apply: (oex_oo P) T _ _=>// x S v K E; rewrite /P -E; apply: H. +Qed. + +Lemma oexF_xo V (U : natmap V) R X (f : R -> X) a ks (h : U) t1 t2 (z0 : R) : + uniq ks -> t1 <=[ks] t2 -> + {in &=ks `[t1, t2], oexF_inv f a ks h z0} -> + f (oexec_le a ks t2 h z0) = f (oexec_lt a ks t1 h z0). +Proof. +move=>Uq T H; pose P := fun x => f x = f (oexec_lt a ks t1 h z0). +by apply: (oex_xo P) T _ _=>// x S v K E; rewrite /P -E; apply: H. +Qed. + +Lemma oexF_xx V (U : natmap V) R X (f : R -> X) a ks (h : U) t1 t2 (z0 : R) : + uniq ks -> t1 <=[ks] t2 -> + {in &=ks `]t1, t2], oexF_inv f a ks h z0} -> + f (oexec_le a ks t2 h z0) = f (oexec_le a ks t1 h z0). +Proof. +move=>Uq T H; pose P := fun x => f x = f (oexec_le a ks t1 h z0). +by apply: (oex_xx P) T _ _=>// x S v K E; rewrite /P -E; apply: H. +Qed. + +Arguments oexF_ox [V U R X] f [a ks h t1 t2 z0]. +Arguments oexF_oo [V U R X] f [a ks h t1 t2 z0]. +Arguments oexF_xo [V U R X] f [a ks h t1 t2 z0]. +Arguments oexF_xx [V U R X] f [a ks h t1 t2 z0]. + +(* one-sided intervals *) + +Lemma oexF_ux V (U : natmap V) R X (f : R -> X) a ks (h : U) t (z0 : R) : + uniq ks -> + {in &=ks `]t, +oo[, oexF_inv f a ks h z0} -> + f (oevalv a ks h z0) = f (oexec_le a ks t h z0). +Proof. +move=>Uq H; pose P := fun x => f x = f (oexec_le a ks t h z0). +by apply: (oex_ux P)=>// x S v K E; rewrite /P -E; apply: H. +Qed. + +Lemma oexF_uo V (U : natmap V) R X (f : R -> X) a ks (h : U) t (z0 : R) : + uniq ks -> + {in &=ks `[t, +oo[, oexF_inv f a ks h z0} -> + f (oevalv a ks h z0) = f (oexec_lt a ks t h z0). +Proof. +move=>Uq H; pose P := fun x => f x = f (oexec_lt a ks t h z0). +by apply: (oex_uo P)=>// x S v K E; rewrite /P -E; apply: H. +Qed. + +Lemma oexF_xu V (U : natmap V) R X (f : R -> X) a ks (h : U) t (z0 : R) : + uniq ks -> + {in &=ks `]-oo, t], oexF_inv f a ks h z0} -> + f (oexec_le a ks t h z0) = f z0. +Proof. +move=>Uq H; pose P := fun x => f x = f z0. +by apply: (oex_xu P)=>// x S v K E; rewrite /P -E; apply: H. +Qed. + +Lemma oexF_ou V (U : natmap V) R X (f : R -> X) a ks (h : U) t (z0 : R) : + uniq ks -> + {in &=ks `]-oo, t[, oexF_inv f a ks h z0} -> + f (oexec_lt a ks t h z0) = f z0. +Proof. +move=>Uq H; pose P := fun x => f x = f z0. +by apply: (oex_ou P)=>// x S v K E; rewrite /P -E; apply: H. +Qed. + +Arguments oexF_ux [V U R X] f [a ks h t z0]. +Arguments oexF_uo [V U R X] f [a ks h t z0]. +Arguments oexF_xu [V U R X] f [a ks h t z0]. +Arguments oexF_ou [V U R X] f [a ks h t z0]. + +(* whole list *) + +Lemma oexF_uu V (U : natmap V) R X (f : R -> X) a ks (h : U) (z0 : R) : + uniq ks -> + {in ks, oexF_inv f a ks h z0} -> + f (oevalv a ks h z0) = f z0. +Proof. +move=>Uq H; pose P := fun x => f x = f z0. +by apply: (oex_uu P)=>// x S v K E; rewrite /P -E; apply: H. +Qed. + +Arguments oexF_uu [V U R X] f [a ks h z0]. + +(* we can now combine the split lemmas with *) +(* cons properties of oeval, to obtain a point-split *) +(* when the middle point is in the map *) + +Lemma oev2_split V (U : natmap V) R a t1 v (h : U) ks (z0 : R) : + t1 \in ks -> (t1, v) \In h -> + oevalv a ks h z0 = + oevalv a (&=ks `]t1, +oo[) h (a (oexec_lt a ks t1 h z0) v). +Proof. +move=>D H; rewrite {1}(eqsl_uoxu ks t1) oev_cat. +by rewrite eqsl_xuL D /=; move/In_find: H=>->. +Qed. + +Arguments oev2_split [V U R a t1 v h ks z0] _ _. + +Lemma oex2_split V (U : natmap V) R a t1 t2 v (h : U) ks (z0 : R) : + t1 <[ks] t2 -> (t1, v) \In h -> + oexec_lt a ks t2 h z0 = + oevalv a (&=ks `]t1, t2[) h (a (oexec_lt a ks t1 h z0) v). +Proof. +move=>T H; rewrite /oexec_lt. +rewrite (eqsl_uoxo (sltW T)) oev_cat eqsl_xoL T /=. +by move/In_find: H=>->. +Qed. + +Arguments oex2_split [V U R a t1 t2 v h ks z0] _ _. + +(* we frequently iterate oex2_split, so the following saves verbiage *) +Lemma oex3_split V (U : natmap V) R a t1 t2 t3 v1 v2 (h : U) ks (z0 : R) : + t1 <[ks] t2 -> (t1, v1) \In h -> + t2 <[ks] t3 -> (t2, v2) \In h -> + oexec_lt a ks t3 h z0 = + oevalv a (&=ks `]t2, t3[) h + (a (oevalv a (&=ks `]t1, t2[) h + (a (oexec_lt a ks t1 h z0) v1)) + v2). +Proof. +move=>T1 H1 T2 H2. +by rewrite (oex2_split T2 H2) (oex2_split T1 H1). +Qed. + +Arguments oex3_split [V U R a t1 t2 t3 v1 v2 h ks z0] _ _ _ _. + +(******************************************) +(* Interaction of consec with oexlt/oexle *) +(******************************************) + +Lemma oexlt_consec V (U : natmap V) R a ks t1 t2 (h : U) (z0 : R) : + uniq ks -> + consec ks t1 t2 -> + oexec_lt a ks t2 h z0 = oexec_le a ks t1 h z0. +Proof. +move=>Uq C; apply: (oexF_ox id)=>//; first by apply: consec_slt. +by move=>x; rewrite consec_oo. +Qed. + +Arguments oexlt_consec [V U R a ks t1 t2 h z0]. + +(* the following is direct *) +Lemma oexlt_split V (U : natmap V) R a x1 t1 t2 x2 (h : U) (z0 : R) : + uniq (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=>Uq; apply: oexlt_consec=>//; apply/consecP_split=>//. +- by rewrite mem_cat !inE eqxx !orbT. +by exists x1, x2. +Qed. + +Lemma oexlt_consec_in V (U : natmap V) R a t1 t2 v (h : U) ks (z0 : R) : + uniq ks -> + (t1, v) \In h -> + consec ks t1 t2 -> + oexec_lt a ks t2 h z0 = a (oexec_lt a ks t1 h z0) v. +Proof. +move=>Uq H C; move/slt_memE: (consec_slt C)=>T. +by rewrite (oexlt_consec Uq C) (oexleE T H). +Qed. + +Lemma oexlt_consec_fst V (U : natmap V) R a t (h : U) k ks (z0 : R) : + uniq (k::ks) -> k \notin dom h -> t \in k::ks -> + consec (k::ks) k t -> + oexec_lt a (k::ks) t h z0 = z0. +Proof. +move=>Uq K T C; case/(consecP_split _ Uq T): C=>xs1 [xs2] X. +case: xs1 X Uq {T}=>[|x xs1]; last first. +- by case=>->-> /=; rewrite mem_cat inE eqxx !orbT. +case=>->{ks}; rewrite /= inE negb_or -andbA; case/and4P=>U1 U2 U3 U4. +rewrite oexlt_cons_notin ?inE 1?negb_or ?(eq_sym t) ?U1 ?U2 //. +by rewrite oexlt_cons_same. +Qed. + +Lemma oexlt_consec_find V (U : natmap V) R a t1 t2 (h : U) ks (z0 : R) : + uniq ks -> + consec ks t1 t2 -> + oexec_lt a ks t2 h z0 = + if find t1 h is Some e + then a (oexec_lt a ks t1 h z0) e + else oexec_lt a ks t1 h z0. +Proof. +move=>Uq C; rewrite (oexlt_consec Uq C). +case E : (find t1 h)=>[e|]; first by move/In_find: E=>/(oexleE (consec_mem C)) ->. +by rewrite oexleNE // orbC; move/In_findN: E=>->. +Qed. + +Arguments oexlt_consec_find [V U R a t1 t2 h ks z0]. + +Lemma oexlt_last V (U : natmap V) R a t2 (h : U) ks (z0 : R) : + uniq ks -> + t2 \notin ks -> + oexec_lt a ks t2 h z0 = + if ks is k :: ks' then + if last k ks' \in dom h then + oexec_le a ks (last k ks') h z0 + else oexec_lt a ks (last k ks') h z0 + else z0. +Proof. +case: ks=>[|k ks] Uq //= T2. +have C : consec (k :: ks) (last k ks) t2. +- by rewrite (consecP_lastE k) // mem_last /=. +rewrite (oexlt_consec_find Uq C). +case: dom_find=>// v /In_find E _. +by rewrite -oexleE // mem_last. +Qed. + +Lemma oexlt_last0 V (U : natmap V) R a t2 (h : U) ks (z0 : R) : + uniq (0::ks) -> + t2 \notin 0::ks -> + oexec_lt a ks t2 h z0 = + if last 0 ks \in dom h then + oexec_le a ks (last 0 ks) h z0 + else oexec_lt a ks (last 0 ks) h z0. +Proof. +move=>Uq T2. +have D : 0 \notin dom h by rewrite cond_dom. +have N : t2 != 0 by case: eqP T2=>// ->. +rewrite -(oexlt_cons_notin D N) oexlt_last //. +case: ifP=>L. +- by rewrite oexle_cons_notin // (dom_cond L). +case: ks Uq {N T2 D L}=>[|k ks] //= Uq. +rewrite oexlt_cons_notin ?cond_dom //. +apply/negP=>/eqP N. +by rewrite -N (mem_last k ks) in Uq. +Qed. + +(* The lemmas past this point are currently used for some examples, *) +(* but will be deprecated and removed in future releases *) + +(*******************) +(*******************) +(* Gapless natmaps *) +(*******************) +(*******************) + +Section Gapless. +Variables (A : Type) (U : natmap A). +Implicit Type h : U. + +Definition gapless h := forall k, 0 < k <= last_key h -> k \in dom h. + +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. +Proof. by move=>k; rewrite lastkey0; case: k. Qed. + +Lemma gp_fresh h u : gapless (pts (fresh h) u \+ h) <-> gapless h. +Proof. +case V : (valid h); last first. +- by move/invalidE: (negbT V)=>->; rewrite join_undef. +split=>H k; move: (V); rewrite -(freshPtUnV u (leqnn _))=>V'; last first. +- rewrite lastkeyPtUn // domPtUn inE V' /= (leq_eqVlt k) eq_sym. + by move: (H k); rewrite -(ltnS k); case: (ltngtP k (last_key h).+1). +rewrite -(ltnS k) -/(fresh h); case/andP=>Z N. +move: (H k); rewrite lastkeyPtUn // domPtUn inE V' Z /= leq_eqVlt eq_sym. +by case: ltngtP N=>//= _ _; apply. +Qed. + +Lemma gpPtUn h k u : + valid (pts k u \+ h) -> + gapless (pts k u \+ h) -> last_key h < k -> k = fresh h. +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 (pts k u \+ h) = k. +- move: (C (fresh h)); rewrite Z (leq_eqVlt _ k) Y orbT andbT; move/(_ (erefl _)). + by rewrite domPtUn inE (negbTE X) /=; case/andP=>_ /dom_fresh; rewrite ltnn. +by rewrite lastkeyPtUn // (validR V). +Qed. + +Lemma gaplessPtUnE u2 u1 h k : + gapless (pts k u1 \+ h) <-> gapless (pts k u2 \+ h). +Proof. +rewrite /gapless (lastkeyPtUnEX 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 U][h] u. + +(*****************************) +(*****************************) +(* Natmaps with binary range *) +(*****************************) +(*****************************) + +Section AtVal. +Variables (A : Type) (U : natmap (A * A)). +Implicit Type h : U. + +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; case/(gapless_pleq G V): {V} H (V)=>h [->{h2} H] V. +by apply: atvalUn. +Qed. + +Definition last_atval v h := atval v (last_key h) h. + +Lemma lastatval0 v : last_atval v Unit = v. +Proof. by rewrite /last_atval /atval lastkey0 find0E. Qed. + +Lemma lastatval_undef v : last_atval v undef = v. +Proof. by rewrite /last_atval /atval lastkey_undef find_undef. Qed. + +Lemma lastatvalPt v p x : p != 0 -> last_atval v (pts p x) = x.2. +Proof. +by move=>V; rewrite /last_atval /atval lastkeyPt // findPt /= V. +Qed. + +Lemma lastatval_fresh v x h : + valid h -> last_atval v (pts (fresh h) x \+ h) = x.2. +Proof. +by move=>V; rewrite /last_atval /atval lastkeyPtUn // findPtUn // freshPtUnV. +Qed. + +Lemma lastatvalUn v h1 h2 : + last_atval v (h1 \+ h2) = + if valid (h1 \+ h2) then + if last_key h1 < last_key h2 then last_atval v h2 else last_atval v h1 + else v. +Proof. +rewrite /last_atval /atval lastkeyUn maxnC /maxn. +case: ifP; last by move/negbT/invalidE=>->; rewrite find_undef. +case: (ltngtP (last_key h1) (last_key h2)) => N V. +- by rewrite findUnR //; case: lastkeyP N. +- by rewrite findUnL //; case: lastkeyP N. +by rewrite !(lastkeyUn0 V N) unitL lastkey0 find0E. +Qed. + +Lemma lastatval_freshUn v x h1 h2 : + valid h1 -> [pcm h2 <= h1] -> + last_atval v (pts (fresh h1) x \+ h2) = x.2. +Proof. +move=>V H; move: (pleqV H V)=>V2; move: (fresh_mono (umpleq_subdom V H))=>N. +by rewrite /last_atval /atval lastkeyPtUn // findPtUn // freshPtUnV. +Qed. + +Lemma lastatval_indom v h : + last_atval v h <> v -> last_key h \in dom h. +Proof. by rewrite /last_atval /atval; case: dom_find=>// ->. Qed. + +End AtVal. + +(* in case A = bool *) +Lemma lastatval_indomb (U : natmap (bool*bool)) (h : U) : + last_atval false h -> last_key h \in dom h. +Proof. by move=>H; apply: (lastatval_indom (v:=false)); rewrite H. Qed. + +(* Continuous maps with binary range *) + +Section Continuous. +Variables (A : Type) (U : natmap (A * A)). +Implicit Type h : U. + +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 undef. +Proof. by move=>x w; rewrite find_undef. Qed. + +Lemma cn0 v : continuous v Unit. +Proof. by move=>x w; rewrite find0E. Qed. + +Lemma cn_fresh v h x : + valid h -> + continuous v (pts (fresh h) x \+ h) <-> + continuous v h /\ last_atval v h = x.1. +Proof. +rewrite -(freshPtUnV x (leqnn _))=>V; split; last first. +- case=>C H k y; rewrite !findPtUn2 // eqSS; case: ltngtP=>N. + - by rewrite ltn_eqF; [apply: C | apply: (ltn_trans N _)]. + - by move/find_some/dom_fresh/(ltn_trans N); rewrite ltnn. + by case=><-; rewrite N ltn_eqF. +move=>C; split; last first. +- move: (C (last_key h) x). + by rewrite !findPtUn2 // eqxx ltn_eqF //; apply. +move=>k w; case: (ltnP k (last_key h))=>N; last first. +- by move/find_some/dom_fresh/(leq_ltn_trans N); rewrite ltnn. +by move: (C k w); rewrite !findPtUn2 // eqSS !ltn_eqF // (ltn_trans N _). +Qed. + +Lemma cn_sub v h x y k : + valid (pts k.+1 (x, y) \+ h) -> continuous v (pts k.+1 (x, y) \+ h) -> + oapp snd v (find k h) = x. +Proof. +by move=>V /(_ k (x, y)); rewrite !findPtUn2 // eqxx ltn_eqF //; apply. +Qed. + +End Continuous. + +Arguments cn_fresh [A U][v][h][x] _. + +(* Complete natmaps with binary range *) + +Section Complete. +Variables (A : Type) (U : natmap (A * A)). +Implicit Type h : U. + +Definition complete v0 h vn := + [/\ valid h, gapless h, continuous v0 h & last_atval v0 h = vn]. + +Lemma cm_valid v0 h vn : complete v0 h vn -> valid h. +Proof. by case. Qed. + +Lemma cm0 v : complete v Unit v. +Proof. by split=>//; [apply: gp0|apply: cn0|rewrite lastatval0]. Qed. + +Lemma cm_fresh v0 vn h x : + complete v0 (pts (fresh h) x \+ h) vn <-> vn = x.2 /\ complete v0 h x.1. +Proof. +split. +- by case=>/validR V /gp_fresh G /(cn_fresh V) []; rewrite lastatval_fresh. +case=>-> [V] /(gp_fresh x) G C E; split=>//; +by [rewrite freshPtUnV|apply/(cn_fresh V)| rewrite lastatval_fresh]. +Qed. + +Lemma cmPtUn v0 vn h k x : + complete v0 (pts k x \+ h) vn -> last_key h < k -> k = fresh h. +Proof. by case=>V /(gpPtUn V). Qed. + +Lemma cmPt v0 vn k x : complete v0 (pts k x) vn -> k = 1 /\ x = (v0, vn). +Proof. +case; rewrite validPt; case: k=>//= k _ /(_ 1). +rewrite lastkeyPt //= domPt inE /= => /(_ (erefl _))/eqP ->. +move/(_ 0 x); rewrite findPt findPt2 /= => -> //. +by rewrite /last_atval lastkeyPt // /atval findPt /= => <-; case: x. +Qed. + +Lemma cmCons v0 vn k x h : + complete v0 (pts k x \+ h) vn -> last_key h < k -> + [/\ k = fresh h, vn = x.2 & complete v0 h x.1]. +Proof. by move=>C H; move: {C} (cmPtUn C H) (C)=>-> /cm_fresh []. Qed. + +End Complete. + +Prenex Implicits cm_valid cmPt. + +(************************) +(************************) +(************************) +(* Surgery on histories *) +(* using leq filtering *) +(************************) +(************************) +(************************) + +Notation le t := (fun '(k, _) => k <= t). +Notation lt t := (fun '(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 + (fun kv : (nat * V) => + let '(k, v) := kv in 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 + (fun kv : nat * V => + let '(k, v) := kv in t1 < k < t2). +Proof. by move=>T [k v] /=; rewrite -ltnNge. Qed. + +Lemma lastkey_umfilt_leq A (U : natmap A) (h : U) t : + last_key (um_filterk (leq^~ t) h) <= t. +Proof. +case V : (valid h); last first. +- by move/negbT/invalidE: V=>->; rewrite pfundef lastkey_undef. +set j := um_filterk _ _. +have V' : valid j by rewrite pfV. +case E : (unitb j); first by move/unitbP: E=>->; rewrite lastkey0. +have : last_key j \in dom j by case: lastkeyP V' E. +by case/In_dom_umfilt=>v []. +Qed. + +Lemma lastkey_umfilt_dom A (U : natmap A) (h : U) t : + last_key (um_filterk (leq^~ t) h) <= last_key h. +Proof. by apply: lastkey_mono; apply: omf_subdom. Qed. + +Lemma umfilt_le_last A (U : natmap A) (h : U) t : + last_key h <= t -> um_filter (le t) h = h. +Proof. +case V : (valid h); last first. +- by move/invalidE: (negbT V)=>->; rewrite pfundef. +move=>N; rewrite -{2}[h]umfilt_predT; apply/eq_in_umfilt. +by case=>k v /In_dom/dom_lastkey/leq_trans; apply. +Qed. + +Lemma umfilt_last A (U : natmap A) (h : U) : + um_filter (le (last_key h)) h = h. +Proof. by apply: umfilt_le_last. Qed. + +Lemma umfilt_le_fresh A (U : natmap A) (h : U) : + um_filter (le (fresh h)) h = h. +Proof. by apply: umfilt_le_last; apply: ltnW. Qed. + +Lemma umfilt_le0 A (U : natmap A) (h : U) 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 eqxx. +Qed. + +Lemma umfilt_le_split A (U : natmap A) (h : U) t1 t2 : + t1 <= t2 -> + um_filter (le t2) h = + um_filter (le t1) h \+ um_filter (fun '(k, _) => t1 < k <= 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 (U : natmap A) (h : U) t1 t2 k : + t1 <= k <= t2 -> + um_filter (fun '(x, _)=>t1 < x <= t2) h = + um_filter (fun '(x, _)=>t1 < x <= k) h \+ + um_filter (fun '(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 (U : natmap A) (h : U) k v : + (k, v) \In h -> + um_filter (fun '(x, _)=>k.-1 < x <= k) h = pts 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 (U : natmap A) (h1 h2 : U) 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 (U : natmap A) (h1 h2 : U) 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 (U : natmap A) (h : U) 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 (U : natmap A) R a (h1 h2 : U) 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 (U : natmap A) R a (h1 h2 : U) 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 (U : natmap A) R a (h : U) (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 (U : natmap A) R a (h : U) t1 t2 (z0 : R) : + t1 <= t2 -> + eval a (le t2) h z0 = + eval a (fun '(k, _)=>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 /In_dom_umfilt [vx][X _] /In_dom_umfilt [wy][/= /andP][]. + by rewrite /= -ltnNge; move/(leq_ltn_trans X). +- by rewrite -(umfilt_predD h (pts_sub T)) pfV. +by rewrite -!eval_umfilt; apply: eq_in_eval; case=>k v _; apply: ptsD. +Qed. + +Lemma eval_lt_split A (U : natmap A) R a (h : U) t1 t2 (z0 : R) : + t1 < t2 -> + eval a (lt t2) h z0 = + eval a (fun '(k, _)=>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 /In_dom_umfilt [vx][X _] /In_dom_umfilt [wy][/= /andP][]. + by rewrite /= -ltnNge; move/(leq_ltn_trans X). +- by rewrite -(umfilt_predD h (pts_sub_lt T)) pfV. +by rewrite -!eval_umfilt; apply: eq_in_eval; case=>k v _; apply: ptsD_lt. +Qed. + +Lemma eval_le_lt_split A (U : natmap A) R a (h : U) t (z0 : R) : + eval a (le t) h z0 = + eval a (fun '(k, _)=>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 /In_dom_umfilt [vx][X _] /In_dom_umfilt [wy][/= /andP][]. + by rewrite /= -ltnNge; move/(leq_ltn_trans X). +- by rewrite -(umfilt_predD h D) pfV. +rewrite -!eval_umfilt; apply: eq_in_eval; case=>k v _ /=. +by case: ltngtP. +Qed. + +Lemma eval_eq A (U : natmap A) R a (h : U) t v (z0 : R) : + (t, v) \In h -> eval a (fun '(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 (fun '(k, _)=>k == t) h = pts t v by rewrite evalPt /= N. +apply/umem_eq=>[||[k w]]; first by rewrite pfV; case: H. +- by rewrite validPt. +rewrite In_umfiltX; split; last by move/InPt =>[[->->]]. +by move=>[/eqP -> H1]; rewrite (In_fun H H1); apply: In_condPt. +Qed. + +Lemma eval_le_last A (U : natmap A) R a (h : U) 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/dom_lastkey/leq_trans; apply. +Qed. + +Lemma eval_fresh_le A (U : natmap A) R a (h : U) t v (z0 : R) : + eval a (le t) (pts (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 freshPtUnV // V ltnNge H -eval_umfilt. +rewrite joinC evalUnPt; last first. +- by apply/allP=>x; apply: dom_lastkey. +- by rewrite joinC freshPtUnV. +rewrite ltnNge H; congr a; apply: eq_in_eval. +case=>k w /In_dom/dom_lastkey /=. +by case: ltngtP H=>// /ltnW H _ /leq_trans; apply. +Qed. + +Lemma eval_fresh_lt A (U : natmap A) R a (h : U) t v (z0 : R) : + eval a (lt t) (pts (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: dom_lastkey. +- by rewrite joinC valid_fresh. +rewrite ltnNge H; congr a; apply: eq_in_eval. +case=>k w /In_dom/dom_lastkey /=; rewrite /fresh -ltnNge in H. +by case: ltngtP H=>// /ltnW H _ /leq_ltn_trans; apply. +Qed. + +Lemma eval_le_fresh A (U : natmap A) R a (h : U) t v (z0 : R) : + t <= last_key h -> + eval a (le t) (pts (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 (U : natmap A) R a (h : U) t v (z0 : R) : + t <= fresh h -> + eval a (lt t) (pts (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 (U : natmap A) R a (P : R -> Prop) (h : U) 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 (U : natmap A) R a (P : R -> Prop) (h : U) 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 (U : natmap A) R X a (f : R -> X) (h : U) 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 (U : natmap A) R X a (f : R -> X) (h : U) 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 (U : natmap A) p (h : U) t1 t2 : + t1 <= t2 -> + um_count (predI p (le t2)) h = + um_count (predI p (le t1)) h + + um_count (predI p (fun '(k, _)=>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 (U : natmap A) p (h : U) 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 (U : natmap A) p (h1 h2 : U) 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 (U : natmap A) p (h1 h2 : U) 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 (U : natmap A) p (h : U) 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 (U : natmap A) p (h : U) t v : + um_count (predI p (le t)) (pts (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 (U : natmap A) p (h : U) t v : + t <= last_key h -> + um_count (predI p (le t)) (pts (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). + +(*********************************) +(* 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) (U : natmap V). + +Lemma exec0 a (h : U) (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) (U : natmap V) (X : ordType) (a : R -> V -> R) (f : R -> X). +Implicit Types p : pred (nat*V). + +Definition growing (h : U) := + 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 (pts k v \+ h) -> + valid (pts k v \+ h) -> + all (ord k) (dom h) -> + p (k, v) -> + f (evalv a p (pts 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 pts 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 \+ (pts k v \+ h2)) -> + valid (h1 \+ (pts k v \+ h2)) -> + {in dom h1, forall x, x < k} -> + all (ord k) (dom h2) -> + p (k, v) -> + f (evalv a p (h1 \+ (pts 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 ?W // in E1; last first. +- move=>x y /D1 X1; rewrite domPtUn inE (validR 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_umfiltX [_ /G]. +- by rewrite -(umfilt_le_split h N) pfV. +by move=>??/In_dom_umfilt[?][/leq_ltn_trans Y _]/In_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 : U := um_filter (le t1) h. +pose h1 : U := um_filter (fun '(x, _)=>t1 < x <= k.-1) h. +pose h2 : U := um_filter (fun '(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 \+ (pts 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 \+ (pts k v \+ h2))) by rewrite -Eh pfV. +rewrite eval_umfilt (umfilt_le_split h K2) evalUn ?(validAL W1) //; last first. +- by move=>??/In_dom_umfilt[?][/leq_ltn_trans Y] _ /In_dom_umfilt[?][] /andP [/Y]. +rewrite -(eval_umfilt (le t1)); apply: helper2 (validR W1) _ _ _ _ =>//. +- by apply: growR W1 _; rewrite -Eh=>k1 v1 z1 /In_umfiltX [] _ /G. +- by move=>x /In_dom_umfilt; rewrite (subX t1 x) //; case=>v0 [] /andP []. +- by apply/allP=>x /In_dom_umfilt; case=>v0 [] /andP []. +rewrite eval_umfilt Eh evalUn ?(validX W1) -?eval_umfilt // in E. +move=>x y /In_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) =>/In_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) (U : natmap V) (X : ordType) (a : R -> V -> R) (f : R -> nat). +Implicit Types p : pred (nat*V). + +Lemma growIn (h : U) 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 : U) 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. diff --git a/pcm/pcm.v b/pcm/pcm.v new file mode 100644 index 0000000..6cc0717 --- /dev/null +++ b/pcm/pcm.v @@ -0,0 +1,1497 @@ +(* +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 defines pcm -- a type equipped with partial commutative *) +(* monoid structure, several subclasses of PCMs, and important *) +(* concrete instances. *) +(******************************************************************************) + +From HB Require Import structures. +From Coq Require Import ssreflect ssrbool ssrfun. +From mathcomp Require Import ssrnat eqtype seq bigop fintype finset finfun. +From pcm Require Import options axioms prelude seqperm pred seqext. + +Declare Scope pcm_scope. +Delimit Scope pcm_scope with pcm. +Open Scope pcm_scope. + +(*******************************) +(* Partial Commutative Monoids *) +(*******************************) + +Definition pcm_axiom T (valid : T -> bool) (join : T -> T -> T) + (unit : T) (unitb : T -> bool) := + Prod6 (commutative join) + (associative join) + (left_id unit join) + (* monotonicity of valid *) + (forall x y, valid (join x y) -> valid x) + (* unit is valid *) + (valid unit) + (* (as a matter of principle, we make points decidable) *) + (forall x, reflect (x = unit) (unitb x)). + +HB.mixin Record isPCM T := { + valid : T -> bool; + join : T -> T -> T; + Unit : T; + unitb : T -> bool; + pcm_subproof : pcm_axiom valid join Unit unitb}. + +#[short(type="pcm")] +HB.structure Definition PCM := {T of isPCM T}. + +Infix "\+" := join (at level 43, left associativity) : pcm_scope. + +Arguments Unit {s}. +Arguments valid {s} : simpl never. +Prenex Implicits join Unit. + +Lemma joinC {U} : commutative (@join U). Proof. by case: (@pcm_subproof U). Qed. +Lemma joinA {U} : associative (@join U). Proof. by case: (@pcm_subproof U). Qed. +Lemma unitL {U} : left_id Unit (@join U). Proof. by case: (@pcm_subproof U). Qed. +Lemma valid_unit {U} : valid (@Unit U). Proof. by case: (@pcm_subproof U). Qed. +Lemma validL {U : pcm} {x y : U} : valid (x \+ y) -> valid x. +Proof. by case: (@pcm_subproof U) x y. Qed. +Lemma unitbP {U : pcm} {x : U} : reflect (x = Unit) (unitb x). +Proof. by case: (@pcm_subproof U). Qed. + +Section DerivedLemmas. +Variables U V : pcm. + +Lemma joinAC (x y z : U) : x \+ y \+ z = x \+ z \+ y. +Proof. by rewrite -joinA (joinC y) joinA. Qed. + +Lemma joinCA (x y z : U) : x \+ (y \+ z) = y \+ (x \+ z). +Proof. by rewrite joinA (joinC x) -joinA. Qed. + +Lemma validR (x y : U) : valid (x \+ y) -> valid y. +Proof. by rewrite joinC; apply: validL. 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 unitR (x : U) : x \+ Unit = x. +Proof. by rewrite joinC unitL. Qed. + +(* some helper lemmas for easier extraction of validity claims *) +Lemma validAR (x y z : U) : valid (x \+ y \+ z) -> valid (y \+ z). +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 DerivedLemmas. + +Arguments unitL {U}. +Arguments unitR {U}. + +#[export] +Hint Resolve valid_unit : core. + +Section UnfoldingRules. +Variable U : pcm. + +Lemma pcm_joinE (x y : U) : x \+ y = isPCM.join (PCM.class U) x y. +Proof. by []. Qed. + +Lemma pcm_validE (x : U) : valid x = isPCM.valid (PCM.class U) x. +Proof. by []. Qed. + +Lemma pcm_unitE : Unit = isPCM.Unit (PCM.class U). +Proof. by []. Qed. + +Lemma unitb0 : unitb (Unit : U). +Proof. by case: unitbP. Qed. + +Definition pcmE := (pcm_joinE, pcm_validE, pcm_unitE, unitb0). + +(* also a simple rearrangment equation *) +Definition pull (x y : U) := (joinC y x, joinCA y x). + +End UnfoldingRules. + + +(*********************) +(* Cancellative PCMs *) +(*********************) + +(* predicate precision for arbitrary PCM U *) + +Definition precise (U : pcm) (P : U -> Prop) := + forall s1 s2 t1 t2, + valid (s1 \+ t1) -> P s1 -> P s2 -> + s1 \+ t1 = s2 \+ t2 -> s1 = s2. + +Definition cpcm_axiom (U : pcm) := + forall x1 x2 x : U, + valid (x1 \+ x) -> x1 \+ x = x2 \+ x -> x1 = x2. + +HB.mixin Record isCPCM U of PCM U := { + cpcm_subproof : cpcm_axiom U}. + +#[short(type="cpcm")] +HB.structure Definition CPCM := {U of PCM U & isCPCM U}. + +Lemma joinKx (U : cpcm) {x1 x2 x : U} : + valid (x1 \+ x) -> x1 \+ x = x2 \+ x -> x1 = x2. +Proof. by apply: cpcm_subproof. Qed. + + +Section CPCMLemmas. +Variable U : cpcm. + +Lemma joinxK (x x1 x2 : U) : valid (x \+ x1) -> x \+ x1 = x \+ x2 -> x1 = x2. +Proof. by rewrite !(joinC x); apply: joinKx. Qed. + +Lemma cancPL (P : U -> Prop) s1 s2 t1 t2 : + precise P -> valid (s1 \+ t1) -> P s1 -> P s2 -> + s1 \+ t1 = s2 \+ t2 -> (s1 = s2) * (t1 = t2). +Proof. +move=>H V H1 H2 E; move: (H _ _ _ _ V H1 H2 E)=>X. +by rewrite -X in E *; rewrite (joinxK V E). +Qed. + +Lemma cancPR (P : U -> Prop) s1 s2 t1 t2 : + precise P -> valid (s1 \+ t1) -> P t1 -> P t2 -> + s1 \+ t1 = s2 \+ t2 -> (s1 = s2) * (t1 = t2). +Proof. +move=>H V H1 H2 E; rewrite (joinC s1) (joinC s2) in E V. +by rewrite !(cancPL H V H1 H2 E). +Qed. + +End CPCMLemmas. + + +(***************) +(* Topped PCMs *) +(***************) + +(* PCM with an explicit undef value *) +(* we and undef elemeent and function undefb *) +(* to test decidably if an element is undef *) + +(* DEVCOMMENT *) +(* obsoleted conditions *) +(* _ : 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)]; *) +(* /DEVCOMMENT *) + +Definition tpcm_axiom (U : pcm) (undef : U) + (undefb : U -> bool) := + Prod3 (forall x : U, reflect (x = undef) (undefb x)) + (~~ valid undef) + (forall x : U, undef \+ x = undef). + +HB.mixin Record isTPCM U of PCM U := { + undef : U; + undefb : U -> bool; + tpcm_subproof : tpcm_axiom undef undefb}. + +#[short(type="tpcm")] +HB.structure Definition TPCM := {U of PCM U & isTPCM U}. + +Lemma undefbP (U : tpcm) {x : U} : reflect (x = undef) (undefb x). +Proof. by case: (@tpcm_subproof U). Qed. + +Lemma valid_undefN {U} : ~~ valid (@undef U). +Proof. by case: (@tpcm_subproof U). Qed. + +Lemma undef_join (U : tpcm) (x : U) : undef \+ x = undef. +Proof. by case: (@tpcm_subproof U). Qed. + + +Section TPCMLemmas. +Variable U : tpcm. + +Lemma undefbE (f : U) : f = undef <-> undefb f. +Proof. by case: undefbP. Qed. + +Lemma undefb_undef : undefb (undef : U). +Proof. by case: undefbP. Qed. + +Lemma valid_undef : valid (undef : U) = false. +Proof. by rewrite (negbTE valid_undefN). Qed. + +Lemma join_undef (x : U) : x \+ undef = undef. +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. + +Lemma unitb_undef : unitb (undef : U) = false. +Proof. by case: unitbP =>// /undef0. Qed. + +Lemma undefD (x : U) : decidable (x = undef). +Proof. +case D: (undefb x). +- by left; move/undefbP: D. +by right; move/undefbP: D. +Qed. + +End TPCMLemmas. + +Definition tpcmE := (undef_join, join_undef, valid_undef, + unitb_undef, undefb_undef). + + +(***************) +(* Normal TPCM *) +(***************) + +(* TPCM whose only invalid element is undef *) + +Definition normal_tpcm_axiom (U : tpcm) := forall x : U, valid x \/ x = undef. + +HB.mixin Record isNormal_TPCM U of TPCM U := { + normal_tpcm_subproof : normal_tpcm_axiom U}. + +#[short(type="normal_tpcm")] +HB.structure Definition Normal_TPCM := {U of TPCM U & isNormal_TPCM U}. + +(* branching on valid x or x = undef *) +Variant normal_spec {U : normal_tpcm} (x : U) : + bool -> bool -> bool -> Type := +| normal_undef of x = undef : normal_spec x false true false +| normal_valid of valid x : normal_spec x true false (unitb x). + +Lemma normalP {U : normal_tpcm} (x : U) : + normal_spec x (valid x) (undefb x) (unitb x). +Proof. +case: undefbP=>[->|N]. +- by rewrite valid_undef unitb_undef; apply: normal_undef. +have V : (valid x) by case: (normal_tpcm_subproof x) N. +by rewrite V; apply: normal_valid. +Qed. + +Lemma undefNV {U : normal_tpcm} (x : U) : undefb x = ~~ valid x. +Proof. by case: normalP. Qed. + +Lemma undefVN {U : normal_tpcm} (x : U) : valid x = ~~ undefb x. +Proof. by case: normalP. Qed. + +(* branching on valid x or x = undef or x = unit *) +Variant normal0_spec {U : normal_tpcm} (x : U) : + bool -> bool -> bool -> Type := +| normal0_undef of x = undef : normal0_spec x false true false +| normal0_unit of x = Unit : normal0_spec x true false true +| normal0_valid of valid x & x <> Unit : normal0_spec x true false false. + +Lemma normalP0 {U : normal_tpcm} (x : U) : + normal0_spec x (valid x) (undefb x) (unitb x). +Proof. +case: (normalP x)=>[->|V]; first by apply: normal0_undef. +case E: (unitb x); move/unitbP: E; +by [apply: normal0_unit|apply: normal0_valid]. +Qed. + + +(***************************************) +(* PCMs with combination of properties *) +(***************************************) + +(* pcm with decidable equality *) +#[short(type="eqpcm")] +HB.structure Definition EQPCM := {U of Equality U & PCM U}. + +(* cancellative pcm with decidable equality *) +#[short(type="eqcpcm")] +HB.structure Definition EQCPCM := {U of CPCM U & EQPCM U}. + +(* tpcm with decidable equality *) +#[short(type="eqtpcm")] +HB.structure Definition EQTPCM := {U of TPCM U & EQPCM U}. + +(* normal tpcm with decidable equality *) +#[short(type="normal_eqtpcm")] +HB.structure Definition Normal_EQTPCM := {U of Normal_TPCM U & EQPCM U}. + +(* cancellative tpcm *) +#[short(type="ctpcm")] +HB.structure Definition CTPCM := {U of CPCM U & TPCM U}. + +(* cancellative tpcm with decidable equality *) +#[short(type="eqctpcm")] +HB.structure Definition EQCTPCM := {U of Equality U & CTPCM U}. + +(* normal cancellative tpcm *) +#[short(type="normal_ctpcm")] +HB.structure Definition Normal_CTPCM := {U of Normal_TPCM U & CPCM U}. + +(* normal cancellative tpcm with decidable equality *) +#[short(type="normal_eqctpcm")] +HB.structure Definition Normal_EQCTPCM := {U of Normal_TPCM U & EQCPCM U}. + + +(***************************************) +(* Support for big operators over PCMs *) +(***************************************) + +(* U has the laws of commutative monoids from bigop *) +HB.instance Definition _ (U : pcm) := + Monoid.isComLaw.Build U Unit join joinA joinC unitL. + +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). + +Lemma big_pmorph I r (P : pred I) F : + valid (\big[join/Unit]_(i <- r | P i) F i) -> + f (\big[join/Unit]_(i <- r | P i) F i) = + \big[op1/id1]_(i <- r | P i) f (F i). +Proof. +rewrite unlock; elim: r=>[|x r IH] //=; case: ifP=>// H V. +by rewrite f_op // IH //; apply: validR V. +Qed. + +End BigPartialMorph. + + +(*********************) +(* PCM constructions *) +(*********************) + +(* nats with addition are a pcm *) + +(* we use isPCM to build a PCM over a type *) +(* if the type has other characteristics, e.g. equality *) +(* the appropriate structure is built automatically *) +(* e.g., the following builds both the PCM and the EQPCM *) +(* eqpcm instance is automatically inferred *) +Lemma nat_is_pcm : pcm_axiom xpredT addn 0 (eq_op^~ 0). +Proof. by split=>//; [apply:addnC|apply:addnA|apply:(@eqP _^~_)]. Qed. +HB.instance Definition natPCM : isPCM nat := isPCM.Build nat nat_is_pcm. +(* Check (PCM.clone nat _). *) + +(* nats are pcm with multiplication too *) +(* but the instance isn't declared canonical as natPCM already is *) +(* DEVCOMMENT *) +(* To have both, we must redo PCM def so that it keys on join op *) +(* (as in bigops), and not on type. But that is drastic and of unclear *) +(* utility in this setting (e.g., we can't have uniform notation \+). *) +(* /DEVCOMMENT *) +Lemma nat_is_mulpcm : pcm_axiom xpredT mult 1 (eq_op^~ 1). +Proof. by split=>//; [apply:mulnC|apply:mulnA|apply:mul1n|apply:(@eqP _^~_)]. Qed. +HB.instance Definition nat_mulPCM : isPCM nat := isPCM.Build nat nat_is_mulpcm. + +(* nats with max too *) +Lemma nat_is_maxpcm : pcm_axiom xpredT maxn 0 (eq_op^~ 0). +Proof. by split=>//; [apply:maxnC|apply:maxnA|apply:max0n|apply:(@eqP _^~_)]. Qed. +HB.instance Definition nat_maxPCM : isPCM nat := isPCM.Build nat nat_is_maxpcm. + +(* bools with conjunction *) +(* eqpcm automatically inferred *) +Lemma bool_is_pcm : pcm_axiom xpredT andb true (eq_op^~ true). +Proof. by split=>//; [apply:andbC|apply:andbA|apply:(@eqP _^~_)]. Qed. +HB.instance Definition boolPCM : isPCM bool := isPCM.Build bool bool_is_pcm. +(* Check (EQPCM.clone bool _). *) + +(* bools with disjunction *) +(* but the instance isn't declared canonical as boolPCM alread is *) +Lemma bool_is_disjpcm : pcm_axiom xpredT orb false (eq_op^~ false). +Proof. by split=>//; [apply:orbC|apply:orbA|apply:(@eqP _^~_)]. Qed. +HB.instance Definition bool_disjPCM : isPCM bool := + isPCM.Build bool bool_is_disjpcm. + +(* positive nats under max *) +Section PosNatMax. +Definition posNat := sig (fun x => x > 0). + +Definition pos_valid := [fun x : posNat => true]. +Definition pos_unit : posNat := Sub 1 (leqnn 1). +Definition pos_unitb := [fun x : posNat => val x == 1]. +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_join := [fun x y : posNat => + Sub (maxn (val x) (val y)) (max_pos x y) : posNat]. + +Lemma posnat_is_pcm : pcm_axiom pos_valid pos_join pos_unit pos_unitb. +Proof. +split=>[x y|x y z|x|||x] //. +- by apply/eqP; rewrite -val_eqE /= maxnC. +- by apply/eqP; rewrite -val_eqE /= maxnA. +- by apply/eqP; rewrite -val_eqE; apply/eqP/maxn_idPr; case: x. +by apply/eqP. +Qed. + +HB.instance Definition _ : isPCM posNat := isPCM.Build posNat posnat_is_pcm. +(* inherit equality type from sig and nat *) +HB.instance Definition _ := Equality.copy posNat _. +End PosNatMax. + +Arguments pos_unit /. + +(* unit is a PCM, but not a TPCM, as there is no undefined element. *) +(* we'll have to lift with option types for that *) +Section UnitPCM. +Definition unit_valid := [fun x : unit => true]. +Definition unit_join := [fun x y : unit => tt]. +Definition unit_unit := tt. +Definition unit_unitb := [fun x : unit => true]. + +Lemma unit_is_pcm : pcm_axiom unit_valid unit_join unit_unit unit_unitb. +Proof. by split=>//; case=>//; apply/eqP. Qed. + +HB.instance Definition _ : isPCM unit := isPCM.Build unit unit_is_pcm. +End UnitPCM. + +Arguments unit_unit /. + +(***********************) +(* Option PCM and TPCM *) +(***********************) + +Section OptionTPCM. +Variables U : pcm. + +Definition ovalid := [fun x : option U => + if x is Some a then valid a else false]. +Definition ojoin := [fun x y : option U => + if x is Some a then + if y is Some b then Some (a \+ b) else None + else None]. +Definition ounit : option U := Some Unit. +Definition ounitb := [fun x : option U => + if x is Some v then unitb v else false]. +Definition oundef : option U := None. +Definition oundefb := [fun x : option U => + if x is Some a then false else true]. + +Lemma option_is_pcm : pcm_axiom ovalid ojoin ounit ounitb. +Proof. +split=>[x y|x y z|x|x y||x]. +- by case: x; case: y=>//=b a; rewrite joinC. +- by case: x; case: y; case: z=>//=c b a; rewrite joinA. +- by case: x=>//=a; rewrite unitL. +- by case x=>//=a; case: y=>//=b /validL. +- by rewrite /= valid_unit. +case: x=>[a|] /=; last by constructor. +by case: unitbP=>[->|H]; constructor=>//; case. +Qed. + +HB.instance Definition _ : isPCM (option U) := + isPCM.Build (option U) option_is_pcm. + +Lemma option_is_tpcm : tpcm_axiom oundef oundefb. +Proof. by split=>//; case=>[a|]; constructor. Qed. + +HB.instance Definition _ : isTPCM (option U) := + isTPCM.Build (option U) option_is_tpcm. +End OptionTPCM. + +Arguments ounit /. +Arguments oundef /. + +(* option preserves equality *) +HB.instance Definition _ (U : eqpcm) := PCM.copy (option U) _. +(* Check (EQPCM.clone (option nat) _). *) + +(* option preserves cancellativity *) +Lemma option_is_cpcm (U : cpcm) : cpcm_axiom (option U). +Proof. by case=>[x|][y|][z|] // V [] /(joinKx V)=>->. Qed. +HB.instance Definition _ (U : cpcm) := + isCPCM.Build (option U) (@option_is_cpcm U). + +(* option is normal if U is all valid *) +Lemma option_is_normal (U : pcm) : + @valid U =1 xpredT -> normal_tpcm_axiom (option U). +Proof. by move=>E [x|]; [left; rewrite /valid /= E|right]. Qed. + +(* case analysis infrastructure *) +CoInductive option_pcm_spec (A : pcm) (x y : option A) : + option A -> Type := +| some_case x' y' of x = some x' & y = some y' : + option_pcm_spec x y (some (x' \+ y')) +| none_case of ~~ valid (x \+ y) : + option_pcm_spec x y None. + +Lemma optionP (A : pcm) (x y : option A) : option_pcm_spec x y (x \+ y). +Proof. +case: x=>[x|]; case: y=>[y|]=>/=; first by [apply: some_case]; +by apply: none_case. +Qed. + +(*****************************) +(* Cartesian product of PCMs *) +(*****************************) + +Section ProdPCM. +Variables U V : pcm. +Local Notation tp := (U * V)%type. + +Definition valid2 := [fun x : tp => valid x.1 && valid x.2]. +Definition join2 := [fun x1 x2 : tp => (x1.1 \+ x2.1, x1.2 \+ x2.2)]. +Definition unit2 : tp := (Unit, Unit). +Definition unitb2 := [fun x : U * V => unitb x.1 && unitb x.2]. + +Lemma prod_is_pcm : pcm_axiom valid2 join2 unit2 unitb2. +Proof. +split=>[x y|x y z|x|x y||x]. +- by rewrite /= (joinC x.1) (joinC x.2). +- by rewrite /= !joinA. +- by rewrite /= !unitL -prod_eta. +- by rewrite /valid2 /=; case/andP=>/validL -> /validL ->. +- by rewrite /valid2 /= !valid_unit. +case: x=>x1 x2 /=; case: andP=>H; constructor. +- by case: H=>/unitbP -> /unitbP ->. +by case=>X1 X2; elim: H; rewrite X1 X2 !pcmE. +Qed. + +HB.instance Definition _ : isPCM (U * V)%type := + isPCM.Build (U * V)%type prod_is_pcm. +End ProdPCM. + +(* explicitly extend to eqpcms *) +HB.instance Definition _ (U V : eqpcm) := PCM.copy (U * V)%type _. + +Arguments unit2 /. + +(* product simplification *) +Section Simplification. +Variable U V : pcm. + +Lemma pcmPJ (x1 y1 : U) (x2 y2 : V) : + (x1, x2) \+ (y1, y2) = (x1 \+ y1, x2 \+ y2). +Proof. by rewrite pcmE. 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 : U * V) : valid x = valid x.1 && valid x.2. +Proof. by []. Qed. + +Lemma pcmPU : Unit = (Unit, Unit) :> U * V. +Proof. by rewrite pcmE. Qed. + +Definition pcmPE := (pcmPJ, pcmFJ, pcmSJ, pcmPV, pcmPU). + +End Simplification. + +(* We often iterate PCM-products so *) +(* we provide primitives for small numbers *) + +Section Prod3PCM. +Variables U1 U2 U3 : pcm. +Notation tp := (Prod3 U1 U2 U3). +Definition valid3 := [fun x : tp => + [&& valid (proj31 x), + valid (proj32 x) & + valid (proj33 x)]]. +Definition join3 := [fun x y : tp => + mk3 (proj31 x \+ proj31 y) + (proj32 x \+ proj32 y) + (proj33 x \+ proj33 y)]. +Definition unit3 : tp := mk3 Unit Unit Unit. +Definition unitb3 := [fun x : tp => + [&& unitb (proj31 x), + unitb (proj32 x) & + unitb (proj33 x)]]. + +Lemma prod3_is_pcm : pcm_axiom valid3 join3 unit3 unitb3. +Proof. +split=>[x y|x y z||x y||x] /=. +- by congr mk3; rewrite joinC. +- by congr mk3; rewrite joinA. +- by case=>*; rewrite /= !unitL. +- by case/and3P; do ![move/validL=>->]. +- by rewrite !valid_unit. +case: x=>x1 x2 x3 /=. +do ![case: unitbP=>[->|H]; last by constructor; case]; +by constructor. +Qed. + +HB.instance Definition _ : isPCM (Prod3 U1 U2 U3) := + isPCM.Build (Prod3 U1 U2 U3) prod3_is_pcm. +End Prod3PCM. + +HB.instance Definition _ (U1 U2 U3 : eqpcm) := + PCM.copy (Prod3 U1 U2 U3) _. + +Arguments unit3 /. + +Section Prod4PCM. +Variables U1 U2 U3 U4 : pcm. +Notation tp := (Prod4 U1 U2 U3 U4). +Definition valid4 := [fun x : tp => + [&& valid (proj41 x), + valid (proj42 x), + valid (proj43 x) & + valid (proj44 x)]]. +Definition join4 := [fun x y : tp => + mk4 (proj41 x \+ proj41 y) + (proj42 x \+ proj42 y) + (proj43 x \+ proj43 y) + (proj44 x \+ proj44 y)]. +Definition unit4 : tp := mk4 Unit Unit Unit Unit. +Definition unitb4 := [fun x : tp => + [&& unitb (proj41 x), + unitb (proj42 x), + unitb (proj43 x) & + unitb (proj44 x)]]. + +Lemma prod4_is_pcm : pcm_axiom valid4 join4 unit4 unitb4. +Proof. +split=>[x y|x y z||x y||x] /=. +- by congr mk4; rewrite joinC. +- by congr mk4; rewrite joinA. +- by case=>*; rewrite /= !unitL. +- by case/and4P; do ![move/validL=>->]. +- by rewrite !valid_unit. +case: x=>x1 x2 x3 x4 /=. +do ![case: unitbP=>[->|H]; last by constructor; case]; +by constructor. +Qed. + +HB.instance Definition _ : isPCM (Prod4 U1 U2 U3 U4) := + isPCM.Build (Prod4 U1 U2 U3 U4) prod4_is_pcm. +End Prod4PCM. + +HB.instance Definition _ (U1 U2 U3 U4 : eqpcm) := + PCM.copy (Prod4 U1 U2 U3 U4) _. + +Arguments unit4 /. + +Section Prod5PCM. +Variables U1 U2 U3 U4 U5 : pcm. +Notation tp := (Prod5 U1 U2 U3 U4 U5). + +Definition valid5 := [fun x : tp => + [&& valid (proj51 x), + valid (proj52 x), + valid (proj53 x), + valid (proj54 x) & + valid (proj55 x)]]. +Definition join5 := [fun x y : tp => + mk5 (proj51 x \+ proj51 y) + (proj52 x \+ proj52 y) + (proj53 x \+ proj53 y) + (proj54 x \+ proj54 y) + (proj55 x \+ proj55 y)]. +Definition unit5 : tp := mk5 Unit Unit Unit Unit Unit. +Definition unitb5 := [fun x : tp => + [&& unitb (proj51 x), + unitb (proj52 x), + unitb (proj53 x), + unitb (proj54 x) & + unitb (proj55 x)]]. + +Lemma prod5_is_pcm : pcm_axiom valid5 join5 unit5 unitb5. +Proof. +split=>[x y|x y z||x y||x] /=. +- by congr mk5; rewrite joinC. +- by congr mk5; rewrite joinA. +- by case=>*; rewrite /= !unitL. +- by case/and5P; do ![move/validL=>->]. +- by rewrite !valid_unit. +case: x=>x1 x2 x3 x4 x5 /=. +do ![case: unitbP=>[->|H]; last by constructor; case]; +by constructor. +Qed. + +HB.instance Definition _ : isPCM (Prod5 U1 U2 U3 U4 U5) := + isPCM.Build (Prod5 U1 U2 U3 U4 U5) prod5_is_pcm. +End Prod5PCM. + +HB.instance Definition _ (U1 U2 U3 U4 U5 : eqpcm) := + PCM.copy (Prod5 U1 U2 U3 U4 U5) _. + +Arguments unit5 /. + +Section Prod6PCM. +Variables U1 U2 U3 U4 U5 U6 : pcm. +Notation tp := (Prod6 U1 U2 U3 U4 U5 U6). + +Definition valid6 := [fun x : tp => + [&& valid (proj61 x), + valid (proj62 x), + valid (proj63 x), + valid (proj64 x), + valid (proj65 x) & + valid (proj66 x)]]. +Definition join6 := [fun x y : tp => + mk6 (proj61 x \+ proj61 y) + (proj62 x \+ proj62 y) + (proj63 x \+ proj63 y) + (proj64 x \+ proj64 y) + (proj65 x \+ proj65 y) + (proj66 x \+ proj66 y)]. +Definition unit6 : tp := mk6 Unit Unit Unit Unit Unit Unit. +Definition unitb6 := [fun x : tp => + [&& unitb (proj61 x), + unitb (proj62 x), + unitb (proj63 x), + unitb (proj64 x), + unitb (proj65 x) & + unitb (proj66 x)]]. + +Lemma prod6_is_pcm : pcm_axiom valid6 join6 unit6 unitb6. +Proof. +split=>[x y|x y z||x y||x] /=. +- by congr mk6; rewrite joinC. +- by congr mk6; rewrite joinA. +- by case=>*; rewrite /= !unitL. +- by case/and6P; do ![move/validL=>->]. +- by rewrite !valid_unit. +case: x=>x1 x2 x3 x4 x5 x6 /=. +do ![case: unitbP=>[->|H]; last by constructor; case]; +by constructor. +Qed. + +HB.instance Definition _ : isPCM (Prod6 U1 U2 U3 U4 U5 U6) := + isPCM.Build (Prod6 U1 U2 U3 U4 U5 U6) prod6_is_pcm. +End Prod6PCM. + +HB.instance Definition _ (U1 U2 U3 U4 U5 U6 : eqpcm) := + PCM.copy (Prod6 U1 U2 U3 U4 U5 U6) _. + +Arguments unit6 /. + +Section Prod7PCM. +Variables U1 U2 U3 U4 U5 U6 U7 : pcm. +Notation tp := (Prod7 U1 U2 U3 U4 U5 U6 U7). + +Definition valid7 := [fun x : tp => + [&& valid (proj71 x), + valid (proj72 x), + valid (proj73 x), + valid (proj74 x), + valid (proj75 x), + valid (proj76 x) & + valid (proj77 x)]]. +Definition join7 := [fun x y : tp => + mk7 (proj71 x \+ proj71 y) + (proj72 x \+ proj72 y) + (proj73 x \+ proj73 y) + (proj74 x \+ proj74 y) + (proj75 x \+ proj75 y) + (proj76 x \+ proj76 y) + (proj77 x \+ proj77 y)]. +Definition unit7 : tp := mk7 Unit Unit Unit Unit Unit Unit Unit. +Definition unitb7 := [fun x : tp => + [&& unitb (proj71 x), + unitb (proj72 x), + unitb (proj73 x), + unitb (proj74 x), + unitb (proj75 x), + unitb (proj76 x) & + unitb (proj77 x)]]. + +Lemma prod7_is_pcm : pcm_axiom valid7 join7 unit7 unitb7. +Proof. +split=>[x y|x y z||x y||x] /=. +- by congr mk7; rewrite joinC. +- by congr mk7; rewrite joinA. +- by case=>*; rewrite /= !unitL. +- by case/and7P; do ![move/validL=>->]. +- by rewrite !valid_unit. +case: x=>x1 x2 x3 x4 x5 x6 x7 /=. +do ![case: unitbP=>[->|H]; last by constructor; case]; +by constructor. +Qed. + +HB.instance Definition _ : isPCM (Prod7 U1 U2 U3 U4 U5 U6 U7) := + isPCM.Build (Prod7 U1 U2 U3 U4 U5 U6 U7) prod7_is_pcm. +End Prod7PCM. + +HB.instance Definition _ (U1 U2 U3 U4 U5 U6 U7 : eqpcm) := + PCM.copy (Prod7 U1 U2 U3 U4 U5 U6 U7) _. + +Arguments unit7 /. + +(* Finite products of PCMs as functions *) +Section FunPCM. +Variables (T : finType) (Us : T -> pcm). +Notation tp := (forall t, Us t). + +Definition fun_valid := [fun f : tp => [forall t, valid (f t)]]. +Definition fun_join := [fun f1 f2 : tp => fun t => f1 t \+ f2 t]. +Definition fun_unit : tp := fun t => Unit. +Definition fun_unitb := [fun f : tp => [forall t, unitb (f t)]]. + +Lemma fun_is_pcm : pcm_axiom fun_valid fun_join fun_unit fun_unitb. +Proof. +split=>[f g|f g h|f|f g||f] /=. +- by apply: fext=>t; rewrite joinC. +- by apply: fext=>t; rewrite joinA. +- by apply: fext=>t; rewrite unitL. +- by move/forallP=>V; apply/forallP=>t; apply: validL (V t). +- by apply/forallP=>t; apply: valid_unit. +case: forallP=>H; constructor. +- by apply: fext=>t; apply/unitbP; apply: H. +by move=>H1; elim: H=>t; apply/unitbP; rewrite H1. +Qed. + +HB.instance Definition _ : isPCM (forall t, Us t) := + isPCM.Build (forall t, Us t) fun_is_pcm. +End FunPCM. + +Arguments fun_unit /. + +(* we won't use fun types for any examples *) +(* so we don't need equality structure on them *) + + +(* Finite products of PCMs as finfuns *) +(* We will work with this second structure, though *) +(* the one above is needed to state that sel and finfun are morphisms *) +(* (between FunPCM and FinPCM) *) +(* dffun used for inheritance (see finfun.v) *) +Section FinPCM. +Variables (T : finType) (Us : T -> pcm). +Notation tp := {dffun forall t, Us t}. + +Definition fin_valid := [fun f : tp => [forall t, valid (sel t f)]]. +Definition fin_join := [fun f g : tp => [ffun t => sel t f \+ sel t g]]. +Definition fin_unit : tp := [ffun => Unit]. +Definition fin_unitb := [fun f : tp => [forall t, unitb (sel t f)]]. + +Lemma ffinprod_is_pcm : pcm_axiom fin_valid fin_join fin_unit fin_unitb. +Proof. +split=>[x y|x y z|x|x y||x] /=. +- by apply/ffunP=>t; rewrite !ffunE joinC. +- by apply/ffunP=>t; rewrite /sel !ffunE joinA. +- by apply/ffunP=>t; rewrite /sel !ffunE unitL. +- move/forallP=>H; apply/forallP=>t; move: (H t). + by rewrite /sel !ffunE=>/validL. +- by apply/forallP=>t; rewrite /sel ffunE valid_unit. +case H: [forall t, _]; constructor. +- move/forallP: H=>H; apply/ffinP=>t. + by rewrite sel_fin; move/unitbP: (H t). +move=>E; move/negP: H; apply; apply/forallP=>t. +by rewrite E sel_fin unitb0. +Qed. + +HB.instance Definition _ := + isPCM.Build {dffun forall t, Us t} ffinprod_is_pcm. +End FinPCM. + +HB.instance Definition _ (T : finType) (Us : T -> eqpcm) := + PCM.copy {dffun forall t, Us t} _. + +Arguments fin_unit /. + +(* product of TPCMs is a TPCM *) + +(* With TPCMs we could remove the pairs where *) +(* one element is valid and the other is not. *) +(* We will do that later with normalization procedure. *) +(* That will give us a new construction for normalized products *) +(* but we first need to introduce PCM morphisms *) +(* Here, we give non-normalized ones. *) +Section ProdTPCM. +Variables U V : tpcm. +Definition undef2 := (@undef U, @undef V). +Definition undefb2 := [fun x : U * V => undefb x.1 && undefb x.2]. + +Lemma prod_is_tpcm : tpcm_axiom undef2 undefb2. +Proof. +split=>/=. +- case=>x1 x2; case: andP=>/= H; constructor. + - by case: H=>/undefbP -> /undefbP ->. + by case=>X1 X2; elim: H; rewrite X1 X2 !tpcmE. +- by rewrite /valid /= !valid_undef. +by case=>x1 x2; rewrite pcmPJ !undef_join. +Qed. + +HB.instance Definition _ := isTPCM.Build (U * V)%type prod_is_tpcm. +End ProdTPCM. + +Arguments undef2 /. + +(* iterated TPCMs *) + +Section Prod3TPCM. +Variables U1 U2 U3 : tpcm. +Notation tp := (Prod3 U1 U2 U3). +Definition undef3 : tp := mk3 undef undef undef. +Definition undefb3 := [fun x : tp => + [&& undefb (proj31 x), + undefb (proj32 x) & + undefb (proj33 x)]]. + +Lemma prod3_is_tpcm : tpcm_axiom undef3 undefb3. +Proof. +split=>[x||x]. +- case: x=>x1 x2 x3 /=. + do ![case: undefbP=>[->|H]; last by constructor; case]. + by constructor. +- by rewrite /valid /= !valid_undef. +by rewrite pcmE /= !undef_join. +Qed. + +HB.instance Definition _ : isTPCM (Prod3 U1 U2 U3) := + isTPCM.Build (Prod3 U1 U2 U3) prod3_is_tpcm. +End Prod3TPCM. + +Arguments undef3 /. + +Section Prod4TPCM. +Variables U1 U2 U3 U4 : tpcm. +Notation tp := (Prod4 U1 U2 U3 U4). +Definition undef4 : tp := mk4 undef undef undef undef. +Definition undefb4 := [fun x : tp => + [&& undefb (proj41 x), + undefb (proj42 x), + undefb (proj43 x) & + undefb (proj44 x)]]. + +Lemma prod4_is_tpcm : tpcm_axiom undef4 undefb4. +Proof. +split=>[x||x] /=. +- case: x=>x1 x2 x3 x4 /=. + do ![case: undefbP=>[->|H]; last by constructor; case]. + by constructor. +- by rewrite /valid /= !valid_undef. +by rewrite pcmE /= !undef_join. +Qed. + +HB.instance Definition _ : isTPCM (Prod4 U1 U2 U3 U4) := + isTPCM.Build (Prod4 U1 U2 U3 U4) prod4_is_tpcm. +End Prod4TPCM. + +Arguments undef4 /. + +Section Prod5TPCM. +Variables U1 U2 U3 U4 U5 : tpcm. +Notation tp := (Prod5 U1 U2 U3 U4 U5). +Definition undef5 : tp := mk5 undef undef undef undef undef. +Definition undefb5 := [fun x : tp => + [&& undefb (proj51 x), + undefb (proj52 x), + undefb (proj53 x), + undefb (proj54 x) & + undefb (proj55 x)]]. + +Lemma prod5_is_tpcm : tpcm_axiom undef5 undefb5. +Proof. +split=>[x||x] /=. +- case: x=>x1 x2 x3 x4 x5 /=. + do ![case: undefbP=>[->|H]; last by constructor; case]. + by constructor. +- by rewrite /valid /= !valid_undef. +by rewrite pcmE /= !undef_join. +Qed. + +HB.instance Definition _ : isTPCM (Prod5 U1 U2 U3 U4 U5) := + isTPCM.Build (Prod5 U1 U2 U3 U4 U5) prod5_is_tpcm. +End Prod5TPCM. + +Arguments undef5 /. + + +Section Prod6TPCM. +Variables U1 U2 U3 U4 U5 U6 : tpcm. +Notation tp := (Prod6 U1 U2 U3 U4 U5 U6). +Definition undef6 : tp := mk6 undef undef undef undef undef undef. +Definition undefb6 := [fun x : tp => + [&& undefb (proj61 x), + undefb (proj62 x), + undefb (proj63 x), + undefb (proj64 x), + undefb (proj65 x) & + undefb (proj66 x)]]. + +Lemma prod6_is_tpcm : tpcm_axiom undef6 undefb6. +Proof. +split=>[x||x] /=. +- case: x=>x1 x2 x3 x4 x5 x6 /=. + do ![case: undefbP=>[->|H]; last by constructor; case]. + by constructor. +- by rewrite /valid /= !valid_undef. +by rewrite pcmE /= !undef_join. +Qed. + +HB.instance Definition _ : isTPCM (Prod6 U1 U2 U3 U4 U5 U6) := + isTPCM.Build (Prod6 U1 U2 U3 U4 U5 U6) prod6_is_tpcm. +End Prod6TPCM. + +Arguments undef6 /. + +Section Prod7TPCM. +Variables U1 U2 U3 U4 U5 U6 U7 : tpcm. +Notation tp := (Prod7 U1 U2 U3 U4 U5 U6 U7). +Definition undef7 : tp := mk7 undef undef undef undef undef undef undef. +Definition undefb7 := [fun x : tp => + [&& undefb (proj71 x), + undefb (proj72 x), + undefb (proj73 x), + undefb (proj74 x), + undefb (proj75 x), + undefb (proj76 x) & + undefb (proj77 x)]]. + +Lemma prod7_is_tpcm : tpcm_axiom undef7 undefb7. +Proof. +split=>[x||x]. +- case: x=>x1 x2 x3 x4 x5 x6 x7 /=. + do ![case: undefbP=>[->|H]; last by constructor; case]. + by constructor. +- by rewrite /valid /= !valid_undef. +by rewrite pcmE /= !undef_join. +Qed. + +HB.instance Definition _ : isTPCM (Prod7 U1 U2 U3 U4 U5 U6 U7) := + isTPCM.Build (Prod7 U1 U2 U3 U4 U5 U6 U7) prod7_is_tpcm. +End Prod7TPCM. + +Arguments undef7 /. + +(* TPCM proofs use function extensionality *) +(* it's TPCM only if T inhabited finite type *) +(* (otherwise valid undef) *) + +Definition fun_undef T (Us : T -> tpcm) : forall t, Us t + := fun t => undef. +Arguments fun_undef {T Us} /. + +Definition fun_undefb (T : finType) (Us : T -> tpcm) := + fun f : forall t, Us t => [forall t, undefb (f t)]. +Arguments fun_undefb {T Us} f /. + +Section FunTPCM. +Variables (T : ifinType) (Us : T -> tpcm). +Notation tp := (forall t, Us t). + +Lemma fun_is_tpcm : tpcm_axiom fun_undef (fun_undefb (Us:=Us)). +Proof. +split=>[f||f] /=. +- case: forallP=>H; constructor. + - by apply: fext=>t; apply/undefbP; apply: H. + by move=>H1; elim: H=>t; apply/undefbP; rewrite H1. +- by apply/negP=>/forallP/(_ inhab); rewrite valid_undef. +by apply: fext=>t; rewrite /join /= undef_join. +Qed. + +HB.instance Definition _ : isTPCM (forall t, Us t) := + isTPCM.Build (forall t, Us t) fun_is_tpcm. + +End FunTPCM. + + +(* TPCM only if T inhabited finite type *) +(* (otherwise valid undef) *) +Definition fin_undef (T : finType) (Us : T -> tpcm) + : {dffun forall t, Us t} := [ffun t => undef]. +Arguments fin_undef {T Us} /. + +Definition fin_undefb (T : finType) (Us : T -> tpcm) + (x : {dffun forall t, Us t}) := + [forall t, undefb (sel t x)]. +Arguments fin_undefb {T Us} x /. + +Section FinTPCM. +Variables (T : ifinType) (Us : T -> tpcm). +Notation tp := {dffun forall t, Us t}. + +Lemma finprod_is_tpcm : tpcm_axiom fin_undef (fin_undefb (Us:=Us)). +Proof. +split=>[x||x] /=. +- case H : [forall t, _]; constructor. + - move/forallP: H=>H; apply/ffinP=>t. + by rewrite sel_fin; move/undefbP: (H t). + move=>E; move/negP: H; apply; apply/forallP=>t. + by rewrite E sel_fin undefb_undef. +- apply/negP=>/forallP/(_ inhab). + by rewrite sel_fin valid_undef. +by apply/ffinP=>t; rewrite !sel_fin undef_join. +Qed. + +HB.instance Definition _ : isTPCM tp := + isTPCM.Build tp finprod_is_tpcm. +End FinTPCM. + +(*************************) +(* PCM-induced pre-order *) +(*************************) + +Definition pcm_preord (U : pcm) (x y : U) := exists z, y = x \+ z. + +Notation "[ 'pcm' x '<=' y ]" := (@pcm_preord _ x y) + (at level 0, x, y at level 69, + format "[ '[hv' 'pcm' x '/ ' <= y ']' ]") : type_scope. + +Section PleqLemmas. +Variable U : pcm. +Implicit Types x y z : U. + +Lemma pleq_unit x : [pcm Unit <= x]. +Proof. by exists x; rewrite unitL. Qed. + +(* preorder lemmas *) + +(* 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]. +Proof. by exists Unit; rewrite unitR. Qed. + +Lemma pleq_trans x y z : [pcm x <= y] -> [pcm y <= z] -> [pcm x <= z]. +Proof. by case=>a -> [b ->]; exists (a \+ b); rewrite joinA. Qed. + +(* monotonicity lemmas *) + +Lemma pleq_join2r x x1 x2 : [pcm x1 <= x2] -> [pcm x1 \+ x <= x2 \+ x]. +Proof. by case=>a ->; exists a; rewrite joinAC. Qed. + +Lemma pleq_join2l x x1 x2 : [pcm x1 <= x2] -> [pcm x \+ x1 <= x \+ x2]. +Proof. by rewrite !(joinC x); apply: pleq_join2r. Qed. + +Lemma pleq_joinr {x1 x2} : [pcm x1 <= x1 \+ x2]. +Proof. by exists x2. Qed. + +Lemma pleq_joinl {x1 x2} : [pcm x2 <= x1 \+ x2]. +Proof. by rewrite joinC; apply: pleq_joinr. Qed. + +(* validity lemmas *) + +Lemma pleqV (x1 x2 : U) : [pcm x1 <= x2] -> valid x2 -> valid x1. +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 -> V; rewrite (validRE3 V). Qed. + +Lemma pleq_validR (x x1 x2 : U) : + [pcm x1 <= x2] -> valid (x2 \+ x) -> valid (x1 \+ x). +Proof. by rewrite -!(joinC x); apply: pleq_validL. Qed. + +(* the next two lemmas only hold for cancellative PCMs *) + +Lemma pleq_joinxK (V : cpcm) (x x1 x2 : V) : + valid (x2 \+ x) -> [pcm x1 \+ x <= x2 \+ x] -> [pcm x1 <= x2]. +Proof. by move=>W [a]; rewrite joinAC=>/(joinKx W) ->; exists a. Qed. + +Lemma pleq_joinKx (V : cpcm) (x x1 x2 : V) : + valid (x \+ x2) -> [pcm x \+ x1 <= x \+ x2] -> [pcm x1 <= x2]. +Proof. by rewrite !(joinC x); apply: pleq_joinxK. Qed. + +End PleqLemmas. + +#[export] +Hint Resolve pleq_unit pleq_refl pleq_joinr pleq_joinl : core. +Prenex Implicits pleq_refl pleq_joinl pleq_joinr. + +(* shorter names *) +Notation pcmR := pleq_refl. +Notation pcmS := pleq_joinr. +Notation pcmO := pleq_joinl. + +Lemma pleq_undef (U : tpcm) (x : U) : [pcm x <= undef]. +Proof. by exists undef; rewrite join_undef. Qed. + +#[export] +Hint Resolve pleq_undef : core. + +(********************) +(* 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. + +Section BigOps. +Variables (U : pcm). +Variables (I : Type) (f : I -> U). + +Lemma big_validV (xs : seq I) i : + valid (\big[join/Unit]_(i <- xs) f i) -> + i \In xs -> valid (f i). +Proof. +elim: xs=>[|x xs IH] in i * => //. +rewrite big_cons InE=>V [->|]; first by apply: (validL V). +by apply: IH; rewrite (validR V). +Qed. + +Lemma big_validVL (xs : seq I) z i : + valid (f z \+ \big[join/Unit]_(i <- xs) f i) -> + i \In xs -> i <> z -> valid (f z \+ f i). +Proof. +elim: xs=>[|x xs IH] in i * => //. +rewrite big_cons InE =>V [->_ |]. +- by move: V; rewrite joinA; apply: validL. +by apply: IH; move: V; rewrite joinCA; apply: validR. +Qed. + +Lemma big_validV2 (xs : seq I) : + valid (\big[join/Unit]_(i <- xs) f i) -> + forall i j, i \In xs -> j \In xs -> i <> j -> valid (f i \+ f j). +Proof. +elim: xs=>[|x xs IH] //=; rewrite big_cons. +move=>V i j; rewrite !InE; case=>[->|Xi][->|Xj] N //; last first. +- by apply: IH=>//; apply: (validR V). +- by rewrite joinC; apply: (big_validVL V). +by apply: (big_validVL V)=>// /esym. +Qed. + +End BigOps. + +(***********************************) +(* separating conjunction aka star *) +(***********************************) + +Section Star. +Variable U : pcm. + +Definition star (p1 p2 : Pred U) : Pred U := + [Pred h | exists h1 h2, [ /\ h = h1 \+ h2, h1 \In p1 & h2 \In p2] ]. +Definition emp : Pred U := eq^~ Unit. +Definition top : Pred U := PredT. + +End Star. + +Arguments emp {U}. +Arguments top {U}. + +Notation "p1 '#' p2" := (star p1 p2) + (at level 57, right associativity) : rel_scope. + +(* iterated star *) + +Module IterStar. +Section IterStar. +Variables (U : pcm) (A : Type). + +Definition seqjoin (s : seq U) : U := + \big[join/Unit]_(i <- s) i. + +Definition sepit (s : seq A) (f : A -> Pred U) : Pred U := + [Pred h | exists hs : seq U, + [ /\ size hs = size s, h = seqjoin hs & + All (fun '(a, h) => h \In f a) (zip s hs)]]. + +Lemma sepit0 f : sepit [::] f =p emp. +Proof. +move=>h; split. +- by case=>/= hs [/size0nil -> -> _]; rewrite /seqjoin !big_nil. +by move=>->; exists [::]; rewrite /seqjoin !big_nil. +Qed. + +Lemma sepit_cons x s f : sepit (x::s) f =p f x # sepit s f. +Proof. +move=>h; split. +- case=>/=; case=>[|h0 hs]; case=>//= /eqP; rewrite eqSS =>/eqP Hs. + rewrite /seqjoin big_cons =>->[H0 H1]. + by exists h0, (seqjoin hs); do!split=>//; exists hs. +case=>h1[_][{h}-> H1][hs][E -> H]. +by exists (h1 :: hs); rewrite /= E /seqjoin !big_cons. +Qed. + +Lemma sepit_cat s1 s2 f : sepit (s1 ++ s2) f =p sepit s1 f # sepit s2 f. +Proof. +elim: s1 s2=>[|x s1 IH] s2 h /=; split. +- case=>hs [E {h}-> H2]. + exists Unit, (seqjoin hs); rewrite unitL. + by split=>//; [rewrite sepit0 | exists hs]. +- by case=>h1[h2][{h}->]; rewrite sepit0=>->; rewrite unitL. +- case=>/=; case=>[|h0 hs]; case=>//= /eqP; rewrite eqSS=>/eqP E. + rewrite /seqjoin !big_cons /= =>->[H0 HS]. + case: (IH s2 (seqjoin hs)).1; first by exists hs. + move=>h1[h2][HJ H1 H2]; rewrite /seqjoin in HJ. + exists (h0 \+ h1), h2; rewrite HJ joinA; split=>//. + by rewrite sepit_cons; exists h0, h1. +case=>h1[h2][{h}->[]]; case=>[|h0 hs1]; case=>//= /eqP; rewrite eqSS=>/eqP E1. +rewrite /seqjoin !big_cons /= =>{h1}->[H0 H1]; case=>hs2[E2 {h2}-> H2]. +exists (h0 :: hs1 ++ hs2); rewrite /seqjoin big_cons big_cat joinA; split=>//=. +- by rewrite !size_cat E1 E2. +rewrite zip_cat //=; split=>//. +by apply/All_cat. +Qed. + +Lemma sepit_perm s1 s2 (f : A -> Pred U) : + perm s1 s2 -> sepit s1 f =p sepit s2 f. +Proof. +elim: s1 s2 =>[|x s1 IH] s2 /=; first by move/pperm_nil=>->. +move=>H; have Hx: x \In s2 by apply/(pperm_in H)/In_cons; left. +case: (In_split Hx)=>s21[s22] E; rewrite {s2 Hx}E in H *. +move/pperm_cons_cat_cons: H=>/IH {}IH. +rewrite sepit_cons sepit_cat /= =>h; split. +- case=>h1[h2][{h}-> H1]; rewrite IH sepit_cat. + case=>_[_][{h2}-> [hs3][E3 -> H3] [hs4][E4 -> H4]]; rewrite joinCA. + exists (seqjoin hs3), (h1 \+ seqjoin hs4); split=>//; first by exists hs3. + by rewrite sepit_cons; exists h1, (seqjoin hs4); split=>//; exists hs4. +case=>_[h2][{h}-> [hs1][Hs1 -> H1]]; rewrite sepit_cons. +case=>h3[_][{h2}-> H3 [hs2][Hs2 -> H2]]; rewrite joinCA. +exists h3, (seqjoin hs1 \+ seqjoin hs2); split=>//. +rewrite IH; exists (hs1 ++ hs2); split. +- by rewrite !size_cat Hs1 Hs2. +- by rewrite /seqjoin big_cat. +by rewrite zip_cat //; apply/All_cat. +Qed. + +Lemma sepitF s (f1 f2 : A -> Pred U) : + (forall x, x \In s -> f1 x =p f2 x) -> sepit s f1 =p sepit s f2. +Proof. +elim: s=>[|x s IH] H h; first by rewrite !sepit0. +have /IH {IH}H': forall x : A, x \In s -> f1 x =p f2 x. + by move=>? H0; apply: H; apply/In_cons; right. +have Hx : x \In x :: s by apply/In_cons; left. +by rewrite !sepit_cons; split; case=>h1[h2][{h}-> H1 H2]; exists h1, h2; +split=>//; [rewrite -H | rewrite -H' | rewrite H | rewrite H']. +Qed. + +End IterStar. + +(* iterated star on eqType *) + +Section IterStarEq. +Variables (U : pcm) (A : eqType). + +Lemma sepitP x s (f : A -> Pred U) : + uniq s -> + sepit s f =p if x \in s + then f x # sepit (filter (predC1 x) s) f + else sepit s f. +Proof. +case E: (x \in s)=>//. +elim: s E=>[|y s IH] //= /[swap]; case/andP=>Hy Hu; rewrite sepit_cons inE; case/orP. +- by move/eqP=>->; rewrite eq_refl filter_predC1. +move=>Hx h0. +have ->: (y != x) by apply/eqP=>Hxy; rewrite Hxy Hx in Hy. +by split; case=>ha[h1][{h0}-> Ha]; [rewrite (IH Hx Hu) | rewrite sepit_cons]; +case=>hb[h][{h1}-> Hb H]; rewrite joinCA; exists hb, (ha \+ h); split=>//; +[rewrite sepit_cons | rewrite (IH Hx Hu)]; exists ha, h. +Qed. + +Corollary eq_sepitF s (f1 f2 : A -> Pred U) : + (forall x, x \in s -> f1 x =p f2 x) -> sepit s f1 =p sepit s f2. +Proof. by move=>H; apply: sepitF=>x Hx; apply/H/mem_seqP. Qed. + +Corollary perm_sepit s1 s2 (f : A -> Pred U) : + perm_eq s1 s2 -> sepit s1 f =p sepit s2 f. +Proof. by move/perm_eq_perm; exact: sepit_perm. Qed. + +End IterStarEq. + +End IterStar. + +(* iterated star on finsets *) + +Section FinIterStar. +Variables (U : pcm) (I : finType). + +Definition sepit (s : {set I}) (Ps : I -> Pred U) := + IterStar.sepit (enum s) Ps. + +Lemma sepit0 f : sepit set0 f =p emp. +Proof. +rewrite /sepit (IterStar.perm_sepit (s2 := filter pred0 (enum I))). +- by rewrite filter_pred0 IterStar.sepit0. +apply: uniq_perm; first by exact: enum_uniq. +- by rewrite filter_uniq // enum_uniq. +by move=>x; rewrite !mem_filter /= in_set0. +Qed. + +Lemma sepitF (s : {set I}) f1 f2 : + (forall x, x \in s -> f1 x =p f2 x) -> sepit s f1 =p sepit s f2. +Proof. +move=>H; apply: IterStar.eq_sepitF=>x H1; apply: H. +by rewrite -mem_enum. +Qed. + +Lemma eq_sepit (s1 s2 : {set I}) f : s1 =i s2 -> sepit s1 f =p sepit s2 f. +Proof. by move/eq_enum=>H; apply: IterStar.perm_sepit; rewrite H. Qed. + +Lemma sepitS x (s : {set I}) f : + sepit s f =p if x \in s then f x # sepit (s :\ x) f + else sepit s f. +Proof. +case E: (x \in s)=>//. +rewrite (IterStar.sepitP x (s:=enum s) f (enum_uniq s)) mem_enum E. +have Hp: perm_eq [seq y <- enum s | predC1 x y] (enum (s :\ x)). +- rewrite -filter_predI. + apply: uniq_perm=>[||y]; try by rewrite enum_uniq. + by rewrite !mem_filter /= in_setD1. +by move=>h0; split; case=>h1[h2][{h0}-> H1 H]; exists h1, h2; split=>//; +rewrite IterStar.perm_sepit; try by [exact: H]; [rewrite perm_sym |]. +Qed. + +Lemma sepitT1 x f : sepit setT f =p f x # sepit (setT :\ x) f. +Proof. by rewrite (sepitS x) in_setT. Qed. + +End FinIterStar. diff --git a/pcm/unionmap.v b/pcm/unionmap.v new file mode 100644 index 0000000..282c703 --- /dev/null +++ b/pcm/unionmap.v @@ -0,0 +1,6853 @@ +(* +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 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 HB Require Import structures. +From Coq Require Import ssreflect ssrbool ssrfun. +From mathcomp Require Import ssrnat eqtype seq path bigop. +From pcm Require Import options axioms prelude finmap seqperm pred seqext. +From pcm Require Export ordtype. +From pcm Require Import pcm morphism. + + +(****************************) +(****************************) +(* Reference Implementation *) +(****************************) +(****************************) + +(* UM.base type is reference implementation of union_maps. *) +(* Type has union_map structure if it exhibits isomorphism with UM.base. *) +(* Tying to reference implementation is alternative to axiomatizing *) +(* union maps, which seems unwieldy. *) +(* Thus, the underlying type of all union_maps instances will be *) +(* (copy of) UM.base. Different copies will *) +(* correspond to different canonical structures. *) +(* Copying further enables using different types for different *) +(* instances, as long as there is isomorphism between them. *) +(* This will allow mixing large and small types and *) +(* storing union maps into union maps (e.g., histories into heaps) *) +(* without universe errors. *) + +Module UM. +Section UM. +Variables (K : ordType) (C : pred K) (V : Type). + +Inductive base := + Undef | Def (f : finMap K V) of all C (supp f). + +Section FormationLemmas. +Variable (f g : finMap K V). + +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 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 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 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' + | 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: eq_irrelevance. +Qed. + +Definition valid f := if f is Def _ _ then true else false. + +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 (C k)) is left pf then + Def (all_supp_insP v pf fpf) + else Undef + else Undef. + +Definition dom f : seq K := + if f is Def fs _ then supp fs else [::]. + +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:=K) k fs else None. + +Definition union f1 f2 := + if (f1, f2) is (Def fs1 pf1, Def fs2 pf2) then + if disj fs1 fs2 then + Def (all_supp_fcatP pf1 pf2) + else Undef + 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. + +Definition pts k v := upd k v empty. + +(* forward induction *) +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) -> + P (union (pts k v) f)) -> + forall f, P f. +Proof. +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] : 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)). +- rewrite /valid /union /pts /upd /=. + case: decP; last by rewrite T1. + by move=>T; case: ifP=>//; rewrite E disjC disj_ins N disj_nil. +move/(H3 k v _ (IH T2)). +rewrite (_ : union (pts k v) (Def T2) = Def H). +- by apply. +rewrite umapE /union /pts /upd /=. +case: decP=>// T; rewrite /disj /= N /=. +by rewrite E fcat_inss // fcat0s. +Qed. + +(* backward induction *) +Lemma base_indb (P : base -> Prop) : + P Undef -> P empty -> + (forall k v f, P f -> valid (union f (pts k v)) -> + (forall x, x \in dom f -> ord x k) -> + P (union (pts k v) f)) -> + forall f, P f. +Proof. +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] : 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 (Def T2) (pts k v)). +- rewrite /valid /union /pts /upd /=. + case: decP; last by rewrite T1. + 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 /=. +case: decP=>// T; rewrite /disj /= N /=. +by rewrite E fcat_inss // fcat0s. +Qed. + +End UM. +End UM. + + +(***********************) +(***********************) +(* Union Map structure *) +(***********************) +(***********************) + +(* type T has union_map structure is it's cancellative TPCM *) +(* with an isomorphism to UM.base *) + +Definition union_map_axiom (K : ordType) (C : pred K) V T + (defined : T -> bool) (empty : T) (undef : T) (upd : K -> V -> T -> T) + (dom : T -> seq K) (assocs : T -> seq (K * V)) (free : T -> K -> T) + (find : K -> T -> option V) (union : T -> T -> T) + (empb : T -> bool) (undefb : T -> bool) (pts : K -> V -> T) + (from : T -> UM.base C V) (to : UM.base C V -> T) := + (((forall b, from (to b) = b) * + (forall f, to (from f) = f)) * + (((forall f, defined f = UM.valid (from f)) * + (undef = to (UM.Undef C V)) * + (empty = to (UM.empty C V))) * + ((forall k v f, upd k v f = to (UM.upd k v (from f))) * + (forall f, dom f = UM.dom (from f)) * + (forall f, assocs f = UM.assocs (from f)) * + (forall f k, free f k = to (UM.free (from f) k))) * + ((forall k f, find k f = UM.find k (from f)) * + (forall f1 f2, union f1 f2 = to (UM.union (from f1) (from f2))) * + (forall f, empb f = UM.empb (from f)) * + (forall f, undefb f = UM.undefb (from f)) * + (forall k v, pts k v = to (UM.pts C k v)))))%type. + +HB.mixin Record isUnionMap_helper (K : ordType) (C : pred K) V T of + Normal_CTPCM T := { + upd : K -> V -> T -> T; + dom : T -> seq K; + assocs : T -> seq (K * V); + free : T -> K -> T; + find : K -> T -> option V; + pts_op : K -> V -> T; + UMC_from : T -> UM.base C V; + UMC_to : UM.base C V -> T; + union_map_helper_subproof : + union_map_axiom valid Unit undef upd dom assocs + free find join unitb undefb pts_op UMC_from UMC_to}. + +(* embed cancellative topped PCM structure to enable inheritance *) +(* we prove later that this isn't needed *) +#[short(type="union_map")] +HB.structure Definition UMC K C V := + {T of @isUnionMap_helper K C V T & Normal_TPCM T & CTPCM T}. + +Lemma ftE K C V (U : @union_map K C V) b : UMC_from (UMC_to b : U) = b. +Proof. by rewrite union_map_helper_subproof. Qed. + +Lemma tfE K C V (U : @union_map K C V) b : UMC_to (UMC_from b) = b :> U. +Proof. by rewrite union_map_helper_subproof. Qed. + +Lemma eqE K C V (U : @union_map K C V) (b1 b2 : UM.base C V) : + UMC_to b1 = UMC_to b2 :> U <-> b1 = b2. +Proof. by split=>[E|->//]; rewrite -(ftE U b1) -(ftE U b2) E. Qed. + +Definition umEX K C V (U : @union_map K C V) := + (@union_map_helper_subproof K C V U, eqE, UM.umapE). + +(* Build CTCPM structure from the rest *) +HB.factory Record isUnion_map (K : ordType) (C : pred K) V T := { + um_defined : T -> bool; + um_empty : T; um_undef : T; + um_upd : K -> V -> T -> T; + um_dom : T -> seq K; + um_assocs : T -> seq (K * V); + um_free : T -> K -> T; + um_find : K -> T -> option V; + um_union : T -> T -> T; + um_empb : T -> bool; + um_undefb : T -> bool; + um_pts : K -> V -> T; + um_from : T -> UM.base C V; + um_to : UM.base C V -> T; + union_map_subproof : + union_map_axiom um_defined um_empty um_undef um_upd um_dom um_assocs + um_free um_find um_union um_empb um_undefb um_pts + um_from um_to}. + +HB.builders Context K C V T of isUnion_map K C V T. + +Definition umEX := snd union_map_subproof. + +Lemma ftE b : um_from (um_to b) = b. +Proof. by rewrite union_map_subproof. Qed. + +Lemma tfE b : um_to (um_from b) = b. +Proof. by rewrite union_map_subproof. Qed. + +Lemma eqE (b1 b2 : UM.base C V) : + um_to b1 = um_to b2 :> T <-> b1 = b2. +Proof. by split=>[E|->//]; rewrite -[b1]ftE -[b2]ftE E. Qed. + +Lemma union_map_is_pcm : pcm_axiom um_defined um_union um_empty um_empb. +Proof. +have joinC f1 f2 : um_union f1 f2 = um_union f2 f1. +- rewrite !umEX !eqE UM.umapE /UM.union. + case: (um_from f1)=>[|f1' H1]; case: (um_from f2)=>[|f2' H2] //. + by case: ifP=>E; rewrite disjC E // fcatC. +have joinCA f1 f2 f3 : um_union f1 (um_union f2 f3) = + um_union f2 (um_union f1 f3). +- rewrite !umEX !eqE !ftE UM.umapE /UM.union. + case: (um_from f1) (um_from f2) (um_from f3)=>[|f1' H1][|f2' H2][|f3' H3] //. + case E1: (disj f2' f3'); last first. + - by case E2: (disj f1' f3')=>//; rewrite disj_fcat E1 andbF. + rewrite disj_fcat andbC. + case E2: (disj f1' f3')=>//; rewrite disj_fcat (disjC f2') E1 /= andbT. + by case E3: (disj f1' f2')=>//; rewrite fcatAC // E1 E2 E3. +split=>[//|f1 f2 f3|f|f1 f2||f]. +- by rewrite (joinC f2) joinCA joinC. +- rewrite !umEX !ftE -[RHS]tfE eqE UM.umapE /UM.union /UM.empty. + by case: (um_from f)=>[//|f' H]; rewrite disjC disj_nil fcat0s. +- by rewrite !umEX !ftE; case: (um_from f1). +- by rewrite !umEX ftE. +rewrite !umEX /= /UM.empty /UM.empb -{1}[f]tfE. +case: (um_from f)=>[|f' F]; first by apply: ReflectF; rewrite eqE. +case: eqP=>E; constructor; rewrite eqE UM.umapE. +- by move/supp_nilE: E=>->. +by move=>H; rewrite H in E. +Qed. + +HB.instance Definition _ := isPCM.Build T union_map_is_pcm. + +Lemma union_map_is_cpcm : cpcm_axiom T. +Proof. +move=>f1 f2 f. +rewrite -{3}[f1]tfE -{2}[f2]tfE !pcmE /= !umEX /= !ftE !eqE. +rewrite !UM.umapE /UM.valid /UM.union. +case: (um_from f) (um_from f1) (um_from f2)=> +[|f' H]; case=>[|f1' H1]; case=>[|f2' H2] //; +case: ifP=>//; case: ifP=>// D1 D2 _ E. +by apply: fcatsK E; rewrite D1 D2. +Qed. + +HB.instance Definition _ := isCPCM.Build T union_map_is_cpcm. + +Lemma union_map_is_tpcm : tpcm_axiom um_undef um_undefb. +Proof. +split=>[/= x||/= x]. +- rewrite !umEX /= /UM.undefb -{1}[x]tfE. + case: (um_from x)=>[|f' F]; first by apply: ReflectT. + by apply: ReflectF; rewrite eqE. +- by rewrite pcmE /= !umEX ftE. +by rewrite pcmE /= !umEX ftE. +Qed. + +HB.instance Definition _ := isTPCM.Build T union_map_is_tpcm. + +Lemma union_map_is_normal : @normal_tpcm_axiom T. +Proof. +move=>/= x; rewrite pcmE /= /undef /= -{2}[x]tfE !umEX /UM.valid. +by case: (um_from x)=>[|f' H]; [right|left]. +Qed. + +HB.instance Definition _ := isNormal_TPCM.Build T union_map_is_normal. +HB.instance Definition _ := isUnionMap_helper.Build K C V T union_map_subproof. +HB.end. + +(* Making pts infer union_map structure *) +Definition ptsx K C V (U : union_map C V) k v of phant U : U := + @pts_op K C V U k v. + +(* use ptsT to pass map type explicitly *) +(* use pts when type inferrable or for printing *) +Notation ptsT U k v := (ptsx k v (Phant U)) (only parsing). +Notation pts k v := (ptsT _ k v). + +(*************************************) +(* Union map with decidable equality *) +(*************************************) + +(* union_map has decidable equality if V does *) + +Section UnionMapEq. +Variables (K : ordType) (C : pred K) (V : eqType). +Variable (U : union_map C V). + +Definition union_map_eq (f1 f2 : U) := + match (UMC_from f1), (UMC_from f2) with + | UM.Undef, UM.Undef => true + | UM.Def f1' _, UM.Def f2' _ => f1' == f2' + | _, _ => false + end. + +Lemma union_map_eqP : Equality.axiom union_map_eq. +Proof. +move=>x y; rewrite -{1}[x]tfE -{1}[y]tfE /union_map_eq. +case: (UMC_from x)=>[|f1 H1]; case: (UMC_from y)=>[|f2 H2] /=. +- by constructor. +- by constructor=>/eqE. +- by constructor=>/eqE. +case: eqP=>E; constructor; rewrite eqE. +- by rewrite UM.umapE. +by case. +Qed. + +(* no canonical projection to latch onto, so no generic inheritance *) +(* but unionmap_eqP can be used for any individual U *) +(* +HB.instance Definition _ := hasDecEq.Build U union_map_eqP. +*) +End UnionMapEq. + + +(*********************************) +(* Instance of union maps *) +(* with trivially true condition *) +(*********************************) + +Record umap K V := UMap {umap_base : @UM.base K xpredT V}. + +Section UmapUMC. +Variables (K : ordType) (V : Type). +Implicit Type f : umap K V. +Local Coercion umap_base : umap >-> UM.base. + +Let um_valid f := @UM.valid K xpredT V f. +Let um_empty := UMap (@UM.empty K xpredT V). +Let um_undef := UMap (@UM.Undef K xpredT V). +Let um_upd k v f := UMap (@UM.upd K xpredT V k v f). +Let um_dom f := @UM.dom K xpredT V f. +Let um_assocs f := @UM.assocs K xpredT V f. +Let um_free f k := UMap (@UM.free K xpredT V f k). +Let um_find k f := @UM.find K xpredT V k f. +Let um_union f1 f2 := UMap (@UM.union K xpredT V f1 f2). +Let um_empb f := @UM.empb K xpredT V f. +Let um_undefb f := @UM.undefb K xpredT V f. +Let um_from (f : umap K V) : UM.base _ _ := f. +Let um_to (b : @UM.base K xpredT V) : umap K V := UMap b. +Let um_pts k v := UMap (@UM.pts K xpredT V k v). + +Lemma umap_is_umc : + union_map_axiom um_valid um_empty um_undef um_upd um_dom + um_assocs um_free um_find um_union um_empb + um_undefb um_pts um_from um_to. +Proof. by split=>//; split=>[|[]]. Qed. + +HB.instance Definition _ := isUnion_map.Build K xpredT V (umap K V) umap_is_umc. +End UmapUMC. + +(* if V is eqtype so is umap K V *) +HB.instance Definition _ (K : ordType) (V : eqType) := + hasDecEq.Build (umap K V) (@union_map_eqP K xpredT V (umap K V)). + +Notation "x \\-> v" := (ptsT (umap _ _) x v) (at level 30). + +(* DEVCOMMENT *) +(* remove these "tests" for release *) +(* 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). +rewrite joinC. +Abort. + +(* can we use the base type? *) +Lemma xx (x : umap nat nat) : x \+ x == x \+ x. +Abort. + +(* can i store maps into maps without universe inconsistencies? *) +(* yes, the idea of the class works *) +Lemma xx : 1 \\-> (1 \\-> 3) = 2 \\-> (7 \\-> 3). +Abort. + +(* /DEVCOMMENT *) + + +(***************) +(* Finite sets *) +(***************) + +(* like umaps with unit range *) +(* but we give them their own copy of the type *) + +Record fset K := FSet {fset_base : @UM.base K xpredT unit}. + +Section FsetUMC. +Variable K : ordType. +Implicit Type f : fset K. +Local Coercion fset_base : fset >-> UM.base. + +Let fs_valid f := @UM.valid K xpredT unit f. +Let fs_empty := FSet (@UM.empty K xpredT unit). +Let fs_undef := FSet (@UM.Undef K xpredT unit). +Let fs_upd k v f := FSet (@UM.upd K xpredT unit k v f). +Let fs_dom f := @UM.dom K xpredT unit f. +Let fs_assocs f := @UM.assocs K xpredT unit f. +Let fs_free f k := FSet (@UM.free K xpredT unit f k). +Let fs_find k f := @UM.find K xpredT unit k f. +Let fs_union f1 f2 := FSet (@UM.union K xpredT unit f1 f2). +Let fs_empb f := @UM.empb K xpredT unit f. +Let fs_undefb f := @UM.undefb K xpredT unit f. +Let fs_from (f : fset K) : UM.base _ _ := f. +Let fs_to (b : @UM.base K xpredT unit) : fset K := FSet b. +Let fs_pts k v := FSet (@UM.pts K xpredT unit k v). + +Lemma fset_is_umc : + union_map_axiom fs_valid fs_empty fs_undef fs_upd fs_dom + fs_assocs fs_free fs_find fs_union fs_empb + fs_undefb fs_pts fs_from fs_to. +Proof. by split=>//; split=>[|[]]. Qed. + +HB.instance Definition _ := + isUnion_map.Build K xpredT unit (fset K) fset_is_umc. +HB.instance Definition _ := + hasDecEq.Build (fset K) (@union_map_eqP K xpredT unit (fset K)). +End FsetUMC. + +Notation "# x" := (ptsT (fset _) x tt) (at level 30, format "# x"). + +(* DEVCOMMENT *) +(* test *) +(* /DECOMMENT *) +Lemma xx : #1 \+ #2 = Unit. +Abort. + +(*******) +(* dom *) +(*******) + +Section DomLemmas. +Variables (K : ordType) (C : pred K) (V : Type) (U : union_map C V). +Implicit Types (k : K) (v : V) (f g : U). + +Lemma dom_undef : dom (undef : U) = [::]. +Proof. by rewrite !umEX. Qed. + +Lemma dom0 : dom (Unit : U) = [::]. +Proof. by rewrite !umEX. Qed. + +Lemma dom0E f : valid f -> dom f =i pred0 -> f = Unit. +Proof. +rewrite !umEX /UM.valid /UM.dom /UM.empty -{3}[f]tfE !eqE UM.umapE. +case: (UMC_from f)=>[|f' S] //= _; rewrite fmapE /= {S}. +by case: f'; case=>[|kv s] //= P /= /(_ kv.1); rewrite inE eq_refl. +Qed. + +Lemma domNE f : reflect (dom f = [::]) (unitb f || undefb f). +Proof. +case: (normalP0 f)=>[->|->|W N]; constructor; rewrite ?dom0 ?dom_undef //. +by move=>M; apply: N; apply: dom0E W _; rewrite M. +Qed. + +Lemma dom0NP f : reflect (exists k, k \in dom f) (dom f != [::]). +Proof. +case: has_nilP=>X; constructor. +- by case/hasP: X=>x; exists x. +by case=>k D; apply: X; apply/hasP; exists k. +Qed. + +Lemma domU k v f : + dom (upd k v f) =i + [pred x | C k & if x == k then valid f else x \in dom f]. +Proof. +rewrite !umEX /UM.dom /UM.upd /UM.valid /=. +case: (UMC_from f)=>[|f' H] /= x. +- by rewrite !inE andbC; case: ifP. +case: decP=>[->|/(introF idP)->//]. +by rewrite supp_ins. +Qed. + +Lemma domF k f : + 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 /=; +by case: ifP=>// E; rewrite supp_rem inE /= eq_sym E. +Qed. + +Lemma subdomF k f : {subset dom (free f k) <= dom f}. +Proof. by move=>x; rewrite domF inE; case: eqP. Qed. + +Lemma domUn f1 f2 : + dom (f1 \+ f2) =i + [pred x | valid (f1 \+ f2) & (x \in dom f1) || (x \in dom f2)]. +Proof. +rewrite !umEX /UM.dom /UM.valid /UM.union. +case: (UMC_from f1) (UMC_from f2)=>[|f1' H1] // [|f2' H2] // x. +by case: ifP=>E //; rewrite supp_fcat. +Qed. + +(* bidirectional version of domUn *) +Lemma domUnE f1 f2 x : + valid (f1 \+ f2) -> + x \in dom (f1 \+ f2) = (x \in dom f1) || (x \in dom f2). +Proof. by move=>W; rewrite domUn inE W. Qed. + +Lemma dom_valid k f : k \in dom f -> valid f. +Proof. by rewrite !umEX; case: (UMC_from f). Qed. + +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. + +Lemma dom_inIR k f1 f2 : + valid (f1 \+ f2) -> k \in dom f2 -> k \in dom (f1 \+ f2). +Proof. by rewrite joinC; apply: dom_inIL. Qed. + +Lemma dom_NNL k f1 f2 : + valid (f1 \+ f2) -> k \notin dom (f1 \+ f2) -> k \notin dom f1. +Proof. by move=> D; apply/contra/dom_inIL. Qed. + +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 f k = f. +Proof. +rewrite -{3}[f]tfE !umEX /UM.dom /UM.free /=. +by case: (UMC_from f)=>[|f' H] //; apply: rem_supp. +Qed. + +Variant dom_find_spec k f : option V -> bool -> Type := +| dom_find_none'' of find k f = None : dom_find_spec k f None false +| dom_find_some'' v of find k f = Some v & + f = upd k v (free f k) : dom_find_spec k f (Some v) true. + +Lemma dom_find k f : dom_find_spec k f (find k f) (k \in dom f). +Proof. +rewrite !umEX /UM.dom -{1}[f]tfE. +case: (UMC_from f)=>[|f' H]. +- by apply: dom_find_none''; rewrite !umEX. +case: suppP (allP H k)=>[v|] /[dup] H1 /= -> I; last first. +- by apply: dom_find_none''; rewrite !umEX. +apply: (dom_find_some'' (v:=v)); rewrite !umEX // /UM.upd /UM.free. +case: decP=>H2; last by rewrite I in H2. +apply/fmapP=>k'; rewrite fnd_ins. +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. + +Lemma find_none k f : find k f = None <-> k \notin dom f. +Proof. by case: dom_find=>// v ->. 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) -> + f1 \+ g1 = f2 \+ g2 -> + dom f1 =i dom f2 -> f1 = f2. +Proof. +rewrite -{4}[f1]tfE -{3}[f2]tfE /= !umEX. +rewrite /UM.valid /UM.union /UM.dom; move=>D E. +case: (UMC_from f1) (UMC_from f2) (UMC_from g1) (UMC_from g2) E D=> +[|f1' F1][|f2' F2][|g1' G1][|g2' G2] //; +case: ifP=>// D1; case: ifP=>// D2 E _ E1; apply/fmapP=>k. +move/(_ k): E1=>E1. +case E2: (k \in supp f2') in E1; last first. +- by move/negbT/fnd_supp: E1=>->; move/negbT/fnd_supp: E2=>->. +suff L: forall m s, k \in supp m -> disj m s -> + fnd k m = fnd k (fcat m s) :> option V. +- by rewrite (L _ _ E1 D1) (L _ _ E2 D2) E. +by move=>m s S; case: disjP=>//; move/(_ _ S)/negbTE; rewrite fnd_fcat=>->. +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. +move=>Vh; apply: uniq_perm; last 1 first. +- by move=>x; rewrite mem_cat domUn inE Vh. +- by apply: uniq_dom. +rewrite cat_uniq !uniq_dom /= andbT; apply/hasPn=>x. +rewrite !umEX /UM.valid /UM.union /UM.dom in Vh *. +case: (UMC_from f1) (UMC_from f2) Vh=>// f1' H1 [//|f2' H2]. +by case: disjP=>// H _; apply: contraL (H x). +Qed. + +Lemma size_domUn f1 f2 : + valid (f1 \+ f2) -> + size (dom (f1 \+ f2)) = size (dom f1) + size (dom f2). +Proof. by move/perm_domUn/seq.perm_size; rewrite size_cat. Qed. + +Lemma size_domF k f : + k \in dom f -> size (dom (free f k)) = (size (dom f)).-1. +Proof. +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 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. +rewrite !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. + +End DomLemmas. + +Arguments subdomF {K C V U k f}. + +#[export] +Hint Resolve sorted_dom uniq_dom all_dom : core. +Prenex Implicits find_some find_none subdomF. + +(* lemmas for comparing doms of two differently-typed maps *) +Section DomLemmas2. +Variables (K : ordType) (C1 C2 : pred K) (V1 V2 : Type). +Variables (U1 : union_map C1 V1) (U2 : union_map C2 V2). + +Lemma domE (f1 : U1) (f2 : U2) : dom f1 =i dom f2 <-> dom f1 = dom f2. +Proof. by split=>[|->] //; apply/ord_sorted_eq/sorted_dom/sorted_dom. Qed. + +Lemma domKi (f1 g1 : U1) (f2 g2 : U2) : + valid (f1 \+ g1) -> valid (f2 \+ g2) -> + dom (f1 \+ g1) =i dom (f2 \+ g2) -> + dom f1 =i dom f2 -> dom g1 =i dom g2. +Proof. +rewrite !umEX /UM.valid /UM.union /UM.dom. +case: (UMC_from f1) (UMC_from f2) (UMC_from g1) (UMC_from g2)=> +[|f1' F1][|f2' F2][|g1' G1][|g2' G2] //. +case: disjP=>// D1 _; case: disjP=>// D2 _ E1 E2 x. +move: {E1 E2} (E2 x) (E1 x); rewrite !supp_fcat !inE /= => E. +move: {D1 D2} E (D1 x) (D2 x)=><- /implyP D1 /implyP D2. +case _ : (x \in supp f1') => //= in D1 D2 *. +by move/negbTE: D1=>->; move/negbTE: D2=>->. +Qed. + +Lemma domK (f1 g1 : U1) (f2 g2 : U2) : + valid (f1 \+ g1) -> valid (f2 \+ g2) -> + dom (f1 \+ g1) = dom (f2 \+ g2) -> + dom f1 = dom f2 -> dom g1 = dom g2. +Proof. by move=>D1 D2 /domE H1 /domE H2; apply/domE; apply: domKi H2. Qed. + +Lemma domKUni (f1 g1 : U1) (f2 g2 : U2) : + valid (f1 \+ g1) = valid (f2 \+ g2) -> + dom f1 =i dom f2 -> dom g1 =i dom g2 -> + dom (f1 \+ g1) =i dom (f2 \+ g2). +Proof. by move=>E D1 D2 x; rewrite !domUn !inE E D1 D2. Qed. + +Lemma domKUn (f1 g1 : U1) (f2 g2 : U2) : + valid (f1 \+ g1) = valid (f2 \+ g2) -> + dom f1 = dom f2 -> dom g1 = dom g2 -> + dom (f1 \+ g1) = dom (f2 \+ g2). +Proof. by move=>E /domE D1 /domE D2; apply/domE/domKUni. Qed. + +Lemma subdom_filter (f1 : U1) (f2 : U2) : + {subset dom f1 <= dom f2} -> + dom f1 = filter (mem (dom f1)) (dom f2). +Proof. +have Tx : transitive (@ord K) by apply: trans. +move=>S; apply: (sorted_eq (leT := @ord K))=>//. +- by apply: antisym. +- by apply: sorted_dom. +- by rewrite sorted_filter // sorted_dom. +apply: uniq_perm; rewrite ?filter_uniq ?uniq_dom // => y. +rewrite mem_filter /=; case E: (y \in dom _)=>//=. +by move/S: E=>->. +Qed. + +End DomLemmas2. + +(*********) +(* valid *) +(*********) + +Section ValidLemmas. +Variables (K : ordType) (C : pred K) (V : Type) (U : union_map C V). +Implicit Types (k : K) (v : V) (f g : U). + +Lemma invalidE f : ~~ valid f <-> f = undef. +Proof. by rewrite undefVN negbK; case: undefbP. Qed. + +Lemma validU k v f : valid (upd k v f) = C k && valid f. +Proof. +rewrite !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 f k) = valid f. +Proof. by rewrite !umEX; case: (UMC_from f). Qed. + +Variant validUn_spec f1 f2 : bool -> Type := +| valid_false1 of ~~ valid f1 : validUn_spec f1 f2 false +| valid_false2 of ~~ valid f2 : validUn_spec f1 f2 false +| valid_false3 k of k \in dom f1 & k \in dom f2 : validUn_spec f1 f2 false +| valid_true of valid f1 & valid f2 & + (forall x, x \in dom f1 -> x \notin dom f2) : validUn_spec f1 f2 true. + +Lemma validUn f1 f2 : validUn_spec f1 f2 (valid (f1 \+ f2)). +Proof. +rewrite !umEX -{1}[f1]tfE -{1}[f2]tfE. +rewrite /UM.valid /UM.union /=. +case: (UMC_from f1) (UMC_from f2)=>[|f1' F1][|f2' F2] /=. +- by apply: valid_false1; rewrite !umEX. +- by apply: valid_false1; rewrite !umEX. +- by apply: valid_false2; rewrite !umEX. +case: ifP=>E. +- apply: valid_true; try by rewrite !umEX. + by move=>k; rewrite !umEX => H; case: disjP E=>//; move/(_ _ H). +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 f1 k \+ f2). +Proof. +case: validUn=>// V1 V2 H _; case: validUn=>//; last 1 first. +- by move=>k'; rewrite domF inE; case: eqP=>// _ /H/negbTE ->. +by rewrite validF V1. +by rewrite V2. +Qed. + +Lemma validUnF k f1 f2 : + valid (f1 \+ f2) -> valid (f1 \+ free f2 k). +Proof. by rewrite !(joinC f1); apply: validFUn. Qed. + +Lemma validUnU k v f1 f2 : + k \in dom f2 -> valid (f1 \+ upd k v f2) = valid (f1 \+ f2). +Proof. +move=>D; apply/esym; move: D; case: validUn. +- by move=>V' _; apply/esym/negbTE; apply: contra V'; move/validL. +- move=>V' _; apply/esym/negbTE; apply: contra V'; move/validR. + by rewrite validU; case/andP. +- move=>k' H1 H2 _; case: validUn=>//; rewrite validU => D1 /andP [/= H D2]. + by move/(_ k' H1); rewrite domU !inE H D2 H2; case: ifP. +move=>V1 V2 H1 H2; case: validUn=>//. +- by rewrite V1. +- by rewrite validU V2 (dom_cond H2). +move=>k'; rewrite domU (dom_cond H2) inE /= V2; move/H1=>H3. +by rewrite (negbTE H3); case: ifP H2 H3=>// /eqP ->->. +Qed. + +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 dom_inNL k f1 f2 : + valid (f1 \+ f2) -> k \in dom f1 -> k \notin dom f2. +Proof. by case: validUn=>//_ _ H _; apply: H. Qed. + +Lemma dom_inNR k f1 f2 : + valid (f1 \+ f2) -> k \in dom f2 -> k \notin dom f1. +Proof. by rewrite joinC; apply: dom_inNL. Qed. + +Lemma dom_inNLX k f1 f2 : + valid (f1 \+ f2) -> k \in dom f1 -> ~ k \in dom f2. +Proof. by move=>W /(dom_inNL W)/negbTE ->. Qed. + +Lemma dom_inNRX k f1 f2 : + valid (f1 \+ f2) -> k \in dom f2 -> ~ k \in dom f1. +Proof. by move=>W /(dom_inNR W)/negbTE ->. Qed. + +(* a more symmetric version of the above *) +Lemma dom_inN k1 k2 f1 f2 : + valid (f1 \+ f2) -> + k1 \in dom f1 -> k2 \in dom f2 -> k1 != k2. +Proof. by case: (k1 =P k2) =>// <- W H /(dom_inNLX W H). Qed. + +Lemma dom_inNX k1 k2 f1 f2 : + valid (f1 \+ f2) -> + k1 \in dom f1 -> k2 \in dom f2 -> k1 <> k2. +Proof. by move=>W K1 K2; apply/eqP; apply: dom_inN K1 K2. Qed. + +End ValidLemmas. + + +(************) +(* um_unitb *) +(************) + +(* AKA empb *) + +Section UnitbLemmas. +Variables (K : ordType) (C : pred K) (V : Type) (U : union_map C V). +Implicit Types (k : K) (v : V) (f : U). + +Lemma um_unitbU k v f : unitb (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 um_unitbUn f1 f2 : unitb (f1 \+ f2) = unitb f1 && unitb f2. +Proof. +rewrite !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 umap0E f1 f2 : f1 \+ f2 = Unit -> f1 = Unit /\ f2 = Unit. +Proof. +by move/unitbP; rewrite um_unitbUn=>/andP [H1 H2]; split; apply/unitbP. +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 !umEX /UM.valid /UM.dom /UM.empty -{3}[f]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/unitbP: E=>->; rewrite unitL valid_unit. +case: validUn=>// W _ H; move/validEb: E; elim; split=>// k. +by apply/negP=>/[dup]/H/negbTE ->. +Qed. + +End UnitbLemmas. + + +(**********) +(* undefb *) +(**********) + +Section UndefbLemmas. +Variables (K : ordType) (C : pred K) (V : Type) (U : union_map C V). +Implicit Types (k : K) (v : V) (f : U). + +Lemma undefb_undef : undefb (undef : U). +Proof. by rewrite !umEX. Qed. + +Lemma undefbP f : reflect (f = undef) (undefb f). +Proof. +rewrite !umEX /UM.undefb -{1}[f]tfE. +by case: (UMC_from f)=>[|f' F]; constructor; rewrite !umEX. +Qed. + +Lemma undefbE f : f = undef <-> undefb f. +Proof. by case: undefbP. Qed. + +Lemma undefb0 : undefb (Unit : U) = false. +Proof. by case: undefbP=>// /esym/undef0. Qed. + +Lemma valid_undefbI f : valid f -> undefb f = false. +Proof. +move=>W; apply/negP=>/undefbP E. +by move: E W=>->; rewrite valid_undef. +Qed. + +End UndefbLemmas. + + +(**********) +(* dom_eq *) +(**********) + +Section DomEqDef. +Variables (K : ordType) (C1 C2 : pred K) (V1 V2 : Type). +Variable U1 : union_map C1 V1. +Variable U2 : union_map C2 V2. + +Definition dom_eq (f1 : U1) (f2 : U2) := + (valid f1 == valid f2) && (dom f1 == dom f2). + +End DomEqDef. + +Section DomEqLemmas. +Variables (K : ordType) (C1 C2 C3 : pred K) (V1 V2 V3 : Type). +Variable U1 : union_map C1 V1. +Variable U2 : union_map C2 V2. +Variable U3 : union_map C3 V3. + +Lemma domeqP (f1 : U1) (f2 : U2) : + reflect (valid f1 = valid f2 /\ dom f1 = dom f2) (dom_eq f1 f2). +Proof. by rewrite /dom_eq; apply/andPP/eqP/eqP. Qed. + +Lemma domeqVE (f1 : U1) (f2 : U2) : dom_eq f1 f2 -> valid f1 = valid f2. +Proof. by case/domeqP. Qed. + +Lemma domeqDE (f1 : U1) (f2 : U2) : dom_eq f1 f2 -> dom f1 = dom f2. +Proof. by case/domeqP. Qed. + +Lemma domeqE (f1 : U1) (f2 : U2) : + dom_eq f1 f2 -> (valid f1 = valid f2) * (dom f1 = dom f2). +Proof. by move=>H; rewrite (domeqVE H) (domeqDE H). Qed. + +Lemma domeq0E (f : U1) : dom_eq f (Unit : U1) -> f = Unit. +Proof. +case/andP=>/eqP; rewrite valid_unit dom0=>H1 H2. +by apply: dom0E=>//; rewrite (eqP H2). +Qed. + +Lemma domeq_unit : dom_eq (Unit : U1) (Unit : U2). +Proof. by rewrite /dom_eq !valid_unit !dom0. Qed. + +Lemma domeq_undef : dom_eq (undef : U1) (undef : U2). +Proof. by rewrite /dom_eq !valid_undef !dom_undef. Qed. + +Lemma domeq_refl (f : U1) : dom_eq f f. +Proof. by rewrite /dom_eq !eqxx. Qed. + +Lemma domeq_sym (f1 : U1) (f2 : U2) : dom_eq f1 f2 -> dom_eq f2 f1. +Proof. by case/andP=>V D; apply/andP; rewrite (eqP V) (eqP D). Qed. + +Lemma domeq_trans (f1 : U1) (f2 : U2) (f3 : U3) : + dom_eq f1 f2 -> dom_eq f2 f3 -> dom_eq f1 f3. +Proof. by rewrite /dom_eq=>/andP [/eqP -> /eqP ->]. Qed. + +Lemma domeqVUnE (f1 f1' : U1) (f2 f2' : U2) : + dom_eq f1 f2 -> dom_eq f1' f2' -> + valid (f1 \+ f1') = valid (f2 \+ f2'). +Proof. +suff L (C1' C2' : pred K) V1' V2' (U1' : union_map C1' V1') + (U2' : union_map C2' V2') (g1 g1' : U1') (g2 g2' : U2') : + dom_eq g1 g2 -> dom_eq g1' g2' -> + valid (g1 \+ g1') -> valid (g2 \+ g2'). +- by move=>D D'; apply/idP/idP; apply: L=>//; apply: domeq_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. + +Lemma domeqVUn (f1 f1' : U1) (f2 f2' : U2) : + dom_eq f1 f2 -> dom_eq f1' f2' -> + valid (f1 \+ f1') -> valid (f2 \+ f2'). +Proof. by move=>D D'; rewrite -(domeqVUnE D D'). Qed. + +Lemma domeqUn (f1 f1' : U1) (f2 f2' : U2) : + dom_eq f1 f2 -> dom_eq f1' f2' -> + dom_eq (f1 \+ f1') (f2 \+ f2'). +Proof. +move/[dup]=>D /domeqP [V E] /[dup] D' /domeqP [V' E']. +move: (domeqVUnE D D')=>Ex; apply/domeqP=>//. +by rewrite Ex (domKUn Ex E E'). +Qed. + +Lemma domeqK (f1 f1' : U1) (f2 f2' : U2) : + valid (f1 \+ f1') -> + dom_eq (f1 \+ f1') (f2 \+ f2') -> + dom_eq f1 f2 -> dom_eq f1' f2'. +Proof. +move=>D1 /domeqP [E1 H1] /domeqP [E2 H2]; move: (D1); rewrite E1=>D2. +apply/domeqP; split; first by rewrite (validE2 D1) (validE2 D2). +by apply: domK D1 D2 H1 H2. +Qed. + +Lemma big_domeqUn A (xs : seq A) (f1 : A -> U1) (f2 : A -> U2) : + (forall x, x \In xs -> dom_eq (f1 x) (f2 x)) -> + dom_eq (\big[join/Unit]_(i <- xs) (f1 i)) + (\big[join/Unit]_(i <- xs) (f2 i)). +Proof. +elim: xs=>[|x xs IH] D; first by rewrite !big_nil domeq_unit. +rewrite !big_cons; apply: domeqUn. +- by apply: D; rewrite InE; left. +by apply: IH=>z Z; apply: D; rewrite InE; right. +Qed. + +End DomEqLemmas. + +#[export] +Hint Resolve domeq_refl : core. + +Section DomEq1Lemmas. +Variables (K : ordType) (C : pred K) (V : Type) (U : union_map C V). + +Lemma domeqfU k v (f : U) : k \in dom f -> dom_eq f (upd k v f). +Proof. +move=>D; rewrite /dom_eq validU (dom_cond D) (dom_valid D) /=. +apply/eqP/domE=>x; rewrite domU inE (dom_cond D) (dom_valid D). +by case: eqP=>// ->. +Qed. + +Lemma domeqUf k v (f : U) : k \in dom f -> dom_eq (upd k v f) f. +Proof. by move=>D; apply/domeq_sym/domeqfU. Qed. + +Lemma domeqfUn (f f1 f2 f1' f2' : U) : + dom_eq (f \+ f1) f2 -> + dom_eq f1' (f \+ f2') -> + dom_eq (f1 \+ f1') (f2 \+ f2'). +Proof. +move=>D1 D2; apply: (domeq_trans (f2:=f1 \+ f \+ f2')). +- by rewrite -joinA; apply: domeqUn. +by rewrite -joinA joinCA joinA; apply: domeqUn. +Qed. + +Lemma domeqUnf (f f1 f2 f1' f2' : U) : + dom_eq f1 (f \+ f2) -> dom_eq (f \+ f1') f2' -> + dom_eq (f1 \+ f1') (f2 \+ f2'). +Proof. +by move=>D1 D2; rewrite (joinC f1) (joinC f2); apply: domeqfUn D2 D1. +Qed. + +End DomEq1Lemmas. + +Section DomEq2Lemmas. +Variables (K : ordType) (C1 C2 : pred K) (V1 V2 : Type). +Variables (U1 : union_map C1 V1) (U2 : union_map C2 V2). + +Lemma domeq_symE (f1 : U1) (f2 : U2) : dom_eq f1 f2 = dom_eq f2 f1. +Proof. by apply/idP/idP; apply: domeq_sym. Qed. + +End DomEq2Lemmas. + + +(**********) +(* update *) +(**********) + +Section UpdateLemmas. +Variable (K : ordType) (C : pred K) (V : Type) (U : union_map C V). +Implicit Types (k : K) (v : V) (f : U). + +Lemma upd_undef k v : upd k v undef = undef :> U. +Proof. by rewrite !umEX. Qed. + +Lemma upd_condN k v f : ~~ C k -> upd k v f = undef. +Proof. +rewrite !umEX /UM.upd. +case: (UMC_from f)=>[|f' H']=>//. +by case: decP=>//= ->. +Qed. + +Lemma upd_inj k v1 v2 f : + valid f -> C k -> upd k v1 f = upd k v2 f -> v1 = v2. +Proof. +rewrite !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. +Qed. + +Lemma updU k1 k2 v1 v2 f : + upd k1 v1 (upd k2 v2 f) = + if k1 == k2 then upd k1 v1 f else upd k2 v2 (upd k1 v1 f). +Proof. +rewrite !umEX /UM.upd. +case: (UMC_from f)=>[|f' H']; case: ifP=>// E; +case: decP=>H1; case: decP=>H2 //; rewrite !umEX. +- by rewrite ins_ins E. +- by rewrite (eqP E) in H2 *. +by rewrite ins_ins E. +Qed. + +Lemma updF k1 k2 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; +by case: decP=>H2 //; rewrite !umEX ins_rem H1. +Qed. + +Lemma updUnL k v f1 f2 : + upd k v (f1 \+ f2) = + if k \in dom f1 then upd k v f1 \+ f2 else f1 \+ upd k v f2. +Proof. +rewrite !umEX /UM.upd /UM.union /UM.dom. +case: (UMC_from f1) (UMC_from f2)=>[|f1' F1][|f2' F2] //=. +- by case: ifP=>//; case: decP. +case: ifP=>// D; case: ifP=>// H1; case: decP=>// H2. +- rewrite disjC disj_ins disjC D !umEX; case: disjP D=>// H _. + by rewrite (H _ H1) /= fcat_inss // (H _ H1). +- by rewrite disj_ins H1 D /= !umEX fcat_sins. +- by rewrite disjC disj_ins disjC D andbF. +by rewrite disj_ins D andbF. +Qed. + +Lemma updUnR k v f1 f2 : + upd 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. + +Lemma domUE k e f : k \in dom f -> dom (upd k e f) = dom f. +Proof. +move=>H; apply/domE=>x; rewrite domU inE (dom_cond H) /=. +by case: eqP=>// ->; rewrite H (dom_valid H). +Qed. + +End UpdateLemmas. + + +(********) +(* free *) +(********) + +Section FreeLemmas. +Variables (K : ordType) (C : pred K) (V : Type) (U : union_map C V). +Implicit Types (k : K) (v : V) (f : U). + +Lemma free0 k : free Unit k = Unit :> U. +Proof. by rewrite !umEX /UM.free /UM.empty rem_empty. Qed. + +Lemma free_undef k : free undef k = undef :> U. +Proof. by rewrite !umEX. Qed. + +Lemma freeU k1 k2 v f : + free (upd k2 v f) k1 = + if k1 == k2 then + if C k2 then free f k1 else undef + else upd k2 v (free f k1). +Proof. +rewrite !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. +- by case: ifP H2. +by rewrite !umEX rem_ins H1. +Qed. + +Lemma freeF k1 k2 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 (f1 \+ f2) k = + if k \in dom (f1 \+ f2) then free f1 k \+ free f2 k + else f1 \+ f2. +Proof. +rewrite !umEX /UM.free /UM.union /UM.dom. +case: (UMC_from f1) (UMC_from f2)=>[|f1' F1][|f2' F2] //. +case: ifP=>// E1; rewrite supp_fcat inE /=. +case: ifP=>E2; last by rewrite !umEX rem_supp // supp_fcat inE E2. +rewrite disj_rem; last by rewrite disjC disj_rem // disjC. +rewrite !umEX; case/orP: E2=>E2. +- suff E3: k \notin supp f2' by rewrite -fcat_rems // (rem_supp E3). + by case: disjP E1 E2=>// H _; move/H. +suff E3: k \notin supp f1' by rewrite -fcat_srem // (rem_supp E3). +by case: disjP E1 E2=>// H _; move/contra: (H k); rewrite negbK. +Qed. + +Lemma freeUnV k f1 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 freeUnR 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 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 => ->. +move=>x H3; move: (H _ H3); rewrite domF inE /=. +by case: ifP H3 V1=>[|_ _ _]; [move/eqP=><- -> | move/negbTE=>->]. +Qed. + +Lemma freeUnL k f1 f2 : + k \notin dom f2 -> + free (f1 \+ f2) k = free f1 k \+ f2. +Proof. by move=>H; rewrite joinC freeUnR // joinC. Qed. + +(* positive re-statement of freeUnL, but requires validity *) +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 freeUnL. Qed. + +(* positive re-statement of freeUnR, but requires validity *) +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. + +Lemma freeND k f : k \notin dom f -> free f k = f. +Proof. by move=>W; rewrite -[f]unitR freeUnR // free0. Qed. + +End FreeLemmas. + + +(********) +(* find *) +(********) + +Section FindLemmas. +Variables (K : ordType) (C : pred K) (V : Type) (U : union_map C V). +Implicit Types (k : K) (v : V) (f : U). + +Lemma find0E k : find k (Unit : U) = None. +Proof. by rewrite !umEX /UM.find /= fnd_empty. Qed. + +Lemma find_undef k : find k (undef : U) = None. +Proof. by rewrite !umEX /UM.find. Qed. + +Lemma find_cond k f : ~~ C k -> find k f = None. +Proof. +rewrite !umEX /UM.find; case: (UMC_from f)=>[|f' F] // H. +suff: k \notin supp f' by case: suppP. +by apply: contra H; case: allP F=>// F _ /F. +Qed. + +Lemma findU k1 k2 v f : + find k1 (upd k2 v f) = + if k1 == k2 then + if C k2 && valid f then Some v else None + else if C k2 then find k1 f else None. +Proof. +rewrite !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 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. +by rewrite fnd_rem. +Qed. + +Lemma findUnL k f1 f2 : + valid (f1 \+ f2) -> + find k (f1 \+ f2) = if k \in dom f1 then find k f1 else find k f2. +Proof. +rewrite !umEX /UM.valid /UM.find /UM.union /UM.dom. +case: (UMC_from f1) (UMC_from f2)=>[|f1' F1][|f2' F2] //. +case: ifP=>// D _; case: ifP=>E1; last first. +- rewrite fnd_fcat; case: ifP=>// E2. + by rewrite fnd_supp ?E1 // fnd_supp ?E2. +suff E2: k \notin supp f2' by rewrite fnd_fcat (negbTE E2). +by case: disjP D E1=>// H _; apply: H. +Qed. + +Lemma findUnR k f1 f2 : + valid (f1 \+ f2) -> + find k (f1 \+ f2) = if k \in dom f2 then find k f2 else find k f1. +Proof. by rewrite joinC=>D; rewrite findUnL. Qed. + +Lemma find_eta f1 f2 : + valid f1 -> valid f2 -> + (forall k, find k f1 = find k f2) -> f1 = f2. +Proof. +rewrite !umEX /UM.valid /UM.find -{2 3}[f1]tfE -{2 3}[f2]tfE. +case: (UMC_from f1) (UMC_from f2)=>[|f1' F1][|f2' F2] // _ _ H. +by rewrite !umEX; apply/fmapP=>k; move: (H k); rewrite !umEX. +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) (C : pred K) (U : union_map C unit) (f1 f2 : U) : + dom_eq f1 f2 -> f1 = f2. +Proof. +case/domeqP=>V E; case V1 : (valid f1); last first. +- by move: V1 (V1); rewrite {1}V; do 2![move/negbT/invalidE=>->]. +apply: find_eta=>//; first by rewrite -V. +move=>k; case D : (k \in dom f1); move: D (D); rewrite {1}E. +- by do 2![case: dom_find=>// [[_ _]]]. +by do 2![case: dom_find=>// _]. +Qed. + +(**********) +(* assocs *) +(**********) + +Section AssocsLemmas. +Variables (K : ordType) (C : pred K) (V : Type) (U : union_map C V). +Implicit Type f : U. + +Lemma assocs_valid k f : k \In assocs f -> valid f. +Proof. by rewrite !umEX; case: (UMC_from f). Qed. + +Lemma assocs0 : assocs (Unit : U) = [::]. +Proof. by rewrite !umEX. Qed. + +Lemma assocs_undef : assocs (undef : U) = [::]. +Proof. by rewrite !umEX. Qed. + +Lemma assocsF f x : + assocs (free f x) = filter (fun kv => kv.1 != x) (assocs f). +Proof. by rewrite !umEX /UM.assocs; case: (UMC_from f)=>//=; case. Qed. + +Lemma assocs_perm f1 f2 : + valid (f1 \+ f2) -> perm (assocs (f1 \+ f2)) (assocs f1 ++ assocs f2). +Proof. +rewrite !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 assocs_dom f : dom f = map fst (assocs f). +Proof. by rewrite !umEX; case: (UMC_from f). Qed. + +Lemma size_assocs f : size (assocs f) = size (dom f). +Proof. by rewrite assocs_dom size_map. Qed. + +End AssocsLemmas. + +Lemma uniq_assocs K C (V : eqType) (U : @union_map K C V) (f : U) : + uniq (assocs f). +Proof. +rewrite !umEX /UM.assocs /=; case: (UMC_from f)=>[|[s H _]] //=. +by move/(sorted_uniq (@trans K) (@irr K)): H; apply: map_uniq. +Qed. + +Lemma assocs_map K C (V : Type) (U : @union_map K C V) (f : U) k v1 v2 : + (k, v1) \In assocs f -> (k, v2) \In assocs f -> v1 = v2. +Proof. +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. + +(*********************************) +(* Interaction of subset and dom *) +(*********************************) + +Section Interaction. +Variables (K : ordType) (C : pred K) (A : Type) (U : union_map C A). +Implicit Types (x y a b : U) (p q : pred K). + +(* Some equivalent forms for subset with dom *) + +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 *) + +Lemma subset_disjE p q : {subset p <= predC q} <-> [predI p & q] =1 pred0. +Proof. +split=>H a /=; first by case X: (a \in p)=>//; move/H/negbTE: X. +by move=>D; move: (H a); rewrite inE /= D; move/negbT. +Qed. + +Lemma valid_subdom x y : valid (x \+ y) -> {subset dom x <= [predC dom y]}. +Proof. by case: validUn. Qed. + +Lemma subdom_valid x y : + {subset dom x <= [predC dom 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_validL x y a : + valid a -> + valid (x \+ y) -> {subset dom a <= dom x} -> valid (a \+ y). +Proof. +rewrite !validUnAE=>-> /and3P [_ -> D] S. +by apply: sub_all D=>z /(contra (S z)). +Qed. + +End Interaction. + + +(****************************) +(* Generic points-to lemmas *) +(****************************) + +Section PointsToLemmas. +Variables (K : ordType) (C : pred K) (V : Type) (U : union_map 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 /ptsx !umEX /UM.pts /UM.upd; case: decP. Qed. + +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 /ptsx !umEX /= /UM.dom /supp /UM.pts /UM.upd /UM.empty /=. +by case D : (decP _)=>[a|a] /=; rewrite ?a ?(introF idP a). +Qed. + +(* a weaker legacy version of 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) = C k. +Proof. by rewrite ptsU validU valid_unit andbT. Qed. + +Lemma domPtV k v : valid (pts k v : U) -> k \in dom (pts k v : U). +Proof. by move=>D; rewrite domPtK -(validPt k v) D inE. Qed. + +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) && 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 & 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 : + C k2 -> + free (pts k2 v) k1 = if k1 == k2 then Unit else pts k2 v :> U. +Proof. by move=>N; rewrite ptsU freeU free0 N. Qed. + +Lemma freePt k v : + C k -> free (pts k v : U) k = Unit. +Proof. by move=>N; rewrite freePt2 // eq_refl. Qed. + +Lemma cancelPt k v1 v2 : + valid (pts k v1 : U) -> pts k v1 = pts k v2 :> U -> v1 = v2. +Proof. by rewrite validPt !ptsU; apply: upd_inj. Qed. + +Lemma cancelPt2 k1 k2 v1 v2 : + valid (pts k1 v1 : U) -> + pts k1 v1 = pts k2 v2 :> U -> (k1, v1) = (k2, v2). +Proof. +move=>H E; have : dom (pts k1 v1 : U) = dom (pts k2 v2 : U) by rewrite E. +rewrite !domPtK -(validPt _ v1) -(validPt _ v2) -E H. +by case=>E'; rewrite -{k2}E' in E *; rewrite (cancelPt H E). +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 um_unitbPt k v : unitb (pts k v : U) = false. +Proof. by rewrite ptsU um_unitbU. Qed. + +(* assocs *) + +Lemma assocsPt k v : + assocs (pts k v : U) = + if C k then [:: (k, v)] else [::]. +Proof. +rewrite /ptsx !umEX /UM.assocs/UM.pts /=. +by case E: decP=>[a|a] /=; [rewrite a | case: ifP]. +Qed. + +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. +rewrite /ptsx !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 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 /ptsx !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. + +(* valid *) + +Lemma validPtUn k v 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) //. + by rewrite domPtK H1 inE. +- by rewrite validPt; move/negbTE=>->. +- by move/negbTE=>->; rewrite andbF. +rewrite domPtK=>x; case: ifP=>// _. +by rewrite inE=>/eqP ->->; rewrite andbF. +Qed. + +Lemma validUnPt k v 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 validPtUnI v1 k f : + valid (pts k v1 \+ f) -> [&& C k, valid f & k \notin dom f]. +Proof. by rewrite validPtUn. 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 validPtUn_cond k v f : valid (pts k v \+ f) -> C k. +Proof. by rewrite validPtUn; case/and3P. 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. + +Lemma validUnPtD k v f : valid (f \+ pts k v) -> k \notin dom f. +Proof. by rewrite joinC; apply: validPtUnD. Qed. + +(* dom *) + +Lemma domPtUn k v f : + dom (pts k v \+ f) =i + [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: (C k). +Qed. + +Lemma domUnPt k v f : + dom (f \+ pts k v) =i + [pred x | valid (f \+ pts k v) & (k == x) || (x \in dom f)]. +Proof. by rewrite joinC; apply: domPtUn. Qed. + +Lemma domPtUnE k v f : k \in dom (pts k v \+ f) = valid (pts k v \+ f). +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: ord_sorted_eq=>//=. +- 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: ord_sorted_eq=>//=. +- 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. + +(* find *) + +Lemma findPtUn k v f : + valid (pts k v \+ f) -> find k (pts k v \+ f) = Some v. +Proof. +move=>V'; rewrite findUnL // domPt inE. +by rewrite ptsU findU eq_refl valid_unit (validPtUn_cond V'). +Qed. + +Lemma findUnPt k v f : + valid (f \+ pts k v) -> find k (f \+ pts k v) = Some v. +Proof. by rewrite joinC; apply: findPtUn. Qed. + +Lemma findPtUn2 k1 k2 v f : + valid (pts k2 v \+ f) -> + find k1 (pts k2 v \+ f) = + if k1 == k2 then Some v else find k1 f. +Proof. +move=>V'; rewrite findUnL // domPt inE eq_sym. +by rewrite findPt2 (validPtUn_cond V') andbT /=; case: ifP=>// ->. +Qed. + +Lemma findUnPt2 k1 k2 v f : + valid (f \+ pts k2 v) -> + find k1 (f \+ pts k2 v) = + if k1 == k2 then Some v else find k1 f. +Proof. by rewrite joinC; apply: findPtUn2. Qed. + +(* free *) + +Lemma freePtUn2 k1 k2 v f : + valid (pts k2 v \+ 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. +case: eqP=>/= E; first by rewrite E unitL (dom_free (validPtUnD V')). +by case: ifP=>// N; rewrite dom_free // N. +Qed. + +Lemma freeUnPt2 k1 k2 v f : + valid (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 (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 (f \+ pts k v) k = f. +Proof. by rewrite joinC; apply: freePtUn. Qed. + +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. +Proof. +case V1 : (valid (pts k v2 \+ f)). +- by rewrite updUnL domPt inE eq_refl updPt (validPtUn_cond V1). +have V2 : valid (pts k v1 \+ f) = false. +- by rewrite !validPtUn in V1 *. +move/negbT/invalidE: V1=>->; move/negbT/invalidE: V2=>->. +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. + +Lemma upd_eta k v f : upd k v f = pts k v \+ free f k. +Proof. +case: (normalP f)=>[->|W]; first by 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. + +Lemma upd_eta2 k v f : k \notin dom f -> upd k v f = pts k v \+ f. +Proof. by move=>D; rewrite upd_eta freeND. Qed. + +(* um_unitb *) + +Lemma um_unitbPtUn k v f : unitb (pts k v \+ f) = false. +Proof. by rewrite um_unitbUn um_unitbPt. Qed. + +Lemma um_unitbUnPt k v f : unitb (f \+ pts k v) = false. +Proof. by rewrite joinC; apply: um_unitbPtUn. Qed. + +(* undef *) + +Lemma pts_undef k v : pts k v = undef :> U <-> ~~ C k. +Proof. by rewrite -(validPt k v) invalidE. Qed. + +Lemma pts_undef2 k v1 v2 : pts k v1 \+ pts k v2 = undef :> U. +Proof. +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 f k. +Proof. +case: dom_find=>// v E1 E2 _; exists v; split=>//. +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 f k. +Proof. by move=>E; case/um_eta: (find_some E)=>v' []; rewrite E; case=><-. Qed. + +Lemma cancel k v1 v2 f1 f2 : + valid (pts k v1 \+ f1) -> + pts k v1 \+ f1 = pts k v2 \+ f2 -> [/\ v1 = v2, valid f1 & f1 = f2]. +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) (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 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 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 f1 k2 = free f2 k1. +- by rewrite -E''; rewrite E' joinCA in E; case/(cancel V1): E. +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 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 /ptsx !umEX => H1 H2 H3 f; rewrite -[f]tfE. +apply: (UM.base_indf (P := (fun b => P (UMC_to b))))=>//. +move=>k v g H V1; move: (H3 k v _ H); rewrite !umEX. +by apply. +Qed. + +(* backward progressing over keys *) +Lemma um_indb (P : U -> Prop) : + 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 /ptsx !umEX => H1 H2 H3 f; rewrite -[f]tfE. +apply: (UM.base_indb (P := (fun b => P (UMC_to b))))=>//. +move=>k v g H V1; move: (H3 k v _ H); rewrite !umEX. +by apply. +Qed. + +(* validity holds pairwise *) +Lemma um_valid3 f1 f2 f3 : + valid (f1 \+ f2 \+ f3) = + [&& valid (f1 \+ f2), valid (f2 \+ f3) & valid (f1 \+ f3)]. +Proof. +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 : + 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 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)). +- by rewrite -(validPt _ v) -E -joinA in X. +rewrite -[pts k v]unitR -joinA in E. +move/(cancel2 W): E; case: ifP. +- by move/eqP=>-> [->] /umap0E [->->]; rewrite unitR; left. +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. + +Prenex Implicits validPtUn_cond findPt_inv um_eta2 um_contra. +Prenex Implicits validPtUnD validUnPtD cancelPt cancelPt2 um_prime. + +(* dom_eq *) + +Section DomeqPtLemmas. +Variables (K : ordType) (C1 C2 : pred K) (V1 V2 : Type). +Variables (U1 : union_map C1 V1) (U2 : union_map C2 V2). + +Lemma domeqPt (k : K) (v1 : V1) (v2 : V2) : + C1 k = C2 k -> + dom_eq (pts k v1 : U1) (pts k v2 : U2). +Proof. by move=>E; apply/domeqP; rewrite !validPt !domPtK E. Qed. + +Lemma domeqPtUn k v1 v2 (f1 : U1) (f2 : U2) : + C1 k = C2 k -> + dom_eq f1 f2 -> dom_eq (pts k v1 \+ f1) (pts k v2 \+ f2). +Proof. by move=>E; apply: domeqUn=>//; apply: domeqPt. Qed. + +Lemma domeqUnPt k v1 v2 (f1 : U1) (f2 : U2) : + C1 k = C2 k -> + dom_eq f1 f2 -> dom_eq (f1 \+ pts k v1) (f2 \+ pts k v2). +Proof. by rewrite (joinC f1) (joinC f2); apply: domeqPtUn. Qed. + +End DomeqPtLemmas. + +#[export] +Hint Resolve domeqPt domeqPtUn domeqUnPt : core. + +(* decidable equality for umaps *) + +Section EqPtLemmas. +Variables (K : ordType) (C : pred K) (V : eqType). +Variables (U : union_map C V). +Notation Ue := + (Equality.pack_ (Equality.Mixin (union_map_eqP (U:=U)))). + +Lemma umPtPtE (k1 k2 : K) (v1 v2 : V) : + (pts k1 v1 == pts k2 v2 :> Ue) = + if C k1 then [&& C k2, k1 == k2 & v1 == v2] else ~~ C k2. +Proof. +case D1 : (C k1); last first. +- case D2 : (C k2); rewrite pts_condN ?D1 //. + - by apply/eqP/nesym; rewrite pts_undef D2. + by rewrite pts_condN ?D2 // eq_refl. +rewrite -(validPt U k1 v1) -(validPt U k2 v2) in D1 *. +apply/idP/idP. +- by case/eqP/(cancelPt2 D1)=><-<-; rewrite D1 !eq_refl. +by case/and3P=>_ /eqP -> /eqP ->. +Qed. + +Lemma umPtPtEq (k1 k2 : K) (v1 v2 : V) : + (pts k1 v1 == pts k2 v2 :> Ue) = + [&& C k1 == C k2, k1 == k2 & v1 == v2] || + (~~ C k1 && ~~ C k2). +Proof. +rewrite umPtPtE. +case D1 : (C k1); case D2 : (C k2)=>//=. +- by rewrite orbF. +by rewrite orbT. +Qed. + +Lemma umPt0E (k : K) (v : V) : (pts k v == Unit :> Ue) = false. +Proof. by apply/eqP=>/unitbP; rewrite um_unitbPt. Qed. + +Lemma um0PtE (k : K) (v : V) : (Unit == pts k v :> Ue) = false. +Proof. by rewrite eq_sym umPt0E. Qed. + +Lemma umPtUndefE (k : K) (v : V) : (pts k v == undef :> Ue) = ~~ C k. +Proof. by apply/idP/idP=>[|H]; [move/eqP/pts_undef|apply/eqP/pts_undef]. Qed. + +Lemma umUndefPtE (k : K) (v : V) : (undef == pts k v :> Ue) = ~~ C k. +Proof. by rewrite eq_sym umPtUndefE. Qed. + +Lemma umUndef0E : (undef == Unit :> Ue) = false. +Proof. by apply/eqP/undef0. Qed. + +Lemma um0UndefE : (Unit == undef :> Ue) = false. +Proof. by rewrite eq_sym umUndef0E. Qed. + +Lemma umPtUE (k : K) (v : V) f : (pts k v \+ f == Unit :> Ue) = false. +Proof. by apply/eqP=>/umap0E/proj1/unitbP; rewrite um_unitbPt. Qed. + +Lemma umUPtE (k : K) (v : V) f : (f \+ pts k v == Unit :> Ue) = false. +Proof. by rewrite joinC umPtUE. Qed. + +Lemma umPtUPtE (k1 k2 : K) (v1 v2 : V) f : + pts k1 v1 \+ f == pts k2 v2 :> Ue = + if C k1 then + if C k2 then [&& k1 == k2, v1 == v2 & unitb f] + else ~~ valid (pts k1 v1 \+ f) + else ~~ C k2. +Proof. +case D1 : (C k1); last first. +- rewrite pts_condN ?D1 //= undef_join eq_sym. + case D2 : (C k2). + - by apply/eqP=>/=; rewrite pts_undef D2. + by rewrite pts_condN ?D2 // eq_refl. +case D2 : (C k2); last first. +- rewrite (pts_condN U v2) ?D2 //=; apply/idP/idP. + - by move/eqP=>/undefbP; rewrite undefNV. + by move=>D; apply/eqP/undefbP; rewrite undefNV. +apply/idP/idP; last first. +- by case/and3P=>/eqP -> /eqP -> /unitbP ->; rewrite unitR. +case/eqP/(um_prime D2); case. +- move/(cancelPt2 _); rewrite validPt=>/(_ D1) [->->->]. + by rewrite !eq_refl unitb0. +by move/unitbP; rewrite um_unitbPt. +Qed. + +Lemma umPtPtUE (k1 k2 : K) (v1 v2 : V) f : + pts k1 v1 == pts k2 v2 \+ f :> Ue = + if C k2 then + if C k1 then [&& k1 == k2, v1 == v2 & unitb f] + else ~~ valid (pts k2 v2 \+ f) + else ~~ C k1. +Proof. by rewrite eq_sym umPtUPtE (eq_sym k1) (eq_sym v1). Qed. + +Lemma umUPtPtE (k1 k2 : K) (v1 v2 : V) f : + f \+ pts k1 v1 == pts k2 v2 :> Ue = + if C k1 then + if C k2 then [&& k1 == k2, v1 == v2 & unitb f] + else ~~ valid (pts k1 v1 \+ f) + else ~~ C k2. +Proof. by rewrite joinC umPtUPtE. Qed. + +Lemma umPtUPt2E (k1 k2 : K) (v1 v2 : V) f : + pts k1 v1 == f \+ pts k2 v2 :> Ue = + if C k2 then + if C k1 then [&& k1 == k2, v1 == v2 & unitb f] + else ~~ valid (pts k2 v2 \+ f) + else ~~ C k1. +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 U, @unitR U), (validPt, @valid_unit U))), + (((eq_refl, unitb0), + (um_unitbPt, undef_join)), (join_undef, unitb_undef))). +End EqPtLemmas. + + +(*****************) +(* Other notions *) +(*****************) + +(************) +(* um_preim *) +(************) + +Section PreimDefLemmas. +Variables (K : ordType) (C : pred K) (V : Type) (U : union_map C V). +Implicit Types (k : K) (v : V) (f : U) (p : pred V). + +Definition um_preim p f k := oapp p false (find k f). + +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. + +Lemma umpreim_cond p f k : um_preim p f k -> C k. +Proof. +rewrite /um_preim; case E: (find k f)=>[v|] //= _. +by move/find_some/dom_cond: E. +Qed. + +Lemma umpreimPt p k v : + C k -> + um_preim p (pts k v) =1 [pred x | (x == k) && p v]. +Proof. +move=>Hk x; rewrite /um_preim /= findPt2. +by case: eqP=>//= _; rewrite Hk. +Qed. + +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. + +Lemma umpreim_pred0 f : um_preim pred0 f =1 pred0. +Proof. by move=>x; rewrite /um_preim; by case: (find x f). Qed. + +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. + +End PreimDefLemmas. + + +(****************************) +(* map membership predicate *) +(****************************) + +Section MapMembership. +Variables (K : ordType) (C : pred K) (V : Type) (U : union_map C V). +Implicit Type f : U. + +Definition Mem_UmMap f := + fun x : K * V => valid f /\ [pcm pts x.1 x.2 <= f]. + +(* DEVCOMMENT *) +(* is this needed? *) +(* Coercion Pred_of_Mem_UmMap (f : U) : Pred_Class := [eta Mem_UmMap f]. *) +(* /DEVCOMMENT *) +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. + +Lemma In_undef x : x \In (undef : U) -> False. +Proof. by case; rewrite valid_undef. Qed. + +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. + +(* for inlined rewriting *) +Lemma In_findE k v f : (k, v) \In f -> find k f = Some v. +Proof. by move/In_find. 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 InPtE k v f : + C k -> f = pts k v -> (k, v) \In f. +Proof. by move=>W ->; apply: In_condPt W. 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_valid x f : x \In f -> valid f. +Proof. by case. Qed. + +Lemma In_domN x f : x \In f -> dom f <> [::]. +Proof. by move/In_dom=>H; apply/eqP/dom0NP; exists x.1. Qed. + +Lemma In_unitN x f : x \In f -> f <> Unit. +Proof. by move/[swap]=>-> /In0. Qed. + +Lemma In_inj_fun k1 k2 v f : + {in dom f, injective (fun x => find x f)} -> + (k1, v) \In f -> (k2, v) \In f -> k1 = k2. +Proof. +move=>H H1 H2; apply: H. +- by move/In_dom: H1. +by move/In_find: H1 H2=>-> /In_find ->. +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. + +(* this is just find_none stated as reflection *) +Lemma In_findN k f : reflect (find k f = None) (k \notin dom f). +Proof. +case: In_domX=>H; constructor. +- by case: H=>x /In_find ->. +by apply/find_none/negP=>/In_domX [v X]; apply: H; exists v. +Qed. + +Lemma In_findNE k f : k \notin dom f -> find k f = None. +Proof. by move/In_findN. 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. + +(* a frequently used pattern *) +Lemma InPtUnEL k v f f' : + valid f -> f = pts k v \+ f' -> (k, v) \In f. +Proof. by move/[swap] ->; apply: InPtUnL. 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 -> + [/\ [&& C k, valid f & k \notin dom 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). +rewrite (validPtUnI (W v)); 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 In_assocs 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. + +(* if we have equality of domains, we can get rid of one direction *) +(* in the hypothesis in umem_eq *) + +Lemma umem_eqD f1 f2 : + valid f1 -> valid f2 -> + dom f1 =i dom f2 -> + (forall x, x \In f1 -> x \In f2) -> f1 = f2. +Proof. +move=>V1 V2 E H; apply: umem_eq=>//; case=>k v; split; first by apply: H. +move=>H2; move: (In_dom H2); rewrite -E /= =>/In_domX [w H1]. +by move/H/(In_fun H2): (H1)=>->. +Qed. + +Lemma umem_eq_assocs f1 f2 : + valid f1 -> valid f2 -> + assocs f1 = assocs f2 -> + f1 = f2. +Proof. +move=>V1 V2 E; apply: umem_eqD V1 V2 _ _. +- by rewrite !assocs_dom E. +by move=>x /In_assocs; rewrite E=>/In_assocs. +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. + +(* different format that avoids conditionals *) +Lemma In_U x k v f : + x \In upd k v f <-> + valid (upd k v f) /\ + [\/ x.1 == k /\ x.2 = v | + x.1 != k /\ x \In f]. +Proof. +rewrite InU; split; case=>W H; split=>//. +- by case: eqP H; [left | right]. +by case: H=>[[->]//|[/negbTE ->]]. +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 InPtUnEL. + +Section MorphMembership. +Variables (K : ordType) (C : pred K) (V : Type). +Variables (U1 : pcm) (U2 : union_map C V). + +Section PlainMorph. +Variable f : pcm_morph U1 U2. + +Lemma InpfL e (x y : U1) : + valid (x \+ y) -> sep f x y -> e \In f x -> e \In f (x \+ y). +Proof. by move=>W H1 H2; rewrite pfjoin //=; apply: InL=>//; rewrite pfV2. Qed. + +Lemma InpfR e (x y : U1) : + valid (x \+ y) -> sep f x y -> e \In f y -> e \In f (x \+ y). +Proof. by move=>W H1 H2; rewrite pfjoin //=; apply: InR=>//; rewrite pfV2. Qed. + +Lemma InpfUn e (x y : U1) : + valid (x \+ y) -> sep f x y -> + e \In f (x \+ y) <-> e \In f x \/ e \In f y. +Proof. +move=>W H; rewrite pfjoin //; split; first by case/InUn; eauto. +by case; [apply: InL|apply: InR]; rewrite pfV2. +Qed. + +End PlainMorph. + +Section FullMorph. +Variable f : full_pcm_morph U1 U2. + +Lemma InpfLT e (x y : U1) : valid (x \+ y) -> e \In f x -> e \In f (x \+ y). +Proof. by move=>W; apply: InpfL. Qed. + +Lemma InpfRT e (x y : U1) : valid (x \+ y) -> e \In f y -> e \In f (x \+ y). +Proof. by move=>W; apply: InpfR. Qed. + +Lemma InpfUnT e (x y : U1) : + valid (x \+ y) -> e \In f (x \+ y) <-> e \In f x \/ e \In f y. +Proof. by move=>W; apply: InpfUn. Qed. +End FullMorph. + +End MorphMembership. + +(*********) +(* range *) +(*********) + +Section Range. +Variables (K : ordType) (C : pred K) (V : Type) (U : union_map 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. + +Lemma rangePtK k v : C k -> range (pts k v) = [:: v]. +Proof. by move=>C'; rewrite /range assocsPt C'. 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. + +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 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: (normalP f)=>[->|W]. +- by rewrite range_undef; 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/In_assocs H1 /mem_seqP/In_assocs 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/In_assocs/mem_seqP: H1=>->; move/In_assocs/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 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 rangeF k (f : U) : {subset range (free f k) <= range f}. +Proof. +case: (normalP f)=>[->|W]; first by rewrite free_undef !range_undef. +case D: (k \in dom f); last by move/negbT/dom_free: D=>->. +case: (um_eta D) W=>x [_] {1 3}-> Vf p. +by rewrite rangePtUn inE Vf=>->; rewrite orbT. +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: (normalP f)=>[->|W]; first by 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 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/In_assocs 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/In_assocs/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 k1 k2 v f : + um_mono f -> (k1, v) \In f -> (k2, v) \In f -> k1 = k2. +Proof. +move/ummonoP=>M H1 H2; case: (ordP k1 k2). +- by move/(M _ _ _ _ H1 H2); rewrite irr. +- by move/eqP. +by move/(M _ _ _ _ H2 H1); rewrite irr. +Qed. + +Lemma In_mono_range v f1 f2 : + valid (f1 \+ f2) -> um_mono (f1 \+ f2) -> + v \in range f1 -> v \in range f2 -> false. +Proof. +move=>W M /mem_seqP/In_rangeX [k1 H1] /mem_seqP/In_rangeX [k2 H2]. +have H1' : (k1, v) \In f1 \+ f2 by apply/InL. +have H2' : (k2, v) \In f1 \+ f2 by apply/InR. +rewrite -{k2 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 ummono_inj_find f : + um_mono f -> {in dom f & predT, injective (fun x => find x f)}. +Proof. +move/ummono_leP=>H k1 k2 /In_domX [x1 F1] _ E. +have /In_domX [x2 F2] : k2 \in dom f. +- by case: (dom_find k2) F1 E=>// _ /In_find ->. +move/In_find: (F1) E=>->; move/In_find: (F2)=>-> [?]; subst x2. +move: (H _ _ _ _ F1 F2) (H _ _ _ _ F2 F1); rewrite orefl=>{H} H1 H2. +case: (equivP idP H1) (@oantisym K k1 k2)=>// _. +by case: (equivP idP H2)=>// _; apply. +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. + +Lemma ummonoF f x : um_mono f -> um_mono (free f x). +Proof. +move/ummonoP=>X; apply/ummonoP=>k k' v v'. +by case/InF=>_ _ F /InF [_ _]; apply: X F. +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 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 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. + +Lemma In_umfoldl z0 d (f : U) (k : K) (v : V) : + (k, v) \In f -> [pcm a Unit k v <= um_foldl a z0 d f]. +Proof. +move=>H; move: (H); rewrite (In_eta H); case=>W _. +by rewrite umfoldlUn_frame // umfoldlPt (validPtUn_cond W) /= -joinA; eauto. +Qed. + +End PCMFold. + +(* Fold when the function produces a map *) + +Section FoldMap. +Variables (K : ordType) (C : pred K) (V : Type) (U : union_map C V). +Variable (a : U -> K -> V -> U). +Hypothesis frame : forall x y k v, a (x \+ y) k v = a x k v \+ y. + +Lemma In_umfoldlMX (f : U) (k : K) (v : V) : + (k, v) \In um_foldl a Unit undef f -> + exists k1 (v1 : V), (k, v) \In a Unit k1 v1 /\ (k1, v1) \In f. +Proof. +elim/um_indf: f. +- by rewrite umfoldl_undef=>/In_undef. +- by rewrite umfoldl0=>/In0. +move=>k1 v1 f IH W P. +rewrite umfoldlUn_frame // umfoldlPt (validPtUn_cond W) unitR. +case/InUn=>[H|/IH [k2][v2][H2 R2]]; first by exists k1, v1. +by exists k2, v2; split=>//; apply/InPtUnE=>//; right. +Qed. + +Lemma In_umfoldlM (f : U) (k : K) (v : V) k1 v1 : + valid (um_foldl a Unit undef f) -> + (k, v) \In a Unit k1 v1 -> (k1, v1) \In f -> + (k, v) \In um_foldl a Unit undef f. +Proof. +move=>W H /(In_umfoldl frame Unit undef) [x E]. +by move: E W=>-> W; apply/InL. +Qed. + +End FoldMap. + +(**************) +(* Map prefix *) +(**************) + +Section MapPrefix. +Variables (K : ordType) (C : pred K) (V : Type). +Variables (U : union_map C V). + +Definition um_prefix (h1 h2 : U) := Prefix (assocs h1) (assocs h2). + +Lemma umpfx_undef h : um_prefix undef h. +Proof. +by rewrite /um_prefix assocs_undef; apply/PrefixE; exists (assocs h). +Qed. + +Lemma umpfx0 h : um_prefix Unit h. +Proof. +by rewrite /um_prefix assocs0; apply/PrefixE; exists (assocs h). +Qed. + +Lemma umpfxD h1 h2 : um_prefix h1 h2 -> Prefix (dom h1) (dom h2). +Proof. +by case/PrefixE=>x E; rewrite !assocs_dom E map_cat; apply: Prefix_cat. +Qed. + +(* we next need a helper lemma, which should be in assocs section *) +(* but couldn't be proved there, because um_indf appears later *) +Lemma assocsUn (h1 h2 : U) : + valid (h1 \+ h2) -> + (forall k1 k2, k1 \in dom h1 -> k2 \in dom h2 -> ord k1 k2) -> + assocs (h1 \+ h2) = assocs h1 ++ assocs h2. +Proof. +move: h1 h2; apply: um_indf=>[h1|h2 W H| + k v f IH W1 /(order_path_min (@trans _)) P h2 W2 H]. +- by rewrite undef_join valid_undef. +- by rewrite assocs0 unitL. +rewrite -joinA in W2; rewrite -joinA !assocsPtUn //= ?IH //. +- by rewrite (validR W2). +- by move=>k1 k2 K1 K2; apply: H=>//; rewrite domPtUn inE K1 orbT W1. +apply/allP=>x; rewrite domUn inE (validR W2) /=. +by case/orP=>[/(allP P)//|]; apply: H; rewrite domPtUnE. +Qed. + +Lemma umpfxI h1 h2 : + valid (h1 \+ h2) -> + (forall x y, x \in dom h1 -> y \in dom h2 -> ord x y) -> + um_prefix h1 (h1 \+ h2). +Proof. by move=>W X; rewrite /um_prefix assocsUn //; apply: Prefix_cat. Qed. + +Lemma umpfxE h1 h : + valid h1 -> + um_prefix h1 h -> + exists2 h2, h = h1 \+ h2 & + forall x y, x \in dom h1 -> y \in dom h2 -> ord x y. +Proof. +move=>V1; case: (normalP h)=>[->|W]. +- by exists undef; rewrite ?join_undef ?dom_undef. +move: h1 h V1 W; apply: um_indf=>[h|h _ W| +k v h1 IH W1 /(order_path_min (@trans _)) P h _ W]. +- by rewrite valid_undef. +- by exists h; [rewrite unitL | rewrite dom0]. +case/PrefixE=>h2'; rewrite assocsPtUn //= => Eh. +set h' := free h k; have W' : valid h' by rewrite validF. +have : um_prefix h1 h'. +- rewrite /um_prefix assocsF Eh /= eq_refl filter_cat /=. + have : forall kv, kv \In assocs h1 -> kv.1 != k. + - move=>kv /In_assocs/In_dom H; move/allP/(_ _ H): P. + by case: eqP=>// ->; rewrite irr. + by move/eq_In_filter=>->; rewrite filter_predT; apply: Prefix_cat. +case/(IH _ (validR W1) W')=>h2 Eh' H; exists h2. +- rewrite -joinA -Eh'; apply/um_eta2/In_find/In_assocs. + by rewrite Eh In_cons; left. +move=>x y; rewrite domPtUn inE W1 /=; case/orP; last by apply: H. +move/eqP=><-{x} Dy. +have : y \in dom h' by rewrite Eh' domUn inE -Eh' W' /= Dy orbT. +rewrite domF inE; case: (k =P y)=>// /eqP N. +rewrite assocs_dom Eh /= inE eq_sym (negbTE N) /=. +case/mem_seqP/MapP; case=>a b X -> /=. +have {}P : path ord k (map fst (assocs h1 ++ h2')). +- by move: (sorted_dom h); rewrite assocs_dom Eh /=. +suff {X} : forall x, x \In assocs h1 ++ h2' -> ord k x.1 by move/(_ _ X). +apply/allPIn/(order_path_min (x:=(k,v)) (leT:=fun k x=>ord k.1 x.1)). +- by move=>???; apply: trans. +by rewrite -path_map. +Qed. + +End MapPrefix. + + +(********) +(* omap *) +(********) + +(* Combines map and filter by having filtering function return option. *) +(* This is very common form, so we also build structure for it. *) + +(* Definition and basic properties *) +Section OmapDefs. +Variables (K : ordType) (C : pred K) (V V' : Type). +Variables (U : union_map C V) (U' : union_map C V'). +Variable f : K * V -> option V'. + +Definition omap (x : U) : U' := + if valid x then + foldl (fun z kv => if f kv is Some v' then z \+ pts kv.1 v' else z) + Unit (assocs x) + else undef. + +Lemma omap0 : omap Unit = Unit. +Proof. by rewrite /omap valid_unit assocs0. Qed. + +Lemma omap_undef : omap undef = undef. +Proof. by rewrite /omap valid_undef. Qed. + +Lemma omapPt k v : + omap (pts k v) = + if C k then + if f (k, v) is Some w then pts k w else Unit + else undef. +Proof. +rewrite /omap validPt assocsPt; case D : (C k)=>//=. +by case: (f _)=>// w; rewrite unitL. +Qed. + +Lemma omapUn x1 x2 : + valid (x1 \+ x2) -> omap (x1 \+ x2) = omap x1 \+ omap x2. +Proof. +move=>W; rewrite /omap W (validL W) (validR W); set o := fun z kv => _. +have H (x y : U') kv : o (x \+ y) kv = o x kv \+ y. +- by rewrite /o; case: (f kv)=>// w; rewrite joinAC. +rewrite (foldl_perm H _ (assocs_perm W)) foldl_cat. +by rewrite joinC -(foldl_init H) unitL. +Qed. + +Lemma omapPtUn k v x : + omap (pts k v \+ x) = + if valid (pts k v \+ x) then + if f (k, v) is Some v' then pts k v' \+ omap x else omap x + else undef. +Proof. +case: ifP=>D; last first. +- by move/invalidE: (negbT D)=>->; rewrite omap_undef. +rewrite omapUn // omapPt (validPtUn_cond D). +by case: (f _)=>//; rewrite unitL. +Qed. + +Lemma omap_subdom x : {subset dom (omap x) <= dom x}. +Proof. +elim/um_indf: x=>[||k v x IH W P] t. +- by rewrite omap_undef dom_undef. +- by rewrite omap0 dom0. +rewrite omapPtUn W domPtUn W !inE /=. +case E : (f _)=>[b|]; last by move/IH=>->; rewrite orbT. +rewrite domPtUn inE; case/andP=>W2. +by case/orP=>[->//|/IH ->]; rewrite orbT. +Qed. + +Lemma valid_omap x : valid (omap x) = valid x. +Proof. +elim/um_indf: x=>[||k v x IH W P]. +- by rewrite omap_undef !valid_undef. +- by rewrite omap0 !valid_unit. +rewrite omapPtUn W; case E : (f _)=>[b|]; last by rewrite IH (validR W). +rewrite validPtUn (validPtUn_cond W) IH (validR W). +by apply: contra (notin_path P); apply: omap_subdom. +Qed. + +End OmapDefs. + +Arguments omap {K C V V' U U'}. + +(* Structure definition *) + +Definition omap_fun_axiom (K : ordType) (C : pred K) (V V' : Type) + (U : union_map C V) (U' : union_map C V') + (f : U -> U') (omf : K * V -> option V') := + f =1 omap omf. + +(* factory to use if full/norm/tpcm morphism property already proved *) +(* (omap_fun isn't binormal as it can drop timestamps) *) +HB.mixin Record isOmapFun_morph (K : ordType) (C : pred K) (V V' : Type) + (U : union_map C V) (U' : union_map C V') (f : U -> U') of + @Full_Norm_TPCM_morphism U U' f := { + omf_op : K * V -> option V'; + omfE_op : omap_fun_axiom f omf_op}. + +#[short(type=omap_fun)] +HB.structure Definition OmapFun (K : ordType) (C : pred K) (V V' : Type) + (U : union_map C V) (U' : union_map C V') := + {f of isOmapFun_morph K C V V' U U' f &}. + +Arguments omfE_op {K C V V' U U'}. +Arguments omf_op {K C V V' U U'} _ _ /. + +(* factory to use when full/norm/tpcm morphism property *) +(* hasn't been established *) +HB.factory Record isOmapFun (K : ordType) (C : pred K) (V V' : Type) + (U : union_map C V) (U' : union_map C V') (f : U -> U') := { + omf : K * V -> option V'; + omfE : omap_fun_axiom f omf}. + +HB.builders Context K C V V' U U' f of isOmapFun K C V V' U U' f. + +Lemma omap_fun_is_pcm_morph : pcm_morph_axiom relT f. +Proof. +split=>[|x y W _]; first by rewrite omfE omap0. +by rewrite !omfE -omapUn // valid_omap. +Qed. + +HB.instance Definition _ := + isPCM_morphism.Build U U' f omap_fun_is_pcm_morph. + +Lemma omap_fun_is_tpcm_morph : tpcm_morph_axiom f. +Proof. by rewrite /tpcm_morph_axiom omfE omap_undef. Qed. + +HB.instance Definition _ := + isTPCM_morphism.Build U U' f omap_fun_is_tpcm_morph. + +Lemma omap_fun_is_full_morph : full_pcm_morph_axiom f. +Proof. by []. Qed. + +HB.instance Definition _ := + isFull_PCM_morphism.Build U U' f omap_fun_is_full_morph. + +Lemma omap_fun_is_norm_morph : norm_pcm_morph_axiom f. +Proof. by move=>x /=; rewrite omfE valid_omap. Qed. + +HB.instance Definition _ := + isNorm_PCM_morphism.Build U U' f omap_fun_is_norm_morph. +HB.instance Definition _ := + isOmapFun_morph.Build K C V V' U U' f omfE. +HB.end. + +(* notation to hide the structure when projecting omf *) +Section OmapFunNotation. +Variables (K : ordType) (C : pred K) (V V' : Type). +Variables (U : union_map C V) (U' : union_map C V'). + +Definition omfx (f : omap_fun U U') of phantom (U -> U') f : + K * V -> option V' := omf_op f. + +Notation omf f := (omfx (Phantom (_ -> _) f)). + +Lemma omfE (f : omap_fun U U') : f =1 omap (omf f). +Proof. exact: omfE_op. Qed. +End OmapFunNotation. + +Notation omf f := (omfx (Phantom (_ -> _) f)). + +(* omap is omap_fun *) +Section OmapOmapFun. +Variables (K : ordType) (C : pred K) (V V' : Type). +Variables (U : union_map C V) (U' : union_map C V'). + +HB.instance Definition _ f := + isOmapFun.Build K C V V' U U' (omap f) (fun => erefl _). + +Lemma omf_omap f : omf (omap f) = f. Proof. by []. Qed. +End OmapOmapFun. + + +(* general omap_fun lemmas *) + +Section OmapFunLemmas. +Variables (K : ordType) (C : pred K) (V V' : Type). +Variables (U : union_map C V) (U' : union_map C V'). +Implicit Types (f : omap_fun U U') (x : U). + +(* different name for pfjoin on full morphisms *) +Lemma omfUn f x1 x2 : valid (x1 \+ x2) -> f (x1 \+ x2) = f x1 \+ f x2. +Proof. exact: pfjoinT. Qed. + +Lemma omfPtE f k v : + f (pts k v) = + if C k then + if omf f (k, v) is Some w then pts k w else Unit + else undef. +Proof. by rewrite omfE omapPt. Qed. + +Lemma omfPt f k v : + C k -> + f (pts k v) = if omf f (k, v) is Some w then pts k w else Unit. +Proof. by rewrite omfPtE=>->. Qed. + +Lemma omfPtUn f k v x : + f (pts k v \+ x) = + if valid (pts k v \+ x) then + if omf f (k, v) is Some v' then pts k v' \+ f x else f x + else undef. +Proof. +case: ifP=>X; last first. +- by move/invalidE: (negbT X)=>->; rewrite pfundef. +rewrite omfUn // omfPtE (validPtUn_cond X). +by case: (omf f _)=>//; rewrite unitL. +Qed. + +Lemma omfUnPt f k v x : + f (x \+ pts k v) = + if valid (x \+ pts k v) then + if omf f (k, v) is Some v' then f x \+ pts k v' else f x + else undef. +Proof. +rewrite joinC omfPtUn; case: ifP=>// W. +by case: (omf f _)=>// w; rewrite joinC. +Qed. + +(* membership *) +Lemma In_omfX f k v x : + (k, v) \In f x <-> + exists2 w, (k, w) \In x & omf f (k, w) = Some v. +Proof. +rewrite omfE; elim/um_indf: x k v=>[||k' v' x IH W P] k v. +- by rewrite pfundef; split=>[/In_undef|[w][]] //; rewrite valid_undef. +- by rewrite pfunit; split=>[/In0|[w] /In0]. +split=>[H|[w] H E]. +- move: (dom_valid (In_dom H)); rewrite omapPtUn W in H *. + case E : (omf f _) H=>[z|] H W1; last first. + - by case/IH: H=>w1; exists w1=>//; apply/InR. + move: H; rewrite InPtUnE //. + case=>[[->->]|]; first by exists v'. + by case/IH=>w; exists w=>//; apply/InR. +have : valid (omap (U':=U') (omf f) (pts k' v' \+ x)) by rewrite pfV. +rewrite omapPtUn W; move/(InPtUnE _ W): H E; case. +- by case=>->->-> V1; apply/InPtUnE=>//; left. +case: (omf _ (k', v'))=>[b|] H E V1; last by apply/IH; exists w. +by rewrite InPtUnE //; right; apply/IH; exists w. +Qed. + +(* one side of the In_omfX can destruct the existential *) +Lemma In_omf f k v w x : + (k, w) \In x -> omf f (k, w) = Some v -> (k, v) \In f x. +Proof. by move=>H1 H2; apply/In_omfX; exists w. Qed. + +(* dom *) +Lemma omf_subdom f x : {subset dom (f x) <= dom x}. +Proof. rewrite omfE; apply: omap_subdom. Qed. + +Arguments omf_subdom {f x}. + +Lemma omf_cond f x k : k \in dom (f x) -> C k. +Proof. by move/omf_subdom/dom_cond. Qed. + +Lemma omf_sorted f x : sorted ord (dom (f x)). +Proof. by apply: sorted_dom. Qed. + +Lemma path_omf f x k : path ord k (dom x) -> path ord k (dom (f x)). +Proof. +apply: subseq_path; first by apply: trans. +apply: (sorted_subset_subseq (ltT := ord)); last by apply: omf_subdom. +- by apply: irr. +- by apply: trans. +- by apply: sorted_dom. +by apply: sorted_dom. +Qed. + +Lemma In_dom_omfX f x k : + reflect (exists v, (k, v) \In x /\ omf f (k, v)) + (k \in dom (f x)). +Proof. +case: In_domX=>H; constructor. +- by case: H=>v /In_omfX [w H1 H2]; exists w; rewrite H2. +case=>v [X]; case Y: (omf _ _)=>[w|] // _. +by elim: H; exists w; apply/In_omfX; exists v. +Qed. + +Lemma dom_omfUn f x1 x2 : + dom (f (x1 \+ x2)) =i + [pred k | valid (x1 \+ x2) & (k \in dom (f x1)) || (k \in dom (f x2))]. +Proof. +move=>k; rewrite !(omfE f) inE. +case: (normalP (x1 \+ x2))=>[->|W]; first by rewrite pfundef dom_undef. +by rewrite pfjoinT //= domUn inE -pfjoinT //= pfV //= W. +Qed. + +Lemma dom_omfPtUn f k v x : + dom (f (pts k v \+ x)) =i + [pred t | valid (pts k v \+ x) & + (omf f (k, v) && (t == k)) || (t \in dom (f x))]. +Proof. +move=>t; rewrite dom_omfUn !inE; case D : (valid _)=>//=. +rewrite omfPtE (validPtUn_cond D); case E: (omf _ _)=>[a|] /=. +- by rewrite domPt inE (validPtUn_cond D) eq_sym. +by rewrite dom0. +Qed. + +(* validity *) + +Lemma In_omfV f kw x : + kw \In f x -> + valid x. +Proof. by case/In_omfX=>v /In_valid. Qed. + +(* stronger than morphism property as uses two different f's *) +Lemma valid_omfUn f1 f2 x1 x2 : + valid (x1 \+ x2) -> + valid (f1 x1 \+ f2 x2). +Proof. +move=>W; rewrite (omfE f1) (omfE f2) validUnAE !pfVE !(validE2 W) /=. +by apply/allP=>x /omf_subdom D2; apply/negP=>/omf_subdom/(dom_inNRX W D2). +Qed. + +Lemma valid_omfPtUn f k v x : + [&& C k, valid x & k \notin dom x] -> + valid (pts k v \+ f x). +Proof. +case/and3P=>H1 H2 H3; rewrite validPtUn H1 pfVE H2 /=. +by apply: contra (omf_subdom _) H3. +Qed. + +Lemma valid_omfUnPt f k v x : + [&& C k, valid x & k \notin dom x] -> + valid (f x \+ pts k v). +Proof. by move=>H; rewrite joinC; apply: valid_omfPtUn. Qed. + +(* range *) + +Lemma In_range_omapUn f x1 x2 v : + v \In range (f (x1 \+ x2)) <-> + valid (x1 \+ x2) /\ (v \In range (f x1) \/ v \In range (f x2)). +Proof. +split=>[|[D]]. +- case: (normalP (x1 \+ x2))=>[->|D]; last first. + - by rewrite pfjoinT // => /In_rangeUn. + by rewrite pfundef range_undef. +rewrite pfjoinT //; case. +- by apply/In_rangeL/pfV2I. +by apply/In_rangeR/pfV2I. +Qed. + +(* other interactions *) + +Lemma find_omf f k x : + find k (f x) = + if find k x is Some v then omf f (k, v) else None. +Proof. +case E1 : (find k _)=>[b|]. +- by move/In_find: E1=>/In_omfX [w] /In_find ->. +move/find_none/negP: E1=>E1. +case E2 : (find k x)=>[c|] //; move/In_find: E2=>E2. +case E3 : (omf f (k, c))=>[d|] //; elim: E1. +by apply/In_dom_omfX; exists c; rewrite E3. +Qed. + +Lemma omfF f k x : f (free x k) = free (f x) k. +Proof. +case: (normalP x)=>[->|D]; first by rewrite pfundef !free_undef pfundef. +apply: umem_eq. +- by rewrite pfVE validF. +- by rewrite validF pfVE. +case=>t v; split. +- case/In_omfX=>w /InF /= [_ N M E]. + apply/InF; rewrite validF pfVE D N; split=>//. + by apply/In_omfX; exists w. +case/InF=>/= _ N /In_omfX [v'] M E. +by apply/In_omfX; exists v'=>//; apply/InF; rewrite validF. +Qed. + +Lemma omfUE f k (v : V) (x : U) : + f (upd k v x) = + if C k then + if omf f (k, v) is Some v' then upd k v' (f x) + else free (f x) k + else undef. +Proof. +case: ifP=>// D; last by rewrite (upd_condN v x (negbT D)) pfundef. +case: (normalP x)=>[->|H]. +- by case: (omf _ _)=>[a|]; rewrite ?(upd_undef,free_undef,pfundef). +rewrite upd_eta omfPtUn validPtUn D validF H domF inE eq_refl /=. +case: (omf _ _)=>[xa|]; last by rewrite omfF. +by rewrite upd_eta omfF. +Qed. + +Lemma omfU f k (v : V) (x : U) : + C k -> + f (upd k v x) = + if omf f (k, v) is Some v' then upd k v' (f x) + else free (f x) k. +Proof. by move=>D; rewrite omfUE D. Qed. + +(* when mapped functions are equal *) + +Lemma eq_in_omf f1 f2 x : + f1 x = f2 x <-> + (forall kv, kv \In x -> omf f1 kv = omf f2 kv). +Proof. +have L h : (forall kv, kv \In h -> omf f1 kv = omf f2 kv) -> + f1 h = f2 h. +- elim/um_indf: h=>[||k v h IH W P H]; rewrite ?pfundef ?pfunit //. + by rewrite !omfPtUn W (H _ (InPtUnL W)) -IH // => kv /(InR W)/H. +split; last by apply: L. +elim/um_indf: x=>[_ kv /In_undef|_ kv /In0|k v x IH W P H kv] //. +have E t : find t (f1 (pts k v \+ x)) = find t (f2 (pts k v \+ x)). +- by rewrite H. +case/(InPtUnE _ W)=>[->|]. +- by move: (E k); rewrite !find_omf findPtUn. +apply: {kv} IH; apply: L; case=>k1 v1 X. +move: (E k1); rewrite !find_omf findPtUn2 //. +case: (k1 =P k) X =>[-> /In_dom|_ /In_find ->] //=. +by move/negbTE: (validPtUnD W)=>->. +Qed. + +(* if mapped function is total, we have some stronger lemmas *) + +Lemma dom_omf_some f x : + (forall kv, kv \In x -> omf f kv) -> dom (f x) = dom x. +Proof. +move=>H; apply/domE=>k; apply/idP/idP=>/In_domX [w]. +- by case/In_omfX=>v /In_dom. +by move/[dup]/H=>E X; apply/In_dom_omfX; exists w. +Qed. + +Lemma domeq_omf_some f x : + (forall kv, kv \In x -> omf f kv) -> dom_eq (f x) x. +Proof. by move/dom_omf_some=>E; rewrite /dom_eq E pfVE !eq_refl. Qed. + +(* if mapped function is total on x1, x2 *) +(* we don't need validity condition for union *) +Lemma omfUn_some f x1 x2 : + (forall k, k \In x1 -> omf f k) -> + (forall k, k \In x2 -> omf f k) -> + f (x1 \+ x2) = f x1 \+ f x2. +Proof. +move=>H1 H2; have Ev : valid (f x1 \+ f x2) = valid (x1 \+ x2). +- by rewrite !validUnAE !pfVE !dom_omf_some. +case: (normalP (x1 \+ x2)) Ev=>[->|D _]; last by apply: pfjoinT. +by rewrite pfundef=>/negbT/invalidE. +Qed. + +Lemma omf_predU f1 f2 x : + (forall kv, omf f1 kv = None \/ omf f2 kv = None) -> + f1 x \+ f2 x = + omap (fun kv => if omf f1 kv is Some v + then Some v else omf f2 kv) x. +Proof. +move=>E; rewrite (omfE f1) (omfE f2). +elim/um_indf: x=>[||k v x IH W /(order_path_min (@trans K)) P]. +- by rewrite !pfundef undef_join. +- by rewrite !pfunit unitL. +rewrite !omapPtUn W -IH. +case E1 : (omf f1 (k, v))=>[b1|]; case E2 : (omf f2 (k, v))=>[b2|] //. +- by move: (E (k, v)); rewrite E1 E2; case. +- by rewrite joinA. +by rewrite joinCA. +Qed. + +Lemma omf_unit f x : + reflect (valid x /\ forall kv, kv \In x -> omf f kv = None) + (unitb (f x)). +Proof. +case: unitbP=>H; constructor; last first. +- case=>D; elim/um_indf: x D H=>[||k v x IH W P _]; + rewrite ?valid_undef ?pfunit ?omfPtUn //= W. + case E : (omf f _)=>[a|]; first by move=>_ /(_ _ (InPtUnL W)); rewrite E. + by move=>H1 H2; apply: (IH (validR W) H1)=>kv H; apply: H2 (InR W H). +split; first by rewrite -(pfVE (f:=f)) /= H valid_unit. +elim/um_indf: x H=>[H kv /In_undef|H kv /In0|k v x IH W P /[swap] kv] //. +rewrite omfPtUn W; case E : (omf f _)=>[a|]. +- by move/unitbP; rewrite um_unitbPtUn. +by move=>H /(InPtUnE _ W) [->//|]; apply: IH H _. +Qed. + +Lemma omf_none f x : + valid x -> + (forall kv, kv \In x -> omf f kv = None) -> + f x = Unit. +Proof. by move=>D H; apply/unitbP/omf_unit. Qed. + +Lemma omf_noneR f x : + f x = Unit -> forall kv, kv \In x -> omf f kv = None. +Proof. by case/unitbP/omf_unit. Qed. + +End OmapFunLemmas. + +Arguments omf_subdom {K C V V' U U' f x}. +Arguments In_omf {K C V V' U U'} _ {k v w x}. + +(* omap_fun's with different ranges *) + +Section OmapFun2. +Variables (K : ordType) (C : pred K) (V V1 V2 : Type). +Variables (U : union_map C V) (U1 : union_map C V1) (U2 : union_map C V2). +Variables (f1 : omap_fun U U1) (f2 : omap_fun U U2). + +(* when one mapped function is included in other *) + +Lemma omf_dom_subseq (x : U) : + (forall kv, kv \In x -> omf f1 kv -> omf f2 kv) -> + subseq (dom (f1 x)) (dom (f2 x)). +Proof. +elim/um_indf: x=>[||k v x IH W P] H; +rewrite ?pfundef ?pfunit ?dom_undef ?dom0 //. +move: (W); rewrite validPtUn=>W'; rewrite !omfPtUn W. +have T : transitive (@ord K) by apply: trans. +have B : (k, v) \In pts k v \+ x by apply: InPtUnL. +case E1 : (omf f1 (k, v))=>[x1|]. +- have /(H _ B): (omf f1 (k, v)) by rewrite E1. + case: (omf f2 (k, v))=>// x2 _. + rewrite !domPtUnK //=; last first. + - by apply/allP=>? /In_dom_omfX [?][] /In_dom Y _; apply: path_mem Y. + - by rewrite valid_omfPtUn. + - by apply/allP=>? /In_dom_omfX [?][] /In_dom Y _; apply: path_mem Y. + - by rewrite valid_omfPtUn. + by rewrite eq_refl; apply: IH=>kx X; apply: H (InR _ _). +case E2 : (omf f2 (k, v))=>[x2|]; last by apply: IH=>kx X; apply: H (InR _ _). +rewrite domPtUnK /=; last first. +- by apply/allP=>? /In_dom_omfX [?][] /In_dom Y _; apply: path_mem Y. +- by rewrite valid_omfPtUn. +case D : (dom (f1 x))=>[//|t ts]. +case: eqP D=>[-> D|_ <-]; last by apply: IH=>kv X; apply: H (InR _ _). +have : k \in dom (f1 x) by rewrite D inE eq_refl. +case/In_dom_omfX=>w1 [] /In_dom /= D1. +by rewrite validPtUn D1 !andbF in W. +Qed. + +End OmapFun2. + +Section OmapFun2Eq. +Variables (K : ordType) (C : pred K) (V V1 V2 : Type). +Variables (U : union_map C V) (U1 : union_map C V1) (U2 : union_map C V2). +Variables (f1 : omap_fun U U1) (f2 : omap_fun U U2). + +Lemma omf_dom_eq (x : U) : + (forall kv, kv \In x -> omf f1 kv <-> omf f2 kv) -> + dom (f1 x) = dom (f2 x). +Proof. +move=>H; apply: subseq_anti. +by rewrite !omf_dom_subseq // => kv /H ->. +Qed. + +End OmapFun2Eq. + +Section OmapMembershipExtra. +Variables (K : ordType) (C : pred K) (V1 V2 : Type). +Variables (U1 : union_map C V1) (U2 : union_map C V2). +Variables f : omap_fun U1 U2. + +Lemma omfL e (x y : U1) : + e.1 \in dom x -> + e \In f (x \+ y) -> + e \In f x. +Proof. +move=>D /[dup] /In_valid/fpV V; rewrite pfjoinT //. +by case/InUn=>// /In_dom/omf_subdom/(dom_inNLX V D). +Qed. + +Lemma omfR e (x y : U1) : + e.1 \in dom y -> + e \In f (x \+ y) -> + e \In f y. +Proof. +move=>D /[dup] /In_valid/fpV V; rewrite pfjoinT //. +by case/InUn=>// /In_dom/omf_subdom/(dom_inNRX V D). +Qed. + +Lemma omfDL k (x y : U1) : + k \in dom x -> + k \in dom (f (x \+ y)) -> + k \in dom (f x). +Proof. by move=>D /In_domX [v] /omfL-/(_ D)/In_dom. Qed. + +Lemma omfDR k (x y : U1) : + k \in dom y -> + k \in dom (f (x \+ y)) -> + k \in dom (f y). +Proof. by move=>D /In_domX [v] /omfR-/(_ D)/In_dom. Qed. + +Lemma InpfLTD k (x y : U1) : + valid (x \+ y) -> + k \in dom (f x) -> + k \in dom (f (x \+ y)). +Proof. by move=>V /In_domX [v] /(InpfLT V)/In_dom. Qed. + +Lemma InpfRTD k (x y : U1) : + valid (x \+ y) -> + k \in dom (f y) -> + k \in dom (f (x \+ y)). +Proof. by move=>V /In_domX [v] /(InpfRT V)/In_dom. Qed. +End OmapMembershipExtra. + +Section OmapMembershipExtra2. +Variables (K : ordType) (C : pred K) (V1 V2 : Type). +Variables (U1 : union_map C V1) (U2 : union_map C V2). +Variables f : omap_fun U1 U2. + +Lemma omfDL00 k a (x y : U1) : + k \in a::dom x -> + k \in a::dom (f (x \+ y)) -> + k \in a::dom (f x). +Proof. by rewrite !inE; case: (k =P a)=>//= _; apply: omfDL. Qed. + +Lemma omfDL0 k a (x y : U1) : + k \in dom x -> + k \in a::dom (f (x \+ y)) -> + k \in a::dom (f x). +Proof. by rewrite !inE; case: (k =P a)=>//= _; apply: omfDL. Qed. + +Lemma omfDR00 k a (x y : U1) : + k \in a::dom y -> + k \in a::dom (f (x \+ y)) -> + k \in a::dom (f y). +Proof. by rewrite !inE; case: (k =P a)=>//= _; apply: omfDR. Qed. + +Lemma omfDR0 k a (x y : U1) : + k \in dom y -> + k \in a::dom (f (x \+ y)) -> + k \in a::dom (f y). +Proof. by rewrite !inE; case: (k =P a)=>//= _; apply: omfDR. Qed. + +Lemma InpfLTD0 k a (x y : U1) : + valid (x \+ y) -> + k \in a::dom (f x) -> + k \in a::dom (f (x \+ y)). +Proof. by move=>V; rewrite !inE; case: (k =P a)=>//= _; apply: InpfLTD. Qed. + +Lemma InpfRTD0 k a (x y : U1) : + valid (x \+ y) -> + k \in a::dom (f y) -> + k \in a::dom (f (x \+ y)). +Proof. by move=>V; rewrite !inE; case: (k =P a)=>//= _; apply: InpfRTD. Qed. +End OmapMembershipExtra2. + +(* categoricals *) + +Section OmapFunId. +Variables (K : ordType) (C : pred K) (V : Type). +Variables (U : union_map C V). + +Lemma omf_some (f : omap_fun U U) x : + (forall kv, kv \In x -> omf f kv = Some kv.2) -> + f x = x. +Proof. +elim/um_indf: x=>[||k v x IH W P] H; rewrite ?pfundef ?pfunit //. +by rewrite omfPtUn W H //= IH // => kv D; apply: H (InR _ D). +Qed. + +Lemma id_is_omap_fun : omap_fun_axiom (@idfun U) (Some \o snd). +Proof. by move=>x; rewrite omf_some. Qed. + +(* we already have that id is full/norm/tpcm morph *) +(* so we use the factory for that situation *) +HB.instance Definition _ := + isOmapFun_morph.Build K C V V U U idfun id_is_omap_fun. + +Lemma omf_id : omf idfun = Some \o snd. Proof. by []. Qed. + +Lemma valid_omfUnL (f : omap_fun U U) x1 x2 : + valid (x1 \+ x2) -> valid (f x1 \+ x2). +Proof. by rewrite {2}(_ : x2 = @idfun U x2) //; apply: valid_omfUn. Qed. + +Lemma valid_omfUnR (f : omap_fun U U) x1 x2 : + valid (x1 \+ x2) -> valid (x1 \+ f x2). +Proof. by rewrite !(joinC x1); apply: valid_omfUnL. Qed. + +End OmapFunId. + + +Section OmapFunComp. +Variables (K : ordType) (C : pred K) (V1 V2 V3 : Type). +Variables (U1 : union_map C V1) (U2 : union_map C V2) (U3 : union_map C V3). +Variables (f1 : omap_fun U1 U2) (f2 : omap_fun U2 U3). + +Lemma comp_is_omap_fun : + omap_fun_axiom (f2 \o f1) + (fun x => obind (fun v => omf f2 (x.1, v)) (omf f1 x)). +Proof. +elim/um_indf=>[||k v x /= IH W P]; rewrite ?pfundef ?pfunit //=. +rewrite !omapPtUn !omfPtUn W /=; case: (omf f1 (k, v))=>[w|//]. +rewrite omfPtUn validPtUn (validPtUn_cond W) pfVE (validR W) /= IH. +by rewrite (contra _ (validPtUnD W)) //; apply: omf_subdom. +Qed. + +HB.instance Definition _ := + isOmapFun_morph.Build K C V1 V3 U1 U3 (f2 \o f1) comp_is_omap_fun. + +Lemma omf_comp : + omf (f2 \o f1) = + fun x => obind (fun v => omf f2 (x.1, v)) (omf f1 x). +Proof. by []. Qed. + +End OmapFunComp. + + +(*********************) +(* omap specifically *) +(*********************) + +(* 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 OmapId. +Variables (K : ordType) (C : pred K) (V : Type) (U : union_map C V). + +Lemma omapv_id : omapv Some =1 @id U. +Proof. by move=>x; apply: omf_some. Qed. + +Lemma mapv_id : mapv id =1 @id U. +Proof. by apply: omapv_id. Qed. + +End OmapId. + +Section OmapComp. +Variables (K : ordType) (C : pred K) (V1 V2 V3 : Type). +Variables (U1 : union_map C V1) (U2 : union_map C V2) (U3 : union_map C V3). + +Lemma omap_omap f1 f2 (x : U1) : + omap f2 (omap f1 x : U2) = + omap (fun kv => obind (fun v => f2 (kv.1, v)) (f1 kv)) x :> U3. +Proof. by rewrite (compE (omap f2)) omfE. Qed. + +Lemma omapv_omapv f1 f2 (x : U1) : + omapv f2 (omapv f1 x : U2) = omapv (obind f2 \o f1) x :> U3. +Proof. by rewrite omap_omap. Qed. + +Lemma omapv_mapv f1 f2 (x : U1) : + omapv f2 (mapv f1 x : U2) = omapv (f2 \o f1) x :> U3. +Proof. by rewrite omapv_omapv. Qed. + +(* RHS is just filtering; will restate with um_filter below *) +Lemma mapv_omapvE f g (x : U1) : + ocancel f g -> + mapv g (omapv f x : U2) = + omapv (fun v => if f v then Some v else None) x :> U1. +Proof. +move=>O; rewrite (compE (mapv g)) eq_in_omf omf_omap omf_comp !omf_omap /=. +by case=>k v H; case X: (f v)=>[a|] //=; rewrite -(O v) X. +Qed. + +End OmapComp. + + +Section OmapMembership. +Variables (K : ordType) (C : pred K) (V V' : Type). +Variables (U : union_map C V) (U' : union_map C V'). + +Lemma In_omapX f k v (x : U) : + (k, v) \In (omap f x : U') <-> + exists2 w, (k, w) \In x & f (k, w) = Some v. +Proof. exact: In_omfX. Qed. + +Lemma In_omap f k v w (x : U) : + (k, w) \In x -> f (k, w) = Some v -> (k, v) \In (omap f x : U'). +Proof. by move=>H1 H2; apply: In_omf H1 _; apply: H2. Qed. + +Lemma In_dom_omapX (f : K * V -> option V') k (x : U) : + reflect (exists w, (k, w) \In x /\ f (k, w)) + (k \in dom (omap f x : U')). +Proof. by case: In_dom_omfX=>H; constructor; exact: H. Qed. + +Lemma In_omapv f k v w (x : U) : + (k, w) \In x -> f w = Some v -> (k, v) \In (omapv f x : U'). +Proof. by move=>H1 H2; apply: In_omf H1 _; rewrite omf_omap. Qed. + +Lemma In_mapv f k v (x : U) : + injective f -> (k, f v) \In (mapv f x : U') <-> (k, v) \In x. +Proof. by move=>I; rewrite In_omfX; split; [case=>w H [] /I <-|exists v]. Qed. + +End OmapMembership. + +Arguments In_omapv {K C V V' U U' f k v w}. + +Section OmapMembership2. +Variables (K : ordType) (C : pred K) (V V' : Type). +Variables (U : union_map C V) (U' : union_map C V'). + +Lemma In_omapV f g k v (x : U) : + ocancel f g -> pcancel g f -> + (k, v) \In (omapv f x : U') <-> (k, g v) \In x. +Proof. +move=>O P; rewrite -(In_mapv U _ _ _ (pcan_inj P)). +rewrite mapv_omapvE // In_omapX /=. +split; first by case=>w H; case: (f w)=>[a|] //= [<-]. +by exists (g v)=>//=; rewrite P. +Qed. + +Lemma In_rangev f g v (x : U) : + ocancel f g -> pcancel g f -> + v \In range (omapv f x : U') <-> g v \In range x. +Proof. +move=>O P; rewrite !In_rangeX; split; case=>k H; exists k. +- by rewrite -(In_omapV _ _ _ O P). +by rewrite (In_omapV _ _ _ O P). +Qed. + +End OmapMembership2. + +Section OmapMembership3. +Variables (K : ordType) (C : pred K) (V V' : eqType). +Variables (U : union_map C V) (U' : union_map C V'). + +(* decidable variant of In_rangev *) +Lemma mem_rangev f g v (x : U) : + ocancel f g -> pcancel g f -> + v \in range (omapv f x : U') = (g v \in range x). +Proof. +by move=>O P; apply/idP/idP; move/mem_seqP/(In_rangev _ _ _ O P)/mem_seqP. +Qed. + +End OmapMembership3. + + +(* freeing k representable as omap *) +Section OmapFree. +Variables (K : ordType) (C : pred K) (V V' : Type). +Variables (U : union_map C V) (U' : union_map C V'). + +Lemma omap_free k x : + free x k = + omap (fun kv => if kv.1 == k then None else Some kv.2) x :> U. +Proof. +case: (normalP x)=>[->|D]; first by rewrite free_undef pfundef. +apply: umem_eq. +- by rewrite validF. +- by rewrite pfVE. +case=>t w; rewrite InF In_omapX /= validF D; split. +- by case=>_ /negbTE ->; exists w. +by case=>w' H; case: eqP=>// N [<-]. +Qed. + +End OmapFree. + +(* eq_in_omap *) +Section EqInOmap. +Variables (K : ordType) (C : pred K) (V V' : Type). +Variables (U : union_map C V) (U' : union_map C V'). + +Lemma eq_in_omap a1 a2 (h : U) : + (forall kv, kv \In h -> a1 kv = a2 kv) <-> + omap a1 h = omap a2 h :> U'. +Proof. by rewrite eq_in_omf. Qed. +End EqInOmap. + + +(*************) +(* 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 C V). +Implicit Types (f : U) (p q : pred (K * V)). + +Definition um_filter p f : U := + omap (fun kv => if p kv then Some kv.2 else None) f. + +(* um_filter is omap_fun *) +HB.instance Definition _ p := OmapFun.copy (um_filter p) (um_filter p). + +Lemma umfiltPtE 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 omfPtE omf_omap; case: (ifP (p _)). Qed. + +Lemma umfiltPt p k v : + C k -> + um_filter p (pts k v) = if p (k, v) then pts k v else Unit. +Proof. by move=>D; rewrite umfiltPtE D. Qed. + +Lemma umfiltUn p f1 f2 : + valid (f1 \+ f2) -> + um_filter p (f1 \+ f2) = um_filter p f1 \+ um_filter p f2. +Proof. exact: pfjoinT. 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 omfPtUn omf_omap; case: (ifP (p _)). 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 omfUnPt omf_omap; case: (ifP (p _)). Qed. + +Lemma In_umfiltX x p f : x \In um_filter p f <-> p x /\ x \In f. +Proof. +case: x => k v; rewrite In_omfX omf_omap; split=>[[w H]|[I H]] /=. +- by case E: (p (k, w))=>//=; case=><-. +by exists v=>//; rewrite I. +Qed. + +(* In_umfilt is really only good in one direction. *) +(* For the other direction, we need the following one *) +Lemma In_umfilt p x f : x \In f -> p x -> x \In um_filter p f. +Proof. by move=>X1 X2; apply/In_umfiltX. Qed. + +Lemma In_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_umfiltX []; exists v. +by case=>v [Hp Hf]; elim: H; exists v; apply/In_umfilt. +Qed. + +Lemma dom_omf_umfilt V' (U' : union_map C V') (f : omap_fun U U') x : + dom (f x) = dom (um_filter (isSome \o omf f) x). +Proof. +apply/domE=>k; apply/idP/idP. +- case/In_dom_omfX=>//= v [H1 H2]. + by apply/In_dom_umfilt; exists v; rewrite /= H2. +case/In_dom_umfilt=>w [/=]. +case E: (omf f (k, w))=>[a|] // _ H. +by move/In_dom: (In_omf _ H E). +Qed. + +Lemma dom_omap_umfilt V' (U' : union_map C V') a x : + dom (omap a x : U') = dom (um_filter (isSome \o a) x). +Proof. exact: dom_omf_umfilt. 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: ord_sorted_eq=>//=. +- by apply: sorted_filter; [apply: trans | apply: sorted_dom]. +move=>k; rewrite mem_filter; apply/idP/idP. +- by case/In_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/In_dom_umfilt; exists v. +Qed. + +Lemma umfilt_pred0 f : valid f -> um_filter pred0 f = Unit. +Proof. by move=>W; apply: omf_none. Qed. + +Lemma umfilt_predT f : um_filter predT f = f. +Proof. by apply: omf_some. Qed. + +Lemma umfilt_predI p1 p2 f : + um_filter (predI p1 p2) f = um_filter p1 (um_filter p2 f). +Proof. +rewrite [RHS]compE eq_in_omf omf_comp !omf_omap /=. +by case=>k v H /=; case: (ifP (p2 _))=>//=; 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. +rewrite eq_in_omf !omf_omap /=. +split=>E [k v] H; first by rewrite E. +by case: (p1 _) (E _ H); case: (p2 _). +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 omf_predU=>[|kv]. +- by rewrite eq_in_omf !omf_omap /= => kv; case: (p1 _). +by rewrite !omf_omap /=; case: (p1 _)=>/=; [right|left]. +Qed. + +(* we put localization back In for xor *) +(* DEVCOMMENT *) +(* TODO: this should be done for all umfilt_?pred? lemmas *) +(* /DEVCOMMENT *) +Lemma umfilt_predX f p q : + (forall kv, kv \In f -> p kv (+) q kv) -> + f = um_filter p f \+ um_filter q f. +Proof. +move=>D; rewrite -{1}[f]umfilt_predT. +have : forall kv, kv \In f -> predT kv = (predU p q) kv. +- by move=>kv /D /=; case: (p kv). +move/eq_in_umfilt=>->; rewrite umfilt_predU; congr (_ \+ _). +by apply/eq_in_umfilt=>kv /D /=; case: (p kv)=>// /negbTE ->. +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. + +Corollary umfilt_predC f p : f = um_filter p f \+ um_filter (predC p) f. +Proof. +rewrite -umfilt_dpredU; last by move=>? /=; rewrite negbK. +rewrite -[LHS]umfilt_predT; apply: eq_in_umfiltE=>kv /=. +by rewrite orbN. +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: omf_subdom. + move=>D; move: (D); rewrite -{1}E domUn inE W' /=. + by case/orP=>// /omf_subdom; 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 umfiltUE 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 omfUE omf_omap; case: (p _). Qed. + +Lemma umfiltU p k v f : + C k -> + um_filter p (upd k v f) = + if p (k, v) then upd k v (um_filter p f) + else free (um_filter p f) k. +Proof. by move=>H; rewrite umfiltUE H. Qed. + +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 !pfundef !umfoldl_undef. +- by rewrite !pfunit !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 /In_dom_umfilt [w][_] /In_dom. + by apply: path_mem (@trans K) P. +have : valid (um_filter p (pts k v \+ f)) by rewrite pfVE 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. by move/omf_noneR=>H k v /H; rewrite omf_omap /=; case: (p _). Qed. + +Lemma umfilt_mem0X p f k v : + (k, v) \In f -> um_filter p f = Unit -> ~ p (k, v). +Proof. by move=>H /umfilt_mem0/(_ _ _ H)/negP. Qed. + +Lemma umfilt_mem0L p f : + valid f -> (forall k v, (k, v) \In f -> ~~ p (k, v)) -> + um_filter p f = Unit. +Proof. by move=>W H; rewrite omf_none // omf_omap; case=>k v /H/negbTE ->. 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 pfundef assocs_undef. +- by rewrite pfunit 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 pfVE. +by apply/allP=>x; move/allP: P=>P; move/omf_subdom/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. by rewrite find_omf. Qed. + +Lemma unitb_umfilt p f : unitb f -> unitb (um_filter p f). +Proof. by move/unitbP=>->; rewrite pfunit 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. + +Lemma dom_umfilt2 p1 p2 f x : + x \in dom (um_filter p1 (um_filter p2 f)) = + (x \in dom (um_filter p1 f)) && (x \in dom (um_filter p2 f)). +Proof. +rewrite -umfilt_predI; apply/idP/idP. +- case/In_dom_umfilt=>v [/andP [X1 X2] H]. + by apply/andP; split; apply/In_dom_umfilt; exists v. +case/andP=>/In_dom_umfilt [v1][X1 H1] /In_dom_umfilt [v2][X2 H2]. +move: (In_fun H1 H2)=>E; rewrite -{v2}E in X2 H2 *. +by apply/In_dom_umfilt; exists v1; split=>//; apply/andP. +Qed. + +End FilterDefLemmas. + +#[export] +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). + +Arguments In_umfilt [K C V U] p x f _ _. + +Section FilterKLemmas. +Variables (K : ordType) (C : pred K) (V : Type) (U : union_map C V). +Implicit Type f : U. +Implicit Type p q : pred K. + +Lemma dom_umfiltkE p f : dom (um_filterk p f) = filter p (dom f). +Proof. +apply: ord_sorted_eq=>//=. +- by apply: sorted_filter; [apply: trans | apply: sorted_dom]. +move=>k; rewrite mem_filter; apply/idP/idP. +- by case/In_dom_umfilt=>v [/= -> /In_dom]. +by case/andP=>H1 /In_domX [v H2]; apply/In_dom_umfilt; exists v. +Qed. + +Lemma valid_umfiltkUn p1 p2 f : + valid f -> + {in dom f, forall x, x \in p1 -> x \in p2 -> false} -> + valid (um_filterk p1 f \+ um_filterk p2 f). +Proof. +move=>W H; rewrite validUnAE !pfVE W /=. +apply/allP=>x; case/In_dom_umfilt=>v1 /= [H2 F1]. +apply/negP; case/In_dom_umfilt=>v2 /= [H1 F2]. +by move: (H x (In_dom F1) H1 H2). +Qed. + +Lemma dom_umfiltk p f : dom (um_filterk p f) =i predI p (mem (dom f)). +Proof. by move=>k; rewrite dom_umfiltkE mem_filter. Qed. + +(* DEVCOMMENT *) +(* this also holds for invalid f1, as the corollary shows *) +(* /DEVCOMMENT *) +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 pfVE. +- by rewrite (validL W). +case=>k v; rewrite In_umfiltX; split=>[|H]. +- by case=>H /InUn [] // /In_dom; rewrite (negbTE (dom_inNL W H)). +by split; [apply: In_dom H | apply: InL]. +Qed. + +Corollary umfiltk_dom' f : um_filterk (mem (dom f)) f = f. +Proof. +case: (normalP f)=>[->|H]; first by rewrite pfundef. +by rewrite -{2}[f]unitR umfiltk_dom // unitR. +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 umfiltk_subdom p f : {subset dom (um_filterk p f) <= p}. +Proof. by move=>a; rewrite dom_umfiltk; case/andP. Qed. + +Lemma umfiltk_subdomE 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/umfiltk_subdomE. Qed. + +Lemma find_umfiltk k (p : pred K) f : + find k (um_filterk p f) = if p k then find k f else None. +Proof. by rewrite find_umfilt /=; case: (find _ _)=>[a|]; case: ifP. Qed. + +Lemma umfiltk_subdom0 p f : + valid f -> {subset dom f <= predC p} <-> um_filterk p f = Unit. +Proof. +move=>W; split=>[H|H k X]. +- rewrite (eq_in_umfiltk (p2:=pred0)) ?umfilt_pred0 //. + by move=>a /H /negbTE ->. +case: dom_find X H=>// v _ -> _; rewrite omfUE !inE omf_omap /=. +case: (ifP (p k))=>// _ /unitbP. +by rewrite fun_if um_unitbU unitb_undef if_same. +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 umfiltUE pfunit 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: (normalP (pts k v \+ f))=>[->|W]; first by rewrite pfundef. +rewrite pfjoinT //= umfiltPtE (validPtUn_cond W) /=. +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. + +(* filters commute with omap_fun *) + +Lemma umfiltk_omf V' (U' : union_map C V') (f : omap_fun U U') p x : + f (um_filterk p x) = um_filterk p (f x). +Proof. +rewrite (compE f) [RHS]compE eq_in_omf !omf_comp !omf_omap /=. +by rewrite /obind/oapp /=; case=>k v; case: ifP; case: (omf f _). +Qed. + +Lemma umfiltk_dom_omf V' (U' : union_map C V') (f : omap_fun U U') x : + um_filterk (mem (dom x)) (f x) = f x. +Proof. by rewrite -umfiltk_omf umfiltk_dom'. Qed. + +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 umfiltUE /=; case: ifP; case: ifP=>//= Hp Hc. +- by rewrite dom_free // 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 omfF; case: ifP=>//= N. +by rewrite dom_free // dom_umfiltk inE N. +Qed. + +Lemma In_umfiltk p x f : x \In f -> p x.1 -> x \In um_filterk p f. +Proof. by apply: In_umfilt. Qed. + +End FilterKLemmas. + +Arguments In_umfiltk [K C V U] p [x f] _ _. + +Section FilterVLemmas. +Variables (K : ordType) (C : pred K) (V : Type) (U : union_map 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. + +Lemma In_umfiltv p x f : x \In f -> p x.2 -> x \In um_filterv p f. +Proof. by apply: In_umfilt. Qed. + +End FilterVLemmas. + +Arguments In_umfiltv [K C V U] p [x f] _ _. + +Section OmapMembershipLemmas. +Variables (K : ordType) (C : pred K) (V V' : Type). +Variables (U : union_map C V) (U' : union_map C V'). + +Lemma mapv_omapv f g (x : U) : + ocancel f g -> + mapv g (omapv f x : U') = um_filterv (isSome \o f) x. +Proof. exact: mapv_omapvE. Qed. + +End OmapMembershipLemmas. + + +(************************) +(* PCM-induced ordering *) +(************************) + +(* Union maps and PCM ordering. *) + +Section UmpleqLemmas. +Variables (K : ordType) (C : pred K) (V : Type) (U : union_map C V). +Implicit Types (x y a b : U) (p : pred K). + +(* PCM-induced preorder is 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/umap0E=>->; 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: (normalP y)=>[->|]. +- by rewrite pfundef; apply: pleq_undef. +by case: H=>z -> D; rewrite pfjoinT. +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. +move=>D /(umfilt_pleq_mono (p \o fst)). +by 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. by move=>W; split=>[|<-] // H; case: H W=>b ->; 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: (normalP f)=>[->???|Df Dab]; first by apply: pleq_undef. +move/(umpleqk_pleqE _ Df)=> <- /(umpleqk_pleqE _ Df) <-. +by rewrite -umfilt_dpredU //; case=>/= k _; apply: dom_inNL Dab. +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/subsetR: (valid_subdom W). +by move/(umfiltk_subdom0 _ (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 move/(dom_inNLX V12)=>X /X. +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 *) +(********************) + +(* DEVCOMMENT *) +(* naturally belongs to dom section, but proofs use lemmas *) +(* that haven't been proved before the dom section *) +(* /DEVCOMMENT *) +Section Precision. +Variables (K : ordType) (C : pred K) (V : Type) (U : union_map 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. + by rewrite E domUn inE -E V1 /= (negbTE (dom_inNL V12 Hs)) orbF. +move=>Hs; case: validUn V12=>//. +- by rewrite (validE2 V1). +- by rewrite E in V1; rewrite (validE2 V1). +by rewrite E in V1; move=>k /Hs /(dom_inNL V1)/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 *) +(****************) + +(* version of eval where the user provides new order of evaluation *) +(* as list of timestamps which are then read in-order. *) + +Section OrdEvalDefLemmas. +Variables (K : ordType) (C : pred K) (V R : Type) (U : union_map 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: (normalP f)=>[-> _|D]; first by apply: oev_undef. +by move/(dom0E D)=>->; apply: oev0. +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. +elim: ks z0=>[|k' ks IH] z0 //=. +by 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. +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_umfiltX /=; 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) -> k \notin ks -> + 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 right. +by case=>// [][??]; subst k0; rewrite H in S. +Qed. + +(* a somewhat different version; makes side conditions easier to discharge *) +Lemma oevPtUn_sub 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 F; apply: oevPtUn=>//; apply/negP=>/F. +by rewrite (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. + +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. + +(* a somewhat stronger lemma making clear that z0' isn't *) +(* arbitrary but always obtained by evaluation starting from z0 *) +Lemma oev_indX {P : R -> Prop} f ks a z0 : + P z0 -> + (forall k ks1 ks2 v z0', (k, v) \In f -> ks = ks1 ++ k :: ks2 -> + z0' = oeval a ks1 f z0 -> 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. +- move=>k' ks1 ks2 v' z' X D E. + by apply: (H k' (k::ks1) ks2 v' z' X _ _)=>//; rewrite D. +case E: (find k f)=>[b|] //; move/In_find: E=>E. +by apply: (H k [::] _ b z0). +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 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 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 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 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_umfiltX. +split; last by case. +split=>//; move: H; move/In_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 umfiltPtE. +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 pfV. +rewrite /eval umfiltPtUn W. +case: (p (k, v))=>W'; last first. +- rewrite oevPtUn //; apply/negP=>/omf_subdom. + by rewrite (negbTE (validPtUnD W)). +rewrite domPtUnK //=; last by apply/allP=>x /omf_subdom /H. +by rewrite findPtUn // oevPtUn // (validPtUnD W'). +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 pfV. +rewrite /eval umfiltUnPt W. +case: (p (k, v))=>W'; last first. +- rewrite joinC oevPtUn //; first by rewrite joinC. + by apply/negP=>/omf_subdom; rewrite (negbTE (validUnPtD W)). +rewrite domUnPtK //=; last by apply/allP=>x /omf_subdom /H. +rewrite (oev_rconsP _ (v:=v)) // joinC oevPtUn //; first by rewrite joinC. +by apply/negP=>/omf_subdom; rewrite (negbTE (validUnPtD W)). +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: (normalP f)=>[->|D]; first by 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 C V) (U' : union_map 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 pfundef !eval_undef. +- by rewrite pfunit !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 /omf_subdom /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 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 C1 V1) (U2 : union_map C2 V2). +Variables (U : union_map 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 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_sub=>//; apply: omf_subdom. +rewrite oevUn // -(oev_sub_filter (p:=mem [:: k])) ?(domPtK,Ck) //. +rewrite -dom_umfiltkE umfiltPtUn /= valid_omfUnR // inE eq_refl. +rewrite umfilt_mem0L ?(inE,pfV,validR W) //=; first last. +- by move=>?? /In_umfiltX [] _ /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_umfiltkE umfiltPtUn /= valid_omfUnR // (negbTE D). +congr (a (oeval a (dom _) f z0) k v). +by rewrite -umfiltk_subdomE; apply: omf_subdom. +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 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 pfunit dom0. Qed. + +Lemma umcnt_undef p : um_count p undef = 0. +Proof. by rewrite /um_count pfundef 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 umfiltPtE; 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. +by move=>W; rewrite /um_count umfiltUn // size_domUn // -umfiltUn // pfV. +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/In_dom_umfilt: E=>w [H1 H2]. + rewrite -{w H2}(In_fun H H2) in H1 *. + rewrite H1 omfF size_domF ?subn1 //=. + by apply/In_dom_umfilt; exists v. +rewrite omfF; case: ifP=>P; last by rewrite dom_free // E. +by move/negP: E; elim; apply/In_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: (normalP f)=>[->|W]; first by rewrite !umcnt_undef. +rewrite /um_count (umfilt_predD _ S) size_domUn ?leq_addr //. +by rewrite -umfilt_predD // pfV. +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: (normalP f)=>[->|D]; first by 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: (normalP f)=>[->|W]; first by rewrite !umcnt_undef. +rewrite /um_count umfilt_predU size_domUn //. +case: validUn=>//; rewrite ?(pfV,W) //. +move=>k /In_dom_umfilt [v1 [P1 H1]] /In_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. by rewrite umcnt_umfiltC umcnt_predU !(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 ?pfV ?D. +Qed. + +Lemma umcnt_eval0 R a p f (z0 : R) : um_count p f = 0 -> eval a p f z0 = z0. +Proof. +case: (normalP f)=>[->|W]; first by 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: (normalP f)=>[->|W]. +- by 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. + + +(*************************************) +(* Filtering maps of tagged elements *) +(*************************************) + +Section SideFilter. +Variables (T : eqType) (Us : T -> Type). + +(* could also be defined as *) +(* Definition side_m t : sigT Us -> option (Us t) := *) +(* fun '(Tag tx ux) => *) +(* if t =P tx is ReflectT pf then Some (cast Us pf ux) *) +(* else None. *) +(* but that doesn't reduce to then/else clause *) +(* if t == tx and t != tx, respectively *) +(* The following definition gets that reduction *) + +Definition side_m t : sigT Us -> option (Us t) := + fun '(Tag tx ux) => + if decP (t =P tx) is left pf then Some (cast Us pf ux) + else None. + +Lemma side_ocancel t : ocancel (side_m t) (Tag t). +Proof. by case=>tx vx /=; case: eqP=>//= pf; subst tx; rewrite eqc. Qed. + +End SideFilter. + +#[export] +Hint Extern 0 (ocancel (side_m _) (Tag _)) => + (apply: side_ocancel) : core. + +(* select all tags in h that equal t *) +(* differs from tags t in that it drops t *) +(* from the result map *) +(* whereas tags keeps the entries in result map unmodified *) + +Section Side. +Variables (K : ordType) (C : pred K) (T : eqType) (Us : T -> Type). +Variables (U : union_map C (sigT Us)). +Variables (Ut : forall t, union_map C (Us t)). +Variables t : T. + +Definition side_map : U -> Ut t := locked (omapv (side_m t)). +Lemma side_unlock : side_map = omapv (side_m t). +Proof. by rewrite /side_map -lock. Qed. +Lemma sidemap_is_omf : omap_fun_axiom side_map (side_m t \o snd). +Proof. by rewrite side_unlock; apply: omfE. Qed. +HB.instance Definition _ := + isOmapFun.Build K C (sigT Us) (Us t) U (Ut t) side_map sidemap_is_omf. + +(* explicit name for validity of side *) +Lemma valid_side (h : U) : valid (side_map h) = valid h. +Proof. exact: pfVE. Qed. + +Lemma In_side x (v : Us t) (h : U) : + (x, v) \In side_map h <-> (x, Tag t v) \In h. +Proof. +rewrite side_unlock In_omapX; split=>[|H]; last first. +- by exists (Tag t v)=>//=; case: eqP=>//= ?; rewrite eqc. +case; case=>t' v' /= H; case: eqP=>//= ?; subst t'. +by rewrite eqc; case=><-. +Qed. + +Lemma side_umfilt p q (h : U) : + (forall k v, + (k, Tag t v) \In h -> p (k, Tag t v) = q (k, v)) -> + side_map (um_filter p h) = um_filter q (side_map h). +Proof. +move=>H; rewrite side_unlock /um_filter !omap_omap eq_in_omf !omf_omap /=. +rewrite /side_m/obind/oapp/=; case=>k; case=>t' v X /=. +by case P : (p _); case: eqP=>//= ?; subst t'; rewrite eqc -H // P. +Qed. + +(* if p can only inspect keys *) +Lemma side_umfiltk (p : pred K) h : + side_map (um_filterk p h) = um_filterk p (side_map h). +Proof. by apply: side_umfilt. Qed. + +Lemma In_side_umfiltX x (v : Us t) p (h : U) : + (x, v) \In side_map (um_filter p h) <-> + p (x, Tag t v) /\ (x, Tag t v) \In h. +Proof. by rewrite In_side In_umfiltX. Qed. + +Lemma In_side_umfilt x (v : Us t) (p : pred (K * sigT Us)) (h : U) : + p (x, Tag t v) -> (x, Tag t v) \In h -> + (x, v) \In side_map (um_filter p h). +Proof. by move=>H1 H2; apply/In_side_umfiltX. Qed. + +Lemma In_dom_sideX x (h : U) : + reflect (exists a, (x, Tag t a) \In h) + (x \in dom (side_map h)). +Proof. +case: In_dom_omfX=>/= H; constructor. +- case: H=>[[t' v']][H]; rewrite /omfx/=. + by case: eqP=>// ?; subst t'; exists v'. +case=>v H'; elim: H. +exists (Tag t v); split=>//; rewrite /omfx/=. +by case: eqP=>// ?; rewrite eqc. +Qed. + +Lemma In_dom_side x a (h : U) : + (x, Tag t a) \In h -> x \in dom (side_map h). +Proof. by move/In_side/In_dom. Qed. + +Lemma sidePtE x e : + side_map (pts x e) = + if C x then + if decP (t =P tag e) is left pf then + pts x (cast Us pf (tagged e)) else Unit + else undef. +Proof. by case: e=>k v; rewrite omfPtE /omfx/=; case: eqP. Qed. + +Lemma dom_sidePt x e : + dom (side_map (pts x e)) = + if C x && (t == tag e) then [:: x] else [::]. +Proof. +rewrite sidePtE; case H : (C x)=>//=; last by rewrite dom_undef. +case: eqP=>[pf|] /=; last by rewrite dom0. +by rewrite domPtK H. +Qed. + +Lemma dom_sidePtUn k e h : + dom (side_map (pts k e \+ h)) =i + [pred x | valid (pts k e \+ h) & + (x == k) && (t == tag e) || (x \in dom (side_map h))]. +Proof. +move=>x; rewrite dom_omfPtUn !inE /omfx/= (andbC (x == k)). +by case: e=>t' v /=; case: eqP. +Qed. + +End Side. + +Arguments side_map {K C T Us U}. +Arguments In_side {K C T Us U} Ut {t x v h}. +Arguments In_dom_sideX {K C T Us U} Ut {t x h}. +Arguments In_dom_side {K C T Us U} Ut {t x a h}. + + +(* lemmas about two different sides *) + +Section Side2. +Variables (K : ordType) (C : pred K) (T : eqType) (Us : T -> Type). +Variables (U : union_map C (sigT Us)). +Variables (Ut : forall t, union_map C (Us t)). +Variables t1 t2 : T. + +Lemma sideN_all_predC (h : U) : + t1 <> t2 -> + all [predC dom (side_map Ut t1 h)] + (dom (side_map Ut t2 h)). +Proof. +move=>N; apply/allP=>x /In_domX [v1] /In_side H /=. +apply/negP=>/In_domX [v2] /In_side /(In_fun H). +by case=>X; rewrite X in N. +Qed. + +Lemma In_side_fun k (v1 : Us t1) (v2 : Us t2) (h : U) : + (k, v1) \In side_map Ut t1 h -> + (k, v2) \In side_map Ut t2 h -> + t1 = t2 /\ jmeq Us v1 v2. +Proof. +move/In_side=>H /In_side/(In_fun H) [?]; subst t2. +by move/inj_pair2=>->. +Qed. + +Lemma dom_sideE k (h : U) : + k \in dom (side_map Ut t1 h) -> + k \in dom (side_map Ut t2 h) -> + t1 = t2. +Proof. by case/In_domX=>v1 H1 /In_domX [v2] /(In_side_fun H1) []. Qed. + +Lemma dom_sideEX k (h : U) : + k \in dom (side_map Ut t1 h) -> + k \in dom (side_map Ut t2 h) = (t1 == t2). +Proof. +case/In_dom_sideX=>v H; case: (t1 =P t2)=>[?|N]. +- by subst t2; apply/In_dom_sideX; exists v. +by apply/In_dom_sideX; case=>w /(In_fun H) [] /N. +Qed. + +Lemma dom_sideN k1 k2 (h : U) : + k1 \in dom (side_map Ut t1 h) -> + k2 \in dom (side_map Ut t2 h) -> + t1 != t2 -> k1 != k2. +Proof. by move=>D1 D2; apply: contra=>E; rewrite -(dom_sideEX D1) (eqP E). Qed. + +End Side2. + +(* iterated tagging = when sigT Ts is used as tag *) + +Definition sliceT T (Ts : T -> Type) Us t := + {x : Ts t & Us (Tag t x)}. +Arguments sliceT {T Ts}. + +Section IterTag. +Variables (T : Type) (Ts : T -> Type) (Us : sigT Ts -> Type). + +(* tag re-association *) + +(* split (i.e., slice) tag in two *) +Definition slice_m : sigT Us -> sigT (sliceT Us) := + fun '(Tag (Tag t k) v) => Tag t (Tag k v). +(* conjoin (i.e. gather) first and second tag *) +Definition gather_m : sigT (sliceT Us) -> sigT Us := + fun '(Tag t (Tag k v)) => Tag (Tag t k) v. + +Variables (K : ordType) (C : pred K). +Variables (U : union_map C (sigT Us)). +Variables (S : union_map C (sigT (sliceT Us))). + +Definition slice : U -> S := mapv slice_m. +HB.instance Definition _ := OmapFun.copy slice slice. +Definition gather : S -> U := mapv gather_m. +HB.instance Definition _ := OmapFun.copy gather gather. + +Lemma In_slice x t (k : Ts t) (v : Us (Tag t k)) h : + (x, Tag t (Tag k v)) \In slice h <-> + (x, Tag (Tag t k) v) \In h. +Proof. +rewrite In_omfX; split=>[|H]; last first. +- by exists (Tag (Tag t k) v). +case; case; case=>t' k' v' H /=. +case=>?; subst t'=>/inj_pair2 ?; subst k'. +by move/inj_pair2/inj_pair2=><-. +Qed. + +Lemma In_gather x t (k : Ts t) (v : Us (Tag t k)) h : + (x, Tag (Tag t k) v) \In gather h <-> + (x, Tag t (Tag k v)) \In h. +Proof. +rewrite In_omfX; split=>[|H]; last first. +- by exists (Tag t (Tag k v)). +case; case=>t' [k' v'] H /=. +case=>?; subst t'=>/inj_pair2 ?; subst k'. +by move/inj_pair2=><-. +Qed. + +Lemma gather_slice h : gather (slice h) = h. +Proof. +case: (normalP h)=>[->|V]; first by rewrite !pfundef. +apply: umem_eq=>//; first by rewrite !pfV. +by case=>x [[t k v]]; rewrite In_gather In_slice. +Qed. + +Lemma slice_gather h : slice (gather h) = h. +Proof. +case: (normalP h)=>[->|V]; first by rewrite !pfundef. +apply: umem_eq=>//; first by rewrite !pfV. +by case=>x [t][k v]; rewrite In_slice In_gather. +Qed. + +End IterTag. + +Arguments slice {T Ts Us K C U}. +Arguments gather {T Ts Us K C U}. + + +(* grafting *) +(* clear side t from h and replace it with ht *) +(* side+graft are for union_maps what sel+splice are for finfuns *) + +Section GraftDef. +Variables (K : ordType) (C : pred K) (T : eqType) (Us : T -> Type). +Variables (U : union_map C (sigT Us)). +Variables (Ut : forall t, union_map C (Us t)). +Variables (h : U) (t : T). + +Definition graft (ht : Ut t) : U := + um_filterv (fun v => t != tag v) h \+ mapv (Tag t) ht. +End GraftDef. + +Arguments graft {K C T Us U Ut}. + +Section GraftLemmas. +Variables (K : ordType) (C : pred K) (T : eqType) (Us : T -> Type). +Variables (U : union_map C (sigT Us)) (Ut : forall t, union_map C (Us t)). +Implicit Type h : U. + +(* ht has disjoint domain with other sides of h *) +Definition disjoint_graft h t (ht : Ut t) := + forall tx x, + x \in dom (side_map Ut tx h) -> + x \in dom ht -> + tx = t. + +Lemma valid_graft h (t : T) (ht : Ut t) : + [/\ valid h, valid ht & disjoint_graft h ht] -> + valid (graft h t ht). +Proof. +case=>Vh V D. +rewrite validUnAE !pfV //=; apply/allP=>x /In_dom_omapX []. +move=>t1 [/In_dom /= H _]; apply/In_dom_umfilt. +case; case=>tx vx [/eqP /= N]. +by move/(In_side Ut)/In_dom/D/(_ H)/esym/N. +Qed. + +Lemma valid_graftUn h1 h2 (t : T) (ht : Ut t) : + [/\ valid (h1 \+ h2), valid ht, + disjoint_graft h1 ht & + forall x, x \in dom h2 -> x \in dom ht -> false] -> + valid (graft h1 t ht \+ h2). +Proof. +case=>Vh V D1 D2. +rewrite validUnAE valid_graft ?(validL Vh, validR Vh) //=. +apply/allP=>x D; apply/In_domX; case; case=>tx vx. +case/InUn; first by case/In_umfiltX=>_ /In_dom /(dom_inNLX Vh). +case/In_omapX=>w /In_dom /= H [?]; subst tx=>/inj_pair2 ?; subst w. +by move: (D2 _ D H). +Qed. + +(* if ht has disjoint domain with other sides of h *) +(* then side t (graft h t ht) = ht *) +Lemma side_graftE h t (ht : Ut t) : + valid h -> + disjoint_graft h ht -> + side_map Ut t (graft h t ht) = ht. +Proof. +move=>Vh D; case: (normalP ht)=>[->|V]. +- by rewrite /graft pfundef join_undef pfundef. +have W : valid (graft h t ht) by apply: valid_graft. +rewrite /graft pfjoin //=; apply/umem_eq=>//=; first by rewrite pfV2. +case=>k v; split=>[|H]. +- case/InUn; first by case/In_side/In_umfiltX; rewrite /= eqxx. + by case/In_side/In_omapX=>w H [] /inj_pair2 <-. +apply: InR; first by rewrite pfV2. +by apply/In_side/(In_omap _ H). +Qed. + +Lemma side_graftN h (t1 t2 : T) (ht2 : Ut t2) : + valid ht2 -> + disjoint_graft h ht2 -> + t1 != t2 -> + side_map Ut t1 (graft h t2 ht2) = side_map Ut t1 h. +Proof. +move=>V D N; case: (normalP h)=>[->|Vh]. +- by rewrite /graft pfundef undef_join. +have W : valid (graft h t2 ht2) by apply: valid_graft. +rewrite /graft pfjoin //=; apply/umem_eq=>/=. +- by rewrite pfV2. +- by rewrite pfV. +case=>k v; split. +- case/InUn; first by case/In_side/In_umfiltX=>_ /(In_side Ut). + case/In_side/In_omapX=>w /In_dom /= H [?]; subst t2. + by rewrite eqxx in N. +move/In_side=>H; apply/InL; first by rewrite pfV2. +by apply/In_side/In_umfiltX; rewrite /= eq_sym N. +Qed. + +End GraftLemmas. + + +(***********) +(* dom_map *) +(***********) + +(* Viewing domain as finite set instead of sequence. *) +(* Because sequences don't have PCM structure, *) +(* dom_maps facilitate proving disjointness of sequences *) +(* by enabling the pcm lemmas. *) + +Section DomMap. +Variables (K : ordType) (C : pred K) (V : Type). +Variables (U : union_map C V) (U' : union_map C unit). + +Definition dom_map (x : U) : U' := omap (fun => Some tt) x. + +HB.instance Definition _ := OmapFun.copy dom_map dom_map. + +Lemma dom_mapE x : dom (dom_map x) = dom x. +Proof. by apply: dom_omf_some. Qed. + +(* like pfjoinT but without validity condition *) +Lemma dom_mapUn (x y : U) : + dom_map (x \+ y) = dom_map x \+ dom_map y. +Proof. +case W : (valid (x \+ y)); first exact: pfjoinT. +move/invalidE: (negbT W)=>->; rewrite pfundef. +case: validUn W=>//. +- by move/invalidE=>->; rewrite pfundef undef_join. +- by move/invalidE=>->; rewrite pfundef 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 domeq_dom_mapL x : dom_eq (dom_map x) x. +Proof. by rewrite /dom_eq pfVE dom_mapE !eq_refl. Qed. + +Lemma domeq_dom_mapR x : dom_eq x (dom_map x). +Proof. by apply: domeq_sym; apply: domeq_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 C V). +Variables (U' : union_map C unit) (U'' : union_map C unit). + +Lemma dom_map_idemp (x : U) : dom_map (dom_map x : U') = dom_map x :> U''. +Proof. by rewrite /dom_map omap_omap. Qed. + +End DomMapIdempotent. + + +(*****************) +(* Map inversion *) +(*****************) + +Section MapInvert. +Variables (K V : ordType) (C : pred K) (C' : pred V). +Variables (U : union_map C V) (U' : union_map C' K). + +(* inverting f = flipping each pts k v in f into pts v k *) +Definition invert (f : U) : U' := + um_foldl (fun r k v => r \+ pts v k) Unit undef f. + +(* invert isn't morphism as pfV doesn't hold *) +(* but invert has several morphism properties *) + +Lemma invert_undef : invert undef = undef. +Proof. by rewrite /invert umfoldl_undef. Qed. + +Lemma invert0 : invert Unit = Unit. +Proof. by rewrite /invert umfoldl0. Qed. + +Lemma invertUn f1 f2 : + valid (f1 \+ f2) -> + invert (f1 \+ f2) = invert f1 \+ invert f2. +Proof. +move=>W; rewrite /invert umfoldlUn_frame ?unitR //. +by move=>*; rewrite joinAC. +Qed. + +Lemma invertPt k v : C k -> invert (pts k v) = pts v k. +Proof. by move=>H; rewrite /invert umfoldlPt unitL H. Qed. + +Lemma invertPtUn k v f : + valid (pts k v \+ f) -> + invert (pts k v \+ f) = pts v k \+ invert f. +Proof. by move=>W; rewrite invertUn // invertPt // (validPtUn_cond W). Qed. + +Lemma dom_invert f : valid (invert f) -> dom (invert f) =i range f. +Proof. +rewrite /invert/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_invert f : + valid (invert 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 invert_undef !valid_undef. +- by rewrite invert0 !valid_unit range0. +rewrite invertPtUn 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_invert ->. +Qed. + +(* The next few lemmas that depend on valid (invert f) *) +(* can be strenghtened to require just uniq (range f) and *) +(* all C' (range f). *) +(* However, the formulation with the hypothesis valid (invert f) *) +(* is packaged somewhat better, and is what's encountered in practice. *) + +Lemma range_invert f : + valid (invert f : U') -> + range (invert f) =i dom f. +Proof. +elim/um_indf: f=>[||k v f IH W P]. +- by rewrite invert_undef valid_undef. +- by rewrite invert0 range0 dom0. +rewrite invertPtUn // => W' x. +rewrite rangePtUn inE W' domPtUn inE W /=. +by case: eqP=>//= _; rewrite IH // (validR W'). +Qed. + +Lemma In_invert k v f : + valid (invert f) -> + (k, v) \In f <-> (v, k) \In invert f. +Proof. +elim/um_indf: f k v=>[||x w f IH W /(order_path_min (@trans K)) P] k v. +- by rewrite invert_undef valid_undef. +- by rewrite invert0; split=>/In0. +move=>W'; rewrite invertPtUn // !InPtUnE //; last by rewrite -invertPtUn. +rewrite IH; first by split; case=>[[->->]|]; auto. +rewrite !valid_invert rangePtUnK // (validR W) in W' *. +by case/and3P: W'=>_ /= /andP [_ ->] /andP [_ ->]. +Qed. + +Lemma uniq_range_invert f : uniq (range (invert f)). +Proof. +case: (normalP (invert f))=>[->|W]; first by 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/In_assocs/(In_invert _ _ W): H1=>H1. +move/mem_seqP/In_assocs/(In_invert _ _ W): H2=>H2. +by rewrite (In_fun H1 H2). +Qed. + +Lemma all_range_invert f : all C (range (invert f)). +Proof. +apply: (@sub_all _ (fun k => k \in dom f)); first by move=>x /dom_cond. +by apply/allP=>x H; rewrite -range_invert //; apply: range_valid H. +Qed. + +Lemma sorted_range_invert f : + valid (invert f) -> + sorted ord (range f) -> + sorted ord (range (invert f)). +Proof. +move=>W /ummonoP X; apply/ummonoP=>v v' k k'. +move/(In_invert _ _ W)=>H1 /(In_invert _ _ 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 MapInvert. + +Arguments invert [K V C C' U U']. + + +Section InvertLaws. +Variables (K V : ordType) (C : pred K) (C' : pred V). +Variables (U : union_map C V) (U' : union_map C' K). + +Lemma valid_invert_idemp (f : U) : + valid (invert (invert f : U') : U) = valid (invert f : U'). +Proof. by rewrite valid_invert uniq_range_invert all_range_invert !andbT. Qed. + +Lemma invert_idemp (f : U) : + valid (invert f : U') -> + invert (invert f : U') = f. +Proof. +move=>W; apply/umem_eq. +- by rewrite valid_invert_idemp. +- by move: W; rewrite valid_invert; case/and3P. +move=>x; split=>H. +- by apply/(In_invert _ _ W); apply/(@In_invert _ _ _ _ _ U) =>//; case: H. +case: x H=>k v /(In_invert _ _ W)/In_invert; apply. +by rewrite valid_invert_idemp. +Qed. + +End InvertLaws. + +Arguments In_invert [K V C C' U U' k v f]. + + +(***************************) +(* Composition of two maps *) +(***************************) + +(* composing maps as functions, rather than PCMs *) + +Section MapComposition. +Variables (K1 K2 : ordType) (C1 : pred K1) (C2 : pred K2) (V : Type). +Variables (Uf : union_map C1 K2) (Ug : union_map C2 V) (U : union_map 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 umcomp_subdom 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: umcomp_subdom. +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 umcompfPtE 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 umcompfPt g k v : + C1 k -> + um_comp g (pts k v) = + if find v g is Some w then pts k w else Unit. +Proof. by move=>H; rewrite umcompfPtE H. 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 // umcompfPtE (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 pfunit. + by case: ifP=>//; rewrite valid_unit. +rewrite umcompfUn // umcompfPtE (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 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). +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) -> 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. + +#[export] +Hint Resolve umall_undef umall0 : core. + + +(***********) +(* um_allb *) +(***********) + +(* decidable um_all *) + +Section MapAllDecidable. +Variables (K : ordType) (C : pred K) (V : Type) (U : union_map C V). +Implicit Types (f : U) (p : pred (K*V)). + +Definition um_allb p f := um_count p f == size (dom f). + +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 umallb_is_pcm_morph p : pcm_morph_axiom relT (um_allb p). +Proof. +rewrite /um_allb; split=>[|x y W _]; first by rewrite umcnt0 dom0. +split=>//; apply/idP/idP. +- by move/umallbP=>H; apply/andP; split; apply/umallbP=>k X; + apply: H; [apply: InL | apply: InR]. +case/andP=>/umallbP X1 /umallbP X2; apply/umallbP=>k. +by case/InUn; [apply: X1 | apply: X2]. +Qed. + +HB.instance Definition _ p := + isPCM_morphism.Build U bool (um_allb p) (umallb_is_pcm_morph p). +HB.instance Definition _ p := + isFull_PCM_morphism.Build U bool (um_allb p) (fun x y => erefl). + +Lemma umallb0 p : um_allb p Unit. +Proof. exact: pfunit. Qed. + +Lemma umallbUn p f1 f2 : + valid (f1 \+ f2) -> + um_allb p (f1 \+ f2) = um_allb p f1 && um_allb p f2. +Proof. exact: pfjoinT. Qed. + +(* bool isn't tpcm, so we can't declare umallb tpcm morphism *) +(* however, we can prove some properties about under *) + +Lemma umallb_undef p : um_allb p undef. +Proof. by rewrite /um_allb umcnt_undef dom_undef. 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 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]; first 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. + +(* big join and union maps *) + +Section BigOpsUM. +Variables (K : ordType) (C : pred K) (T : Type). +Variables (U : union_map C T). +Variables (I : Type) (f : I -> U). + +Lemma big_domUn (xs : seq I) : + dom (\big[join/Unit]_(i <- xs) f i) =i + [pred x | valid (\big[join/Unit]_(i <- xs) f i) & + has (fun i => x \in dom (f i)) xs]. +Proof. +elim: xs=>[|x xs IH] i; first by rewrite big_nil inE /= dom0 valid_unit. +rewrite big_cons /= inE domUn inE IH inE /=. +by case V : (valid _)=>//=; rewrite (validR V) /=. +Qed. + +Lemma big_domUnE (xs : seq I) a : + valid (\big[join/Unit]_(i <- xs) f i) -> + a \in dom (\big[join/Unit]_(i <- xs) f i) = + has (fun i => a \in dom (f i)) xs. +Proof. by move=>V; rewrite big_domUn inE V. Qed. + +(* we can construct a big validity, if we know that *) +(* the list contains unique elements *) +Lemma big_validV2I (xs : seq I) : + Uniq xs -> + (forall i, i \In xs -> valid (f i)) -> + (forall i j, i \In xs -> j \In xs -> i <> j -> valid (f i \+ f j)) -> + valid (\big[join/Unit]_(i <- xs) f i). +Proof. +elim: xs=>[|x xs IH] /=; first by rewrite big_nil valid_unit. +case=>X Uq H1 H2; rewrite big_cons validUnAE. +rewrite H1 /=; last by rewrite InE; left. +rewrite IH //=; last first. +- by move=>i j Xi Xj; apply: H2; rewrite InE; right. +- by move=>i Xi; apply: H1; rewrite InE; right. +apply/allP=>a /=; apply: contraL=>Dx; apply/negP. +rewrite big_domUn inE=>/andP [V] /hasPIn [y Y Dy]. +have : x <> y by move=>N; rewrite N in X; move/X: Y. +move/(H2 x y (or_introl erefl) (or_intror Y)). +by case: validUn=>// _ _ /(_ a Dx); rewrite Dy. +Qed. + +Lemma big_size_dom (xs : seq I) : + valid (\big[join/Unit]_(i <- xs) f i) -> + size (dom (\big[join/Unit]_(i <- xs) f i)) = + \sum_(i <- xs) (size (dom (f i))). +Proof. +elim: xs=>[|x xs IH] /=; first by rewrite !big_nil dom0. +rewrite !big_cons=>V; rewrite size_domUn //; congr (size _ + _). +by apply/IH/(validR V). +Qed. + +Lemma big_find_some (xs : seq I) a i v : + valid (\big[join/Unit]_(i <- xs) f i) -> + i \In xs -> + find a (f i) = some v -> + find a (\big[join/Unit]_(i <- xs) f i) = some v. +Proof. +elim: xs=>[|x xs IH /[swap]] //; rewrite big_cons InE. +case=>[<-{x}|Xi] V E; first by rewrite findUnL // (find_some E). +rewrite findUnR // big_domUnE ?(validR V) //=. +rewrite ifT; first by apply: IH (validR V) Xi E. +by apply/hasPIn; exists i=>//; apply: find_some E. +Qed. + +Lemma big_find_someD (xs : seq I) a i v : + i \In xs -> + a \in dom (f i) -> + find a (\big[join/Unit]_(x <- xs) f x) = Some v -> + find a (f i) = Some v. +Proof. +elim: xs v=>[|y xs IH] v //=; rewrite big_cons InE. +case=>[->|Xi] Da /[dup]/In_find/In_valid V; first by rewrite findUnL // Da. +rewrite findUnR // big_domUnE ?(validR V) //=. +by rewrite ifT; [apply: IH | apply/hasPIn; exists i]. +Qed. + +Lemma big_find_someX (xs : seq I) a v : + find a (\big[join/Unit]_(i <- xs) f i) = Some v -> + exists2 i, i \In xs & find a (f i) = Some v. +Proof. +elim: xs=>[|x xs IH]; first by rewrite big_nil find0E. +move/[dup]=>E1; rewrite big_cons=>/[dup] E /find_some. +rewrite domUn inE=>/andP [V] /orP [] D; last first. +- move: E; rewrite findUnR // D=>/IH [j Dj Xj]. + by exists j=>//; rewrite InE; right. +case/In_domX: D=>w /In_find /[dup] X. +move/(big_find_some (dom_valid (find_some E1)) (or_introl erefl)). +by rewrite E1; case=>->; exists x=>//; rewrite InE; left. +Qed. + +Lemma bigIn (xs : seq I) a i v : + valid (\big[join/Unit]_(i <- xs) f i) -> + i \In xs -> + (a, v) \In f i -> + (a, v) \In \big[join/Unit]_(i <- xs) f i. +Proof. by move=>V H /In_find/(big_find_some V H)/In_find. Qed. + +Lemma bigInD (xs : seq I) a x v : + x \In xs -> + a \in dom (f x) -> + (a, v) \In \big[join/Unit]_(i <- xs) f i -> + (a, v) \In f x. +Proof. by move=>X D /In_find/(big_find_someD X D)/In_find. Qed. + +Lemma bigInX (xs : seq I) a v : + (a, v) \In \big[join/Unit]_(i <- xs) f i -> + exists2 i, i \In xs & (a, v) \In f i. +Proof. by case/In_find/big_find_someX=>x X /In_find; exists x. Qed. + +End BigOpsUM. + +Prenex Implicits big_find_some big_find_someD. +Prenex Implicits big_find_someX bigIn bigInD bigInX. + +(* if the index type is eqtype, we can have a somewhat better *) +(* big validity lemma *) + +Section BigOpsUMEq. +Variables (K : ordType) (C : pred K) (T : Type) (U : union_map C T). +Variables (I : eqType) (f : I -> U). + +Lemma big_validP (xs : seq I) : + reflect + [/\ forall i, i \in xs -> valid (f i), + forall i k, i \in xs -> k \in dom (f i) -> count_mem i xs = 1 & + forall i j, i \in xs -> j \in xs -> i <> j -> valid (f i \+ f j)] + (valid (\big[join/Unit]_(i <- xs) f i)). +Proof. +elim: xs=>[|x xs IH] /=; first by rewrite big_nil valid_unit; constructor. +rewrite big_cons validUnAE. +case H1 : (valid (f x))=>/=; last first. +- by constructor; case=>/(_ x); rewrite inE eqxx H1=>/(_ erefl). +rewrite andbC; case: allP=>/= H2; last first. +- constructor; case=>H3 H4 H5; apply: H2=>k; rewrite big_domUn inE. + case/andP=>V /hasP [y]; case : (x =P y)=>[<- /[swap]|N X D]. + - by move/(H4 x); rewrite inE eqxx add1n=>/(_ erefl) []/count_memPn/negbTE ->. + by apply: dom_inNR (H5 _ _ _ _ N) D; rewrite inE ?eqxx ?X ?orbT. +case: {2 3}(valid _) / IH (erefl (valid (\big[join/Unit]_(j <- xs) f j))) +=> H V; constructor; last first. +- case=>H3 H4 H5; apply: H; split; last 1 first. + - by move=>i k X1 X2; apply: H5; rewrite inE ?X1 ?X2 orbT. + - by move=>i X; apply: H3; rewrite inE X orbT. + move=>i k X1 X2; move: (H4 i k); rewrite inE X1 orbT=>/(_ erefl X2). + case: (x =P i) X1 X2=>[<-{i} X1 X2|]; last by rewrite add0n. + by rewrite add1n; case=>/count_memPn; rewrite X1. +case: H=>H3 H4 H5; split; first by move=>i; rewrite inE=>/orP [/eqP ->|/H3]. +- move=>i k; rewrite inE eq_sym; case: (x =P i)=>[<- _|X] /= D; last first. + - by rewrite add0n; apply: H4 D. + rewrite add1n; congr (_.+1); apply/count_memPn. + apply: contraL (D)=>X; apply: H2; rewrite big_domUn V inE. + by apply/hasP; exists x. +move=>i j; rewrite !inE=>/orP [/eqP ->{i}| Xi] /orP [/eqP ->{j}|Xj] // N. +- rewrite validUnAE H1 H3 //=; apply/allP=>k D; apply: H2. + by rewrite big_domUn V inE; apply/hasP; exists j. +- rewrite validUnAE H3 ?H1 //=; apply/allP=>k; apply: contraL=>D; apply: H2. + by rewrite big_domUn V inE; apply/hasP; exists i. +by apply: H5 Xi Xj N. +Qed. + +End BigOpsUMEq. + +Section BigCatSeqUM. +Variables (K : ordType) (C : pred K) (T : Type) (U : union_map C T). +Variables (I : eqType) (g : I -> U). + +Lemma big_valid_dom_seq (xs : seq I) : + (valid (\big[join/Unit]_(i <- xs) g i) = + all (fun i => valid (g i)) xs && + uniq (\big[cat/[::]]_(i <- xs) dom (g i))) * + (dom (\big[join/Unit]_(i <- xs) g i) =i + if valid (\big[join/Unit]_(i <- xs) g i) then + \big[cat/[::]]_(i <- xs) dom (g i) + else [::]). +Proof. +elim: xs=>[|x xs [IH1 IH2]]; first by rewrite !big_nil valid_unit /= dom0. +rewrite !big_cons cat_uniq /=; split=>[|z]; last first. +- rewrite domUn inE; case V : (valid _)=>//=. + by rewrite mem_cat IH2 (validR V). +rewrite validUnAE IH1 -all_predC -!andbA uniq_dom /=. +case V : (valid (g x))=>//=; case A : (all _)=>//=. +rewrite [RHS]andbC; case Uq : (uniq _)=>//=. +by apply: eq_all_r=>i; rewrite IH2 IH1 A Uq. +Qed. + +Lemma big_valid_seq (xs : seq I) : + valid (\big[join/Unit]_(i <- xs) g i) = + all (fun i => valid (g i)) xs && + uniq (\big[cat/[::]]_(i <- xs) dom (g i)). +Proof. by rewrite big_valid_dom_seq. Qed. + +Lemma big_dom_seq (xs : seq I) : + dom (\big[join/Unit]_(i <- xs) g i) =i + if valid (\big[join/Unit]_(i <- xs) g i) then + \big[cat/[::]]_(i <- xs) dom (g i) + else [::]. +Proof. by move=>x; rewrite big_valid_dom_seq. Qed. + +End BigCatSeqUM. + +(* DEVCOMMENT: *) +(* remove "tests" for release *) +Lemma xx (f : umap nat nat) : 3 \in dom (free f 2). +rewrite domF -domF. +Abort. +(* /DEVCOMMENT *) From c4ba78e51d3a0819a6a9b31113a53233edf38868 Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Thu, 1 Aug 2024 18:29:29 +0200 Subject: [PATCH 03/93] renaming --- {theories => htt}/domain.v | 0 {theories => htt}/dune | 0 {theories => htt}/heapauto.v | 0 {theories => htt}/interlude.v | 0 {theories => htt}/model.v | 0 {theories => htt}/options.v | 0 6 files changed, 0 insertions(+), 0 deletions(-) rename {theories => htt}/domain.v (100%) rename {theories => htt}/dune (100%) rename {theories => htt}/heapauto.v (100%) rename {theories => htt}/interlude.v (100%) rename {theories => htt}/model.v (100%) rename {theories => htt}/options.v (100%) diff --git a/theories/domain.v b/htt/domain.v similarity index 100% rename from theories/domain.v rename to htt/domain.v diff --git a/theories/dune b/htt/dune similarity index 100% rename from theories/dune rename to htt/dune diff --git a/theories/heapauto.v b/htt/heapauto.v similarity index 100% rename from theories/heapauto.v rename to htt/heapauto.v diff --git a/theories/interlude.v b/htt/interlude.v similarity index 100% rename from theories/interlude.v rename to htt/interlude.v diff --git a/theories/model.v b/htt/model.v similarity index 100% rename from theories/model.v rename to htt/model.v diff --git a/theories/options.v b/htt/options.v similarity index 100% rename from theories/options.v rename to htt/options.v From d0cf27743c43883e8d5a06d813db9134c53da59e Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Thu, 1 Aug 2024 20:12:02 +0200 Subject: [PATCH 04/93] blah --- _CoqProject | 43 +- core/ordtype.v | 52 +- core/prelude.v | 214 ++++--- core/seqext.v | 337 ++++++++--- core/seqperm.v | 67 ++- core/slice.v | 11 + htt/domain.v | 1488 ++++++++++++++++++----------------------------- htt/interlude.v | 28 + htt/options.v | 7 +- 9 files changed, 1150 insertions(+), 1097 deletions(-) diff --git a/_CoqProject b/_CoqProject index b2abb00..233c167 100644 --- a/_CoqProject +++ b/_CoqProject @@ -1,11 +1,40 @@ --Q . htt --arg "-w -notation-overridden,-local-declaration,-redundant-canonical-projection,-projection-no-head-constant" +-Q core pcm +-Q pcm pcm +-Q htt htt +-Q examples htt +-docroot docs # where the documentation should go -theories/options.v -theories/interlude.v -theories/domain.v -theories/model.v -theories/heapauto.v +-arg -w -arg -notation-overridden +-arg -w -arg -redundant-canonical-projection + +core/options.v +core/axioms.v +core/prelude.v +core/pred.v +core/auto.v +core/seqext.v +core/slice.v +core/uslice.v +core/useqord.v +core/uconsec.v +core/ordtype.v +core/seqperm.v +core/finmap.v +core/rtc.v +pcm/pcm.v +pcm/autopcm.v +pcm/morphism.v +pcm/invertible.v +pcm/unionmap.v +pcm/natmap.v +pcm/automap.v +pcm/heap.v +pcm/mutex.v +htt/options.v +htt/interlude.v +htt/domain.v +htt/model.v +htt/heapauto.v examples/exploit.v examples/gcd.v examples/llist.v diff --git a/core/ordtype.v b/core/ordtype.v index 13072a1..1a83f23 100644 --- a/core/ordtype.v +++ b/core/ordtype.v @@ -21,7 +21,7 @@ limitations under the License. From HB Require Import structures. From Coq Require Import ssreflect ssrbool ssrfun. From mathcomp Require Import ssrnat eqtype seq path fintype. -From pcm Require Import options. +From pcm Require Import options seqext. Definition connected (T : eqType) (ord : rel T) := forall x y, x != y -> ord x y || ord y x. @@ -223,7 +223,7 @@ Definition oleq (T : ordType) (t1 t2 : T) := ord t1 t2 || (t1 == t2). Prenex Implicits ord oleq. Section Lemmas. -Variable T : ordType. +Context {T : ordType}. Implicit Types x y : T. Lemma ord_total x y : [|| ord x y, x == y | ord y x]. @@ -232,7 +232,17 @@ case E: (x == y)=>/=; first by rewrite orbT. by apply: connex; rewrite negbT. Qed. -Lemma nsym x y : ord x y -> ~ ord y x. +Lemma ord_neq (x y : T) : + ord x y -> + x != y. +Proof. +move=>H; apply/negP=>/eqP E. +by rewrite E irr in H. +Qed. + +Lemma nsym x y : + ord x y -> + ~ ord y x. Proof. by move=>E1 E2; move: (trans E1 E2); rewrite irr. Qed. Lemma antisym : antisymmetric (@ord T). @@ -258,30 +268,37 @@ Qed. Lemma subrel_ord : subrel (@ord T) oleq. Proof. exact: subrelUl. Qed. -Lemma sorted_oleq s : sorted (@ord T) s -> sorted (@oleq T) s. +Lemma sorted_oleq s : + sorted (@ord T) s -> + sorted (@oleq T) s. Proof. 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). + path r x s -> + path r x (filter p s). Proof. exact: (subseq_path tr (filter_subseq p s)). Qed. Lemma ord_sorted_eq (s1 s2 : seq T) : - sorted ord s1 -> sorted ord s2 -> s1 =i s2 -> s1 = s2. + sorted ord s1 -> + sorted ord s2 -> + s1 =i s2 -> + s1 = s2. Proof. exact/irr_sorted_eq/irr/(@trans _). Qed. Lemma oleq_ord_trans (n m p : T) : - oleq m n -> ord n p -> ord m p. + oleq m n -> + ord n p -> + ord m p. Proof. by case/orP; [apply: trans | move/eqP=>->]. Qed. Lemma ord_oleq_trans (n m p : T) : - ord m n -> oleq n p -> ord m p. + ord m n -> + oleq n p -> + ord m p. Proof. by move=>H /orP [|/eqP <- //]; apply: trans. Qed. End Lemmas. -Arguments orefl {T}. -Arguments otrans {T}. - #[export] Hint Resolve orefl irr trans connex otrans oantisym oleq_ord_trans : core. @@ -352,7 +369,6 @@ Proof. by case:oleqP=>// /ordW ->//. Qed. End Weakening. - (**********************************************) (* Building hierarchies for some basic orders *) (**********************************************) @@ -468,6 +484,14 @@ Canonical seq_ordType : ordType := Ordered.Pack (sort:=seq T) (Ordered.Class seq_ord_mix). End SeqOrdType. -#[deprecated(since="fcsl-pcm 1.4.0", note="Use ord_sorted_eq instead.")] -Notation eq_sorted_ord := ord_sorted_eq (only parsing). +Lemma sorted_cat_cons_cat (A : ordType) (l r : seq A) x : + sorted ord (l ++ x :: r) = + sorted ord (l ++ [::x]) && sorted ord (x::r). +Proof. +rewrite !(sorted_pairwise (@trans A)) cats1 pairwise_cat. +rewrite pairwise_rcons allrel_consr !pairwise_cons. +case/boolP: (all (ord^~ x) l)=>//= Hl. +case/boolP: (all (ord x) r)=>/= [Hr|_]; last by rewrite !andbF. +by rewrite (allrel_trans (@trans A) Hl Hr). +Qed. diff --git a/core/prelude.v b/core/prelude.v index f0fe244..46fac1b 100644 --- a/core/prelude.v +++ b/core/prelude.v @@ -18,7 +18,8 @@ limitations under the License. From HB Require Import structures. From Coq Require Import ssreflect ssrbool ssrfun Eqdep. -From mathcomp Require Import ssrnat seq eqtype choice fintype finfun. +From mathcomp Require Import ssrnat seq eqtype choice fintype finfun tuple. +From mathcomp Require Import perm fingroup. From pcm Require Import options axioms. (***********) @@ -30,23 +31,25 @@ From pcm Require Import options axioms. (* 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. +Arguments inj_pair2 {U P p x y}. (* 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). + (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. + +Arguments inj_pair {A B a1 a2 b1 b2}. (* 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). +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. @@ -57,7 +60,10 @@ Lemma idfun0E (U V : Type) (f : U -> V): (idfun \o f = f) * (f \o idfun = f). Proof. by []. Qed. -Lemma trans_eq A (x y z : A) : x = y -> x = z -> y = z. +Lemma trans_eq A (x y z : A) : + x = y -> + x = z -> + y = z. Proof. by move/esym; apply: eq_trans. Qed. (* Triples *) @@ -65,51 +71,60 @@ 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. +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. +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. +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). +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). +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). +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. (************) (* Products *) (************) -Inductive Prod3 (U1 U2 U3 : Type) := +Inductive Prod3 U1 U2 U3 := mk3 {proj31 : U1; proj32 : U2; proj33 : U3}. -Prenex Implicits proj31 proj32 proj33. - -Inductive Prod4 (U1 U2 U3 U4 : Type) := +Inductive Prod4 U1 U2 U3 U4 := mk4 {proj41 : U1; proj42 : U2; proj43 : U3; proj44 : U4}. -Prenex Implicits proj41 proj42 proj43 proj44. - -Inductive Prod5 (U1 U2 U3 U4 U5 : Type) := +Inductive Prod5 U1 U2 U3 U4 U5 := mk5 {proj51 : U1; proj52 : U2; proj53 : U3; proj54 : U4; proj55 : U5}. -Prenex Implicits proj51 proj52 proj53 proj54 proj55. - -Inductive Prod6 (U1 U2 U3 U4 U5 U6 : Type) := +Inductive Prod6 U1 U2 U3 U4 U5 U6 := mk6 {proj61 : U1; proj62 : U2; proj63 : U3; proj64 : U4; proj65 : U5; proj66 : U6}. -Prenex Implicits proj61 proj62 proj63 proj64 proj65 proj66. - -Inductive Prod7 (U1 U2 U3 U4 U5 U6 U7 : Type) := +Inductive Prod7 U1 U2 U3 U4 U5 U6 U7 := mk7 {proj71 : U1; proj72 : U2; proj73 : U3; proj74 : U4; proj75 : U5; proj76 : U6; proj77 : U7}. +Prenex Implicits proj31 proj32 proj33. +Prenex Implicits proj41 proj42 proj43 proj44. +Prenex Implicits proj51 proj52 proj53 proj54 proj55. +Prenex Implicits proj61 proj62 proj63 proj64 proj65 proj66. Prenex Implicits proj71 proj72 proj73 proj74 proj75 proj76 proj77. Definition eq3 (U1 U2 U3 : eqType) (x y : Prod3 U1 U2 U3) := @@ -128,31 +143,36 @@ Definition eq7 (U1 U2 U3 U4 U5 U6 U7 : eqType) (x y : Prod7 U1 U2 U3 U4 U5 U6 U7 proj74 x == proj74 y, proj75 x == proj75 y, proj76 x == proj76 y & proj77 x == proj77 y]. -Lemma eq3P U1 U2 U3 : Equality.axiom (@eq3 U1 U2 U3). +Lemma eq3P U1 U2 U3 : + Equality.axiom (@eq3 U1 U2 U3). Proof. rewrite /eq3; case=>x1 x2 x3 [y1 y2 y3] /=. by do ![case: eqP=>[->|]]; constructor=>//; case. Qed. -Lemma eq4P U1 U2 U3 U4 : Equality.axiom (@eq4 U1 U2 U3 U4). +Lemma eq4P U1 U2 U3 U4 : + Equality.axiom (@eq4 U1 U2 U3 U4). Proof. rewrite /eq4; case=>x1 x2 x3 x4 [y1 y2 y3 y4] /=. by do ![case: eqP=>[->|]]; constructor=>//; case. Qed. -Lemma eq5P U1 U2 U3 U4 U5 : Equality.axiom (@eq5 U1 U2 U3 U4 U5). +Lemma eq5P U1 U2 U3 U4 U5 : + Equality.axiom (@eq5 U1 U2 U3 U4 U5). Proof. rewrite /eq5; case=>x1 x2 x3 x4 x5 [y1 y2 y3 y4 y5] /=. by do ![case: eqP=>[->|]]; constructor=>//; case. Qed. -Lemma eq6P U1 U2 U3 U4 U5 U6 : Equality.axiom (@eq6 U1 U2 U3 U4 U5 U6). +Lemma eq6P U1 U2 U3 U4 U5 U6 : + Equality.axiom (@eq6 U1 U2 U3 U4 U5 U6). Proof. rewrite /eq6; case=>x1 x2 x3 x4 x5 x6 [y1 y2 y3 y4 y5 y6] /=. by do ![case: eqP=>[->|]]; constructor=>//; case. Qed. -Lemma eq7P U1 U2 U3 U4 U5 U6 U7 : Equality.axiom (@eq7 U1 U2 U3 U4 U5 U6 U7). +Lemma eq7P U1 U2 U3 U4 U5 U6 U7 : + Equality.axiom (@eq7 U1 U2 U3 U4 U5 U6 U7). Proof. rewrite /eq7; case=>x1 x2 x3 x4 x5 x6 x7 [y1 y2 y3 y4 y5 y6 y7] /=. by do ![case: eqP=>[->|]]; constructor=>//; case. @@ -160,20 +180,15 @@ Qed. HB.instance Definition _ (U1 U2 U3 : eqType) := hasDecEq.Build (Prod3 U1 U2 U3) (@eq3P U1 U2 U3). - HB.instance Definition _ (U1 U2 U3 U4 : eqType) := hasDecEq.Build (Prod4 U1 U2 U3 U4) (@eq4P U1 U2 U3 U4). - HB.instance Definition _ (U1 U2 U3 U4 U5 : eqType) := hasDecEq.Build (Prod5 U1 U2 U3 U4 U5) (@eq5P U1 U2 U3 U4 U5). - HB.instance Definition _ (U1 U2 U3 U4 U5 U6 : eqType) := hasDecEq.Build (Prod6 U1 U2 U3 U4 U5 U6) (@eq6P U1 U2 U3 U4 U5 U6). - HB.instance Definition _ (U1 U2 U3 U4 U5 U6 U7 : eqType) := hasDecEq.Build (Prod7 U1 U2 U3 U4 U5 U6 U7) (@eq7P U1 U2 U3 U4 U5 U6 U7). - (***************************) (* operations on functions *) (***************************) @@ -189,13 +204,15 @@ 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. -Lemma compE A B C (g : B -> C) (f : A -> B) x : g (f x) = (g \o f) x. +Lemma compE A B C (g : B -> C) (f : A -> B) x : + g (f x) = (g \o f) x. 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. +Notation "f1 \* f2" := (fprod f1 f2) + (at level 42, left associativity) : fun_scope. (* product morphism, aka. fork/fanout/fsplice *) Definition pmorphism A B1 B2 (f1 : A -> B1) (f2 : A -> B2) := @@ -227,29 +244,20 @@ Notation "[ ** f1 , f2 , f3 , f4 & f5 ]" := ((((f1 \** f2) \** f3) \** f4) \** f 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 , P6 , P7 , P8 , P9 , P10 & P11 ]" (at level 0, format "'[hv' [ /\ '[' P1 , '/' P2 , '/' P3 , '/' P4 , '/' P5 , '/' P6 , '/' P7 , '/' P8 , '/' P9 , '/' P10 ']' '/ ' & P11 ] ']'"). - Reserved Notation "[ /\ P1 , P2 , P3 , P4 , P5 , P6 , P7 , P8 , P9 , P10 , P11 & P12 ]" (at level 0, format "'[hv' [ /\ '[' P1 , '/' P2 , '/' P3 , '/' P4 , '/' P5 , '/' P6 , '/' P7 , '/' P8 , '/' P9 , '/' P10 , '/' P11 ']' '/ ' & P12 ] ']'"). - - 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 @@ -526,19 +534,23 @@ Lemma iffE (b1 b2 : bool) : b1 = b2 <-> (b1 <-> b2). Proof. by split=>[->|] //; move/iffPb/eqP. Qed. Lemma subsetC T (p q : mem_pred T) : - {subset p <= q} -> {subset predC q <= predC p}. + {subset p <= q} -> + {subset predC q <= predC p}. Proof. by move=>H x; apply: contra (H x). Qed. Lemma subsetR T (p q : mem_pred T) : - {subset p <= predC q} -> {subset q <= predC p}. + {subset p <= predC q} -> + {subset q <= predC p}. Proof. by move=>H x; apply: contraL (H x). Qed. Lemma subsetL T (p q : mem_pred T) : - {subset predC p <= q} -> {subset predC q <= p}. + {subset predC p <= q} -> + {subset predC q <= p}. Proof. by move=>H x; apply: contraR (H x). Qed. Lemma subsetLR T (p q : mem_pred T) : - {subset predC p <= predC q} -> {subset q <= p}. + {subset predC p <= predC q} -> + {subset q <= p}. Proof. by move=>H x; apply: contraLR (H x). Qed. Lemma subset_disj T (p q r : mem_pred T) : @@ -557,6 +569,12 @@ move=>H1 H2 D1; apply: subset_disj H1 _ => x H /H2. by apply: D1 H. Qed. +Lemma implyb_trans a b c : + a ==> b -> + b ==> c -> + a ==> c. +Proof. by case: a=>//=->. Qed. + (**************) (* empty type *) (**************) @@ -690,42 +708,43 @@ HB.instance Definition _ : isFinite Side := isFinite.Build Side Side_enumP. HB.instance Definition _ := isInhabited.Build Side (inhabits LL). (*************) -(* sequences *) +(* Sequences *) (*************) +(* folds *) (* TODO upstream to mathcomp *) -Section Fold. Lemma map_foldr {T1 T2} (f : T1 -> T2) xs : - map f xs = foldr (fun x ys => f x :: ys) [::] xs. + map f xs = foldr (fun x ys => f x :: ys) [::] xs. Proof. by []. Qed. Lemma fusion_foldr {T R Q} (g : R -> Q) f0 f1 z0 z1 (xs : seq T) : - (forall x y, g (f0 x y) = f1 x (g y)) -> g z0 = z1 -> - g (foldr f0 z0 xs) = foldr f1 z1 xs. + (forall x y, g (f0 x y) = f1 x (g y)) -> + g z0 = z1 -> + g (foldr f0 z0 xs) = foldr f1 z1 xs. Proof. by move=>Hf Hz; elim: xs=>//= x xs <-. Qed. Lemma fusion_foldl {T R Q} (g : R -> Q) f0 f1 z0 z1 (xs : seq T) : - (forall x y, g (f0 x y) = f1 (g x) y) -> g z0 = z1 -> - g (foldl f0 z0 xs) = foldl f1 z1 xs. + (forall x y, g (f0 x y) = f1 (g x) y) -> + g z0 = z1 -> + g (foldl f0 z0 xs) = foldl f1 z1 xs. Proof. move=>Hf Hz; elim: xs z0 z1 Hz =>//= x xs IH z0 z1 Hz. by apply: IH; rewrite Hf Hz. Qed. Lemma foldl_foldr {T R} (f : R -> T -> R) z xs : - foldl f z xs = foldr (fun b g x => g (f x b)) id xs z. + foldl f z xs = foldr (fun b g x => g (f x b)) id xs z. Proof. by elim: xs z=>/=. Qed. Lemma foldr_foldl {T R} (f : T -> R -> R) z xs : - foldr f z xs = foldl (fun g b x => g (f b x)) id xs z. + foldr f z xs = foldl (fun g b x => g (f b x)) id xs z. Proof. elim/last_ind: xs z=>//= xs x IH z. by rewrite foldl_rcons -IH foldr_rcons. Qed. -End Fold. - +(* pmap *) (* TODO upstream to mathcomp *) Lemma pmap_pcomp {S T U} (f : T -> option U) (g : S -> option T) s : pmap (pcomp f g) s = pmap f (pmap g s). @@ -734,12 +753,10 @@ Proof. by elim: s=>//= x s ->; rewrite /pcomp; case: (g x). Qed. (* sequence prefixes *) (* Two helper concepts for searching in sequences: *) -(* *) (* - onth: like nth, but returns None when the element is not found *) (* - Prefix: a prefix relation on sequences, used for growing *) (* interpretation contexts *) - Fixpoint onth A (s : seq A) n : option A := if s is x::sx then if n is nx.+1 then onth sx nx else Some x else None. @@ -789,7 +806,9 @@ elim: s n=>/= [|a s IH] n /=; first by apply: nth_nil. by case: n. Qed. -Lemma onth_nth x0 n s : n < size s -> onth s n = Some (nth x0 s n). +Lemma onth_nth x0 n s : + n < size s -> + onth s n = Some (nth x0 s n). Proof. elim: s n=>//= a s IH n. by rewrite ltnS; case: n. @@ -845,25 +864,38 @@ End SeqPrefix. #[export] Hint Resolve Prefix_refl : core. -Lemma onth_mem (A : eqType) (s : seq A) n x : - onth s n = Some x -> x \in s. +(* when A : eqType *) + +Section SeqPrefixEq. +Variable A : eqType. +Implicit Type s : seq A. + +Lemma onth_mem s n x : + onth s n = Some x -> + x \in s. Proof. by elim: s n=>//= a s IH [[->]|n /IH]; rewrite inE ?eq_refl // orbC =>->. Qed. -Lemma onth_index (A : eqType) (s : seq A) x : - onth s (index x s) = if x \in s then Some x else None. +Lemma onth_index (s : seq A) x : + onth s (index x s) = + if x \in s then Some x else None. Proof. by elim: s=>//=h s IH; rewrite inE eq_sym; case: eqP=>//= ->. Qed. -Lemma PrefixP (A : eqType) (s1 s2 : seq A) : +Lemma PrefixP (s1 s2 : seq A) : reflect (Prefix s1 s2) (prefix s1 s2). Proof. apply/(equivP (prefixP (s1:=s1) (s2:=s2))). by apply: iff_sym; exact: PrefixE. Qed. +End SeqPrefixEq. + + + + (******************************) (* Some commuting conversions *) (******************************) @@ -1057,6 +1089,46 @@ Qed. End FinFunMap. +(* surgery on tuples and finfuns *) + +Section OnthCodom. +Variable A : Type. + +Lemma onth_tnth {n} (s : n.-tuple A) (i : 'I_n) : + onth s i = Some (tnth s i). +Proof. +elim: n s i =>[|n IH] s i; first by case: i. +case/tupleP: s=>/=x s; case: (unliftP ord0 i)=>[j|]-> /=. +- by rewrite tnthS. +by rewrite tnth0. +Qed. + +Lemma onth_codom {n} (i : 'I_n) (f: {ffun 'I_n -> A}) : + onth (fgraph f) i = Some (f i). +Proof. +pose i' := cast_ord (esym (card_ord n)) i. +move: (@tnth_fgraph _ _ f i'); rewrite (enum_val_ord) {2}/i' cast_ordKV=><-. +by rewrite (onth_tnth (fgraph f) i'). +Qed. + +End OnthCodom. + +(* ffun and permutation *) +Section PermFfun. +Variables (I : finType) (A : Type). + +Definition pffun (p : {perm I}) (f : {ffun I -> A}) := + [ffun i => f (p i)]. + +Lemma pffunE1 (f : {ffun I -> A}) : pffun 1%g f = f. +Proof. by apply/ffunP=>i; rewrite !ffunE permE. Qed. + +Lemma pffunEM (p p' : {perm I}) (f : {ffun I -> A}) : + pffun (p * p') f = pffun p (pffun p' f). +Proof. by apply/ffunP => i; rewrite !ffunE permM. Qed. + +End PermFfun. + (* Tagging *) diff --git a/core/seqext.v b/core/seqext.v index 75457cf..6a1eb77 100644 --- a/core/seqext.v +++ b/core/seqext.v @@ -12,13 +12,24 @@ limitations under the License. *) From Coq Require Import ssreflect ssrbool ssrfun. -From mathcomp Require Import ssrnat seq path eqtype choice fintype bigop. +From mathcomp Require Import ssrnat seq eqtype path choice fintype bigop perm. From pcm Require Import options prelude pred. (*********************) (* Extensions to seq *) (*********************) +(* TODO upstream to mathcomp *) + +Lemma rcons_nseq {A} n (x : A) : + rcons (nseq n x) x = nseq n.+1 x. +Proof. by elim: n=>//=n ->. Qed. + +Lemma behead_rcons {A} (xs : seq A) x : + 0 < size xs -> + behead (rcons xs x) = rcons (behead xs) x. +Proof. by case: xs. Qed. + Lemma nilp_hasPn {A} (s : seq A) : nilp s = ~~ has predT s. Proof. by case: s. Qed. @@ -31,10 +42,12 @@ Lemma filter_swap {A} (s : seq A) p1 p2 : Proof. by rewrite -!filter_predI filter_predIC. Qed. (* TODO contribute to mathcomp? *) -Lemma map_nilp {A B} (f : A -> B) (s : seq A) : nilp (map f s) = nilp s. +Lemma map_nilp {A B} (f : A -> B) (s : seq A) : + nilp (map f s) = nilp s. Proof. by rewrite /nilp; case: s. Qed. -Lemma filter_nilp {A} (p : pred A) (s : seq A) : nilp (filter p s) = ~~ has p s. +Lemma filter_nilp {A} (p : pred A) (s : seq A) : + nilp (filter p s) = ~~ has p s. Proof. by rewrite /nilp size_filter -leqn0 leqNgt has_count. Qed. Lemma head_map {P Q} (f : P -> Q) z (s : seq P) : @@ -49,14 +62,14 @@ elim: s1 s2=>/= [|x1 s1 IH] [|x2 s2] //=. by congr cons. Qed. -Corollary zip_mapl {P Q R} (f : P -> R) (s1 : seq P) (s2 : seq Q) : - zip (map f s1) s2 = - map (fun '(x1,x2) => (f x1,x2)) (zip s1 s2). +Lemma zip_mapl {P Q R} (f : P -> R) (s1 : seq P) (s2 : seq Q) : + zip (map f s1) s2 = + map (fun '(x1,x2) => (f x1,x2)) (zip s1 s2). Proof. by rewrite -{1}(map_id s2) zip_map2. Qed. -Corollary zip_mapr {P Q S} (g : Q -> S) (s1 : seq P) (s2 : seq Q) : - zip s1 (map g s2) = - map (fun '(x1,x2) => (x1,g x2)) (zip s1 s2). +Lemma zip_mapr {P Q S} (g : Q -> S) (s1 : seq P) (s2 : seq Q) : + zip s1 (map g s2) = + map (fun '(x1,x2) => (x1,g x2)) (zip s1 s2). Proof. by rewrite -{1}(map_id s1) zip_map2. Qed. Lemma drop_take_id {A} x (s : seq A) : drop x (take x s) = [::]. @@ -78,7 +91,7 @@ by rewrite drop_oversize // size_take_min geq_min Hxy. Qed. Section LemmasEq. -Variables (A : eqType). +Variables A : eqType. Implicit Type xs : seq A. (* With A : Type, we have the In_split lemma. *) @@ -86,23 +99,66 @@ Implicit Type xs : seq A. (* not only return the split of xs, but the split of xs *) (* that uses the first occurrence of x is xs *) Lemma in_split xs x : - x \in xs -> exists xs1 xs2, xs = xs1 ++ x :: xs2 /\ x \notin xs1. + x \in xs -> + exists xs1 xs2, + xs = xs1 ++ x :: xs2 /\ x \notin xs1. Proof. rewrite -has_pred1; case/split_find=>_ s1 s2 /eqP ->. by rewrite has_pred1=>H; exists s1, s2; rewrite -cats1 -catA. Qed. +(* a weaker form of in_split *) +Lemma mem_split (x : A) s : + x \in s -> + exists s1 s2, s = s1 ++ [:: x] ++ s2. +Proof. by case/in_split=>s1 [s2][H _]; exists s1, s2. Qed. + +Lemma mem_split_uniq (x : A) s : + x \in s -> uniq s -> + exists s1 s2, [/\ s = s1 ++ [:: x] ++ s2, + uniq (s1 ++ s2) & + x \notin s1 ++ s2]. +Proof. +move=>/[swap] Hu /mem_split [s1 [s2 H]]. +exists s1, s2; move: Hu. +by rewrite H uniq_catCA cons_uniq; case/andP. +Qed. + +Lemma perm_cons2 (x y : A) s : + perm_eq [:: x, y & s] [:: y, x & s]. +Proof. +by rewrite (_ : [::x,y & s] = [::x] ++ [::y] ++ s) // + (_ : [::y,x & s] = [::y] ++ [::x] ++ s) // perm_catCA. +Qed. + +Lemma all_notin (p : pred A) xs y : + all p xs -> + ~~ p y -> + y \notin xs. +Proof. by move/allP=>Ha; apply/contra/Ha. Qed. + +Lemma subset_all a (s1 s2 : seq A) : + {subset s1 <= s2} -> + all a s2 -> + all a s1. +Proof. by move=>Hs /allP Ha1; apply/allP=>s /Hs /Ha1. Qed. + Lemma subset_nilR xs : - {subset xs <= [::]} -> xs = [::]. + {subset xs <= [::]} -> + xs = [::]. Proof. by case: xs=>// x xs /(_ x); rewrite inE eqxx=>/(_ erefl). Qed. Lemma subset_nil xs ys : - {subset xs <= ys} -> ys = [::] -> xs = [::]. + {subset xs <= ys} -> + ys = [::] -> + xs = [::]. Proof. by move=>X E; move: E X=>->; apply: subset_nilR. Qed. -Lemma all_mem xs ys : +Lemma all_mem xs ys : reflect {subset ys <= xs} (all [mem xs] ys). -Proof. by case: allP=>H; constructor; [move=>x /H | move=>X; apply: H=>x /X]. Qed. +Proof. +by case: allP=>H; constructor; [move=>x /H | move=>X; apply: H=>x /X]. +Qed. Lemma all_predC_sym xs ys : all [predC xs] ys = all [predC ys] xs. @@ -137,8 +193,8 @@ split; first by exact: memNindex. by move=>E; rewrite -index_mem E ltnn. Qed. -Corollary index_sizeE x xs : - reflect (index x xs = size xs) (x \notin xs). +Lemma index_sizeE x xs : + reflect (index x xs = size xs) (x \notin xs). Proof. by apply/(equivP idP)/index_memN. Qed. Lemma size0nilP xs : @@ -172,14 +228,16 @@ by rewrite eqn0Ngt -has_count; apply/hasPn=>x /H; case: (p x)=>//; apply. Qed. Lemma filter_pred1 x xs : - x \notin xs -> filter (pred1 x) xs = [::]. + x \notin xs -> + filter (pred1 x) xs = [::]. Proof. move=>H; apply/eqP; apply/filter_nilP=>z /eqP ->. by rewrite (negbTE H). Qed. Lemma filter_predC1 x xs : - x \notin xs -> filter (predC1 x) xs = xs. + x \notin xs -> + filter (predC1 x) xs = xs. Proof. by move=>H; apply/all_filterP/allP=>y /=; case: eqP=>// ->; apply/contraL. Qed. @@ -189,7 +247,9 @@ Lemma filter_mem_sym (s1 s2 : seq A) : Proof. by move=>x; rewrite !mem_filter andbC. Qed. Lemma index_inj xs x y : - x \in xs -> index x xs = index y xs -> x = y. + x \in xs -> + index x xs = index y xs -> + x = y. Proof. elim: xs=>[|k xs IH] //=; rewrite inE eq_sym. case: eqP=>[->{k} _|_ /= S]; case: eqP=>// _ []; apply: IH S. @@ -215,7 +275,8 @@ Qed. (* if the list is not empty, the default value in head doesn't matter *) Lemma head_dflt (x1 x2 x : A) xs : - x \in xs -> head x1 xs = head x2 xs. + x \in xs -> + head x1 xs = head x2 xs. Proof. by case: xs. Qed. Lemma mem_head (x : A) xs : head x xs \in x :: xs. @@ -223,25 +284,30 @@ Proof. by case: xs=>[|y ys]; rewrite !inE //= eqxx orbT. Qed. (* a common pattern of using mem_head that avoids forward reasoning *) Lemma mem_headI (x : A) xs a : - a = head x xs -> a \in x :: xs. + a = head x xs -> + a \in x :: xs. Proof. by move=>->; apply: mem_head. Qed. Lemma head_nilp (x : A) xs : - x \notin xs -> head x xs = x -> nilp xs. + x \notin xs -> + head x xs = x -> + nilp xs. Proof. elim: xs=>[|y ys IH] //= H1 H2. by rewrite H2 inE eqxx /= in H1. Qed. Lemma head_notin (x y : A) xs : - y \in xs -> x \notin xs -> head x xs != x. + y \in xs -> + x \notin xs -> + head x xs != x. Proof. move=>Y X; apply/negP=>/eqP; move/(head_nilp X)/nilP=>E. by rewrite E in Y. Qed. (* weaker form of in_mask *) -(* TODO contribute to mathcomp *) +(* TODO upstream to mathcomp *) Lemma in_mask_count x m xs : count_mem x xs <= 1 -> x \in mask m xs = (x \in xs) && nth false m (index x xs). @@ -292,7 +358,8 @@ Lemma findlast_auxE oi0 p s : (if k == size s then oi0.1 else Some (oi0.2 + (size s - k).-1), oi0.2 + size s). Proof. -rewrite /findlast_aux; elim: s oi0=>/= [|x s IH] [o0 i0] /=; first by rewrite addn0. +rewrite /findlast_aux; elim: s oi0=>/= [|x s IH] [o0 i0] /=. +- by rewrite addn0. rewrite IH /= rev_cons -cats1 find_cat /= has_find. move: (find_size p (rev s)); rewrite size_rev; case: ltngtP=>// H _. - case: eqP=>[E|_]; first by rewrite E ltnNge leqnSn in H. @@ -306,29 +373,32 @@ Definition findlast p s : nat := let '(o, i) := findlast_aux (None, 0) p s in odflt i o. (* finding last is finding first in reversed list and flipping indices *) -Corollary findlastE p s : - findlast p s = - if has p s then (size s - seq.find p (rev s)).-1 - else size s. +Lemma findlastE p s : + findlast p s = + if has p s then (size s - seq.find p (rev s)).-1 else size s. Proof. rewrite /findlast findlast_auxE /= !add0n -has_rev; case/boolP: (has p (rev s)). - by rewrite has_find size_rev; case: ltngtP. by move/hasNfind=>->; rewrite size_rev eqxx. Qed. -Lemma findlast_size p s : findlast p s <= size s. +Lemma findlast_size p s : + findlast p s <= size s. Proof. rewrite findlastE; case: ifP=>// _. by rewrite -subnS; apply: leq_subr. Qed. -Lemma has_findlast p s : has p s = (findlast p s < size s). +Lemma has_findlast p s : + has p s = (findlast p s < size s). Proof. symmetry; rewrite findlastE; case: ifP=>H /=; last by rewrite ltnn. by rewrite -subnS /= ltn_subrL /=; case: s H. Qed. -Lemma hasNfindlast p s : ~~ has p s -> findlast p s = size s. +Lemma hasNfindlast p s : + ~~ has p s -> + findlast p s = size s. Proof. by rewrite has_findlast; case: ltngtP (findlast_size p s). Qed. Lemma findlast0 p x : findlast p [::] = 0. @@ -373,14 +443,18 @@ case: (p x)=>/=; first by rewrite andbF addn0. by rewrite andbT addn1. Qed. -Lemma nth_findlast x0 p s : has p s -> p (nth x0 s (findlast p s)). +Lemma nth_findlast x0 p s : + has p s -> + p (nth x0 s (findlast p s)). Proof. rewrite findlastE=>/[dup] E ->; rewrite -has_rev in E. rewrite -subnS -nth_rev; last by rewrite -size_rev -has_find. by apply: nth_find. Qed. -Lemma has_drop p s i : has p s -> has p (drop i s) = (i <= findlast p s). +Lemma has_drop p s i : + has p s -> + has p (drop i s) = (i <= findlast p s). Proof. rewrite findlastE=>/[dup] E ->; rewrite -has_rev in E. have Hh: 0 < size s - find p (rev s). @@ -398,7 +472,8 @@ rewrite -ltnNge subnS prednK //. by apply/leq_trans/Hi; exact: leq_subr. Qed. -Lemma find_geq p s i : has p (drop i s) -> i <= findlast p s. +Lemma find_geq p s i : + has p (drop i s) -> i <= findlast p s. Proof. case/boolP: (has p s)=>Hp; first by rewrite (has_drop _ Hp). suff: ~~ has p (drop i s) by move/negbTE=>->. @@ -406,7 +481,8 @@ move: Hp; apply: contra; rewrite -{2}(cat_take_drop i s) has_cat=>->. by rewrite orbT. Qed. -Lemma find_leq_last p s : find p s <= findlast p s. +Lemma find_leq_last p s : + find p s <= findlast p s. Proof. rewrite findlastE. case/boolP: (has p s)=>[|_]; last by apply: find_size. @@ -448,7 +524,7 @@ Qed. End FindLast. Section FindLastEq. -Variables (T : eqType). +Variables T : eqType. Implicit Type s : seq T. Definition indexlast (x : T) : seq T -> nat := findlast (pred1 x). @@ -487,13 +563,16 @@ Lemma indexlast_rcons x y s : else (size s).+1. Proof. by rewrite /indexlast findlast_rcons has_pred1. Qed. -Lemma index_geq x s i : x \in drop i s -> i <= indexlast x s. +Lemma index_geq x s i : + x \in drop i s -> + i <= indexlast x s. Proof. by rewrite -has_pred1; apply: find_geq. Qed. Lemma index_leq_last x s : index x s <= indexlast x s. Proof. by apply: find_leq_last. Qed. -Lemma indexlast_count x s : count_mem x s <= 1 <-> index x s = indexlast x s. +Lemma indexlast_count x s : + count_mem x s <= 1 <-> index x s = indexlast x s. Proof. elim: s=>//= h t IH; rewrite indexlast_cons. case: eqP=>/= ?; last first. @@ -503,13 +582,17 @@ rewrite add1n ltnS leqn0; split. by case: ifP=>//= /count_memPn->. Qed. -Corollary index_lt_last x s : 1 < count_mem x s -> index x s < indexlast x s. +Lemma index_lt_last x s : + 1 < count_mem x s -> + index x s < indexlast x s. Proof. move=>H; move: (index_leq_last x s); rewrite leq_eqVlt. by case: eqP=>//= /indexlast_count; case: leqP H. Qed. -Corollary indexlast_uniq x s : uniq s -> index x s = indexlast x s. +Lemma indexlast_uniq x s : + uniq s -> + index x s = indexlast x s. Proof. move=>H; apply/indexlast_count. by rewrite count_uniq_mem //; apply: leq_b1. @@ -523,7 +606,8 @@ by move=>E; rewrite -indexlast_mem E ltnn. Qed. Lemma index_last_inj x y s : - (x \in s) || (y \in s) -> index x s = indexlast y s -> x = y. + (x \in s) || (y \in s) -> + index x s = indexlast y s -> x = y. Proof. elim: s=>[|k s IH] //=; rewrite !inE indexlast_cons !(eq_sym k). case: eqP=>[{k}<- _|_ /= S]; first by case: eqP=>//=. @@ -534,7 +618,9 @@ by case: eqP=>//; rewrite orbF=>_ H []; apply: IH; rewrite H. Qed. Lemma indexlast_inj x y s : - x \in s -> indexlast x s = indexlast y s -> x = y. + x \in s -> + indexlast x s = indexlast y s -> + x = y. Proof. elim: s=>[|k s IH] //=; rewrite inE eq_sym !indexlast_cons. case: eqP=>[->{k} _|_ /= S] /=. @@ -569,8 +655,7 @@ End FindLastEq. (* finding all occurrences *) Section FindAll. - -Variables (T : Type). +Variables T : Type. Implicit Types (x : T) (p : pred T) (s : seq T). (* helper function for finding all occurrences in a single pass with difference lists *) @@ -581,7 +666,8 @@ Lemma findall_auxE oi0 p s : (forall s1 s2 : seq nat, oi0.1 (s1 ++ s2) = oi0.1 s1 ++ s2) -> let: (rs, ix) := findall_aux oi0 p s in (forall s' : seq nat, - rs s' = oi0.1 (unzip1 (filter (p \o snd) (zip (iota oi0.2 (size s)) s))) ++ s') + rs s' = oi0.1 (unzip1 (filter (p \o snd) + (zip (iota oi0.2 (size s)) s))) ++ s') /\ ix = oi0.2 + size s. Proof. move; rewrite /findall_aux; elim: s oi0=>[|x s IH] [o0 i0] /= H0. @@ -617,14 +703,14 @@ rewrite (eq_filter (a2:=(p \o snd))); last by case. by apply: eq_map; case. Qed. -Corollary findall_cons p x s : - findall p (x::s) = - if p x then 0 :: map S (findall p s) else map S (findall p s). +Lemma findall_cons p x s : + findall p (x::s) = + if p x then 0 :: map S (findall p s) else map S (findall p s). Proof. by rewrite -cat1s findall_cat /= findallE /=; case: (p x). Qed. -Corollary findall_rcons p x s : - findall p (rcons s x) = - if p x then rcons (findall p s) (size s) else findall p s. +Lemma findall_rcons p x s : + findall p (rcons s x) = + if p x then rcons (findall p s) (size s) else findall p s. Proof. rewrite -cats1 findall_cat /= !findallE /=; case: (p x)=>/=. - by rewrite addn0 cats1. @@ -671,8 +757,7 @@ Qed. End FindAll. Section FindAllEq. - -Variables (T : eqType). +Variables T : eqType. Implicit Type s : seq T. Definition indexall (x : T) : seq T -> seq nat := findall (pred1 x). @@ -694,12 +779,12 @@ Proof. by rewrite /indexlast /indexall findall_last. Qed. Lemma indexall_count1 x s : count_mem x s <= 1 -> - indexall x s = if x \in s then [::index x s] else [::]. + indexall x s = if x \in s then [:: index x s] else [::]. Proof. by rewrite /indexall /index -has_pred1; apply: findall_count1. Qed. Corollary indexall_uniq x s : uniq s -> - indexall x s = if x \in s then [::index x s] else [::]. + indexall x s = if x \in s then [:: index x s] else [::]. Proof. by move=>U; apply: indexall_count1; rewrite (count_uniq_mem _ U) leq_b1. Qed. End FindAllEq. @@ -707,10 +792,12 @@ End FindAllEq. (* Interaction of filter/last/index *) Section FilterLastIndex. -Variables (A : eqType). +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. +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. @@ -719,48 +806,64 @@ 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. +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. -Lemma last_notin_nilp k (s : seq A) : ~~ nilp s -> k \notin s -> last k s != k. +Lemma last_notin_nilp k (s : seq A) : + ~~ nilp s -> + k \notin s -> + last k s != k. Proof. move=>N; apply: (last_notin (x := head k s)). by case: s N=>//= x s _; rewrite inE eqxx. 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. +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. + last k s != k -> + forall x, last x s = last k s. Proof. by elim: s k=>[|k s IH] y //=; rewrite eqxx. Qed. Lemma last_changeE2 k (s : seq A) : - last k s != k -> forall x, x \notin s -> last x s != x. + 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 == [::]). +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 = [::]. +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 and rcons *) Lemma rcons_lastX x y (s : seq A) : - x \in s -> exists s', s = rcons s' (last y s). + x \in s -> + exists s', s = rcons s' (last y s). Proof. elim/last_ind: s=>[|ks k IH] //=. by rewrite last_rcons; exists ks. Qed. Lemma rcons_lastP x (s : seq A) : - reflect (exists s', s = rcons s' (last x s)) (last x s \in s). + reflect (exists s', s = rcons s' (last x s)) + (last x s \in s). Proof. case X : (last x s \in s); constructor; first by apply: rcons_lastX X. case=>s' E; move/negP: X; elim. @@ -768,7 +871,8 @@ by rewrite E last_rcons mem_rcons inE eqxx. Qed. Lemma rcons_lastXP x y (s : seq A) : - reflect (exists s', s = rcons s' x) ((x == last y s) && (x \in s)). + reflect (exists s', s = rcons s' x) + ((x == last y s) && (x \in s)). Proof. case: eqP=>[->|N]; first by apply: rcons_lastP. by constructor; case=>s' E; elim: N; rewrite E last_rcons. @@ -787,7 +891,9 @@ Qed. (* last has bigger index than anything in x *) Lemma index_last_mono x k (s : seq A) : - uniq s -> x \in s -> index x s <= index (last k s) s. + uniq s -> + x \in s -> + index x s <= index (last k s) s. Proof. elim: s k=>[|k s IH] //= k'; rewrite inE !(eq_sym k). case/andP=>K U; case: (x =P k)=>//= /eqP N X. @@ -798,8 +904,10 @@ 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. + 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. @@ -1065,7 +1173,7 @@ End FilterLastIndex. (* index and mapping *) Section IndexPmap. -Variables (A B : eqType). +Variables A B : eqType. Lemma index_pmap_inj (s : seq A) (f : A -> option B) a1 a2 b1 b2 : injective f -> @@ -1137,12 +1245,77 @@ Qed. End IndexPmap. +Section Allrel. +Variables S T : Type. + +Lemma allrel_rconsl (r : T -> S -> bool) x xs ys : + allrel r (rcons xs x) ys = allrel r xs ys && all (r x) ys. +Proof. by rewrite -cats1 allrel_catl allrel1l. Qed. + +Lemma allrel_rconsr (r : T -> S -> bool) y xs ys : + allrel r xs (rcons ys y) = allrel r xs ys && all (r^~ y) xs. +Proof. by rewrite -cats1 allrel_catr allrel1r. Qed. + +End Allrel. + + +Section AllrelEq. +Variables S T : eqType. + +Lemma allrel_in_l (r : T -> S -> bool) (xs xs' : seq T) (ys : seq S) : + xs =i xs' -> + allrel r xs ys = allrel r xs' ys. +Proof. +by move=>H; rewrite !allrel_allpairsE; apply/eq_all_r/mem_allpairs_dep. +Qed. + +Lemma allrel_in_r (r : T -> S -> bool) (xs : seq T) (ys ys' : seq S) : + ys =i ys' -> + allrel r xs ys = allrel r xs ys'. +Proof. +by move=>H; rewrite !allrel_allpairsE; apply/eq_all_r/mem_allpairs_dep. +Qed. + +Lemma allrel_sub_l (r : T -> S -> bool) (xs xs' : seq T) (ys : seq S) : + {subset xs' <= xs} -> + allrel r xs ys -> allrel r xs' ys. +Proof. +move=>H Ha; apply/allrelP=>x y Hx Hy. +by move/allrelP: Ha; apply=>//; apply: H. +Qed. + +Lemma allrel_sub_r (r : T -> S -> bool) (xs : seq T) (ys ys' : seq S) : + {subset ys' <= ys} -> + allrel r xs ys -> allrel r xs ys'. +Proof. +move=>H Ha; apply/allrelP=>x y Hx Hy. +by move/allrelP: Ha; apply=>//; apply: H. +Qed. + +Lemma allrel_trans (xs ys : seq S) z r : + transitive r -> + all (r^~ z) xs -> all (r z) ys -> allrel r xs ys. +Proof. +move=>Ht /allP Ha /allP Hp; apply/allrelP=>x y + Hy. +by move/Ha/Ht; apply; apply: Hp. +Qed. + +End AllrelEq. + Section SeqRel. -Variable (A : eqType). -Implicit Type (ltT leT : rel A). +Variable A : eqType. +Implicit Type ltT leT : rel A. (* ordering with path, seq and last *) +(* renaming *) +#[deprecated(since="fcsl-pcm 1.4.0", note="Use order_path_min instead.")] +Lemma path_all (xs : seq A) x r : + transitive r -> + path r x xs -> + all (r x) xs. +Proof. exact: order_path_min. Qed. + Lemma eq_last (s : seq A) x y : x \in s -> last y s = last x s. @@ -1286,6 +1459,20 @@ elim: s=>[|a s IH] //= P H; rewrite rcons_path P /=. by apply: H (mem_last _ _). Qed. +Lemma sorted_rconsE (leT : rel A) xs x : + transitive leT -> + sorted leT (rcons xs x) = + all (leT^~ x) xs && sorted leT xs. +Proof. +move/rev_trans=>Ht; rewrite -(revK (rcons _ _)) rev_rcons rev_sorted /=. +by rewrite path_sortedE // all_rev rev_sorted. +Qed. + +Lemma sorted1 (r : rel A) xs : + size xs == 1 -> + sorted r xs. +Proof. by case: xs=>// x; case. Qed. + Lemma sorted_subset_subseq (s1 s2 : seq A) ltT : irreflexive ltT -> transitive ltT -> diff --git a/core/seqperm.v b/core/seqperm.v index 0f8a4ad..012b36f 100644 --- a/core/seqperm.v +++ b/core/seqperm.v @@ -33,14 +33,16 @@ 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 = [::]. +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. +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] *; @@ -48,10 +50,16 @@ 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. +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. +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). @@ -65,10 +73,14 @@ Qed. Hint Resolve pperm_catC : core. -Lemma pperm_cat2lL s s1 s2 : perm s1 s2 -> perm (s ++ s1) (s ++ s2). +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). +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. @@ -77,37 +89,47 @@ by apply: pperm_cat2lL. Qed. Lemma pperm_catL s1 t1 s2 t2 : - perm s1 s2 -> perm t1 t2 -> perm (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). + 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). +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). +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). + 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. +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). + perm (s1 ++ x :: t1) (s2 ++ x :: t2) -> + perm (s1 ++ t1) (s2 ++ t2). Proof. move: s1 t1 s2 t2 x. suff H: @@ -143,24 +165,29 @@ case: s1 E2=>/=[|a s1][->]=>E2; case: s2 E4=>/=[|d s2][->]=>E4; by apply: permutation_swap; apply: IH. Qed. -Lemma pperm_cons x s1 s2 : perm (x :: s1) (x :: s2) <-> perm s1 s2. +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. +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. +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). +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). +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 : @@ -184,7 +211,8 @@ End Permutations. (* 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). + 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)]. @@ -206,3 +234,4 @@ 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/core/slice.v b/core/slice.v index 6c71b66..9f053ce 100644 --- a/core/slice.v +++ b/core/slice.v @@ -13,6 +13,7 @@ limitations under the License. From Coq Require Import ssreflect ssrbool ssrfun. From mathcomp Require Import ssrnat eqtype seq path interval order. +From mathcomp Require Import fintype finfun tuple. From pcm Require Import options prelude seqext. Open Scope order_scope. @@ -714,3 +715,13 @@ by apply/drop_sorted/take_sorted. Qed. End LemmasEq. + +(* slicing and rcons *) +Lemma codom_ux_rcons A n (f : {ffun 'I_n -> A}) (i : 'I_n) : + &:(fgraph f) `]-oo, i : nat] = + rcons (&:(fgraph f) `]-oo, i : nat[) (f i). +Proof. by rewrite slice_xR // slice_uu onth_codom. Qed. + + + + diff --git a/htt/domain.v b/htt/domain.v index e287984..75d80de 100644 --- a/htt/domain.v +++ b/htt/domain.v @@ -1,104 +1,85 @@ -From mathcomp Require Import ssreflect ssrbool ssrfun ssrnat eqtype seq. +(* +Copyright 2019 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 HB Require Import structures. +From Coq Require Import ssreflect ssrbool ssrfun. +From mathcomp Require Import ssrnat eqtype. From pcm Require Import options axioms pred prelude. +(**************************************************) +(* This file develops some basic domain theory of *) +(* posets, complete lattices, complete partial *) +(* orders (CPOs). It developes the standard fixed *) +(* point theorems: Knaster-Tarski's for complete *) +(* lattices, and Kleene's for CPOs. *) +(**************************************************) + (**********) (* Posets *) (**********) -(* We put bottom in posets right away, instead of adding it later both in *) -(* lattices and in cpos. Since we never consider bottom-less posets, this *) -(* saves some tedium and name space. *) - -Module Poset. - -Section RawMixin. - -Record mixin_of (T : Type) := Mixin { - mx_leq : T -> T -> Prop; - mx_bot : T; - _ : forall x, mx_leq mx_bot x; - _ : forall x, mx_leq x x; - _ : forall x y, mx_leq x y -> mx_leq y x -> x = y; - _ : forall x y z, mx_leq x y -> mx_leq y z -> mx_leq x z}. - -End RawMixin. - -Section ClassDef. +Definition poset_axiom T (leq : T -> T -> Prop) := + [/\ forall x, leq x x, + forall x y, leq x y -> leq y x -> x = y & + forall x y z, leq x y -> leq y z -> leq x z]. -Record class_of T := Class {mixin : mixin_of T}. +HB.mixin Record isPoset T := { + poset_leq : T -> T -> Prop; + poset_subproof : poset_axiom poset_leq}. -Structure type : Type := Pack {sort : Type; _ : class_of sort}. -Local Coercion sort : type >-> Sortclass. +#[short(type="poset")] +HB.structure Definition Poset := {T of isPoset T}. -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 a poset type out of the mixin *) -(* equalize m0 and m by means of a phantom *) -Definition pack (m0 : mixin_of T) := - fun m & phant_id m0 m => Pack (@Class T m). - -Definition leq := mx_leq (mixin class). -Definition bot := mx_bot (mixin class). - -End ClassDef. - -Module Exports. -Coercion sort : type >-> Sortclass. -Notation poset := Poset.type. -Notation PosetMixin := Poset.Mixin. -Notation Poset T m := (@pack T _ m id). +Notation "x <== y" := (poset_leq x y) (at level 70). -Notation "[ 'poset' 'of' T 'for' cT ]" := (@clone T cT _ id) - (at level 0, format "[ 'poset' 'of' T 'for' cT ]") : form_scope. -Notation "[ 'poset' 'of' T ]" := (@clone T _ _ id) - (at level 0, format "[ 'poset' 'of' T ]") : form_scope. - -Notation "x <== y" := (Poset.leq x y) (at level 70). -Notation bot := Poset.bot. - -Arguments Poset.bot {cT}. -Prenex Implicits Poset.bot. - -(* re-state lattice properties using the exported notation *) -Section Laws. +Section Repack. Variable T : poset. - -Lemma botP (x : T) : bot <== x. -Proof. by case: T x=>s [[leq b B]]. Qed. - Lemma poset_refl (x : T) : x <== x. -Proof. by case: T x=>S [[leq b B R]]. Qed. - +Proof. by case: (@poset_subproof T). Qed. Lemma poset_asym (x y : T) : x <== y -> y <== x -> x = y. -Proof. by case: T x y=>S [[l b B R A Tr]] *; apply: (A). Qed. - +Proof. by case: (@poset_subproof T)=>_ H _; apply: H. Qed. Lemma poset_trans (y x z : T) : x <== y -> y <== z -> x <== z. -Proof. by case: T y x z=>S [[l b B R A Tr]] ? x y z; apply: (Tr). Qed. - -End Laws. +Proof. by case: (@poset_subproof T)=>_ _; apply. Qed. +End Repack. -#[export] -Hint Resolve botP poset_refl : core. +#[export] Hint Resolve poset_refl : core. -Add Parametric Relation (T : poset) : T (@Poset.leq T) +Add Parametric Relation (T : poset) : T (@poset_leq T) reflexivity proved by (@poset_refl _) transitivity proved by (fun x y => @poset_trans _ y x) as poset_rel. -End Exports. +(**********************) +(* monotone functions *) +(**********************) + +Definition monofun_axiom (T1 T2 : poset) (f : T1 -> T2) := + forall x y, x <== y -> f x <== f y. + +HB.mixin Record isMonoFun (T1 T2 : poset) (f : T1 -> T2) := { + monofun_subproof : monofun_axiom f}. -End Poset. +#[short(type="mono_fun")] +HB.structure Definition MonoFun (T1 T2 : poset) := + {f of isMonoFun T1 T2 f}. -Export Poset.Exports. +Lemma monofunP (T1 T2 : poset) (f : mono_fun T1 T2) x y : + x <== y -> f x <== f y. +Proof. exact: monofun_subproof. Qed. (**************************) (* some basic definitions *) (**************************) -Definition monotone (T1 T2 : poset) (f : T1 -> T2) := - forall x y, x <== y -> f x <== f y. - Section IdealDef. Variable T : poset. @@ -119,100 +100,55 @@ End IdealDef. (***********************) Section SubPoset. -Variables (T : poset) (s : Pred T) (C : bot \In s). - +Variables (T : poset) (s : Pred T). Local Notation tp := {x : T | x \In s}. -Definition sub_bot : tp := exist _ bot C. Definition sub_leq (p1 p2 : tp) := sval p1 <== sval p2. -Lemma sub_botP x : sub_leq sub_bot x. -Proof. by apply: botP. Qed. - -Lemma sub_refl x : sub_leq x x. -Proof. by rewrite /sub_leq. Qed. - -Lemma sub_asym x y : sub_leq x y -> sub_leq y x -> x = y. -Proof. -move: x y=>[x Hx][y Hy]; rewrite /sub_leq /= => H1 H2. -move: (poset_asym H1 H2) Hy=> <- Hy; congr exist. -by apply: pf_irr. -Qed. - -Lemma sub_trans x y z : sub_leq x y -> sub_leq y z -> sub_leq x z. -Proof. -move: x y z=>[x Hx][y Hy][z Hz]; rewrite /sub_leq /=. +Lemma sub_is_poset : poset_axiom sub_leq. +Proof. +split=>[x|[x Hx][y Hy]|[x Hx][y Hy][z Hz]]; rewrite /sub_leq //=. +- move=>H1 H2; rewrite -(poset_asym H1 H2) in Hy *. + by rewrite (pf_irr Hx). by apply: poset_trans. Qed. - -(* no need to put canonical here, because the system won't be *) -(* able to get the proof C from the {x : T | x \In s} syntax *) -Definition subPosetMixin := PosetMixin sub_botP sub_refl sub_asym sub_trans. -Definition subPoset := Eval hnf in Poset tp subPosetMixin. - +HB.instance Definition _ := isPoset.Build tp sub_is_poset. End SubPoset. (* pairing of posets *) -Section PairPoset. -Variable (A B : poset). +Section ProdPoset. +Variables A B : poset. Local Notation tp := (A * B)%type. -Definition pair_bot : tp := (bot, bot). -Definition pair_leq (p1 p2 : tp) := p1.1 <== p2.1 /\ p1.2 <== p2.2. - -Lemma pair_botP x : pair_leq pair_bot x. -Proof. by split; apply: botP. Qed. - -Lemma pair_refl x : pair_leq x x. -Proof. by []. Qed. - -Lemma pair_asym x y : pair_leq x y -> pair_leq y x -> x = y. -Proof. -move: x y=>[x1 x2][y1 y2][/= H1 H2][/= H3 H4]. -by congr (_, _); apply: poset_asym. -Qed. +Definition poset_prod_leq := + [fun p1 p2 : tp => p1.1 <== p2.1 /\ p1.2 <== p2.2]. -Lemma pair_trans x y z : pair_leq x y -> pair_leq y z -> pair_leq x z. -Proof. -move: x y z=>[x1 x2][y1 y2][z1 z2][/= H1 H2][/= H3 H4]; split=>/=. -- by apply: poset_trans H3. -by apply: poset_trans H4. +Lemma prod_is_poset : poset_axiom poset_prod_leq. +Proof. +split=>[//|[x1 x2][y1 y2]|[x1 x2][y1 y2][z1 z2]] /=. +- by case=>H1 H2 [H3 H4]; congr (_, _); apply: poset_asym. +case=>H1 H2 [H3 H4]; split; +by [apply: poset_trans H3|apply: poset_trans H4]. Qed. +HB.instance Definition _ := isPoset.Build tp prod_is_poset. +End ProdPoset. -Definition pairPosetMixin := - PosetMixin pair_botP pair_refl pair_asym pair_trans. -Canonical pairPoset := Eval hnf in Poset tp pairPosetMixin. - -End PairPoset. - -(* functions into a poset form a poset *) +(* functions into poset form a poset *) Section FunPoset. Variable (A : Type) (B : poset). Local Notation tp := (A -> B). -Definition fun_bot : tp := fun x => bot. -Definition fun_leq (p1 p2 : tp) := forall x, p1 x <== p2 x. - -Lemma fun_botP x : fun_leq fun_bot x. -Proof. by move=>y; apply: botP. Qed. - -Lemma fun_refl x : fun_leq x x. -Proof. by []. Qed. +Definition poset_fun_leq := [fun p1 p2 : tp => forall x, p1 x <== p2 x]. -Lemma fun_asym x y : fun_leq x y -> fun_leq y x -> x = y. +Lemma fun_is_poset : poset_axiom poset_fun_leq. Proof. -move=>H1 H2; apply: fext=>z; -by apply: poset_asym; [apply: H1 | apply: H2]. +split=>[x|x y H1 H2|x y z H1 H2 t] //=. +- by apply: fext=>z; apply: poset_asym; [apply: H1|apply: H2]. +by apply: poset_trans (H2 t). Qed. - -Lemma fun_trans x y z : fun_leq x y -> fun_leq y z -> fun_leq x z. -Proof. by move=>H1 H2 t; apply: poset_trans (H2 t). Qed. - -Definition funPosetMixin := PosetMixin fun_botP fun_refl fun_asym fun_trans. -Canonical funPoset := Eval hnf in Poset tp funPosetMixin. - +HB.instance Definition _ := isPoset.Build tp fun_is_poset. End FunPoset. (* dependent functions into a poset form a poset *) @@ -221,286 +157,151 @@ Section DFunPoset. Variables (A : Type) (B : A -> poset). Local Notation tp := (forall x, B x). -Definition dfun_bot : tp := fun x => bot. -Definition dfun_leq (p1 p2 : tp) := forall x, p1 x <== p2 x. - -Lemma dfun_botP x : dfun_leq dfun_bot x. -Proof. by move=>y; apply: botP. Qed. +Definition poset_dfun_leq := [fun p1 p2 : tp => forall x, p1 x <== p2 x]. -Lemma dfun_refl x : dfun_leq x x. -Proof. by []. Qed. - -Lemma dfun_asym x y : dfun_leq x y -> dfun_leq y x -> x = y. +Lemma dfun_is_poset : poset_axiom poset_dfun_leq. Proof. -move=>H1 H2; apply: fext=>z; -by apply: poset_asym; [apply: H1 | apply: H2]. +split=>[x|x y H1 H2|x y z H1 H2 t] //. +- by apply: fext=>z; apply: poset_asym; [apply: H1|apply: H2]. +by apply: poset_trans (H2 t). Qed. -Lemma dfun_trans x y z : dfun_leq x y -> dfun_leq y z -> dfun_leq x z. -Proof. by move=>H1 H2 t; apply: poset_trans (H2 t). Qed. - -(* no point in declaring this canonical, since it's keyed on forall *) +(* no instance declaration, since it's keyed on forall *) (* and forall is not a symbol *) -Definition dfunPosetMixin := - PosetMixin dfun_botP dfun_refl dfun_asym dfun_trans. -Definition dfunPoset := Eval hnf in Poset tp dfunPosetMixin. - +Definition dfun_poset_mix := isPoset.Build tp dfun_is_poset. +Definition dfun_poset : poset := Poset.pack_ dfun_poset_mix. End DFunPoset. -(* ideal of a poset is a poset *) +(* ideal of poset is poset *) Section IdealPoset. Variable (T : poset) (P : T). -Definition ideal_bot := Ideal (botP P). -Definition ideal_leq (p1 p2 : ideal P) := id_val p1 <== id_val p2. - -Lemma ideal_botP x : ideal_leq ideal_bot x. -Proof. by apply: botP. Qed. - -Lemma ideal_refl x : ideal_leq x x. -Proof. by case: x=>x H; rewrite /ideal_leq. Qed. +Definition poset_ideal_leq := + [fun p1 p2 : ideal P => id_val p1 <== id_val p2]. -Lemma ideal_asym x y : ideal_leq x y -> ideal_leq y x -> x = y. +Lemma ideal_is_poset : poset_axiom poset_ideal_leq. Proof. -move: x y=>[x1 H1][x2 H2]; rewrite /ideal_leq /= => H3 H4; move: H1 H2. -rewrite (poset_asym H3 H4)=>H1 H2. -congr Ideal; apply: pf_irr. +split=>[[x]|[x1 H1][x2 H2]|[x1 H1][x2 H2][x3 H3]] //=. +- move=>H3 H4; rewrite (poset_asym H3 H4) in H1 H2 *. + by rewrite (pf_irr H1). +by apply: poset_trans. Qed. - -Lemma ideal_trans x y z : ideal_leq x y -> ideal_leq y z -> ideal_leq x z. -Proof. by apply: poset_trans. Qed. - -Definition idealPosetMixin := - PosetMixin ideal_botP ideal_refl ideal_asym ideal_trans. -Canonical idealPoset := Eval hnf in Poset (ideal P) idealPosetMixin. - +HB.instance Definition _ := isPoset.Build (ideal P) ideal_is_poset. End IdealPoset. -(* Prop is a poset (Sierpinski poset) *) - -Section PropPoset. - -Definition prop_bot := False. -Definition prop_leq (p1 p2 : Prop) := p1 -> p2. - -Lemma prop_botP x : prop_leq prop_bot x. -Proof. by []. Qed. - -Lemma prop_refl x : prop_leq x x. -Proof. by []. Qed. - -Lemma prop_asym x y : prop_leq x y -> prop_leq y x -> x = y. -Proof. by move=>H1 H2; apply: pext. Qed. - -Lemma prop_trans x y z : prop_leq x y -> prop_leq y z -> prop_leq x z. -Proof. by move=>H1 H2; move/H1. Qed. - -Definition propPosetMixin := - PosetMixin prop_botP prop_refl prop_asym prop_trans. -Canonical propPoset := Eval hnf in Poset Prop propPosetMixin. - -Lemma prop_topP x : x <== True. -Proof. by []. Qed. - -End PropPoset. - -(* Pred is a poset *) - -(* Can be inherited from posets of -> and Prop, but we declare a *) -(* dedicated structure to keep the infix notation of Pred. Otherwise, *) -(* poset inference turns Pred A into A -> Prop. *) - -Section PredPoset. -Variable A : Type. - -Definition predPosetMixin : Poset.mixin_of (Pred A) := - funPosetMixin A propPoset. -Canonical predPoset := Eval hnf in Poset (Pred A) predPosetMixin. - -Lemma pred_topP (x : Pred A) : x <== PredT. -Proof. by []. Qed. - -End PredPoset. - -(* nat is a poset *) -Section NatPoset. - -Lemma nat_botP x : 0 <= x. Proof. by []. Qed. -Lemma nat_refl x : x <= x. Proof. by []. Qed. - -Lemma nat_asym x y : x <= y -> y <= x -> x = y. -Proof. by move=>H1 H2; apply: anti_leq; rewrite H1 H2. Qed. - -Lemma nat_trans x y z : x <= y -> y <= z -> x <= z. -Proof. by apply: leq_trans. Qed. - -Definition natPosetMixin := PosetMixin nat_botP nat_refl nat_asym nat_trans. -Canonical natPoset := Eval hnf in Poset nat natPosetMixin. - -End NatPoset. +(* Prop is poset *) +Definition poset_prop_leq := [fun p1 p2 : Prop => p1 -> p2]. +Lemma prop_is_poset : poset_axiom poset_prop_leq. +Proof. by split=>[x|x y H1 H2|x y z H1 H2 /H1] //; apply: pext. Qed. +HB.instance Definition _ := isPoset.Build Prop prop_is_poset. (*********************) (* Complete lattices *) (*********************) -Module Lattice. - -Section RawMixin. - -Variable T : poset. - -Record mixin_of := Mixin { - mx_sup : Pred T -> T; - _ : forall (s : Pred T) x, x \In s -> x <== mx_sup s; - _ : forall (s : Pred T) x, - (forall y, y \In s -> y <== x) -> mx_sup s <== x}. - -End RawMixin. - -Section ClassDef. - -Record class_of (T : Type) := Class { - base : Poset.class_of T; - mixin : mixin_of (Poset.Pack base)}. - -Local Coercion base : class_of >-> Poset.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. +Definition lattice_axiom (T : poset) (sup : Pred T -> T) := + [/\ (* sup is upper bound *) + forall (s : Pred T) x, x \In s -> x <== sup s & + (* sup is least upper bound *) + forall (s : Pred T) x, + (forall y, y \In s -> y <== x) -> sup s <== x]. -(* produce a lattice type out of the mixin *) -(* equalize m0 and m by means of a phantom *) -Definition pack b0 (m0 : mixin_of (Poset.Pack b0)) := - fun m & phant_id m0 m => Pack (@Class T b0 m). +HB.mixin Record isLattice T of Poset T := { + sup : Pred T -> T; + lattice_subproof : lattice_axiom sup}. -Definition sup (s : Pred cT) : cT := mx_sup (mixin class) s. +#[short(type="lattice")] +HB.structure Definition Lattice := {T of Poset T & isLattice T}. -Definition poset := Poset.Pack class. - -End ClassDef. - -Module Exports. -Coercion base : class_of >-> Poset.class_of. -Coercion sort : type >-> Sortclass. -Coercion poset : type >-> Poset.type. -Canonical Structure poset. - -Notation lattice := Lattice.type. -Notation LatticeMixin := Lattice.Mixin. -Notation Lattice T m := (@pack T _ _ m id). - -Notation "[ 'lattice' 'of' T 'for' cT ]" := (@clone T cT _ id) - (at level 0, format "[ 'lattice' 'of' T 'for' cT ]") : form_scope. -Notation "[ 'lattice' 'of' T ]" := (@clone T _ _ id) - (at level 0, format "[ 'lattice' 'of' T ]") : form_scope. - -Arguments Lattice.sup [cT] s. -Prenex Implicits Lattice.sup. -Notation sup := Lattice.sup. - -(* re-state lattice properties using the exported notation *) -Section Laws. +Section Repack. Variable T : lattice. - Lemma supP (s : Pred T) x : x \In s -> x <== sup s. -Proof. by case: T s x=>S [[p]][/= s H1 *]; apply: H1. Qed. - +Proof. by case: (@lattice_subproof T)=>H _; apply: H. Qed. Lemma supM (s : Pred T) x : (forall y, y \In s -> y <== x) -> sup s <== x. -Proof. by case: T s x=>S [[p]][/= s H1 H2 *]; apply: (H2). Qed. - -End Laws. - -End Exports. - -End Lattice. - -Export Lattice.Exports. +Proof. by case: (@lattice_subproof T)=>_; apply. Qed. +End Repack. -(* we have greatest lower bounds too *) +(* infima follow *) Section Infimum. Variable T : lattice. -Definition inf (s : Pred T) := - sup [Pred x : T | forall y, y \In s -> x <== y]. +Definition inf (s : Pred T) := sup [Pred x : T | forall y, y \In s -> x <== y]. -Lemma infP s : forall x, x \In s -> inf s <== x. -Proof. by move=>x H; apply: supM=>y; apply. Qed. +Lemma infP s x : x \In s -> inf s <== x. +Proof. by move=>H; apply: supM=>y; apply. Qed. -Lemma infM s : forall x, (forall y, y \In s -> x <== y) -> x <== inf s. +Lemma infM s x : (forall y, y \In s -> x <== y) -> x <== inf s. Proof. by apply: supP. Qed. End Infimum. -(* we can compute least and greatest fixed points *) - -Section Lat. +(* least and greatest fixed points by Knaster-Tarski construction *) +Section KnasterTarskiDefs. Variable T : lattice. +Definition lbot : T := sup Pred0. + Definition tarski_lfp (f : T -> T) := inf [Pred x : T | f x <== x]. Definition tarski_gfp (f : T -> T) := sup [Pred x : T | x <== f x]. Definition sup_closed (T : lattice) := - [Pred s : Pred T | - bot \In s /\ forall d, d <=p s -> sup d \In s]. + [Pred s : Pred T | forall d, d <=p s -> sup d \In s]. Definition sup_closure (T : lattice) (s : Pred T) := [Pred p : T | forall t : Pred T, s <=p t -> t \In sup_closed T -> p \In t]. -End Lat. +End KnasterTarskiDefs. +Arguments lbot {T}. Arguments sup_closed {T}. -Arguments sup_closure [T]. +Arguments sup_closure [T] s. Prenex Implicits sup_closed sup_closure. Section BasicProperties. Variable T : lattice. Lemma sup_mono (s1 s2 : Pred T) : s1 <=p s2 -> sup s1 <== sup s2. -Proof. by move=>H; apply: supM=>y; move/H; apply: supP. Qed. +Proof. by move=>H; apply: supM=>y /H/supP. Qed. Lemma supE (s1 s2 : Pred T) : s1 =p s2 -> sup s1 = sup s2. -Proof. by move=>H; apply: poset_asym; apply: supM=>y; move/H; apply: supP. Qed. +Proof. by move=>H; apply: poset_asym; apply: supM=>y /H/supP. Qed. -(* Knaster-Tarski *) -Lemma tarski_lfp_fixed (f : T -> T) : - monotone f -> f (tarski_lfp f) = tarski_lfp f. +(* Knaster-Tarski theorem *) +Lemma tarski_lfp_fixed (f : mono_fun T T) : + f (tarski_lfp f) = tarski_lfp f. Proof. -move=>M; suff L: f (tarski_lfp f) <== tarski_lfp f. -- by apply: poset_asym=>//; apply: infP; apply: M L. -by apply: infM=>x L; apply: poset_trans (L); apply: M; apply: infP. +suff L : f (tarski_lfp f) <== tarski_lfp f. +- by apply: poset_asym=>//; apply/infP/monofunP. +by apply: infM=>x L; apply: poset_trans (L); apply/monofunP/infP. Qed. -Lemma tarski_lfp_least f : forall x : T, f x = x -> tarski_lfp f <== x. -Proof. by move=>x H; apply: infP; rewrite InE /= H. Qed. +Lemma tarski_lfp_least f (x : T) : f x <== x -> tarski_lfp f <== x. +Proof. by move=>H; apply: infP. Qed. -Lemma tarski_gfp_fixed (f : T -> T) : - monotone f -> f (tarski_gfp f) = tarski_gfp f. +Lemma tarski_gfp_fixed (f : mono_fun T T) : + f (tarski_gfp f) = tarski_gfp f. Proof. -move=>M; suff L: tarski_gfp f <== f (tarski_gfp f). -- by apply: poset_asym=>//; apply: supP; apply: M L. -by apply: supM=>x L; apply: poset_trans (L) _; apply: M; apply: supP. +suff L: tarski_gfp f <== f (tarski_gfp f). +- by apply: poset_asym=>//; apply/supP/monofunP. +by apply: supM=>x L; apply: poset_trans (L) _; apply/monofunP/supP. Qed. -Lemma tarski_gfp_greatest f : forall x : T, f x = x -> x <== tarski_gfp f. -Proof. by move=>x H; apply: supP; rewrite InE /= H. Qed. +Lemma tarski_gfp_greatest f (x : T) : x <== f x -> x <== tarski_gfp f. +Proof. by move=>H; apply: supP. Qed. -Lemma tarski_lfp_mono (f1 f2 : T -> T) : - monotone f2 -> f1 <== f2 -> tarski_lfp f1 <== tarski_lfp f2. +Lemma tarski_lfp_mono (f1 : T -> T) (f2 : mono_fun T T) : + f1 <== f2 -> tarski_lfp f1 <== tarski_lfp f2. Proof. -move=>M H; apply/infP/(poset_trans (H (tarski_lfp f2))). +move=>H; apply: infP; apply: poset_trans (H (tarski_lfp f2)) _. by rewrite tarski_lfp_fixed. Qed. -Lemma tarski_gfp_mono (f1 f2 : T -> T) : - monotone f1 -> f1 <== f2 -> tarski_gfp f1 <== tarski_gfp f2. +Lemma tarski_gfp_mono (f1 : mono_fun T T) (f2 : T -> T) : + (f1 : T -> T) <== f2 -> tarski_gfp f1 <== tarski_gfp f2. Proof. -move=>M H; apply/supP/poset_trans/(H (tarski_gfp f1)). +move=>H; apply/supP/poset_trans/(H (tarski_gfp f1)). by rewrite tarski_gfp_fixed. Qed. @@ -516,9 +317,8 @@ Proof. by move=>t H1 H2 p; move/(_ _ H1 H2). Qed. (* closure is closed *) Lemma sup_closP (s : Pred T) : sup_closed (sup_closure s). Proof. -split; first by move=>t _ []. -move=>d H1 t H3 H4; move: (sup_clos_min H3 H4)=>{H3} -H3. -by case: H4=>_; apply=>// x; move/H1; move/H3. +move=>d H1 t H3 H4; move: (sup_clos_min H3 H4)=>{}H3. +by apply: H4=>x; move/H1; move/H3. Qed. Lemma sup_clos_idemp (s : Pred T) : sup_closed s -> sup_closure s =p s. @@ -527,591 +327,408 @@ Proof. by move=>p; split; [apply: sup_clos_min | apply: sup_clos_sub]. Qed. Lemma sup_clos_mono (s1 s2 : Pred T) : s1 <=p s2 -> sup_closure s1 <=p sup_closure s2. Proof. -move=>H1; apply: sup_clos_min (sup_closP s2)=>p H2. -by apply: sup_clos_sub; apply: H1. +move=>H1; apply: sup_clos_min (@sup_closP s2). +by move=>t /H1; apply: sup_clos_sub. Qed. End BasicProperties. (* lattice constructions *) +(* subset lattice *) Section SubLattice. Variables (T : lattice) (s : Pred T) (C : sup_closed s). -Local Notation tp := (subPoset (proj1 C)). +Local Notation tp := {x : T | x \In s}. -Definition sub_sup' (u : Pred tp) : T := +Definition lattice_sub_sup' (u : Pred tp) : T := sup [Pred x : T | exists y, y \In u /\ x = sval y]. +Lemma lattice_sub_supX (u : Pred tp) : lattice_sub_sup' u \In s. +Proof. by apply: C=>x [][y H][_] ->. Qed. +Definition lattice_sub_sup (u : Pred tp) : tp := + exist _ (lattice_sub_sup' u) (lattice_sub_supX u). -Lemma sub_supX (u : Pred tp) : sub_sup' u \In s. -Proof. by case: C u=>P /= H u; apply: H=>t [[y]] H1 [_] ->. Qed. - -Definition sub_sup (u : Pred tp) : tp := - exist _ (sub_sup' u) (sub_supX u). - -Lemma sub_supP (u : Pred tp) (x : tp) : x \In u -> x <== sub_sup u. -Proof. by move=>H; apply: supP; exists x. Qed. - -Lemma sub_supM (u : Pred tp) (x : tp) : - (forall y, y \In u -> y <== x) -> sub_sup u <== x. -Proof. by move=>H; apply: supM=>y [z][H1] ->; apply: H H1. Qed. - -Definition subLatticeMixin := LatticeMixin sub_supP sub_supM. -Definition subLattice := Eval hnf in Lattice {x : T | x \In s} subLatticeMixin. - +Lemma sub_is_lattice : lattice_axiom lattice_sub_sup. +Proof. +split=>u x H; first by apply: supP; exists x. +by apply: supM=>y [z][H1] ->; apply: H H1. +Qed. +HB.instance Definition _ := isLattice.Build tp sub_is_lattice. End SubLattice. -(* pairing *) - -Section PairLattice. +(* product *) +Section ProdLattice. Variables (A B : lattice). Local Notation tp := (A * B)%type. -Definition pair_sup (s : Pred tp) : tp := +Definition lattice_prod_sup (s : Pred tp) : tp := (sup [Pred p | exists f, p = f.1 /\ f \In s], sup [Pred p | exists f, p = f.2 /\ f \In s]). -Lemma pair_supP (s : Pred tp) (p : tp) : p \In s -> p <== pair_sup s. -Proof. by move=>H; split; apply: supP; exists p. Qed. - -Lemma pair_supM (s : Pred tp) (p : tp) : - (forall q, q \In s -> q <== p) -> pair_sup s <== p. -Proof. by move=>H; split; apply: supM=>y [f][->]; case/H. Qed. - -Definition pairLatticeMixin := LatticeMixin pair_supP pair_supM. -Canonical pairLattice := Eval hnf in Lattice tp pairLatticeMixin. - -End PairLattice. - -(* functions into a latice form a lattice *) +Lemma prod_is_lattice : lattice_axiom lattice_prod_sup. +Proof. +split=>s p H; first by split; apply: supP; exists p. +by split; apply: supM=>y [f][->]; case/H. +Qed. +HB.instance Definition _ := isLattice.Build tp prod_is_lattice. +End ProdLattice. +(* functions into latice form lattice *) Section FunLattice. Variables (A : Type) (B : lattice). Local Notation tp := (A -> B). -Definition fun_sup (s : Pred tp) : tp := +Definition lattice_fun_sup (s : Pred tp) : tp := fun x => sup [Pred p | exists f, f \In s /\ p = f x]. -Lemma fun_supP (s : Pred tp) (p : tp) : p \In s -> p <== fun_sup s. -Proof. by move=>H x; apply: supP; exists p. Qed. - -Lemma fun_supM (s : Pred tp) (p : tp) : - (forall q, q \In s -> q <== p) -> fun_sup s <== p. -Proof. by move=>H t; apply: supM=>x [f][H1] ->; apply: H. Qed. - -Definition funLatticeMixin := LatticeMixin fun_supP fun_supM. -Canonical funLattice := Eval hnf in Lattice tp funLatticeMixin. - +Lemma fun_is_lattice : lattice_axiom lattice_fun_sup. +Proof. +split=>s p H x; first by apply: supP; exists p. +by apply: supM=>y [f][H1] ->; apply: H. +Qed. +HB.instance Definition _ := isLattice.Build tp fun_is_lattice. End FunLattice. (* dependent functions into a lattice form a lattice *) - Section DFunLattice. Variables (A : Type) (B : A -> lattice). -Local Notation tp := (dfunPoset B). +Local Notation tp := (dfun_poset B). -Definition dfun_sup (s : Pred tp) : tp := +Definition lattice_dfun_sup (s : Pred tp) : tp := fun x => sup [Pred p | exists f, f \In s /\ p = f x]. -Lemma dfun_supP (s : Pred tp) (p : tp) : - p \In s -> p <== dfun_sup s. -Proof. by move=>H x; apply: supP; exists p. Qed. - -Lemma dfun_supM (s : Pred tp) (p : tp) : - (forall q, q \In s -> q <== p) -> dfun_sup s <== p. -Proof. by move=>H t; apply: supM=>x [f][H1] ->; apply: H. Qed. - -Definition dfunLatticeMixin := LatticeMixin dfun_supP dfun_supM. -Definition dfunLattice := Eval hnf in Lattice (forall x, B x) dfunLatticeMixin. +Lemma dfun_is_lattice : lattice_axiom lattice_dfun_sup. +Proof. +split=>s p H x; first by apply: supP; exists p. +by apply: supM=>y [f][H1] ->; apply: H. +Qed. +Definition dfun_lattice_mix := isLattice.Build tp dfun_is_lattice. +Definition dfun_lattice : lattice := Lattice.pack_ dfun_lattice_mix. End DFunLattice. (* applied sup equals the sup of applications *) - Lemma sup_appE A (B : lattice) (s : Pred (A -> B)) (x : A) : sup s x = sup [Pred y : B | exists f, f \In s /\ y = f x]. Proof. by []. Qed. -Lemma sup_dappE A (B : A -> lattice) (s : Pred (dfunLattice B)) (x : A) : +Lemma sup_dappE A (B : A -> lattice) (s : Pred (dfun_lattice B)) (x : A) : sup s x = sup [Pred y : B x | exists f, f \In s /\ y = f x]. Proof. by []. Qed. (* ideal of a lattice forms a lattice *) - Section IdealLattice. Variables (T : lattice) (P : T). -Definition ideal_sup' (s : Pred (ideal P)) := +Definition lattice_ideal_sup' (s : Pred (ideal P)) := sup [Pred x | exists p, p \In s /\ id_val p = x]. - -Lemma ideal_supP' (s : Pred (ideal P)) : ideal_sup' s <== P. +Lemma lattice_ideal_supP' (s : Pred (ideal P)) : lattice_ideal_sup' s <== P. Proof. by apply: supM=>y [[x]] H /= [_] <-. Qed. +Definition lattice_ideal_sup (s : Pred (ideal P)) := + Ideal (lattice_ideal_supP' s). -Definition ideal_sup (s : Pred (ideal P)) := Ideal (ideal_supP' s). - -Lemma ideal_supP (s : Pred (ideal P)) p : - p \In s -> p <== ideal_sup s. -Proof. by move=>H; apply: supP; exists p. Qed. - -Lemma ideal_supM (s : Pred (ideal P)) p : - (forall q, q \In s -> q <== p) -> ideal_sup s <== p. -Proof. by move=>H; apply: supM=>y [q][H1] <-; apply: H. Qed. - -Definition idealLatticeMixin := LatticeMixin ideal_supP ideal_supM. -Canonical idealLattice := Eval hnf in Lattice (ideal P) idealLatticeMixin. - +Lemma ideal_is_lattice : lattice_axiom lattice_ideal_sup. +Proof. +split=>s p H; first by apply: supP; exists p. +by apply: supM=>y [q][H1] <-; apply: H. +Qed. +HB.instance Definition _ := isLattice.Build (ideal P) ideal_is_lattice. End IdealLattice. (* Prop is a lattice *) +Definition lattice_prop_sup (s : Pred Prop) : Prop := + exists p, p \In s /\ p. +Lemma prop_is_lattice : lattice_axiom lattice_prop_sup. +Proof. by split=>s p; [exists p|move=>H [r][]; move/H]. Qed. +HB.instance Definition _ := isLattice.Build Prop prop_is_lattice. -Section PropLattice. +(****************) +(* Monotonicity *) +(****************) -Definition prop_sup (s : Pred Prop) : Prop := exists p, p \In s /\ p. +Section MonoCanonicals. +Variables T T1 T2 T3 S1 S2 : poset. -Lemma prop_supP (s : Pred Prop) p : p \In s -> p <== prop_sup s. -Proof. by exists p. Qed. +Lemma idfun_is_mono : monofun_axiom (@idfun T). Proof. by []. Qed. +HB.instance Definition _ := isMonoFun.Build T T idfun idfun_is_mono. -Lemma prop_supM (s : Pred Prop) p : - (forall q, q \In s -> q <== p) -> prop_sup s <== p. -Proof. by move=>H [r][]; move/H. Qed. +Definition const_fun (y : T2) (_ : T1) := y. +Lemma const_is_mono (y : T2) : monofun_axiom (const_fun y). Proof. by []. Qed. +HB.instance Definition _ y := + isMonoFun.Build T1 T2 (const_fun y) (const_is_mono y). -Definition propLatticeMixin := LatticeMixin prop_supP prop_supM. -Canonical propLattice := Eval hnf in Lattice Prop propLatticeMixin. +Lemma fst_is_mono : monofun_axiom (@fst T1 T2). +Proof. by case=>x1 x2 [y1 y2][]. Qed. +HB.instance Definition _ := isMonoFun.Build (T1*T2)%type T1 fst fst_is_mono. -End PropLattice. +Lemma snd_is_mono : monofun_axiom (@snd T1 T2). +Proof. by case=>x1 x2 [y1 y2][]. Qed. +HB.instance Definition _ := isMonoFun.Build (T1*T2)%type T2 snd snd_is_mono. -(* Pred is a lattice *) +Definition diag (x : T) := (x, x). +Lemma diag_is_mono : monofun_axiom diag. Proof. by []. Qed. +HB.instance Definition _ := isMonoFun.Build T (T*T)%type diag diag_is_mono. -Section PredLattice. -Variable A : Type. +Lemma sval_is_mono (P : Pred T) : monofun_axiom (@sval T P). +Proof. by []. Qed. +HB.instance Definition _ P := + isMonoFun.Build {x : T | P x} T sval (@sval_is_mono P). -Definition predLatticeMixin := funLatticeMixin A propLattice. -Canonical predLattice := Eval hnf in Lattice (Pred A) predLatticeMixin. +Definition app A B (x : A) := fun f : A -> B => f x. +Lemma app_is_mono A (x : A) : monofun_axiom (@app A T x). +Proof. by move=>f1 f2; apply. Qed. +HB.instance Definition _ A x := + isMonoFun.Build (A -> T) T (@app A T x) (app_is_mono x). -End PredLattice. +Definition dapp A (B : A -> poset) (x : A) := fun f : dfun_poset B => f x. +Lemma dapp_is_mono A (B : A -> poset) (x : A) : monofun_axiom (@dapp A B x). +Proof. by move=>f1 f2; apply. Qed. +HB.instance Definition _ A B x := + isMonoFun.Build (dfun_poset B) (B x) (@dapp A B x) (dapp_is_mono x). + +Section Comp. +Variables (f1 : mono_fun T2 T1) (f2 : mono_fun T3 T2). +Lemma comp_is_mono : monofun_axiom (f1 \o f2). +Proof. by move=>x y H; apply/monofunP/monofunP. Qed. +HB.instance Definition _ := isMonoFun.Build T3 T1 (f1 \o f2) comp_is_mono. +End Comp. + +Section Prod. +Variables (f1 : mono_fun S1 T1) (f2 : mono_fun S2 T2). +Lemma funprod_is_mono : monofun_axiom (f1 \* f2). +Proof. by case=>x1 x2 [y1 y2][/= H1 H2]; split=>/=; apply: monofunP. Qed. +HB.instance Definition _ := + isMonoFun.Build (S1 * S2)%type (T1 * T2)%type (f1 \* f2) funprod_is_mono. +End Prod. +End MonoCanonicals. + +Prenex Implicits diag app. (**********) (* Chains *) (**********) -Section Chains. -Variable T : poset. - -Definition chain_axiom (s : Pred T) := +Definition chain_axiom (T : poset) (s : Pred T) := (exists d, d \In s) /\ (forall x y, x \In s -> y \In s -> x <== y \/ y <== x). -Structure chain := Chain { - pred_of :> Pred T; - _ : chain_axiom pred_of}. - -Canonical chainPredType := @mkPredType T chain pred_of. - -End Chains. - -Lemma chainE (T : poset) (s1 s2 : chain T) : - s1 = s2 <-> pred_of s1 =p pred_of s2. -Proof. -split=>[->//|]; move: s1 s2=>[s1 H1][s2 H2] /= E; move: H1 H2. -suff: s1 = s2 by move=>-> H1 H2; congr Chain; apply: pf_irr. -by apply: fext=>x; apply: pext; split; move/E. -Qed. +HB.mixin Record isChain (T : poset) (s : Pred T) := { + chain_subproof : chain_axiom s}. + +#[short(type="chain")] +HB.structure Definition Chain (T : poset) := {s of isChain T s}. -(* common chain constructions *) - -(* adding bot to a chain *) +Arguments chain T%_type. -Section LiftChain. -Variable (T : poset) (s : chain T). +(* \In notation *) +Definition chain_pred (T : poset) (x : chain T) : Pred T := x. +Canonical chainPredType (T : poset) := + Eval hnf in @mkPredType T (@chain T) (@chain_pred T). -Lemma lift_chainP : chain_axiom [Pred x : T | x = bot \/ x \In s]. +Lemma chainE (T : poset) (s1 s2 : chain T) : + s1 = s2 <-> chain_pred s1 =p chain_pred s2. Proof. -case: s=>p [[d H1] H2] /=; split=>[|x y]; first by exists d; right. -by case=>[->|H3][->|H4]; auto. +split=>[->//|]; case: s1 s2=>s1 H1 [s2 H2] /= E; move: H1 H2. +suff: s1 = s2. +- move=>-> [[H1]][[H2]]; congr (Chain.Pack (Chain.Class _)). + by rewrite (pf_irr H1). +by apply: fext=>x; apply: pext; split=> /E. Qed. -Definition lift_chain := Chain lift_chainP. - -End LiftChain. - -(* mapping monotone function over a chain *) - Section ImageChain. -Variables (T1 T2 : poset) (s : chain T1) (f : T1 -> T2) (M : monotone f). - -Lemma image_chainP : - chain_axiom [Pred x2 : T2 | exists x1, x2 = f x1 /\ x1 \In s]. -Proof. -case: s=>p [[d H1] H2]; split=>[|x y]; first by exists (f d); exists d. -case=>x1 [->] H3 [y1][->] H4; rewrite -!toPredE /= in H3 H4. -by case: (H2 x1 y1 H3 H4)=>L; [left | right]; apply: M L. +Variables (T1 T2 : poset) (f : mono_fun T1 T2) (s : chain T1). + +Lemma image_is_chain : chain_axiom (Image f s). +Proof. +case: s=>p [[[[d H1] H2]]] /=. +split=>[|x y]; first by exists (f d); exists d. +case=>x1 -> H3 [y1 -> H4]; rewrite -!toPredE /= in H3 H4. +by case: (H2 x1 y1 H3 H4)=>L; [left | right]; apply: monofunP L. Qed. - -Definition image_chain := Chain image_chainP. - +HB.instance Definition _ := isChain.Build T2 (Image f s) image_is_chain. End ImageChain. -Notation "[ f '^^' s 'by' M ]" := (@image_chain _ _ s f M) - (at level 0, format "[ f '^^' s 'by' M ]") : form_scope. - -Section ChainId. -Variables (T : poset) (s : chain T). - -Lemma id_mono : monotone (@id T). -Proof. by []. Qed. - -Lemma id_chainE (M : monotone id) : [id ^^ s by M] = s. -Proof. by apply/chainE=>x; split; [case=>y [<-]|exists x]. Qed. - -End ChainId. - -Arguments id_mono [T]. -Prenex Implicits id_mono. +Lemma id_chainE T (s : chain T) f : f =1 idfun -> Image f s =p s. +Proof. +move=>H x; split; first by case=>y ->; rewrite H. +by exists x=>//; rewrite H. +Qed. Section ChainConst. Variables (T1 T2 : poset) (y : T2). -Lemma const_mono : monotone (fun x : T1 => y). -Proof. by []. Qed. - -Lemma const_chainP : chain_axiom (Pred1 y). -Proof. by split; [exists y | move=>x1 x2 ->->; left]. Qed. +Definition const_chain : Pred _ := Pred1 y. +Lemma const_is_chain : chain_axiom const_chain. +Proof. by split; [exists y | move=>x1 x2 <-<-; left]. Qed. +HB.instance Definition _ := isChain.Build T2 const_chain const_is_chain. -Definition const_chain := Chain const_chainP. - -Lemma const_chainE s : [_ ^^ s by const_mono] = const_chain. +Lemma const_chainE (s : chain T1) : Image (const_fun y) s =p const_chain. Proof. -apply/chainE=>z1; split; first by case=>z2 [->]. -by case: s=>s [[d] H1] H2 <-; exists d. +move=>z1; split; first by case=>z2 ->. +by case: s=>s [[[[d H1]] H2]] /= <-; exists d. Qed. - End ChainConst. -Arguments const_mono [T1 T2 y]. -Prenex Implicits const_mono. - Section ChainCompose. -Variables (T1 T2 T3 : poset) (f1 : T2 -> T1) (f2 : T3 -> T2). -Variables (s : chain T3) (M1 : monotone f1) (M2 : monotone f2). - -Lemma comp_mono : monotone (f1 \o f2). -Proof. by move=>x y H; apply: M1; apply: M2. Qed. - -Lemma comp_chainE : - [f1 ^^ [f2 ^^ s by M2] by M1] = [f1 \o f2 ^^ s by comp_mono]. -Proof. -apply/chainE=>x1; split; first by case=>x2 [->][x3][->]; exists x3. -by case=>x3 [->]; exists (f2 x3); split=>//; exists x3. -Qed. - +Variables T1 T2 T3 : poset. +Variables (f1 : mono_fun T2 T1) (f2 : mono_fun T3 T2). +Variable s : chain T3. +Lemma comp_chainE : Image f1 (Image f2 s) =p Image (f1 \o f2) s. +Proof. by move=>x; apply: ImageIm. Qed. End ChainCompose. -Arguments comp_mono [T1 T2 T3 f1 f2]. -Prenex Implicits comp_mono. - -(* projections out of a chain *) - Section ProjChain. -Variables (T1 T2 : poset) (s : chain [poset of T1 * T2]). - -Lemma proj1_mono : monotone (@fst T1 T2). -Proof. by case=>x1 x2 [y1 y2][]. Qed. - -Lemma proj2_mono : monotone (@snd T1 T2). -Proof. by case=>x1 x2 [y1 y2][]. Qed. - -Definition proj1_chain := [@fst _ _ ^^ s by proj1_mono]. -Definition proj2_chain := [@snd _ _ ^^ s by proj2_mono]. - +Variables T1 T2 : poset. +Variables s : chain (T1 * T2). +Definition proj1_chain : Pred T1 := Image fst s. +HB.instance Definition _ := Chain.copy proj1_chain proj1_chain. +Definition proj2_chain : Pred T2 := Image snd s. +HB.instance Definition _ := Chain.copy proj2_chain proj2_chain. End ProjChain. -Arguments proj1_mono [T1 T2]. -Arguments proj2_mono [T1 T2]. -Prenex Implicits proj1_mono proj2_mono. - -(* diagonal chain *) - Section DiagChain. -Variable (T : poset) (s : chain T). - -Lemma diag_mono : monotone (fun x : T => (x, x)). -Proof. by []. Qed. - -Definition diag_chain := [_ ^^ s by diag_mono]. - -Lemma proj1_diagE : proj1_chain diag_chain = s. -Proof. by rewrite /proj1_chain /diag_chain comp_chainE id_chainE. Qed. - -Lemma proj2_diagE : proj2_chain diag_chain = s. -Proof. by rewrite /proj2_chain /diag_chain comp_chainE id_chainE. Qed. - +Variables (T : poset) (s : chain T). +Definition diag_chain : Pred (T * T) := Image diag s. +HB.instance Definition _ := Chain.copy diag_chain diag_chain. End DiagChain. -Arguments diag_mono [T]. -Prenex Implicits diag_mono. - -(* applying functions from a chain of functions *) - -Section AppChain. -Variables (A : Type) (T : poset) (s : chain [poset of A -> T]). - -Lemma app_mono x : monotone (fun f : A -> T => f x). -Proof. by move=>f1 f2; apply. Qed. - -Definition app_chain x := [_ ^^ s by app_mono x]. - -End AppChain. - -Arguments app_mono [A T]. -Prenex Implicits app_mono. - -(* ditto for dependent functions *) - -Section DAppChain. -Variables (A : Type) (T : A -> poset) (s : chain (dfunPoset T)). - -Lemma dapp_mono x : monotone (fun f : dfunPoset T => f x). -Proof. by move=>f1 f2; apply. Qed. - -Definition dapp_chain x := [_ ^^ s by dapp_mono x]. - -End DAppChain. - -Arguments dapp_mono [A T]. -Prenex Implicits dapp_mono. - -(* pairing chain applications *) - -Section ProdChain. -Variables (S1 S2 T1 T2 : poset) (f1 : S1 -> T1) (f2 : S2 -> T2). -Variables (M1 : monotone f1) (M2 : monotone f2). -Variable (s : chain [poset of S1 * S2]). - -Lemma prod_mono : monotone (f1 \* f2). -Proof. by case=>x1 x2 [y1 y2][/= H1 H2]; split; [apply: M1 | apply: M2]. Qed. - -Definition prod_chain := [f1 \* f2 ^^ s by prod_mono]. - -Lemma proj1_prodE : proj1_chain prod_chain = [f1 ^^ proj1_chain s by M1]. -Proof. -rewrite /proj1_chain /prod_chain !comp_chainE !/comp /=. -by apply/chainE. -Qed. - -Lemma proj2_prodE : proj2_chain prod_chain = [f2 ^^ proj2_chain s by M2]. -Proof. -rewrite /proj2_chain /prod_chain !comp_chainE !/comp /=. -by apply/chainE. -Qed. - -End ProdChain. - -Arguments prod_mono [S1 S2 T1 T2 f1 f2]. -Prenex Implicits prod_mono. - -(* chain of all nats *) - -Section NatChain. - -Lemma nat_chain_axiom : chain_axiom (@PredT nat). -Proof. -split=>[|x y _ _]; first by exists 0. -rewrite /Poset.leq /= [y <= x]leq_eqVlt. -by case: leqP; [left | rewrite orbT; right]. -Qed. - -Definition nat_chain := Chain nat_chain_axiom. - -End NatChain. (*********) (* CPO's *) (*********) -Module CPO. - -Section RawMixin. - -Record mixin_of (T : poset) := Mixin { - mx_lim : chain T -> T; - _ : forall (s : chain T) x, x \In s -> x <== mx_lim s; - _ : forall (s : chain T) x, - (forall y, y \In s -> y <== x) -> mx_lim s <== x}. - -End RawMixin. - -Section ClassDef. +Definition cpo_axiom (T : poset) (bot : T) (lim : chain T -> T) := + [/\ forall x, bot <== x, + forall (s : chain T) x, x \In s -> x <== lim s & + forall (s : chain T) x, + (forall y, y \In s -> y <== x) -> lim s <== x]. -Record class_of (T : Type) := Class { - base : Poset.class_of T; - mixin : mixin_of (Poset.Pack base)}. +HB.mixin Record isCPO T of Poset T := { + bot : T; + lim_op : chain T -> T; + cpo_subproof: cpo_axiom bot lim_op}. -Local Coercion base : class_of >-> Poset.class_of. +#[short(type="cpo")] +HB.structure Definition CPO := {T of Poset T & isCPO T}. -Structure type : Type := Pack {sort; _ : class_of sort; _ : Type}. -Local Coercion sort : type >-> Sortclass. +Definition limx (D : cpo) (s : chain D) of phantom (Pred _) s := lim_op s. +Notation lim s := (limx (Phantom (Pred _) s)). -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 T. - -Definition pack b0 (m0 : mixin_of (Poset.Pack b0)) := - fun m & phant_id m0 m => Pack (@Class T b0 m) T. - -Definition poset := Poset.Pack class. -Definition lim (s : chain poset) : cT := mx_lim (mixin class) s. - -End ClassDef. - -Module Import Exports. -Coercion base : class_of >-> Poset.class_of. -Coercion sort : type >-> Sortclass. -Coercion poset : type >-> Poset.type. -Canonical Structure poset. - -Notation cpo := type. -Notation CPOMixin := Mixin. -Notation CPO T m := (@pack T _ _ m id). - -Notation "[ 'cpo' 'of' T 'for' cT ]" := (@clone T cT _ idfun) - (at level 0, format "[ 'cpo' 'of' T 'for' cT ]") : form_scope. -Notation "[ 'cpo' 'of' T ]" := (@clone T _ _ id) - (at level 0, format "[ 'cpo' 'of' T ]") : form_scope. +Section Repack. +Variable D : cpo. +Lemma botP (x : D) : bot <== x. +Proof. by case: (@cpo_subproof D). Qed. +Lemma limP (s : chain D) x : x \In s -> x <== lim s. +Proof. by case: (@cpo_subproof D)=>_ H _; apply: H. Qed. +Lemma limM (s : chain D) x : (forall y, y \In s -> y <== x) -> lim s <== x. +Proof. by case: (@cpo_subproof D)=>_ _; apply. Qed. +End Repack. -Arguments CPO.lim [cT]. -Prenex Implicits CPO.lim. -Notation lim := CPO.lim. +#[export] Hint Resolve botP : core. -Section Laws. -Variable D : cpo. +Lemma limE (D : cpo) (s1 s2 : chain D) : s1 =p s2 -> lim s1 = lim s2. +Proof. by move/chainE=>->. Qed. -Lemma limP (s : chain D) x : x \In s -> x <== lim s. -Proof. by case: D s x=>S [[p]][/= l H1 *]; apply: (H1). Qed. +(* common chain constructions *) -Lemma limM (s : chain D) x : (forall y, y \In s -> y <== x) -> lim s <== x. -Proof. by case: D s x=>S [[p]][/= l H1 H2 *]; apply: (H2). Qed. +(* adding bot to a chain *) -End Laws. +Definition lift (D : cpo) (s : Pred D) : Pred D := + fun x => x = bot \/ x \In s. -End Exports. +Lemma lift_is_chain (D : cpo) (s : chain D) : chain_axiom (lift s). +Proof. +case: s=>p [[[[d H1]] H2]] /=. +split=>[|x y]; first by exists d; right. +by case=>[->|H3][->|H4]; auto. +Qed. -End CPO. +HB.instance Definition _ (D : cpo) (s : chain D) := + isChain.Build D (lift s) (@lift_is_chain D s). -Export CPO.Exports. (****************************) (* common cpo constructions *) (****************************) -(* pairs *) - -Section PairCPO. +(* products *) +Section ProdCPO. Variables (A B : cpo). -Local Notation tp := [poset of A * B]. - -Definition pair_lim (s : chain tp) : tp := - (lim (proj1_chain s), lim (proj2_chain s)). - -Lemma pair_limP (s : chain tp) x : x \In s -> x <== pair_lim s. -Proof. by split; apply: limP; exists x. Qed. +Definition cpo_prod_bot := (@bot A, @bot B). +Definition cpo_prod_lim (s : chain (A * B)) := + (lim (Image fst s), lim (Image snd s)). -Lemma pair_limM (s : chain tp) x : - (forall y, y \In s -> y <== x) -> pair_lim s <== x. -Proof. by split; apply: limM=>y [z][->]; case/H. Qed. - -Definition pairCPOMixin := CPOMixin pair_limP pair_limM. -Canonical pairCPO := Eval hnf in CPO (A * B) pairCPOMixin. - -End PairCPO. +Lemma prod_is_cpo : cpo_axiom cpo_prod_bot cpo_prod_lim. +Proof. +split=>[x|s x|s x]. +- by split; apply: botP. +- by split; apply: limP; exists x. +by split; apply: limM=>y [z ->]; case/H. +Qed. +HB.instance Definition _ := isCPO.Build (A * B)%type prod_is_cpo. +End ProdCPO. (* functions *) - Section FunCPO. Variable (A : Type) (B : cpo). -Local Notation tp := [poset of A -> B]. - -Definition fun_lim (s : chain tp) : tp := - fun x => lim (app_chain s x). - -Lemma fun_limP (s : chain tp) x : x \In s -> x <== fun_lim s. -Proof. by move=>H t; apply: limP; exists x. Qed. - -Lemma fun_limM (s : chain tp) x : - (forall y, y \In s -> y <== x) -> fun_lim s <== x. -Proof. by move=>H1 t; apply: limM=>y [f][->] H2; apply: H1. Qed. - -Definition funCPOMixin := CPOMixin fun_limP fun_limM. -Canonical funCPO := Eval hnf in CPO (A -> B) funCPOMixin. +Definition cpo_fun_bot (x : A) := @bot B. +Definition cpo_fun_lim (s : chain (A -> B)) x := + lim (Image (app x) s). +Lemma fun_is_cpo : cpo_axiom cpo_fun_bot cpo_fun_lim. +Proof. +split=>[f|s f|s f]. +- by move=>y; apply: botP. +- by move=>H t; apply: limP; exists f. +by move=>H1 t; apply: limM=>y [g ->] H2; apply: H1. +Qed. +HB.instance Definition _ := isCPO.Build (A -> B) fun_is_cpo. End FunCPO. (* dependent functions *) - Section DFunCPO. -Variable (A : Type) (B : A -> cpo). -Local Notation tp := (dfunPoset B). - -Definition dfun_lim (s : chain tp) : tp := - fun x => lim (dapp_chain s x). +Variables (A : Type) (B : A -> cpo). -Lemma dfun_limP (s : chain tp) x : x \In s -> x <== dfun_lim s. -Proof. by move=>H t; apply: limP; exists x. Qed. +Definition cpo_dfun_bot : dfun_poset B := fun x => bot. +Definition cpo_dfun_lim (s : chain (dfun_poset B)) : dfun_poset B := + fun x => lim (Image (dapp x) s). -Lemma dfun_limM (s : chain tp) x : - (forall y, y \In s -> y <== x) -> dfun_lim s <== x. -Proof. by move=>H1 t; apply: limM=>y [f][->] H2; apply: H1. Qed. - -Definition dfunCPOMixin := CPOMixin dfun_limP dfun_limM. -Definition dfunCPO := Eval hnf in CPO (forall x, B x) dfunCPOMixin. +Lemma dfun_is_cpo : cpo_axiom cpo_dfun_bot cpo_dfun_lim. +Proof. +split=>[x|s x|s x]. +- by move=>y; apply: botP. +- by move=>H t; apply: limP; exists x. +by move=>H1 t; apply: limM=>y [f ->] H2; apply: H1. +Qed. +Definition dfun_cpo_mix := isCPO.Build (dfun_poset B) dfun_is_cpo. +Definition dfun_cpo : cpo := CPO.pack_ dfun_cpo_mix. End DFunCPO. (* Prop *) - -Section PropCPO. -Local Notation tp := [poset of Prop]. - -Definition prop_lim (s : chain tp) : tp := exists p, p \In s /\ p. - -Lemma prop_limP (s : chain tp) p : p \In s -> p <== prop_lim s. -Proof. by exists p. Qed. - -Lemma prop_limM (s : chain tp) p : - (forall q, q \In s -> q <== p) -> prop_lim s <== p. -Proof. by move=>H [r][]; move/H. Qed. - -Definition propCPOMixin := CPOMixin prop_limP prop_limM. -Canonical propCPO := Eval hnf in CPO Prop propCPOMixin. - -End PropCPO. +Definition cpo_prop_bot : Prop := False. +Definition cpo_prop_lim (s : chain Prop) : Prop := + exists p, p \In s /\ p. +Lemma prop_is_cpo : cpo_axiom cpo_prop_bot cpo_prop_lim. +Proof. by split=>[x|s p|s p H] //=; [exists p|case=>r []; move/H]. Qed. +HB.instance Definition _ := isCPO.Build Prop prop_is_cpo. (* Pred *) +HB.instance Definition _ A := CPO.copy (Pred A) _. -Section PredCPO. -Variable A : Type. - -Definition predCPOMixin := funCPOMixin A propCPO. -Canonical predCPO := Eval hnf in CPO (Pred A) predCPOMixin. - -End PredCPO. - -(* every complete lattice is a cpo *) - +(* every complete lattice is cpo *) Section LatticeCPO. -Variable A : lattice. -Local Notation tp := (Lattice.poset A). +Variable A : lattice. -Definition lat_lim (s : chain tp) : tp := sup s. +Definition lat_bot : A := lbot. +Definition lat_lim (s : Pred A) : A := sup s. -Lemma lat_limP (s : chain tp) x : x \In s -> x <== lat_lim s. -Proof. by apply: supP. Qed. - -Lemma lat_limM (s : chain tp) x : - (forall y, y \In s -> y <== x) -> lat_lim s <== x. -Proof. by apply: supM. Qed. - -Definition latCPOMixin := CPOMixin lat_limP lat_limM. -Definition latCPO := Eval hnf in CPO tp latCPOMixin. +Lemma lat_is_cpo : cpo_axiom lat_bot lat_lim. +Proof. by split=>[x|s x|s x]; [apply: supM|apply: supP|apply: supM]. Qed. +(* no instance declaration as nothing to *) +(* hang the new structure onto *) +Definition lat_cpo_mix := isCPO.Build A lat_is_cpo. +Definition lat_cpo := CPO.pack_ lat_cpo_mix. End LatticeCPO. (* sub-CPO's *) @@ -1143,7 +760,7 @@ Proof. by move=>H1 H2 p; move/(_ _ H1 H2). Qed. Lemma chain_closP (s : Pred T) : chain_closed (chain_closure s). Proof. split; first by move=>t _ []. -move=>d H1 t H3 H4; move: (chain_clos_min H3 H4)=>{H3} -H3. +move=>d H1 t H3 H4; move: (chain_clos_min H3 H4)=>{}H3. by case: H4=>_; apply=>// x; move/H1; move/H3. Qed. @@ -1151,7 +768,7 @@ Lemma chain_clos_idemp (s : Pred T) : chain_closed s -> chain_closure s =p s. Proof. move=>p; split; last by apply: chain_clos_sub. -by apply: chain_clos_min=>//; apply: chain_closP. +by apply: chain_clos_min=>// /chain_closP. Qed. Lemma chain_clos_mono (s1 s2 : Pred T) : @@ -1178,35 +795,30 @@ Prenex Implicits chain_closed. Lemma chain_clos_diag (T : cpo) (s : Pred (T * T)) : chain_closed s -> chain_closed [Pred t : T | (t, t) \In s]. Proof. -move=>[B H1]; split=>// d H2. -rewrite InE /= -{1}(proj1_diagE d) -{2}(proj2_diagE d). -by apply: H1; case=>x1 x2 [x][[<- <-]]; apply: H2. +case=>B H1; split=>// d H2; move: (H1 (Image diag d)). +rewrite /limx/=/lim_op /= /cpo_prod_lim /=. +rewrite !(limE (comp_chainE _ _ _)) !(limE (id_chainE _ _)) //=. +by apply; case=>x1 x2 [z ->]; apply: H2. Qed. Section SubCPO. Variables (D : cpo) (s : Pred D) (C : chain_closed s). +Local Notation tp := {x : D | x \In s}. -Local Notation tp := (subPoset (proj1 C)). - -Lemma sval_mono : monotone (sval : tp -> D). -Proof. by move=>[x1 H1][x2 H2]; apply. Qed. - -Lemma sub_limX (u : chain tp) : lim [sval ^^ u by sval_mono] \In s. -Proof. by case: C u=>P H u; apply: (H)=>t [[y]] H1 [->]. Qed. - -Definition sub_lim (u : chain tp) : tp := - exist _ (lim [sval ^^ u by sval_mono]) (sub_limX u). - -Lemma sub_limP (u : chain tp) x : x \In u -> x <== sub_lim u. -Proof. by move=>H; apply: limP; exists x. Qed. - -Lemma sub_limM (u : chain tp) x : - (forall y, y \In u -> y <== x) -> sub_lim u <== x. -Proof. by move=>H; apply: limM=>y [z][->]; apply: H. Qed. - -Definition subCPOMixin := CPOMixin sub_limP sub_limM. -Definition subCPO := Eval hnf in CPO {x : D | x \In s} subCPOMixin. +Definition subcpo_bot := exist _ (@bot D) (proj1 C). +Lemma subcpo_limX (u : chain tp) : lim (Image sval u) \In s. +Proof. by case: C u=>P H u; apply: (H)=>t [[y H1 ->]]. Qed. +Definition subcpo_lim (u : chain tp) : tp := + exist _ (lim (Image sval u)) (subcpo_limX u). +Lemma subcpo_is_cpo : cpo_axiom subcpo_bot subcpo_lim. +Proof. +split=>[x|u x H|u x H] //=. +- by apply: botP. +- by apply: limP; exists x. +by apply: limM=>y [z ->]; apply: H. +Qed. +HB.instance Definition _ := isCPO.Build tp subcpo_is_cpo. End SubCPO. (***********************************************) @@ -1215,14 +827,9 @@ End SubCPO. Lemma lim_mono (D : cpo) (s1 s2 : chain D) : s1 <=p s2 -> lim s1 <== lim s2. -Proof. by move=>H; apply: limM=>y; move/H; apply: limP. Qed. +Proof. by move=>H; apply: limM=>y /H/limP. Qed. -Lemma limE (D : cpo) (s1 s2 : chain D) : - s1 =p s2 -> lim s1 = lim s2. -Proof. by move=>H; apply: poset_asym; apply: lim_mono=>x; rewrite H. Qed. - -Lemma lim_liftE (D : cpo) (s : chain D) : - lim s = lim (lift_chain s). +Lemma lim_liftE (D : cpo) (s : chain D) : lim s = lim (lift s). Proof. apply: poset_asym; apply: limM=>y H; first by apply: limP; right. by case: H; [move=>-> | apply: limP]. @@ -1231,137 +838,200 @@ Qed. (* applied lim equals the lim of applications *) (* ie., part of continuity of application *) (* but is so often used, I give it a name *) - -Lemma lim_appE A (D : cpo) (s : chain [cpo of A -> D]) (x : A) : - lim s x = lim (app_chain s x). +Lemma lim_appE A (D : cpo) (s : chain (A -> D)) (x : A) : + lim s x = lim (Image (app x) s). Proof. by []. Qed. -Lemma lim_dappE A (D : A -> cpo) (s : chain (dfunCPO D)) (x : A) : - lim s x = lim (dapp_chain s x). +Lemma lim_dappE A (D : A -> cpo) (s : chain (dfun_cpo D)) (x : A) : + lim s x = lim (Image (dapp x) s). Proof. by []. Qed. -Section Continuity. -Variables (D1 D2 : cpo) (f : D1 -> D2). +(************************) +(* continuous functions *) +(************************) -Definition continuous := - exists M : monotone f, - forall s : chain D1, f (lim s) = lim [f ^^ s by M]. +Definition contfun_axiom (D1 D2 : cpo) (f : mono_fun D1 D2) := + forall s : chain D1, f (lim s) <== lim (Image f s). -Lemma cont_mono : continuous -> monotone f. -Proof. by case. Qed. +HB.mixin Record isContFun (D1 D2 : cpo) (f : D1 -> D2) + of @MonoFun D1 D2 f := {contfun_subproof : contfun_axiom f}. -Lemma contE (s : chain D1) (C : continuous) : - f (lim s) = lim [f ^^ s by cont_mono C]. -Proof. -case: C=>M E; rewrite E; congr (lim (image_chain _ _)). -by apply: pf_irr. +#[short(type="cont_fun")] +HB.structure Definition ContFun (D1 D2 : cpo) := + {f of @MonoFun D1 D2 f & isContFun D1 D2 f}. + +Lemma contE (D1 D2 : cpo) (f : cont_fun D1 D2) (s : chain D1) : + f (lim s) = lim (Image f s). +Proof. +apply: poset_asym; first by apply: contfun_subproof. +by apply: limM=>y [x1 ->] /limP/monofunP. Qed. -End Continuity. +(* just for this proof, we want that nat is a poset under <= *) +(* in other scopes, we'll use the equality as ordering *) +(* Thus, we encapsulate under module. *) + +Module Kleene. + +Lemma nat_is_poset : poset_axiom leq. +Proof. +split=>[|x y H1 H2|x y z] //; last by apply: leq_trans. +by apply: anti_leq; rewrite H1 H2. +Qed. +HB.instance Definition _ := isPoset.Build nat nat_is_poset. + +(* nats form a chain *) +Definition nats : Pred nat := PredT. +Lemma nats_chain_axiom : chain_axiom nats. +Proof. +split=>[|x y _ _]; first by exists 0. +rewrite /poset_leq /= [y <= x]leq_eqVlt. +by case: leqP; [left | rewrite orbT; right]. +Qed. +HB.instance Definition _ := isChain.Build nat nats nats_chain_axiom. Section Kleene. -Variables (D : cpo) (f : D -> D) (C : continuous f). +Variables (D : cpo) (f : cont_fun D D). Fixpoint pow m := if m is n.+1 then f (pow n) else bot. -Lemma pow_mono : monotone pow. +Lemma pow_is_mono : monofun_axiom pow. Proof. move=>m n; elim: n m=>[|n IH] m /=; first by case: m. -rewrite {1}/Poset.leq /= leq_eqVlt ltnS. +rewrite {1}/poset_leq /= leq_eqVlt ltnS. case/orP; first by move/eqP=>->. move/IH=>{IH} H; apply: poset_trans H _. -by elim: n=>[|n IH] //=; apply: cont_mono IH. +by elim: n=>[|n IH] //=; apply: monofunP IH. Qed. +HB.instance Definition _ := isMonoFun.Build nat D pow pow_is_mono. -Definition pow_chain := [pow ^^ nat_chain by pow_mono]. - -Lemma reindex : pow_chain =p lift_chain [f ^^ pow_chain by cont_mono C]. +Lemma reindex : Image pow nats =p lift (Image f (Image pow nats)). Proof. move=>x; split. -- case; case=>[|n][->] /=; first by left. - by right; exists (pow n); split=>//; exists n. +- case; case=>[|n] -> _; first by left. + by right; exists (pow n)=>//; exists n. case=>/=; first by move=>->; exists 0. -by case=>y [->][n][->]; exists n.+1. +by case=>y -> [n ->]; exists n.+1. Qed. -Definition kleene_lfp := lim pow_chain. +Definition kleene_lfp := lim (Image pow nats). Lemma kleene_lfp_fixed : f kleene_lfp = kleene_lfp. -Proof. by rewrite (@contE _ _ f) lim_liftE; apply: limE; rewrite reindex. Qed. +Proof. by rewrite /kleene_lfp contE /= (limE reindex) /= lim_liftE. Qed. -Lemma kleene_lfp_least : forall x, f x = x -> kleene_lfp <== x. +Lemma kleene_lfp_least x : f x = x -> kleene_lfp <== x. Proof. -move=>x H; apply: limM=>y [n][->] _. -by elim: n=>[|n IH] //=; rewrite -H; apply: cont_mono IH. +move=>H; apply: limM=>y [n ->] _. +by elim: n=>[|n IH] //=; rewrite -H; apply: monofunP IH. Qed. End Kleene. -(**********************************) -(* Continuity of common functions *) -(**********************************) - -Lemma id_cont (D : cpo) : continuous (@id D). -Proof. by exists id_mono; move=>d; rewrite id_chainE. Qed. - -Arguments id_cont {D}. -Prenex Implicits id_cont. - -Lemma const_cont (D1 D2 : cpo) (y : D2) : continuous (fun x : D1 => y). +Lemma kleene_lfp_mono (D : cpo) (f1 f2 : cont_fun D D) : + (f1 : D -> D) <== f2 -> + kleene_lfp f1 <== kleene_lfp f2. Proof. -exists const_mono; move=>s; apply: poset_asym. -- by apply: limP; case: s=>[p][[d H1] H2]; exists d. -by apply: limM=>_ [x][->]. +move=>H; apply: limM=>x [n ->] X. +have N : forall n, pow f1 n <== pow f2 n. +- by elim=>//= m /(monofunP f2); apply: poset_trans (H _). +by apply: poset_trans (N n) _; apply: limP; exists n. Qed. -Arguments const_cont {D1 D2 y}. -Prenex Implicits const_cont. +Module Exports. +Notation kleene_lfp := kleene_lfp. +Notation kleene_lfp_fixed := kleene_lfp_fixed. +Notation kleene_lfp_least := kleene_lfp_least. +Notation kleene_lfp_mono := kleene_lfp_mono. +End Exports. -Lemma comp_cont (D1 D2 D3 : cpo) (f1 : D2 -> D1) (f2 : D3 -> D2) : - continuous f1 -> continuous f2 -> continuous (f1 \o f2). -Proof. -case=>M1 H1 [M2 H2]; exists (comp_mono M1 M2); move=>d. -by rewrite /= H2 H1 comp_chainE. -Qed. +End Kleene. -Arguments comp_cont [D1 D2 D3 f1 f2]. -Prenex Implicits comp_cont. +Export Kleene.Exports. -Lemma proj1_cont (D1 D2 : cpo) : continuous (@fst D1 D2). -Proof. by exists proj1_mono. Qed. -Lemma proj2_cont (D1 D2 : cpo) : continuous (@snd D1 D2). -Proof. by exists proj2_mono. Qed. +(**********************************) +(* Continuity of common functions *) +(**********************************) -Arguments proj1_cont {D1 D2}. -Arguments proj2_cont {D1 D2}. -Prenex Implicits proj1_cont proj2_cont. +Lemma idfun_is_cont (D : cpo) : contfun_axiom (@idfun D). +Proof. by move=>d; rewrite (limE (id_chainE _ _)). Qed. +HB.instance Definition _ (D : cpo) := + isContFun.Build D D idfun (@idfun_is_cont D). -Lemma diag_cont (D : cpo) : continuous (fun x : D => (x, x)). +Lemma const_is_cont (D1 D2 : cpo) (y : D2) : + contfun_axiom (@const_fun D1 D2 y). Proof. -exists diag_mono=>d; apply: poset_asym; -by split=>/=; [rewrite proj1_diagE | rewrite proj2_diagE]. +by move=>s; apply: limP; case: s=>[p][[[[d H1] H2]]]; exists d. Qed. +HB.instance Definition _ (D1 D2 : cpo) (y : D2) := + isContFun.Build D1 D2 (const_fun y) (const_is_cont y). + +Section ContCompose. +Variables D1 D2 D3 : cpo. +Variable (f1 : cont_fun D2 D1) (f2 : cont_fun D3 D2). +Lemma comp_is_cont : contfun_axiom (f1 \o f2). +Proof. by move=>s; rewrite /= !contE (limE (comp_chainE _ _ _)). Qed. +HB.instance Definition _ := isContFun.Build D3 D1 (f1 \o f2) comp_is_cont. +End ContCompose. + +Lemma fst_is_cont (D1 D2 : cpo) : contfun_axiom (@fst D1 D2). +Proof. by []. Qed. +HB.instance Definition _ (D1 D2 : cpo) := + isContFun.Build (D1*D2)%type D1 fst (@fst_is_cont D1 D2). -Arguments diag_cont {D}. -Prenex Implicits diag_cont. +Lemma snd_is_cont (D1 D2 : cpo) : contfun_axiom (@snd D1 D2). +Proof. by []. Qed. +HB.instance Definition _ (D1 D2 : cpo) := + isContFun.Build (D1*D2)%type D2 snd (@snd_is_cont D1 D2). -Lemma app_cont A (D : cpo) x : continuous (fun f : A -> D => f x). -Proof. by exists (app_mono x). Qed. +Lemma diag_is_cont (D : cpo) : contfun_axiom (@diag D). +Proof. +move=>d; split=>/=; +by rewrite (limE (comp_chainE _ _ _)) (limE (id_chainE _ _)). +Qed. +HB.instance Definition _ (D : cpo) := + isContFun.Build D (D*D)%type diag (@diag_is_cont D). -Lemma dapp_cont A (D : A -> cpo) x : continuous (fun f : dfunCPO D => f x). -Proof. by exists (dapp_mono x). Qed. +Lemma app_is_cont A (D : cpo) x : contfun_axiom (@app A D x). +Proof. by []. Qed. +HB.instance Definition _ A (D : cpo) x := + isContFun.Build (A -> D) D (@app A D x) (app_is_cont x). + +(* can't reuse dapp and it's proof of monotonicity *) +(* because inheritance doesn't propagate properly *) +(* for dependent functions; so we duplicate *) +(* See mathcomp/finfun.v for solution (dfun vs. dffun). *) +Definition dapp2 A (B : A -> cpo) (x : A) := + fun f : dfun_cpo B => f x. +Lemma dapp2_is_mono A (B : A -> cpo) x : monofun_axiom (@dapp2 A B x). +Proof. by move=>f1 f2; apply. Qed. +HB.instance Definition _ A (B : A -> cpo) x := + isMonoFun.Build (dfun_poset B) (B x) (@dapp2 A B x) (dapp2_is_mono x). +Lemma dapp2_is_cont A (B : A -> cpo) (x : A) : contfun_axiom (@dapp2 A B x). +Proof. +move=>s; have <- // : dapp2 x (lim s) = lim [Image dapp2 x i | i <- s]. +by apply: limE=>/=. +Qed. +HB.instance Definition _ A (B : A -> cpo) x := + isContFun.Build (dfun_cpo B) (B x) (@dapp2 A B x) (dapp2_is_cont x). -Arguments app_cont [A D]. -Arguments dapp_cont [A D]. -Prenex Implicits app_cont dapp_cont. +Section ProdCont. +Variables S1 S2 T1 T2 : cpo. +Variables (f1 : cont_fun S1 T1) (f2 : cont_fun S2 T2). -Lemma prod_cont (S1 S2 T1 T2 : cpo) (f1 : S1 -> T1) (f2 : S2 -> T2) : - continuous f1 -> continuous f2 -> continuous (f1 \* f2). +Lemma prod_is_cont : contfun_axiom (f1 \* f2). Proof. -case=>M1 H1 [M2 H2]; exists (prod_mono M1 M2)=>d; apply: poset_asym; -by (split=>/=; [rewrite proj1_prodE H1 | rewrite proj2_prodE H2]). +move=>d; split=>//=; +by rewrite contE !(limE (comp_chainE _ _ _)); apply: lim_mono. Qed. +HB.instance Definition _ := + isContFun.Build (S1*S2)%type (T1*T2)%type (f1 \* f2) prod_is_cont. +End ProdCont. + +(* DEVCOMMENT *) +(* TODO: *) +(* 1. limit of a chain of continuous functions is a continuous function *) +(* 2. CPO of continuous functions should be the subCPO of functions *) +(* 3. kleene_lfp is a continuous operation *) +(* /DEVCOMMENT *) -Arguments prod_cont [S1 S2 T1 T2 f1 f2]. -Prenex Implicits prod_cont. diff --git a/htt/interlude.v b/htt/interlude.v index bf75a71..364243a 100644 --- a/htt/interlude.v +++ b/htt/interlude.v @@ -2,12 +2,15 @@ From Coq Require Import ssreflect ssrbool ssrfun. From mathcomp Require Import ssrnat seq eqtype path fintype finfun tuple interval perm fingroup. From pcm Require Import options prelude ordtype seqext slice. +(* Lemma implyb_trans a b c : a ==> b -> b ==> c -> a ==> c. Proof. by case: a=>//=->. Qed. +*) Section Seq. Variable (A : Type). +(* Lemma rcons_nseq n (x : A) : rcons (nseq n x) x = nseq n.+1 x. Proof. by elim: n=>//=n ->. Qed. @@ -16,12 +19,16 @@ Lemma behead_rcons (xs : seq A) x : 0 < size xs -> behead (rcons xs x) = rcons (behead xs) x. Proof. by case: xs. Qed. +*) +(* Lemma path_all (xs : seq A) x r : transitive r -> path r x xs -> all (r x) xs. Proof. by move=>Ht; rewrite path_sortedE; [case/andP | exact: Ht]. Qed. +*) +(* Lemma sorted_rconsE (leT : rel A) xs x : transitive leT -> sorted leT (rcons xs x) = all (leT^~ x) xs && sorted leT xs. @@ -29,21 +36,27 @@ Proof. move/rev_trans=>Ht; rewrite -(revK (rcons _ _)) rev_rcons rev_sorted /=. by rewrite path_sortedE // all_rev rev_sorted. Qed. +*) +(* Lemma sorted1 (r : rel A) xs : size xs == 1 -> sorted r xs. Proof. by case: xs=>// x; case. Qed. +*) End Seq. Section SeqEq. Variable (A : eqType). +(* Lemma perm_cons2 (x y : A) s : perm_eq [:: x, y & s] [:: y, x & s]. Proof. by rewrite (_ : [::x,y & s] = [::x] ++ [::y] ++ s) // (_ : [::y,x & s] = [::y] ++ [::x] ++ s) // perm_catCA. Qed. +*) +(* (* a weaker form of in_split *) Lemma mem_split (x : A) s : x \in s -> exists s1 s2, s = s1 ++ [:: x] ++ s2. @@ -59,16 +72,20 @@ move=>/[swap] Hu /mem_split [s1 [s2 H]]. exists s1, s2; move: Hu. by rewrite H uniq_catCA cons_uniq; case/andP. Qed. +*) +(* Lemma all_notin (p : pred A) xs y : all p xs -> ~~ p y -> y \notin xs. Proof. by move/allP=>Ha; apply/contra/Ha. Qed. Lemma subset_all a (s1 s2 : seq A) : {subset s1 <= s2} -> all a s2 -> all a s1. Proof. by move=>Hs /allP Ha1; apply/allP=>s /Hs /Ha1. Qed. +*) End SeqEq. +(* Section Allrel. Variables (S T : Type). @@ -81,7 +98,9 @@ Lemma allrel_rconsr (r : T -> S -> bool) y xs ys : Proof. by rewrite -cats1 allrel_catr allrel1r. Qed. End Allrel. +*) +(* Section AllrelEq. Variables (S T : eqType). @@ -124,7 +143,9 @@ by move/Ha/Ht; apply; apply: Hp. Qed. End AllrelEq. +*) +(* Section SeqOrd. Variable (A : ordType). @@ -143,9 +164,11 @@ by rewrite (allrel_trans (@trans A) Hl Hr). Qed. End SeqOrd. +*) (* surgery on finfuns: slicing & permuting *) +(* Section OnthCodom. Variable (A : Type). @@ -165,7 +188,9 @@ by rewrite (onth_tnth (fgraph f) i'). Qed. End OnthCodom. +*) +(* Section CodomWS. Variable (n : nat) (A : Type). @@ -174,7 +199,9 @@ Lemma codom_ux_rcons (f : {ffun 'I_n -> A}) (i : 'I_n) : Proof. by rewrite slice_xR // slice_uu onth_codom. Qed. End CodomWS. +*) +(* Section PermFfun. Variables (I : finType) (A : Type). @@ -190,3 +217,4 @@ Lemma pffunEM (p p' : {perm I}) (f : {ffun I -> A}) : Proof. by apply/ffunP => i; rewrite !ffunE permM. Qed. End PermFfun. +*) diff --git a/htt/options.v b/htt/options.v index ddcdc8f..c519867 100644 --- a/htt/options.v +++ b/htt/options.v @@ -1,3 +1,6 @@ -Export Unset Program Cases. (* turn off the automation of Program *) -Obligation Tactic := auto. +#[export] Obligation Tactic := auto. +(* turn off other Program extensions *) +#[export] Unset Program Cases. +#[export] Unset Program Generalized Coercion. +#[export] Unset Program Mode. From fd96996ff9eb0c4d14f05acbbf2f7877ab618c27 Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Thu, 1 Aug 2024 20:16:15 +0200 Subject: [PATCH 05/93] removed interlude.v --- _CoqProject | 1 - examples/bst.v | 3 +- examples/bubblesort.v | 2 +- examples/congmath.v | 1 - examples/counter.v | 2 +- examples/cyclic.v | 4 +- examples/kvmaps.v | 2 +- examples/quicksort.v | 2 +- examples/tree.v | 1 - examples/union_find.v | 2 +- htt/interlude.v | 220 ------------------------------------------ 11 files changed, 8 insertions(+), 232 deletions(-) delete mode 100644 htt/interlude.v diff --git a/_CoqProject b/_CoqProject index 233c167..a4e3600 100644 --- a/_CoqProject +++ b/_CoqProject @@ -31,7 +31,6 @@ pcm/automap.v pcm/heap.v pcm/mutex.v htt/options.v -htt/interlude.v htt/domain.v htt/model.v htt/heapauto.v diff --git a/examples/bst.v b/examples/bst.v index e61fabf..37ff69a 100644 --- a/examples/bst.v +++ b/examples/bst.v @@ -2,8 +2,7 @@ From Coq Require Import ssreflect ssrbool ssrfun. From mathcomp Require Import eqtype ssrnat seq path. From pcm Require Import options axioms pred ordtype. From pcm Require Import pcm unionmap heap autopcm automap. -From htt Require Import interlude model heapauto. -From htt Require Import bintree. +From htt Require Import model heapauto bintree. Section BST. Context {T : ordType}. diff --git a/examples/bubblesort.v b/examples/bubblesort.v index fe2731f..6fe6e19 100644 --- a/examples/bubblesort.v +++ b/examples/bubblesort.v @@ -2,7 +2,7 @@ From Coq Require Import ssreflect ssrbool ssrfun. From mathcomp Require Import ssrnat eqtype seq path fintype tuple finfun fingroup perm interval. From pcm Require Import options axioms prelude pred ordtype slice. From pcm Require Import pcm unionmap heap. -From htt Require Import options interlude model heapauto. +From htt Require Import options model heapauto. From htt Require Import array. (* hack to avoid "_ *p _" notation clash *) diff --git a/examples/congmath.v b/examples/congmath.v index fde0566..64f72b9 100755 --- a/examples/congmath.v +++ b/examples/congmath.v @@ -5,7 +5,6 @@ From Coq Require Import Recdef ssreflect ssrbool ssrfun. From mathcomp Require Import eqtype choice ssrnat seq bigop fintype finfun. From pcm Require Import options prelude ordtype finmap pred seqext. -From htt Require Import interlude. Ltac add_morphism_tactic := SetoidTactics.add_morphism_tactic. diff --git a/examples/counter.v b/examples/counter.v index fdb4081..cc57ebf 100755 --- a/examples/counter.v +++ b/examples/counter.v @@ -2,7 +2,7 @@ From mathcomp Require Import ssreflect ssrbool ssrfun. From mathcomp Require Import ssrnat eqtype. From pcm Require Import options axioms prelude pred. From pcm Require Import pcm unionmap heap. -From htt Require Import options interlude model heapauto. +From htt Require Import options model heapauto. (* counter that hides local state with an abstract predicate *) diff --git a/examples/cyclic.v b/examples/cyclic.v index 7b8501a..02705dc 100644 --- a/examples/cyclic.v +++ b/examples/cyclic.v @@ -2,7 +2,7 @@ From Coq Require Import ssreflect ssrbool ssrfun. From mathcomp Require Import eqtype seq ssrnat. From pcm Require Import options axioms pred. From pcm Require Import pcm unionmap heap automap autopcm. -From htt Require Import options interlude model heapauto. +From htt Require Import options model heapauto. From htt Require Import llist. (* a queue variant that has a fixed capacity and can overwrite data in a circular way *) @@ -315,4 +315,4 @@ by rewrite unitR. Qed. End Buffer. -End Buffer. \ No newline at end of file +End Buffer. diff --git a/examples/kvmaps.v b/examples/kvmaps.v index 8b9f1cd..1909fe6 100755 --- a/examples/kvmaps.v +++ b/examples/kvmaps.v @@ -6,7 +6,7 @@ From Coq Require Import ssreflect ssrbool ssrfun. From mathcomp Require Import eqtype seq path. From pcm Require Import options axioms pred ordtype finmap. From pcm Require Import pcm unionmap heap autopcm automap. -From htt Require Import options interlude model heapauto. +From htt Require Import options model heapauto. (***********************) (* stateful KV map ADT *) diff --git a/examples/quicksort.v b/examples/quicksort.v index ff27146..0e6bbba 100644 --- a/examples/quicksort.v +++ b/examples/quicksort.v @@ -4,7 +4,7 @@ From mathcomp Require Import fingroup perm. From mathcomp Require Import interval. From pcm Require Import options axioms prelude pred ordtype slice. From pcm Require Import pcm unionmap heap. -From htt Require Import options interlude model heapauto. +From htt Require Import options model heapauto. From htt Require Import array. (* hack to avoid "_ *p _" notation clash *) diff --git a/examples/tree.v b/examples/tree.v index 23c4027..1a99b79 100644 --- a/examples/tree.v +++ b/examples/tree.v @@ -1,7 +1,6 @@ From mathcomp Require Import ssreflect ssrbool ssrfun. From mathcomp Require Import eqtype ssrnat seq bigop choice. From pcm Require Import pred prelude seqext. -From htt Require Import interlude. (* arbitrarily branching tree *) Inductive tree A := TNode of A & seq (tree A). diff --git a/examples/union_find.v b/examples/union_find.v index 2c002aa..ec4e775 100644 --- a/examples/union_find.v +++ b/examples/union_find.v @@ -2,7 +2,7 @@ From mathcomp Require Import ssreflect ssrbool ssrfun fintype. From mathcomp Require Import eqtype ssrnat seq bigop choice. From pcm Require Import options axioms pred seqext. From pcm Require Import prelude pcm unionmap natmap heap autopcm automap. -From htt Require Import options interlude model heapauto tree. +From htt Require Import options model heapauto tree. (* TODO: remove these after upgrading FCSL-PCM to 1.9+ *) Corollary foldr_join A (U : pcm) h (s : seq A) (a : A -> U): diff --git a/htt/interlude.v b/htt/interlude.v deleted file mode 100644 index 364243a..0000000 --- a/htt/interlude.v +++ /dev/null @@ -1,220 +0,0 @@ -From Coq Require Import ssreflect ssrbool ssrfun. -From mathcomp Require Import ssrnat seq eqtype path fintype finfun tuple interval perm fingroup. -From pcm Require Import options prelude ordtype seqext slice. - -(* -Lemma implyb_trans a b c : a ==> b -> b ==> c -> a ==> c. -Proof. by case: a=>//=->. Qed. -*) - -Section Seq. -Variable (A : Type). - -(* -Lemma rcons_nseq n (x : A) : - rcons (nseq n x) x = nseq n.+1 x. -Proof. by elim: n=>//=n ->. Qed. - -Lemma behead_rcons (xs : seq A) x : - 0 < size xs -> - behead (rcons xs x) = rcons (behead xs) x. -Proof. by case: xs. Qed. -*) - -(* -Lemma path_all (xs : seq A) x r : - transitive r -> - path r x xs -> all (r x) xs. -Proof. by move=>Ht; rewrite path_sortedE; [case/andP | exact: Ht]. Qed. -*) - -(* -Lemma sorted_rconsE (leT : rel A) xs x : - transitive leT -> - sorted leT (rcons xs x) = all (leT^~ x) xs && sorted leT xs. -Proof. -move/rev_trans=>Ht; rewrite -(revK (rcons _ _)) rev_rcons rev_sorted /=. -by rewrite path_sortedE // all_rev rev_sorted. -Qed. -*) - -(* -Lemma sorted1 (r : rel A) xs : size xs == 1 -> sorted r xs. -Proof. by case: xs=>// x; case. Qed. -*) - -End Seq. - -Section SeqEq. -Variable (A : eqType). - -(* -Lemma perm_cons2 (x y : A) s : perm_eq [:: x, y & s] [:: y, x & s]. -Proof. -by rewrite (_ : [::x,y & s] = [::x] ++ [::y] ++ s) // - (_ : [::y,x & s] = [::y] ++ [::x] ++ s) // perm_catCA. -Qed. -*) - -(* -(* a weaker form of in_split *) -Lemma mem_split (x : A) s : - x \in s -> exists s1 s2, s = s1 ++ [:: x] ++ s2. -Proof. by case/in_split=>s1 [s2][H _]; exists s1, s2. Qed. - -Lemma mem_split_uniq (x : A) s : - x \in s -> uniq s -> - exists s1 s2, [/\ s = s1 ++ [:: x] ++ s2, - uniq (s1 ++ s2) & - x \notin s1 ++ s2]. -Proof. -move=>/[swap] Hu /mem_split [s1 [s2 H]]. -exists s1, s2; move: Hu. -by rewrite H uniq_catCA cons_uniq; case/andP. -Qed. -*) - -(* -Lemma all_notin (p : pred A) xs y : - all p xs -> ~~ p y -> y \notin xs. -Proof. by move/allP=>Ha; apply/contra/Ha. Qed. - -Lemma subset_all a (s1 s2 : seq A) : {subset s1 <= s2} -> all a s2 -> all a s1. -Proof. by move=>Hs /allP Ha1; apply/allP=>s /Hs /Ha1. Qed. -*) - -End SeqEq. - -(* -Section Allrel. -Variables (S T : Type). - -Lemma allrel_rconsl (r : T -> S -> bool) x xs ys : - allrel r (rcons xs x) ys = allrel r xs ys && all (r x) ys. -Proof. by rewrite -cats1 allrel_catl allrel1l. Qed. - -Lemma allrel_rconsr (r : T -> S -> bool) y xs ys : - allrel r xs (rcons ys y) = allrel r xs ys && all (r^~ y) xs. -Proof. by rewrite -cats1 allrel_catr allrel1r. Qed. - -End Allrel. -*) - -(* -Section AllrelEq. -Variables (S T : eqType). - -Lemma allrel_in_l (r : T -> S -> bool) (xs xs' : seq T) (ys : seq S) : - xs =i xs' -> - allrel r xs ys = allrel r xs' ys. -Proof. -by move=>H; rewrite !allrel_allpairsE; apply/eq_all_r/mem_allpairs_dep. -Qed. - -Lemma allrel_in_r (r : T -> S -> bool) (xs : seq T) (ys ys' : seq S) : - ys =i ys' -> - allrel r xs ys = allrel r xs ys'. -Proof. -by move=>H; rewrite !allrel_allpairsE; apply/eq_all_r/mem_allpairs_dep. -Qed. - -Lemma allrel_sub_l (r : T -> S -> bool) (xs xs' : seq T) (ys : seq S) : - {subset xs' <= xs} -> - allrel r xs ys -> allrel r xs' ys. -Proof. -move=>H Ha; apply/allrelP=>x y Hx Hy. -by move/allrelP: Ha; apply=>//; apply: H. -Qed. - -Lemma allrel_sub_r (r : T -> S -> bool) (xs : seq T) (ys ys' : seq S) : - {subset ys' <= ys} -> - allrel r xs ys -> allrel r xs ys'. -Proof. -move=>H Ha; apply/allrelP=>x y Hx Hy. -by move/allrelP: Ha; apply=>//; apply: H. -Qed. - -Lemma allrel_trans (xs ys : seq S) z r : - transitive r -> - all (r^~ z) xs -> all (r z) ys -> allrel r xs ys. -Proof. -move=>Ht /allP Ha /allP Hp; apply/allrelP=>x y + Hy. -by move/Ha/Ht; apply; apply: Hp. -Qed. - -End AllrelEq. -*) - -(* -Section SeqOrd. -Variable (A : ordType). - -Lemma ord_neq (x y : A) : ord x y -> x != y. -Proof. -move=>H; apply/negP=>/eqP E. -by rewrite E irr in H. -Qed. - -Lemma sorted_cat_cons_cat (l r : seq A) x : - sorted ord (l ++ x :: r) = sorted ord (l ++ [::x]) && sorted ord (x::r). -Proof. -rewrite !(sorted_pairwise (@trans A)) cats1 pairwise_cat pairwise_rcons allrel_consr !pairwise_cons. -case/boolP: (all (ord^~ x) l)=>//= Hl; case/boolP: (all (ord x) r)=>/= [Hr|_]; last by rewrite !andbF. -by rewrite (allrel_trans (@trans A) Hl Hr). -Qed. - -End SeqOrd. -*) - -(* surgery on finfuns: slicing & permuting *) - -(* -Section OnthCodom. -Variable (A : Type). - -Lemma onth_tnth {n} (s : n.-tuple A) (i : 'I_n) : onth s i = Some (tnth s i). -Proof. -elim: n s i =>[|n IH] s i; first by case: i. -case/tupleP: s=>/=x s; case: (unliftP ord0 i)=>[j|]-> /=. -- by rewrite tnthS. -by rewrite tnth0. -Qed. - -Lemma onth_codom {n} (i : 'I_n) (f: {ffun 'I_n -> A}) : onth (fgraph f) i = Some (f i). -Proof. -pose i' := cast_ord (esym (card_ord n)) i. -move: (@tnth_fgraph _ _ f i'); rewrite (enum_val_ord) {2}/i' cast_ordKV=><-. -by rewrite (onth_tnth (fgraph f) i'). -Qed. - -End OnthCodom. -*) - -(* -Section CodomWS. -Variable (n : nat) (A : Type). - -Lemma codom_ux_rcons (f : {ffun 'I_n -> A}) (i : 'I_n) : - &:(fgraph f) `]-oo, i : nat] = rcons (&:(fgraph f) `]-oo, i : nat[) (f i). -Proof. by rewrite slice_xR // slice_uu onth_codom. Qed. - -End CodomWS. -*) - -(* -Section PermFfun. -Variables (I : finType) (A : Type). - -Definition pffun (p : {perm I}) (f : {ffun I -> A}) := - [ffun i => f (p i)]. - -Lemma pffunE1 (f : {ffun I -> A}) : - pffun 1%g f = f. -Proof. by apply/ffunP=>i; rewrite !ffunE permE. Qed. - -Lemma pffunEM (p p' : {perm I}) (f : {ffun I -> A}) : - pffun (p * p') f = pffun p (pffun p' f). -Proof. by apply/ffunP => i; rewrite !ffunE permM. Qed. - -End PermFfun. -*) From b4d305346ae8c6516081f87dfaea732a2f4ad6d4 Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Thu, 1 Aug 2024 20:57:08 +0200 Subject: [PATCH 06/93] started refactoring model.v --- htt/model.v | 245 +++++++++++++++++++++++++++++++--------------------- 1 file changed, 147 insertions(+), 98 deletions(-) diff --git a/htt/model.v b/htt/model.v index c2fe817..a9146fc 100644 --- a/htt/model.v +++ b/htt/model.v @@ -1,8 +1,30 @@ +(* +Copyright 2012 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 HB Require Import structures. From mathcomp Require Import ssreflect ssrbool ssrfun ssrnat eqtype seq. From pcm Require Import options axioms pred prelude. From pcm Require Import pcm unionmap heap. From htt Require Import domain. +(*********************************************) +(* Denotational model *) +(* denoting programs as relations *) +(* over input/output heaps and output result *) +(*********************************************) + +(* Helper definitions *) + (* Exceptions are an equality type *) Inductive exn : Type := exn_from_nat of nat. @@ -15,21 +37,23 @@ Definition eqexn := Lemma eqexnP : Equality.axiom eqexn. Proof. by move=>[x][y]/=; case: eqP=>[->|*]; constructor=>//; case. Qed. -Canonical Structure exn_eqMixin := EqMixin eqexnP. -Canonical Structure exn_eqType := EqType exn exn_eqMixin. +HB.instance Definition _ := hasDecEq.Build exn eqexnP. (* Answer type *) Inductive ans (A : Type) : Type := Val of A | Exn of exn. Arguments Exn [A]. -(* A set of heaps *) -Notation pre := (Pred heap). - -(* A set of (ans A * heap) *) -(* This models the fact that programs can hang, returning nothing, *) -(* or produce nondeterministic results (e.g. alloc). *) -Notation post A := (ans A -> heap -> Prop). - +(* precondition is predicate over heaps *) +Definition pre := Pred heap. +(* postcondition is relates return result and output heap *) +(* because it's a relation, it models that programs can produce *) +(* no results (by looping forever) or produce more than one result *) +(* non-deterministically. *) +(* The latter will be used to model the alloc program *) +Definition post A := ans A -> heap -> Prop. +(* specification is a pair of precondition and postcondition *) +(* possibly parameterized by logical variable of type G (ghost) *) +(* (G may have product type, to allow for context of logical variables) *) Definition spec G A := G -> pre * post A : Type. (*************************************************************) @@ -37,12 +61,10 @@ Definition spec G A := G -> pre * post A : Type. (*************************************************************) Module Type VrfSig. - Parameter ST : Type -> Type. - Parameter ret : forall A, A -> ST A. Parameter throw : forall A, exn -> ST A. -Parameter bind : forall A B, ST A -> (A -> ST B) -> ST B. +Parameter bnd : forall A B, ST A -> (A -> ST B) -> ST B. Parameter try : forall A B, ST A -> (A -> ST B) -> (exn -> ST B) -> ST B. Parameter read : forall A, ptr -> ST A. Parameter write : forall A, ptr -> A -> ST unit. @@ -50,58 +72,62 @@ Parameter alloc : forall A, A -> ST ptr. Parameter allocb : forall A, A -> nat -> ST ptr. Parameter dealloc : ptr -> ST unit. -Arguments throw [A] e. -Arguments read [A] x. +Arguments throw {A}. +Arguments read {A}. -(* we need program to come first in the argument list - so that automation can match on it *) +(* weakest-pre transformer *) +(* the order of args is: program, heap, post *) +(* program must come first so that automation can match on it *) Parameter vrf' : forall A, ST A -> heap -> post A -> Prop. -(* recover the usual [pre]prog[post] order with a notation *) +(* recover the usual order of: pre-heap, prog, post *) +(* with a notation *) Notation vrf i e Q := (vrf' e i Q). Parameter vrfV : forall A e i (Q : post A), - (valid i -> vrf i e Q) -> vrf i e Q. + (valid i -> vrf i e Q) -> vrf i e Q. Parameter vrf_post : forall A e i (Q1 Q2 : post A), - (forall y m, valid m -> Q1 y m -> Q2 y m) -> - vrf i e Q1 -> vrf i e Q2. + (forall y m, valid m -> Q1 y m -> Q2 y m) -> + vrf i e Q1 -> vrf i e Q2. Parameter vrf_frame : forall A e i j (Q : post A), - vrf i e (fun y m => valid (m \+ j) -> Q y (m \+ j)) -> - vrf (i \+ j) e Q. + vrf i e (fun y m => valid (m \+ j) -> Q y (m \+ j)) -> + vrf (i \+ j) e Q. Parameter vrf_ret : forall A x i (Q : post A), - (valid i -> Q (Val x) i) -> vrf i (ret x) Q. + (valid i -> Q (Val x) i) -> vrf i (ret x) Q. Parameter vrf_throw : forall A e i (Q : post A), - (valid i -> Q (Exn e) i) -> vrf i (throw e) Q. -Parameter vrf_bind : forall A B (e1 : ST A) (e2 : A -> ST B) i (Q : post B), - vrf i e1 (fun x m => - match x with - | Val x' => vrf m (e2 x') Q - | Exn e => valid m -> Q (Exn e) m - end) -> - vrf i (bind e1 e2) Q. -Parameter vrf_try : forall A B (e : ST A) (e1 : A -> ST B) (e2 : exn -> ST B) i (Q : post B), - vrf i e (fun x m => - match x with - | Val x' => vrf m (e1 x') Q - | Exn ex => vrf m (e2 ex) Q - end) -> - vrf i (try e e1 e2) Q. + (valid i -> Q (Exn e) i) -> vrf i (throw e) Q. +Parameter vrf_bind : forall A B (e1 : ST A) (e2 : A -> ST B) + i (Q : post B), + vrf i e1 (fun x m => + match x with + | Val x' => vrf m (e2 x') Q + | Exn e => valid m -> Q (Exn e) m + end) -> + vrf i (bnd e1 e2) Q. +Parameter vrf_try : forall A B (e : ST A) (e1 : A -> ST B) + (e2 : exn -> ST B) i (Q : post B), + vrf i e (fun x m => + match x with + | Val x' => vrf m (e1 x') Q + | Exn ex => vrf m (e2 ex) Q + end) -> + vrf i (try e e1 e2) Q. Parameter vrf_read : forall A x j (v : A) (Q : post A), - (valid (x :-> v \+ j) -> Q (Val v) (x :-> v \+ j)) -> - vrf (x :-> v \+ j) (read x) Q. + (valid (x :-> v \+ j) -> Q (Val v) (x :-> v \+ j)) -> + vrf (x :-> v \+ j) (read x) Q. Parameter vrf_write : forall A x (v : A) B (u : B) j (Q : post unit), - (valid (x :-> v \+ j) -> Q (Val tt) (x :-> v \+ j)) -> - vrf (x :-> u \+ j) (write x v) Q. + (valid (x :-> v \+ j) -> Q (Val tt) (x :-> v \+ j)) -> + vrf (x :-> u \+ j) (write x v) Q. Parameter vrf_alloc : forall A (v : A) i (Q : post ptr), - (forall x, valid (x :-> v \+ i) -> Q (Val x) (x :-> v \+ i)) -> - vrf i (alloc v) Q. + (forall x, valid (x :-> v \+ i) -> Q (Val x) (x :-> v \+ i)) -> + vrf i (alloc v) Q. Parameter vrf_allocb : forall A (v : A) n i (Q : post ptr), - (forall x, valid (updi x (nseq n v) \+ i) -> - Q (Val x) (updi x (nseq n v) \+ i)) -> - vrf i (allocb v n) Q. + (forall x, valid (updi x (nseq n v) \+ i) -> + Q (Val x) (updi x (nseq n v) \+ i)) -> + vrf i (allocb v n) Q. Parameter vrf_dealloc : forall x A (v : A) j (Q : post unit), - (x \notin dom j -> valid j -> Q (Val tt) j) -> - vrf (x :-> v \+ j) (dealloc x) Q. + (x \notin dom j -> valid j -> Q (Val tt) j) -> + vrf (x :-> v \+ j) (dealloc x) Q. Definition has_spec G A (s : spec G A) (e : ST A) := forall g i, (s g).1 i -> vrf i e (s g).2. @@ -114,17 +140,31 @@ Arguments STspec G [A] s. Notation "'Do' e" := (@STprog _ _ _ e _) (at level 80). -Notation "x '<--' c1 ';' c2" := (bind c1 (fun x => c2)) - (at level 81, right associativity). -Notation "c1 ';;' c2" := (bind c1 (fun _ => c2)) - (at level 81, right associativity). +(* some notation *) + +Notation "x '<--' c1 ';' c2" := (bnd c1 (fun x => c2)) + (at level 201, right associativity, format + "'[v' x '<--' c1 ';' '//' c2 ']'"). +Notation "' x '<--' c1 ';' c2" := (bnd c1 (fun x => c2)) + (at level 201, right associativity, x strict pattern, format + "'[v' ' x '<--' c1 ';' '//' c2 ']'"). +Notation "c1 ';;' c2" := (bnd c1 (fun _ => c2)) + (at level 201, right associativity). Notation "'!' x" := (read x) (at level 50). Notation "x '::=' e" := (write x e) (at level 60). +(* Fixed point constructor *) +(* We shall make fix work over *monotone closure* of argument function. *) +(* This is a mathematical hack to avoid spurious monotonicity proofs *) +(* that's necessary in shallow embedding, as we have no guarantee *) +(* that the argument function is actually monotone. *) +(* We *do* prove later on that all constructors are monotone. *) +(* We also hide model definition so that all functions one can *) +(* apply ffix to will be built out of monotonicity-preserving *) +(* constructors (and hence will be monotone). *) Parameter Fix : forall G A (B : A -> Type) (s : forall x : A, spec G (B x)), ((forall x : A, STspec G (s x)) -> forall x : A, STspec G (s x)) -> forall x : A, STspec G (s x). - End VrfSig. @@ -137,8 +177,9 @@ Module Vrf : VrfSig. Section BasePrograms. Variables (P : pre) (A : Type). -(* we carve out the model out of the following base type *) -Definition prog : Type := forall i : heap, valid i -> i \In P -> post A. +(* the model is carved out of the following base type *) +Definition prog : Type := + forall i : heap, valid i -> i \In P -> post A. (* we take only preconditions and progs with special properties *) (* which we define next *) @@ -149,7 +190,7 @@ Definition safe_mono := (* defined heaps map to defined heaps *) Definition def_strict (e : prog) := - forall i p v x, Heap.Undef \Notin e i v p x. + forall i p v x, undef \Notin e i v p x. (* frame property *) Definition frameable (e : prog) := @@ -159,35 +200,44 @@ Definition frameable (e : prog) := End BasePrograms. -Section STDef. -Variable (A : Type). - -Structure ST' := Prog { +(* the model considers programs of base type, given a precondition *) +(* and adds several properties: *) +(* - safety monotonicity *) +(* - strictness of definedness *) +(* - frameability *) +Structure ST' (A : Type) := Prog { pre_of : pre; prog_of : prog pre_of A; _ : safe_mono pre_of; _ : def_strict prog_of; _ : frameable prog_of}. +Arguments prog_of {A}. + (* module field must be a definition, not structure *) Definition ST := ST'. -Lemma sfm_st e : safe_mono (pre_of e). -Proof. by case: e. Qed. +(* projections *) -Arguments prog_of : clear implicits. +Lemma sfm_st A (e : ST A) : safe_mono (pre_of e). +Proof. by case: e. Qed. -Lemma dstr_st e : def_strict (prog_of e). +Lemma dstr_st A (e : ST A) : def_strict (prog_of e). Proof. by case: e. Qed. -Corollary dstr_valid e i p v x m : - m \In prog_of e i p v x -> valid m. +Lemma dstr_valid A (e : ST A) i p v x m : + m \In prog_of e i p v x -> + valid m. Proof. by case: m=>// /dstr_st. Qed. -Lemma fr_st e : frameable (prog_of e). +Lemma fr_st A (e : ST A) : frameable (prog_of e). Proof. by case: e. Qed. -Arguments fr_st [e i j]. +Arguments fr_st {A e i j}. + +Section STDef. +Variable A : Type. +Implicit Type e : ST A. (* poset structure on ST *) @@ -196,30 +246,24 @@ Definition st_leq e1 e2 := forall i (v : valid i) (p : i \In pre_of e2), prog_of e1 _ v (pf _ p) <== prog_of e2 _ v p. -Lemma st_refl e : st_leq e e. -Proof. -exists (poset_refl _)=>i V P y m. -by rewrite (pf_irr (poset_refl (pre_of e) i P) P). -Qed. - -Lemma st_asym e1 e2 : st_leq e1 e2 -> st_leq e2 e1 -> e1 = e2. +Lemma st_is_poset : poset_axiom st_leq. Proof. -move: e1 e2=>[p1 e1 S1 D1 F1][p2 e2 S2 D2 F2]; rewrite /st_leq /=. -case=>E1 R1 [E2 R2]. -move: (poset_asym E1 E2)=>?; subst p2. -have : e1 = e2. -- apply: fext=>i; apply: fext=>Vi; apply: fext=>Pi; apply: fext=>y; apply: fext=>m. - move: (R2 i Vi Pi y m)=>{}R2; move: (R1 i Vi Pi y m)=>{}R1. - apply: pext; split. - - by move=>H1; apply: R1; rewrite (pf_irr (E1 i Pi) Pi). - by move=>H2; apply: R2; rewrite (pf_irr (E2 i Pi) Pi). -move=>?; subst e2. -by congr Prog; apply: pf_irr. -Qed. - -Lemma st_trans e1 e2 e3 : st_leq e1 e2 -> st_leq e2 e3 -> st_leq e1 e3. -Proof. -move: e1 e2 e3=>[p1 e1 S1 D1 F1][p2 e2 S2 D2 F2][p3 e3 S3 D3 F3]. +split=>[e|e1 e2|e1 e2 e3]. +- exists (poset_refl _)=>i V P y m. + by rewrite (pf_irr (poset_refl (pre_of e) i P) P). +- case: e1 e2=>p1 e1 S1 D1 F1 [p2 e2 S2 D2 F2]. + rewrite /st_leq /=; case=>E1 R1 [E2 R2]. + move: (poset_asym E1 E2)=>?; subst p2. + have : e1 = e2. + - apply: fext=>i; apply: fext=>Vi; apply: fext=>Pi. + apply: fext=>y; apply: fext=>m. + move: (R2 i Vi Pi y m)=>{}R2; move: (R1 i Vi Pi y m)=>{}R1. + apply: pext; split. + - by move=>H1; apply: R1; rewrite (pf_irr (E1 i Pi) Pi). + by move=>H2; apply: R2; rewrite (pf_irr (E2 i Pi) Pi). + move=>?; subst e2. + by congr Prog; apply: pf_irr. +case: e1 e2 e3=>p1 e1 S1 D1 F1 [p2 e2 S2 D2 F2][p3 e3 S3 D3 F3]. case=>/= E1 R1 [/= E2 R2]; rewrite /st_leq /=. have E3 := poset_trans E2 E1; exists E3=>i V P y m. set P' := E2 i P. @@ -228,7 +272,10 @@ move=>H1; apply/R2/R1. by rewrite (pf_irr (E1 i P') (E3 i P)). Qed. -(* bottom is a program that can always run but never returns (an endless loop) *) +HB.instance Definition _ := isPoset.Build (ST A) st_is_poset. + +(* bottom is a program that can always run *) +(* but never returns (infinite loop) *) Definition pre_bot : pre := top. @@ -247,13 +294,15 @@ Proof. by []. Qed. Definition st_bot := Prog sfmono_bot dstrict_bot frame_bot. Lemma st_botP e : st_leq st_bot e. -Proof. by case: e=>p e S D F; exists (@pred_topP _ _)=>???; apply: botP. Qed. - -Definition stPosetMixin := PosetMixin st_botP st_refl st_asym st_trans. -Canonical stPoset := Eval hnf in Poset ST stPosetMixin. +Proof. by case: e=>p e S D F; exists (fun _ _ => I)=>???; apply: botP. Qed. (* lattice structure on ST *) +Here need that bot is bottom. + + + + (* intersection of preconditions *) Definition pre_sup (u : Pred ST) : pre := fun h => forall e, e \In u -> h \In pre_of e. From daf3122e929398c923f1049332fe58dc1d80b630 Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Thu, 1 Aug 2024 21:04:16 +0200 Subject: [PATCH 07/93] minor --- htt/model.v | 44 ++++++++++++++++++++------------------------ 1 file changed, 20 insertions(+), 24 deletions(-) diff --git a/htt/model.v b/htt/model.v index a9146fc..f3f1c23 100644 --- a/htt/model.v +++ b/htt/model.v @@ -205,12 +205,14 @@ End BasePrograms. (* - safety monotonicity *) (* - strictness of definedness *) (* - frameability *) + +Definition st_axiom A (p : pre) (e : prog p A) := + [/\ safe_mono p, def_strict e & frameable e]. + Structure ST' (A : Type) := Prog { pre_of : pre; prog_of : prog pre_of A; - _ : safe_mono pre_of; - _ : def_strict prog_of; - _ : frameable prog_of}. + _ : st_axiom prog_of}. Arguments prog_of {A}. @@ -220,10 +222,10 @@ Definition ST := ST'. (* projections *) Lemma sfm_st A (e : ST A) : safe_mono (pre_of e). -Proof. by case: e. Qed. +Proof. by case: e=>?? []. Qed. Lemma dstr_st A (e : ST A) : def_strict (prog_of e). -Proof. by case: e. Qed. +Proof. by case: e=>?? []. Qed. Lemma dstr_valid A (e : ST A) i p v x m : m \In prog_of e i p v x -> @@ -231,10 +233,11 @@ Lemma dstr_valid A (e : ST A) i p v x m : Proof. by case: m=>// /dstr_st. Qed. Lemma fr_st A (e : ST A) : frameable (prog_of e). -Proof. by case: e. Qed. +Proof. by case: e=>?? []. Qed. Arguments fr_st {A e i j}. + Section STDef. Variable A : Type. Implicit Type e : ST A. @@ -251,7 +254,7 @@ Proof. split=>[e|e1 e2|e1 e2 e3]. - exists (poset_refl _)=>i V P y m. by rewrite (pf_irr (poset_refl (pre_of e) i P) P). -- case: e1 e2=>p1 e1 S1 D1 F1 [p2 e2 S2 D2 F2]. +- case: e1 e2=>p1 e1 [S1 D1 F1][p2 e2 [S2 D2 F2]]. rewrite /st_leq /=; case=>E1 R1 [E2 R2]. move: (poset_asym E1 E2)=>?; subst p2. have : e1 = e2. @@ -263,7 +266,7 @@ split=>[e|e1 e2|e1 e2 e3]. by move=>H2; apply: R2; rewrite (pf_irr (E2 i Pi) Pi). move=>?; subst e2. by congr Prog; apply: pf_irr. -case: e1 e2 e3=>p1 e1 S1 D1 F1 [p2 e2 S2 D2 F2][p3 e3 S3 D3 F3]. +case: e1 e2 e3=>p1 e1 [S1 D1 F1][p2 e2 [S2 D2 F2]][p3 e3 [S3 D3 F3]]. case=>/= E1 R1 [/= E2 R2]; rewrite /st_leq /=. have E3 := poset_trans E2 E1; exists E3=>i V P y m. set P' := E2 i P. @@ -277,28 +280,21 @@ HB.instance Definition _ := isPoset.Build (ST A) st_is_poset. (* bottom is a program that can always run *) (* but never returns (infinite loop) *) -Definition pre_bot : pre := top. - -Definition prog_bot : prog pre_bot A := - fun _ _ _ _ => bot. +(* DO WE NEED THIS? If we have lattice sturcture *) +(* existence of bot follows *) -Lemma sfmono_bot : safe_mono pre_bot. -Proof. by []. Qed. - -Lemma dstrict_bot : def_strict prog_bot. -Proof. by move=>*. Qed. - -Lemma frame_bot : frameable prog_bot. -Proof. by []. Qed. - -Definition st_bot := Prog sfmono_bot dstrict_bot frame_bot. +Definition pre_bot : pre := top. +Definition prog_bot : prog pre_bot A := fun _ _ _ _ => bot. +Lemma progbot_is_st : st_axiom prog_bot. +Proof. by split=>//; move=>*. Qed. +Definition st_bot := Prog progbot_is_st. Lemma st_botP e : st_leq st_bot e. -Proof. by case: e=>p e S D F; exists (fun _ _ => I)=>???; apply: botP. Qed. +Proof. by case: e=>p e [???]; exists (fun _ _ => I)=>???; apply: botP. Qed. (* lattice structure on ST *) -Here need that bot is bottom. +UP TO HERE From 5af642058685dcf96c7bd4756d3817e436f8d2eb Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Fri, 2 Aug 2024 15:51:32 +0200 Subject: [PATCH 08/93] started refactoring examples --- core/pred.v | 1 + examples/array.v | 43 ++- examples/congmath.v | 0 examples/counter.v | 0 examples/exploit.v | 13 + examples/gcd.v | 25 +- examples/hashtab.v | 0 examples/kvmaps.v | 0 examples/llist.v | 121 ++++--- examples/queue.v | 0 examples/stack.v | 0 htt/domain.v | 3 +- htt/heapauto.v | 154 ++++----- htt/model.v | 759 +++++++++++++++++++------------------------- 14 files changed, 546 insertions(+), 573 deletions(-) mode change 100755 => 100644 examples/congmath.v mode change 100755 => 100644 examples/counter.v mode change 100755 => 100644 examples/exploit.v mode change 100755 => 100644 examples/hashtab.v mode change 100755 => 100644 examples/kvmaps.v mode change 100755 => 100644 examples/queue.v mode change 100755 => 100644 examples/stack.v diff --git a/core/pred.v b/core/pred.v index 0d458b0..fc711ee 100644 --- a/core/pred.v +++ b/core/pred.v @@ -39,6 +39,7 @@ 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. +Open Scope fun_scope. (**************************************************************************) (* We follow ssrbool, and provide four different types of predicates. *) diff --git a/examples/array.v b/examples/array.v index 42cbe10..61c7de4 100644 --- a/examples/array.v +++ b/examples/array.v @@ -1,3 +1,16 @@ +(* +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. From mathcomp Require Import ssrnat eqtype seq path fintype tuple finfun finset. From pcm Require Import options axioms prelude pred. @@ -6,8 +19,6 @@ From htt Require Import options model heapauto. Definition indx {I : finType} (x : I) := index x (enum I). -Prenex Implicits indx. - (***********************************) (* Arrays indexed by a finite type *) (***********************************) @@ -15,7 +26,7 @@ Prenex Implicits indx. (* an array is a pointer to a contiguous memory region holding its values *) Record array (I : finType) (T : Type) : Type := Array {orig :> ptr}. -Arguments Array [I T]. +Arguments Array {I T}. Definition array_for (I : finType) (T : Type) of phant (I -> T) := array I T. Identity Coercion array_for_array : array_for >-> array. @@ -70,11 +81,11 @@ Program Definition newf (f : {ffun I -> T}) : (* preallocate a new array *) x <-- new (f v); (* fill it with values of f on I *) - let fill := Fix (fun (loop : fill_loop x f) s => - Do (if s is k::t return _ then + let fill := ffix (fun (loop : fill_loop x f) s => + Do (if s is k::t return _ then x .+ (indx k) ::= f k;; loop t - else ret (Array x))) + else ret (Array x))) in fill (enum I) else ret (Array null)). (* first the loop *) @@ -124,10 +135,10 @@ Definition free_loop (a : array) : Type := Program Definition free (a : array) : STsep (fun i => exists f, i \In shape a f, [vfun _ : unit => emp]) := - Do (let go := Fix (fun (loop : free_loop a) k => - Do (if k == #|I| then ret tt - else dealloc a.+k;; - loop k.+1)) + Do (let go := ffix (fun (loop : free_loop a) k => + Do (if k == #|I| then ret tt + else dealloc a.+k;; + loop k.+1)) in go 0). (* first the loop *) Next Obligation. @@ -136,8 +147,8 @@ move=>a loop k [] i /= [[|v xs]][->] /= _; first by rewrite add0n=>/eqP ->; step (* the suffix is non-empty so k < #|I| *) case: eqP=>[->|_ H]; first by move/eqP; rewrite -{2}(add0n #|I|) eqn_add2r. (* run the program, simplify *) -step; apply: vrfV=>V; apply: [gE]=>//=. -by exists xs; rewrite V ptrA addn1 -addSnnS unitL. +step; apply: vrfV=>V; apply: [gE]=>//=; exists xs. +by rewrite V unitL -addSnnS H -addnS. Qed. (* now the outer program *) Next Obligation. @@ -151,8 +162,8 @@ Qed. (* reading from an array, doesn't modify the heap *) Program Definition read (a : array) (k : I) : - {f h}, STsep (fun i => i = h /\ i \In shape a f, - [vfun (y : T) m => m = h /\ y = f k]) := + STsep {f h} (fun i => i = h /\ i \In shape a f, + [vfun (y : T) m => m = h /\ y = f k]) := Do (!a .+ (indx k)). Next Obligation. (* pull out ghost vars *) @@ -164,8 +175,8 @@ Qed. (* writing to an array, updates the spec function with a new value *) Program Definition write (a : array) (k : I) (x : T) : - {f}, STsep (shape a f, - [vfun _ : unit => shape a (finfun [eta f with k |-> x])]) := + STsep {f} (shape a f, + [vfun _ : unit => shape a (finfun [eta f with k |-> x])]) := Do (a .+ (indx k) ::= x). Next Obligation. (* pull out ghost vars, split the heap *) diff --git a/examples/congmath.v b/examples/congmath.v old mode 100755 new mode 100644 diff --git a/examples/counter.v b/examples/counter.v old mode 100755 new mode 100644 diff --git a/examples/exploit.v b/examples/exploit.v old mode 100755 new mode 100644 index 0a6c959..f42ffa3 --- a/examples/exploit.v +++ b/examples/exploit.v @@ -1,3 +1,16 @@ +(* +Copyright 2009 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 Logic.Hurkens. (* This file shows the unsoundness of the axiom pack_injective assumed in *) diff --git a/examples/gcd.v b/examples/gcd.v index bd85688..9d718c9 100644 --- a/examples/gcd.v +++ b/examples/gcd.v @@ -1,6 +1,19 @@ +(* +Copyright 2021 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 div. -From pcm Require Import axioms pred pcm heap. +From pcm Require Import axioms pred ordtype pcm heap. From htt Require Import options model heapauto. (* classical mutable Euclid's algorithm for computing GCD with subtractions *) @@ -13,12 +26,12 @@ Definition shape (p q : ptr) (x y : nat) := (* (`unit` is needed because `Fix` always requires a single parameter) *) Definition gcd_loopT (p q : ptr) : Type := unit -> - {x y : nat}, STsep (shape p q x y, + STsep {x y : nat} (shape p q x y, [vfun (_ : unit) h => h \In shape p q (gcdn x y) (gcdn x y)]). Program Definition gcd_loop (p q : ptr) := - Fix (fun (go : gcd_loopT p q) _ => + ffix (fun (go : gcd_loopT p q) _ => Do (x <-- !p; y <-- !q; if (x : nat) != y then @@ -48,11 +61,11 @@ suff {p q go m}: gcdn x (y - x) = gcdn x y by move=>->. by rewrite -gcdnDr subnK //; apply: ltnW. Qed. -(* note that there's no guarantee on termination, so we have only partial correctness *) -(* i.e. the algorithm will get stuck if u = 0 /\ v != 0 or vice versa *) +(* There's no guarantee on termination, as this is partial correctness. *) +(* The algorithm loops if u = 0 /\ v != 0 or vice versa *) Program Definition gcd u v : STsep (PredT, [vfun r _ => r = gcdn u v]) := - (* we allocate in the reverse order because the symbolic heap behaves as a stack *) + (* allocate in the reverse order because the symbolic heap behaves as a stack *) (* this way it'll match the specification perfectly, removing a bit of bureaucracy *) Do (q <-- alloc v; p <-- alloc u; diff --git a/examples/hashtab.v b/examples/hashtab.v old mode 100755 new mode 100644 diff --git a/examples/kvmaps.v b/examples/kvmaps.v old mode 100755 new mode 100644 diff --git a/examples/llist.v b/examples/llist.v index 87eca58..6c6f0ab 100644 --- a/examples/llist.v +++ b/examples/llist.v @@ -1,3 +1,15 @@ +(* +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 mathcomp Require Import ssreflect ssrbool ssrnat eqtype ssrfun seq. From pcm Require Import options axioms pred. @@ -11,7 +23,7 @@ From htt Require Import options model heapauto. Fixpoint lseg {A} (p q : ptr) (xs : seq A) := if xs is hd::tl then [Pred h | exists r h', - h = p :-> hd \+ (p .+ 1 :-> r \+ h') /\ h' \In lseg r q tl] + h = p :-> hd \+ (p.+1 :-> r \+ h') /\ h' \In lseg r q tl] else [Pred h | p = q /\ h = Unit]. Definition EmptyList : exn := exn_from_nat 15. @@ -23,22 +35,24 @@ Variable A : Type. Lemma lseg_rcons (xs : seq A) x p r h : h \In lseg p r (rcons xs x) <-> - exists q h', h = h' \+ (q :-> x \+ q .+ 1 :-> r) /\ h' \In lseg p q xs. + exists q h', h = h' \+ (q :-> x \+ q.+1 :-> r) /\ + h' \In lseg p q xs. Proof. move: xs x p r h; elim=>[|x xs IH] y p r h /=. - by split; case=>x [h'][->][<- ->]; [exists p | exists r]; hhauto. split. - case=>z [h1][->]; case/IH=>w [h2][->] H1. - by exists w, (p :-> x \+ (p .+ 1 :-> z \+ h2)); hhauto. + by exists w, (p :-> x \+ (p.+1 :-> z \+ h2)); hhauto. case=>q [h1][->][z][h2][->] H1. -exists z, (h2 \+ q :-> y \+ q .+ 1 :-> r). +exists z, (h2 \+ q :-> y \+ q.+1 :-> r). by rewrite -!joinA; split=>//; apply/IH; eauto. Qed. (* null pointer represents an empty segment *) Lemma lseg_null (xs : seq A) q h : - valid h -> h \In lseg null q xs -> + valid h -> + h \In lseg null q xs -> [/\ q = null, xs = [::] & h = Unit]. Proof. case: xs=>[|x xs] D /= H; first by case: H=><- ->. @@ -47,9 +61,11 @@ Qed. (* empty heap represents an empty segment *) -Lemma lseg_empty (xs : seq A) p q : Unit \In lseg p q xs -> p = q /\ xs = [::]. +Lemma lseg_empty (xs : seq A) p q : + Unit \In lseg p q xs -> + p = q /\ xs = [::]. Proof. -by case: xs=>[|x xs][] // r [h][/esym/join0E][/unitbE]; rewrite /heap_pts ptsU um_unitbU. +by case: xs=>[|x xs][] //= r [h][/esym/umap0E][/unitbP]; rewrite um_unitbU. Qed. (* reformulation of the specification *) @@ -59,7 +75,7 @@ Lemma lseg_case (xs : seq A) p q h : [/\ p = q, xs = [::] & h = Unit] \/ exists x r h', [/\ xs = x :: behead xs, - h = p :-> x \+ (p .+ 1 :-> r \+ h') & + h = p :-> x \+ (p.+1 :-> r \+ h') & h' \In lseg r q (behead xs)]. Proof. case: xs=>[|x xs] /=; first by case=>->->; left. @@ -72,7 +88,7 @@ Corollary lseg_neq (xs : seq A) p q h : p != q -> h \In lseg p q xs -> exists x r h', [/\ xs = x :: behead xs, - h = p :-> x \+ (p .+ 1 :-> r \+ h') & + h = p :-> x \+ (p.+1 :-> r \+ h') & h' \In lseg r q (behead xs)]. Proof. move=>H /lseg_case; case=>//; case=>E. @@ -81,11 +97,11 @@ Qed. (* non-empty list is represented by a non-trivial segment *) -Corollary lseg_lt0n (xs : seq A) p q h : +Lemma lseg_lt0n (xs : seq A) p q h : 0 < size xs -> h \In lseg p q xs -> exists x r h', [/\ xs = x :: behead xs, - h = p :-> x \+ (p .+ 1 :-> r \+ h') & + h = p :-> x \+ (p.+1 :-> r \+ h') & h' \In lseg r q (behead xs)]. Proof. move=>H /lseg_case; case=>//; case=>_ E. @@ -103,7 +119,7 @@ elim: xs h p=>/=. by case=>j[_][h2][{h}-> [->->]]; rewrite unitL. move=>x xs IH h p; split. - case=>r[_][{h}-> /IH][j][h1][h2][-> H1 H2]. - exists j, (p :-> x \+ p.+ 1 :-> r \+ h1), h2; rewrite !joinA; split=>//. + exists j, (p :-> x \+ p.+1 :-> r \+ h1), h2; rewrite !joinA; split=>//. by exists r, h1; rewrite joinA. case=>j[_][h2][{h}-> [r][h1][-> H1 H2]]. exists r, (h1 \+ h2); rewrite !joinA; split=>//. @@ -121,21 +137,26 @@ Variable (A : Type). (* specializing the null and neq lemmas *) Lemma lseq_null (xs : seq A) h : - valid h -> h \In lseq null xs -> xs = [::] /\ h = Unit. + valid h -> + h \In lseq null xs -> + xs = [::] /\ h = Unit. Proof. by move=>D; case/(lseg_null D)=>_ ->. Qed. Lemma lseq_pos (xs : seq A) p h : - p != null -> h \In lseq p xs -> + p != null -> + h \In lseq p xs -> exists x r h', [/\ xs = x :: behead xs, - h = p :-> x \+ (p .+ 1 :-> r \+ h') & + h = p :-> x \+ (p.+1 :-> r \+ h') & h' \In lseq r (behead xs)]. Proof. by apply: lseg_neq. Qed. (* a valid heap cannot match two different specs *) Lemma lseq_func (l1 l2 : seq A) p h : - valid h -> h \In lseq p l1 -> h \In lseq p l2 -> l1 = l2. + valid h -> + h \In lseq p l1 -> + h \In lseq p l2 -> l1 = l2. Proof. elim: l1 l2 p h => [|x1 xt IH] /= l2 p h V. - by case=>->->; case/lseq_null. @@ -156,7 +177,7 @@ Next Obligation. by move=>[] /= _ ->; step. Qed. (* prepending a value *) Program Definition insert p (x : A) : - {l}, STsep (lseq p l, [vfun p' => lseq p' (x::l)]) := + STsep {l} (lseq p l, [vfun p' => lseq p' (x::l)]) := Do (q <-- allocb p 2; q ::= x;; ret q). @@ -164,14 +185,14 @@ Next Obligation. (* pull out ghost var + precondition, run the first step *) move=>p x [l][] i /= H; step=>q. (* run the last 2 steps, guess the final pointer and heap from the goal *) -rewrite unitR -joinA; heval. +by rewrite unitR -joinA; heval. Qed. (* getting the head element *) (* an example of a partial program, doesn't modify the heap *) Program Definition head p : - {l}, STsep (lseq p l, + STsep {l} (lseq p l, fun (y : ans A) h => h \In lseq p l /\ match y with Val v => l = v :: behead l | Exn e => e = EmptyList /\ l = [::] end) := @@ -193,11 +214,11 @@ Qed. (* removing the head element, no-op for an empty list *) Program Definition remove p : - {xs : seq A}, STsep (lseq p xs, [vfun p' => lseq p' (behead xs)]) := + STsep {xs : seq A} (lseq p xs, [vfun p' => lseq p' (behead xs)]) := Do (if p == null then ret p - else pnext <-- !(p .+ 1); + else pnext <-- !p.+1; dealloc p;; - dealloc p .+ 1;; + dealloc p.+1;; ret pnext). Next Obligation. (* pull out ghost + precondition, branch *) @@ -215,16 +236,16 @@ Qed. (* 1. heap is unchanged *) (* 2. total length is accumulator + the length of unprocessed list *) Definition lenT : Type := forall (pl : ptr * nat), - {xs : seq A}, STsep (lseq pl.1 xs, + STsep {xs : seq A} (lseq pl.1 xs, [vfun l h => l == pl.2 + length xs /\ lseq pl.1 xs h]). Program Definition len p : - {xs : seq A}, STsep (lseq p xs, + STsep {xs : seq A} (lseq p xs, [vfun l h => l == length xs /\ lseq p xs h]) := - Do (let len := Fix (fun (go : lenT) '(p, l) => - Do (if p == null then ret l - else pnext <-- !(p .+ 1); - go (pnext, l + 1))) + Do (let len := ffix (fun (go : lenT) '(p, l) => + Do (if p == null then ret l + else pnext <-- !p.+1; + go (pnext, l + 1))) in len (p, 0)). (* first, the loop *) Next Obligation. @@ -247,19 +268,19 @@ Qed. (* the loop invariant: *) (* the first list should not be empty and not overlap the second *) Definition catT (p2 : ptr) : Type := - forall (p1 : ptr), {xs1 xs2 : seq A}, - STsep (fun h => p1 != null /\ (lseq p1 xs1 # lseq p2 xs2) h, - [vfun _ : unit => lseq p1 (xs1 ++ xs2)]). + forall (p1 : ptr), STsep {xs1 xs2 : seq A} + (fun h => p1 != null /\ (lseq p1 xs1 # lseq p2 xs2) h, + [vfun _ : unit => lseq p1 (xs1 ++ xs2)]). Program Definition concat p1 p2 : - {xs1 xs2 : seq A}, STsep (lseq p1 xs1 # lseq p2 xs2, + STsep {xs1 xs2 : seq A} (lseq p1 xs1 # lseq p2 xs2, [vfun a => lseq a (xs1 ++ xs2)]) := - Do (let cat := Fix (fun (go : catT p2) q => - Do (next <-- !(q .+ 1); - if (next : ptr) == null - then q .+ 1 ::= p2;; - ret tt - else go next)) + Do (let cat := ffix (fun (go : catT p2) q => + Do (next <-- !q.+1; + if (next : ptr) == null + then q.+1 ::= p2;; + ret tt + else go next)) in if p1 == null then ret p2 else cat p1;; @@ -303,16 +324,16 @@ Qed. (* 1. the processed and remaining parts should not overlap *) (* 2. the result is processed part + a reversal of remainder *) Definition revT : Type := forall (p : ptr * ptr), - {i done : seq A}, STsep (lseq p.1 i # lseq p.2 done, + STsep {i done : seq A} (lseq p.1 i # lseq p.2 done, [vfun y => lseq y (catrev i done)]). Program Definition reverse p : - {xs : seq A}, STsep (lseq p xs, [vfun p' => lseq p' (rev xs)]) := - Do (let reverse := Fix (fun (go : revT) '(i, done) => - Do (if i == null then ret done - else next <-- !i .+ 1; - i .+ 1 ::= done;; - go (next, i))) + STsep {xs : seq A} (lseq p xs, [vfun p' => lseq p' (rev xs)]) := + Do (let reverse := ffix (fun (go : revT) '(i, done) => + Do (if i == null then ret done + else next <-- !i.+1; + i.+1 ::= done;; + go (next, i))) in reverse (p, null)). (* first, the loop *) Next Obligation. @@ -341,16 +362,16 @@ Variable B : Type. (* the result is a mapped list *) Definition lmapT (f : A -> B) := forall (p : ptr), - {xs : seq A}, STsep (lseq p xs, + STsep {xs : seq A} (lseq p xs, [vfun _ : unit => lseq p (map f xs)]). Program Definition lmap f : lmapT f := - Fix (fun (lmap : lmapT f) p => + ffix (fun (lmap : lmapT f) p => Do (if p == null then ret tt else t <-- !p; p ::= f t;; - nxt <-- !p .+ 1; + nxt <-- !p.+1; lmap nxt)). Next Obligation. (* pull out ghost + precondition, branch *) @@ -365,5 +386,5 @@ Qed. End LList. -Arguments head {A} p. -Arguments remove {A} p. +Arguments head {A}. +Arguments remove {A}. diff --git a/examples/queue.v b/examples/queue.v old mode 100755 new mode 100644 diff --git a/examples/stack.v b/examples/stack.v old mode 100755 new mode 100644 diff --git a/htt/domain.v b/htt/domain.v index 75d80de..8a14f09 100644 --- a/htt/domain.v +++ b/htt/domain.v @@ -227,7 +227,8 @@ End Repack. Section Infimum. Variable T : lattice. -Definition inf (s : Pred T) := sup [Pred x : T | forall y, y \In s -> x <== y]. +Definition inf (s : Pred T) := + sup [Pred x : T | forall y, y \In s -> x <== y]. Lemma infP s x : x \In s -> inf s <== x. Proof. by move=>H; apply: supM=>y; apply. Qed. diff --git a/htt/heapauto.v b/htt/heapauto.v index 994935c..8503353 100644 --- a/htt/heapauto.v +++ b/htt/heapauto.v @@ -1,3 +1,16 @@ +(* +Copyright 2012 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 mathcomp Require Import ssreflect ssrbool ssrnat eqtype seq ssrfun. From pcm Require Import options axioms pred prelude. From pcm Require Import pcm autopcm unionmap heap. @@ -23,7 +36,7 @@ Import Prenex Implicits. (* Currently, it doesn't deal with weak pointers. I.e. if it sees terms *) (* like x :-> v1 and x :-> v2, it will reduce to v1 = v2 only if v1, v2 *) (* are of the same type. A more general tactic would emit obligation *) -(* dyn v1 = dyn v2, but I don't bother with this now. *) +(* dyn v1 = dyn v2, but that's left for future work. *) (* *) (**************************************************************************) @@ -108,8 +121,8 @@ Lemma try_retR e1 e2 (v : A) i (r : post B) : Proof. by move=>H; apply/vrf_try/val_retR. Qed. Lemma bnd_retR e (v : A) i (r : post B) : - vrf i (e v) r -> vrf i (bind (ret v) e) r. -Proof. by move=>H; apply/vrf_bind/val_retR. Qed. + vrf i (e v) r -> vrf i (bnd (ret v) e) r. +Proof. by move=>H; apply/vrf_bnd/val_retR. Qed. End EvalRetR. @@ -126,8 +139,8 @@ Proof. by move=>H; apply/vrf_try/val_throwR. Qed. Lemma bnd_throwR e e1 i (r : post B) : (valid i -> r (Exn e) i) -> - vrf i (bind (@throw A e) e1) r. -Proof. by move=>H; apply/vrf_bind/val_throwR. Qed. + vrf i (bnd (@throw A e) e1) r. +Proof. by move=>H; apply/vrf_bnd/val_throwR. Qed. End EvalThrowR. @@ -147,8 +160,8 @@ Proof. by move=>H; apply/vrf_try/val_readR. Qed. Lemma bnd_readR e v x i (f : form (x :-> v) i) (r : post B) : vrf f (e v) r -> - vrf f (bind (@read A x) e) r. -Proof. by move=>H; apply/vrf_bind/val_readR. Qed. + vrf f (bnd (@read A x) e) r. +Proof. by move=>H; apply/vrf_bnd/val_readR. Qed. End EvalReadR. @@ -169,8 +182,8 @@ Proof. by move=>H; apply/vrf_try/val_writeR. Qed. Lemma bnd_writeR e (v : A) (w : B) x i (f : forall k, form k i) (r : post C) : vrf (f (x :-> v)) (e tt) r -> - vrf (f (x :-> w)) (bind (write x v) e) r. -Proof. by move=>H; apply/vrf_bind/val_writeR. Qed. + vrf (f (x :-> w)) (bnd (write x v) e) r. +Proof. by move=>H; apply/vrf_bnd/val_writeR. Qed. End EvalWriteR. @@ -187,8 +200,8 @@ Proof. by move=>H; apply/vrf_try/val_allocR. Qed. Lemma bnd_allocR e (v : A) i (r : post B) : (forall x, vrf (x :-> v \+ i) (e x) r) -> - vrf i (bind (alloc v) e) r. -Proof. by move=>H; apply/vrf_bind/val_allocR. Qed. + vrf i (bnd (alloc v) e) r. +Proof. by move=>H; apply/vrf_bnd/val_allocR. Qed. End EvalAllocR. @@ -205,8 +218,8 @@ Proof. by move=>H; apply/vrf_try/val_allocbR. Qed. Lemma bnd_allocbR e (v : A) n i (r : post B) : (forall x, vrf (updi x (nseq n v) \+ i) (e x) r) -> - vrf i (bind (allocb v n) e) r. -Proof. by move=>H; apply/vrf_bind/val_allocbR. Qed. + vrf i (bnd (allocb v n) e) r. +Proof. by move=>H; apply/vrf_bnd/val_allocbR. Qed. End EvalAllocbR. @@ -226,8 +239,8 @@ Proof. by move=>H; apply/vrf_try/val_deallocR. Qed. Lemma bnd_deallocR e (v : A) x i (f : forall k, form k i) (r : post B) : vrf (f Unit) (e tt) r -> - vrf (f (x :-> v)) (bind (dealloc x) e) r. -Proof. by move=>H; apply/vrf_bind/val_deallocR. Qed. + vrf (f (x :-> v)) (bnd (dealloc x) e) r. +Proof. by move=>H; apply/vrf_bnd/val_deallocR. Qed. End EvalDeallocR. @@ -246,7 +259,7 @@ Structure val_form A i r (p : Prop):= Structure bnd_form A B i (e : A -> ST B) r (p : Prop) := BndForm {bnd_pivot :> ST A; - _ : p -> vrf i (bind bnd_pivot e) r}. + _ : p -> vrf i (bnd bnd_pivot e) r}. Structure try_form A B i (e1 : A -> ST B) (e2 : exn -> ST B) r (p : Prop) := @@ -264,7 +277,7 @@ Proof. by case:e=>[?]; apply. Qed. (* continue searching for a val_form. *) Lemma bnd_case_pf A B i (s : A -> ST B) r p (e : bnd_form i s r p) : - p -> vrf i (bind e s) r. + p -> vrf i (bnd e s) r. Proof. by case:e=>[?]; apply. Qed. Canonical Structure @@ -407,20 +420,20 @@ End Uncons. (* These issues can be circumvented by falling back to simpler (automated) *) (* lemmas `gU`, `gR` and their variants. *) -Lemma gX G A (s : spec G A) g (m : heapPCM) m0 j tm k wh r2 - (e : STspec G s) - (fm : Syntactify.form (empx _) j tm) - (fu : uform m0 (Syntactify.untag fm)) - (f : forall q, form q r2) - (fg : PullX.rform j k tm wh (Some (untag (heap_of (f m0))))) - (Q : post A) : +Lemma gX G A (s : spec G A) g (m : heap) m0 j tm k wh r2 + (e : STspec G s) + (fm : Syntactify.form (empx _) j tm) + (fu : uform m0 (Syntactify.untag fm)) + (f : forall q, form q r2) + (fg : PullX.rform j k tm wh (Some (untag (heap_of (f m0))))) + (Q : post A) : untagU fu = m -> (valid m -> (s g).1 m) -> (forall v n, (s g).2 (Val v) n -> valid (untag (f n)) -> Q (Val v) (f n)) -> (forall x n, (s g).2 (Exn x) n -> valid (untag (f n)) -> Q (Exn x) (f n)) -> - vrf (fg : heapPCM) e Q. + vrf (PullX.unpack fg) e Q. Proof. case: e=>e /= H; case: fu=>_ ->/= Em Hp Hv Hx; rewrite -{}Em in Hp. rewrite (pullX (Syntactify.untag fm)) formE joinCA joinA. @@ -429,7 +442,7 @@ apply/vrf_frame/vrf_post/V. by case=>[v|x] n _=>[/Hv|/Hx]; rewrite formE. Qed. -Arguments gX [G A s] g m {m0 j tm k wh r2 e fm fu f fg Q} _ _ _ _. +Arguments gX {G A s} g m {m0 j tm k wh r2 e fm fu f fg Q}. Notation "[gX] @ m" := (gX tt m erefl) (at level 0). @@ -437,55 +450,55 @@ Notation "[ 'gX' x1 , .. , xn ] @ m" := (gX (existT _ x1 .. (existT _ xn tt) ..) m erefl) (at level 0, format "[ 'gX' x1 , .. , xn ] @ m"). -(* a combination of gX + vrf_bind *) -Lemma stepX G A B (s : spec G A) g (m : heapPCM) m0 j tm k wh r2 - (e : STspec G s) (e2 : A -> ST B) - (fm : Syntactify.form (empx _) j tm) - (fu : uform m0 (Syntactify.untag fm)) - (f : forall q, form q r2) - (fg : PullX.rform j k tm wh (Some (untag (heap_of (f m0))))) - (Q : post B) : +(* combination of gX + vrf_bind *) +Lemma stepX G A B (s : spec G A) g (m : heap) m0 j tm k wh r2 + (e : STspec G s) (e2 : A -> ST B) + (fm : Syntactify.form (empx _) j tm) + (fu : uform m0 (Syntactify.untag fm)) + (f : forall q, form q r2) + (fg : PullX.rform j k tm wh (Some (untag (heap_of (f m0))))) + (Q : post B) : untagU fu = m -> (valid m -> (s g).1 m) -> (forall v n, (s g).2 (Val v) n -> vrf (f n) (e2 v) Q) -> (forall x n, (s g).2 (Exn x) n -> valid (untag (f n)) -> Q (Exn x) (f n)) -> - vrf (fg : heapPCM) (bind e e2) Q. + vrf (PullX.unpack fg) (bnd e e2) Q. Proof. -move=>Hm Hp Hv Hx. -by apply/vrf_bind/(gX _ _ Hm Hp)=>[v n H V|x n H V _]; [apply: Hv| apply: Hx]. +move=>Hm Hp Hv Hx; apply/vrf_bnd/(gX _ _ Hm Hp). +- by move=>v n H V; apply: Hv. +by move=>x n H V _; apply: Hx. Qed. -Arguments stepX [G A B s] g m {m0 j tm k wh r2 e e2 fm fu f fg Q} _ _ _ _. +Arguments stepX {G A B s} g m {m0 j tm k wh r2 e e2 fm fu f fg Q}. Notation "[stepX] @ m" := (stepX tt m erefl) (at level 0). - Notation "[ 'stepX' x1 , .. , xn ] @ m" := (stepX (existT _ x1 .. (existT _ xn tt) ..) m erefl) (at level 0, format "[ 'stepX' x1 , .. , xn ] @ m"). -(* a combination of gX + vrf_try *) -Lemma tryX G A B (s : spec G A) g (m : heapPCM) m0 j tm k wh r2 - (e : STspec G s) (e1 : A -> ST B) (e2 : exn -> ST B) - (fm : Syntactify.form (empx _) j tm) - (fu : uform m0 (Syntactify.untag fm)) - (f : forall q, form q r2) - (fg : PullX.rform j k tm wh (Some (untag (heap_of (f m0))))) - (Q : post B) : +(* combination of gX + vrf_try *) +Lemma tryX G A B (s : spec G A) g (m : heap) m0 j tm k wh r2 + (e : STspec G s) (e1 : A -> ST B) (e2 : exn -> ST B) + (fm : Syntactify.form (empx _) j tm) + (fu : uform m0 (Syntactify.untag fm)) + (f : forall q, form q r2) + (fg : PullX.rform j k tm wh (Some (untag (heap_of (f m0))))) + (Q : post B) : untagU fu = m -> (valid m -> (s g).1 m) -> (forall v n, (s g).2 (Val v) n -> vrf (f n) (e1 v) Q) -> (forall x n, (s g).2 (Exn x) n -> vrf (f n) (e2 x) Q) -> - vrf (fg : heapPCM) (try e e1 e2) Q. + vrf (PullX.unpack fg) (try e e1 e2) Q. Proof. -move=>Hm Hp Hv Hx. -by apply/vrf_try/(gX _ _ Hm Hp)=>[x|ex] n H V; [apply: Hv|apply: Hx]. +move=>Hm Hp Hv Hx; apply/vrf_try/(gX _ _ Hm Hp). +- by move=>x n H V; apply: Hv. +by move=>ex n H V; apply: Hx. Qed. -Arguments tryX [G A B s] g m {m0 j tm k wh r2 e e1 e2 fm fu f fg Q} _ _ _ _. +Arguments tryX {G A B s} g m {m0 j tm k wh r2 e e1 e2 fm fu f fg Q}. Notation "[tryX] @ m" := (stepX tt m erefl) (at level 0). - Notation "[ 'tryX' x1 , .. , xn ] @ m" := (tryX (existT _ x1 .. (existT _ xn tt) ..) m erefl) (at level 0, format "[ 'tryX' x1 , .. , xn ] @ m"). @@ -494,7 +507,7 @@ Notation "[ 'tryX' x1 , .. , xn ] @ m" := (* Simplified ghost lemma automations *) (**************************************) -(* A simpler version of an automated framing+ghost lemma which only works on *) +(* simpler version of an automated framing+ghost lemma which only works on *) (* literal heap subexpressions (here `m`) but preserves associativity. *) Lemma gR G A (s : spec G A) g m r (e : STspec G s) (f : forall k, form k r) (Q : post A) : @@ -511,40 +524,38 @@ apply/vrf_frame/vrf_post/V. by case=>[x|ex] n _ =>[/Hv|/Hx]; rewrite formE. Qed. -Arguments gR [G A s] g m {r e f Q} _ _ _. +Arguments gR {G A s} g m {r e f Q}. Notation "[gR] @ m" := (gR tt m) (at level 0). - Notation "[ 'gR' x1 , .. , xn ] @ m" := (gR (existT _ x1 .. (existT _ xn tt) ..) m) (at level 0, format "[ 'gR' x1 , .. , xn ] @ m"). -(* a combination of gR + vrf_bind *) +(* combination of gR + vrf_bind *) Lemma stepR G A B (s : spec G A) g i j (e : STspec G s) (e2 : A -> ST B) - (f : forall k, form k j) (Q : post B) : + (f : forall k, form k j) (Q : post B) : (valid i -> (s g).1 i) -> (forall x m, (s g).2 (Val x) m -> vrf (f m) (e2 x) Q) -> (forall x m, (s g).2 (Exn x) m -> valid (untag (f m)) -> Q (Exn x) (f m)) -> - vrf (f i) (bind e e2) Q. + vrf (f i) (bnd e e2) Q. Proof. -move=>Hi H1 H2. -apply/vrf_bind/(gR _ _ Hi)=>[x m H V|ex m H V _]. -- by apply: H1 H. -by apply: H2. +move=>Hi H1 H2; apply/vrf_bnd/(gR _ _ Hi). +- by move=>x m H V; apply: H1 H. +by move=>ex m H V _; apply: H2. Qed. -Arguments stepR [G A B s] g i [j e e2 f Q] _ _ _. +Arguments stepR {G A B s} g i {j e e2 f Q}. Notation "[stepR] @ i" := (stepR tt i) (at level 0). - Notation "[ 'stepR' x1 , .. , xn ] @ i" := (stepR (existT _ x1 .. (existT _ xn tt) ..) i) (at level 0, format "[ 'stepR' x1 , .. , xn ] @ i"). -(* a combination of gR + vrf_try *) -Lemma tryR G A B (s : spec G A) g i j (e : STspec G s) (e1 : A -> ST B) (e2 : exn -> ST B) - (f : forall k, form k j) (Q : post B) : +(* combination of gR + vrf_try *) +Lemma tryR G A B (s : spec G A) g i j (e : STspec G s) + (e1 : A -> ST B) (e2 : exn -> ST B) + (f : forall k, form k j) (Q : post B) : (valid i -> (s g).1 i) -> (forall x m, (s g).2 (Val x) m -> vrf (f m) (e1 x) Q) -> (forall x m, (s g).2 (Exn x) m -> vrf (f m) (e2 x) Q) -> @@ -556,17 +567,16 @@ apply/vrf_try/(gR _ _ Hi)=>[x|ex] m H V. by apply: H2. Qed. -Arguments tryR [G A B s] g i [j e e1 e2 f Q] _ _ _. +Arguments tryR {G A B s} g i {j e e1 e2 f Q}. Notation "[tryR] @ i" := (tryR tt i) (at level 0). - Notation "[ 'tryR' x1 , .. , xn ] @ i" := (tryR (existT _ x1 .. (existT _ xn tt) ..) i) (at level 0, format "[ 'tryR' x1 , .. , xn ] @ i"). -(* This is really brittle, but I didn't get around yet to substitute it *) -(* with Mtac or overloaded lemmas. So for now, let's stick with the hack *) -(* in order to support the legacy proofs *) +(* The following is brittle, and should eventually be substituted *) +(* with overloaded lemmas. For now, sticking with the hack *) +(* for the purpose of supporting legacy proofs *) (* First cancelation in hypotheses *) diff --git a/htt/model.v b/htt/model.v index f3f1c23..c7f4ff8 100644 --- a/htt/model.v +++ b/htt/model.v @@ -75,13 +75,15 @@ Parameter dealloc : ptr -> ST unit. Arguments throw {A}. Arguments read {A}. -(* weakest-pre transformer *) -(* the order of args is: program, heap, post *) -(* program must come first so that automation can match on it *) +(* given program e : ST A, input heap i, postcondition Q *) +(* vrf' judges if running e in i produces output heap and result *) +(* that satisfy Q *) +(* related to predicate transformers *) +(* vrf' has program as its first argument to facilitate automation *) Parameter vrf' : forall A, ST A -> heap -> post A -> Prop. -(* recover the usual order of: pre-heap, prog, post *) -(* with a notation *) +(* in practice it's more convenient to order the arguments as *) +(* initial heap, program, postcondition *) Notation vrf i e Q := (vrf' e i Q). Parameter vrfV : forall A e i (Q : post A), @@ -96,7 +98,7 @@ Parameter vrf_ret : forall A x i (Q : post A), (valid i -> Q (Val x) i) -> vrf i (ret x) Q. Parameter vrf_throw : forall A e i (Q : post A), (valid i -> Q (Exn e) i) -> vrf i (throw e) Q. -Parameter vrf_bind : forall A B (e1 : ST A) (e2 : A -> ST B) +Parameter vrf_bnd : forall A B (e1 : ST A) (e2 : A -> ST B) i (Q : post B), vrf i e1 (fun x m => match x with @@ -138,18 +140,16 @@ Structure STspec G A (s : spec G A) := STprog { Arguments STspec G [A] s. -Notation "'Do' e" := (@STprog _ _ _ e _) (at level 80). - (* some notation *) - +Notation "'Do' e" := (@STprog _ _ _ e _) (at level 80). Notation "x '<--' c1 ';' c2" := (bnd c1 (fun x => c2)) - (at level 201, right associativity, format + (at level 81, right associativity, format "'[v' x '<--' c1 ';' '//' c2 ']'"). Notation "' x '<--' c1 ';' c2" := (bnd c1 (fun x => c2)) - (at level 201, right associativity, x strict pattern, format + (at level 81, right associativity, x strict pattern, format "'[v' ' x '<--' c1 ';' '//' c2 ']'"). Notation "c1 ';;' c2" := (bnd c1 (fun _ => c2)) - (at level 201, right associativity). + (at level 81, right associativity). Notation "'!' x" := (read x) (at level 50). Notation "x '::=' e" := (write x e) (at level 60). @@ -162,7 +162,7 @@ Notation "x '::=' e" := (write x e) (at level 60). (* We also hide model definition so that all functions one can *) (* apply ffix to will be built out of monotonicity-preserving *) (* constructors (and hence will be monotone). *) -Parameter Fix : forall G A (B : A -> Type) (s : forall x : A, spec G (B x)), +Parameter ffix : forall G A (B : A -> Type) (s : forall x : A, spec G (B x)), ((forall x : A, STspec G (s x)) -> forall x : A, STspec G (s x)) -> forall x : A, STspec G (s x). End VrfSig. @@ -235,9 +235,10 @@ Proof. by case: m=>// /dstr_st. Qed. Lemma fr_st A (e : ST A) : frameable (prog_of e). Proof. by case: e=>?? []. Qed. +Arguments sfm_st {A e i j}. +Arguments dstr_st {A e i}. Arguments fr_st {A e i j}. - Section STDef. Variable A : Type. Implicit Type e : ST A. @@ -277,54 +278,31 @@ Qed. HB.instance Definition _ := isPoset.Build (ST A) st_is_poset. -(* bottom is a program that can always run *) -(* but never returns (infinite loop) *) - -(* DO WE NEED THIS? If we have lattice sturcture *) -(* existence of bot follows *) - -Definition pre_bot : pre := top. -Definition prog_bot : prog pre_bot A := fun _ _ _ _ => bot. -Lemma progbot_is_st : st_axiom prog_bot. -Proof. by split=>//; move=>*. Qed. -Definition st_bot := Prog progbot_is_st. - -Lemma st_botP e : st_leq st_bot e. -Proof. by case: e=>p e [???]; exists (fun _ _ => I)=>???; apply: botP. Qed. - (* lattice structure on ST *) -UP TO HERE - - - - +(* suprema of programs: *) +(* - precondition is intersection of preconditions of programs *) +(* - denotation is union of denotations of programs *) + (* intersection of preconditions *) -Definition pre_sup (u : Pred ST) : pre := +Definition pre_sup (u : Pred (ST A)) : pre := fun h => forall e, e \In u -> h \In pre_of e. Definition pre_sup_leq u e (pf : e \In u) : pre_sup u <== pre_of e := fun h (pf1 : pre_sup u h) => pf1 e pf. (* union of postconditions *) -Definition prog_sup (u : Pred ST) : prog (pre_sup u) A := +Definition prog_sup (u : Pred (ST A)) : prog (pre_sup u) A := fun i V P y m => exists e (pf : e \In u), prog_of e _ V (pre_sup_leq pf P) y m. Arguments prog_sup : clear implicits. -Lemma pre_sup_sfmono u : safe_mono (pre_sup u). -Proof. -move=>i j Pi Vij e He. -by apply: sfm_st=>//; apply: Pi. -Qed. - -Lemma prog_sup_dstrict u : def_strict (prog_sup u). -Proof. by move=>i P V y; case; case=>p e S D F [H1] /D. Qed. - -Lemma prog_sup_frame u : frameable (prog_sup u). +Lemma progsup_is_st u : st_axiom (prog_sup u). Proof. -move=>i j Pi Vij Pij y m [e][He]Pe. +split=>[i j Pi Vij E He|i P V y|i j Pi Vij Pij y m [e][He Pe]]. +- by apply: sfm_st=>//; apply: Pi. +- by case; case=>p e [S D F][H1] /D. have Pi' := Pi e He; have Pij' := Pij e He. move: Pe; rewrite (pf_irr (pre_sup_leq He Pij) Pij'). case/(fr_st Pi' Vij Pij')=>h [{m}-> Vhj Ph]. @@ -332,43 +310,29 @@ exists h; split=>//; exists e, He. by rewrite (pf_irr (pre_sup_leq He Pi) Pi'). Qed. -Definition st_sup u : ST := - Prog (@pre_sup_sfmono u) (@prog_sup_dstrict u) (@prog_sup_frame u). - -Lemma st_supP u e : e \In u -> e <== st_sup u. -Proof. -case: e=>p e' S D F R. -exists (pre_sup_leq R)=>/=p0 y m H. -by exists (Prog S D F), R. -Qed. +Definition st_sup u : ST A := Prog (progsup_is_st u). -Lemma st_supM u e : - (forall e1, e1 \In u -> e1 <== e) -> st_sup u <== e. +Lemma st_is_lattice : lattice_axiom st_sup. Proof. -case: e=>p e S D F R. +split=>/= u e; case: e=>p e' X R. +- exists (pre_sup_leq R)=>/=p0 y m H. + by exists (Prog X), R. have J : p <== pre_sup u. -- by move=>/= x Px e' pf; case: (R _ pf)=>/= + _; apply. +- by move=>/= x Px ex pf; case: (R _ pf)=>/= + _; apply. exists J=>i V P y m [e0][H0 Hm]. case: (R _ H0)=>/= Hx; apply. by rewrite (pf_irr (Hx i P) (pre_sup_leq H0 (J i P))). Qed. -Definition stLatticeMixin := LatticeMixin st_supP st_supM. -Canonical stLattice := Lattice ST stLatticeMixin. - +HB.instance Definition _ := isLattice.Build (ST A) st_is_lattice. End STDef. -Arguments prog_of [A]. -Arguments sfm_st [A e i j]. -Arguments dstr_st [A e i]. -Arguments fr_st [A e i j]. +(* types that embed the specs *) Section STspecDef. Variables (G A : Type) (s : spec G A). -(* strongest postcondition predicate transformer *) - -Definition vrf' (e : ST A) i (Q : post A) := +Definition vrf' (e : ST A) (i : heap) (Q : post A) := forall (V : valid i), exists (pf : i \In pre_of e), forall y m, prog_of e _ V pf y m -> Q y m. @@ -392,35 +356,23 @@ Qed. Definition stsp_leq e1 e2 := model e1 <== model e2. -Lemma stsp_refl e : stsp_leq e e. -Proof. by case: e=>e He; apply: poset_refl. Qed. - -Lemma stsp_asym e1 e2 : stsp_leq e1 e2 -> stsp_leq e2 e1 -> e1 = e2. -Proof. -move: e1 e2=>[e1 H1][e2 H2]; rewrite /stsp_leq /= =>E1 E2. -have E := poset_asym E1 E2; subst e2. -by congr STprog; apply: pf_irr. -Qed. - -Lemma stsp_trans e1 e2 e3 : stsp_leq e1 e2 -> stsp_leq e2 e3 -> stsp_leq e1 e3. +Lemma stspleq_is_poset : poset_axiom stsp_leq. Proof. -move: e1 e2 e3=>[e1 H1][e2 H2][e3 H3]. +split=>[e|e1 e2|e1 e2 e3]. +- by case: e=>e He; apply: poset_refl. +- case: e1 e2=>e1 H1 [e2 H2]; rewrite /stsp_leq /= =>E1 E2. + have E := poset_asym E1 E2; subst e2. + by congr STprog; apply: pf_irr. +case: e1 e2 e3=>e1 H1 [e2 H2][e3 H3]. by apply: poset_trans. Qed. -Lemma st_bot_has_spec : @st_bot A \In has_spec. -Proof. by move=>g i H V /=; exists I. Qed. - -Definition stsp_bot := STprog st_bot_has_spec. - -Lemma stsp_botP e : stsp_leq stsp_bot e. -Proof. by case: e=>*; apply: botP. Qed. - -Definition stspPosetMixin := PosetMixin stsp_botP stsp_refl stsp_asym stsp_trans. -Canonical stspPoset := Eval hnf in Poset STspec stspPosetMixin. +HB.instance Definition _ := isPoset.Build STspec stspleq_is_poset. (* lattice structure on STspec *) +(* denotation of the union of STspec programs is *) +(* the union of their denotations *) Definition st_sup' (u : Pred STspec) : ST A := sup [Pred p | exists e, p = model e /\ e \In u]. @@ -437,53 +389,20 @@ Qed. Definition stsp_sup u := STprog (@st_sup_has_spec' u). -Lemma stsp_supP u e : e \In u -> e <== stsp_sup u. -Proof. by case: e=>p S R; apply: supP; exists (STprog S). Qed. - -Lemma stsp_supM u e : - (forall e1, e1 \In u -> e1 <== e) -> stsp_sup u <== e. -Proof. by case: e=>p S R; apply: supM=>/= y[q][->]; apply: R. Qed. - -Definition stspLatticeMixin := LatticeMixin stsp_supP stsp_supM. -Canonical stspLattice := Lattice STspec stspLatticeMixin. - -End STspecDef. - -Notation vrf i e Q := (vrf' e i Q). - -(************************************) -(* modeling the language primitives *) -(************************************) - -(* recursion *) -Section Fix. -Variables (G A : Type) (B : A -> Type) (s : forall x, spec G (B x)). -Notation tp := (forall x, STspec (s x)). -Notation lat := (dfunLattice (fun x => [lattice of STspec (s x)])). -Variable (f : tp -> tp). - -(* we take a fixpoint not of f, but of its monotone completion f' *) - -Definition f' (e : lat) := - sup [Pred t : lat | exists e', e' <== e /\ t = f e']. - -Lemma f'_mono : monotone f'. +Lemma stspsup_is_lattice : lattice_axiom stsp_sup. Proof. -move=>x y H; apply: sup_mono=>fz; case=>z [Hz {fz}->]. -by exists z; split=>//; apply/poset_trans/H. +split=>/= u [p S R]. +- by apply: supP; exists (STprog S). +by apply: supM=>/= y[q][->]; apply: R. Qed. -Definition Fix : tp := tarski_lfp f'. +HB.instance Definition _ := isLattice.Build STspec stspsup_is_lattice. -(* fixed point constructor which requires explicit proof of monotonicity *) -Definition Fix' (pf : monotone (f : lat -> lat)) : tp := - tarski_lfp (f : lat -> lat). - -End Fix. +End STspecDef. -Arguments Fix [G A B s] f x. -Arguments Fix' [G A B s] f pf x. +Notation vrf i e Q := (vrf' e i Q). +(* required vrf lemmas *) Section VrfLemmas. Variables (A : Type) (e : ST A). @@ -525,102 +444,110 @@ Qed. End VrfLemmas. +(************************************) +(* modeling the language primitives *) +(************************************) + +(* recursion *) + +Section Fix. +Variables (G A : Type) (B : A -> Type) (pq : forall x, spec G (B x)). +Notation tp := (forall x, STspec (pq x)). +Notation lat := (dfun_lattice (fun x => STspec (pq x))). +Variable f : tp -> tp. + +(* fixed point constructor over monotone closure *) +Definition f' (c : lat) := sup [Pred t : lat | exists c', c' <== c /\ t = f c']. +Definition ffix := tarski_lfp f'. + +(* fixed point constructor which requires explicit proof of monotonicity *) +Definition monotone' := monofun_axiom (f : lat -> lat). +Definition ffix2 (pf : monotone') : tp := tarski_lfp (f : lat -> lat). +End Fix. + +Arguments ffix {G A B pq}. +Arguments ffix2 {G A B pq}. + +(* monadic unit *) Section Return. Variables (A : Type) (x : A). Definition ret_pre : pre := top. Definition ret_prog : prog ret_pre A := - fun i _ _ y m => - m = i /\ y = Val x. - -Lemma ret_sfmono : safe_mono ret_pre. -Proof. by []. Qed. + fun i _ _ y m => m = i /\ y = Val x. -Lemma ret_dstrict : def_strict ret_prog. -Proof. by move=>i [] V _ /= [E _]; rewrite -E in V. Qed. - -Lemma ret_frame : frameable ret_prog. -Proof. by move=>i j [Vij []] _ _ [-> ->]; exists i. Qed. +Lemma retprog_is_st : st_axiom ret_prog. +Proof. +split=>[//|i [] V _ /= [E _]|i j [Vij []] _ _ [->->]]. +- by rewrite -E in V. +by exists i. +Qed. -Definition ret := Prog ret_sfmono ret_dstrict ret_frame. +Definition ret := Prog retprog_is_st. Lemma vrf_ret i (Q : post A) : (valid i -> Q (Val x) i) -> vrf i ret Q. Proof. by move=>H V; exists I=>_ _ [->->]; apply: H. Qed. - End Return. +(* exception throwing *) Section Throw. Variables (A : Type) (e : exn). Definition throw_pre : pre := top. Definition throw_prog : prog throw_pre A := - fun i _ _ y m => - m = i /\ y = @Exn A e. - -Lemma throw_sfmono : safe_mono throw_pre. -Proof. by []. Qed. - -Lemma throw_dstrict : def_strict throw_prog. -Proof. by move=>i [] V _ /= [E _]; rewrite -E in V. Qed. + fun i _ _ y m => m = i /\ y = @Exn A e. -Lemma throw_frame : frameable throw_prog. -Proof. by move=>i j [Vij []] _ _ [-> ->]; exists i. Qed. +Lemma throwprog_is_st : st_axiom throw_prog. +Proof. +split=>[//|i [V] _ /= [E _]|i j [Vij []] _ _ [->->]]. +- by rewrite -E in V. +by exists i. +Qed. -Definition throw := Prog throw_sfmono throw_dstrict throw_frame. +Definition throw := Prog throwprog_is_st. Lemma vrf_throw i (Q : post A) : (valid i -> Q (Exn e) i) -> vrf i throw Q. Proof. by move=>H V; exists I=>_ _ [->->]; apply: H. Qed. - End Throw. -Section Bind. +(* monadic bind *) +Section Bnd. Variables (A B : Type). Variables (e1 : ST A) (e2 : A -> ST B). -Definition bind_pre : pre := +Definition bnd_pre : pre := fun i => exists (Vi : valid i) (Pi : i \In pre_of e1), forall x m, prog_of e1 _ Vi Pi (Val x) m -> pre_of (e2 x) m. -Definition bind_pre_proj i : i \In bind_pre -> i \In pre_of e1 := +Definition bnd_pre_proj i : i \In bnd_pre -> i \In pre_of e1 := fun '(ex_intro _ (ex_intro p _)) => p. -Definition bind_prog : prog bind_pre B := +Definition bnd_prog : prog bnd_pre B := fun i V P y m => - exists x h (Ph : h \In prog_of e1 _ V (bind_pre_proj P) x), + exists x h (Ph : h \In prog_of e1 _ V (bnd_pre_proj P) x), match x with | Val x' => exists Ph' : h \In pre_of (e2 x'), m \In prog_of (e2 x') _ (dstr_valid Ph) Ph' y | Exn e => y = Exn e /\ m = h end. -Lemma bind_sfmono : safe_mono bind_pre. -Proof. -move=>i j [Vi][Pi]P Vij. -have Pij := sfm_st Pi Vij. -exists Vij, Pij=>x m. -case/(fr_st Pi Vij Pij)=>h [{m}-> Vhj]. -rewrite (pf_irr (validL Vij) Vi)=>/P Ph. -by apply: sfm_st=>//; apply: (validL Vhj). -Qed. - -Lemma bind_dstrict : def_strict bind_prog. -Proof. -move=>i [Vi][Pi P] Vi' y[x][h][/=]. -case: x=>[x|e]Ph. -- by case=>Ph' /dstr_st. -by case=>_; move: Ph=>/[swap]<- /dstr_st. -Qed. - -Lemma bind_frame : frameable bind_prog. +Lemma bndprog_is_st : st_axiom bnd_prog. Proof. +split. +- move=>i j [Vi][Pi] P Vij; have Pij := sfm_st Pi Vij. + exists Vij, Pij=>x m; case/(fr_st Pi Vij Pij)=>h [{m}-> Vhj]. + rewrite (pf_irr (validL Vij) Vi)=>/P Ph. + by apply: sfm_st=>//; apply: (validL Vhj). +- move=>i [Vi][Pi P] Vi' y [x][h][/=]. + case: x=>[x|e]Ph; first by case=>Ph' /dstr_st. + by case=>_; move: Ph=>/[swap]<- /dstr_st. move=>i j [Vi][Pi P] Vij [_ [Pij _]] y m [x][h][/=]. -move: (fr_st Pi Vij Pij)=>H. -case: x=>[x|e] Ph. +move: (fr_st Pi Vij Pij)=>H; case: x=>[x|e] Ph. - case=>Ph'. case: (H _ _ Ph)=>h1[Eh V1 Ph1]; subst h. rewrite (pf_irr (validL Vij) Vi) in Ph1 *; move: (P _ _ Ph1)=> P21. @@ -633,27 +560,27 @@ case/H: Ph=>h1[Eh Vh Ph1]. by exists h1; split=>//; exists (Exn e), h1, Ph1. Qed. -Definition bind := Prog bind_sfmono bind_dstrict bind_frame. +Definition bnd := Prog bndprog_is_st. -Lemma vrf_bind i (Q : post B) : +Lemma vrf_bnd i (Q : post B) : vrf i e1 (fun x m => match x with | Val x' => vrf m (e2 x') Q | Exn e => valid m -> Q (Exn e) m end) -> - vrf i bind Q. + vrf i bnd Q. Proof. move=>H Vi; case: (H Vi)=>Hi {}H /=. -have Hi' : i \In bind_pre. +have Hi' : i \In bnd_pre. - by exists Vi, Hi=>x m Pm; case: (H _ _ Pm (dstr_valid Pm)). exists Hi'=>y j /= [x][m][Pm] C. -rewrite (pf_irr Hi (bind_pre_proj Hi')) in H. +rewrite (pf_irr Hi (bnd_pre_proj Hi')) in H. case: x Pm C=>[x|e] Pm; move: (H _ _ Pm (dstr_valid Pm))=>{}H. - by case=>Pm2 Pj; case: H=>Pm2'; apply; rewrite (pf_irr Pm2' Pm2). by case=>->->. Qed. -End Bind. +End Bnd. Section Try. Variables (A B : Type). @@ -678,32 +605,23 @@ Definition try_prog : prog try_pre B := m \In prog_of (e2 ex) _ (dstr_valid Ph) Ph' y end. -Lemma try_sfmono : safe_mono try_pre. +Lemma tryprog_is_st : st_axiom try_prog. Proof. -move=>i j [Vi [Pi][E1 E2]] Vij. -have Pij := sfm_st Pi Vij. -exists Vij, Pij; split. -- move=>y m. +split. +- move=>i j [Vi [Pi][E1 E2]] Vij; have Pij := sfm_st Pi Vij. + exists Vij, Pij; split. + - move=>y m. + case/(fr_st Pi Vij Pij)=>h [{m}-> Vhj]. + rewrite (pf_irr (validL Vij) Vi)=>/E1 Ph. + by apply: sfm_st=>//; apply: (validL Vhj). + move=>ex m. case/(fr_st Pi Vij Pij)=>h [{m}-> Vhj]. - rewrite (pf_irr (validL Vij) Vi)=>/E1 Ph. + rewrite (pf_irr (validL Vij) Vi)=>/E2 Ph. by apply: sfm_st=>//; apply: (validL Vhj). -move=>ex m. -case/(fr_st Pi Vij Pij)=>h [{m}-> Vhj]. -rewrite (pf_irr (validL Vij) Vi)=>/E2 Ph. -by apply: sfm_st=>//; apply: (validL Vhj). -Qed. - -Lemma try_dstrict : def_strict try_prog. -Proof. -move=>i [Vi [Pi][E1 E2]] Vi' y[x][h][/=]. -by case: x=>[x|ex] Eh; case=>Ph /dstr_st. -Qed. - -Lemma try_frame : frameable try_prog. -Proof. +- move=>i [Vi [Pi][E1 E2]] Vi' y [x][h][/=]. + by case: x=>[x|ex] Eh; case=>Ph /dstr_st. move=>i j [Vi [Pi][E1 E2]] Vij [_ [Pij _]] y m [x][h][/=]. -move: (fr_st Pi Vij Pij)=>H. -case: x=>[x|ex] Ph. +move: (fr_st Pi Vij Pij)=>H; case: x=>[x|ex] Ph. - case=>Ph'. case: (H _ _ Ph)=>h1[Eh V1 Ph1]; subst h. rewrite (pf_irr (validL Vij) Vi) in Ph1 *; move: (E1 _ _ Ph1)=>P21. @@ -720,7 +638,7 @@ exists h2; split=>//; exists (Exn ex), h1, Ph1, P21. by rewrite (pf_irr (dstr_valid Ph1) (validL V1)). Qed. -Definition try := Prog try_sfmono try_dstrict try_frame. +Definition try := Prog tryprog_is_st. Lemma vrf_try i (Q : post B) : vrf i e (fun x m => @@ -742,18 +660,18 @@ Qed. End Try. -(* don't export, just for fun *) +(* bnd follows from try; not to export, just for fun *) Lemma bnd_is_try A B (e1 : ST A) (e2 : A -> ST B) i r : vrf i (try e1 e2 (throw B)) r -> - vrf i (bind e1 e2) r. + vrf i (bnd e1 e2) r. Proof. move=>H Vi; case: (H Vi)=>[[Vi'][P1 /= [E1 E2]]] {}H. -have J : i \In pre_of (bind e1 e2). +have J : i \In pre_of (bnd e1 e2). - exists Vi, P1=>y m. by rewrite (pf_irr Vi Vi')=>/E1. exists J=>y m /= [x][h][Ph]C. apply: H; exists x, h=>/=. -rewrite (pf_irr P1 (bind_pre_proj J)) in E2 *; exists Ph. +rewrite (pf_irr P1 (bnd_pre_proj J)) in E2 *; exists Ph. move: Ph C; case: x=>// e Ph [{y}-> {m}->]. rewrite (pf_irr Vi' Vi) in E2. by exists (E2 _ _ Ph). @@ -762,8 +680,6 @@ Qed. Section Read. Variable (A : Type) (x : ptr). -Local Notation idyn v := (@dyn _ id _ v). - Definition read_pre : pre := fun i => x \in dom i /\ exists v : A, find x i = Some (idyn v). @@ -771,27 +687,23 @@ Definition read_prog : prog read_pre A := fun i _ _ y m => exists w, [/\ m = i, y = Val w & find x m = Some (idyn w)]. -Lemma read_sfmono : safe_mono read_pre. -Proof. -move=>i j [Hx [v E]] Vij; split. -- by rewrite domUn inE Vij Hx. -by exists v; rewrite findUnL // Hx. -Qed. - -Lemma read_dstrict : def_strict read_prog. -Proof. by move=>i _ Vi _ [_ [E _ _]]; rewrite -E in Vi. Qed. - -Lemma read_frame : frameable read_prog. +Lemma readprog_is_st : st_axiom read_prog. Proof. +split. +- move=>i j [Hx [v E]] Vij; split. + - by rewrite domUn inE Vij Hx. + by exists v; rewrite findUnL // Hx. +- by move=>i _ Vi _ [_ [E _ _]]; rewrite -E in Vi. move=>i j [Hx [v E]] Vij _ _ _ [w [-> -> H]]. exists i; split=>//; exists w; split=>{v E}//. by rewrite findUnL // Hx in H. Qed. -Definition read := Prog read_sfmono read_dstrict read_frame. +Definition read := Prog readprog_is_st. Lemma vrf_read j (v : A) (Q : post A) : - (valid (x :-> v \+ j) -> Q (Val v) (x :-> v \+ j)) -> + (valid (x :-> v \+ j) -> + Q (Val v) (x :-> v \+ j)) -> vrf (x :-> v \+ j) read Q. Proof. move=>H Vi /=. @@ -808,30 +720,20 @@ End Read. Section Write. Variable (A : Type) (x : ptr) (v : A). -Local Notation idyn v := (@dyn _ id _ v). - -Definition write_pre : pre := - fun i => x \in dom i. +Definition write_pre : pre := fun i => x \in dom i. Definition write_prog : prog write_pre unit := fun i _ _ y m => [/\ y = Val tt, m = upd x (idyn v) i & x \in dom i]. -Lemma write_sfmono : safe_mono write_pre. -Proof. -move=>i j; rewrite /write_pre -!toPredE /= => Hx Vij. -by rewrite domUn inE Vij Hx. -Qed. - -Lemma write_dstrict : def_strict write_prog. -Proof. -move=>i Hx _ _ [_ E _]; rewrite /write_pre -toPredE /= in Hx. -suff {E}: valid (upd x (idyn v) i) by rewrite -E. -by rewrite validU (dom_cond Hx) (dom_valid Hx). -Qed. - -Lemma write_frame : frameable write_prog. +Lemma writeprog_is_st : st_axiom write_prog. Proof. +split. +- move=>i j; rewrite /write_pre -!toPredE /= => Hx Vij. + by rewrite domUn inE Vij Hx. +- move=>i Hx _ _ [_ E _]; rewrite /write_pre -toPredE /= in Hx. + suff {E}: valid (upd x (idyn v) i) by rewrite -E. + by rewrite validU (dom_cond Hx) (dom_valid Hx). move=>i j Hx Vij _ _ _ [-> -> _]. exists (upd x (idyn v) i); split=>//; rewrite /write_pre -toPredE /= in Hx. @@ -839,10 +741,11 @@ rewrite /write_pre -toPredE /= in Hx. by rewrite validUUn. Qed. -Definition write := Prog write_sfmono write_dstrict write_frame. +Definition write := Prog writeprog_is_st. Lemma vrf_write B (u : B) j (Q : post unit) : - (valid (x :-> v \+ j) -> Q (Val tt) (x :-> v \+ j)) -> + (valid (x :-> v \+ j) -> + Q (Val tt) (x :-> v \+ j)) -> vrf (x :-> u \+ j) write Q. Proof. move=>H Vi /=. @@ -858,8 +761,6 @@ End Write. Section Allocation. Variables (A : Type) (v : A). -Local Notation idyn v := (@dyn _ id _ v). - Definition alloc_pre : pre := top. Definition alloc_prog : prog alloc_pre ptr := @@ -867,18 +768,12 @@ Definition alloc_prog : prog alloc_pre ptr := exists l, [/\ y = Val l, m = l :-> v \+ i, l != null & l \notin dom i]. -Lemma alloc_sfmono : safe_mono alloc_pre. -Proof. by []. Qed. - -Lemma alloc_dstrict : def_strict alloc_prog. -Proof. -move=>i [] Vi _ [l][_ E Hl Hl2]. -suff {E}: valid (l :-> v \+ i) by rewrite -E. -by rewrite validPtUn; apply/and3P. -Qed. - -Lemma alloc_frame : frameable alloc_prog. +Lemma allocprog_is_st : st_axiom alloc_prog. Proof. +split=>//. +- move=>i [] Vi _ [l][_ E Hl Hl2]. + suff {E}: valid (l :-> v \+ i) by rewrite -E. + by rewrite validPtUn; apply/and3P. move=>i j [] Vij [] _ _ [l][->-> Hl Hl2]. exists (l :-> v \+ i); rewrite -joinA; split=>//. - rewrite validUnAE validPt domPtK Hl Vij /= all_predC. @@ -888,10 +783,11 @@ exists l; split=>//. by apply/dom_NNL/Hl2. Qed. -Definition alloc := Prog alloc_sfmono alloc_dstrict alloc_frame. +Definition alloc := Prog allocprog_is_st. Lemma vrf_alloc i (Q : post ptr) : - (forall x, valid (x :-> v \+ i) -> Q (Val x) (x :-> v \+ i)) -> + (forall x, valid (x :-> v \+ i) -> + Q (Val x) (x :-> v \+ i)) -> vrf i alloc Q. Proof. move=>H Vi /=. @@ -910,21 +806,17 @@ Definition allocb_prog : prog allocb_pre ptr := fun i _ _ y m => exists l, [/\ y = Val l, m = updi l (nseq n v) \+ i & valid m]. -Lemma allocb_sfmono : safe_mono allocb_pre. -Proof. by []. Qed. - -Lemma allocb_dstrict : def_strict allocb_prog. -Proof. by move=>i [] Vi y [l][]. Qed. - -Lemma allocb_frame : frameable allocb_prog. +Lemma allocbprog_is_st : st_axiom allocb_prog. Proof. +split=>//. +- by move=>i [] Vi y [l][]. move=>i j [] Vij [] _ _ [l][->-> V]. exists (updi l (nseq n v) \+ i); rewrite -joinA; split=>//. exists l; split=>//. by rewrite joinA in V; apply: (validL V). Qed. -Definition allocb := Prog allocb_sfmono allocb_dstrict allocb_frame. +Definition allocb := Prog allocbprog_is_st. Lemma vrf_allocb i (Q : post ptr) : (forall x, valid (updi x (nseq n v) \+ i) -> @@ -948,33 +840,27 @@ Definition dealloc_prog : prog dealloc_pre unit := fun i _ _ y m => [/\ y = Val tt, m = free i x & x \in dom i]. -Lemma dealloc_sfmono : safe_mono dealloc_pre. -Proof. -move=>i j; rewrite /dealloc_pre -!toPredE /= => Hx Vij. -by rewrite domUn inE Vij Hx. -Qed. - -Lemma dealloc_dstrict : def_strict dealloc_prog. -Proof. -move=>i _ Vi _ [_ E _]. -suff {E}: valid (free i x) by rewrite -E. -by rewrite validF. -Qed. - -Lemma dealloc_frame : frameable dealloc_prog. +Lemma deallocprog_is_st : st_axiom dealloc_prog. Proof. +split. +- move=>i j; rewrite /dealloc_pre -!toPredE /= => Hx Vij. + by rewrite domUn inE Vij Hx. +- move=>i _ Vi _ [_ E _]. + suff {E}: valid (free i x) by rewrite -E. + by rewrite validF. move=>i j Hx Vij Hx' _ _ [->-> _]. exists (free i x); split=>//; rewrite /dealloc_pre -!toPredE /= in Hx Hx'. -- by apply/freeUnR/dom_inNL/Hx. +- by apply/freeUnL/dom_inNL/Hx. by apply: validFUn. Qed. -Definition dealloc := - Prog dealloc_sfmono dealloc_dstrict dealloc_frame. +Definition dealloc := Prog deallocprog_is_st. Lemma vrf_dealloc A (v : A) j (Q : post unit) : - (x \notin dom j -> valid j -> Q (Val tt) j) -> + (x \notin dom j -> + valid j -> + Q (Val tt) j) -> vrf (x :-> v \+ j) dealloc Q. Proof. move=>H Vi /=. @@ -990,39 +876,44 @@ End Deallocation. (* Monotonicity of the constructors *) Section Monotonicity. - -Variables (A B : Type). +Variables A B : Type. Lemma do_mono G (e1 e2 : ST A) (s : spec G A) - (pf1 : has_spec s e1) (pf2 : has_spec s e2) : - e1 <== e2 -> @STprog _ _ _ e1 pf1 <== @STprog _ _ _ e2 pf2. + (pf1 : has_spec s e1) (pf2 : has_spec s e2) : + e1 <== e2 -> + @STprog _ _ _ e1 pf1 <== @STprog _ _ _ e2 pf2. Proof. by []. Qed. Lemma bind_mono (e1 e2 : ST A) (k1 k2 : A -> ST B) : - e1 <== e2 -> k1 <== k2 -> (bind e1 k1 : ST B) <== bind e2 k2. + e1 <== e2 -> + k1 <== k2 -> + (bnd e1 k1 : ST B) <== bnd e2 k2. Proof. move=>[H1 H2] pf2. -have pf: bind_pre e2 k2 <== bind_pre e1 k1. +have pf: bnd_pre e2 k2 <== bnd_pre e1 k1. - move=>h [Vh][Pi] H. exists Vh, (H1 _ Pi)=>x m /H2/H H'. by case: (pf2 x)=>+ _; apply. exists pf=>i Vi /[dup][[Vi'][Pi']P'] Pi x h. case; case=>[a|e][h0][Ph][Ph'] H. - exists (Val a), h0. - move: (H2 i Vi (bind_pre_proj Pi))=>H'. - rewrite (pf_irr (H1 i (bind_pre_proj Pi)) (bind_pre_proj (pf i Pi))) in H'. + move: (H2 i Vi (bnd_pre_proj Pi))=>H'. + rewrite (pf_irr (H1 i (bnd_pre_proj Pi)) (bnd_pre_proj (pf i Pi))) in H'. have Ph0 := (H' _ _ Ph); exists Ph0. - move: (P' a h0); rewrite (pf_irr Vi' Vi) (pf_irr Pi' (bind_pre_proj Pi))=>H''. + move: (P' a h0); rewrite (pf_irr Vi' Vi) (pf_irr Pi' (bnd_pre_proj Pi))=>H''. exists (H'' Ph0); case: (pf2 a)=>Pr; apply. - by rewrite (pf_irr (dstr_valid Ph0) (dstr_valid Ph)) (pf_irr (Pr h0 (H'' Ph0)) Ph'). + rewrite (pf_irr (dstr_valid Ph0) (dstr_valid Ph)). + by rewrite (pf_irr (Pr h0 (H'' Ph0)) Ph'). rewrite Ph' H; exists (Exn e), h0. -move: (H2 i Vi (bind_pre_proj Pi))=>H'. -rewrite (pf_irr (H1 i (bind_pre_proj Pi)) (bind_pre_proj (pf i Pi))) in H'. +move: (H2 i Vi (bnd_pre_proj Pi))=>H'. +rewrite (pf_irr (H1 i (bnd_pre_proj Pi)) (bnd_pre_proj (pf i Pi))) in H'. by exists (H' _ _ Ph). Qed. Lemma try_mono (e1 e2 : ST A) (k1 k2 : A -> ST B) (h1 h2 : exn -> ST B) : - e1 <== e2 -> k1 <== k2 -> h1 <== h2 -> + e1 <== e2 -> + k1 <== k2 -> + h1 <== h2 -> (try e1 k1 h1 : ST B) <== try e2 k2 h2. Proof. move=>[H1 H2] pf2 pf3. @@ -1039,81 +930,93 @@ case; case=>[a|e][h0][Ph][Ph'] H. move: (H2 i Vi (try_pre_proj Pi))=>H'. rewrite (pf_irr (H1 i (try_pre_proj Pi)) (try_pre_proj (pf i Pi))) in H'. have Ph1 := (H' _ _ Ph); exists Ph1. - move: (Pk0 a h0); rewrite (pf_irr Vi' Vi) (pf_irr Pi' (try_pre_proj Pi))=>H''. + move: (Pk0 a h0); rewrite (pf_irr Vi' Vi). + rewrite (pf_irr Pi' (try_pre_proj Pi))=>H''. exists (H'' Ph1); case: (pf2 a)=>Pr; apply. - by rewrite (pf_irr (dstr_valid Ph1) (dstr_valid Ph)) (pf_irr (Pr h0 (H'' Ph1)) Ph'). + rewrite (pf_irr (dstr_valid Ph1) (dstr_valid Ph)). + by rewrite (pf_irr (Pr h0 (H'' Ph1)) Ph'). exists (Exn e), h0. move: (H2 i Vi (try_pre_proj Pi))=>H'. rewrite (pf_irr (H1 i (try_pre_proj Pi)) (try_pre_proj (pf i Pi))) in H'. have Ph1 := (H' _ _ Ph); exists Ph1. move: (Ph0 e h0); rewrite (pf_irr Vi' Vi) (pf_irr Pi' (try_pre_proj Pi))=>H''. exists (H'' Ph1); case: (pf3 e)=>Pr; apply. -by rewrite (pf_irr (dstr_valid Ph1) (dstr_valid Ph)) (pf_irr (Pr h0 (H'' Ph1)) Ph'). +rewrite (pf_irr (dstr_valid Ph1) (dstr_valid Ph)). +by rewrite (pf_irr (Pr h0 (H'' Ph1)) Ph'). Qed. (* the rest of the constructors are trivial *) Lemma ret_mono (v1 v2 : A) : - v1 = v2 -> (ret v1 : ST A) <== ret v2. + v1 = v2 -> + (ret v1 : ST A) <== ret v2. Proof. by move=>->. Qed. Lemma throw_mono (e1 e2 : exn) : - e1 = e2 -> (throw A e1 : ST A) <== throw A e2. + e1 = e2 -> + (throw A e1 : ST A) <== throw A e2. Proof. by move=>->. Qed. Lemma read_mono (p1 p2 : ptr) : - p1 = p2 -> (read A p1 : ST A) <== read A p2. + p1 = p2 -> + (read A p1 : ST A) <== read A p2. Proof. by move=>->. Qed. Lemma write_mono (p1 p2 : ptr) (x1 x2 : A) : - p1 = p2 -> x1 = x2 -> (write p1 x1 : ST unit) <== write p2 x2. + p1 = p2 -> + x1 = x2 -> + (write p1 x1 : ST unit) <== write p2 x2. Proof. by move=>->->. Qed. Lemma alloc_mono (x1 x2 : A) : - x1 = x2 -> (alloc x1 : ST ptr) <== alloc x2. + x1 = x2 -> + (alloc x1 : ST ptr) <== alloc x2. Proof. by move=>->. Qed. Lemma allocb_mono (x1 x2 : A) (n1 n2 : nat) : - x1 = x2 -> n1 = n2 -> (allocb x1 n1 : ST ptr) <== allocb x2 n2. + x1 = x2 -> + n1 = n2 -> + (allocb x1 n1 : ST ptr) <== allocb x2 n2. Proof. by move=>->->. Qed. Lemma dealloc_mono (p1 p2 : ptr) : - p1 = p2 -> (dealloc p1 : ST unit) <== dealloc p2. + p1 = p2 -> + (dealloc p1 : ST unit) <== dealloc p2. Proof. by move=>->. Qed. -Variables (G : Type) (C : A -> Type) (s : forall x, spec G (C x)). -Notation lat := (dfunLattice (fun x => [lattice of STspec (s x)])). +Variables (G : Type) (C : A -> Type) (pq : forall x, spec G (C x)). +Notation lat := (dfun_lattice (fun x => STspec (pq x))). -Lemma fix_mono (f1 f2 : lat -> lat) : f1 <== f2 -> (Fix f1 : lat) <== Fix f2. +Lemma fix_mono (f1 f2 : lat -> lat) : + f1 <== f2 -> + (ffix f1 : lat) <== ffix f2. Proof. -move=>Hf; apply: tarski_lfp_mono. -- move=>x1 x2 Hx; apply: supM=>z [x][H ->]; apply: supP; exists x. +move=>Hf; rewrite /ffix. +(* f' f2 is monotone closure of f2; hence it's monotone *) +have M : monofun_axiom (f' f2). +- move=>x1 x2 Hx; apply: supM=>z [x][H ->]; apply: supP; exists x. by split=>//; apply: poset_trans H Hx. +set ff := MonoFun.pack_ (isMonoFun.Build _ _ (f' f2) M). +apply: (tarski_lfp_mono (f2:=ff)). move=>y; apply: supM=>_ [x][H1 ->]. by apply: poset_trans (Hf x) _; apply: supP; exists x. Qed. End Monotonicity. -Section MonotonicityExamples. - Notation "'Do' e" := (@STprog _ _ _ e _) (at level 80). - -Notation "x '<--' c1 ';' c2" := (bind c1 (fun x => c2)) +Notation "x '<--' c1 ';' c2" := (bnd c1 (fun x => c2)) (at level 81, right associativity). -Notation "c1 ';;' c2" := (bind c1 (fun _ => c2)) +Notation "c1 ';;' c2" := (bnd c1 (fun _ => c2)) (at level 81, right associativity). Notation "'!' x" := (read x) (at level 50). Notation "x '::=' e" := (write x e) (at level 60). -Fixpoint fact (x : nat) := if x is x'.+1 then x * fact x' else 1. - -Definition facttp := forall x : nat, @STspec unit nat - (fun => (fun i => True, - fun v m => if v is Val w then w = fact x else False)). +(* Some examples *) -(* for the example, we need a quick lemma for function calls *) +(* the examples require a lemma for instantiating ghosts *) +(* when doing function calls *) (* this need not be exported, as it follows from the signature *) -(* will reprove it outside of the module *) +(* and will be reproven outside of the Vrf module *) Lemma gE G A (s : spec G A) g i (e : @STspec G A s) (Q : post A) : (s g).1 i -> @@ -1127,34 +1030,40 @@ case: e=>e /= /[apply] Hp Hv He; apply: vrfV=>V /=. by apply/vrf_post/Hp; case=>[v|ex] m Vm H; [apply: Hv | apply: He]. Qed. -(* now we can do the example with monotonicity explicitly *) +(* factorial *) +Fixpoint fact (x : nat) := if x is x'.+1 then x * fact x' else 1. +Definition facttp := forall x : nat, @STspec unit nat + (fun => (fun i => True, + fun v m => if v is Val w then w = fact x else False)). + +(* proof using ffix2, with monotonicity proved explicitly *) Program Definition factx := - Fix' (fun (loop : facttp) (x : nat) => + ffix2 (fun (loop : facttp) (x : nat) => Do (if x is x'.+1 then t <-- loop x'; ret (x * t) else ret 1)) _. Next Obligation. case=>i /= _; case: x; first by apply: vrf_ret. -by move=>n; apply: vrf_bind=>//; apply: gE=>// x m -> _; apply: vrf_ret. +by move=>n; apply: vrf_bnd=>//; apply: gE=>// x m -> _; apply: vrf_ret. Qed. Next Obligation. move=>x1 x2 H; case=>[|n]; first by apply: poset_refl. by apply: do_mono; apply: bind_mono=>//; apply: H. Qed. -(* Monotonocity works even if we compute *) +(* monotonocity can be proved even if we compute *) (* with propositions, and deliberately invert the polarity. *) Definition proptp := unit -> @STspec unit Prop (fun => (fun i => True, fun R m => True)). Program Definition propx := - Fix' (fun (loop : proptp) (x : unit) => + ffix2 (fun (loop : proptp) (x : unit) => Do (R <-- loop tt; ret (not R))) _. Next Obligation. -case=>i _; apply: vrf_bind=>//. +case=>i _; apply: vrf_bnd=>//. by apply: gE=>//= R m _ C; apply: vrf_ret. Qed. Next Obligation. @@ -1166,10 +1075,10 @@ Definition proptp2 := Prop -> @STspec unit Prop (fun => (fun i => True, fun (Ro : ans Prop) m => True)). Program Definition propx2 := - Fix' (fun (loop : proptp2) (x : Prop) => + ffix2 (fun (loop : proptp2) (x : Prop) => Do (R <-- loop (x -> x); ret R)) _ True. Next Obligation. -case=>i /= _; apply: vrf_bind=>//. +case=>i /= _; apply: vrf_bnd=>//. by apply: gE=>//= R m _ _; apply: vrf_ret. Qed. Next Obligation. @@ -1182,10 +1091,10 @@ Definition proptp3 := forall R : Prop, @STspec unit Prop fun (Ro : ans Prop) m => if Ro is Val Ro' then Ro' else False)). Program Definition propx3 := - Fix' (fun (loop : proptp3) (x : Prop) => + ffix2 (fun (loop : proptp3) (x : Prop) => Do (R <-- loop (x -> x); ret R)) _ True. Next Obligation. -case=>i /= _; apply: vrf_bind=>//. +case=>i /= _; apply: vrf_bnd=>//. by apply: gE=>//= R m X _; apply: vrf_ret. Qed. Next Obligation. @@ -1195,16 +1104,13 @@ Qed. (* It seems safe to say that monotonicity is always easy *) (* to prove for all the programs that we expect to write. *) -(* Thus, we will export Fix over monotone closure, but not Fix'. *) -(* Exporting Fix has the advantage of eliding the always simple *) -(* and syntax-directed proofs of monotonicity, which we just don't *) -(* want to bother with. *) +(* Thus, the module will export ffix over monotone closure, but not ffix2. *) +(* Exporting ffix has the advantage of eliding the always simple *) +(* and syntax-directed proofs of monotonicity. *) (* This is in line with the existing proofs of soundness of HTT as a *) -(* standalone theory, where you can show that all functions that can be *) -(* produced using the syntax are monotone in the model. That proof *) -(* relied on a significant subset of Coq, but still a subset. *) - +(* standalone theory, which shows that all functions that can be *) +(* produced using the declared syntax are monotone in the model. *) (* The paper is: *) (* Partiality, State and Dependent Types *) (* Kasper Svendsen, Lars Birkedal and Aleksandar Nanevski *) @@ -1212,31 +1118,28 @@ Qed. (* pages 198-212, 2011. *) (* The model in that paper, and the one we use here, are quite different *) -(* however, and we can't exclude concocting a *) -(* non-monotone program using the syntax we export here, *) +(* however, and one can't exclude concocting a *) +(* non-monotone program using the syntax exported here, *) (* based on the results of that other paper. *) -(* That said, we have so far failed to do produce such a non-monotone function. *) -(* But if such a case arises, we should switch to using Fix' *) +(* That said, so far there we failed to produce such a non-monotone function. *) +(* But if such a case arises, we should switch to using ffix2 *) (* and ask for explicit proofs of monotonicity. *) -(* For all the examples we did so far, such a proof is easily *) +(* For all the examples done so far, such a proof is easily *) (* constructed, along the above lines (moreover, such proofs ought to be automatable) *) (* Perhaps a proof, in the theory of Coq, that all functions are monotone *) -(* can be produced if we worked from parametricity properties of Coq, *) +(* can be produced if one worked from parametricity properties of Coq, *) (* which have been established in the past work, and even internalized *) (* into Coq by means of parametricity axioms. But that remains future work. *) -End MonotonicityExamples. - End Vrf. - Export Vrf. +(* some helpers *) Definition skip := ret tt. -Corollary vrf_mono A (e : ST A) i : monotone (vrf' e i). +Lemma vrf_mono A (e : ST A) i : monofun_axiom (vrf' e i). Proof. by move=>/= Q1 Q2 H; apply: vrf_post=>y m _; apply: H. Qed. - (******************************************) (* Notation for logical variable postexts *) (******************************************) @@ -1248,11 +1151,17 @@ Definition logvar {B A} (G : A -> Type) (s : forall x : A, spec (G x) B) : spec {x : A & G x} B := fun '(existT x g) => s x g. -Notation "'STsep' ( p , q ) " := (STspec unit (logbase p q)) (at level 0). +Notation "'STsep' ( p , q ) " := + (STspec unit (fun => (p, q))) (at level 0, + format "'[hv ' STsep ( '[' p , '/' q ']' ) ']'"). -Notation "{ x .. y }, 'STsep' ( p , q ) " := +Notation "'STsep' { x .. y } ( p , q )" := (STspec _ (logvar (fun x => .. (logvar (fun y => logbase p q)) .. ))) - (at level 0, x binder, y binder, right associativity). + (at level 0, x binder, y binder, right associativity, + format "'[hv ' STsep { x .. y } '/ ' ( '[' p , '/' q ']' ) ']'"). + +Notation "[> h1 , .. , hn <]" := + (existT _ h1 .. (existT _ hn tt) .. ) (at level 0). (************************************************************) (* Lemmas for pulling out and instantiating ghost variables *) @@ -1261,11 +1170,11 @@ Notation "{ x .. y }, 'STsep' ( p , q ) " := (* Lemmas without framing, i.e. they pass the entire heap to the *) (* routine being invoked. *) -Lemma gE G A (s : spec G A) g i (e : STspec G s) (Q : post A) : - (s g).1 i -> - (forall v m, (s g).2 (Val v) m -> +Lemma gE G A (pq : spec G A) (e : STspec G pq) g (Q : post A) i : + (pq g).1 i -> + (forall v m, (pq g).2 (Val v) m -> valid m -> Q (Val v) m) -> - (forall x m, (s g).2 (Exn x) m -> + (forall x m, (pq g).2 (Exn x) m -> valid m -> Q (Exn x) m) -> vrf i e Q. Proof. @@ -1273,49 +1182,50 @@ case: e=>e /= /[apply] Hp Hv He; apply: vrfV=>V /=. by apply/vrf_post/Hp; case=>[v|ex] m Vm H; [apply: Hv | apply: He]. Qed. -Arguments gE [G A s] g [i e Q] _ _. - -Notation "[gE]" := (gE tt) (at level 0). +Arguments gE {G A pq e} g {Q}. +Notation "[gE]" := (gE tt) (at level 0). Notation "[ 'gE' x1 , .. , xn ]" := (gE (existT _ x1 .. (existT _ xn tt) ..)) - (at level 0, format "[ 'gE' x1 , .. , xn ]"). - -(* a combination of gE + vrf_bind, for "stepping over" the call *) -Lemma stepE G A B (s : spec G A) g i (e : STspec G s) (e2 : A -> ST B) (Q : post B) : - (s g).1 i -> - (forall x m, (s g).2 (Val x) m -> vrf m (e2 x) Q) -> - (forall x m, (s g).2 (Exn x) m -> +(at level 0, format "[ 'gE' x1 , .. , xn ]"). + +(* combination of gE + vrf_bind, for "stepping over" the call *) +Lemma stepE G A B (pq : spec G A) (e : STspec G pq) + (e2 : A -> ST B) i (Q : post B) (g : G) : + (pq g).1 i -> + (forall x m, (pq g).2 (Val x) m -> vrf m (e2 x) Q) -> + (forall x m, (pq g).2 (Exn x) m -> valid m -> Q (Exn x) m) -> - vrf i (bind e e2) Q. + vrf i (bnd e e2) Q. Proof. -move=>Hp Hv He. -by apply/vrf_bind/(gE _ Hp)=>[v m P|x m P _] V; [apply: Hv | apply: He]. +move=>Hp Hv He; apply/vrf_bnd/(gE _ _ Hp). +- by move=>v m P V; apply: Hv. +by move=>x m P _ V; apply: He. Qed. -Arguments stepE [G A B s] g [i e e2 Q] _ _. +Arguments stepE {G A B pq e e2 i Q}. Notation "[stepE]" := (stepE tt) (at level 0). - Notation "[ 'stepE' x1 , .. , xn ]" := (stepE (existT _ x1 .. (existT _ xn tt) ..)) - (at level 0, format "[ 'stepE' x1 , .. , xn ]"). - -(* a combination of gE + vrf_try *) -Lemma tryE G A B (s : spec G A) g i (e : STspec G s) (e1 : A -> ST B) (e2 : exn -> ST B) (Q : post B) : - (s g).1 i -> - (forall x m, (s g).2 (Val x) m -> vrf m (e1 x) Q) -> - (forall x m, (s g).2 (Exn x) m -> vrf m (e2 x) Q) -> +(at level 0, format "[ 'stepE' x1 , .. , xn ]"). + +(* combination of gE + vrf_try *) +Lemma tryE G A B (pq : spec G A) (e : STspec G pq) (e1 : A -> ST B) + (e2 : exn -> ST B) i (Q : post B) (g : G) : + (pq g).1 i -> + (forall x m, (pq g).2 (Val x) m -> vrf m (e1 x) Q) -> + (forall x m, (pq g).2 (Exn x) m -> vrf m (e2 x) Q) -> vrf i (try e e1 e2) Q. Proof. -move=>Hp Hv Hx. -by apply/vrf_try/(gE _ Hp)=>[x|ex] m Vm P; [apply: Hv | apply: Hx]. +move=>Hp Hv Hx; apply/vrf_try/(gE _ _ Hp). +- by move=>x m Vm P; apply: Hv. +by move=>ex m Vm P; apply: Hx. Qed. -Arguments tryE [G A B s] g [i e e1 e2 Q] _ _. +Arguments tryE {G A B pq e e1 e2 i Q}. Notation "[tryE]" := (tryE tt) (at level 0). - Notation "[ 'tryE' x1 , .. , xn ]" := (tryE (existT _ x1 .. (existT _ xn tt) ..)) (at level 0, format "[ 'tryE' x1 , .. , xn ]"). @@ -1324,11 +1234,11 @@ Notation "[ 'tryE' x1 , .. , xn ]" := (* empty heap to the routine. For more sophisticated framing *) (* variants see the `heapauto` module. *) -Lemma gU G A (s : spec G A) g i (e : STspec G s) (Q : post A) : - (s g).1 Unit -> - (forall v m, (s g).2 (Val v) m -> +Lemma gU G A (pq : spec G A) (e : STspec G pq) i (Q : post A) g : + (pq g).1 Unit -> + (forall v m, (pq g).2 (Val v) m -> valid (m \+ i) -> Q (Val v) (m \+ i)) -> - (forall x m, (s g).2 (Exn x) m -> + (forall x m, (pq g).2 (Exn x) m -> valid (m \+ i) -> Q (Exn x) (m \+ i)) -> vrf i e Q. Proof. @@ -1338,69 +1248,62 @@ by case=>[x|ex] n _ =>[/Hv|/Hx]. Qed. Notation "[gU]" := (gU tt) (at level 0). - Notation "[ 'gU' x1 , .. , xn ]" := (gU (existT _ x1 .. (existT _ xn tt) ..)) (at level 0, format "[ 'gU' x1 , .. , xn ]"). -(* a combination of gU + vrf_bind *) -Lemma stepU G A B (s : spec G A) g i (e : STspec G s) (e2 : A -> ST B) - (Q : post B) : - (s g).1 Unit -> - (forall x m, (s g).2 (Val x) m -> vrf (m \+ i) (e2 x) Q) -> - (forall x m, (s g).2 (Exn x) m -> +(* combination of gU + vrf_bind *) +Lemma stepU G A B (pq : spec G A) (e : STspec G pq) (e2 : A -> ST B) + i (Q : post B) g : + (pq g).1 Unit -> + (forall x m, (pq g).2 (Val x) m -> vrf (m \+ i) (e2 x) Q) -> + (forall x m, (pq g).2 (Exn x) m -> valid (m \+ i) -> Q (Exn x) (m \+ i)) -> - vrf i (bind e e2) Q. + vrf i (bnd e e2) Q. Proof. -move=>Hp Hv Hx. -apply/vrf_bind/(gU _ Hp)=>[x m H V|ex m H V _]. -- by apply: Hv H. -by apply: Hx. +move=>Hp Hv Hx; apply/vrf_bnd/(gU _ Hp). +- by move=>x m H V; apply: Hv H. +by move=>ex m H V _; apply: Hx. Qed. -Arguments stepU [G A B s] g i [e e2 Q] _ _ _. +Arguments stepU {G A B pq e e2 i Q}. Notation "[stepU]" := (stepU tt) (at level 0). - Notation "[ 'stepU' x1 , .. , xn ]" := (stepU (existT _ x1 .. (existT _ xn tt) ..)) (at level 0, format "[ 'stepU' x1 , .. , xn ]"). -(* a combination of gU + vrf_try *) -Lemma tryU G A B (s : spec G A) g i (e : STspec G s) - (e1 : A -> ST B) (e2 : exn -> ST B) (Q : post B) : - (s g).1 Unit -> - (forall x m, (s g).2 (Val x) m -> vrf (m \+ i) (e1 x) Q) -> - (forall x m, (s g).2 (Exn x) m -> vrf (m \+ i) (e2 x) Q) -> +(* combination of gU + vrf_try *) +Lemma tryU G A B (pq : spec G A) (e : STspec G pq) + (e1 : A -> ST B) (e2 : exn -> ST B) i (Q : post B) g : + (pq g).1 Unit -> + (forall x m, (pq g).2 (Val x) m -> vrf (m \+ i) (e1 x) Q) -> + (forall x m, (pq g).2 (Exn x) m -> vrf (m \+ i) (e2 x) Q) -> vrf i (try e e1 e2) Q. Proof. -move=>Hi H1 H2. -apply/vrf_try/(gU _ Hi)=>[x|ex] m H V. -- by apply: H1 H. -by apply: H2. +move=>Hi H1 H2; apply/vrf_try/(gU _ Hi). +- by move=>x m H V; apply: H1 H. +by move=>ex m H V; apply: H2. Qed. -Arguments tryU [G A B s] g i [e e1 e2 Q] _ _ _. +Arguments tryU {G A B pq e e1 e2 i Q}. Notation "[tryU]" := (tryU tt) (at level 0). - Notation "[ 'tryU' x1 , .. , xn ]" := (tryU (existT _ x1 .. (existT _ xn tt) ..)) (at level 0, format "[ 'tryU' x1 , .. , xn ]"). -(* some notation for writing posts that signify no exceptions are raised *) +(* notation for writing posts that signify *) +(* that no exception is raised *) Definition vfun' A (f : A -> heap -> Prop) : post A := fun y i => if y is Val v then f v i else False. Notation "[ 'vfun' x => p ]" := (vfun' (fun x => p)) (at level 0, x name, format "[ 'vfun' x => p ]") : fun_scope. - Notation "[ 'vfun' x : aT => p ]" := (vfun' (fun (x : aT) => p)) (at level 0, x name, only parsing) : fun_scope. - Notation "[ 'vfun' x i => p ]" := (vfun' (fun x i => p)) (at level 0, x name, format "[ 'vfun' x i => p ]") : fun_scope. - Notation "[ 'vfun' ( x : aT ) i => p ]" := (vfun' (fun (x : aT) i => p)) (at level 0, x name, only parsing) : fun_scope. From cd47cf28f4c38a5efd07c5d35d903893d8c7ee8c Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Fri, 2 Aug 2024 18:33:10 +0200 Subject: [PATCH 09/93] continuing refactoring --- _CoqProject | 2 +- examples/bubblesort.v | 114 ++++++++++++++++++++++++------------------ examples/quicksort.v | 98 +++++++++++++++++++++++------------- examples/stack.v | 29 ++++++++--- htt/heapauto.v | 12 +++-- htt/options.v | 10 +++- pcm/autopcm.v | 8 ++- 7 files changed, 173 insertions(+), 100 deletions(-) diff --git a/_CoqProject b/_CoqProject index a4e3600..f94a95a 100644 --- a/_CoqProject +++ b/_CoqProject @@ -1,7 +1,7 @@ -Q core pcm -Q pcm pcm --Q htt htt -Q examples htt +-Q htt htt -docroot docs # where the documentation should go -arg -w -arg -notation-overridden diff --git a/examples/bubblesort.v b/examples/bubblesort.v index 6fe6e19..79e73cd 100644 --- a/examples/bubblesort.v +++ b/examples/bubblesort.v @@ -1,7 +1,21 @@ +(* +Copyright 2022 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 path fintype tuple finfun fingroup perm interval. +From mathcomp Require Import ssrnat eqtype seq path fintype tuple. +From mathcomp Require Import finfun fingroup perm interval. From pcm Require Import options axioms prelude pred ordtype slice. -From pcm Require Import pcm unionmap heap. +From pcm Require Import seqext pcm unionmap heap. From htt Require Import options model heapauto. From htt Require Import array. @@ -95,7 +109,7 @@ set i0 := Wo i; set i1 := So i. rewrite {1}codomE (enum_split i0) /= {2}(enum_split i1) /= /heap.indx (index_enum_ord i0) (index_enum_ord i1) drop_cat size_take size_enum_ord ltn_ord So_eq ltnn subnn /= map_cat /= map_take map_drop -codomE. -by rewrite /slice /= addn0 addn1 /= drop0 take_size. +by rewrite /slice.slice/= addn0 addn1 /= drop0 take_size. Qed. (* TODO generalize? *) @@ -229,8 +243,8 @@ Lemma swnx_split (f : {ffun 'I_n.+1 -> A}) (i : 'I_n) : [:: f (So i); f (Wo i)] ++ &:(fgraph f) `]i.+1, +oo[. Proof. -rewrite /= (codom1_split _ i) /=; set i0 := Wo i; set i1 := So i. -by rewrite swnx_ao // swnx_oa //= swnxE1 swnxE2. +rewrite [LHS](codom1_split _ i) /=; set i0 := Wo i; set i1 := So i. +by rewrite swnx_ao // swnx_oa // swnxE1 swnxE2. Qed. Lemma swnx_ux_rcons (f : {ffun 'I_n.+1 -> A}) (i : 'I_n) : @@ -319,8 +333,8 @@ Opaque Array.write Array.read. (*****************************************************) Program Definition cas_next (a : {array 'I_n.+2 -> A}) (i : 'I_n.+1) : - {f : {ffun 'I_n.+2 -> A}}, - STsep (Array.shape a f, + STsep {f : {ffun 'I_n.+2 -> A}} + (Array.shape a f, [vfun (b : bool) h => exists (p : 'S_n.+2), let i0 := Wo i in let i1 := So i in h \In Array.shape a (pffun p f) /\ @@ -353,8 +367,8 @@ Opaque cas_next. Definition bubble_pass_loop (a : {array 'I_n.+2 -> A}) := forall isw : 'I_n.+1 * bool, - {f : {ffun 'I_n.+2 -> A}}, - STsep (fun h => h \In Array.shape a f /\ + STsep {f : {ffun 'I_n.+2 -> A}} + (fun h => h \In Array.shape a f /\ (* if no swaps have happened yet, the prefix is sorted *) ~~ isw.2 ==> sorted oleq (&:(fgraph f) `]-oo, isw.1 : nat]), [vfun (b : bool) h => @@ -364,13 +378,13 @@ Definition bubble_pass_loop (a : {array 'I_n.+2 -> A}) := ~~ b ==> (p == 1%g) && sorted oleq (fgraph f)]]). Program Definition bubble_pass (a : {array 'I_n.+2 -> A}) : - {f : {ffun 'I_n.+2 -> A}}, - STsep (Array.shape a f, + STsep {f : {ffun 'I_n.+2 -> A}} + (Array.shape a f, [vfun (b : bool) h => exists (p : 'S_n.+2), h \In Array.shape a (pffun p f) /\ ~~ b ==> (p == 1%g) && sorted oleq (fgraph f)]) := - Do (let go := Fix (fun (loop : bubble_pass_loop a) '(i, sw) => + Do (let go := ffix (fun (loop : bubble_pass_loop a) '(i, sw) => Do (sw0 <-- cas_next a i; let sw1 := sw || sw0 in if decP (b := i < n) idP is left pf @@ -417,15 +431,15 @@ Opaque bubble_pass. Definition bubble_sort_loop (a : {array 'I_n.+2 -> A}) : Type := unit -> - {f : {ffun 'I_n.+2 -> A}}, - STsep (Array.shape a f, + STsep {f : {ffun 'I_n.+2 -> A}} + (Array.shape a f, [vfun (_ : unit) h => exists (p : 'S_n.+2), h \In Array.shape a (pffun p f) /\ sorted oleq (fgraph (pffun p f))]). Program Definition bubble_sort (a : {array 'I_n.+2 -> A}) : bubble_sort_loop a := - Fix (fun (go : bubble_sort_loop a) _ => + ffix (fun (go : bubble_sort_loop a) _ => Do (sw <-- bubble_pass a; if sw then go tt : ST _ @@ -473,8 +487,8 @@ Opaque cas_next. Definition bubble_pass_opt_loop (a : {array 'I_n.+2 -> A}) (k : 'I_n.+1) := forall isw : 'I_n.+1 * bool, - {f : {ffun 'I_n.+2 -> A}}, - STsep (fun h => [/\ h \In Array.shape a f, + STsep {f : {ffun 'I_n.+2 -> A}} + (fun h => [/\ h \In Array.shape a f, (* all elements before f(k+2) are smaller than it *) allrel oleq (&:(fgraph f) `]-oo, k.+1]) (&:(fgraph f) `]k.+1, +oo[), @@ -496,8 +510,8 @@ Definition bubble_pass_opt_loop (a : {array 'I_n.+2 -> A}) (k : 'I_n.+1) := else (p == 1%g) && sorted oleq (fgraph f)]]). Program Definition bubble_pass_opt (a : {array 'I_n.+2 -> A}) (k : 'I_n.+1) : - {f : {ffun 'I_n.+2 -> A}}, - STsep (fun h => [/\ h \In Array.shape a f, + STsep {f : {ffun 'I_n.+2 -> A}} + (fun h => [/\ h \In Array.shape a f, allrel oleq (&:(fgraph f) `]-oo, k.+1]) (&:(fgraph f) `]k.+1, +oo[) & sorted oleq (&:(fgraph f) `]k.+1, +oo[)], @@ -509,17 +523,16 @@ Program Definition bubble_pass_opt (a : {array 'I_n.+2 -> A}) (k : 'I_n.+1) : (&:(fgraph f') `]k : nat, +oo[) && sorted oleq (&:(fgraph f') `]k : nat, +oo[) else (p == 1%g) && sorted oleq (fgraph f)]) := - Do (let go := Fix (fun (loop : bubble_pass_opt_loop a k) '(i, sw) => + Do (let go := ffix (fun (loop : bubble_pass_opt_loop a k) '(i, sw) => Do (sw0 <-- cas_next a i; let sw1 := sw || sw0 in if decP (b := i < k) idP is left pf - then loop (Sbo (i:=i) pf, sw1) - else ret sw1)) + then loop (Sbo (i:=i) pf, sw1) + else ret sw1)) in go (ord0, false)). Next Obligation. (* Sbo is always safe *) -move=>_ k _ _ i _ _ /= E; rewrite E ltnS; symmetry. -by apply: (leq_trans E); rewrite -ltnS; apply: ltn_ord. +by move=>_ k _ _ i _ _ /= E; rewrite E ltnS (leq_trans E) // -ltnS ltn_ord. Qed. Next Obligation. move=>a k loop _ i sw [f][] h /= [H Hak Hsk Hik Hai Hsi]. @@ -587,8 +600,8 @@ Opaque bubble_pass_opt. Definition bubble_sort_opt_loop (a : {array 'I_n.+2 -> A}) : Type := forall (k : 'I_n.+1), - {f : {ffun 'I_n.+2 -> A}}, - STsep (fun h => [/\ h \In Array.shape a f, + STsep {f : {ffun 'I_n.+2 -> A}} + (fun h => [/\ h \In Array.shape a f, allrel oleq (&:(fgraph f) `]-oo, k.+1]) (&:(fgraph f) `]k.+1, +oo[) & sorted oleq (&:(fgraph f) `]k.+1, +oo[)], @@ -598,13 +611,13 @@ Definition bubble_sort_opt_loop (a : {array 'I_n.+2 -> A}) : Type := sorted oleq (fgraph f')]). Program Definition bubble_sort_opt (a : {array 'I_n.+2 -> A}) : - {f : {ffun 'I_n.+2 -> A}}, - STsep (Array.shape a f, + STsep {f : {ffun 'I_n.+2 -> A}} + (Array.shape a f, [vfun (_ : unit) h => exists (p : 'S_n.+2), h \In Array.shape a (pffun p f) /\ sorted oleq (fgraph (pffun p f))]) := - Do (let go := Fix (fun (loop : bubble_sort_opt_loop a) k => + Do (let go := ffix (fun (loop : bubble_sort_opt_loop a) k => Do (sw <-- bubble_pass_opt a k; if sw then loop (Po k) : ST _ @@ -668,8 +681,8 @@ Opaque cas_next. Definition bubble_pass_opt2_loop (a : {array 'I_n.+2 -> A}) (k : 'I_n.+1) := forall ils : 'I_n.+1 * 'I_n.+2, - {f : {ffun 'I_n.+2 -> A}}, - STsep (fun h => [/\ h \In Array.shape a f, + STsep {f : {ffun 'I_n.+2 -> A}} + (fun h => [/\ h \In Array.shape a f, allrel oleq (&:(fgraph f) `]-oo, k.+1]) (&:(fgraph f) `]k.+1, +oo[), sorted oleq (&:(fgraph f) `]k.+1, +oo[), @@ -686,8 +699,8 @@ Definition bubble_pass_opt2_loop (a : {array 'I_n.+2 -> A}) (k : 'I_n.+1) := sorted oleq (&:(fgraph f') `[j : nat, +oo[)]]). Program Definition bubble_pass_opt2 (a : {array 'I_n.+2 -> A}) (k : 'I_n.+1) : - {f : {ffun 'I_n.+2 -> A}}, - STsep (fun h => [ /\ h \In Array.shape a f, + STsep {f : {ffun 'I_n.+2 -> A}} + (fun h => [ /\ h \In Array.shape a f, allrel oleq (&:(fgraph f) `]-oo, k.+1]) (&:(fgraph f) `]k.+1, +oo[) & sorted oleq (&:(fgraph f) `]k.+1, +oo[)], @@ -697,7 +710,7 @@ Program Definition bubble_pass_opt2 (a : {array 'I_n.+2 -> A}) (k : 'I_n.+1) : allrel oleq (&:(fgraph f') `]-oo, j : nat[) (&:(fgraph f') `[j : nat, +oo[) & sorted oleq (&:(fgraph f') `[j : nat, +oo[)]]) := - Do (let go := Fix (fun (loop : bubble_pass_opt2_loop a k) '(i, ls) => + Do (let go := ffix (fun (loop : bubble_pass_opt2_loop a k) '(i, ls) => Do (sw <-- cas_next a i; let ls1 := if sw then So i else ls in if decP (b := i < k) idP is left pf @@ -706,8 +719,7 @@ Program Definition bubble_pass_opt2 (a : {array 'I_n.+2 -> A}) (k : 'I_n.+1) : in go (ord0, ord0)). Next Obligation. (* show that Sbo is always safe *) -move=>_ k _ _ i _ _ /= E; rewrite E ltnS; symmetry. -by apply: (leq_trans E); rewrite -ltnS; apply: ltn_ord. +by move=>_ k _ _ i _ _ /= E; rewrite E ltnS (leq_trans E) // -ltnS ltn_ord. Qed. Next Obligation. move=>a k loop _ i ls [f][] h /= [Hs Hak Hsk /andP [Hils Hik] Hai Hsi]. @@ -741,7 +753,8 @@ apply: [stepE f]=>//= sw m [p][Hm Hsw]; case: decP=>H. suff: all (oleq^~ (f (Wo i))) (&:(codom f) `]-oo, ls:nat[). - by apply/sub_all=>z /otrans; apply. by move: Hai; rewrite codom1_ax_rcons // allrel_rconsr; case/andP. - rewrite codom1_ax_rcons2 // (sorted_rconsE _ _ (@otrans A)) -{2}codom1_ax_rcons // Hsi andbT. + rewrite codom1_ax_rcons2 //. + rewrite (sorted_rconsE _ _ (@otrans A)) -{2}codom1_ax_rcons // Hsi andbT. rewrite all_rcons Hf /=. move: Hsi; rewrite codom1_ax_rcons // (sorted_rconsE _ _ (@otrans A)). by case/andP=>+ _; apply/sub_all=>z Hz; apply/otrans/Hf. @@ -761,8 +774,9 @@ step=>Vm; exists p; case: sw Hsw=>/=; case=>Ep Hf. - rewrite Ep slice_oSR swnx_ux_rcons swnx_xu_cons allrel_rconsl allrel_consr /=. move: Hak; rewrite codom1_ax_rcons2 // !allrel_rconsl -/i0 -/i1 -andbA. case/and3P=>-> _ ->; rewrite (ordW Hf) /= !andbT. - move: Hai; rewrite codom1_ax_rcons //= allrel_rconsr; case/andP=>_ Hals. - move: Hsi; rewrite codom1_ax_rcons // sorted_rconsE //; case/andP =>Halsi _. + move: Hai; rewrite codom1_ax_rcons //= allrel_rconsr=>/andP [_ Hals]. + move: Hsi; rewrite codom1_ax_rcons // sorted_rconsE; last by apply: otrans. + case/andP => Halsi _. move: Hils; rewrite leq_eqVlt; case/orP=>[/eqP <-|Hlsi] //. rewrite (slice_split (x:=ls) _ true); last by rewrite in_itv. by rewrite /= all_cat Hals. @@ -770,18 +784,22 @@ step=>Vm; exists p; case: sw Hsw=>/=; case=>Ep Hf. by move: Hak; rewrite codom1_ax_rcons2 // !allrel_rconsl -!andbA; case/and3P. (* swap didn't happen on last iteration *) rewrite {p}Ep pffunE1 /= in Hm *. -rewrite (slice_split (x:=i) (i:=`[ls:nat, +oo[) _ false); last by rewrite in_itv /= andbT. +rewrite (slice_split (x:=i) (i:=`[ls:nat, +oo[) _ false); last first. +- by rewrite in_itv /= andbT. split=>//=. -- rewrite allrel_catr Hai /= -slice_oSL codom1_xu_cons allrel_consr; apply/andP; split. +- rewrite allrel_catr Hai /= -slice_oSL codom1_xu_cons allrel_consr. + apply/andP; split. - suff: all (oleq^~ (f (Wo i))) (&:(codom f) `]-oo, ls:nat[). - by apply/sub_all=>z /otrans; apply. apply/sub_all/Hai=>z /allP; apply. by rewrite codom1_ax_rcons // mem_rcons inE eqxx. apply/allrel_sub_l/Hak/slice_subset. by rewrite subitvE /= bnd_simp leEnat; apply/ltnW. -rewrite (sorted_pairwise (@otrans A)) pairwise_cat -!(sorted_pairwise (@otrans A)) Hsi /=. +rewrite (sorted_pairwise (@otrans A)) pairwise_cat. +rewrite -!(sorted_pairwise (@otrans A)) Hsi /=. rewrite -slice_oSL codom1_xu_cons /= allrel_consr -andbA; apply/and3P; split. -- move: Hsi; rewrite (sorted_pairwise (@otrans A)) codom1_ax_rcons // pairwise_rcons. +- move: Hsi. + rewrite (sorted_pairwise (@otrans A)) codom1_ax_rcons // pairwise_rcons. case/andP=>H' _; rewrite all_rcons Hf /=. by apply/sub_all/H'=>z Hz; apply/otrans/Hf. - apply/allrel_sub_l/Hak/slice_subset. @@ -800,8 +818,8 @@ Opaque bubble_pass_opt2. Definition bubble_sort_opt2_loop (a : {array 'I_n.+2 -> A}) : Type := forall (k : 'I_n.+1), - {f : {ffun 'I_n.+2 -> A}}, - STsep (fun h => [/\ h \In Array.shape a f, + STsep {f : {ffun 'I_n.+2 -> A}} + (fun h => [/\ h \In Array.shape a f, allrel oleq (&:(fgraph f) `]-oo, k.+1]) (&:(fgraph f) `]k.+1, +oo[) & sorted oleq (&:(fgraph f) `]k.+1, +oo[)], @@ -811,13 +829,13 @@ Definition bubble_sort_opt2_loop (a : {array 'I_n.+2 -> A}) : Type := sorted oleq (fgraph f')]). Program Definition bubble_sort_opt2 (a : {array 'I_n.+2 -> A}) : - {f : {ffun 'I_n.+2 -> A}}, - STsep (Array.shape a f, + STsep {f : {ffun 'I_n.+2 -> A}} + (Array.shape a f, [vfun (_ : unit) h => exists (p : 'S_n.+2), let f' := pffun p f in h \In Array.shape a f' /\ sorted oleq (fgraph f')]) := - Do (let go := Fix (fun (loop : bubble_sort_opt2_loop a) k => + Do (let go := ffix (fun (loop : bubble_sort_opt2_loop a) k => Do (k1 <-- bubble_pass_opt2 a k; if decP (b := 1 < k1) idP is left pf then loop (M2lo (i:=k1) pf) : ST _ diff --git a/examples/quicksort.v b/examples/quicksort.v index 0e6bbba..6e0af19 100644 --- a/examples/quicksort.v +++ b/examples/quicksort.v @@ -1,8 +1,21 @@ +(* +Copyright 2022 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 mathcomp Require Import ssreflect ssrbool ssrfun. From mathcomp Require Import ssrnat eqtype seq path fintype tuple finfun finset. From mathcomp Require Import fingroup perm. From mathcomp Require Import interval. -From pcm Require Import options axioms prelude pred ordtype slice. +From pcm Require Import options axioms prelude seqext pred ordtype slice. From pcm Require Import pcm unionmap heap. From htt Require Import options model heapauto. From htt Require Import array. @@ -13,7 +26,9 @@ Import order.Order.NatOrder order.Order.TTheory. (* TODO added to mathcomp > 1.16 *) Lemma perm_on_id {T : finType} (u : {perm T}) (S : {set T}) : - perm_on S u -> #|S| <= 1 -> u = 1%g. + perm_on S u -> + #|S| <= 1 -> + u = 1%g. Proof. rewrite leq_eqVlt ltnS leqn0 => pSu S10; apply/permP => t; rewrite perm1. case/orP : S10; last first. @@ -24,9 +39,10 @@ by apply/eqP; have := perm_closed x pSu; rewrite Sx !inE => ->. Qed. Lemma perm_onC {T : finType} (S1 S2 : {set T}) (u1 u2 : {perm T}) : - perm_on S1 u1 -> perm_on S2 u2 -> - [disjoint S1 & S2] -> - commute u1 u2. + perm_on S1 u1 -> + perm_on S2 u2 -> + [disjoint S1 & S2] -> + commute u1 u2. Proof. move=> pS1 pS2 S12; apply/permP => t; rewrite !permM. case/boolP : (t \in S1) => tS1. @@ -38,19 +54,21 @@ case/boolP : (t \in S2) => tS2. by rewrite (out_perm pS1) // (out_perm pS2) // (out_perm pS1). Qed. -Lemma tperm_on {T : finType} (x y : T) : perm_on [set x; y] (tperm x y). +Lemma tperm_on {T : finType} (x y : T) : + perm_on [set x; y] (tperm x y). Proof. by apply/subsetP => z /[!inE]; case: tpermP => [->|->|]; rewrite eqxx // orbT. Qed. (* TODO added to fcsl-pcm *) -Corollary slice_oPR {A : Type} a x (s : seq A) : +Lemma slice_oPR {A : Type} a x (s : seq A) : 0 < x -> &:s (Interval a (BRight x.-1)) = &:s (Interval a (BLeft x)). Proof. by move=>Hx; rewrite -{2}(prednK Hx) slice_oSR. Qed. -Lemma slice_FR {A : Type} (s : seq A) x : &:s (Interval x +oo) = &:s (Interval x (BLeft (size s))). -Proof. by rewrite /slice /= addn0. Qed. +Lemma slice_FR {A : Type} (s : seq A) x : + &:s (Interval x +oo) = &:s (Interval x (BLeft (size s))). +Proof. by rewrite /slice.slice /= addn0. Qed. Lemma slice_extrude {A : Type} (s : seq A) (i : interval nat) : order.Order.lt i.1 i.2 -> @@ -258,8 +276,8 @@ Variable (n : nat) (A : ordType). Opaque Array.write Array.read. Program Definition swap (a : {array 'I_n -> A}) (i j : 'I_n) : - {f : {ffun 'I_n -> A}}, - STsep (Array.shape a f, + STsep {f : {ffun 'I_n -> A}} + (Array.shape a f, [vfun _ h => h \In Array.shape a (pffun (tperm i j) f)]) := Do (if i == j then skip else @@ -281,12 +299,13 @@ Qed. Opaque swap. -Definition partition_lm_loop (a : {array 'I_n -> A}) (pivot : A) (lo hi : 'I_n) := +Definition partition_lm_loop (a : {array 'I_n -> A}) (pivot : A) + (lo hi : 'I_n) := forall ij : sigT (fun i : 'I_n => sig (fun j : 'I_n => i <= j /\ j < hi)), let i := projT1 ij in let j := proj1_sig (projT2 ij) in - {f : {ffun 'I_n -> A}}, - STsep (fun h => [/\ h \In Array.shape a f, + STsep {f : {ffun 'I_n -> A}} + (fun h => [/\ h \In Array.shape a f, lo <= i, all (oleq^~ pivot) (&:(fgraph f) `[lo : nat, i : nat[) & all (ord pivot) (&:(fgraph f) `[i : nat, j : nat[)], @@ -298,9 +317,10 @@ Definition partition_lm_loop (a : {array 'I_n -> A}) (pivot : A) (lo hi : 'I_n) all (oleq^~ pivot) (&:(fgraph f') `[lo : nat, v : nat[) & all (ord pivot) (&:(fgraph f') `[v : nat, hi : nat[)]]). -Program Definition partition_lm_pass (a : {array 'I_n -> A}) (pivot : A) (lo hi : 'I_n) (Hlo : lo < hi): - {f : {ffun 'I_n -> A}}, - STsep (Array.shape a f, +Program Definition partition_lm_pass (a : {array 'I_n -> A}) (pivot : A) + (lo hi : 'I_n) (Hlo : lo < hi): + STsep {f : {ffun 'I_n -> A}} + (Array.shape a f, [vfun (v : 'I_n) h => lo <= v <= hi /\ exists p, let f' := pffun p f in @@ -309,7 +329,7 @@ Program Definition partition_lm_pass (a : {array 'I_n -> A}) (pivot : A) (lo hi all (oleq^~ pivot) (&:(fgraph f') `[lo : nat, v : nat[) & all (ord pivot) (&:(fgraph f') `[v : nat, hi : nat[)]]) := Do (let go := - Fix (fun (loop : partition_lm_loop a pivot lo hi) '(existT i (exist j (conj Hi Hj))) => + ffix (fun (loop : partition_lm_loop a pivot lo hi) '(existT i (exist j (conj Hi Hj))) => Do (x <-- Array.read a j; if oleq x pivot then swap a i j;; @@ -402,9 +422,10 @@ Qed. Opaque partition_lm_pass. -Program Definition partition_lm (a : {array 'I_n -> A}) (lo hi : 'I_n) (Hlo : lo < hi): - {f : {ffun 'I_n -> A}}, - STsep (Array.shape a f, +Program Definition partition_lm (a : {array 'I_n -> A}) + (lo hi : 'I_n) (Hlo : lo < hi): + STsep {f : {ffun 'I_n -> A}} + (Array.shape a f, [vfun (v : 'I_n) h => lo <= v <= hi /\ exists p, let f' := pffun p f in @@ -457,8 +478,8 @@ Definition quicksort_lm_loop (a : {array 'I_n.+1 -> A}) := forall (lohi : 'I_n.+1 * 'I_n.+1), let lo := lohi.1 in let hi := lohi.2 in - {f : {ffun 'I_n.+1 -> A}}, - STsep (Array.shape a f, + STsep {f : {ffun 'I_n.+1 -> A}} + (Array.shape a f, [vfun (_ : unit) h => exists p, let f' := pffun p f in [/\ perm_on [set ix : 'I_n.+1 | lo <= ix <= hi] p, @@ -466,14 +487,14 @@ Definition quicksort_lm_loop (a : {array 'I_n.+1 -> A}) := sorted oleq (&:(fgraph f') `[lo : nat, hi : nat])]]). Program Definition quicksort_lm (a : {array 'I_n.+1 -> A}) : - {f : {ffun 'I_n.+1 -> A}}, - STsep (Array.shape a f, + STsep {f : {ffun 'I_n.+1 -> A}} + (Array.shape a f, [vfun (_ : unit) h => exists p, let f' := pffun p f in h \In Array.shape a f' /\ sorted oleq (fgraph f')]) := Do (let go := - Fix (fun (loop : quicksort_lm_loop a) '(l,h) => + ffix (fun (loop : quicksort_lm_loop a) '(l,h) => Do (if decP (b:=(l : nat) < h) idP isn't left pf then skip else v <-- partition_lm a pf; loop (l, Po v);; @@ -541,13 +562,16 @@ move: (ltn_ord v); rewrite ltnS leq_eqVlt; case/orP=>[/eqP Ev|Nv]. by rewrite !inE -eqn_leq =>/eqP E; apply/eqP/ord_inj. move: Sl Hpl; rewrite Eh Ev Epr mul1g => Sl Hpl. rewrite slice_xR; last by rewrite bnd_simp leEnat; move: Hvl; rewrite Ev. - rewrite {22}(_ : n = (ord_max : 'I_n.+1)) // onth_codom /= sorted_rconsE //. + rewrite {22}(_ : n = (ord_max : 'I_n.+1)) // onth_codom /= sorted_rconsE //; + last by apply: otrans. move: Sl; rewrite slice_oPR; last by rewrite lt0n -Ev. move=>->; rewrite andbT; move: Al; rewrite Ev. have ->: pffun (pl * p) f ord_max = pffun p f ord_max. - - by rewrite !ffunE permM (out_perm Hpl) // inE negb_and -!ltnNge /= ltn_predL lt0n -{3}Ev Nv0 orbT. - have ->: v = ord_max by apply/ord_inj. - rewrite (perm_all (s2:=&:(codom (pffun (pl * p) f)) `[l: nat, n[)) // pffunEM perm_sym. + - rewrite !ffunE permM (out_perm Hpl) // inE negb_and -!ltnNge /=. + by rewrite ltn_predL lt0n -{3}Ev Nv0 orbT. + have ->: v = ord_max by apply/ord_inj. + rewrite [in X in X -> _](perm_all (s2:=&:(codom (pffun (pl * p) f)) `[l: nat, n[)) //. + rewrite pffunEM perm_sym. rewrite {8 15}(_ : n = (ord_max : 'I_n.+1)) //; apply: perm_on_fgraph. apply/(subset_trans Hpl)/subsetP=>/= z; rewrite 2!inE in_itv /= leEnat ltEnat /=. by case/andP=>->/= Hz; apply: (leq_ltn_trans Hz); rewrite ltn_predL lt0n -Ev. @@ -559,11 +583,13 @@ rewrite (slice_xL (x:=v)) // onth_codom /=. have -> : pffun (pr * (pl * p)) f v = pffun p f v. - rewrite !ffunE mulgA; suff ->: (pr * pl * p)%g v = p v by []. rewrite permM. - have Hmul: perm_on [set ix : 'I_n.+1| (l <= ix <= v.-1) || (v < ix <= h)] (pr * pl)%g. + have Hmul : perm_on [set ix : 'I_n.+1| (l <= ix <= v.-1) || (v < ix <= h)] + (pr * pl)%g. - apply: perm_onM. - by apply/(subset_trans Hpr)/subsetP=>/= z; rewrite !inE=>->; rewrite orbT. by apply/(subset_trans Hpl)/subsetP=>/= z; rewrite !inE=>->. - by rewrite (out_perm Hmul) // inE negb_or !negb_and -leqNgt -!ltnNge leqnn /= andbT ltn_predL lt0n Nv0 orbT. + rewrite (out_perm Hmul) // inE negb_or !negb_and -leqNgt -!ltnNge leqnn /=. + by rewrite andbT ltn_predL lt0n Nv0 orbT. rewrite {1}pffunEM (perm_on_notin _ Hpr); last first. - rewrite disjoint_subset; apply/subsetP=>/= z. rewrite 3!inE in_itv /= negb_and leEnat ltEnat /= -leqNgt -ltnNge. @@ -577,11 +603,13 @@ rewrite -mulgA (pffunEM _ (pr * p)%g) (perm_on_notin _ Hpl) in Sr *; last first. - rewrite disjoint_subset; apply/subsetP=>/= z. rewrite 3!inE in_itv /= negb_and leEnat /= -leqNgt -ltnNge. case/andP=>_ Hz; apply/orP; left; apply: (leq_trans Hz). - by exact: leq_pred. -rewrite sorted_pairwise // pairwise_cat /= allrel_consr -andbA -!sorted_pairwise //. + exact: leq_pred. +rewrite sorted_pairwise // pairwise_cat /=. +rewrite allrel_consr -andbA -!sorted_pairwise //. apply/and5P; split=>//. - move: Al; congr (_ = _); apply/esym/perm_all; rewrite pffunEM. - apply/perm_on_fgraph/(subset_trans Hpl)/subsetP=>/= z; rewrite 2!inE in_itv /= leEnat ltEnat /=. + apply/perm_on_fgraph/(subset_trans Hpl)/subsetP=>/= z. + rewrite 2!inE in_itv /= leEnat ltEnat /=. by case/andP=>->/= Hz; apply: (leq_ltn_trans Hz); rewrite ltn_predL lt0n. - apply/allrelP=>x y Hx Hy; apply: (otrans (y:=pffun p f v)). - move/allP: Al=>/(_ x); apply. diff --git a/examples/stack.v b/examples/stack.v index 6d0892f..931e088 100644 --- a/examples/stack.v +++ b/examples/stack.v @@ -1,3 +1,16 @@ +(* +Copyright 2022 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 seq. From pcm Require Import options axioms pred. @@ -49,19 +62,19 @@ Program Definition free s : STsep (shape s [::], [vfun _ h => h = Unit]) := Do (dealloc s). Next Obligation. -by move=>s [] ? /= [?][?][V -> [_ ->]]; step=>_; rewrite unitR. +by case=>i [?][?][V ->][_ ->]; step=>_; rewrite unitR. Qed. (* pushing to the stack is inserting into the list and updating the pointer *) -Program Definition push s x : {xs}, STsep (shape s xs, +Program Definition push s x : STsep {xs} (shape s xs, [vfun _ => shape s (x :: xs)]) := Do (l <-- !s; l' <-- insert l x; s ::= l'). Next Obligation. (* pull out ghost + precondition, get the list *) -move=>s x [xs][] _ /= [l][h][V -> H]; step. +case=>xs [] _ /= [l][h][V -> H]; step. (* run the insert procedure with the ghost, deconstruct the new list *) apply: [stepX xs]@h=>//= x0 _ [r][h'][-> H']. (* store the new list *) @@ -73,10 +86,10 @@ Qed. (* 2. removing it from the list and updating the pointer on success *) Program Definition pop s : - {xs}, STsep (shape s xs, - fun y h => shape s (behead xs) h /\ - match y with Val v => xs = v :: behead xs - | Exn e => e = EmptyStack /\ xs = [::] end) := + STsep {xs} (shape s xs, + fun y h => shape s (behead xs) h /\ + match y with Val v => xs = v :: behead xs + | Exn e => e = EmptyStack /\ xs = [::] end) := Do (l <-- !s; try (head l) (fun x => @@ -86,7 +99,7 @@ Program Definition pop s : (fun _ => throw EmptyStack)). Next Obligation. (* pull out ghost vars and precondition *) -move=>s [xs][] _ /= [p][h0][V -> H]. +case=>xs [] _ /= [p][h0][V -> H]. (* get the list and invoke head on it, deal with exception first *) step; apply/[tryX xs]@h0=>//= [x|ex] m [Hm]; last first. - (* throw the stack exception *) diff --git a/htt/heapauto.v b/htt/heapauto.v index 8503353..8ae51d9 100644 --- a/htt/heapauto.v +++ b/htt/heapauto.v @@ -13,7 +13,7 @@ limitations under the License. From mathcomp Require Import ssreflect ssrbool ssrnat eqtype seq ssrfun. From pcm Require Import options axioms pred prelude. -From pcm Require Import pcm autopcm unionmap heap. +From pcm Require Import auto pcm autopcm unionmap heap. From htt Require Import model. Import Prenex Implicits. @@ -450,8 +450,10 @@ Notation "[ 'gX' x1 , .. , xn ] @ m" := (gX (existT _ x1 .. (existT _ xn tt) ..) m erefl) (at level 0, format "[ 'gX' x1 , .. , xn ] @ m"). +Definition heapPCM : pcm := heap. + (* combination of gX + vrf_bind *) -Lemma stepX G A B (s : spec G A) g (m : heap) m0 j tm k wh r2 +Lemma stepX G A B (s : spec G A) g (m : heapPCM) (m0 : heap) (j : ctx heapPCM) tm k wh r2 (e : STspec G s) (e2 : A -> ST B) (fm : Syntactify.form (empx _) j tm) (fu : uform m0 (Syntactify.untag fm)) @@ -463,19 +465,19 @@ Lemma stepX G A B (s : spec G A) g (m : heap) m0 j tm k wh r2 (forall v n, (s g).2 (Val v) n -> vrf (f n) (e2 v) Q) -> (forall x n, (s g).2 (Exn x) n -> valid (untag (f n)) -> Q (Exn x) (f n)) -> - vrf (PullX.unpack fg) (bnd e e2) Q. + vrf (fg : heapPCM) (bnd e e2) Q. Proof. move=>Hm Hp Hv Hx; apply/vrf_bnd/(gX _ _ Hm Hp). - by move=>v n H V; apply: Hv. by move=>x n H V _; apply: Hx. Qed. -Arguments stepX {G A B s} g m {m0 j tm k wh r2 e e2 fm fu f fg Q}. +Arguments stepX [G A B s] g m {m0 j tm k wh r2 e e2 fm fu f fg Q}. Notation "[stepX] @ m" := (stepX tt m erefl) (at level 0). Notation "[ 'stepX' x1 , .. , xn ] @ m" := (stepX (existT _ x1 .. (existT _ xn tt) ..) m erefl) - (at level 0, format "[ 'stepX' x1 , .. , xn ] @ m"). + (at level 0, format "[ 'stepX' x1 , .. , xn ] @ m"). (* combination of gX + vrf_try *) Lemma tryX G A B (s : spec G A) g (m : heap) m0 j tm k wh r2 diff --git a/htt/options.v b/htt/options.v index c519867..94dba88 100644 --- a/htt/options.v +++ b/htt/options.v @@ -1,6 +1,12 @@ (* turn off the automation of Program *) #[export] Obligation Tactic := auto. (* turn off other Program extensions *) -#[export] Unset Program Cases. -#[export] Unset Program Generalized Coercion. +#[export] Unset Program Cases. +(* commenting this out allows a bit more convenience *) +(* when working with tuples, as one can pass a proof *) +(* of a "wrong" fact, and Program would emit obligation *) +(* that wrong fact equals the right fact. *) +(* If left uncommented, the equality will have to be *) +(* witnessed manually in the program *) +(* #[export] Unset Program Generalized Coercion. *) #[export] Unset Program Mode. diff --git a/pcm/autopcm.v b/pcm/autopcm.v index 8372161..e3b5f0f 100644 --- a/pcm/autopcm.v +++ b/pcm/autopcm.v @@ -13,7 +13,8 @@ limitations under the License. From Coq Require Import ssreflect ssrbool ssrfun. From mathcomp Require Import ssrnat eqtype seq. -From pcm Require Import options pred prelude auto pcm. +From pcm Require Import options pred prelude pcm. +From pcm Require Export auto. (**************************************************************************) (**************************************************************************) @@ -355,3 +356,8 @@ End PullX. Export PullX.Exports. + + + + + From 5270a214c38a24f3724916c6635688c0408c89b3 Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Mon, 5 Aug 2024 19:51:23 +0200 Subject: [PATCH 10/93] refactored a bunch of examples. going on with graph.v --- examples/bintree.v | 72 ++++++++++------- examples/bst.v | 55 ++++++++----- examples/bubblesort.v | 2 +- examples/congmath.v | 54 ++++++++----- examples/counter.v | 29 +++++-- examples/cyclic.v | 100 +++++++++++++---------- examples/dlist.v | 88 ++++++++++++-------- examples/hashtab.v | 43 ++++++---- examples/kvmaps.v | 182 ++++++++++++++++++++++-------------------- examples/queue.v | 59 +++++++++----- examples/quicksort.v | 3 +- examples/tree.v | 21 ++++- examples/union_find.v | 108 ++++++++++++++----------- pcm/heap.v | 13 ++- pcm/unionmap.v | 53 +++++++++--- 15 files changed, 540 insertions(+), 342 deletions(-) diff --git a/examples/bintree.v b/examples/bintree.v index 0b41d66..0cbe427 100644 --- a/examples/bintree.v +++ b/examples/bintree.v @@ -1,3 +1,15 @@ +(* +Copyright 2021 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 seq ssrnat. @@ -34,14 +46,15 @@ Fixpoint addr {A} (y : A) (t : tree A) : tree A := Fixpoint shape {A} (p : ptr) (t : tree A) := if t is Node l a r then [Pred h | exists l' r' h', - h = p :-> l' \+ (p .+ 1 :-> a \+ (p .+ 2 :-> r' \+ h')) + h = p :-> l' \+ (p.+1 :-> a \+ (p.+2 :-> r' \+ h')) /\ h' \In shape l' l # shape r' r] else [Pred h | p = null /\ h = Unit]. (* Null pointer represents a leaf *) Lemma shape_null {A} (t : tree A) h : - valid h -> h \In shape null t -> + valid h -> + h \In shape null t -> t = leaf /\ h = Unit. Proof. move=>V; case: t=>/= [|l a r]; first by case=>_->. @@ -51,10 +64,11 @@ Qed. (* Non-null pointer represents a node *) Lemma shape_cont {A} (t : tree A) p h : - p != null -> h \In shape p t -> + p != null -> + h \In shape p t -> exists l a r l' r' h', [/\ t = Node l a r, - h = p :-> l' \+ (p .+ 1 :-> a \+ (p .+ 2 :-> r' \+ h')) + h = p :-> l' \+ (p.+1 :-> a \+ (p.+2 :-> r' \+ h')) & h' \In shape l' l # shape r' r]. Proof. move=>E; case: t=>/= [|l a r]. @@ -72,13 +86,13 @@ Program Definition mknode (x : A) : STsep (emp, [vfun p => shape p (Node leaf x leaf)]) := Do (n <-- allocb null 3; - n.+ 1 ::= x;; + n.+1 ::= x;; ret n). Next Obligation. (* the starting heap is empty *) move=>x [] _ /= ->. (* run all the steps *) -step=>n; rewrite !unitR ptrA; step; step=>_. +step=>n; rewrite !unitR; step; step=>_. (* the postcondition is satisfied *) by exists null, null, Unit; vauto; rewrite unitR. Qed. @@ -87,18 +101,18 @@ Qed. (* We start from a well-formed tree and arrive at an empty heap *) Definition disposetreeT : Type := - forall p, {t : tree A}, STsep (shape p t, [vfun _ : unit => emp]). + forall p, STsep {t : tree A} (shape p t, [vfun _ : unit => emp]). Program Definition disposetree : disposetreeT := - Fix (fun (loop : disposetreeT) p => + ffix (fun (loop : disposetreeT) p => Do (if p == null then ret tt else l <-- !p; - r <-- !(p.+ 2); + r <-- !p.+2; loop l;; loop r;; dealloc p;; - dealloc p.+ 1;; - dealloc p.+ 2 + dealloc p.+1;; + dealloc p.+2 )). Next Obligation. (* pull out ghost var + precondition, branch on null check *) @@ -120,16 +134,16 @@ Qed. (* loop invariant: *) (* the subtree size is added to the accumulator *) Definition treesizeT : Type := forall (ps : ptr * nat), - {t : tree A}, STsep (shape ps.1 t, + STsep {t : tree A} (shape ps.1 t, [vfun s h => s == ps.2 + size_tree t /\ shape ps.1 t h]). Program Definition treesize p : - {t : tree A}, STsep (shape p t, + STsep {t : tree A} (shape p t, [vfun s h => s == size_tree t /\ shape p t h]) := - Do (let len := Fix (fun (go : treesizeT) '(p, s) => + Do (let len := ffix (fun (go : treesizeT) '(p, s) => Do (if p == null then ret s else l <-- !p; - r <-- !(p.+ 2); + r <-- !p.+2; ls <-- go (l, s); go (r, ls.+1))) in len (p, 0)). @@ -160,18 +174,18 @@ Opaque insert new. (* loop invariant: *) (* the subtree is unchanged, its values are prepended to the accumulator list *) Definition inordertravT : Type := forall (ps : ptr * ptr), - {(t : tree A) (l : seq A)}, - STsep (shape ps.1 t # lseq ps.2 l, - [vfun s h => h \In shape ps.1 t # lseq s (inorder t ++ l)]). + STsep {(t : tree A) (l : seq A)} + (shape ps.1 t # lseq ps.2 l, + [vfun s h => h \In shape ps.1 t # lseq s (inorder t ++ l)]). Program Definition inordertrav p : - {t : tree A}, STsep (shape p t, + STsep {t : tree A} (shape p t, [vfun s h => h \In shape p t # lseq s (inorder t)]) := - Do (let loop := Fix (fun (go : inordertravT) '(p, s) => + Do (let loop := ffix (fun (go : inordertravT) '(p, s) => Do (if p == null then ret s else l <-- !p; - a <-- !(p.+ 1); - r <-- !(p.+ 2); + a <-- !p.+1; + r <-- !p.+2; s1 <-- go (r, s); s2 <-- insert s1 (a : A); go (l, s2))) @@ -193,13 +207,13 @@ move=>s1 _ [hr'][hs][-> Hr' Hs]. (* prepend the node value to the list *) apply: [stepX (inorder r ++ xs)]@hs=>//= pa _ [s2][h'][-> H']. (* run the traversal on the left branch with updated list *) -apply: [gX l, (a::inorder r ++ xs)]@(hl \+ pa :-> a \+ pa.+ 1 :-> s2 \+ h')=>//=. +apply: [gX l, (a::inorder r ++ xs)]@(hl \+ pa :-> a \+ pa.+1 :-> s2 \+ h')=>//=. - (* the precondition is satisfied by simple heap manipulation *) - move=>_; exists hl, (pa :-> a \+ (pa.+ 1 :-> s2 \+ h')). + move=>_; exists hl, (pa :-> a \+ (pa.+1 :-> s2 \+ h')). by split=>//=; [rewrite !joinA | vauto]. (* the postcondition is also satisfied by simple massaging *) move=>s3 _ [hl''][hs'][-> Hl'' Hs'] _. -exists (p :-> l' \+ (p.+ 1 :-> a \+ (p.+ 2 :-> r' \+ (hl'' \+ hr')))), hs'. +exists (p :-> l' \+ (p.+1 :-> a \+ (p.+2 :-> r' \+ (hl'' \+ hr')))), hs'. split; try by hhauto. by rewrite -catA. Qed. @@ -221,17 +235,17 @@ Opaque mknode. (* loop invariant: the value is added to the rightmost branch *) Definition expandrightT x : Type := forall (p : ptr), - {t : tree A}, STsep (shape p t, + STsep {t : tree A} (shape p t, [vfun p' => shape p' (addr x t)]). Program Definition expandright x : expandrightT x := - Fix (fun (go : expandrightT x) p => + ffix (fun (go : expandrightT x) p => Do (if p == null then n <-- mknode x; ret n - else pr <-- !(p.+ 2); + else pr <-- !p.+2; p' <-- go pr; - p.+ 2 ::= p';; + p.+2 ::= p';; ret p)). Next Obligation. (* pull out ghost + precondition, branch on null check *) diff --git a/examples/bst.v b/examples/bst.v index 37ff69a..eae0dfe 100644 --- a/examples/bst.v +++ b/examples/bst.v @@ -1,8 +1,21 @@ +(* +Copyright 2021 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 ssrnat seq path. -From pcm Require Import options axioms pred ordtype. +From pcm Require Import options axioms pred ordtype seqext. From pcm Require Import pcm unionmap heap autopcm automap. -From htt Require Import model heapauto bintree. +From htt Require Import options model heapauto bintree. Section BST. Context {T : ordType}. @@ -47,7 +60,7 @@ Lemma bst_to_sorted (t : tree T) : bst t = sorted ord (inorder t). Proof. elim: t=>//=l -> a r ->. -by rewrite sorted_cat_cons_cat /= cats1 sorted_rconsE // +by rewrite sorted_cat_cons_cat /= cats1 sorted_rconsE //= (path_sortedE (@trans T)) andbACA -andbA. Qed. @@ -123,9 +136,9 @@ by apply: (perm_trans Hi2); rewrite Hx2. Qed. (* Moreover, such representations are equal *) -Corollary insert_insert x1 x2 (t : tree T) : - bst t -> - inorder (insert x1 (insert x2 t)) = inorder (insert x2 (insert x1 t)). +Lemma insert_insert x1 x2 (t : tree T) : + bst t -> + inorder (insert x1 (insert x2 t)) = inorder (insert x2 (insert x1 t)). Proof. move=>H; apply: ord_sorted_eq. - by rewrite -bst_to_sorted; do 2!apply: bst_insert. @@ -157,24 +170,24 @@ Opaque mknode. Definition inserttreeT x : Type := forall p, - {t : tree T}, STsep (shape p t, - [vfun p' => shape p' (insert x t)]). + STsep {t : tree T} (shape p t, + [vfun p' => shape p' (insert x t)]). Program Definition inserttree x : inserttreeT x := - Fix (fun (go : inserttreeT x) p => + ffix (fun (go : inserttreeT x) p => Do (if p == null then n <-- mknode x; ret n - else a <-- !(p.+ 1); + else a <-- !p.+1; if x == a then ret p else if ord x a then l <-- !p; l' <-- go l; p ::= l';; ret p - else r <-- !(p.+ 2); + else r <-- !p.+2; r' <-- go r; - p.+ 2 ::= r';; + p.+2 ::= r';; ret p)). Next Obligation. (* pull out ghost + precondition, branch on null check *) @@ -202,19 +215,19 @@ Qed. (* lopp invariant: the tree is unchanged, return true if the element is found *) Definition searchtreeT x : Type := forall p, - {t : tree T}, STsep (shape p t, - [vfun b h => shape p t h /\ b == search t x]). + STsep {t : tree T} (shape p t, + [vfun b h => shape p t h /\ b == search t x]). Program Definition searchtree x : searchtreeT x := - Fix (fun (go : searchtreeT x) p => + ffix (fun (go : searchtreeT x) p => Do (if p == null then ret false - else a <-- !(p.+ 1); + else a <-- !p.+1; if x == a then ret true else if ord x a then l <-- !p; go l - else r <-- !(p.+ 2); + else r <-- !p.+2; go r)). Next Obligation. (* pull out ghost + precondition, branch on null check *) @@ -236,10 +249,10 @@ Qed. (* test that the API is consistent, i.e. BST invariant is preserved *) (* and lookup finds previously inserted elements *) Program Definition test p x1 x2 : - {t : tree T}, STsep (fun h => shape p t h /\ bst t, - [vfun (pb : ptr * bool) h => - let t' := insert x2 (insert x1 t) in - [/\ shape pb.1 t' h, bst t' & pb.2]]) := + STsep {t : tree T} (fun h => shape p t h /\ bst t, + [vfun (pb : ptr * bool) h => + let t' := insert x2 (insert x1 t) in + [/\ shape pb.1 t' h, bst t' & pb.2]]) := Do (p1 <-- inserttree x1 p; p2 <-- inserttree x2 p1; b <-- searchtree x1 p2; diff --git a/examples/bubblesort.v b/examples/bubblesort.v index 79e73cd..501aa51 100644 --- a/examples/bubblesort.v +++ b/examples/bubblesort.v @@ -775,7 +775,7 @@ step=>Vm; exists p; case: sw Hsw=>/=; case=>Ep Hf. move: Hak; rewrite codom1_ax_rcons2 // !allrel_rconsl -/i0 -/i1 -andbA. case/and3P=>-> _ ->; rewrite (ordW Hf) /= !andbT. move: Hai; rewrite codom1_ax_rcons //= allrel_rconsr=>/andP [_ Hals]. - move: Hsi; rewrite codom1_ax_rcons // sorted_rconsE; last by apply: otrans. + move: Hsi; rewrite codom1_ax_rcons // sorted_rconsE //=. case/andP => Halsi _. move: Hils; rewrite leq_eqVlt; case/orP=>[/eqP <-|Hlsi] //. rewrite (slice_split (x:=ls) _ true); last by rewrite in_itv. diff --git a/examples/congmath.v b/examples/congmath.v index 64f72b9..46cdc88 100644 --- a/examples/congmath.v +++ b/examples/congmath.v @@ -1,7 +1,21 @@ +(* +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. +*) + (**********************) (* Congruence closure *) (**********************) +From HB Require Import structures. From Coq Require Import Recdef ssreflect ssrbool ssrfun. From mathcomp Require Import eqtype choice ssrnat seq bigop fintype finfun. From pcm Require Import options prelude ordtype finmap pred seqext. @@ -17,8 +31,8 @@ Inductive constant : Type := const_with_arity of nat & nat. (* constants are an equality type *) Definition eqcnst (c1 c2 : constant) : bool := match c1, c2 with - | const_with_arity s1 n1, const_with_arity s2 n2 => - (s1 == s2) && (n1 == n2) + | const_with_arity s1 n1, const_with_arity s2 n2 => + (s1 == s2) && (n1 == n2) end. Lemma eqconstantP : Equality.axiom eqcnst. @@ -28,8 +42,7 @@ do ![case: eqP=>/=[->|H]; last by apply: ReflectF; case]. by apply: ReflectT. Qed. -Canonical Structure constant_eqMixin := EqMixin eqconstantP. -Canonical Structure constant_eqType := EqType _ constant_eqMixin. +HB.instance Definition _ := hasDecEq.Build constant eqconstantP. (* constants are a countable type *) @@ -42,25 +55,27 @@ Definition consttag (u : {i : nat & nat}) := Lemma tagconstK : cancel tagconst consttag. Proof. by case. Qed. +(* Lemma consttagK : cancel consttag tagconst. Proof. by case. Qed. +*) -Definition const_countMixin := CanCountMixin tagconstK. -Definition const_choiceMixin := CountChoiceMixin const_countMixin. -Canonical Structure const_choiceType := Eval hnf in ChoiceType _ const_choiceMixin. -Canonical Structure const_countType := Eval hnf in CountType _ const_countMixin. +HB.instance Definition _ := Countable.copy constant (can_type tagconstK). (***********************************************************) (* The symbols are a predetermined finite set of constants *) (***********************************************************) - Variable symbs : seq constant. -Definition symb := @seq_sub const_choiceType symbs. -Canonical Structure symb_eqType := [eqType of symb]. -Canonical Structure symb_finType := [finType of symb for @seq_sub_finType _ _]. -(* symb is an ordered type *) -Definition symb_ordMixin := [fin_ordMixin of symb]. -Canonical Structure symb_ordType := OrdType _ symb_ordMixin. +Definition symb := seq_sub symbs. + +HB.instance Definition _ := Equality.on symb. +HB.instance Definition _ := Finite.on symb. + +HB.instance Definition symb_ord_mix := + isOrdered.Build symb (fintype_is_ordtype symb). +(* manual canonical declaration, as HB fails to declare it *) +Canonical symb_ordType : ordType := + Ordered.Pack (sort:=symb) (Ordered.Class symb_ord_mix). (****************************) (* Expressions over symbols *) @@ -86,8 +101,7 @@ case: IH2=>[->|H]; last by apply: ReflectF; case. by apply: ReflectT. Qed. -Canonical Structure exp_eqMixin := EqMixin eqexpP. -Canonical Structure exp_eqType := EqType _ exp_eqMixin. +HB.instance Definition _ := hasDecEq.Build exp eqexpP. (****************************************) (* Congruence relations over expression *) @@ -258,8 +272,7 @@ do ![case: eqP=>[->|H]; last by apply: ReflectF; case]; by apply: ReflectT. Qed. -Canonical Structure Eq_eqMixin := EqMixin eqEqP. -Canonical Structure Eq_eqType := EqType _ Eq_eqMixin. +HB.instance Definition _ := hasDecEq.Build Eq eqEqP. (* interpreting equations as relations *) Definition eq2rel (eq : Eq) : rel_exp := @@ -320,8 +333,7 @@ do ![case: eqP=>[->|H]; last by apply: ReflectF; case]; by apply: ReflectT. Qed. -Canonical Structure pend_eqMixin := EqMixin eq_pendP. -Canonical Structure pend_eqType := EqType _ pend_eqMixin. +HB.instance Definition _ := hasDecEq.Build pend eq_pendP. (* pending equations are interpreted as follows for *) (* purposes of computing the relations they encode *) diff --git a/examples/counter.v b/examples/counter.v index cc57ebf..f85522c 100644 --- a/examples/counter.v +++ b/examples/counter.v @@ -1,3 +1,16 @@ +(* +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 mathcomp Require Import ssreflect ssrbool ssrfun. From mathcomp Require Import ssrnat eqtype. From pcm Require Import options axioms prelude pred. @@ -10,8 +23,8 @@ Prenex Implicits exist. Record cnt : Type := Counter {inv : nat -> Pred heap; - method : {v : nat}, STsep (inv v, - [vfun y m => y = v /\ m \In inv v.+1])}. + method : STsep {v : nat} (inv v, + [vfun y m => y = v /\ m \In inv v.+1])}. Program Definition new : STsep (emp, [vfun y m => m \In inv y 0]) := Do (x <-- alloc 0; @@ -28,7 +41,7 @@ move=>[] _ /= ->. (* but it doesn't; we suspect because universe levels are off *) (* but error messages aren't helpful *) (* one can still finish the step manually *) -apply/vrf_bind/vrf_alloc=>x; rewrite unitR=>_. +apply/vrf_bnd/vrf_alloc=>x; rewrite unitR=>_. (* after which, automation proceeds to work *) by step. Qed. @@ -53,8 +66,8 @@ Definition interp (r : rep) : Pred heap := [Pred h | h = rptr r :-> rval r]. Record scnt : Type := SCount {sinv : nat -> rep; - smethod : {v : nat}, STsep (interp (sinv v), - [vfun y m => y = v /\ m \In interp (sinv v.+1)])}. + smethod : STsep {v : nat} (interp (sinv v), + [vfun y m => y = v /\ m \In interp (sinv v.+1)])}. Program Definition snew : STsep (emp, [vfun y m => m \In interp (sinv y 0)]) := @@ -69,7 +82,7 @@ by do 3!step. Qed. Next Obligation. move=>[] _ /= ->. -apply/vrf_bind/vrf_alloc=>x; rewrite unitR=>_. +apply/vrf_bnd/vrf_alloc=>x; rewrite unitR=>_. by step. Qed. @@ -93,7 +106,7 @@ Definition pinterp (R : rep -> Prop) : Pred heap := Record pcnt : Type := PCount {pinv : nat -> rep -> Prop; - pmethod : {v : nat}, STsep (pinterp (pinv v), + pmethod : STsep {v : nat} (pinterp (pinv v), [vfun y m => y = v /\ m \In pinterp (pinv v.+1)])}. Program Definition pnew : STsep (emp, @@ -109,7 +122,7 @@ by do 3!step; move=>?; split=>// _ ->. Qed. Next Obligation. move=>[] _ /= ->. -apply/vrf_bind/vrf_alloc=>x; rewrite unitR=>_. +apply/vrf_bnd/vrf_alloc=>x; rewrite unitR=>_. by step=>_ /= _ ->. Qed. diff --git a/examples/cyclic.v b/examples/cyclic.v index 02705dc..76c22bf 100644 --- a/examples/cyclic.v +++ b/examples/cyclic.v @@ -1,7 +1,20 @@ +(* +Copyright 2021 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 seq ssrnat. -From pcm Require Import options axioms pred. -From pcm Require Import pcm unionmap heap automap autopcm. +From pcm Require Import options axioms pred seqext. +From pcm Require Import pcm unionmap heap auto automap autopcm. From htt Require Import options model heapauto. From htt Require Import llist. @@ -46,14 +59,14 @@ Definition shape (b : buffer) (xs : seq T) := (* loop invariant for allocating the inner structure *) Definition new_loopT (n : nat) (init : T) : Type := forall (pk : ptr * nat), - {q}, STsep (fun h => pk.2 < n /\ h \In lseg pk.1 q (nseq pk.2 init), + STsep {q} (fun h => pk.2 < n /\ h \In lseg pk.1 q (nseq pk.2 init), [vfun p' => lseg p' q (nseq n.-1 init)]). (* allocate the buffer with capacity n prefilled by init *) Program Definition new (n : nat) (init : T) : STsep (emp, [vfun v => shape v [::]]) := (* allocate the buffer in a loop *) - Do (let run := Fix (fun (go : new_loopT n init) '(r,k) => + Do (let run := ffix (fun (go : new_loopT n init) '(r,k) => Do (if k < n.-1 then p' <-- allocb r 2; p' ::= init;; @@ -67,7 +80,7 @@ Program Definition new (n : nat) (init : T) : (* allocate the remaining n-1 nodes by prepending to it *) q <-- run (p, 0); (* form the cycle *) - p.+ 1 ::= q;; + p.+1 ::= q;; (* allocate the structure *) m <-- alloc 0; pi <-- alloc q; @@ -93,23 +106,23 @@ Qed. (* now the outer program *) Next Obligation. (* simplify, match on n *) -move=>n init [] _ /= ->; case: ifP. +move=>n init [] _ /= ->; case: ifP=>[N0|_]. - (* 0 < n, allocate the initial node *) - move=>H0; step=>p; step; rewrite !unitR. + step=>p; step; rewrite !unitR. (* run the loop by framing on an empty heap *) apply: [stepU p]=>//= q h H. (* run the rest of the program *) step; step=>m; step=>pi; step=>pa; step=>V. (* the structure is well-formed if the buffer is *) - exists q, q, 0, (h \+ (p :-> init \+ p.+ 1 :-> q)); split=>//. + exists q, q, 0, (h \+ (p :-> init \+ p.+1 :-> q)); split=>//. (* most of the conditions hold trivially or by simple arithmetics *) - exists (nseq n init), Unit, (h \+ (p :-> init \+ p.+ 1 :-> q)); split=>//=. + exists (nseq n init), Unit, (h \+ (p :-> init \+ p.+1 :-> q)); split=>//=. - by rewrite unitL. - by rewrite add0n size_nseq. (* the cycle is well-formed *) - by rewrite -(ltn_predK H0) -rcons_nseq; apply/lseg_rcons; exists p, h. + by rewrite -(ltn_predK N0) -rcons_nseq; apply/lseg_rcons; exists p, h. (* n = 0, allocate a trivial structure with a null buffer *) -move=>_; step=>m; step=>pi; step=>pa; step=>V. +step=>m; step=>pi; step=>pa; step=>V. (* it is well-formed *) exists null, null, 0, Unit=>/=; split=>//. by exists [::], Unit, Unit; split=>//; rewrite unitR. @@ -117,16 +130,16 @@ Qed. (* "polite" write, checks the buffer length, throws an exception on overflow *) Program Definition write (x : T) (b : buffer) : - {xs}, STsep (shape b xs, - fun y h => match y with - | Val _ => h \In shape b (rcons xs x) - | Exn e => e = BufferFull /\ h \In shape b xs - end) := + STsep {xs} (shape b xs, + fun y h => match y with + | Val _ => h \In shape b (rcons xs x) + | Exn e => e = BufferFull /\ h \In shape b xs + end) := Do (m <-- !len b; if m < capacity b then i <-- !inactive b; i ::= x;; - r <-- !i.+ 1; + r <-- !i.+1; inactive b ::= (r : ptr);; len b ::= m.+1 else throw BufferFull). @@ -138,11 +151,11 @@ rewrite Ec; step; case: ltnP. - (* buffer is not full, so the inactive part is non-empty *) rewrite {h}E -{1}(addn0 (size xs)) ltn_add2l => Hys. (* read the inactive pointer, unroll the inactive heap so that we can proceed *) - step; case/(lseg_lt0n Hys): Hi=>y[r][h][_ {hi}-> H]; do 4![step]=>{y}V. + step; case/(lseg_lt0n Hys): Hi=>y [r][h][_ {hi}-> H]; do 4![step]=>{y}V. (* the new structure is well-formed if the buffer is *) - exists a, r, (size xs).+1, (ha \+ (i :-> x \+ (i.+ 1 :-> r \+ h))); split=>//. + exists a, r, (size xs).+1, (ha \+ (i :-> x \+ (i.+1 :-> r \+ h))); split=>//. (* most of the conditions hold trivially or by simple arithmetics *) - exists (behead ys), (ha \+ (i :-> x \+ i.+ 1 :-> r)), h; split=>//. + exists (behead ys), (ha \+ (i :-> x \+ i.+1 :-> r)), h; split=>//. - by rewrite !joinA. - by rewrite Ec size_rcons size_behead -subn1 addnABC // subn1. - by rewrite size_rcons. @@ -160,11 +173,12 @@ Qed. (* a version that overwrites data in a cyclic fashion *) (* we make checking for capacity != 0 the client's problem so it can be dealt with globally *) Program Definition overwrite (x : T) (b : buffer) : - {xs}, STsep (fun h => 0 < capacity b /\ h \In shape b xs, - [vfun _ => shape b (drop ((size xs).+1 - capacity b) (rcons xs x))]) := + STsep {xs} (fun h => 0 < capacity b /\ h \In shape b xs, + [vfun _ => shape b (drop ((size xs).+1 - capacity b) + (rcons xs x))]) := Do (i <-- !inactive b; i ::= x;; - r <-- !i.+ 1; + r <-- !i.+1; inactive b ::= (r : ptr);; m <-- !len b; if m < capacity b then @@ -184,9 +198,9 @@ rewrite Ec in Hc *; step; case: ys Ec Hi Hc=>/=. (* run the rest of the program, the postcondition simplifies to beheading the extended xs *) do 4!step; rewrite ltnn subSnn drop1; step=>V. (* the new structure is well-formed if the buffer is *) - exists r, r, (size xs), (a :-> x \+ (a.+ 1 :-> r \+ h')); split=>//. + exists r, r, (size xs), (a :-> x \+ (a.+1 :-> r \+ h')); split=>//. (* most of the conditions hold trivially or by simple arithmetics *) - exists [::], (a :-> x \+ (a.+ 1 :-> r \+ h')), Unit; split=>//=. + exists [::], (a :-> x \+ (a.+1 :-> r \+ h')), Unit; split=>//=. - by rewrite unitR. - by rewrite addn0 size_behead size_rcons. - by rewrite size_behead size_rcons. @@ -200,9 +214,9 @@ move=>y ys Ec [r][h'][{hi}-> H'] _. do 4![step]=>{y}; rewrite -{1}(addn0 (size xs)) ltn_add2l /=; step=>V. rewrite addnS subSS subnDA subnn sub0n drop0. (* the new structure is well-formed if the buffer is *) -exists a, r, (size xs).+1, (ha \+ (i :-> x \+ (i.+ 1 :-> r \+ h'))); split=>//. +exists a, r, (size xs).+1, (ha \+ (i :-> x \+ (i.+1 :-> r \+ h'))); split=>//. (* the buffer is well-formed by simple arithmetics *) -exists ys, (ha \+ (i :-> x \+ i.+ 1 :-> r)), h'; split=>//. +exists ys, (ha \+ (i :-> x \+ i.+1 :-> r)), h'; split=>//. - by rewrite !joinA. - by rewrite Ec size_rcons addnS addSn. - by rewrite size_rcons. @@ -211,16 +225,16 @@ Qed. (* reading (or rather popping) a value from the buffer *) Program Definition read (b : buffer) : - {xs}, STsep (shape b xs, - fun y h => match y with - | Val x => h \In shape b (behead xs) /\ ohead xs = Some x - | Exn e => e = BufferEmpty /\ xs = [::] - end) := + STsep {xs} (shape b xs, + fun y h => match y with + | Val x => h \In shape b (behead xs) /\ ohead xs = Some x + | Exn e => e = BufferEmpty /\ xs = [::] + end) := Do (m <-- !len b; if 0 < m then a <-- !active b; x <-- !a; - r <-- !a.+ 1; + r <-- !a.+1; active b ::= (r : ptr);; len b ::= m.-1;; ret x @@ -237,9 +251,9 @@ step; case: ltnP. (* run the rest of the program *) do 5![step]=>V; split; last by rewrite Exs. (* the new structure is well-formed if the buffer is *) - exists r, i, (size xs).-1, (a :-> x \+ (a.+ 1 :-> r \+ h') \+ hi); split=>//. + exists r, i, (size xs).-1, (a :-> x \+ (a.+1 :-> r \+ h') \+ hi); split=>//. (* the buffer is well-formed by simple arithmetics *) - exists (rcons ys x), h', (hi \+ (a :-> x \+ a.+ 1 :-> r)); split=>//. + exists (rcons ys x), h', (hi \+ (a :-> x \+ a.+1 :-> r)); split=>//. - by rewrite [LHS](pullX (h' \+ hi)) !joinA. - by rewrite size_rcons Hc {1}Exs /= addnS addSn. - by rewrite size_behead. @@ -255,17 +269,17 @@ Qed. (* loop invariant for deallocating the inner structure *) Definition free_loopT (n : nat) : Type := forall (pk : ptr * nat), - {q (xs : seq T)}, STsep (fun h => h \In lseg pk.1 q xs /\ size xs = n - pk.2, - [vfun _ : unit => emp]). + STsep {q (xs : seq T)} (fun h => h \In lseg pk.1 q xs /\ size xs = n - pk.2, + [vfun _ : unit => emp]). Program Definition free (b : buffer) : - {xs}, STsep (fun h => h \In shape b xs, - [vfun _ => emp]) := - Do (let run := Fix (fun (go : free_loopT (capacity b)) '(r,k) => + STsep {xs} (fun h => h \In shape b xs, + [vfun _ => emp]) := + Do (let run := ffix (fun (go : free_loopT (capacity b)) '(r,k) => Do (if k < capacity b then - p' <-- !r.+ 1; + p' <-- !r.+1; dealloc r;; - dealloc r.+ 1;; + dealloc r.+1;; go (p', k.+1) else ret tt)) in (if 0 < capacity b then @@ -300,7 +314,7 @@ Next Obligation. move=>b [xs][] _ /= [a][i][m][_][_ -> [ys][ha][hi][-> Hc -> Ha Hi]]; case: ltnP. - (* 0 < capacity, first apply the vrf_bind primitive to "step into" the `if` block *) (* run the loop starting from the active part, frame on the corresponding heaps *) - move=>_; apply/vrf_bind; step; apply: [stepX a, xs++ys]@(ha \+ hi)=>//=. + move=>_; apply/vrf_bnd; step; apply: [stepX a, xs++ys]@(ha \+ hi)=>//=. - (* concatenate the active and inactive parts to satisfy the loop precondition *) move=>_; split; last by rewrite size_cat subn0. by apply/lseg_cat; exists i, ha, hi. diff --git a/examples/dlist.v b/examples/dlist.v index b54cefd..4657590 100644 --- a/examples/dlist.v +++ b/examples/dlist.v @@ -1,3 +1,15 @@ +(* +Copyright 2021 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. @@ -13,7 +25,7 @@ From htt Require Import options model heapauto. Fixpoint dseg {A} (pp p q qn : ptr) (xs : seq A) := if xs is hd::tl then [Pred h | exists r h', - h = p :-> hd \+ (p.+ 1 :-> r \+ (p.+ 2:-> pp \+ h')) + h = p :-> hd \+ (p.+1 :-> r \+ (p.+2:-> pp \+ h')) /\ h' \In dseg p r q qn tl] else [Pred h | [/\ p = qn, pp = q & h = Unit]]. @@ -27,7 +39,7 @@ Variable A : Type. Lemma dseg_rcons (xs : seq A) x pp p q qn h : h \In dseg pp p q qn (rcons xs x) <-> exists r h', - h = h' \+ (q :-> x \+ (q.+ 1 :-> qn \+ q.+ 2 :-> r)) + h = h' \+ (q :-> x \+ (q.+1 :-> qn \+ q.+2 :-> r)) /\ h' \In dseg pp p r q xs. Proof. elim: xs pp p h => [|y xs IH] pp p h/=. @@ -36,18 +48,19 @@ elim: xs pp p h => [|y xs IH] pp p h/=. by exists qn, Unit; rewrite unitR. split. - case=>r[_][-> /IH][s][h'][-> H']. - exists s, (p :-> y \+ (p.+ 1 :-> r \+ (p.+ 2 :-> pp \+ h'))). + exists s, (p :-> y \+ (p.+1 :-> r \+ (p.+2 :-> pp \+ h'))). rewrite !joinA; split=>//. by exists r, h'; rewrite !joinA. case=>r[_][->][s][h'][-> H']. -exists s, (h' \+ (q :-> x \+ (q.+ 1 :-> qn \+ q.+ 2 :-> r))). +exists s, (h' \+ (q :-> x \+ (q.+1 :-> qn \+ q.+2 :-> r))). rewrite !joinA; split=>//; apply/IH. by exists r, h'; rewrite !joinA. Qed. (* first node being null means the list is empty *) Lemma dseg_nullL (xs : seq A) pp q qn h : - valid h -> h \In dseg pp null q qn xs -> + valid h -> + h \In dseg pp null q qn xs -> [/\ qn = null, pp = q, xs = [::] & h = Unit]. Proof. case: xs=>[|x xs] /= D H; first by case: H. @@ -56,7 +69,8 @@ Qed. (* last node being null also means the list is empty *) Lemma dseg_nullR (xs : seq A) pp p qn h : - valid h -> h \In dseg pp p null qn xs -> + valid h -> + h \In dseg pp p null qn xs -> [/\ p = qn, pp = null, xs = [::] & h = Unit]. Proof. case/lastP: xs=>[|xs x] D /=; first by case. @@ -66,10 +80,11 @@ Qed. (* deconstruct a non-trivial segment from the left *) Lemma dseg_neqL (xs : seq A) (pp p q qn : ptr) h : - p != qn -> h \In dseg pp p q qn xs -> + p != qn -> + h \In dseg pp p q qn xs -> exists x r h', [/\ xs = x :: behead xs, - h = p :-> x \+ (p.+ 1 :-> r \+ (p.+ 2 :-> pp \+ h')) & + h = p :-> x \+ (p.+1 :-> r \+ (p.+2 :-> pp \+ h')) & h' \In dseg p r q qn (behead xs)]. Proof. case: xs=>[|x xs] /= H; last first. @@ -79,10 +94,11 @@ Qed. (* deconstruct a non-trivial segment from the right *) Lemma dseg_neqR (xs : seq A) (pp p q qn : ptr) h : - pp != q -> h \In dseg pp p q qn xs -> + pp != q -> + h \In dseg pp p q qn xs -> exists s x r h', [/\ xs = rcons s x, - h = h' \+ (q :-> x \+ (q.+ 1 :-> qn \+ q.+ 2 :-> r)) & + h = h' \+ (q :-> x \+ (q.+1 :-> qn \+ q.+2 :-> r)) & h' \In dseg pp p r q s]. Proof. case/lastP: xs=>[|xs x] /= H. @@ -94,14 +110,14 @@ Qed. (* splitting the list corresponds to splitting the heap *) Lemma dseg_cat (xs ys : seq A) pp p q qn h : h \In dseg pp p q qn (xs++ys) <-> - exists jn j, h \In dseg pp p j jn xs # dseg j jn q qn ys. + exists jn j, h \In dseg pp p j jn xs # dseg j jn q qn ys. Proof. elim: xs pp p h=>/=. - move=>pp p h; split; first by move=>H; exists p, pp, Unit, h; rewrite unitL. by case=>jn [j][h1][h2][{h}-> [->->->]]; rewrite unitL. move=>x xs IH pp p h; split. - case=>r [_][{h}-> /IH][jn][j][h1][h2][-> H1 H2]. - exists jn, j, (p :-> x \+ p.+ 1 :-> r \+ p.+ 2 :-> pp \+ h1), h2. + exists jn, j, (p :-> x \+ p.+1 :-> r \+ p.+2 :-> pp \+ h1), h2. by rewrite !joinA; split=>//; exists r, h1; rewrite !joinA. case=>jn[j][_][h2][{h}->][r][h1][-> H1 H2]. exists r, (h1 \+ h2); rewrite !joinA; split=>//. @@ -120,28 +136,32 @@ Variable A : Type. (* specializing the segment lemmas *) Lemma dseq_nullL (xs : seq A) q h : - valid h -> h \In dseq null q xs -> + valid h -> + h \In dseq null q xs -> [/\ q = null, xs = [::] & h = Unit]. Proof. by move=>D; case/(dseg_nullL D). Qed. Lemma dseq_nullR (xs : seq A) p h : - valid h -> h \In dseq p null xs -> + valid h -> + h \In dseq p null xs -> [/\ p = null, xs = [::] & h = Unit]. Proof. by move=>D; case/(dseg_nullR D). Qed. Lemma dseq_posL (xs : seq A) p q h : - p != null -> h \In dseq p q xs -> + p != null -> + h \In dseq p q xs -> exists x r h', [/\ xs = x :: behead xs, - h = p :-> x \+ (p.+ 1 :-> r \+ (p.+ 2 :-> null \+ h')) & + h = p :-> x \+ (p.+1 :-> r \+ (p.+2 :-> null \+ h')) & h' \In dseg p r q null (behead xs)]. Proof. by apply: dseg_neqL. Qed. Lemma dseq_posR (xs : seq A) p q h : - q != null -> h \In dseq p q xs -> + q != null -> + h \In dseq p q xs -> exists s x r h', [/\ xs = rcons s x, - h = h' \+ (q :-> x \+ (q.+ 1 :-> null \+ q.+ 2 :-> r)) & + h = h' \+ (q :-> x \+ (q.+1 :-> null \+ q.+2 :-> r)) & h' \In dseg null p r q s]. Proof. by rewrite eq_sym=>Hq /(dseg_neqR Hq). Qed. @@ -150,50 +170,50 @@ Proof. by rewrite eq_sym=>Hq /(dseg_neqR Hq). Qed. (* prepend a value, return pointers to new first and last nodes *) Program Definition insertL p q (x : A) : - {l}, STsep (dseq p q l, [vfun pq => dseq pq.1 pq.2 (x::l)]) := + STsep {l} (dseq p q l, [vfun pq => dseq pq.1 pq.2 (x :: l)]) := Do (r <-- allocb x 3; - r.+ 1 ::= p;; - r.+ 2 ::= null;; + r.+1 ::= p;; + r.+2 ::= null;; if p == null then ret (r,r) - else p.+ 2 ::= r;; + else p.+2 ::= r;; ret (r,q)). Next Obligation. (* pull out ghost + precondition *) move=>p q x [l][] i /= H. (* create a new node in first 3 steps (+ rearrange pointer arith), branch *) -step=>r; rewrite unitR ptrA; do 2!step; case: ifP H=>[/eqP->|/negbT N]. +step=>r; rewrite unitR; do 2!step; case: ifP H=>[/eqP ->|/negbT N]. - (* the list is empty, so new first node = last node *) move/dseq_nullL=>H; step; rewrite joinA=>/validR/H [_->->] /=. by exists null, Unit; rewrite !joinA. (* deconstruct non-empty list, run the rest *) case/(dseq_posL N)=>y[z][h'][E {i}-> H']; do 2![step]=>V. (* massage the heap to fit the goal *) -exists p, (p :-> y \+ (p.+ 1 :-> z \+ (p.+ 2 :-> r \+ h'))). +exists p, (p :-> y \+ (p.+1 :-> z \+ (p.+2 :-> r \+ h'))). by rewrite !joinA; split=>//; rewrite E /=; exists z, h'; rewrite !joinA. Qed. (* append a value (in constant), return pointers to new first and last nodes *) Program Definition insertR p q (x : A) : - {l}, STsep (dseq p q l, [vfun pq => dseq pq.1 pq.2 (rcons l x)]) := + STsep {l} (dseq p q l, [vfun pq => dseq pq.1 pq.2 (rcons l x)]) := Do (r <-- allocb x 3; - r.+ 1 ::= null;; - r.+ 2 ::= q;; + r.+1 ::= null;; + r.+2 ::= q;; if q == null then ret (r,r) - else q.+ 1 ::= r;; + else q.+1 ::= r;; ret (p,r)). Next Obligation. (* pull out ghost + precondition *) move=>p q x [l []] i /= H. (* create a new node in first 3 steps (+ rearrange pointer arith), branch *) -step=>r; rewrite unitR ptrA; do 2!step; case: ifP H=>[/eqP->|/negbT N]. +step=>r; rewrite unitR; do 2!step; case: ifP H=>[/eqP->|/negbT N]. - (* the list is empty, so new first node = last node *) move/dseq_nullR=>H; step; rewrite joinA=>/validR/H [_->->] /=. by exists null, Unit; rewrite !joinA. (* deconstruct non-empty list, run the rest, restructure the goal *) case/(dseq_posR N)=>s[y][z][h'][{l}-> {i}-> H']; do 2![step]=>_; apply/dseg_rcons. (* massage the heap and simplify *) -exists q, (h' \+ (q :-> y \+ (q.+ 1 :-> r \+ q.+ 2 :-> z))). +exists q, (h' \+ (q :-> y \+ (q.+1 :-> r \+ q.+2 :-> z))). split; first by rewrite joinC. (* restructure the goal once more *) by apply/dseg_rcons; vauto. @@ -208,18 +228,18 @@ Qed. (* we carry the pointer to the accumulator as a ghost var *) Definition traverse_backT (p : ptr) : Type := forall (qs : ptr * seq A), - {l nx}, STsep (dseg null p qs.1 nx l, + STsep {l nx} (dseg null p qs.1 nx l, [vfun r h => h \In dseg null p qs.1 nx l /\ r = l ++ qs.2]). Program Definition traverse_back p q : - {l}, STsep (dseq p q l, + STsep {l} (dseq p q l, [vfun r h => h \In dseq p q l /\ r = l]) := Do (let tb := - Fix (fun (go : traverse_backT p) '(r, acc) => + ffix (fun (go : traverse_backT p) '(r, acc) => Do (if r == null then ret acc else x <-- !r; - rnxt <-- !(r.+ 2); + rnxt <-- !r.+2; go (rnxt, x :: acc))) in tb (q, [::])). (* first, the loop *) diff --git a/examples/hashtab.v b/examples/hashtab.v index 2e7f552..6689359 100644 --- a/examples/hashtab.v +++ b/examples/hashtab.v @@ -1,8 +1,22 @@ +(* +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 HB Require Import structures. From Coq Require Import ssreflect ssrbool ssrfun. From mathcomp Require Import ssrnat eqtype seq fintype tuple finfun finset. From pcm Require Import options axioms prelude pred ordtype finmap. From pcm Require Import pcm unionmap heap autopcm. -From htt Require Import model heapauto. +From htt Require Import options model heapauto. From htt Require Import array kvmaps. Module HashTab. @@ -41,7 +55,7 @@ Definition new_loopinv x := forall k, Program Definition new : STsep (emp, [vfun y => shape y nil]) := Do (t <-- Array.new _ (KVmap.default buckets); - let go := Fix (fun (loop : new_loopinv t) k => + let go := ffix (fun (loop : new_loopinv t) k => Do (if decP (b := k < n) idP is left pf then b <-- KVmap.new buckets; Array.write t (Ordinal pf) b;; @@ -69,7 +83,8 @@ exists [ffun z => if z == Ordinal pf then b else tab z], m2, (m \+ h2); split=>/ (* remove the new bucket from the heap *) rewrite (sepitS (Ordinal pf)) in_set leqnn {1}/table ffunE eq_refl; exists m, h2; do!split=>{m Hm}//. apply: tableP2 H2=>{h2}//. -- by case=>x Hx; rewrite !in_set -val_eqE /= ltnS (leq_eqVlt x); case: ltngtP. +- case=>x Hx; rewrite !in_set in_set1 -val_eqE /= ltnS (leq_eqVlt x). + by case: ltngtP. (* removing k from the domain of the new table gives the old table back *) by move=>x _; rewrite in_set ffunE; case: eqP=>//->; rewrite ltnn. Qed. @@ -97,10 +112,10 @@ Definition free_loopinv x := forall k, sepit [set x:'I_n | x >= k] (table t b), [vfun _ : unit => emp]). -Program Definition free x : {s}, STsep (shape x s, - [vfun _ : unit => emp]) := +Program Definition free x : STsep {s} (shape x s, + [vfun _ : unit => emp]) := (* we add an extra Do here so we can derive the precondition from the loop *) - Do (Fix (fun (loop : free_loopinv x) k => + Do (ffix (fun (loop : free_loopinv x) k => Do (if decP (b := k < n) idP is left pf then b <-- Array.read x (Ordinal pf); KVmap.free b;; @@ -127,7 +142,7 @@ apply: [stepX (bf (Ordinal pf))] @ h3=>{h3 H3}//= _ _ ->; rewrite unitL. apply: [gE]=>//=; split=>//; exists tf, bf, h1, h4; split=>//. (* drop the k-th entry from the table *) apply/tableP2/H4=>//. -move=>z; rewrite !in_set; case: eqP=>/=. +move=>z; rewrite !in_set in_set1; case: eqP=>/=. - by move=>->/=; rewrite ltnn. by move/eqP; rewrite -val_eqE /=; case: ltngtP. Qed. @@ -142,8 +157,8 @@ Qed. (* returning the pointer is technically not needed, as the array is not moved *) (* but we need to fit the KV map API *) -Program Definition insert x k v : {s}, STsep (shape x s, - [vfun y => shape y (ins k v s)]) := +Program Definition insert x k v : STsep {s} (shape x s, + [vfun y => shape y (ins k v s)]) := Do (let hk := hash k in b <-- Array.read x hk; b' <-- KVmap.insert b k v; @@ -177,14 +192,14 @@ exists m3, (m2 \+ h4); split=>{Hf Hh m3 E3 V3}//. (* split out the modified bucket *) rewrite (sepitT1 (hash k)) /table /= ffunE eq_refl; exists m2, h4; split=>{m2 H2}//. (* the table fits too *) -by apply/tableP/H4=>/= x0; rewrite !in_set andbT ?ffunE =>/negbTE->. +by apply/tableP/H4=>/= x0; rewrite !in_set in_set1 andbT ?ffunE =>/negbTE->. Qed. (* removing from a hashmap is removing from corresponding bucket + updating the array *) (* returning the pointer is again not needed except for the API fit *) Program Definition remove x k : - {s}, STsep (shape x s, + STsep {s} (shape x s, [vfun y => shape y (rem k s)]) := Do (let hk := hash k in b <-- Array.read x hk; @@ -219,13 +234,13 @@ exists m3, (m2\+ h4); split=>{m3 E3 V3 Hf Hh}//. (* split out the modified bucket *) rewrite (sepitT1 (hash k)) /table /= ffunE eq_refl; exists m2, h4; split=>{m2 H2}//. (* the table fits too *) -by apply/tableP/H4=>/= x0; rewrite !in_set andbT ?ffunE =>/negbTE->. +by apply/tableP/H4=>/= x0; rewrite !in_set in_set1 andbT ?ffunE =>/negbTE->. Qed. (* looking up in a hashmap is looking up in the corresponging bucket *) Program Definition lookup x k : - {s}, STsep (shape x s, + STsep {s} (shape x s, [vfun y m => m \In shape x s /\ y = fnd k s]) := Do (b <-- Array.read x (hash k); KVmap.lookup b k). @@ -251,4 +266,4 @@ Definition HashTab := KVmap.make (Array null) new free insert remove lookup. End HashTab. End HashTab. -Definition HT K V := HashTab.HashTab (AssocList.AssocList K V). \ No newline at end of file +Definition HT K V := HashTab.HashTab (AssocList.AssocList K V). diff --git a/examples/kvmaps.v b/examples/kvmaps.v index 1909fe6..5ad60b7 100644 --- a/examples/kvmaps.v +++ b/examples/kvmaps.v @@ -1,10 +1,23 @@ +(* +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. +*) + (******************) (* Key-Value maps *) (******************) From Coq Require Import ssreflect ssrbool ssrfun. -From mathcomp Require Import eqtype seq path. -From pcm Require Import options axioms pred ordtype finmap. +From mathcomp Require Import eqtype seq path ssrnat. +From pcm Require Import options axioms pred ordtype finmap seqext. From pcm Require Import pcm unionmap heap autopcm automap. From htt Require Import options model heapauto. @@ -15,27 +28,20 @@ From htt Require Import options model heapauto. Module KVmap. Record Sig (K : ordType) (V : Type) : Type := make {tp :> Type; - default : tp; - shape : tp -> {finMap K -> V} -> Pred heap; - new : STsep (emp, [vfun x => shape x (nil K V)]); - - free : forall x, {s}, STsep (shape x s, - [vfun _ : unit => emp]); - + free : forall x, STsep {s} (shape x s, + [vfun _ : unit => emp]); insert : forall x k v, - {s}, STsep (shape x s, - [vfun y => shape y (ins k v s)]); - + STsep {s} (shape x s, + [vfun y => shape y (ins k v s)]); remove : forall x k, - {s}, STsep (shape x s, - [vfun y => shape y (rem k s)]); - + STsep {s} (shape x s, + [vfun y => shape y (rem k s)]); lookup : forall x k, - {s}, STsep (shape x s, - [vfun y m => m \In shape x s /\ y = fnd k s])}. + STsep {s} (shape x s, + [vfun y m => m \In shape x s /\ y = fnd k s])}. End KVmap. (**********************************************************) @@ -44,8 +50,7 @@ End KVmap. Module AssocList. Section AssocList. - -Variable (K : ordType) (V : Set). +Variables (K : ordType) (V : Set). Notation fmap := (finMap K V). Notation nil := (nil K V). @@ -72,8 +77,9 @@ Definition shape (x : ptr) (s : fmap) : Pred heap := (* null pointer represents an empty map *) Lemma shape_null (s : fmap) h : - valid h -> h \In shape null s -> - s = nil /\ h = Unit. + valid h -> + h \In shape null s -> + s = nil /\ h = Unit. Proof. move=>D; case: s; case=>/=. - by move=>? [_ ->] //; rewrite fmapE. @@ -85,30 +91,34 @@ Qed. (* a head entry respecting the key order can be pulled out of the map *) Lemma shape_cont (s : fmap) p h : - p != null -> h \In shape p s -> - exists k v q h', - [/\ s = ins k v (behd s), (* k:->v is the head entry *) - all (ord k) (supp (behd s)), - h = entry p q k v \+ h' & - h' \In shape q (behd s)]. + p != null -> + h \In shape p s -> + exists k v q h', + [/\ s = ins k v (behd s), (* k:->v is the head entry *) + all (ord k) (supp (behd s)), + h = entry p q k v \+ h' & + h' \In shape q (behd s)]. Proof. move=>E; case: s=>xs srt /=. elim: xs srt=>/=. - by move=>? [E0]; rewrite E0 in E. move=>[k v] /= l IH srt [q][h'][-> H]. -exists k,v,q,h'; split=>//; last by apply: path_all. +exists k,v,q,h'; split=>//; last by apply: order_path_min. by rewrite fmapE /= last_ins'. Qed. (* TODO move to finmap? *) -Lemma all_path (s : fmap) k : all (ord k) (supp s) -> path ord k (supp s). +Lemma all_path (s : fmap) k : + all (ord k) (supp s) -> + path ord k (supp s). Proof. by rewrite path_sortedE // =>->/=; case: s. Qed. (* inserting an entry with minimal key is prepending to the list *) Lemma shape_cons (s : fmap) p q h k v : - all (ord k) (supp s) -> h \In shape q s -> - entry p q k v \+ h \In shape p (ins k v s). + all (ord k) (supp s) -> + h \In shape q s -> + entry p q k v \+ h \In shape p (ins k v s). Proof. move/all_path=>S H. suff: ins k v s = @FinMap _ _ ((k,v)::seq_of s) S by move=>->; exists q,h. @@ -119,10 +129,10 @@ Qed. (* inserting an entry with maximal key is appending to the list *) Lemma shape_seg_rcons (s : fmap) x p q h k v : - (* conceptually last s < k *) - all (ord^~ k) (supp s) -> - h \In shape_seg x p s -> - (h \+ entry p q k v) \In shape_seg x q (ins k v s). + (* conceptually last s < k *) + all (ord^~ k) (supp s) -> + h \In shape_seg x p s -> + (h \+ entry p q k v) \In shape_seg x q (ins k v s). Proof. case: s=>xs; elim: xs h x=>/=. - move=>??? _ [->->]; rewrite unitL. @@ -137,10 +147,11 @@ Qed. (* disjoint maps can be concatenated if the order is respected *) Lemma shape_fcat s1 s2 h1 h2 x y : - (* conceptually last s1 < head s2 *) - allrel ord (supp s1) (supp s2) -> - h1 \In shape_seg x y s1 -> h2 \In shape y s2 -> - h1 \+ h2 \In shape x (fcat s1 s2). + (* conceptually last s1 < head s2 *) + allrel ord (supp s1) (supp s2) -> + h1 \In shape_seg x y s1 -> + h2 \In shape y s2 -> + h1 \+ h2 \In shape x (fcat s1 s2). Proof. move=>O1 H1. case: s2 O1=>xs; elim: xs h1 y h2 s1 H1=>/=. @@ -149,7 +160,7 @@ move=>[k' v'] xs; rewrite /fcat /= => IH /= h1 y h2 s1 H1 srt O2 H2. case: H2=>z[h'][-> H']; rewrite joinA; apply: IH; first 1 last. - by apply/path_sorted/srt. - move=>H0; rewrite (allrel_in_l (xs':=k'::supp s1) _); last by apply: supp_ins. - rewrite allrel_consl path_all //=. + rewrite allrel_consl order_path_min //=. by apply/allrel_sub_r/O2=>?; rewrite inE orbC=>->. - by move=>?; apply: H'. apply: shape_seg_rcons=>//. @@ -167,15 +178,15 @@ Next Obligation. by move=>[] /= _ ->; step. Qed. (* freeing a map is deallocating all its elements *) Definition freeT : Type := - forall p, {fm}, STsep (shape p fm, [vfun _ : unit => emp]). + forall p, STsep {fm} (shape p fm, [vfun _ : unit => emp]). Program Definition free : freeT := - Fix (fun (loop : freeT) p => + ffix (fun (loop : freeT) p => Do (if p == null then ret tt - else pnext <-- !(p.+2); + else pnext <-- !p.+2; dealloc p;; - dealloc (p.+1);; - dealloc (p.+2);; + dealloc p.+1;; + dealloc p.+2;; loop pnext)). Next Obligation. (* pull out ghost var, precondition, branch *) @@ -193,21 +204,21 @@ Qed. (* looking up an element in the map *) Definition lookupT k : Type := - forall p, {fm}, STsep (shape p fm, - [vfun y m => m \In shape p fm /\ y = fnd k fm]). + forall p, STsep {fm} (shape p fm, + [vfun y m => m \In shape p fm /\ y = fnd k fm]). Program Definition lookup x (k : K) : - {fm}, STsep (shape x fm, - [vfun y m => m \In shape x fm /\ y = fnd k fm]) := - Fix (fun (loop : lookupT k) (cur : ptr) => + STsep {fm} (shape x fm, + [vfun y m => m \In shape x fm /\ y = fnd k fm]) := + ffix (fun (loop : lookupT k) (cur : ptr) => Do (if cur == null then ret None else k' <-- !cur; if k == k' - then v <-- !(cur.+1); + then v <-- !cur.+1; ret (Some v) else if ord k' k - then next <-- !(cur.+2); + then next <-- !cur.+2; loop next else ret None)) x. Next Obligation. @@ -228,7 +239,7 @@ case: ifP=>Ho'. step; apply/[gR (behd fm)] @ h'=>//= v0 h0' [H0 E0] _. by rewrite Ef fnd_ins (negbTE Ek); split=>//; exact: shape_cons. (* head key is bigger than the needed one, abort *) -move: (semiconnex Ek); rewrite {}Ho' orbC /= =>Ho'; step=>_; split=>//. +move: (connex Ek); rewrite {}Ho' orbC /= =>Ho'; step=>_; split=>//. (* k is not in the head entry *) apply/esym/fnd_supp; rewrite Ef supp_ins inE negb_or; apply/andP; split=>//. (* nor it is in the tail *) @@ -244,7 +255,7 @@ Qed. (* the focus is needed so that we can connect the remainder of the map to it after removal *) Definition removeT p k : Type := forall (prevcur : ptr * ptr), - {fm}, STsep (fun h => exists fml fmr k' v', + STsep {fm} (fun h => exists fml fmr k' v', [/\ fm = fcat (ins k' v' fml) fmr, all (ord^~ k') (supp fml) /\ all (ord k') (supp fmr), k \notin supp fml /\ k != k' & @@ -255,33 +266,33 @@ Definition removeT p k : Type := [vfun _ : unit => shape p (rem k fm)]). Program Definition remove x (k : K) : - {fm}, STsep (shape x fm, - [vfun y => shape y (rem k fm)]) := - Do (let go := Fix (fun (loop : removeT x k) '(prev, cur) => + STsep {fm} (shape x fm, + [vfun y => shape y (rem k fm)]) := + Do (let go := ffix (fun (loop : removeT x k) '(prev, cur) => Do (if cur == null then ret tt else k' <-- !cur; if k == k' - then next <-- !(cur.+2); + then next <-- !cur.+2; dealloc cur;; - dealloc (cur.+1);; - dealloc (cur.+2);; + dealloc cur.+1;; + dealloc cur.+2;; prev.+2 ::= (next : ptr);; ret tt else if ord k' k - then next <-- !(cur.+2); + then next <-- !cur.+2; loop (cur, next) else ret tt)) in if x == null then ret null else k' <-- !x; if k == k' - then next <-- !(x.+2); + then next <-- !x.+2; dealloc x;; - dealloc (x.+1);; - dealloc (x.+2);; + dealloc x.+1;; + dealloc x.+2;; ret next else if ord k' k - then next <-- !(x.+2); + then next <-- !x.+2; go (x, next);; ret x else ret x). @@ -313,8 +324,8 @@ step; case: eqP=>[|/eqP] Ek. rewrite joinC; apply/shape_fcat/Hr'; last by apply: shape_seg_rcons. (* the ordering is respected as well *) rewrite (allrel_in_l (xs':=k::supp fml) _); last by apply: supp_ins. - rewrite allrel_consl path_all //=. - by apply/(allrel_trans (z:=k))/path_all=>//; exact: trans. + rewrite allrel_consl order_path_min //=. + by apply/(allrel_trans (z:=k))/order_path_min=>//=. (* k <> k', now branch on order comparison *) case: ifP=>Ho0. - (* k' is less than k, invoke the loop, shifting the focus to the right *) @@ -325,7 +336,7 @@ case: ifP=>Ho0. rewrite (eq_all_r (s2:=k::supp fml)) /= ?Ho' /=; last by apply: supp_ins. by apply/sub_all/Ol=>? /trans; apply. - (* new focus comes before the new suffix *) - by apply: path_all. + by apply: order_path_min. - (* the needed key is not in fml ++ old focus *) by rewrite supp_ins inE negb_or E. (* heap shape is respected *) @@ -336,12 +347,12 @@ case: ifP=>Ho0. move=>_ m Hm Vm; rewrite Efr. by rewrite fcat_inss // -?fcat_sins // in Hm; apply: notin_path. (* k' is bigger than k, abort *) -move: (semiconnex Ek); rewrite {}Ho0 orbC /= =>Ho0. +move: (connex Ek); rewrite {}Ho0 orbC /= =>Ho0. step=>_; rewrite rem_supp. - (* the shape is preserved *) rewrite joinC; apply: shape_fcat; first 1 last. - by apply: shape_seg_rcons. - - by rewrite Efr; apply: shape_cons=>//; apply: path_all. + - by rewrite Efr; apply: shape_cons=>//; apply: order_path_min. (* ordering is preserved *) rewrite (allrel_in_l (xs':=k::supp fml) _); last by apply: supp_ins. rewrite allrel_consl Or /=. @@ -375,7 +386,7 @@ case: ifP=>Ho0. - by rewrite fcat_inss; [rewrite fcat0s|apply/notin_path/all_path]. by exists Unit, (entry x next k v \+ h'); split=>//; [rewrite unitL | vauto]. (* k' is bigger than k, abort *) -move: (semiconnex Ek); rewrite {}Ho0 orbC /= =>Ho0. +move: (connex Ek); rewrite {}Ho0 orbC /= =>Ho0. (* return the old head, invariants are preserved *) step=>_; rewrite -Eh rem_supp // Ef supp_ins !inE negb_or Ek /=. by apply/notin_path/path_le/all_path/Of. @@ -390,7 +401,7 @@ Qed. (* the focus is needed so that we can connect the new element to it after insertion *) Definition insertT p k v : Type := forall (prevcur : ptr * ptr), - {fm}, STsep (fun h => exists fml fmr k' v', + STsep {fm} (fun h => exists fml fmr k' v', [/\ fm = fcat (ins k' v' fml) fmr, all (ord^~ k') (supp fml) /\ all (ord k') (supp fmr), k \notin supp fml /\ ord k' k & @@ -401,9 +412,9 @@ Definition insertT p k v : Type := [vfun _ : unit => shape p (ins k v fm)]). Program Definition insert x (k : K) (v : V) : - {fm}, STsep (shape x fm, + STsep {fm} (shape x fm, [vfun y => shape y (ins k v fm)]) := - Do (let go := Fix (fun (loop : insertT x k v) '(prev, cur) => + Do (let go := ffix (fun (loop : insertT x k v) '(prev, cur) => Do (if cur == null then new <-- allocb k 3; new.+1 ::= v;; @@ -415,7 +426,7 @@ Program Definition insert x (k : K) (v : V) : then cur.+1 ::= v;; ret tt else if ord k' k - then next <-- !(cur.+2); + then next <-- !cur.+2; loop (cur, next) else new <-- allocb k 3; new.+1 ::= v;; @@ -433,7 +444,7 @@ Program Definition insert x (k : K) (v : V) : then x.+1 ::= v;; ret x else if ord k' k - then next <-- !(x.+2); + then next <-- !x.+2; go (x, next);; ret x else new <-- allocb k 3; @@ -449,7 +460,7 @@ move=>x k0 v0 loop _ prev cur [_][] _ /= [fml][fmr][k][v][-> [Ol Or][El Ho0]][hl case: eqP=>[|/eqP] Ec. - (* cur = null, insert as the last element *) rewrite {}Ec in Hr; apply: vrfV=>Vh. - step=>p; rewrite ptrA unitR; do 2!step; rewrite joinC; do 2![step]=>_. + step=>p; rewrite unitR; do 2!step; rewrite joinC; do 2![step]=>_. (* fmr is empty *) case: (shape_null (validX Vh) Hr)=>/=->->. rewrite fcats0 unitR [X in _ \+ entry _ _ _ _ \+ X]joinA. @@ -460,7 +471,8 @@ case: eqP=>[|/eqP] Ec. (* cur <> null, pull out the head entry from fmr *) case: (shape_cont Ec Hr)=>k'[v'][next][hr'][Efr Or' {hr Hr Ec}-> Hr']; rewrite joinA joinC. (* derive ordering facts *) -move/all_path: (Or); rewrite {1}Efr; case/(path_supp_ins_inv (all_path Or'))/andP=>Ho' /(path_all (@trans _)) Or''. +move/all_path: (Or); rewrite {1}Efr; case/(path_supp_ins_inv (all_path Or'))/andP=>Ho'. +move/(order_path_min (@trans _))=>Or''. (* get new key, branch on equality comparison *) step; case: eqP=>[|/eqP] Ek. - (* k = k', update the value at this position *) @@ -494,15 +506,15 @@ case: ifP=>Ho'0. rewrite fcat_inss // in Hm; first by rewrite -fcat_sins in Hm. by apply/notin_path/all_path. (* k' is bigger than k, insert at this position *) -move: (semiconnex Ek); rewrite {}Ho'0 orbC /= =>Ho0'. +move: (connex Ek); rewrite {}Ho'0 orbC /= =>Ho0'. (* run the allocation and assignments *) -step=>new; rewrite ptrA unitR; do 2!step; rewrite joinA joinC; do 2![step]=>_. +step=>new; rewrite unitR; do 2!step; rewrite joinA joinC; do 2![step]=>_. rewrite Efr -fcat_sins; apply: shape_fcat; first 1 last. - (* shape is respected for the prefix fml ++ old focus *) by apply: shape_seg_rcons. - (* shape is satisfed for new element + suffix *) rewrite [X in X \+ (entry _ _ _ _ \+ _)]joinA; apply/shape_cons/shape_cons=>//. - by apply/path_all/path_supp_ins=>//; apply/path_le/all_path/Or'. + by apply/order_path_min=>//; apply/path_supp_ins=>//; apply/path_le/all_path/Or'. (* ordering is respected *) rewrite (allrel_in_l (xs':=k::supp fml) _); last by apply: supp_ins. rewrite (allrel_in_r (ys':=k0::k'::supp (behd fmr)) _ _); last first. @@ -521,7 +533,7 @@ move=>/= x k0 v0 [fm][]h /= H; case: eqP=>[|/eqP] Ex. - (* x = null, insert as the only element, heap and spec are empty *) apply: vrfV=>Vh; move: H; rewrite Ex=>/(shape_null Vh) [->->]. (* run *) - step=>p; rewrite ptrA !unitR; do 3![step]=>_. + step=>p; rewrite !unitR; do 3![step]=>_. by exists null, Unit; rewrite unitR joinA. (* x <> null, pull out the head entry *) case: (shape_cont Ex H)=>{Ex}k[v][next][h'][Ef Of {h H}-> H]. @@ -539,12 +551,12 @@ case: ifP=>Ho0. - by rewrite fcat_inss; [rewrite fcat0s|apply/notin_path/all_path]. by exists Unit, (entry x next k v \+ h'); split=>//; [rewrite unitL | vauto]. (* k' is bigger than k, insert after the head *) -move: (semiconnex Ek); rewrite {}Ho0 orbC /= =>Ho0. +move: (connex Ek); rewrite {}Ho0 orbC /= =>Ho0. (* run allocation and assignments, return the old head *) -step=>q; rewrite ptrA unitR; do 3![step]=>_. +step=>q; rewrite unitR; do 3![step]=>_. (* invariants are preserved *) rewrite Ef [X in X \+ (entry _ _ _ _ \+ _)]joinA; apply/shape_cons/shape_cons=>//. -by apply/path_all/path_supp_ins=>//; apply/path_le/all_path/Of. +by apply/order_path_min=>//; apply/path_supp_ins=>//; apply/path_le/all_path/Of. Qed. (* ordered association list is a KV map *) diff --git a/examples/queue.v b/examples/queue.v index 3007251..0e4a15b 100644 --- a/examples/queue.v +++ b/examples/queue.v @@ -1,8 +1,21 @@ +(* +Copyright 2021 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 seq. +From mathcomp Require Import eqtype ssrnat seq. From pcm Require Import options axioms pred. From pcm Require Import pcm unionmap heap automap. -From htt Require Import model heapauto. +From htt Require Import options model heapauto. From htt Require Import llist. Record queue (T : Type) : Type := Queue {front: ptr; back: ptr}. @@ -18,8 +31,8 @@ Definition is_queue (fr bq : ptr) (xs : seq T) := if fr == null then [Pred h | [/\ bq = null, xs = [::] & h = Unit]] else [Pred h | exists xt x h', [/\ xs = rcons xt x, - valid (h' \+ (bq :-> x \+ bq .+ 1 :-> null)), - h = h' \+ (bq :-> x \+ bq .+ 1 :-> null) & + valid (h' \+ (bq :-> x \+ bq.+1 :-> null)), + h = h' \+ (bq :-> x \+ bq.+1 :-> null) & h' \In lseg fr bq xt]]. (* the structure itself is a pair of pointers to body + last node *) @@ -38,7 +51,8 @@ Proof. by case=>h1[bq][h'] [] D ->. Qed. (* empty queue is a pair of null pointers *) Lemma is_queue_nil fr bq h : - h \In is_queue fr bq [::] -> [/\ fr = null, bq = null & h = Unit]. + h \In is_queue fr bq [::] -> + [/\ fr = null, bq = null & h = Unit]. Proof. by rewrite /is_queue; case: eqP=>[->[-> _ ->] | _ [[|y xt][x][h'][]]]. Qed. @@ -47,8 +61,8 @@ Qed. Lemma is_queue_rcons fr bq xt x h : h \In is_queue fr bq (rcons xt x) <-> - (exists h', [/\ valid (h' \+ (bq :-> x \+ bq .+ 1 :-> null)), - h = h' \+ (bq :-> x \+ bq .+ 1 :-> null) & + (exists h', [/\ valid (h' \+ (bq :-> x \+ bq.+1 :-> null)), + h = h' \+ (bq :-> x \+ bq.+1 :-> null) & h' \In lseg fr bq xt]). Proof. rewrite /is_queue; split. @@ -62,7 +76,8 @@ Qed. (* pointers should agree in a well-formed queue *) Lemma backfront fr bq xs h : - h \In is_queue fr bq xs -> (fr == null) = (bq == null). + h \In is_queue fr bq xs -> + (fr == null) = (bq == null). Proof. rewrite /is_queue; case: ifP=>[E [->]_ _| E [xt][x][h'][_] D] //. by case: eqP D=>// -> /validR; rewrite validPtUn. @@ -92,7 +107,7 @@ Program Definition free (q : queue) : dealloc (back q)). Next Obligation. (* pull out ghosts and precondition *) -move=>q [] _ /= [fr][bq][h][/[swap]->/[swap]]. +move=>q [_][fr][bq][h][/[swap] -> /[swap]]. (* both pointers are null *) case/is_queue_nil=>->->->; rewrite unitR=>V. (* run the program *) @@ -104,15 +119,15 @@ Qed. (* enqueuing is adding a node at the end *) Program Definition enq (q : queue) (x : T) : - {xs}, STsep (shape q xs, - [vfun _ => shape q (rcons xs x)]) := + STsep {xs} (shape q xs, + [vfun _ => shape q (rcons xs x)]) := Do (next <-- allocb null 2; next ::= x;; ba <-- !back q; back q ::= next;; (if (ba : ptr) == null then front q - else ba .+ 1) ::= next). + else ba.+1) ::= next). Next Obligation. (* pull out ghosts + precondition *) move=>q x [xs][] _ /= [fr][bq][h'][D -> H]. @@ -123,17 +138,17 @@ rewrite -(backfront H) unitR; case: ifP H=>Ef; rewrite /is_queue ?Ef. - (* the queue was empty, set the front pointer to new node *) case=>_->->; step; rewrite unitR=>V. (* massage the heap and restructure the goal *) - exists next, next, (next :-> x \+ next.+ 1 :-> null). + exists next, next, (next :-> x \+ next.+1 :-> null). rewrite joinA joinC; split=>//; apply/(@is_queue_rcons _ _ [::]). by exists Unit; rewrite unitL; split=>//; exact: (validL V). (* the queue wasn't empty, link the new node to the last one *) case=>s2[x2][i2][->] {}D -> H2; step=>V. (* massage the heap and simplify the goal *) -exists fr, next, (i2 \+ bq :-> x2 \+ bq.+ 1 :-> next \+ next :-> x \+ next.+ 1 :-> null). +exists fr, next, (i2 \+ bq :-> x2 \+ bq.+1 :-> next \+ next :-> x \+ next.+1 :-> null). split; first by apply: (validX V). - by rewrite joinC !joinA. (* the new node conforms to the queue spec *) -apply/is_queue_rcons; exists (i2 \+ bq :-> x2 \+ bq.+ 1 :-> next). +apply/is_queue_rcons; exists (i2 \+ bq :-> x2 \+ bq.+1 :-> next). rewrite joinA; split=>//; first by apply: (validX V). (* assemble the old queue back *) by apply/lseg_rcons; exists bq, i2; rewrite joinA. @@ -142,18 +157,18 @@ Qed. (* dequeuing is removing the head node and adjusting pointers *) Program Definition deq (q : queue) : - {xs}, STsep (shape q xs, - fun y h => shape q (behead xs) h /\ - match y with Val v => xs = v :: behead xs - | Exn e => e = EmptyQueue /\ xs = [::] end) := + STsep {xs} (shape q xs, + fun y h => shape q (behead xs) h /\ + match y with Val v => xs = v :: behead xs + | Exn e => e = EmptyQueue /\ xs = [::] end) := Do (fr <-- !front q; if (fr : ptr) == null then throw EmptyQueue else x <-- !fr; - next <-- !fr .+ 1; + next <-- !fr.+1; front q ::= next;; dealloc fr;; - dealloc fr .+ 1;; + dealloc fr.+1;; if (next : ptr) == null then back q ::= null;; ret x @@ -182,7 +197,7 @@ case: ifP H=>[/eqP ->|N] H. case/(lseg_null (validX V2)): H D=>/=-> _ _ /validR. by rewrite validPtUn. (* return the segment head and simplify *) -step=>V; split=>//; exists next, bq, (h2 \+ (bq :-> x \+ bq .+ 1 :-> null)). +step=>V; split=>//; exists next, bq, (h2 \+ (bq :-> x \+ bq.+1 :-> null)). by rewrite N; split=>//; vauto; apply: (validX V). Qed. diff --git a/examples/quicksort.v b/examples/quicksort.v index 6e0af19..586a164 100644 --- a/examples/quicksort.v +++ b/examples/quicksort.v @@ -562,8 +562,7 @@ move: (ltn_ord v); rewrite ltnS leq_eqVlt; case/orP=>[/eqP Ev|Nv]. by rewrite !inE -eqn_leq =>/eqP E; apply/eqP/ord_inj. move: Sl Hpl; rewrite Eh Ev Epr mul1g => Sl Hpl. rewrite slice_xR; last by rewrite bnd_simp leEnat; move: Hvl; rewrite Ev. - rewrite {22}(_ : n = (ord_max : 'I_n.+1)) // onth_codom /= sorted_rconsE //; - last by apply: otrans. + rewrite {22}(_ : n = (ord_max : 'I_n.+1)) // onth_codom /= sorted_rconsE //=. move: Sl; rewrite slice_oPR; last by rewrite lt0n -Ev. move=>->; rewrite andbT; move: Al; rewrite Ev. have ->: pffun (pl * p) f ord_max = pffun p f ord_max. diff --git a/examples/tree.v b/examples/tree.v index 1a99b79..542437f 100644 --- a/examples/tree.v +++ b/examples/tree.v @@ -1,3 +1,17 @@ +(* +Copyright 2023 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 HB Require Import structures. From mathcomp Require Import ssreflect ssrbool ssrfun. From mathcomp Require Import eqtype ssrnat seq bigop choice. From pcm Require Import pred prelude seqext. @@ -94,8 +108,9 @@ Qed. End EncodeDecodeTree. -Definition tree_eqMixin (A : eqType) := PcanEqMixin (@pcancel_tree A). -Canonical tree_eqType (A : eqType) := Eval hnf in EqType _ (@tree_eqMixin A). + +HB.instance Definition _ (A : eqType) := + Equality.copy (tree A) (pcan_type (pcancel_tree A)). Section TreeEq. Context {A : eqType}. @@ -134,7 +149,7 @@ Lemma in_preorder t : preorder t =i t. Proof. elim/tree_ind1: t=>t cs IH x. rewrite preorderE in_tnode /= inE; case: eqVneq=>//= N. -rewrite big_cat_mem_seq; apply: eq_in_has=>z Hz. +rewrite big_cat_mem_has; apply: eq_in_has=>z Hz. by rewrite IH //; apply/mem_seqP. Qed. diff --git a/examples/union_find.v b/examples/union_find.v index ec4e775..e8eae48 100644 --- a/examples/union_find.v +++ b/examples/union_find.v @@ -1,3 +1,16 @@ +(* +Copyright 2023 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 mathcomp Require Import ssreflect ssrbool ssrfun fintype. From mathcomp Require Import eqtype ssrnat seq bigop choice. From pcm Require Import options axioms pred seqext. @@ -15,18 +28,18 @@ Qed. Section BigOpsUM. Variables (K : ordType) (C : pred K) (T : Type). -Variables (U : union_map_class C T). +Variables (U : union_map C T). Variables (I : Type) (f : I -> U). Lemma big_find_someXP (xs : seq I) P a v : - find a (\big[PCM.join/Unit]_(i <- xs | P i) f i) = Some v -> + find a (\big[join/Unit]_(i <- xs | P i) f i) = Some v -> exists i, [/\ i \In xs, P i & find a (f i) = Some v]. Proof. by rewrite -big_filter=>/big_find_someX [i] /Mem_filter [H1 H2 H3]; exists i. Qed. Lemma bigInXP (xs : seq I) P a v : - (a, v) \In \big[PCM.join/Unit]_(i <- xs | P i) f i -> + (a, v) \In \big[join/Unit]_(i <- xs | P i) f i -> exists i, [/\ i \In xs, P i & (a, v) \In f i]. Proof. by case/In_find/big_find_someXP=>x [X1 X2 /In_find]; exists x. Qed. @@ -34,17 +47,17 @@ End BigOpsUM. Section OMapBig. Variables (K : ordType) (C : pred K) (T T' : Type). -Variables (U : union_map_class C T) (U' : union_map_class C T'). +Variables (U : union_map C T) (U' : union_map C T'). Variables (I : Type) (f : I -> U). Lemma omapVUnBig (a : K * T -> option T') ts : - omap a (\big[PCM.join/Unit]_(x <- ts) f x) = - if valid (\big[PCM.join/Unit]_(x <- ts) f x) - then \big[PCM.join/Unit]_(x <- ts) omap a (f x) - else um_undef : U'. + omap a (\big[join/Unit]_(x <- ts) f x) = + if valid (\big[join/Unit]_(x <- ts) f x) + then \big[join/Unit]_(x <- ts) omap a (f x) + else undef : U'. Proof. elim: ts=>[|t ts IH]; first by rewrite !big_nil omap0 valid_unit. -by rewrite !big_cons omapVUn IH; case: ifP=>//= /validR=>->. +by rewrite !big_cons omapVUn IH; case: ifP=>// /validR ->. Qed. End OMapBig. @@ -68,7 +81,7 @@ End OMapBig. Fixpoint tlay (t : tree ptr) (r : ptr) : heap := foldr (fun x h => h \+ tlay x (rt t)) (rt t :-> r) (ch t). -Fixpoint tset (t : tree ptr) (r : ptr) : union_mapUMC ptr_ordType ptr := +Fixpoint tset (t : tree ptr) (r : ptr) : umap ptr ptr := foldr (fun t h => h \+ tset t r) (rt t \\-> r) (ch t). (* explicit equations for expanding the defs of tlayout and tset *) @@ -77,7 +90,7 @@ Fixpoint tset (t : tree ptr) (r : ptr) : union_mapUMC ptr_ordType ptr := Lemma tlayE t r : tlay t r = - rt t :-> r \+ \big[PCM.join/Unit]_(x <- ch t) (tlay x (rt t)). + rt t :-> r \+ \big[join/Unit]_(x <- ch t) (tlay x (rt t)). Proof. case: t=>a ts /=; rewrite foldr_join; congr (_ \+ _). elim: ts=>[|t ts IH] /=; first by rewrite big_nil. @@ -86,7 +99,7 @@ Qed. Lemma tsetE t r : tset t r = - rt t \\-> r \+ \big[PCM.join/Unit]_(x <- ch t) (tset x r). + rt t \\-> r \+ \big[join/Unit]_(x <- ch t) (tset x r). Proof. case: t=>a ts /=; rewrite foldr_join; congr (_ \+ _). elim: ts=>[|t ts IH] /=; first by rewrite big_nil. @@ -105,25 +118,25 @@ Lemma valid_dom_tset (t : tree ptr) r : if valid (tset t r) then preorder t else [::]). Proof. elim/tree_ind2: t r=>a ts IH r. rewrite tsetE preorderE. -rewrite validPtUn !big_valid_dom_seq /= big_cat_mem_seq. +rewrite validPtUn !big_valid_dom_seq /= big_cat_mem_has. case: allP=>H /=; last first. - split=>[|x]; last first. - by rewrite domUn inE validPtUn !big_valid_dom_seq; case: allP. - rewrite andbC; case: uniq_big_catP=>//=; case=>H1 _ _. + rewrite andbC; case: uniq_big_catE=>//=; case=>H1 _ _. by case: H=>t T; rewrite IH // H1. case U1: (uniq _)=>/=; last first. - rewrite andbC; case U2: (uniq _)=>/=; last first. - by split=>// x; rewrite domPtUn inE validPtUn /= !big_valid_dom_seq U1 andbF. - case/uniq_big_catP: U2=>H1 H2 H3. - case: uniq_big_catP U1=>//; case; split. + case/uniq_big_catE: U2=>H1 H2 H3. + case: uniq_big_catE U1=>//; case; split. - by move=>i; rewrite uniq_dom. - move=>i k X D; apply: (H2 i k X). by rewrite (IH i X r) H in D. move=>i j k X Y Di Dj. apply: (H3 i j k)=>//; first by rewrite (IH i X r) H in Di. by rewrite (IH j Y r) H in Dj. -rewrite big_cat_mem_seq andbC; case: uniq_big_catP=>/=; last first. -- case/uniq_big_catP: U1=>H1 H2 H3; case; split. +rewrite big_cat_mem_has andbC; case: uniq_big_catE=>/=; last first. +- case/uniq_big_catE: U1=>H1 H2 H3; case; split. - by move=>i X; rewrite -(IH i X r) H. - by move=>i k X D; apply: (H2 i k X); rewrite (IH i X r) H. move=>i j k X Y Di Dj; apply: (H3 i j k X Y); first by rewrite (IH i X r) H. @@ -131,8 +144,8 @@ rewrite big_cat_mem_seq andbC; case: uniq_big_catP=>/=; last first. case=>K1 K2 K3; split=>[|x]. - by rewrite -!all_predC; apply: eq_in_all=>i X; rewrite /= IH ?H. rewrite domPtUn inE validPtUn /= !big_valid_dom_seq U1 andbT. -case: allP=>//= _; rewrite !big_cat_mem_seq -all_predC. -case: allP=>//= A; rewrite inE big_cat_mem_seq eq_sym. +case: allP=>//= _; rewrite !big_cat_mem_has -all_predC. +case: allP=>//= A; rewrite inE big_cat_mem_has eq_sym. by case: (x =P a)=>//= _; apply: eq_in_has=>i X; rewrite IH ?H. Qed. @@ -174,10 +187,10 @@ Qed. Lemma domeq_tlay_tset (t : tree ptr) r1 r2 : {in t, forall x, x != null} -> - dom_eq2 (tlay t r1) (tset t r2). + dom_eq (tlay t r1) (tset t r2). Proof. elim/tree_ind1: t r1 r2=>a ts IH r1 r2 Tn. -rewrite tlayE tsetE /= domeq2PtUn ?Tn // big_domeq2Un //. +rewrite tlayE tsetE /= domeqPtUn ?Tn // big_domeqUn //. by move=>x X; apply: IH=>// z Z; apply: Tn (in_tnode2 X Z). Qed. @@ -194,9 +207,9 @@ Lemma valid_tlay (t : tree ptr) r : valid (tset t r) && all (fun x => x != null) (preorder t). Proof. apply/idP/idP; last first. -- by case/andP=>V /tallP W; rewrite (domeq2VE (domeq_tlay_tset r r W)). +- by case/andP=>V /tallP W; rewrite (domeqVE (domeq_tlay_tset r r W)). move/[dup]=>V /valid_tlayE /[dup] N /tallP ->. -by rewrite -(domeq2VE (domeq_tlay_tset r r N)) V. +by rewrite -(domeqVE (domeq_tlay_tset r r N)) V. Qed. Lemma valid_tlayN (t : tree ptr) r : @@ -207,7 +220,7 @@ Lemma dom_tlay (t : tree ptr) r : valid (tlay t r) -> dom (tlay t r) =i t. Proof. rewrite valid_tlay=>/andP [V /tallP A] x. -by rewrite -(dom_tset V) (domeq2DE (domeq_tlay_tset r r A)). +by rewrite -(dom_tset V) (domeqDE (domeq_tlay_tset r r A)). Qed. Lemma dom_tlayE (t : tree ptr) r : @@ -242,8 +255,8 @@ case V : (valid (tset t r))=>/=; last first. rewrite -(dom_tset V); case: dom_find=>// v E _. elim/tree_ind1: t V E=>a ts IH; rewrite tsetE /= => V. rewrite !findPtUn2 //; case: (x =P a)=>//= _. -case/big_find_someX=>i X D; apply: big_find_some (validX V) (X) (IH _ X _ D). -by apply: big_validV (validR V) X. +case/big_find_someX=>i X /[dup] D /In_find/In_valid W. +by apply: IH X W (D). Qed. Lemma In_tsetP x r t: @@ -293,10 +306,10 @@ Proof. by move/tlay_rt; rewrite eqxx. Qed. Definition flay ts := foldr (fun t h => tlay t (rt t) \+ h) Unit ts. Definition fset ts := foldr (fun t h => tset t (rt t) \+ h) Unit ts. -Lemma flayE ts : flay ts = \big[PCM.join/Unit]_(t <- ts) (tlay t (rt t)). +Lemma flayE ts : flay ts = \big[join/Unit]_(t <- ts) (tlay t (rt t)). Proof. by elim: ts=>[|t ts IH] /=; [rewrite big_nil|rewrite big_cons IH]. Qed. -Lemma fsetE ts : fset ts = \big[PCM.join/Unit]_(t <- ts) (tset t (rt t)). +Lemma fsetE ts : fset ts = \big[join/Unit]_(t <- ts) (tset t (rt t)). Proof. by elim: ts=>[|t ts IH] /=; [rewrite big_nil|rewrite big_cons IH]. Qed. Lemma find_flayTp (x : ptr) (ts : seq (tree ptr)) (p : dynamic id) : @@ -322,7 +335,7 @@ Lemma dom_fset (ts : seq (tree ptr)) x : valid (fset ts) -> x \in dom (fset ts) = has (fun t => x \in t) ts. Proof. -elim: ts=>[|t ts IH] //= V; first by rewrite dom0. +elim: ts=>[|t ts IH] //= V. by rewrite domUn V inE /= IH ?dom_tset ?(validX V). Qed. @@ -339,7 +352,7 @@ Lemma range_fsetE ts : range (fset ts) =i [pred x | valid (fset ts) && has (fun t => x == rt t) ts]. Proof. -elim: ts=>[|t ts IH] /= x; first by rewrite range0 andbF. +elim: ts=>[|t ts IH] //= x. rewrite rangeUn inE; case V : (valid _)=>//=. by rewrite range_tset ?inE ?IH ?(validX V). Qed. @@ -379,7 +392,7 @@ Lemma valid_flay_fset (ts : seq (tree ptr)) : valid (flay ts) = valid (fset ts) && all (fun t => all (fun i => i != null) (preorder t)) ts. Proof. -elim: ts=>[|t ts IH] //=; first by rewrite valid_unit. +elim: ts=>[|t ts IH] //=. rewrite !validUnAE IH -!andbA valid_tlay /=. case V1 : (valid _)=>//=. case V2 : (valid _)=>/=; last by rewrite andbF. @@ -395,7 +408,7 @@ Lemma dom_flay_fset (ts : seq (tree ptr)) : all (fun t => all (fun i => i != null) (preorder t)) ts -> dom (flay ts) = dom (fset ts). Proof. -move=>A; apply/dom2E=>x; rewrite dom_flayE dom_fsetE !inE valid_flay_fset A /=. +move=>A; apply/domE=>x; rewrite dom_flayE dom_fsetE !inE valid_flay_fset A /=. by rewrite andbT. Qed. @@ -462,7 +475,7 @@ Lemma dom_flay_big (ts : seq (tree ptr)) : valid (flay ts) -> dom (flay ts) =i \big[cat/[::]]_(t <- ts) preorder t. Proof. -move=>V x; rewrite flayE big_domUn inE -flayE V big_cat_mem_seq /=. +move=>V x; rewrite flayE big_domUn inE -flayE V big_cat_mem_has /=. by apply: eq_in_has=>i H; rewrite dom_tlay ?in_preorder ?(valid_flay_tlay V H). Qed. @@ -528,17 +541,17 @@ Proof. by rewrite flayE big_cons big_nil unitR. Qed. Lemma flay_uniq_ts ts : valid (flay ts) -> uniq ts. Proof. -move=>/flay_uniq/uniq_big_catP [_ H _]; apply: count_mem_uniq=>t. +move=>/flay_uniq/uniq_big_catE [_ H _]; apply: count_mem_uniq=>t. case T : (t \in ts); last by apply/count_memPn; apply: negbT. by apply: (H t (rt t))=>//; rewrite in_preorder rt_in. Qed. -Lemma nochange_mapv (K : ordType) (V : eqType) (m : union_mapPCM K V) b x : +Lemma nochange_mapv (K : ordType) (V : eqType) (m : umap K V) b x : valid m -> x \notin range m -> mapv [fun v => v with x |-> b] m = m. Proof. -move=>W /negP R; apply: umem_eq=>[|//|[k v]]; first by rewrite valid_omf. +move=>W /negP R; apply: umem_eq=>[|//|[k v]]; first by rewrite pfV. rewrite In_omapX /=; split=>[[w]|H]. - by case: (w =P x)=>[->{w} /mem_range/R|_ /[swap][[]->]//]. by exists v=>//; case: (v =P x) H=>// -> /mem_range/R. @@ -571,7 +584,7 @@ Proof. by case=>ts [->->]; rewrite valid_flay_fset=>/andP []. Qed. (* Creates a new equivalence class with a single element *) Program Definition newT : - {m}, STsep (shape m, [vfun r => shape (r \\-> r \+ m)]) := + STsep {m} (shape m, [vfun r => shape (r \\-> r \+ m)]) := Do (p <-- alloc null; p ::= p;; ret p). @@ -588,13 +601,13 @@ Opaque newT. (* Returns the canonical representative of the equivalence class of an element*) Definition find_tp (x : ptr) := - {rs r}, STsep (fun h => shape rs h /\ (x, r) \In rs, + STsep {rs r} (fun h => shape rs h /\ (x, r) \In rs, [vfun res h => shape rs h /\ res = r]). Program Definition find1 (x : ptr) : find_tp x := - Do (let root := Fix (fun (go : forall x, find_tp x) (x : ptr) => + Do (let root := ffix (fun (go : forall x, find_tp x) (x : ptr) => Do (p <-- !x; if x == p then ret p else go p)) in root x). @@ -619,10 +632,10 @@ Opaque find1. (*********) (* Joins the equivalence classes of the two arguments *) -Definition union_tp (x y : ptr) := {rx ry m}, - STsep (fun h => [/\ shape m h, (x, rx) \In m & (y, ry) \In m], - [vfun res h => shape (mapv [fun v => v with rx |-> ry] m) h /\ - res = ry]). +Definition union_tp (x y : ptr) := STsep {rx ry m} + (fun h => [/\ shape m h, (x, rx) \In m & (y, ry) \In m], + [vfun res h => shape (mapv [fun v => v with rx |-> ry] m) h /\ + res = ry]). Program Definition union (x y : ptr) : union_tp x y := @@ -683,8 +696,7 @@ rewrite range_tset // mem_seq1; case: eqP =>// eqAB. rewrite tsetE -rtB joinA big_filter //; congr (_ \+ _). (*Case 2.3: tset of trees different than a and b don't change *) apply: nochange_mapv =>//; apply/negP; move/mem_rangeX; case=>k H. -move: (@bigInXP ptr_ordType _ _ (union_mapUMC ptr_ordType ptr) (tree ptr) (fun i0 => tset i0 (rt i0)))=>//=. -move/(_ _ _ _ _ H); case=>j [/mem_seqP X1 /andP [X2 X3]]. +case: (bigInXP H)=>j [/mem_seqP X1 /andP [X2 X3]]. move/mem_range; rewrite range_tset //; last by apply: valid_fset_tset X1. rewrite inE => /eqP X4; move/negP: X2; apply; apply/eqP. by apply: flay_mem_eq X1 B J=>//; rewrite X4 rt_in. @@ -719,8 +731,8 @@ by move/shapeV: X; rewrite invalidX. Qed. Program Definition test2 (x: ptr): - {y}, STsep (fun h => shape (x \\-> y \+ y \\-> y) h, - [vfun res h => shape (x \\-> y \+ y \\-> y) h /\ res = y]) := + STsep {y} (fun h => shape (x \\-> y \+ y \\-> y) h, + [vfun res h => shape (x \\-> y \+ y \\-> y) h /\ res = y]) := Do (res <-- find1 x; ret res). Next Obligation. diff --git a/pcm/heap.v b/pcm/heap.v index 3477902..a1280be 100644 --- a/pcm/heap.v +++ b/pcm/heap.v @@ -197,10 +197,18 @@ Lemma hcancelPtV2 A x1 x2 (v1 v2 : A) : 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 h x. + x \in dom h -> + exists A (v : A), + 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. +(* restatement of um_eta2, to avoid showing idyn's *) +Lemma heap_eta2 A x h (v : A) : + find x h = Some (idyn v) -> + h = x :-> v \+ free h x. +Proof. exact: um_eta2. Qed. + Lemma hcancelT A1 A2 x (v1 : A1) (v2 : A2) h1 h2 : valid (x :-> v1 \+ h1) -> x :-> v1 \+ h1 = x :-> v2 \+ h2 -> A1 = A2. @@ -222,6 +230,7 @@ Proof. by move=>V /(cancel2 V); case: ifP=>// _ [/dyn_inj]. Qed. End HeapPointsToLemmas. +Prenex Implicits heap_eta heap_eta2. (*******************************************) (* Block update for reasoning about arrays *) diff --git a/pcm/unionmap.v b/pcm/unionmap.v index 282c703..e02fcd0 100644 --- a/pcm/unionmap.v +++ b/pcm/unionmap.v @@ -97,7 +97,6 @@ From pcm Require Import options axioms prelude finmap seqperm pred seqext. From pcm Require Export ordtype. From pcm Require Import pcm morphism. - (****************************) (****************************) (* Reference Implementation *) @@ -2500,7 +2499,20 @@ 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 InPtUnEL. +Prenex Implicits InPt In_eta InPtUn In_dom InPtUnEL In_findE. + +(* Umap and fset are special cases of union_map but are *) +(* defined before membership structure is declared. *) +(* Their membership structures aren't directly inheritted from that *) +(* of union_map, but must be declared separately *) +Canonical umap_PredType (K : ordType) V : PredType (K * V) := + Mem_UmMap_PredType (umap K V). +Coercion Pred_of_umap K V (x : umap K V) : Pred_Class := + [eta Mem_UmMap x]. +Canonical fset_PredType (K : ordType) : PredType (K * unit) := + Mem_UmMap_PredType (fset K). +Coercion Pred_of_fset K (x : fset K) : Pred_Class := + [eta Mem_UmMap x]. Section MorphMembership. Variables (K : ordType) (C : pred K) (V : Type). @@ -2510,15 +2522,22 @@ Section PlainMorph. Variable f : pcm_morph U1 U2. Lemma InpfL e (x y : U1) : - valid (x \+ y) -> sep f x y -> e \In f x -> e \In f (x \+ y). + valid (x \+ y) -> + sep f x y -> + e \In f x -> + e \In f (x \+ y). Proof. by move=>W H1 H2; rewrite pfjoin //=; apply: InL=>//; rewrite pfV2. Qed. Lemma InpfR e (x y : U1) : - valid (x \+ y) -> sep f x y -> e \In f y -> e \In f (x \+ y). + valid (x \+ y) -> + sep f x y -> + e \In f y -> + e \In f (x \+ y). Proof. by move=>W H1 H2; rewrite pfjoin //=; apply: InR=>//; rewrite pfV2. Qed. Lemma InpfUn e (x y : U1) : - valid (x \+ y) -> sep f x y -> + valid (x \+ y) -> + sep f x y -> e \In f (x \+ y) <-> e \In f x \/ e \In f y. Proof. move=>W H; rewrite pfjoin //; split; first by case/InUn; eauto. @@ -2530,14 +2549,21 @@ End PlainMorph. Section FullMorph. Variable f : full_pcm_morph U1 U2. -Lemma InpfLT e (x y : U1) : valid (x \+ y) -> e \In f x -> e \In f (x \+ y). +Lemma InpfLT e (x y : U1) : + valid (x \+ y) -> + e \In f x -> + e \In f (x \+ y). Proof. by move=>W; apply: InpfL. Qed. -Lemma InpfRT e (x y : U1) : valid (x \+ y) -> e \In f y -> e \In f (x \+ y). +Lemma InpfRT e (x y : U1) : + valid (x \+ y) -> + e \In f y -> + e \In f (x \+ y). Proof. by move=>W; apply: InpfR. Qed. Lemma InpfUnT e (x y : U1) : - valid (x \+ y) -> e \In f (x \+ y) <-> e \In f x \/ e \In f y. + valid (x \+ y) -> + e \In f (x \+ y) <-> e \In f x \/ e \In f y. Proof. by move=>W; apply: InpfUn. Qed. End FullMorph. @@ -3267,7 +3293,8 @@ by case: (f _)=>// w; rewrite unitL. Qed. Lemma omapUn x1 x2 : - valid (x1 \+ x2) -> omap (x1 \+ x2) = omap x1 \+ omap x2. + valid (x1 \+ x2) -> + omap (x1 \+ x2) = omap x1 \+ omap x2. Proof. move=>W; rewrite /omap W (validL W) (validR W); set o := fun z kv => _. have H (x y : U') kv : o (x \+ y) kv = o x kv \+ y. @@ -3276,6 +3303,14 @@ rewrite (foldl_perm H _ (assocs_perm W)) foldl_cat. by rewrite joinC -(foldl_init H) unitL. Qed. +Lemma omapVUn x1 x2 : + omap (x1 \+ x2) = + if valid (x1 \+ x2) then omap x1 \+ omap x2 else undef. +Proof. +case: (normalP (x1 \+ x2))=>[->|W]; +by [rewrite omap_undef|apply: omapUn]. +Qed. + Lemma omapPtUn k v x : omap (pts k v \+ x) = if valid (pts k v \+ x) then From 23ce5abaf4085802e1db66eefa9dd7eca7f5a1a1 Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Mon, 5 Aug 2024 19:51:51 +0200 Subject: [PATCH 11/93] started refactoring graph.v --- examples/graph.v | 1034 ++++++++++++++++++++++++++++++++++++++++++++++ 1 file changed, 1034 insertions(+) create mode 100644 examples/graph.v diff --git a/examples/graph.v b/examples/graph.v new file mode 100644 index 0000000..afe16dc --- /dev/null +++ b/examples/graph.v @@ -0,0 +1,1034 @@ +(* +Copyright 2022 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 HB Require Import structures. +From Coq Require Import ssreflect ssrbool ssrfun. +From mathcomp Require Import eqtype ssrnat seq path. +From pcm Require Import options axioms pred prelude. +From pcm Require Import pcm unionmap natmap autopcm automap. +From pcm Require Import seqext. + +(* Pregraphs are natmaps, mapping nodes into *) +(* seq node listing the children of the node (adjacency list) *) +(* Technically pregraph's a "prequiver", because this representation *) +(* allows loops and parallel edges *) +(* Pregraph differs from graph, in that edges can *dangle*; that is *) +(* originate or terminate (or both) with a node that isn't in the graph *) + +Notation node := nat. +Record pregraph := Pregraph {pregraph_base : @UM.base node nat_pred (seq node)}. + +Section PregraphUMC. +Implicit Type f : pregraph. +Local Coercion pregraph_base : pregraph >-> UM.base. +Let pg_valid f := @UM.valid nat nat_pred (seq nat) f. +Let pg_empty := Pregraph (@UM.empty nat nat_pred (seq nat)). +Let pg_undef := Pregraph (@UM.Undef nat nat_pred (seq nat)). +Let pg_upd k v f := Pregraph (@UM.upd nat nat_pred (seq nat) k v f). +Let pg_dom f := @UM.dom nat nat_pred (seq nat) f. +Let pg_assocs f := @UM.assocs nat nat_pred (seq nat) f. +Let pg_free f k := Pregraph (@UM.free nat nat_pred (seq nat) f k). +Let pg_find k f := @UM.find nat nat_pred (seq nat) k f. +Let pg_union f1 f2 := Pregraph (@UM.union nat nat_pred (seq nat) f1 f2). +Let pg_empb f := @UM.empb nat nat_pred (seq nat) f. +Let pg_undefb f := @UM.undefb nat nat_pred (seq nat) f. +Let pg_from (f : pregraph) : UM.base _ _ := f. +Let pg_to (b : @UM.base nat nat_pred (seq nat)) : pregraph := Pregraph b. +Let pg_pts k v := Pregraph (@UM.pts nat nat_pred (seq nat) k v). + +Lemma pregraph_is_umc : + union_map_axiom pg_valid pg_empty pg_undef pg_upd pg_dom + pg_assocs pg_free pg_find pg_union pg_empb + pg_undefb pg_pts pg_from pg_to. +Proof. by split=>//; split=>[|[]]. Qed. + +HB.instance Definition _ := + isUnion_map.Build node nat_pred (seq node) pregraph pregraph_is_umc. +End PregraphUMC. + +HB.instance Definition _ := isNatMap.Build (seq node) pregraph. +HB.instance Definition _ := + hasDecEq.Build pregraph (@union_map_eqP node _ (seq node) pregraph). +Canonical pregraph_PredType : PredType (node * (seq node)) := + Mem_UmMap_PredType pregraph. +Coercion Pred_of_history (x : pregraph) : Pred_Class := + [eta Mem_UmMap x]. + +Notation "x &-> v" := (ptsT pregraph x v) (at level 30). + +(* p is out-node if no edge goes into it *) +Definition out_node (g : pregraph) (p : node) := + all (fun s => p \notin s) (range g). + +(* pregraph is graph if valid, and *) +(* all nodes in all adjacency lists are good *) +Definition graph_axiom (g : pregraph) := + valid g && all (all (fun p : node => (p == null) || (p \in dom g))) (range g). + +HB.mixin Record isGraph (g : pregraph) := { + graph_subproof : graph_axiom g}. +#[short(type=graph)] +HB.structure Definition Graph := {g of isGraph g }. + +(* removing out-node causes no dangling edges; *) +(* hence preserves graph axiom *) +Lemma graphF g p : + out_node g p -> + graph_axiom g -> + graph_axiom (free g p). +Proof. +move/allP=>/= Hna /andP [V /allP/= Ha]. +rewrite /graph_axiom validF V /=. +apply/allP=>/= s /rangeF Hs; apply/allP=>q Hq. +move/allP: (Ha _ Hs) (Hna _ Hs)=>/(_ _ Hq) {}Hs. +by rewrite domF inE; case: (p =P q) Hq=>//= ->->. +Qed. + +(* children of node p obtained by non-dangingle edges *) +Definition children (g : pregraph) p : seq node := + oapp (filter (mem (dom g))) [::] (find p g). + +Lemma children_undef p : children undef p = [::]. +Proof. by []. Qed. + +Lemma children_unit p : children Unit p = [::]. +Proof. by []. Qed. + +Lemma childrenND g x : + x \notin dom g -> + children g x = [::]. +Proof. by rewrite /children/oapp; case: dom_find. Qed. + +Lemma childrenD g x : + {subset children g x <= dom g}. +Proof. +move=>y; rewrite /children/oapp/=. +case D : (find x g)=>[a|//]. +by rewrite mem_filter; case/andP. +Qed. + +Lemma childrenUnL g1 g2 x : + valid (g1 \+ g2) -> + {subset children g1 x <= children (g1 \+ g2) x}. +Proof. +move=>V y; rewrite /children/oapp findUnL //. +case: dom_find=>// ys /In_find/In_dom /= D _. +rewrite !mem_filter /= domUn inE V /=. +by case/andP=>->->. +Qed. + +Lemma childrenUnR g1 g2 x : + valid (g1 \+ g2) -> + {subset children g2 x <= children (g1 \+ g2) x}. +Proof. by rewrite joinC; apply: childrenUnL. Qed. + +(* edge relation is just the applicative variant of children *) +(* thus, consider removing one of them *) + +Definition edge g : rel node := mem \o children g. +Arguments edge g x y : simpl never. + +Lemma edge_undef x y : edge undef x y = false. +Proof. by rewrite /edge/= children_undef. Qed. + +Lemma edge_unit x y : edge Unit x y = false. +Proof. by rewrite /edge/= children_unit. Qed. + +Lemma edge_children g x y : + edge g x y = (y \in children g x). +Proof. by []. Qed. + +Lemma edgeUnL g1 g2 x y : + valid (g1 \+ g2) -> + edge g1 x y -> + edge (g1 \+ g2) x y. +Proof. by move=>V; apply: childrenUnL. Qed. + +Lemma edgeUnR g1 g2 x y : + valid (g1 \+ g2) -> + edge g2 x y -> + edge (g1 \+ g2) x y. +Proof. by move=>V; apply: childrenUnR. Qed. + +Lemma edge_dom g x y : + edge g x y -> + (x \in dom g) * (y \in dom g). +Proof. +rewrite /edge/= => H; split; last by apply: childrenD H. +by apply: contraLR H=>/childrenND ->. +Qed. + +Lemma find_edge g x y ys : + find x g = Some ys -> + edge g x y = (y \in dom g) && (y \in ys). +Proof. +rewrite /edge/children/oapp/= => ->. +by rewrite mem_filter. +Qed. + +Lemma path_dom g x xs : + path (edge g) x xs -> + all (mem (dom g)) xs. +Proof. +elim: xs x=>[|a xs IH] x //= /andP [He Hp]. +by apply/andP; split; [case/edge_dom: He|apply: IH Hp]. +Qed. + +(* lifting the theory of finite graphs to unbounded pregraphs *) + +(* list of nodes traversed by depth-first search of g *) +(* at depth n, starting from x, and avoiding v *) +Fixpoint dfs (g : pregraph) (n : nat) (v : seq node) x := + if (x \notin dom g) || (x \in v) then v else + if n is n'.+1 then foldl (dfs g n') (x :: v) (children g x) else v. + +Lemma subset_dfs g n v x : + {subset v <= foldl (dfs g n) v x}. +Proof. +elim: n x v => [|n IHn] /=; elim=>[|x xs IHx] v //=. +- by case: ifP. +move=>y Hy; apply: IHx; case: ifP=>//= _. +by apply: IHn; rewrite inE Hy orbT. +Qed. + +(* accumulator is bound by g *) +Lemma subset_foldl_dfs_dom g n v x : + {subset v <= dom g} -> + {subset foldl (dfs g n) v x <= dom g}. +Proof. +elim: n x v=>[|n IHn]; elim=>[|x xs IHx] v //=. +- by case: ifP=>_; apply: IHx. +case: ifP; first by case: (x \in dom g)=>//= H; apply: IHx. +case Dx: (x \in dom g)=>//= H Gx; apply/IHx/IHn. +by move=>z; rewrite inE; case/orP=>[/eqP ->|/Gx]. +Qed. + +Lemma subset_dfs_dom g n v x : + {subset v <= dom g} -> + {subset dfs g n v x <= dom g}. +Proof. +case: n=>[|n] H /=; case: ifP=>//=. +case Dx : (x \in dom g)=>//= _; apply: subset_foldl_dfs_dom. +by move=>z; rewrite inE; case/orP=>[/eqP ->|/H]. +Qed. + +Lemma uniq_dfs_foldl g n v x : + uniq v -> + uniq (foldl (dfs g n) v x). +Proof. +elim: n x v=>[|n IHn]; elim=>[|x xs IHx] v U //=; apply: IHx. +- by rewrite if_same. +case: (x \in dom g)=>//=; case: ifP=>// Xv. +by rewrite IHn //= Xv. +Qed. + +Lemma uniq_dfs g n v x : + uniq v -> + uniq (dfs g n v x). +Proof. +case: n=>[|n] U /=; first by rewrite if_same. +case: (x \in dom g)=>//=; case: ifP=>// Xv. +by rewrite uniq_dfs_foldl //= Xv. +Qed. + +(* there's a path in g from x to y avoiding v *) +Inductive dfs_path g (v : seq node) x y : Prop := + DfsPath xs of + path (edge g) x xs & + y = last x xs & + (* v and x :: xs are disjoint sequences *) + all (fun z => z \notin v) (x :: xs) && + all (fun z => z \notin x :: xs) v. + +Lemma dfs_path_id g v x : + x \in dom g -> + x \notin v -> + dfs_path g v x x. +Proof. +move=>Gx Vx; apply: (DfsPath (xs:=[::]))=>//=. +rewrite andbT Vx /=; apply/allP=>z; rewrite inE. +by case: (z =P x) Vx=>//= -> /negbTE ->. +Qed. + +Lemma dfs_pathP g n x y v : + size (dom g) <= size v + n -> + uniq v -> + {subset v <= dom g} -> + y \notin v -> + x \in dom g -> + reflect (dfs_path g v x y) (y \in dfs g n v x). +Proof. +elim: n=>[|n IHn] /= in x y v * => Hv Uv Sv Ny Dx. +- rewrite addn0 in Hv. + case: (uniq_min_size Uv Sv Hv) Ny=>_ Ev /negbTE Ny. + suff: ~ dfs_path g v x y. + - by rewrite Dx /= if_same Ny; apply: ReflectF. + by case=>/= xs E Ey; rewrite Ev Dx. + +UP TO HERE + +case Epl: (links g x)=>[ls|] /=; last first. +- rewrite if_same (negbTE Ny); apply: ReflectF; case=>/= xs. + case: xs=>[|??]/=; last by rewrite (negbTE (links_edge_not _ Epl)). + by move=>_ Ey Hy; move/linksF/negbTE: Epl; rewrite Hy. +have [Vx|] := ifPn. +- rewrite (negbTE Ny); apply: ReflectF; case=>/= xs. + by rewrite Vx. +rewrite negb_or; case/andP=>Nx Px. +set v1 := x :: v; have [-> | Nyx] := eqVneq y x. +- rewrite subset_dfs; last by rewrite inE eq_refl. + apply: ReflectT; apply: dfs_path_id=>//. + by case/linksT: Epl. +apply: (@equivP (exists2 x1, x1 \in ls & dfs_path p g v1 x1 y)); last first. +- split=> {IHn} [[x1 Hx1 [p1 P1 Ey Py]] | [p1 /shortenP []]] /=. + - case/andP; rewrite !inE !negb_or -andbA; case/and3P=>Ex1 Ex1v Nx1 Ha1; apply: (DfsPath (xs:=x1::p1))=>//=. + - by rewrite P1 andbT (links_edge _ Epl). + rewrite !inE !negb_or Nx Px Ex1v Nx1 /=. + by apply/sub_all/Ha1=>z; rewrite !inE !negb_or -andbA; case/and3P=>_->->. + case=>[_ _ _ /= Eyx | x1 xs /=]; first by case/eqP: Nyx. + rewrite (links_edge _ Epl) /=; case/andP=>Hx1 Hp1 /and3P [H11 H12 H13] H2 H3 H4 /andP [H5 H6]. + exists x1=>//; apply: (DfsPath (xs:=xs))=>//. + - by apply/implyP=>_; case/linksT: Epl=>_ /allP /(_ _ Hx1). + apply/allP=>z /[dup] Hz /H2; move/allP: H6=>/(_ z) /[apply]. + rewrite !inE !negb_or; case/andP=>->->/=; rewrite !andbT; apply/negP=>/eqP Ez. + by rewrite -Ez Hz in H11. +have{Nyx Ny}: y \notin v1 by apply/norP. +have{Nx Uv}: uniq v1 by rewrite /= Nx. +have{Sv}: {subset v1 <= dom g}. +- move=>z; rewrite inE; case/orP; last by move/Sv. + by move/eqP=>{z}->; case/linksT: Epl. +have{Hv}: size (dom g) <= size v1 + n by rewrite addSnnS. +case/linksT: Epl=>_ Hl. +elim: {x Px v}ls Hl v1 => /= [_|x a IHa /andP [Ha Hl]] v /= Hv Sv Uv Nv. +- by rewrite (negPf Nv); apply: ReflectF; case. +set v2 := dfs p g n v x; have v2v: {subset v <= v2} := @subset_dfs p g n v [:: x]. +have [Hy2 | Ny2] := boolP (y \in v2). +- rewrite subset_dfs //; apply: ReflectT; exists x; first by rewrite inE eq_refl. + by apply/IHn. +apply: {IHa}(equivP (IHa Hl _ _ _ _ Ny2)). +- by apply: (leq_trans Hv); rewrite leq_add2r; apply: uniq_leq_size. +- by apply: subset_dfs_dom. +- by apply: uniq_dfs. +split=> [] [x1 Hx1 [p1 P1 Ey Py Nay]]. + exists x1; first by rewrite inE Hx1 orbT. + apply: (DfsPath (xs:=p1))=>//; apply/sub_all/Nay=>z; apply: contra. + by rewrite !inE; case/orP; [move/v2v=>->|move=>->; rewrite orbT]. +suff Nv2: all (fun z => z \notin v2) (x1 :: p1). +- move: Hx1; rewrite inE; case/orP=>[/eqP Ex1|Hx1]; last first. + - exists x1=>//; apply: (DfsPath (xs:=p1))=>//. + apply/allP=>z Hz; rewrite inE negb_or; apply/andP; split. + - by move/allP: Nv2; apply. + by move/allP: Nay=>/(_ _ Hz); rewrite inE negb_or; case/andP. + rewrite {x1}Ex1 in P1 Py Ey Nay Nv2. + exfalso; move: Nv2=>/=; case/andP=>+ _; move/negbTE/negP; apply. + suff [Nx Nxp]: (x \notin v) /\ (x \notin dom p). + - by apply/IHn=>//; apply: dfs_path_id. + by move: Nay=>/=; case/andP; rewrite inE negb_or; case/andP. +apply: contraR Ny2; case/allPn => x2 Hx2 /negbNE Hx2v. +case/splitPl: Hx2 Ey P1 Nay => p0 p2 p0x2. +rewrite last_cat cat_path -cat_cons lastI cat_rcons {}p0x2 => p2y /andP[_ g_p2]. +rewrite all_cat /=; case/and3P=> {p0}_; rewrite inE negb_or; case/andP=>Nx2v Nx2p Np2. +have{Nx2v Hx2v} [p3 g_p1 p1_x2 V2 not_p1v] := (IHn _ _ v Hv Uv Sv Nx2v Hx2v). +apply/IHn=>//; exists (p3 ++ p2)=>//. +- by rewrite cat_path g_p1 -p1_x2. +- by rewrite last_cat -p1_x2. +- by rewrite cat_nilp; apply/implyP; case/andP=>+ _; apply/implyP. +by rewrite -cat_cons all_cat not_p1v. +Qed. + + + + + + + +(*******************) + + + +Variant dfs_path (p : nodeset) g (v : seq ptr) x y : Prop := + DfsPath xs of + path (edge g) x xs & + y = last x xs & + nilp xs ==> (x \in dom g) & + all (fun z => z \notin [predU v & (dom p)]) (x::xs). + + + + +(* ps : nodes that stop the traversal *) +(* g : pregraph to traverse in dfs *) +(* n : traversal depth *) +(* acc : accumulator; currently observed sequence of nodes *) +(* x : starting node *) + +Fixpoint dfs (ps : fset node) (g : pregraph) (n : nat) (acc : seq node) x := + if [|| x \notin dom g, x \in acc | x \in dom ps] then acc else + if n is n'.+1 then foldl (dfs ps g n') (x :: acc) (children g x) else acc. + +(* +Fixpoint dfs (ps : fset node) (g : pregraph) (n : nat) (acc : seq node) x := + if (x \in acc) || (x \in dom ps) then acc + else + if x \in dom g then + if n is n'.+1 then foldl (dfs' ps g n') (x :: acc) (children g x) + else acc + else acc. + +Fixpoint dfs (p : fset node) (g : pregraph) n (v : seq node) x := + if x \in [predU v & (dom p)] then v else + if x \in dom g then + if n is n'.+1 then foldl (dfs p g n') (x :: v) (children g x) else v + else v. +*) + +(* accumulator grows *) +Lemma subset_dfs ps g n acc x : + {subset acc <= foldl (dfs ps g n) acc x}. +Proof. +elim: n x acc => [|n IHn] /=; elim=>[|x xs IHx] acc //=. +- by case: ifP. +move=>y Hy; apply: IHx; case: ifP=>//= _. +by apply: IHn; rewrite inE Hy orbT. +Qed. + +(* accumulator is bounded by g *) +Lemma subset_foldl_dfs_dom ps g n acc x : + {subset acc <= dom g} -> + {subset foldl (dfs ps g n) acc x <= dom g}. +Proof. +elim: n x acc=>[|n IHn]; elim=>[|x xs IHx] acc //=. +- by case: ifP=>_; apply: IHx. +case: ifP; first by case: (x \in dom g)=>//= H; apply: IHx. +case Dx: (x \in dom g)=>//= H Gx; apply/IHx/IHn. +by move=>z; rewrite inE; case/orP=>[/eqP ->|/Gx]. +Qed. + +Lemma subset_dfs_dom ps g n acc x : + {subset acc <= dom g} -> + {subset dfs ps g n acc x <= dom g}. +Proof. +case: n=>[|n] H /=; case: ifP=>//=. +case Dx : (x \in dom g)=>//= _; apply: subset_foldl_dfs_dom. +by move=>z; rewrite inE; case/orP=>[/eqP ->|/H]. +Qed. + +Inductive dfs_path g v x y : Prop := + DfsPath xs of path (edge g) x xs & y = last x xs & [disjoint x :: xs & v]. + + + +(*******************) + + + +Variant dfs_path (p : nodeset) g (v : seq ptr) x y : Prop := + DfsPath xs of + path (edge g) x xs & + y = last x xs & + nilp xs ==> (x \in dom g) & + all (fun z => z \notin [predU v & (dom p)]) (x::xs). + +Lemma dfs_path_id p g v x : + x \in dom g -> + x \notin dom p -> x \notin v -> dfs_path p g v x x. +Proof. +move=>Hx Nx Nv; apply (DfsPath (xs:=[::]))=>//=. +by rewrite andbT inE negb_or Nv. +Qed. + +Lemma path_dom g x xs : + path (edge g) x xs -> + all (fun z => z \in dom g) xs. +Proof. +elim: xs x=>//=a xs IH x; case/andP=>He Hp. +by apply/andP; split; [case/edge_dom: He | apply: (IH _ Hp)]. +Qed. + +Lemma dfs_pathP p g n x y v : + size (dom g) <= size v + n -> + uniq v -> + {subset v <= dom g} -> + y \notin v -> + reflect (dfs_path p g v x y) (y \in dfs p g n v x). +Proof. +elim: n => [|n IHn] /= in x y v * => Hv Uv Sv Ny. +- rewrite addn0 in Hv. + case: (uniq_min_size Uv Sv Hv) Ny=>_ /(_ y) Ey. + move/negbTE=>Ny; rewrite Ey in Ny. + suff: ~dfs_path p g v x y by case: (links g x)=>[_|]; rewrite if_same Ey Ny; apply: ReflectF. + case; case; first by move=>/= _ <-; rewrite Ny. + by move=>a l /path_dom/allP /(_ y) + /= Eyl; rewrite Ny Eyl mem_last; move/[apply]. +case Epl: (links g x)=>[ls|] /=; last first. +- rewrite if_same (negbTE Ny); apply: ReflectF; case=>/= xs. + case: xs=>[|??]/=; last by rewrite (negbTE (links_edge_not _ Epl)). + by move=>_ Ey Hy; move/linksF/negbTE: Epl; rewrite Hy. +have [Vx|] := ifPn. +- rewrite (negbTE Ny); apply: ReflectF; case=>/= xs. + by rewrite Vx. +rewrite negb_or; case/andP=>Nx Px. +set v1 := x :: v; have [-> | Nyx] := eqVneq y x. +- rewrite subset_dfs; last by rewrite inE eq_refl. + apply: ReflectT; apply: dfs_path_id=>//. + by case/linksT: Epl. +apply: (@equivP (exists2 x1, x1 \in ls & dfs_path p g v1 x1 y)); last first. +- split=> {IHn} [[x1 Hx1 [p1 P1 Ey Py]] | [p1 /shortenP []]] /=. + - case/andP; rewrite !inE !negb_or -andbA; case/and3P=>Ex1 Ex1v Nx1 Ha1; apply: (DfsPath (xs:=x1::p1))=>//=. + - by rewrite P1 andbT (links_edge _ Epl). + rewrite !inE !negb_or Nx Px Ex1v Nx1 /=. + by apply/sub_all/Ha1=>z; rewrite !inE !negb_or -andbA; case/and3P=>_->->. + case=>[_ _ _ /= Eyx | x1 xs /=]; first by case/eqP: Nyx. + rewrite (links_edge _ Epl) /=; case/andP=>Hx1 Hp1 /and3P [H11 H12 H13] H2 H3 H4 /andP [H5 H6]. + exists x1=>//; apply: (DfsPath (xs:=xs))=>//. + - by apply/implyP=>_; case/linksT: Epl=>_ /allP /(_ _ Hx1). + apply/allP=>z /[dup] Hz /H2; move/allP: H6=>/(_ z) /[apply]. + rewrite !inE !negb_or; case/andP=>->->/=; rewrite !andbT; apply/negP=>/eqP Ez. + by rewrite -Ez Hz in H11. +have{Nyx Ny}: y \notin v1 by apply/norP. +have{Nx Uv}: uniq v1 by rewrite /= Nx. +have{Sv}: {subset v1 <= dom g}. +- move=>z; rewrite inE; case/orP; last by move/Sv. + by move/eqP=>{z}->; case/linksT: Epl. +have{Hv}: size (dom g) <= size v1 + n by rewrite addSnnS. +case/linksT: Epl=>_ Hl. +elim: {x Px v}ls Hl v1 => /= [_|x a IHa /andP [Ha Hl]] v /= Hv Sv Uv Nv. +- by rewrite (negPf Nv); apply: ReflectF; case. +set v2 := dfs p g n v x; have v2v: {subset v <= v2} := @subset_dfs p g n v [:: x]. +have [Hy2 | Ny2] := boolP (y \in v2). +- rewrite subset_dfs //; apply: ReflectT; exists x; first by rewrite inE eq_refl. + by apply/IHn. +apply: {IHa}(equivP (IHa Hl _ _ _ _ Ny2)). +- by apply: (leq_trans Hv); rewrite leq_add2r; apply: uniq_leq_size. +- by apply: subset_dfs_dom. +- by apply: uniq_dfs. +split=> [] [x1 Hx1 [p1 P1 Ey Py Nay]]. + exists x1; first by rewrite inE Hx1 orbT. + apply: (DfsPath (xs:=p1))=>//; apply/sub_all/Nay=>z; apply: contra. + by rewrite !inE; case/orP; [move/v2v=>->|move=>->; rewrite orbT]. +suff Nv2: all (fun z => z \notin v2) (x1 :: p1). +- move: Hx1; rewrite inE; case/orP=>[/eqP Ex1|Hx1]; last first. + - exists x1=>//; apply: (DfsPath (xs:=p1))=>//. + apply/allP=>z Hz; rewrite inE negb_or; apply/andP; split. + - by move/allP: Nv2; apply. + by move/allP: Nay=>/(_ _ Hz); rewrite inE negb_or; case/andP. + rewrite {x1}Ex1 in P1 Py Ey Nay Nv2. + exfalso; move: Nv2=>/=; case/andP=>+ _; move/negbTE/negP; apply. + suff [Nx Nxp]: (x \notin v) /\ (x \notin dom p). + - by apply/IHn=>//; apply: dfs_path_id. + by move: Nay=>/=; case/andP; rewrite inE negb_or; case/andP. +apply: contraR Ny2; case/allPn => x2 Hx2 /negbNE Hx2v. +case/splitPl: Hx2 Ey P1 Nay => p0 p2 p0x2. +rewrite last_cat cat_path -cat_cons lastI cat_rcons {}p0x2 => p2y /andP[_ g_p2]. +rewrite all_cat /=; case/and3P=> {p0}_; rewrite inE negb_or; case/andP=>Nx2v Nx2p Np2. +have{Nx2v Hx2v} [p3 g_p1 p1_x2 V2 not_p1v] := (IHn _ _ v Hv Uv Sv Nx2v Hx2v). +apply/IHn=>//; exists (p3 ++ p2)=>//. +- by rewrite cat_path g_p1 -p1_x2. +- by rewrite last_cat -p1_x2. +- by rewrite cat_nilp; apply/implyP; case/andP=>+ _; apply/implyP. +by rewrite -cat_cons all_cat not_p1v. +Qed. + +Definition component p g x : seq ptr := dfs p g (size (dom g)) [::] x. + +Definition connect p g : rel ptr := [rel x y | y \in component p g x]. +Canonical connect_app_pred p g x := ApplicativePred (connect p g x). + +Corollary connectP p g x y : + reflect (exists xs, [/\ path (edge g) x xs, + y = last x xs, + nilp xs ==> (x \in dom g) & + all (fun z => z \notin (dom p)) (x::xs)]) + (y \in connect p g x). +Proof. +apply: (iffP (dfs_pathP _ _ _ _ _ _))=>//. +- by case=>xs P Ey Pv Ha; exists xs. +by case=>xs [P Ey Pv Ha]; apply: (DfsPath (xs:=xs)). +Qed. + +Corollary connectQ p g x y : + reflect [\/ [/\ x = y, x \in dom g & x \notin dom p] | + [/\ x \notin dom p, y \notin dom p & + exists xs, [/\ path (edge g) x xs, + edge g (last x xs) y & + forall z, z \in xs -> z \notin dom p]]] + (y \in connect p g x). +Proof. +case: connectP=>H; constructor. +- case: H; case/lastP=>[|zs z][/= H1 H2 H3 H4]; [left|right]. + - by rewrite andbT in H4; rewrite H2 H3 H4. + case/andP: H4=>H41 /allP H42; rewrite last_rcons in H2. + split=>//. + - by rewrite H2; apply: H42; rewrite mem_rcons inE eqxx. + rewrite rcons_path in H1; case/andP: H1=>H11 H12. + exists zs; split=>//=; first by rewrite H2. + by move=>w W; apply: H42; rewrite mem_rcons inE W orbT. +move=>X; apply: H; case: X. +- by case=>H1 H2 H3; exists [::]; split=>//=; rewrite H3. +case=>H1 H2 [xs][H3 H4 H5]; exists (rcons xs y). +rewrite rcons_path H3 H4 last_rcons /nilp size_rcons; split=>//=. +rewrite H1; apply/allP=>z; rewrite mem_rcons inE. +by case/orP=>[/eqP ->//|/H5]. +Qed. + +Lemma connect_trans p g : transitive (connect p g). +Proof. +move=> x y z /connectP [xs][Hxs -> Hp Ha] /connectP [ys][Hys -> Hp' Ha']; apply/connectP. +exists (xs ++ ys); rewrite cat_path last_cat Hxs /=; split=>//. +- by rewrite cat_nilp; apply/implyP; case/andP=>+ _; apply/implyP. +rewrite all_cat andbA; apply/andP; split=>//. +by move: Ha'=>/=; case/andP. +Qed. + +Lemma connect_refl p g x : + x \in dom g -> + x \in connect p g x = (x \notin dom p). +Proof. +move=>Gx; apply/connectQ/idP; first by case; case. +by move=>Px; left. +Qed. + +Lemma connect_mem p g x y : + x \in dom g -> + x \notin dom p -> + y \notin connect p g x -> + x != y. +Proof. by move=>Gx Px; apply: contra=>/eqP <-; rewrite connect_refl. Qed. + +Lemma connect_undef p x : connect p undef x =i pred0. +Proof. +move=>y; apply/connectP; case; case=>[|k xs][] /=; first by rewrite dom_undef. +by rewrite edge_undef. +Qed. + +Lemma connect0 p g x : x \in dom g -> x \notin dom p -> x \in connect p g x. +Proof. by move=>Hd Hp; apply/connectP; exists [::]; rewrite /= andbT. Qed. + +Lemma connectUn p g0 g x : + valid (g \+ g0) -> + {subset connect p g x <= connect p (g \+ g0) x}. +Proof. +move=>V y; case/connectP=>xs [Hp {y}-> /implyP Nxs Ha]. +apply/connectP; exists xs; split=>//; last first. +- rewrite domUn inE; rewrite V /=. + by apply/implyP=>/Nxs->. +elim: xs {Nxs Ha}x Hp=>//= y xs IH x; case/andP=>Hxy Hxs. +by apply/andP; split; [apply: edgeUn | apply: IH]. +Qed. + +Lemma connect_dom p g x y : + y \in connect p g x -> (x \in dom g) * (y \in dom g). +Proof. +case/connectP; case=>[|h t]; case; first by move=>/= _ ->. +move/[dup]/path_dom=>Ha /=; case/andP=>/edge_dom [Hx _] _ Ey _ _; split=>//. +by apply: (allP Ha); rewrite Ey; exact: mem_last. +Qed. + +Corollary connect_notd p g x : x \notin dom g -> connect p g x =i pred0. +Proof. +move=>Hx y; rewrite [RHS]inE; apply/negP; case/connect_dom=>E. +by rewrite E in Hx. +Qed. + +Lemma connect_notp p g x : x \in dom p -> connect p g x =i pred0. +Proof. +move=>Hx y; rewrite [RHS]inE; apply/negbTE/connectP. +by case=>xs /= [_ _ _]; case/andP; rewrite Hx. +Qed. + +Lemma connect_sub g p1 p2 x : + {subset dom p2 <= dom p1} -> + {subset connect p1 g x <= connect p2 g x}. +Proof. +move=>Hp z; case/connectP=>xs [Hxs {z}-> H Ha]. +apply/connectP; exists xs; split=>//. +by apply/sub_all/Ha=>z; apply/contra/Hp. +Qed. + +Lemma connect_links_sub p g x ns : + find x g = Some ns -> + forall y, y \in connect p g x -> (y == x) || has (fun n => y \in connect p g n) ns. +Proof. +move=>Hx y. +case/connectP=>xs [Hp {y}-> H]. +case: xs Hp H=>[|h xs]/=; first by move=>_ _; rewrite eq_refl. +case/andP; rewrite (find_edge _ Hx); case/andP=>Hh Hns Hxs _; case/and3P=>Hxp Hcp Ha. +apply/orP; right; apply/hasP; exists h=>//. +by apply/connectP; exists xs; rewrite Hh /= Hcp Ha; split=>//; apply: implybT. +Qed. + +Lemma connect_eq_links p g x ns : + find x g = Some ns -> + x \notin dom p -> + forall y, y \in connect p g x = (y == x) || has (fun n => y \in connect p g n) ns. +Proof. +move=>Hx Hd y; apply/iffE; split; first by apply: connect_links_sub. +case/orP. +- move/eqP=>->; apply: connect0=>//. + by move/find_some: Hx. +case/hasP=>n Hn. +case/connectP=>xs [Hp {y}-> H] /=; case/andP=>Hn1 Ha; apply/connectP. +exists (n::xs)=>/=; rewrite Hp Hd Hn1 Ha /= andbT; split=>//. +rewrite (find_edge _ Hx) Hn andbT. +by case: {Ha}xs Hp H=>//= ??; case/andP; case/edge_dom. +Qed. + +(* TODO generalize to arbitrary boundary? *) +Definition subconnect p g := um_filterk (fun z => z \in connect Unit g p) g. + +Definition subdisconnect p g := um_filterk (fun z => z \notin connect Unit g p) g. + +Lemma connect_split p g : g = subconnect p g \+ subdisconnect p g. +Proof. by apply: umfilt_predC. Qed. + +Lemma good_subconnect p g : + good_graph g -> good_graph (subconnect p g). +Proof. +have E := connect_split p g. +case/andP=>V Ha; apply/andP; split; first by rewrite valid_umfilt. +apply/allP=>ns Hns; move/allP: Ha=>/(_ ns). +rewrite {1}E rangeUn inE -E V Hns /= => /(_ erefl) Ha. +case/mem_rangeX: Hns=>k; case/In_umfilt=>/= Hk /In_find Hf. +apply/allP=>n Hn; move/allP: Ha=>/(_ _ Hn). +rewrite /good_ptr; case/boolP: (n == null)=>//= Hnn. +rewrite /subconnect dom_umfiltk inE /= =>En; rewrite En andbT. +apply: (connect_trans Hk); apply/connectP. +exists [::n]; split=>//=; last by rewrite !dom0. +by rewrite andbT (find_edge _ Hf) En. +Qed. + +Lemma edge_subconnect p g : subrel (edge (subconnect p g)) (edge g). +Proof. +move=>x y; rewrite /edge /links find_umfiltk; case: ifP=>//= Hx. +by case Ef: (find x g)=>[v|] //=; rewrite !mem_filter dom_umfiltk inE /= -andbA; case/andP. +Qed. + +Lemma connect_subconnect p g : connect Unit (subconnect p g) p =i connect Unit g p. +Proof. +move=>z; apply/iffE; split. +- case/connectP=>xs [Hxs {z}-> /implyP Nxs _]. + apply/connectP; exists xs; split=>//=; last by rewrite dom0 /= all_predT. + - by apply/sub_path/Hxs/edge_subconnect. + by apply/implyP=>/Nxs; rewrite dom_umfiltk inE /=; case/andP. +case/connectP=>xs [Hxs {z}-> /implyP Nxs _]. +have Hpd : p \in dom g. +- case: xs Hxs Nxs=>/=; first by move=>_; apply. + by move=>h t; case/andP; case/edge_dom. +apply/connectP; exists xs; split=>//=; last by rewrite dom0 /= all_predT. +- apply/(sub_in_path (P:=fun q => q \in connect Unit g p))/Hxs. + - move=>x y; rewrite !inE=>Hx Hy. + rewrite /edge /links find_umfiltk Hx. + case Ef: (find x g)=>[v|] //=; rewrite !mem_filter dom_umfiltk inE /= -andbA =>->. + by rewrite andbT. + move=>/=; apply/andP; split; first by apply/connect0=>//; rewrite dom0. + apply/allP=>x Hx; case/splitPr: Hx Hxs Nxs=>xs1 xs2. + rewrite -cat1s catA cats1 cat_path /=; case/andP=>H1 H2 _. + apply/connectP; exists (rcons xs1 x); split=>//=. + - by rewrite last_rcons. + - by rewrite -cats1 cat_nilp andbF. + by rewrite dom0 /= all_predT. +apply/implyP=>/Nxs; rewrite dom_umfiltk inE /==>->; rewrite andbT. +by apply/connect0=>//; rewrite dom0. +Qed. + +End GraphOps. + +Section Marking. + +Lemma connectMPtUn p m g x (cs : seq ptr) : + valid m -> + p \notin dom m -> + find p g = Some cs -> + forall z, z != p -> + x \in cs -> z \in connect m g x -> + exists2 y, y \in cs & z \in connect (#p \+ m) g y. +Proof. +move=>Vm Npm Hc z Hzp Hx; case/connectP=>xs [Hp Ez Nxs] Ha. +rewrite {z}Ez in Hzp *. +case Hpxs: (p \in x::xs); last first. +- (* there's no loop, path stays the same *) + exists x=>//; apply/connectP; exists xs; split=>//. + apply/allP=>z Hz; rewrite domPtUn inE validPtUn !negb_and negb_or /= negbK Vm (negbTE Npm) /=. + move/allP: Ha=>/(_ _ Hz) ->; rewrite andbT. + by apply/negP=>/eqP E; rewrite E Hz in Hpxs. +(* there's at least one loop, find the last p *) +case: {-1}_ _ _ / (splitLastP Hpxs) (erefl (x::xs)) =>{Hpxs} xs1 xs2 Hxs2. +case: xs2 Hxs2=>/=. +- move=>_; rewrite cats0=>E. + have /=: last x (x :: xs) = last x (rcons xs1 p) by rewrite E. + by rewrite last_rcons=>{}E; rewrite E eq_refl in Hzp. +move=>ch xs2; rewrite inE negb_or; case/andP => Npch Nxs2 E. +case: xs1 E=>/=. +- (* p links to itself *) + case=>Exp E2; rewrite {x Nxs}Exp {xs}E2 /= in Hx Hp Ha Hzp *. + case/andP: Hp=>He Hp. + exists ch; first by move: He; rewrite (find_edge _ Hc); case/andP. + apply/connectP; exists xs2; split=>//. + - by case/edge_dom: He=>_ ->; exact: implybT. + case/and3P: Ha=>_ Hch Ha. + apply/allP=>z; rewrite domPtUn inE validPtUn !negb_and negb_or /= negbK Vm (negbTE Npm) /=. + case/orP=>[/eqP {z}->| Hz]; first by rewrite Npch. + move/allP: Ha=>/(_ _ Hz) ->; rewrite andbT. + by apply/negP=>/eqP E; rewrite E Hz in Nxs2. +(* there's a non-trivial path before the loop *) +move=>_ xs1 [_ Exs]; rewrite {xs}Exs in Hp Nxs Ha Hzp *. +move: Hp; rewrite cat_path last_cat !last_rcons rcons_path -andbA /=. +case/and4P=>_ _ He Hp2. +exists ch; first by move: He; rewrite (find_edge _ Hc); case/andP. +apply/connectP; exists xs2; split=>//. +- by case/edge_dom: He=>_ ->; exact: implybT. +move: Ha; rewrite -cat_cons all_cat; case/andP=>_ Ha. +apply/allP=>z Hz; rewrite domPtUn inE validPtUn !negb_and negb_or /= negbK Vm (negbTE Npm) /=. +move/allP: Ha=>/(_ _ Hz) ->; rewrite andbT. +move: Hz; rewrite inE; case/orP=>[/eqP->|Hz] //. +by apply/negP=>/eqP E; rewrite E Hz in Nxs2. +Qed. + +Lemma connectMUnSub m1 m2 g x : + valid (m1 \+ m2) -> + forall z, z \in connect (m1 \+ m2) g x -> + z \in connect m2 g x. +Proof. +move=>Vm z; case/connectP=>xs [Hp Ez Nxs] Ha. +apply/connectP; exists xs; split=>//. +apply/sub_all/Ha=>q; apply: contra. +rewrite domUn inE Vm => ->. +by rewrite orbT. +Qed. + +Corollary connectMPtUnHas p m g cs : + valid m -> + p \notin dom m -> + find p g = Some cs -> + forall z, z \in connect m g p = (z == p) || has (fun x => z \in connect (#p \+ m) g x) cs. +Proof. +move=>Vm Npm Hc z. +rewrite (connect_eq_links Hc) //; case: eqP=>/= // /eqP Hzp. +have Vpm : valid (#p \+ m) by rewrite validPtUn Vm. +apply/iffE; split; case/hasP=>x Hx. +- move=>Hz; apply/hasP. + by apply: (connectMPtUn (x:=x)). +move/(connectMUnSub Vpm)=>Hz. +by apply/hasP; exists x. +Qed. + +Lemma connectMUn p m1 m2 g x c (cs : seq ptr) : + valid (m2 \+ m1) -> + find p g = Some cs -> + c \in cs -> good_ptr g c -> + dom m2 =i connect m1 g c -> + forall z, x \in cs -> z \in connect m1 g x -> + z \in connect m1 g c + \/ + exists2 y, y \in filter (predC1 c) cs & z \in connect (m2 \+ m1) g y. +Proof. +move=>V21 Hf Hc Hcd Hm z Hxc. case/connectP=>xs [Hp Ez Nxs] Ha. +case/boolP: (all (fun z => z \notin dom m2) (x::xs))=>Hpxs. +- (* path doesn't go through the marked component *) + right; exists x. + - rewrite mem_filter Hxc andbT /=. + case/orP: Hcd=>[/eqP->|Hcd]. + - suff: x \in dom g by move/dom_cond. + by case: {Ez Ha Hpxs}xs Hp Nxs=>//= h t; case/andP; case/edge_dom. + apply/negP=>/eqP Exc; move: Hpxs=>/=; case/andP=>+ _; rewrite Exc Hm. + move/negP; apply; apply: connect0=>//. + by move: Ha=>/=; case/andP=>+ _; rewrite Exc. + apply/connectP; exists xs; split=>//. + apply/allP=>q Hq; rewrite domUn inE negb_and negb_or V21 /=. + by move/allP: Ha=>/(_ _ Hq)->; move/allP: Hpxs=>/(_ _ Hq)->. +(* path goes through the marked component, so z is connected to c *) +left; rewrite -has_predC (eq_has (a2:=fun z=> z \in dom m2)) in Hpxs; last first. +- by move=>q /=; rewrite negbK. +(* q is the last vertex in marked component, xs2 is the free path *) +case: {-1}_ _ _ / (split_findlast Hpxs) (erefl (x::xs))=>{Hpxs} q xs1 xs2 Hq. +rewrite -all_predC (eq_all (a2:=fun z=> z \notin dom m2)) // => Hxs2 Heq. +apply: (connect_trans (y:=q)); rewrite app_predE; first by rewrite -Hm. +case: xs1 Heq=>/=. +- case=>Eq Exs2; rewrite -{q}Eq in Hq *; rewrite -{xs2}Exs2 in Hxs2. + by apply/connectP; exists xs. +move=>_ t [_ Exs]. +apply/connectP; exists xs2; split=>//. +- by move: Hp; rewrite Exs cat_path last_rcons; case/andP. +- by move: Ez; rewrite Exs last_cat last_rcons. +- suff: (edge g (last x t) q) by case/edge_dom=>_ ->; apply: implybT. + rewrite Exs cat_path rcons_path in Hp. + by case/andP: Hp; case/andP. +rewrite Exs -cats1 -catA cat1s -cat_cons all_cat in Ha. +by case/andP: Ha. +Qed. + +Corollary connectMUnHas p m1 m2 g c (cs : seq ptr) : + valid (m2 \+ m1) -> + find p g = Some cs -> + c \in cs -> good_ptr g c -> + dom m2 =i connect m1 g c -> + forall z, + has (fun x => z \in connect m1 g x) cs = + (z \in connect m1 g c) || + has (fun x => z \in connect (m2 \+ m1) g x) (filter (predC1 c) cs). +Proof. +move=>V Hf Hc Hcd Hm z; apply/iffE; split. +- case/hasP=>x Hx Hz. + case: (connectMUn V Hf Hc Hcd Hm Hx Hz); first by move=>->. + by case=>y Hy1 Hy2; apply/orP; right; apply/hasP; exists y. +case/orP; first by move=>Hz; apply/hasP; exists c. +case/hasP=>q; rewrite mem_filter /=; case/andP=>_ Hq /(connectMUnSub V) Hz. +by apply/hasP; exists q. +Qed. + +End Marking. + +Section NGraph. + +Definition n_graph (n : nat) (g : pregraph) : bool := + all (fun ns => size ns == n) (range g). + +Lemma n_graphUnL n g0 g : + valid (g \+ g0) -> + n_graph n (g \+ g0) -> n_graph n g. +Proof. +by move=>V; apply/subset_all=>z; rewrite rangeUn inE V=>->. +Qed. + +(* corollary? *) +Lemma n_graphF n g p : n_graph n g -> n_graph n (free g p). +Proof. by apply/subset_all/rangeF. Qed. + +Definition get_nth (ps : seq ptr) (n : nat) : ptr := + nth null ps n. + +Lemma all_good_get p g ps m : + good_graph g -> + find p g = Some ps -> + good_ptr g (get_nth ps m). +Proof. +case/andP=>_ Hg Hf. +case: (ltnP m (size ps))=>Hm; last first. +- by apply/orP; left; rewrite /get_nth; rewrite nth_default. +have /allP : all (good_ptr g) ps. +- move/allP: Hg; apply. + apply/mem_rangeX; exists p. + by move/In_find: Hf. +by apply; apply: mem_nth. +Qed. + +Lemma all_nth n g : + n_graph n g -> + all (fun ns => ns == map (get_nth ns) (iota 0 n)) (range g). +Proof. +move=>H; apply/sub_all/H=>ns /eqP Hns. +apply/eqP/(eq_from_nth (x0:=null)). +- by rewrite size_map size_iota. +by rewrite Hns=>i Hi; rewrite map_nth_iota0 -Hns // take_size. +Qed. + +End NGraph. + + +Section Acyclic. + +Definition symconnect p g x y := connect p g x y && connect p g y x. + +Lemma symconnect0 p g x : x \in dom g -> x \notin dom p -> symconnect p g x x. +Proof. by move=>Hx Hp; apply/andP; split; apply/connect0. Qed. + +Lemma symconnectUn p g0 g x y : + valid (g \+ g0) -> + symconnect p g x y -> symconnect p (g \+ g0) x y. +Proof. by move=>V; case/andP=>Hxy Hyx; apply/andP; split; apply: connectUn. Qed. + +(* TODO should probably generalize all of this to arbitrary boundary, not Unit *) +Lemma connect_cycle g p : cycle (edge g) p -> {in p &, forall x y, connect Unit g x y}. +Proof. +move=>Hp x y /rot_to[i q Hr]; rewrite -(mem_rot i) Hr => Hy. +have /= Hp1: cycle (edge g) (x :: q) by rewrite -Hr rot_cycle. +have Hd: x \in dom g by move: Hp1; rewrite rcons_path; case/andP=>_ /edge_dom; case. +move: Hp1; case/splitPl: Hy =>r s Hl; rewrite rcons_cat cat_path => /andP[Hxr Hlx]. +apply/connectP; exists r; split=>//; first by rewrite Hd implybT. +by rewrite dom0 all_predT. +Qed. + +Lemma symconnect_cycle g p : cycle (edge g) p -> + {in p &, forall x y, symconnect Unit g x y}. +Proof. by move=>Hp x y Hx Hy; rewrite /symconnect !(connect_cycle Hp). Qed. + +Lemma symconnect_cycle2P g x y : x != y -> + reflect (exists2 p, y \in p & cycle (edge g) (x :: p)) (symconnect Unit g x y). +Proof. +move=> Nxy; apply: (iffP idP) => [|[p yp]]; last first. + by move=> /symconnect_cycle->//; rewrite inE ?eqxx ?yp ?orbT. +move: Nxy =>/[swap]/andP [/connectP[p][Hp {y}-> Np _] /connectP[]]. +elim/last_ind => /= [[] _ <-|q z _]; first by rewrite eqxx. +case; rewrite rcons_path last_rcons => /[swap]{z}<- /andP[gq gzq] _ _ Nxy. +have := mem_last x p; rewrite in_cons eq_sym (negPf Nxy)/= => yp. +exists (p ++ q); first by rewrite mem_cat yp. +by rewrite rcons_path cat_path Hp gq last_cat gzq. +Qed. + +Definition preacyclic g := all2rel (fun x y => symconnect Unit g x y ==> (x == y)) (dom g). + +Lemma preacyclicUn g0 g : + valid (g \+ g0) -> + preacyclic (g \+ g0) -> preacyclic g. +Proof. +move=>V /allrelP H; apply/allrelP=>x y Hx Hy; apply/implyP=>Hs. +have Hx1: x \in dom (g \+ g0) by rewrite domUn inE V Hx. +have Hy1: y \in dom (g \+ g0) by rewrite domUn inE V Hy. +move/implyP: (H _ _ Hx1 Hy1); apply. +by apply: symconnectUn. +Qed. + +Definition acyclic g := all (fun x => ~~ edge g x x) (dom g) && preacyclic g. + +Lemma acyclicUn g0 g : + valid (g \+ g0) -> + acyclic (g \+ g0) -> acyclic g. +Proof. +move=>V; case/andP=>Ha Hp; apply/andP; split; last by apply: (preacyclicUn V Hp). +apply/allP=>x Hx. +suff: ~~ edge (g \+ g0) x x by apply/contra/edgeUn. +by move/allP: Ha; apply; rewrite domUn inE V Hx. +Qed. + +(* TODO all is overkill here, we only need the head of the path in dom g (?) *) +Lemma acyclic_cycleP g : + reflect (forall p, ~~ nilp p -> sorted (edge g) p -> all (fun x => x \in dom g) p -> ~~ cycle (edge g) p) + (acyclic g). +Proof. +rewrite /acyclic; apply: (iffP andP)=>[|Hn]; last first. +- split; first by apply/allP=>x Hx; move: (Hn [::x])=>/=; rewrite !andbT; apply. + apply/allrelP=>x y Hx _; apply/implyP/contraLR => neqxy. + apply/symconnect_cycle2P => // -[p Hp] /[dup]. + rewrite [X in X -> _]/= rcons_path => /andP[/[dup] /path_dom Hd Hg Ha]. + by apply/negP/Hn=>//=; rewrite Hx. +case=>/allP Ne /allrelP Ng. +case=>//= x p _; rewrite rcons_path =>/[dup] E ->/=; case/andP=>Hx. +case: p E=>/=; first by move=>_ _; apply: Ne. +move=>y p; case/andP=>He Hp; case/andP=>Hy Ha. +have: ~~ symconnect Unit g x y. +- apply/negP=>Hs; move: (Ng _ _ Hx Hy); rewrite Hs /= =>/eqP Exy. + by rewrite Exy in He; move: (Ne _ Hy); rewrite He. +apply: contra=>He1; apply: (symconnect_cycle (p:=x::y::p))=>/=; try by rewrite ?(in_cons, eqxx, orbT). +by rewrite rcons_path He Hp. +Qed. + +Lemma acyclic_find g n ns : + acyclic g -> + find n g = Some ns -> + n \notin ns. +Proof. +case/andP=>Ha _ /[dup]Hf /find_some Hn. +move/allP: Ha=>/(_ _ Hn). +by rewrite (find_edge _ Hf) negb_and Hn. +Qed. + +End Acyclic. From 2632e78a4f264bbccdf587c4686c5cfe4891bbe7 Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Mon, 5 Aug 2024 20:04:25 +0200 Subject: [PATCH 12/93] minor --- examples/graph.v | 13 ++++++++++--- 1 file changed, 10 insertions(+), 3 deletions(-) diff --git a/examples/graph.v b/examples/graph.v index afe16dc..c05851c 100644 --- a/examples/graph.v +++ b/examples/graph.v @@ -274,12 +274,19 @@ elim: n=>[|n IHn] /= in x y v * => Hv Uv Sv Ny Dx. suff: ~ dfs_path g v x y. - by rewrite Dx /= if_same Ny; apply: ReflectF. by case=>/= xs E Ey; rewrite Ev Dx. +rewrite Dx /=; have [Vx|Vx] := ifPn. +- rewrite (negbTE Ny); apply: ReflectF. + by case=>xs /=; rewrite Vx. +set v1 := x :: v; set a := children g x; have [->|Nyx] := eqVneq y x. +- rewrite subset_dfs ?inE ?eqxx //. + by apply/ReflectT/dfs_path_id. +apply: (@equivP (exists2 x1, x1 \in a & dfs_path g v1 x1 y))=>/=; last first. + UP TO HERE -case Epl: (links g x)=>[ls|] /=; last first. -- rewrite if_same (negbTE Ny); apply: ReflectF; case=>/= xs. - case: xs=>[|??]/=; last by rewrite (negbTE (links_edge_not _ Epl)). + + by move=>_ Ey Hy; move/linksF/negbTE: Epl; rewrite Hy. have [Vx|] := ifPn. - rewrite (negbTE Ny); apply: ReflectF; case=>/= xs. From 7c38ffa0aadbb21b4a5d4003a8a0e111ab8e6f29 Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Mon, 5 Aug 2024 22:34:36 +0200 Subject: [PATCH 13/93] wip --- examples/graph.v | 162 +++++++++++++++++++++++++++-------------------- 1 file changed, 93 insertions(+), 69 deletions(-) diff --git a/examples/graph.v b/examples/graph.v index c05851c..5cd9cc3 100644 --- a/examples/graph.v +++ b/examples/graph.v @@ -275,84 +275,108 @@ elim: n=>[|n IHn] /= in x y v * => Hv Uv Sv Ny Dx. - by rewrite Dx /= if_same Ny; apply: ReflectF. by case=>/= xs E Ey; rewrite Ev Dx. rewrite Dx /=; have [Vx|Vx] := ifPn. -- rewrite (negbTE Ny); apply: ReflectF. - by case=>xs /=; rewrite Vx. +- by rewrite (negbTE Ny); apply: ReflectF; case=>xs /=; rewrite Vx. set v1 := x :: v; set a := children g x; have [->|Nyx] := eqVneq y x. -- rewrite subset_dfs ?inE ?eqxx //. - by apply/ReflectT/dfs_path_id. +- by rewrite subset_dfs ?inE ?eqxx //; apply/ReflectT/dfs_path_id. apply: (@equivP (exists2 x1, x1 \in a & dfs_path g v1 x1 y))=>/=; last first. - - -UP TO HERE - - - - by move=>_ Ey Hy; move/linksF/negbTE: Epl; rewrite Hy. -have [Vx|] := ifPn. -- rewrite (negbTE Ny); apply: ReflectF; case=>/= xs. - by rewrite Vx. -rewrite negb_or; case/andP=>Nx Px. -set v1 := x :: v; have [-> | Nyx] := eqVneq y x. -- rewrite subset_dfs; last by rewrite inE eq_refl. - apply: ReflectT; apply: dfs_path_id=>//. - by case/linksT: Epl. -apply: (@equivP (exists2 x1, x1 \in ls & dfs_path p g v1 x1 y)); last first. -- split=> {IHn} [[x1 Hx1 [p1 P1 Ey Py]] | [p1 /shortenP []]] /=. - - case/andP; rewrite !inE !negb_or -andbA; case/and3P=>Ex1 Ex1v Nx1 Ha1; apply: (DfsPath (xs:=x1::p1))=>//=. - - by rewrite P1 andbT (links_edge _ Epl). - rewrite !inE !negb_or Nx Px Ex1v Nx1 /=. - by apply/sub_all/Ha1=>z; rewrite !inE !negb_or -andbA; case/and3P=>_->->. - case=>[_ _ _ /= Eyx | x1 xs /=]; first by case/eqP: Nyx. - rewrite (links_edge _ Epl) /=; case/andP=>Hx1 Hp1 /and3P [H11 H12 H13] H2 H3 H4 /andP [H5 H6]. - exists x1=>//; apply: (DfsPath (xs:=xs))=>//. - - by apply/implyP=>_; case/linksT: Epl=>_ /allP /(_ _ Hx1). - apply/allP=>z /[dup] Hz /H2; move/allP: H6=>/(_ z) /[apply]. - rewrite !inE !negb_or; case/andP=>->->/=; rewrite !andbT; apply/negP=>/eqP Ez. - by rewrite -Ez Hz in H11. -have{Nyx Ny}: y \notin v1 by apply/norP. -have{Nx Uv}: uniq v1 by rewrite /= Nx. -have{Sv}: {subset v1 <= dom g}. -- move=>z; rewrite inE; case/orP; last by move/Sv. - by move/eqP=>{z}->; case/linksT: Epl. -have{Hv}: size (dom g) <= size v1 + n by rewrite addSnnS. -case/linksT: Epl=>_ Hl. -elim: {x Px v}ls Hl v1 => /= [_|x a IHa /andP [Ha Hl]] v /= Hv Sv Uv Nv. -- by rewrite (negPf Nv); apply: ReflectF; case. -set v2 := dfs p g n v x; have v2v: {subset v <= v2} := @subset_dfs p g n v [:: x]. +- split=> {IHn} [[x1 Hx1 [p1 P1 E1]] | [p /shortenP[]]] /=. + - case/and3P=>/andP [H1 H2 H3 H4]; apply: (DfsPath (xs:=x1::p1))=>//=. + - by rewrite edge_children -/a Hx1. + rewrite Vx /= -!andbA; apply/and3P; split. + - by apply: contra H1; rewrite /v1 inE orbC=>->. + by apply: sub_all H2=>z; apply: contra; rewrite /v1 inE orbC=>->. + apply/allP=>z /= V; move/allP/(_ z V): H4=>/= H4. + rewrite inE (negbTE H4) orbF. + by case: (z =P x) V=>// ->; rewrite (negbTE Vx). + case=>[_ _ _ /= Eyx|x1 xs /=]; first by case/eqP: Nyx. + case/andP=>Hx1 Hp1 /and3P [H1 H2 H3] H4 H5 /andP [/andP [_ /allP H6] H7]. + exists x1=>//; apply: (DfsPath (xs:=xs))=>//=. + rewrite inE negb_or in H1; case/andP: H1=>H11 H12. + rewrite -!andbA; apply/and4P; split. + - by rewrite /v1 inE negb_or eq_sym H11; apply/H6/H4; rewrite inE eqxx. + - apply/allP=>/= z Z; rewrite /v1 inE negb_or. + have /H6 -> : z \in p by apply/H4; rewrite inE Z orbT. + by case: (z =P x) Z H12=>// ->->. + - by rewrite inE negb_or H11 H12. + by apply: sub_all H7=>z; apply: contra=>/H4; rewrite inE orbC=>->. +move: (Dx). +have {Nyx Ny}: y \notin v1 by apply/norP. +have {Sv Dx}: {subset v1 <= dom g}. +- by move=>z; rewrite inE; case/orP=>[/eqP ->//|/Sv]. +have {Vx Uv}: uniq v1 by rewrite /= Vx. +have {Hv}: size (dom g) <= size v1 + n by rewrite addSnnS. +have : {subset a <= dom g} by apply: childrenD. +elim: {x v}a (x) v1=>[|x a IHa] x' v /= Dxa leq_v'_n U Sv Nv Dx'. +- by rewrite (negbTE Nv); apply: ReflectF; case. +have Dx : x \in dom g by apply: Dxa; rewrite inE eqxx. +have Da : {subset a <= dom g} by move=>z Z; apply: Dxa; rewrite inE Z orbT. +set v2 := dfs g n v x. +have v2v : {subset v <= v2} := @subset_dfs g n v [:: x]. have [Hy2 | Ny2] := boolP (y \in v2). -- rewrite subset_dfs //; apply: ReflectT; exists x; first by rewrite inE eq_refl. - by apply/IHn. -apply: {IHa}(equivP (IHa Hl _ _ _ _ Ny2)). -- by apply: (leq_trans Hv); rewrite leq_add2r; apply: uniq_leq_size. -- by apply: subset_dfs_dom. +- rewrite subset_dfs //; apply: ReflectT. + by exists x; [rewrite inE eq_refl|apply/IHn]. +apply: {IHa} (equivP (IHa _ _ _ _ _ _ Ny2 Dx))=>//. +- by apply/(leq_trans leq_v'_n); rewrite leq_add2r; apply: uniq_leq_size. - by apply: uniq_dfs. -split=> [] [x1 Hx1 [p1 P1 Ey Py Nay]]. +- by apply: subset_dfs_dom. +split=>[][x1 Hx1 [p1 P1 Ey]]. +- rewrite -!andbA=>/and3P [H1 H2 H3]. exists x1; first by rewrite inE Hx1 orbT. - apply: (DfsPath (xs:=p1))=>//; apply/sub_all/Nay=>z; apply: contra. - by rewrite !inE; case/orP; [move/v2v=>->|move=>->; rewrite orbT]. -suff Nv2: all (fun z => z \notin v2) (x1 :: p1). -- move: Hx1; rewrite inE; case/orP=>[/eqP Ex1|Hx1]; last first. - - exists x1=>//; apply: (DfsPath (xs:=p1))=>//. - apply/allP=>z Hz; rewrite inE negb_or; apply/andP; split. - - by move/allP: Nv2; apply. - by move/allP: Nay=>/(_ _ Hz); rewrite inE negb_or; case/andP. - rewrite {x1}Ex1 in P1 Py Ey Nay Nv2. - exfalso; move: Nv2=>/=; case/andP=>+ _; move/negbTE/negP; apply. - suff [Nx Nxp]: (x \notin v) /\ (x \notin dom p). - - by apply/IHn=>//; apply: dfs_path_id. - by move: Nay=>/=; case/andP; rewrite inE negb_or; case/andP. -apply: contraR Ny2; case/allPn => x2 Hx2 /negbNE Hx2v. -case/splitPl: Hx2 Ey P1 Nay => p0 p2 p0x2. -rewrite last_cat cat_path -cat_cons lastI cat_rcons {}p0x2 => p2y /andP[_ g_p2]. -rewrite all_cat /=; case/and3P=> {p0}_; rewrite inE negb_or; case/andP=>Nx2v Nx2p Np2. -have{Nx2v Hx2v} [p3 g_p1 p1_x2 V2 not_p1v] := (IHn _ _ v Hv Uv Sv Nx2v Hx2v). -apply/IHn=>//; exists (p3 ++ p2)=>//. + apply: (DfsPath (xs:=p1))=>//=. + rewrite -!andbA; apply/and3P; split. + - by apply: contra H1=>/v2v. + - by apply: sub_all H2=>z; apply: contra=>/v2v. + by apply/allP=>/= z /v2v; apply: (allP H3). +rewrite /= -!andbA; case/and3P=>H1 H2 H3. +suff [N1 N2] : all (fun z : node => z \notin v2) (x1 :: p1) /\ + all (fun z : node => z \notin x1 :: p1) v2. +- move: Hx1; rewrite inE; case/orP=>[/eqP Ex1|Hx1]. + - subst x1; exfalso; move: N1=>/= /andP [+ _]. + by move/negbTE/negP; apply; apply/IHn=>//; apply: dfs_path_id. + exists x1=>//; apply: (DfsPath (xs:=p1))=>//=. + rewrite -!andbA; apply/and3P; split=>//; last by case/andP: N1. + by rewrite (allP N1) // inE eqxx. +split. +- apply: contraR Ny2; case/allPn => x2 Hx2 /negbNE Hx2v. + have {H1 H2} H12 : all (fun z : node => z \notin v) (x1 :: p1). + - by simpl; rewrite H1 H2. + case/splitPl: Hx2 Ey P1 H12 H3=>p0 p2 p0x2. + rewrite last_cat cat_path -cat_cons lastI cat_rcons {}p0x2 => p2y /andP [_ g_p2]. + rewrite all_cat /=. case/and3P=>H2 H3 H4 H5. + have [p3 g_p1 p1_x2 /= /andP [/andP [V1 V2] V3]] := (IHn _ _ v leq_v'_n U Sv H3 Dx Hx2v). + apply/IHn=>//; exists (p3 ++ p2). + - by rewrite cat_path g_p1 -p1_x2. + - by rewrite last_cat -p1_x2. + simpl. rewrite V1 /=. + rewrite all_cat V2 H4 /=. + rewrite -cat_cons. apply/allP=>z V. + rewrite mem_cat negb_or. apply/andP. split. + - by apply: (allP V3). + move/allP: H5. move/(_ z V). simpl. rewrite mem_cat negb_or. simpl. + rewrite inE negb_or. case/and3P. by []. + +apply: contraR Ny2; case/allPn=>x2 Hx2 /negbNE Hx2v. +have {H1 H2} H12: all (fun z : node => z \notin v) (x1 :: p1). +- by simpl; rewrite H1 H2. +case/splitPl: Hx2v Ey P1 H12 H3=>p0 p2 p0x2. +rewrite last_cat cat_path -cat_cons lastI cat_rcons {}p0x2 => p2y /andP [_ g_p2]. +rewrite all_cat /=. case/and3P=>H2 H3 H4 H5. +have [p3 g_p1 p1_x2 /= /andP [/andP [V1 V2] V3]] := (IHn _ _ v leq_v'_n U Sv H3 Dx Hx2). +apply/IHn=>//; exists (p3 ++ p2). - by rewrite cat_path g_p1 -p1_x2. - by rewrite last_cat -p1_x2. -- by rewrite cat_nilp; apply/implyP; case/andP=>+ _; apply/implyP. -by rewrite -cat_cons all_cat not_p1v. +simpl. +rewrite V1 /=. apply/andP. split. +- by rewrite all_cat V2 /= H4. +rewrite -cat_cons. apply/allP=>z V. +rewrite mem_cat negb_or. apply/andP; split. +- by apply: (allP V3). +move/allP: H5. move/(_ z V). rewrite mem_cat negb_or. simpl. +rewrite inE negb_or. by case/and3P. Qed. +UP TO HERE + From 35b30ff434253ef0f07028e1cc0986128b44353c Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Tue, 6 Aug 2024 19:15:31 +0200 Subject: [PATCH 14/93] continuing with refactoring graph.v. --- examples/graph.v | 808 +++++++++++++++++++++++------------------------ 1 file changed, 402 insertions(+), 406 deletions(-) diff --git a/examples/graph.v b/examples/graph.v index 5cd9cc3..a86c713 100644 --- a/examples/graph.v +++ b/examples/graph.v @@ -18,6 +18,147 @@ From pcm Require Import options axioms pred prelude. From pcm Require Import pcm unionmap natmap autopcm automap. From pcm Require Import seqext. + +(* disjointness and subset stuff *) + +Definition disjoint {A : eqType} (s1 s2 : seq A) := + all (fun x => x \notin s1) s2. + +Arguments disjoint {A} : simpl never. + +Lemma disjointC (A : eqType) (s1 s2 : seq A) : + disjoint s1 s2 = disjoint s2 s1. +Proof. +apply/idP/idP=>/allP S; apply/allP=>x X; +by apply/negP=>/S; rewrite X. +Qed. + +Lemma disjoint_catR (A : eqType) (s s1 s2 : seq A) : + disjoint s (s1 ++ s2) = + disjoint s s1 && disjoint s s2. +Proof. by rewrite /disjoint all_cat. Qed. + +Lemma disjoint_catL (A : eqType) (s s1 s2 : seq A) : + disjoint (s1 ++ s2) s = + disjoint s1 s && disjoint s2 s. +Proof. by rewrite -!(disjointC s) disjoint_catR. Qed. + +Lemma disjoint1L (A : eqType) x (s : seq A) : + disjoint [:: x] s = (x \notin s). +Proof. +apply/idP/idP. +- by apply: contraL=>X; apply/allPn; exists x=>//; rewrite negbK inE. +by apply: contraR=>/allPn [y H]; rewrite inE negbK =>/eqP <-. +Qed. + +Lemma disjoint1R (A : eqType) x (s : seq A) : + disjoint s [:: x] = (x \notin s). +Proof. by rewrite disjointC disjoint1L. Qed. + +Lemma disjoint_consL (A : eqType) x (s1 s2 : seq A) : + disjoint (x :: s1) s2 = + (x \notin s2) && disjoint s1 s2. +Proof. by rewrite -cat1s disjoint_catL disjoint1L. Qed. + +Lemma disjoint_consR (A : eqType) x (s1 s2 : seq A) : + disjoint s1 (x :: s2) = + (x \notin s1) && disjoint s1 s2. +Proof. by rewrite -cat1s disjoint_catR disjoint1R. Qed. + +Lemma disjoint_consLI (A : eqType) x (s1 s2 : seq A) : + x \notin s2 -> + disjoint s1 s2 -> + disjoint (x :: s1) s2. +Proof. by rewrite disjoint_consL=>->->. Qed. + +Lemma disjoint_consRI (A : eqType) x (s1 s2 : seq A) : + x \notin s1 -> + disjoint s1 s2 -> + disjoint s1 (x :: s2). +Proof. by rewrite disjoint_consR=>->->. Qed. + +Lemma disjoint_consLE (A : eqType) x (s1 s2 : seq A) : + disjoint (x :: s1) s2 -> + (x \notin s2) * (disjoint s1 s2). +Proof. by rewrite disjoint_consL=>/andX. Qed. + +Lemma disjoint_consRE (A : eqType) x (s1 s2 : seq A) : + disjoint s1 (x :: s2) -> + (x \notin s1) * (disjoint s1 s2). +Proof. by rewrite disjoint_consR=>/andX. Qed. + +Lemma disjoint_subL (A : eqType) (s s1 s2 : seq A) : + {subset s2 <= s1} -> + disjoint s s1 -> + disjoint s s2. +Proof. by move=>X /allP H; apply/allP=>z /X /H. Qed. + +Lemma disjoint_subR (A : eqType) (s s1 s2 : seq A) : + {subset s2 <= s1} -> + disjoint s1 s -> + disjoint s2 s. +Proof. +move=>X; rewrite disjointC=>/(disjoint_subL X). +by rewrite disjointC. +Qed. + +Lemma disjoint_eqL {A : eqType} {s s1 s2 : seq A} : + s1 =i s2 -> + disjoint s1 s = disjoint s2 s. +Proof. by move=>X; apply/idP/idP; apply: disjoint_subR=>z; rewrite X. Qed. + +Lemma disjoint_eqR {A : eqType} {s s1 s2 : seq A} : + s1 =i s2 -> + disjoint s s1 = disjoint s s2. +Proof. by move=>X; apply/idP/idP; apply: disjoint_subL=>z; rewrite X. Qed. + +Lemma disjointN (A : eqType) (s1 s2 : seq A) : + ~~ disjoint s1 s2 -> + exists2 x, x \in s1 & x \in s2. +Proof. by case/allPn=>x; rewrite negbK; exists x. Qed. + +Lemma subset_consL (A : eqType) x (s1 s2 : seq A) : + {subset x :: s1 <= s2} <-> + x \in s2 /\ {subset s1 <= s2}. +Proof. +split=>[S|[X S]]. +- by split=>[|z Z]; apply: S; rewrite inE ?eqxx ?Z ?orbT. +by move=>z; rewrite inE; case/orP=>[/eqP ->|/S//]. +Qed. + +Lemma subset_consLI (A : eqType) x (s1 s2 : seq A) : + x \in s2 -> + {subset s1 <= s2} -> + {subset x :: s1 <= s2}. +Proof. by move=>H1 H2; rewrite subset_consL. Qed. + +Lemma subset_consR (A : eqType) x (s : seq A) : + {subset s <= x :: s}. +Proof. by move=>z E; rewrite inE E orbT. Qed. + +Lemma subset_consLR (A : eqType) x (s1 s2 : seq A) : + {subset s1 <= s2} -> + {subset x :: s1 <= x :: s2}. +Proof. +move=>X z; rewrite !inE; case/orP=>[|/X] -> //. +by rewrite orbT. +Qed. + +Hint Resolve subset_consR : core. + +Lemma subset_refl (A : Type) (s : pred A) : + {subset s <= s}. +Proof. by []. Qed. + +Lemma subset_trans (A : eqType) (s1 s2 s3 : pred A) : + {subset s1 <= s2} -> + {subset s2 <= s3} -> + {subset s1 <= s3}. +Proof. by move=>S1 S2 z /S1/S2. Qed. + + + + (* Pregraphs are natmaps, mapping nodes into *) (* seq node listing the children of the node (adjacency list) *) (* Technically pregraph's a "prequiver", because this representation *) @@ -66,14 +207,14 @@ Coercion Pred_of_history (x : pregraph) : Pred_Class := Notation "x &-> v" := (ptsT pregraph x v) (at level 30). -(* p is out-node if no edge goes into it *) -Definition out_node (g : pregraph) (p : node) := - all (fun s => p \notin s) (range g). +(* x is out-node if no edge goes into it *) +Definition out_node (g : pregraph) (x : node) := + all (fun s => x \notin s) (range g). (* pregraph is graph if valid, and *) (* all nodes in all adjacency lists are good *) Definition graph_axiom (g : pregraph) := - valid g && all (all (fun p : node => (p == null) || (p \in dom g))) (range g). + valid g && all (all (fun x : node => (x == null) || (x \in dom g))) (range g). HB.mixin Record isGraph (g : pregraph) := { graph_subproof : graph_axiom g}. @@ -82,26 +223,26 @@ HB.structure Definition Graph := {g of isGraph g }. (* removing out-node causes no dangling edges; *) (* hence preserves graph axiom *) -Lemma graphF g p : - out_node g p -> +Lemma graphF g x : + out_node g x -> graph_axiom g -> - graph_axiom (free g p). + graph_axiom (free g x). Proof. move/allP=>/= Hna /andP [V /allP/= Ha]. rewrite /graph_axiom validF V /=. apply/allP=>/= s /rangeF Hs; apply/allP=>q Hq. move/allP: (Ha _ Hs) (Hna _ Hs)=>/(_ _ Hq) {}Hs. -by rewrite domF inE; case: (p =P q) Hq=>//= ->->. +by rewrite domF inE; case: (x =P q) Hq=>//= ->->. Qed. -(* children of node p obtained by non-dangingle edges *) -Definition children (g : pregraph) p : seq node := - oapp (filter (mem (dom g))) [::] (find p g). +(* children of node x obtained by non-dangingle edges *) +Definition children (g : pregraph) x : seq node := + oapp (filter (mem (dom g))) [::] (find x g). -Lemma children_undef p : children undef p = [::]. +Lemma children_undef x : children undef x = [::]. Proof. by []. Qed. -Lemma children_unit p : children Unit p = [::]. +Lemma children_unit x : children Unit x = [::]. Proof. by []. Qed. Lemma childrenND g x : @@ -178,10 +319,10 @@ Qed. Lemma path_dom g x xs : path (edge g) x xs -> - all (mem (dom g)) xs. + {subset xs <= dom g}. Proof. -elim: xs x=>[|a xs IH] x //= /andP [He Hp]. -by apply/andP; split; [case/edge_dom: He|apply: IH Hp]. +elim: xs x=>[|a xs IH] x //= /andP [/edge_dom [_ He] /IH]. +by apply: subset_consLI He. Qed. (* lifting the theory of finite graphs to unbounded pregraphs *) @@ -192,6 +333,11 @@ Fixpoint dfs (g : pregraph) (n : nat) (v : seq node) x := if (x \notin dom g) || (x \in v) then v else if n is n'.+1 then foldl (dfs g n') (x :: v) (children g x) else v. +Lemma dfs_notin g n v x : + x \notin dom g -> + dfs g n v x = v. +Proof. by elim: n=>[|n IH] /= ->. Qed. + Lemma subset_dfs g n v x : {subset v <= foldl (dfs g n) v x}. Proof. @@ -201,7 +347,7 @@ move=>y Hy; apply: IHx; case: ifP=>//= _. by apply: IHn; rewrite inE Hy orbT. Qed. -(* accumulator is bound by g *) +(* avoidance set is bound by g *) Lemma subset_foldl_dfs_dom g n v x : {subset v <= dom g} -> {subset foldl (dfs g n) v x <= dom g}. @@ -246,18 +392,14 @@ Inductive dfs_path g (v : seq node) x y : Prop := DfsPath xs of path (edge g) x xs & y = last x xs & - (* v and x :: xs are disjoint sequences *) - all (fun z => z \notin v) (x :: xs) && - all (fun z => z \notin x :: xs) v. + disjoint v (x :: xs). Lemma dfs_path_id g v x : - x \in dom g -> x \notin v -> dfs_path g v x x. Proof. -move=>Gx Vx; apply: (DfsPath (xs:=[::]))=>//=. -rewrite andbT Vx /=; apply/allP=>z; rewrite inE. -by case: (z =P x) Vx=>//= -> /negbTE ->. +move=>Vx; apply: (DfsPath (xs:=[::]))=>//=. +by rewrite disjoint1R. Qed. Lemma dfs_pathP g n x y v : @@ -269,449 +411,303 @@ Lemma dfs_pathP g n x y v : reflect (dfs_path g v x y) (y \in dfs g n v x). Proof. elim: n=>[|n IHn] /= in x y v * => Hv Uv Sv Ny Dx. -- rewrite addn0 in Hv. - case: (uniq_min_size Uv Sv Hv) Ny=>_ Ev /negbTE Ny. - suff: ~ dfs_path g v x y. - - by rewrite Dx /= if_same Ny; apply: ReflectF. - by case=>/= xs E Ey; rewrite Ev Dx. -rewrite Dx /=; have [Vx|Vx] := ifPn. -- by rewrite (negbTE Ny); apply: ReflectF; case=>xs /=; rewrite Vx. -set v1 := x :: v; set a := children g x; have [->|Nyx] := eqVneq y x. +- rewrite addn0 in Hv; case: (uniq_min_size Uv Sv Hv) Ny=>_ Ev /negbTE Ny. + suff: ~ dfs_path g v x y by rewrite Dx if_same Ny; apply: ReflectF. + by case=>xs E _; rewrite disjoint_consR Ev Dx. +rewrite Dx /=; have [Vx|Vx] := ifPn. +- by rewrite (negbTE Ny); apply: ReflectF=>[[xs]]; rewrite disjoint_consR Vx. +set v1 := x :: v; set a := children g x; have [->|/eqP Nyx] := eqVneq y x. - by rewrite subset_dfs ?inE ?eqxx //; apply/ReflectT/dfs_path_id. -apply: (@equivP (exists2 x1, x1 \in a & dfs_path g v1 x1 y))=>/=; last first. -- split=> {IHn} [[x1 Hx1 [p1 P1 E1]] | [p /shortenP[]]] /=. - - case/and3P=>/andP [H1 H2 H3 H4]; apply: (DfsPath (xs:=x1::p1))=>//=. - - by rewrite edge_children -/a Hx1. - rewrite Vx /= -!andbA; apply/and3P; split. - - by apply: contra H1; rewrite /v1 inE orbC=>->. - by apply: sub_all H2=>z; apply: contra; rewrite /v1 inE orbC=>->. - apply/allP=>z /= V; move/allP/(_ z V): H4=>/= H4. - rewrite inE (negbTE H4) orbF. - by case: (z =P x) V=>// ->; rewrite (negbTE Vx). - case=>[_ _ _ /= Eyx|x1 xs /=]; first by case/eqP: Nyx. - case/andP=>Hx1 Hp1 /and3P [H1 H2 H3] H4 H5 /andP [/andP [_ /allP H6] H7]. - exists x1=>//; apply: (DfsPath (xs:=xs))=>//=. - rewrite inE negb_or in H1; case/andP: H1=>H11 H12. - rewrite -!andbA; apply/and4P; split. - - by rewrite /v1 inE negb_or eq_sym H11; apply/H6/H4; rewrite inE eqxx. - - apply/allP=>/= z Z; rewrite /v1 inE negb_or. - have /H6 -> : z \in p by apply/H4; rewrite inE Z orbT. - by case: (z =P x) Z H12=>// ->->. - - by rewrite inE negb_or H11 H12. - by apply: sub_all H7=>z; apply: contra=>/H4; rewrite inE orbC=>->. +apply: (@equivP (exists2 x1, x1 \in a & dfs_path g v1 x1 y))=>/=; last first. +- split=>{IHn} [[x1 Hx1 [p1 P1 E1 D1]]|[p /shortenP []]]. + - apply: (DfsPath (xs:=x1::p1))=>//=; first by rewrite edge_children -/a Hx1. + by rewrite disjoint_consR Vx (disjoint_consLE D1). + case=>[_ _ _ /Nyx|] //= x1 xs /andP [Hx1 Hp1] /and3P [N1 _ _] S1 E1 D1. + exists x1=>//; apply: (DfsPath (xs:=xs))=>//; rewrite disjoint_consL N1. + by rewrite (disjoint_consRE (disjoint_subL (subset_consLR S1) D1)). move: (Dx). -have {Nyx Ny}: y \notin v1 by apply/norP. -have {Sv Dx}: {subset v1 <= dom g}. -- by move=>z; rewrite inE; case/orP=>[/eqP ->//|/Sv]. -have {Vx Uv}: uniq v1 by rewrite /= Vx. -have {Hv}: size (dom g) <= size v1 + n by rewrite addSnnS. +have {Nyx Ny} : y \notin v1 by apply/norP; move/eqP: Nyx. +have {Sv Dx} : {subset v1 <= dom g} by apply: subset_consLI. +have {Vx Uv} : uniq v1 by rewrite /= Vx. +have {Hv} : size (dom g) <= size v1 + n by rewrite addSnnS. have : {subset a <= dom g} by apply: childrenD. -elim: {x v}a (x) v1=>[|x a IHa] x' v /= Dxa leq_v'_n U Sv Nv Dx'. +elim: {x v}a (x) v1=>[|x xs IHa] x' v /= Dxs Hv U Sv Nv Dx'. - by rewrite (negbTE Nv); apply: ReflectF; case. -have Dx : x \in dom g by apply: Dxa; rewrite inE eqxx. -have Da : {subset a <= dom g} by move=>z Z; apply: Dxa; rewrite inE Z orbT. +have Dx : x \in dom g by apply: Dxs; rewrite inE eqxx. +have Da : {subset xs <= dom g} by apply/subset_trans/Dxs/subset_consR. set v2 := dfs g n v x. -have v2v : {subset v <= v2} := @subset_dfs g n v [:: x]. -have [Hy2 | Ny2] := boolP (y \in v2). +have Sv2 : {subset v <= v2} := @subset_dfs g n v [:: x]. +have [Hy2|Ny2] := boolP (y \in v2). - rewrite subset_dfs //; apply: ReflectT. by exists x; [rewrite inE eq_refl|apply/IHn]. apply: {IHa} (equivP (IHa _ _ _ _ _ _ Ny2 Dx))=>//. -- by apply/(leq_trans leq_v'_n); rewrite leq_add2r; apply: uniq_leq_size. -- by apply: uniq_dfs. +- by rewrite (leq_trans Hv) ?leq_add2r ?uniq_leq_size. +- by rewrite uniq_dfs. - by apply: subset_dfs_dom. -split=>[][x1 Hx1 [p1 P1 Ey]]. -- rewrite -!andbA=>/and3P [H1 H2 H3]. - exists x1; first by rewrite inE Hx1 orbT. - apply: (DfsPath (xs:=p1))=>//=. - rewrite -!andbA; apply/and3P; split. - - by apply: contra H1=>/v2v. - - by apply: sub_all H2=>z; apply: contra=>/v2v. - by apply/allP=>/= z /v2v; apply: (allP H3). -rewrite /= -!andbA; case/and3P=>H1 H2 H3. -suff [N1 N2] : all (fun z : node => z \notin v2) (x1 :: p1) /\ - all (fun z : node => z \notin x1 :: p1) v2. -- move: Hx1; rewrite inE; case/orP=>[/eqP Ex1|Hx1]. - - subst x1; exfalso; move: N1=>/= /andP [+ _]. - by move/negbTE/negP; apply; apply/IHn=>//; apply: dfs_path_id. - exists x1=>//; apply: (DfsPath (xs:=p1))=>//=. - rewrite -!andbA; apply/and3P; split=>//; last by case/andP: N1. - by rewrite (allP N1) // inE eqxx. -split. -- apply: contraR Ny2; case/allPn => x2 Hx2 /negbNE Hx2v. - have {H1 H2} H12 : all (fun z : node => z \notin v) (x1 :: p1). - - by simpl; rewrite H1 H2. - case/splitPl: Hx2 Ey P1 H12 H3=>p0 p2 p0x2. - rewrite last_cat cat_path -cat_cons lastI cat_rcons {}p0x2 => p2y /andP [_ g_p2]. - rewrite all_cat /=. case/and3P=>H2 H3 H4 H5. - have [p3 g_p1 p1_x2 /= /andP [/andP [V1 V2] V3]] := (IHn _ _ v leq_v'_n U Sv H3 Dx Hx2v). - apply/IHn=>//; exists (p3 ++ p2). - - by rewrite cat_path g_p1 -p1_x2. - - by rewrite last_cat -p1_x2. - simpl. rewrite V1 /=. - rewrite all_cat V2 H4 /=. - rewrite -cat_cons. apply/allP=>z V. - rewrite mem_cat negb_or. apply/andP. split. - - by apply: (allP V3). - move/allP: H5. move/(_ z V). simpl. rewrite mem_cat negb_or. simpl. - rewrite inE negb_or. case/and3P. by []. - -apply: contraR Ny2; case/allPn=>x2 Hx2 /negbNE Hx2v. -have {H1 H2} H12: all (fun z : node => z \notin v) (x1 :: p1). -- by simpl; rewrite H1 H2. -case/splitPl: Hx2v Ey P1 H12 H3=>p0 p2 p0x2. -rewrite last_cat cat_path -cat_cons lastI cat_rcons {}p0x2 => p2y /andP [_ g_p2]. -rewrite all_cat /=. case/and3P=>H2 H3 H4 H5. -have [p3 g_p1 p1_x2 /= /andP [/andP [V1 V2] V3]] := (IHn _ _ v leq_v'_n U Sv H3 Dx Hx2). -apply/IHn=>//; exists (p3 ++ p2). -- by rewrite cat_path g_p1 -p1_x2. -- by rewrite last_cat -p1_x2. -simpl. -rewrite V1 /=. apply/andP. split. -- by rewrite all_cat V2 /= H4. -rewrite -cat_cons. apply/allP=>z V. -rewrite mem_cat negb_or. apply/andP; split. -- by apply: (allP V3). -move/allP: H5. move/(_ z V). rewrite mem_cat negb_or. simpl. -rewrite inE negb_or. by case/and3P. +split=>[][x1 Hx1 [p1 P1 Ey D1]]. +- exists x1; first by rewrite inE Hx1 orbT. + by apply: DfsPath (disjoint_subR Sv2 D1). +have Nx1 : x1 \notin v by rewrite (disjoint_consRE D1). +suff D2 : disjoint v2 (x1 :: p1). +- move: Hx1; rewrite inE; case/orP=>[/eqP ?|Hx1]; last first. + - by exists x1=>//; apply: DfsPath D2. + subst x1; have : x \notin v2 by rewrite (disjoint_consRE D2). + by move/negP; elim; apply/IHn=>//; apply: dfs_path_id. +apply: contraR Ny2=>/disjointN [/= x2 Hx2v Hx2]. +case/splitPl: Hx2 Ey P1 D1=>/= pl pr Ex2. +rewrite last_cat cat_path -cat_cons lastI cat_rcons {}Ex2. +move=>Ey /andP [_ P1]; rewrite disjoint_catR=>/andP [D1 D2]. +have Nx2 : x2 \notin v by rewrite (disjoint_consRE D2). +have [p P E D] := IHn _ _ v Hv U Sv Nx2 Dx Hx2v. +apply/IHn=>//; exists (p ++ pr). +- by rewrite cat_path P -E P1. +- by rewrite last_cat -E. +by rewrite -cat_cons disjoint_catR D (disjoint_consRE D2). Qed. -UP TO HERE - - - - - - - -(*******************) - +(* start dfs with p as avoidance set, then filter out p. *) +(* this computes only the nodes visited during dfs. *) +(* not for client use, hence primed name *) +Definition component' (p : pred node) g x : seq node := + filter (predC p) (dfs g (size (dom g)) (filter p (dom g)) x). + +(* y is connected from x if y is visited during dfs *) +(* under assumption that x is in dom g *) +Definition connect p (g : pregraph) := + [rel x y | (x \in dom g) && (y \in component' p g x)]. + +Lemma connectP p (g : pregraph) x y : + reflect [/\ x \in dom g, ~~ p x & exists xs, + [/\ path (edge g) x xs, y = last x xs & + {in xs, forall z, ~~ p z}]] + (y \in connect p g x). +Proof. +rewrite /connect/component' inE mem_filter /= andbA. +case: (boolP (x \in dom g))=>Dx; last by constructor; case. +case: (boolP (p y))=>Py. +- constructor; case=>_ Px [xs][P E Pxs]. + suff : ~~ p y by rewrite Py. + have : y \in x :: xs by rewrite E mem_last. + by rewrite inE=>/orP [/eqP ->|/Pxs]. +apply: (iffP (dfs_pathP _ _ _ _ _))=>//. +- by rewrite leq_addl. +- by rewrite filter_uniq. +- by move=>z; rewrite mem_filter=>/andP []. +- by rewrite mem_filter negb_and Py. +- case=>xs P E /disjoint_consRE []. + rewrite mem_filter Dx andbT => Px D. + split=>//; exists xs; split=>// z Xz; move: (allP D z Xz). + by rewrite mem_filter (path_dom P Xz) andbT. +case=>_ Px [xs][P E Pxs]; exists xs=>//. +rewrite disjoint_consR mem_filter negb_and Px /=. +by apply/allP=>z /Pxs; rewrite mem_filter negb_and=>->. +Qed. -Variant dfs_path (p : nodeset) g (v : seq ptr) x y : Prop := - DfsPath xs of - path (edge g) x xs & - y = last x xs & - nilp xs ==> (x \in dom g) & - all (fun z => z \notin [predU v & (dom p)]) (x::xs). +Lemma connect_undef p x : connect p undef x =i pred0. +Proof. by move=>y; apply/connectP; case; rewrite dom_undef. Qed. +Lemma connect_unit p x : connect p Unit x =i pred0. +Proof. by move=>y; apply/connectP; case; rewrite dom0. Qed. +Lemma connectDx p g x y : + y \in connect p g x -> + (x \in dom g) * ~~ p x. +Proof. by case/connectP. Qed. -(* ps : nodes that stop the traversal *) -(* g : pregraph to traverse in dfs *) -(* n : traversal depth *) -(* acc : accumulator; currently observed sequence of nodes *) -(* x : starting node *) +Lemma connectDy p g x y : + y \in connect p g x -> + (y \in dom g) * ~~ p y. +Proof. +case/connectP=>Dx Px [xs][P E Pxs]. +have : y \in x :: xs by rewrite E mem_last. +rewrite inE; case/orP=>[/eqP ->|Dy]. +- by rewrite Dx Px. +by rewrite (path_dom P) // Pxs. +Qed. -Fixpoint dfs (ps : fset node) (g : pregraph) (n : nat) (acc : seq node) x := - if [|| x \notin dom g, x \in acc | x \in dom ps] then acc else - if n is n'.+1 then foldl (dfs ps g n') (x :: acc) (children g x) else acc. +Lemma connectD p g x y : + y \in connect p g x -> + (x \in dom g) * (y \in dom g). +Proof. by move=>C; rewrite (connectDx C) (connectDy C). Qed. -(* -Fixpoint dfs (ps : fset node) (g : pregraph) (n : nat) (acc : seq node) x := - if (x \in acc) || (x \in dom ps) then acc - else - if x \in dom g then - if n is n'.+1 then foldl (dfs' ps g n') (x :: acc) (children g x) - else acc - else acc. - -Fixpoint dfs (p : fset node) (g : pregraph) n (v : seq node) x := - if x \in [predU v & (dom p)] then v else - if x \in dom g then - if n is n'.+1 then foldl (dfs p g n') (x :: v) (children g x) else v - else v. -*) +Lemma connectX p g x y : + y \in connect p g x -> + (~~ p x) * (~~ p y). +Proof. by move=>C; rewrite (connectDx C) (connectDy C). Qed. -(* accumulator grows *) -Lemma subset_dfs ps g n acc x : - {subset acc <= foldl (dfs ps g n) acc x}. +Lemma connectDN p g x : + x \notin dom g -> + connect p g x =i pred0. Proof. -elim: n x acc => [|n IHn] /=; elim=>[|x xs IHx] acc //=. -- by case: ifP. -move=>y Hy; apply: IHx; case: ifP=>//= _. -by apply: IHn; rewrite inE Hy orbT. +move=>Dx y; apply/negP=>/connectD []. +by rewrite (negbTE Dx). Qed. -(* accumulator is bounded by g *) -Lemma subset_foldl_dfs_dom ps g n acc x : - {subset acc <= dom g} -> - {subset foldl (dfs ps g n) acc x <= dom g}. +Lemma connectXN (p : pred node) g x : + p x -> + connect p g x =i pred0. Proof. -elim: n x acc=>[|n IHn]; elim=>[|x xs IHx] acc //=. -- by case: ifP=>_; apply: IHx. -case: ifP; first by case: (x \in dom g)=>//= H; apply: IHx. -case Dx: (x \in dom g)=>//= H Gx; apply/IHx/IHn. -by move=>z; rewrite inE; case/orP=>[/eqP ->|/Gx]. -Qed. - -Lemma subset_dfs_dom ps g n acc x : - {subset acc <= dom g} -> - {subset dfs ps g n acc x <= dom g}. -Proof. -case: n=>[|n] H /=; case: ifP=>//=. -case Dx : (x \in dom g)=>//= _; apply: subset_foldl_dfs_dom. -by move=>z; rewrite inE; case/orP=>[/eqP ->|/H]. +move=>Hx y; apply/negP=>/connectX []. +by rewrite Hx. Qed. -Inductive dfs_path g v x y : Prop := - DfsPath xs of path (edge g) x xs & y = last x xs & [disjoint x :: xs & v]. - - - -(*******************) - +Lemma connect0 p g x : + x \in connect p g x = (x \in dom g) && ~~ p x. +Proof. by apply/connectP/andP; case=>// H1 H2; split=>//; exists [::]. Qed. +Lemma connect0I p g x : + x \in dom g -> + ~~ p x -> + x \in connect p g x. +Proof. by rewrite connect0=>->->. Qed. -Variant dfs_path (p : nodeset) g (v : seq ptr) x y : Prop := - DfsPath xs of - path (edge g) x xs & - y = last x xs & - nilp xs ==> (x \in dom g) & - all (fun z => z \notin [predU v & (dom p)]) (x::xs). +Lemma connect0N p g x y : + x \in dom g -> + ~~ p x -> + y \notin connect p g x -> + x != y. +Proof. by move=>Gx Px; apply: contra=>/eqP <-; rewrite connect0 Gx. Qed. -Lemma dfs_path_id p g v x : - x \in dom g -> - x \notin dom p -> x \notin v -> dfs_path p g v x x. +Lemma connect_trans p g : transitive (connect p g). Proof. -move=>Hx Nx Nv; apply (DfsPath (xs:=[::]))=>//=. -by rewrite andbT inE negb_or Nv. +move=>x y z /connectP [Dy Hy][ys][Py Ey Hys]. +case/connectP=>Dx Hx [xs][Px Ex Hxs]. +apply/connectP; split=>//. +exists (ys ++ xs); split=>[||w]. +- by rewrite cat_path -Ey Py Px. +- by rewrite last_cat -Ey. +by rewrite mem_cat; case/orP; [apply: Hys|apply: Hxs]. Qed. -Lemma path_dom g x xs : - path (edge g) x xs -> - all (fun z => z \in dom g) xs. +Lemma connectUnL p g1 g2 x : + valid (g1 \+ g2) -> + {subset connect p g1 x <= + connect p (g1 \+ g2) x}. Proof. -elim: xs x=>//=a xs IH x; case/andP=>He Hp. -by apply/andP; split; [case/edge_dom: He | apply: (IH _ Hp)]. +move=>V y /connectP [Dx Hx][xs][Px Ey Hxs]. +apply/connectP; split=>//; first by rewrite domUn inE V Dx. +exists xs; split=>//. +by apply: sub_path Px=>z w; apply: edgeUnL V. Qed. -Lemma dfs_pathP p g n x y v : - size (dom g) <= size v + n -> - uniq v -> - {subset v <= dom g} -> - y \notin v -> - reflect (dfs_path p g v x y) (y \in dfs p g n v x). +Lemma connectUnR p g1 g2 x : + valid (g1 \+ g2) -> + {subset connect p g2 x <= + connect p (g1 \+ g2) x}. +Proof. by rewrite joinC; apply: connectUnL. Qed. + +Lemma connect_sub g p1 p2 x : + {subset p2 <= p1} -> + {subset connect p1 g x <= connect p2 g x}. Proof. -elim: n => [|n IHn] /= in x y v * => Hv Uv Sv Ny. -- rewrite addn0 in Hv. - case: (uniq_min_size Uv Sv Hv) Ny=>_ /(_ y) Ey. - move/negbTE=>Ny; rewrite Ey in Ny. - suff: ~dfs_path p g v x y by case: (links g x)=>[_|]; rewrite if_same Ey Ny; apply: ReflectF. - case; case; first by move=>/= _ <-; rewrite Ny. - by move=>a l /path_dom/allP /(_ y) + /= Eyl; rewrite Ny Eyl mem_last; move/[apply]. -case Epl: (links g x)=>[ls|] /=; last first. -- rewrite if_same (negbTE Ny); apply: ReflectF; case=>/= xs. - case: xs=>[|??]/=; last by rewrite (negbTE (links_edge_not _ Epl)). - by move=>_ Ey Hy; move/linksF/negbTE: Epl; rewrite Hy. -have [Vx|] := ifPn. -- rewrite (negbTE Ny); apply: ReflectF; case=>/= xs. - by rewrite Vx. -rewrite negb_or; case/andP=>Nx Px. -set v1 := x :: v; have [-> | Nyx] := eqVneq y x. -- rewrite subset_dfs; last by rewrite inE eq_refl. - apply: ReflectT; apply: dfs_path_id=>//. - by case/linksT: Epl. -apply: (@equivP (exists2 x1, x1 \in ls & dfs_path p g v1 x1 y)); last first. -- split=> {IHn} [[x1 Hx1 [p1 P1 Ey Py]] | [p1 /shortenP []]] /=. - - case/andP; rewrite !inE !negb_or -andbA; case/and3P=>Ex1 Ex1v Nx1 Ha1; apply: (DfsPath (xs:=x1::p1))=>//=. - - by rewrite P1 andbT (links_edge _ Epl). - rewrite !inE !negb_or Nx Px Ex1v Nx1 /=. - by apply/sub_all/Ha1=>z; rewrite !inE !negb_or -andbA; case/and3P=>_->->. - case=>[_ _ _ /= Eyx | x1 xs /=]; first by case/eqP: Nyx. - rewrite (links_edge _ Epl) /=; case/andP=>Hx1 Hp1 /and3P [H11 H12 H13] H2 H3 H4 /andP [H5 H6]. - exists x1=>//; apply: (DfsPath (xs:=xs))=>//. - - by apply/implyP=>_; case/linksT: Epl=>_ /allP /(_ _ Hx1). - apply/allP=>z /[dup] Hz /H2; move/allP: H6=>/(_ z) /[apply]. - rewrite !inE !negb_or; case/andP=>->->/=; rewrite !andbT; apply/negP=>/eqP Ez. - by rewrite -Ez Hz in H11. -have{Nyx Ny}: y \notin v1 by apply/norP. -have{Nx Uv}: uniq v1 by rewrite /= Nx. -have{Sv}: {subset v1 <= dom g}. -- move=>z; rewrite inE; case/orP; last by move/Sv. - by move/eqP=>{z}->; case/linksT: Epl. -have{Hv}: size (dom g) <= size v1 + n by rewrite addSnnS. -case/linksT: Epl=>_ Hl. -elim: {x Px v}ls Hl v1 => /= [_|x a IHa /andP [Ha Hl]] v /= Hv Sv Uv Nv. -- by rewrite (negPf Nv); apply: ReflectF; case. -set v2 := dfs p g n v x; have v2v: {subset v <= v2} := @subset_dfs p g n v [:: x]. -have [Hy2 | Ny2] := boolP (y \in v2). -- rewrite subset_dfs //; apply: ReflectT; exists x; first by rewrite inE eq_refl. - by apply/IHn. -apply: {IHa}(equivP (IHa Hl _ _ _ _ Ny2)). -- by apply: (leq_trans Hv); rewrite leq_add2r; apply: uniq_leq_size. -- by apply: subset_dfs_dom. -- by apply: uniq_dfs. -split=> [] [x1 Hx1 [p1 P1 Ey Py Nay]]. - exists x1; first by rewrite inE Hx1 orbT. - apply: (DfsPath (xs:=p1))=>//; apply/sub_all/Nay=>z; apply: contra. - by rewrite !inE; case/orP; [move/v2v=>->|move=>->; rewrite orbT]. -suff Nv2: all (fun z => z \notin v2) (x1 :: p1). -- move: Hx1; rewrite inE; case/orP=>[/eqP Ex1|Hx1]; last first. - - exists x1=>//; apply: (DfsPath (xs:=p1))=>//. - apply/allP=>z Hz; rewrite inE negb_or; apply/andP; split. - - by move/allP: Nv2; apply. - by move/allP: Nay=>/(_ _ Hz); rewrite inE negb_or; case/andP. - rewrite {x1}Ex1 in P1 Py Ey Nay Nv2. - exfalso; move: Nv2=>/=; case/andP=>+ _; move/negbTE/negP; apply. - suff [Nx Nxp]: (x \notin v) /\ (x \notin dom p). - - by apply/IHn=>//; apply: dfs_path_id. - by move: Nay=>/=; case/andP; rewrite inE negb_or; case/andP. -apply: contraR Ny2; case/allPn => x2 Hx2 /negbNE Hx2v. -case/splitPl: Hx2 Ey P1 Nay => p0 p2 p0x2. -rewrite last_cat cat_path -cat_cons lastI cat_rcons {}p0x2 => p2y /andP[_ g_p2]. -rewrite all_cat /=; case/and3P=> {p0}_; rewrite inE negb_or; case/andP=>Nx2v Nx2p Np2. -have{Nx2v Hx2v} [p3 g_p1 p1_x2 V2 not_p1v] := (IHn _ _ v Hv Uv Sv Nx2v Hx2v). -apply/IHn=>//; exists (p3 ++ p2)=>//. -- by rewrite cat_path g_p1 -p1_x2. -- by rewrite last_cat -p1_x2. -- by rewrite cat_nilp; apply/implyP; case/andP=>+ _; apply/implyP. -by rewrite -cat_cons all_cat not_p1v. +move=>S y /connectP [Dx Hx][xs][Px Ey Hxs]. +apply/connectP; split=>//; first by apply: contra (S x) Hx. +by exists xs; split=>// z /Hxs; apply/contra/S. Qed. -Definition component p g x : seq ptr := dfs p g (size (dom g)) [::] x. - -Definition connect p g : rel ptr := [rel x y | y \in component p g x]. -Canonical connect_app_pred p g x := ApplicativePred (connect p g x). - -Corollary connectP p g x y : - reflect (exists xs, [/\ path (edge g) x xs, - y = last x xs, - nilp xs ==> (x \in dom g) & - all (fun z => z \notin (dom p)) (x::xs)]) - (y \in connect p g x). +(* exteding connecting path from left *) +Lemma edge_connectI p g x y z : + x \in dom g -> + ~~ p x -> + edge g x y -> + z \in connect p g y -> + z \in connect p g x. Proof. -apply: (iffP (dfs_pathP _ _ _ _ _ _))=>//. -- by case=>xs P Ey Pv Ha; exists xs. -by case=>xs [P Ey Pv Ha]; apply: (DfsPath (xs:=xs)). +move=>Dx Hx H /connectP [Dy Hy][ys][Py Ez Hys]. +apply/connectP; split=>//; exists (y::ys). +split=>[||w] //=; first by rewrite H Py. +by rewrite inE; case/orP=>[/eqP ->|/Hys]. Qed. -Corollary connectQ p g x y : - reflect [\/ [/\ x = y, x \in dom g & x \notin dom p] | - [/\ x \notin dom p, y \notin dom p & - exists xs, [/\ path (edge g) x xs, - edge g (last x xs) y & - forall z, z \in xs -> z \notin dom p]]] - (y \in connect p g x). +(* exteding connecting path from right *) +Lemma connect_edgeI p g x y z : + y \in connect p g x -> + edge g y z -> + ~~ p z -> + z \in connect p g x. Proof. -case: connectP=>H; constructor. -- case: H; case/lastP=>[|zs z][/= H1 H2 H3 H4]; [left|right]. - - by rewrite andbT in H4; rewrite H2 H3 H4. - case/andP: H4=>H41 /allP H42; rewrite last_rcons in H2. - split=>//. - - by rewrite H2; apply: H42; rewrite mem_rcons inE eqxx. - rewrite rcons_path in H1; case/andP: H1=>H11 H12. - exists zs; split=>//=; first by rewrite H2. - by move=>w W; apply: H42; rewrite mem_rcons inE W orbT. -move=>X; apply: H; case: X. -- by case=>H1 H2 H3; exists [::]; split=>//=; rewrite H3. -case=>H1 H2 [xs][H3 H4 H5]; exists (rcons xs y). -rewrite rcons_path H3 H4 last_rcons /nilp size_rcons; split=>//=. -rewrite H1; apply/allP=>z; rewrite mem_rcons inE. -by case/orP=>[/eqP ->//|/H5]. +case/connectP=>Dx Hx [xs][Px Ey Hxs] H Hz. +apply/connectP; split=>//. +exists (rcons xs z); split=>/=. +- by rewrite rcons_path Px -Ey H. +- by rewrite last_rcons. +by move=>w; rewrite mem_rcons inE; case/orP=>[/eqP ->|/Hxs]. Qed. -Lemma connect_trans p g : transitive (connect p g). +(* destructing connecting path from left *) +Lemma edge_connectE p g x y : + y \in connect p g x -> + y = x \/ + exists2 z, edge g x z & y \in connect p g z. Proof. -move=> x y z /connectP [xs][Hxs -> Hp Ha] /connectP [ys][Hys -> Hp' Ha']; apply/connectP. -exists (xs ++ ys); rewrite cat_path last_cat Hxs /=; split=>//. -- by rewrite cat_nilp; apply/implyP; case/andP=>+ _; apply/implyP. -rewrite all_cat andbA; apply/andP; split=>//. -by move: Ha'=>/=; case/andP. +case/connectP=>Dx H [[|a xs]][/= Px Ey Hxs]; first by left. +right; case/andP: Px=>Px Pxs; exists a=>//. +apply/connectP; split. +- by rewrite (edge_dom Px). +- by apply: Hxs; rewrite inE eqxx. +exists xs; split=>// z Z; apply: Hxs. +by rewrite inE Z orbT. Qed. -Lemma connect_refl p g x : - x \in dom g -> - x \in connect p g x = (x \notin dom p). +(* deconstructing connecting path from right *) +Lemma connect_edgeE p g x y : + y \in connect p g x -> + y = x \/ + exists2 z, z \in connect p g x & edge g z y. Proof. -move=>Gx; apply/connectQ/idP; first by case; case. -by move=>Px; left. +case/connectP=>Dx H [xs]. +case: {xs}(lastP xs)=>[|xs a][/= Px Ey Hxs]; first by left. +right; rewrite last_rcons in Ey; rewrite -{a}Ey in Px Hxs. +rewrite rcons_path in Px; case/andP: Px=>Pxs Px. +exists (last x xs)=>//; apply/connectP; split=>//. +exists xs; split=>// z Z; apply: Hxs. +by rewrite mem_rcons inE Z orbT. Qed. -Lemma connect_mem p g x y : - x \in dom g -> - x \notin dom p -> - y \notin connect p g x -> - x != y. -Proof. by move=>Gx Px; apply: contra=>/eqP <-; rewrite connect_refl. Qed. +(* part of g reachable from x (avoiding nothing) *) +Definition subconnect g x : pregraph := + um_filterk (connect pred0 g x) g. -Lemma connect_undef p x : connect p undef x =i pred0. -Proof. -move=>y; apply/connectP; case; case=>[|k xs][] /=; first by rewrite dom_undef. -by rewrite edge_undef. -Qed. +(* part of g unreachable from x (avoiding nothing) *) +Definition subdisconnect g x : pregraph := + um_filterk (negb \o connect pred0 g x) g. -Lemma connect0 p g x : x \in dom g -> x \notin dom p -> x \in connect p g x. -Proof. by move=>Hd Hp; apply/connectP; exists [::]; rewrite /= andbT. Qed. +Lemma connect_split g x : + g = subconnect g x \+ subdisconnect g x. +Proof. by apply: umfilt_predC. Qed. -Lemma connectUn p g0 g x : - valid (g \+ g0) -> - {subset connect p g x <= connect p (g \+ g0) x}. +(* reachable component of a graph is a graph *) +Lemma graph_subconnect g x : + graph_axiom g -> + graph_axiom (subconnect g x). Proof. -move=>V y; case/connectP=>xs [Hp {y}-> /implyP Nxs Ha]. -apply/connectP; exists xs; split=>//; last first. -- rewrite domUn inE; rewrite V /=. - by apply/implyP=>/Nxs->. -elim: xs {Nxs Ha}x Hp=>//= y xs IH x; case/andP=>Hxy Hxs. -by apply/andP; split; [apply: edgeUn | apply: IH]. +have E := connect_split g x. +case/andP=>V Ha; apply/andP; split; first by rewrite pfV. +apply/allP=>/= xs Hxs; move/allP: Ha=>/(_ xs). +rewrite {1}E rangeUn inE -E V Hxs /= => /(_ erefl) Ha. +case/mem_rangeX: Hxs=>k /In_umfiltX [/= Hk] /In_find Hf. +apply/allP=>n Hn; move/allP: Ha=>/(_ _ Hn). +case/boolP: (n == null)=>//= Hnn. +rewrite /subconnect dom_umfiltk inE /= => En. +rewrite En andbT /=; apply: (connect_trans Hk). +apply/connectP; split=>//=; first by rewrite (connectD Hk). +exists [:: n]; split=>//=. +by rewrite (find_edge _ Hf) En Hn. Qed. -Lemma connect_dom p g x y : - y \in connect p g x -> (x \in dom g) * (y \in dom g). -Proof. -case/connectP; case=>[|h t]; case; first by move=>/= _ ->. -move/[dup]/path_dom=>Ha /=; case/andP=>/edge_dom [Hx _] _ Ey _ _; split=>//. -by apply: (allP Ha); rewrite Ey; exact: mem_last. -Qed. -Corollary connect_notd p g x : x \notin dom g -> connect p g x =i pred0. -Proof. -move=>Hx y; rewrite [RHS]inE; apply/negP; case/connect_dom=>E. -by rewrite E in Hx. -Qed. +UP TO HERE -Lemma connect_notp p g x : x \in dom p -> connect p g x =i pred0. -Proof. -move=>Hx y; rewrite [RHS]inE; apply/negbTE/connectP. -by case=>xs /= [_ _ _]; case/andP; rewrite Hx. -Qed. -Lemma connect_sub g p1 p2 x : - {subset dom p2 <= dom p1} -> - {subset connect p1 g x <= connect p2 g x}. -Proof. -move=>Hp z; case/connectP=>xs [Hxs {z}-> H Ha]. -apply/connectP; exists xs; split=>//. -by apply/sub_all/Ha=>z; apply/contra/Hp. -Qed. -Lemma connect_links_sub p g x ns : - find x g = Some ns -> - forall y, y \in connect p g x -> (y == x) || has (fun n => y \in connect p g n) ns. -Proof. -move=>Hx y. -case/connectP=>xs [Hp {y}-> H]. -case: xs Hp H=>[|h xs]/=; first by move=>_ _; rewrite eq_refl. -case/andP; rewrite (find_edge _ Hx); case/andP=>Hh Hns Hxs _; case/and3P=>Hxp Hcp Ha. -apply/orP; right; apply/hasP; exists h=>//. -by apply/connectP; exists xs; rewrite Hh /= Hcp Ha; split=>//; apply: implybT. -Qed. -Lemma connect_eq_links p g x ns : - find x g = Some ns -> - x \notin dom p -> - forall y, y \in connect p g x = (y == x) || has (fun n => y \in connect p g n) ns. -Proof. -move=>Hx Hd y; apply/iffE; split; first by apply: connect_links_sub. -case/orP. -- move/eqP=>->; apply: connect0=>//. - by move/find_some: Hx. -case/hasP=>n Hn. -case/connectP=>xs [Hp {y}-> H] /=; case/andP=>Hn1 Ha; apply/connectP. -exists (n::xs)=>/=; rewrite Hp Hd Hn1 Ha /= andbT; split=>//. -rewrite (find_edge _ Hx) Hn andbT. -by case: {Ha}xs Hp H=>//= ??; case/andP; case/edge_dom. -Qed. + + + +(*******************) + (* TODO generalize to arbitrary boundary? *) Definition subconnect p g := um_filterk (fun z => z \in connect Unit g p) g. From 3c21f95655f998230bcd7dfae762317e7482c9b8 Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Wed, 7 Aug 2024 20:06:08 +0200 Subject: [PATCH 15/93] blah --- examples/graph.v | 699 +++++++++++++++++++++++------------------------ 1 file changed, 342 insertions(+), 357 deletions(-) diff --git a/examples/graph.v b/examples/graph.v index a86c713..f14d604 100644 --- a/examples/graph.v +++ b/examples/graph.v @@ -18,8 +18,46 @@ From pcm Require Import options axioms pred prelude. From pcm Require Import pcm unionmap natmap autopcm automap. From pcm Require Import seqext. +(* subset *) -(* disjointness and subset stuff *) +Lemma subset_refl (A : Type) (s : pred A) : + {subset s <= s}. +Proof. by []. Qed. + +Lemma subset_trans (A : eqType) (s1 s2 s3 : pred A) : + {subset s1 <= s2} -> + {subset s2 <= s3} -> + {subset s1 <= s3}. +Proof. by move=>S1 S2 z /S1/S2. Qed. + +Lemma subset_consL (A : eqType) x (s1 s2 : seq A) : + {subset x :: s1 <= s2} <-> + x \in s2 /\ {subset s1 <= s2}. +Proof. +split=>[S|[X S]]. +- by split=>[|z Z]; apply: S; rewrite inE ?eqxx ?Z ?orbT. +by move=>z; rewrite inE; case/orP=>[/eqP ->|/S//]. +Qed. + +Lemma subset_consLI (A : eqType) x (s1 s2 : seq A) : + x \in s2 -> + {subset s1 <= s2} -> + {subset x :: s1 <= s2}. +Proof. by move=>H1 H2; rewrite subset_consL. Qed. + +Lemma subset_consR (A : eqType) x (s : seq A) : + {subset s <= x :: s}. +Proof. by move=>z E; rewrite inE E orbT. Qed. + +Lemma subset_consLR (A : eqType) x (s1 s2 : seq A) : + {subset s1 <= s2} -> + {subset x :: s1 <= x :: s2}. +Proof. +move=>X z; rewrite !inE; case/orP=>[|/X] -> //. +by rewrite orbT. +Qed. + +(* disjointness *) Definition disjoint {A : eqType} (s1 s2 : seq A) := all (fun x => x \notin s1) s2. @@ -117,47 +155,12 @@ Lemma disjointN (A : eqType) (s1 s2 : seq A) : exists2 x, x \in s1 & x \in s2. Proof. by case/allPn=>x; rewrite negbK; exists x. Qed. -Lemma subset_consL (A : eqType) x (s1 s2 : seq A) : - {subset x :: s1 <= s2} <-> - x \in s2 /\ {subset s1 <= s2}. -Proof. -split=>[S|[X S]]. -- by split=>[|z Z]; apply: S; rewrite inE ?eqxx ?Z ?orbT. -by move=>z; rewrite inE; case/orP=>[/eqP ->|/S//]. -Qed. - -Lemma subset_consLI (A : eqType) x (s1 s2 : seq A) : - x \in s2 -> - {subset s1 <= s2} -> - {subset x :: s1 <= s2}. -Proof. by move=>H1 H2; rewrite subset_consL. Qed. - -Lemma subset_consR (A : eqType) x (s : seq A) : - {subset s <= x :: s}. -Proof. by move=>z E; rewrite inE E orbT. Qed. - -Lemma subset_consLR (A : eqType) x (s1 s2 : seq A) : - {subset s1 <= s2} -> - {subset x :: s1 <= x :: s2}. -Proof. -move=>X z; rewrite !inE; case/orP=>[|/X] -> //. -by rewrite orbT. -Qed. - -Hint Resolve subset_consR : core. - -Lemma subset_refl (A : Type) (s : pred A) : - {subset s <= s}. -Proof. by []. Qed. - -Lemma subset_trans (A : eqType) (s1 s2 s3 : pred A) : - {subset s1 <= s2} -> - {subset s2 <= s3} -> - {subset s1 <= s3}. -Proof. by move=>S1 S2 z /S1/S2. Qed. - - +(*************) +(*************) +(* Pregraphs *) +(*************) +(*************) (* Pregraphs are natmaps, mapping nodes into *) (* seq node listing the children of the node (adjacency list) *) @@ -207,37 +210,13 @@ Coercion Pred_of_history (x : pregraph) : Pred_Class := Notation "x &-> v" := (ptsT pregraph x v) (at level 30). -(* x is out-node if no edge goes into it *) -Definition out_node (g : pregraph) (x : node) := - all (fun s => x \notin s) (range g). - -(* pregraph is graph if valid, and *) -(* all nodes in all adjacency lists are good *) -Definition graph_axiom (g : pregraph) := - valid g && all (all (fun x : node => (x == null) || (x \in dom g))) (range g). +(* internal definition *) +Definition links (g : pregraph) x := oapp id [::] (find x g). -HB.mixin Record isGraph (g : pregraph) := { - graph_subproof : graph_axiom g}. -#[short(type=graph)] -HB.structure Definition Graph := {g of isGraph g }. - -(* removing out-node causes no dangling edges; *) -(* hence preserves graph axiom *) -Lemma graphF g x : - out_node g x -> - graph_axiom g -> - graph_axiom (free g x). -Proof. -move/allP=>/= Hna /andP [V /allP/= Ha]. -rewrite /graph_axiom validF V /=. -apply/allP=>/= s /rangeF Hs; apply/allP=>q Hq. -move/allP: (Ha _ Hs) (Hna _ Hs)=>/(_ _ Hq) {}Hs. -by rewrite domF inE; case: (x =P q) Hq=>//= ->->. -Qed. - -(* children of node x obtained by non-dangingle edges *) +(* children of node x are those obtained by *) +(* non-dangling edges *) Definition children (g : pregraph) x : seq node := - oapp (filter (mem (dom g))) [::] (find x g). + filter (mem (dom g)) (links g x). Lemma children_undef x : children undef x = [::]. Proof. by []. Qed. @@ -248,12 +227,12 @@ Proof. by []. Qed. Lemma childrenND g x : x \notin dom g -> children g x = [::]. -Proof. by rewrite /children/oapp; case: dom_find. Qed. +Proof. by rewrite /children/links/oapp; case: dom_find. Qed. Lemma childrenD g x : {subset children g x <= dom g}. Proof. -move=>y; rewrite /children/oapp/=. +move=>y; rewrite /children/links/oapp/=. case D : (find x g)=>[a|//]. by rewrite mem_filter; case/andP. Qed. @@ -262,7 +241,7 @@ Lemma childrenUnL g1 g2 x : valid (g1 \+ g2) -> {subset children g1 x <= children (g1 \+ g2) x}. Proof. -move=>V y; rewrite /children/oapp findUnL //. +move=>V y; rewrite /children/links/oapp findUnL //. case: dom_find=>// ys /In_find/In_dom /= D _. rewrite !mem_filter /= domUn inE V /=. by case/andP=>->->. @@ -273,6 +252,34 @@ Lemma childrenUnR g1 g2 x : {subset children g2 x <= children (g1 \+ g2) x}. Proof. by rewrite joinC; apply: childrenUnL. Qed. +Lemma size_links g x : + size (links g x) > 0 -> + x \in dom g. +Proof. by rewrite /links/oapp; case: dom_find. Qed. + +Lemma range_links (g : pregraph) x : + x \in dom g -> + links g x \in range g. +Proof. +rewrite /links/oapp. +by case: dom_find=>// xs /In_find/mem_range. +Qed. + +Lemma children_links (g : pregraph) x : + {subset children g x <= links g x}. +Proof. by move=>z; rewrite /children mem_filter=>/andP []. Qed. + +(* if x is node in g then g x contains all children of x *) +(* and maybe some more nodes that aren't in g *) +Lemma range_children (g : pregraph) x : + x \in dom g -> + exists2 xs, xs \in range g & + {subset children g x <= xs}. +Proof. +move=>Dx; exists (links g x); first by apply: range_links. +by apply: children_links. +Qed. + (* edge relation is just the applicative variant of children *) (* thus, consider removing one of them *) @@ -313,7 +320,7 @@ Lemma find_edge g x y ys : find x g = Some ys -> edge g x y = (y \in dom g) && (y \in ys). Proof. -rewrite /edge/children/oapp/= => ->. +rewrite /edge/children/links/oapp/= => ->. by rewrite mem_filter. Qed. @@ -477,11 +484,11 @@ Definition connect p (g : pregraph) := [rel x y | (x \in dom g) && (y \in component' p g x)]. Lemma connectP p (g : pregraph) x y : - reflect [/\ x \in dom g, ~~ p x & exists xs, + reflect [/\ x \in dom g, x \notin p & exists xs, [/\ path (edge g) x xs, y = last x xs & - {in xs, forall z, ~~ p z}]] + {in xs, forall z, z \notin p}]] (y \in connect p g x). -Proof. +Proof. rewrite /connect/component' inE mem_filter /= andbA. case: (boolP (x \in dom g))=>Dx; last by constructor; case. case: (boolP (p y))=>Py. @@ -503,6 +510,31 @@ rewrite disjoint_consR mem_filter negb_and Px /=. by apply/allP=>z /Pxs; rewrite mem_filter negb_and=>->. Qed. +(* there's path from x to y iff *) +(* there's path that doesn't loop through x *) +Lemma connectX p (g : pregraph) x y : + reflect [/\ x \in dom g, x \notin p & exists xs, + [/\ path (edge g) x xs, y = last x xs, + x \notin xs & + {in xs, forall z, z \notin p}]] + (y \in connect p g x). +Proof. +case: connectP=>H; constructor; last first. +- case=>Dx Hx [xs][Px Ey Nx Hxs]. + by apply: H; split=>//; exists xs. +case: H=>Dx Hx [xs][Px Ey Hxs]; split=>//. +(* if x doesn't appear in the path xs, then done *) +have [Mx|Nx] := boolP (x \in xs); last by exists xs. +(* path xs loops to x; find the last occurrence of x *) +(* and use the part of the path from that occurrence *) +case: {-1} _ _ _ / (splitLastP Mx) (erefl xs)=>/= {Mx} p1 p2 Nxp2 E. +rewrite {xs}E /= in Px Ey Hxs. +rewrite last_cat cat_path rcons_path last_rcons -andbA in Px Ey Hxs. +case/and3P: Px=>Px Ex Pp2; exists p2; split=>//. +by move=>z Z; rewrite Hxs // mem_cat Z orbT. +Qed. + + Lemma connect_undef p x : connect p undef x =i pred0. Proof. by move=>y; apply/connectP; case; rewrite dom_undef. Qed. @@ -513,12 +545,12 @@ Proof. by move=>y; apply/connectP; case; rewrite dom0. Qed. Lemma connectDx p g x y : y \in connect p g x -> - (x \in dom g) * ~~ p x. + (x \in dom g) * (x \notin p). Proof. by case/connectP. Qed. Lemma connectDy p g x y : y \in connect p g x -> - (y \in dom g) * ~~ p y. + (y \in dom g) * (y \notin p). Proof. case/connectP=>Dx Px [xs][P E Pxs]. have : y \in x :: xs by rewrite E mem_last. @@ -532,9 +564,9 @@ Lemma connectD p g x y : (x \in dom g) * (y \in dom g). Proof. by move=>C; rewrite (connectDx C) (connectDy C). Qed. -Lemma connectX p g x y : +Lemma connectDp p g x y : y \in connect p g x -> - (~~ p x) * (~~ p y). + (x \notin p) * (y \notin p). Proof. by move=>C; rewrite (connectDx C) (connectDy C). Qed. Lemma connectDN p g x : @@ -545,8 +577,8 @@ move=>Dx y; apply/negP=>/connectD []. by rewrite (negbTE Dx). Qed. -Lemma connectXN (p : pred node) g x : - p x -> +Lemma connectDNp (p : pred node) g x : + x \in p -> connect p g x =i pred0. Proof. move=>Hx y; apply/negP=>/connectX []. @@ -554,18 +586,18 @@ by rewrite Hx. Qed. Lemma connect0 p g x : - x \in connect p g x = (x \in dom g) && ~~ p x. + x \in connect p g x = (x \in dom g) && (x \notin p). Proof. by apply/connectP/andP; case=>// H1 H2; split=>//; exists [::]. Qed. Lemma connect0I p g x : x \in dom g -> - ~~ p x -> + x \notin p -> x \in connect p g x. Proof. by rewrite connect0=>->->. Qed. Lemma connect0N p g x y : x \in dom g -> - ~~ p x -> + x \notin p -> y \notin connect p g x -> x != y. Proof. by move=>Gx Px; apply: contra=>/eqP <-; rewrite connect0 Gx. Qed. @@ -607,10 +639,47 @@ apply/connectP; split=>//; first by apply: contra (S x) Hx. by exists xs; split=>// z /Hxs; apply/contra/S. Qed. -(* exteding connecting path from left *) +Lemma connect_eq g p1 p2 x : + p1 =i p2 -> + connect p1 g x =i connect p2 g x. +Proof. by move=>S y; apply/idP/idP; apply/connect_sub=>z; rewrite S. Qed. + + + + +(* deconstructing connecting path from left *) +Lemma edge_connect p g x y : + y != x -> + y \in connect p g x -> + exists2 z, edge g x z & y \in connect p g z. +Proof. +move/eqP=>Nxy /connectP [Dx H][[|a xs]][/= Px Ey Hxs]; first by move/Nxy: Ey. +case/andP: Px=>Px Pxs; exists a=>//; apply/connectP; split. +- by rewrite (edge_dom Px). +- by apply: Hxs; rewrite inE eqxx. +exists xs; split=>// z Z; apply: Hxs. +by rewrite inE Z orbT. +Qed. + +(* deconstructing connecting path from right *) +Lemma connect_edge p g x y : + y != x -> + y \in connect p g x -> + exists2 z, z \in connect p g x & edge g z y. +Proof. +move/eqP=>Nxy /connectP [Dx H][xs]. +case: {xs}(lastP xs)=>[|xs a][/= Px Ey Hxs]; first by move/Nxy: Ey. +rewrite last_rcons in Ey; rewrite -{a}Ey in Px Hxs. +rewrite rcons_path in Px; case/andP: Px=>Pxs Px. +exists (last x xs)=>//; apply/connectP; split=>//. +exists xs; split=>// z Z; apply: Hxs. +by rewrite mem_rcons inE Z orbT. +Qed. + +(* extending connecting path from left *) Lemma edge_connectI p g x y z : x \in dom g -> - ~~ p x -> + x \notin p -> edge g x y -> z \in connect p g y -> z \in connect p g x. @@ -621,11 +690,11 @@ split=>[||w] //=; first by rewrite H Py. by rewrite inE; case/orP=>[/eqP ->|/Hys]. Qed. -(* exteding connecting path from right *) +(* extending connecting path from right *) Lemma connect_edgeI p g x y z : y \in connect p g x -> edge g y z -> - ~~ p z -> + z \notin p -> z \in connect p g x. Proof. case/connectP=>Dx Hx [xs][Px Ey Hxs] H Hz. @@ -636,317 +705,233 @@ exists (rcons xs z); split=>/=. by move=>w; rewrite mem_rcons inE; case/orP=>[/eqP ->|/Hxs]. Qed. -(* destructing connecting path from left *) -Lemma edge_connectE p g x y : - y \in connect p g x -> - y = x \/ - exists2 z, edge g x z & y \in connect p g z. -Proof. -case/connectP=>Dx H [[|a xs]][/= Px Ey Hxs]; first by left. -right; case/andP: Px=>Px Pxs; exists a=>//. -apply/connectP; split. -- by rewrite (edge_dom Px). -- by apply: Hxs; rewrite inE eqxx. -exists xs; split=>// z Z; apply: Hxs. -by rewrite inE Z orbT. -Qed. +(**********) +(**********) +(* Graphs *) +(**********) +(**********) -(* deconstructing connecting path from right *) -Lemma connect_edgeE p g x y : - y \in connect p g x -> - y = x \/ - exists2 z, z \in connect p g x & edge g z y. +(* x is in_node if it's in the graph or is null *) +Definition in_node (g : pregraph) (x : node) := + (x == null) || (x \in dom g). + +(* x is out-node if no edge goes into it *) +Definition out_node (g : pregraph) (x : node) := + {in range g, forall xs, x \notin xs}. + +(* pregraph is graph if *) +(* all nodes in all adjacency lists are in-nodes *) +Definition graph_axiom (g : pregraph) := + forall (xs : seq node) (x : node), xs \in range g -> x \in xs -> in_node g x. + +HB.mixin Record isGraph (g : pregraph) := { + graph_subproof : graph_axiom g}. +#[short(type=graph)] +HB.structure Definition Graph := {g of isGraph g }. + +(* removing out-node causes no dangling edges; *) +(* hence preserves graph axiom *) +Lemma graphF g x : + out_node g x -> + graph_axiom g -> + graph_axiom (free g x). Proof. -case/connectP=>Dx H [xs]. -case: {xs}(lastP xs)=>[|xs a][/= Px Ey Hxs]; first by left. -right; rewrite last_rcons in Ey; rewrite -{a}Ey in Px Hxs. -rewrite rcons_path in Px; case/andP: Px=>Pxs Px. -exists (last x xs)=>//; apply/connectP; split=>//. -exists xs; split=>// z Z; apply: Hxs. -by rewrite mem_rcons inE Z orbT. +move=>Hna Ha xs q /rangeF Hs Hq. +move: (Ha xs q Hs Hq) (Hna _ Hs)=>{}Hs. +rewrite /in_node domF inE in Hs *. +by case: (x =P q) Hq=>//= ->->. Qed. (* part of g reachable from x (avoiding nothing) *) Definition subconnect g x : pregraph := um_filterk (connect pred0 g x) g. -(* part of g unreachable from x (avoiding nothing) *) -Definition subdisconnect g x : pregraph := - um_filterk (negb \o connect pred0 g x) g. - -Lemma connect_split g x : - g = subconnect g x \+ subdisconnect g x. -Proof. by apply: umfilt_predC. Qed. - (* reachable component of a graph is a graph *) Lemma graph_subconnect g x : + valid g -> graph_axiom g -> graph_axiom (subconnect g x). Proof. -have E := connect_split g x. -case/andP=>V Ha; apply/andP; split; first by rewrite pfV. -apply/allP=>/= xs Hxs; move/allP: Ha=>/(_ xs). -rewrite {1}E rangeUn inE -E V Hxs /= => /(_ erefl) Ha. -case/mem_rangeX: Hxs=>k /In_umfiltX [/= Hk] /In_find Hf. -apply/allP=>n Hn; move/allP: Ha=>/(_ _ Hn). -case/boolP: (n == null)=>//= Hnn. -rewrite /subconnect dom_umfiltk inE /= => En. -rewrite En andbT /=; apply: (connect_trans Hk). -apply/connectP; split=>//=; first by rewrite (connectD Hk). -exists [:: n]; split=>//=. -by rewrite (find_edge _ Hf) En Hn. +have E : g = subconnect g x \+ um_filterk + (negb \o connect pred0 g x) g by apply: umfilt_predC. +move=>V Ha xs /= n Hxs Hn; have {}Dn : in_node g n. +- by apply: Ha Hn; rewrite E rangeUn inE -E V Hxs. +case/mem_rangeX: Hxs=>k /In_umfiltX [/= Ck] /In_find Hf. +rewrite /in_node in Dn *; case/boolP: (n == null) Dn=>//= Hnn Dn. +case: (connectD Ck)=>_ Dk. +rewrite /subconnect dom_umfiltk inE /= Dn andbT. +apply: connect_trans Ck _; apply/connectP; split=>//. +by exists [:: n]; split=>//=; rewrite (find_edge _ Hf) Dn Hn. Qed. +Lemma edge_subconnect g x y z : + edge (subconnect g x) y z = + [&& y \in connect pred0 g x, + z \in connect pred0 g x & + edge g y z]. +Proof. +rewrite /edge/children/links/oapp/= find_umfiltk topredE. +case: ifP=>// _; case: (find y g)=>[ys|]; last by rewrite andbF. +by rewrite !mem_filter /= dom_umfiltk inE -andbA. +Qed. -UP TO HERE - - - - - - +(* connectivity relation out of x in a connected subcomponent *) +(* is that same connectivity relation out of x in the graph *) +Lemma connect_subconnect g x : + connect pred0 (subconnect g x) x =i + connect pred0 g x. +Proof. +move=>w; apply/iffE; split; case/connectP. +- rewrite dom_umfiltk=>/andP [Cx /= Dx] _ [xs][Px Ew _]. + apply/connectP; split=>//; exists xs; split=>//. + by apply/sub_path/Px=>y z; rewrite edge_subconnect=>/and3P []. +move=>Dx _ [xs][Px Ew _]; apply/connectP; split=>//. +- by rewrite dom_umfiltk inE topredE connect0 Dx. +exists xs; split=>//; apply/(sub_in_path (P:=connect pred0 g x))/Px. +- by move=>y z Cy Cz E; rewrite edge_subconnect Cy Cz E. +apply/allP=>z; rewrite inE; case/orP=>[/eqP ->|Hz]. +- by rewrite topredE connect0 Dx. +case/splitPr: Hz Px=>xs1 xs2. +rewrite -cat1s catA cats1 cat_path=>/andP [Px _]. +apply/connectP; split=>//. +by exists (rcons xs1 z); rewrite last_rcons. +Qed. -(*******************) +(****************************) +(* Avoidance sets (marking) *) +(****************************) -(* TODO generalize to arbitrary boundary? *) -Definition subconnect p g := um_filterk (fun z => z \in connect Unit g p) g. +(* deconstructing connecting path from left *) +(* strenghtening avoidance *) +Lemma edge_connectX g p (x y : node) : + y != x -> + y \in connect p g x -> + exists2 z, + edge g x z & y \in connect [predU pred1 x & p] g z. +Proof. +move/eqP=>Nxy /connectX [Dx Hx][[|a xs]] /= [Px Ey Nx Hxs]. +- by move/Nxy: Ey. +case/andP: Px=>Exa Pa; exists a=>//; apply/connectP; split. +- by rewrite (edge_dom Exa). +- rewrite !inE negb_or Hxs ?inE ?eqxx // andbT. + by case: (a =P x) Nx=>//= ->; rewrite inE eqxx. +exists xs; split=>// z Z. +rewrite !inE negb_or Hxs ?inE ?Z ?orbT // andbT. +by case: (z =P x) Nx=>// <-; rewrite inE Z orbT. +Qed. -Definition subdisconnect p g := um_filterk (fun z => z \notin connect Unit g p) g. +(* extending connecting path from left *) +(* weakening avoidance *) +Lemma edge_connectXI p g x y z : + x \in dom g -> + x \notin p -> + edge g x y -> + z \in connect [predU pred1 x & p] g y -> + z \in connect p g x. +Proof. +move=>Dx Px Ex /(connect_sub (p2:=p)) H. +apply: edge_connectI Ex (H _)=>//. +by move=>w W; rewrite inE /= W orbT. +Qed. -Lemma connect_split p g : g = subconnect p g \+ subdisconnect p g. -Proof. by apply: umfilt_predC. Qed. +(* the right-way lemmas can't be stated usefully *) -Lemma good_subconnect p g : - good_graph g -> good_graph (subconnect p g). +(* equivalence variant *) +Lemma edge_connectXE p g x y : + y \in connect p g x <-> + [/\ x \in dom g, x \notin p & y = x \/ exists2 z, + edge g x z & y \in connect [predU pred1 x & p] g z]. Proof. -have E := connect_split p g. -case/andP=>V Ha; apply/andP; split; first by rewrite valid_umfilt. -apply/allP=>ns Hns; move/allP: Ha=>/(_ ns). -rewrite {1}E rangeUn inE -E V Hns /= => /(_ erefl) Ha. -case/mem_rangeX: Hns=>k; case/In_umfilt=>/= Hk /In_find Hf. -apply/allP=>n Hn; move/allP: Ha=>/(_ _ Hn). -rewrite /good_ptr; case/boolP: (n == null)=>//= Hnn. -rewrite /subconnect dom_umfiltk inE /= =>En; rewrite En andbT. -apply: (connect_trans Hk); apply/connectP. -exists [::n]; split=>//=; last by rewrite !dom0. -by rewrite andbT (find_edge _ Hf) En. +split=>[C|[Dx Hx]]; last first. +- case=>[->|[z E C]]; first by rewrite connect0 Dx Hx. + by apply: edge_connectXI E C. +rewrite !(connectDx C); split=>//. +case: (y =P x)=>[->|/eqP Nxy]; first by left. +by right; apply: edge_connectX Nxy C. Qed. -Lemma edge_subconnect p g : subrel (edge (subconnect p g)) (edge g). +(* if y is reachable from x, but not from x', then *) +(* y is reachable from x by a path that avoids whole *) +(* subcomponent of x' *) +Lemma connect_avoid p g x y x' : + y \in connect p g x -> + y \notin connect p g x' -> + y \in connect [predU p & connect p g x'] g x. Proof. -move=>x y; rewrite /edge /links find_umfiltk; case: ifP=>//= Hx. -by case Ef: (find x g)=>[v|] //=; rewrite !mem_filter dom_umfiltk inE /= -andbA; case/andP. +case/connectP=>Dx Hx [xs][Px Ey Hxs] Ny. +have [Nx|Mx] := boolP (has (connect p g x') (x :: xs)); last first. +(* if path contains no nodes reachable from x', nothing to do *) +- move/hasPn: Mx=>Mx; apply/connectP; split=>//. + - by rewrite inE negb_or Hx Mx // inE eqxx. + by exists xs; split=>// z Z; rewrite inE negb_or Hxs // Mx // inE Z orbT. +(* if path contains node reachable from x' *) +(* let a be the last such node in the path *) +case: {-1} _ _ _ / (split_findlast Nx) (erefl (x::xs)). +move=>{Nx} a p1 p2; rewrite topredE=>Ca /hasPn Nx' X. +(* suffices to show that y is reachable from a *) +(* because then y is also reachable from x'; contradiction *) +suff Cy : y \in connect p g a. +- by move: (connect_trans Ca Cy) Ny; rewrite -!topredE /= =>->. +case: p1 X Ey Px Hxs=>[[<- _]|b p1 [_ ->]] /= Ey Px Hxs. +- by apply/connectP; split=>//; exists xs. +rewrite cat_path last_cat last_rcons rcons_path -andbA in Ey Px. +case/and3P: Px=>Px Ea P2; apply/connectP; rewrite !(connectDy Ca); split=>//. +by exists p2; split=>// z Z; rewrite Hxs // mem_cat Z orbT. Qed. -Lemma connect_subconnect p g : connect Unit (subconnect p g) p =i connect Unit g p. -Proof. -move=>z; apply/iffE; split. -- case/connectP=>xs [Hxs {z}-> /implyP Nxs _]. - apply/connectP; exists xs; split=>//=; last by rewrite dom0 /= all_predT. - - by apply/sub_path/Hxs/edge_subconnect. - by apply/implyP=>/Nxs; rewrite dom_umfiltk inE /=; case/andP. -case/connectP=>xs [Hxs {z}-> /implyP Nxs _]. -have Hpd : p \in dom g. -- case: xs Hxs Nxs=>/=; first by move=>_; apply. - by move=>h t; case/andP; case/edge_dom. -apply/connectP; exists xs; split=>//=; last by rewrite dom0 /= all_predT. -- apply/(sub_in_path (P:=fun q => q \in connect Unit g p))/Hxs. - - move=>x y; rewrite !inE=>Hx Hy. - rewrite /edge /links find_umfiltk Hx. - case Ef: (find x g)=>[v|] //=; rewrite !mem_filter dom_umfiltk inE /= -andbA =>->. - by rewrite andbT. - move=>/=; apply/andP; split; first by apply/connect0=>//; rewrite dom0. - apply/allP=>x Hx; case/splitPr: Hx Hxs Nxs=>xs1 xs2. - rewrite -cat1s catA cats1 cat_path /=; case/andP=>H1 H2 _. - apply/connectP; exists (rcons xs1 x); split=>//=. - - by rewrite last_rcons. - - by rewrite -cats1 cat_nilp andbF. - by rewrite dom0 /= all_predT. -apply/implyP=>/Nxs; rewrite dom_umfiltk inE /==>->; rewrite andbT. -by apply/connect0=>//; rewrite dom0. -Qed. +(***************) +(* N-pregraphs *) +(***************) -End GraphOps. +(* pregraphs where there is global bound n *) +(* on the number of links that each node may have *) -Section Marking. +Definition n_pregraph_axiom (n : nat) (g : pregraph) := + {in range g, forall xs, size xs = n}. -Lemma connectMPtUn p m g x (cs : seq ptr) : - valid m -> - p \notin dom m -> - find p g = Some cs -> - forall z, z != p -> - x \in cs -> z \in connect m g x -> - exists2 y, y \in cs & z \in connect (#p \+ m) g y. -Proof. -move=>Vm Npm Hc z Hzp Hx; case/connectP=>xs [Hp Ez Nxs] Ha. -rewrite {z}Ez in Hzp *. -case Hpxs: (p \in x::xs); last first. -- (* there's no loop, path stays the same *) - exists x=>//; apply/connectP; exists xs; split=>//. - apply/allP=>z Hz; rewrite domPtUn inE validPtUn !negb_and negb_or /= negbK Vm (negbTE Npm) /=. - move/allP: Ha=>/(_ _ Hz) ->; rewrite andbT. - by apply/negP=>/eqP E; rewrite E Hz in Hpxs. -(* there's at least one loop, find the last p *) -case: {-1}_ _ _ / (splitLastP Hpxs) (erefl (x::xs)) =>{Hpxs} xs1 xs2 Hxs2. -case: xs2 Hxs2=>/=. -- move=>_; rewrite cats0=>E. - have /=: last x (x :: xs) = last x (rcons xs1 p) by rewrite E. - by rewrite last_rcons=>{}E; rewrite E eq_refl in Hzp. -move=>ch xs2; rewrite inE negb_or; case/andP => Npch Nxs2 E. -case: xs1 E=>/=. -- (* p links to itself *) - case=>Exp E2; rewrite {x Nxs}Exp {xs}E2 /= in Hx Hp Ha Hzp *. - case/andP: Hp=>He Hp. - exists ch; first by move: He; rewrite (find_edge _ Hc); case/andP. - apply/connectP; exists xs2; split=>//. - - by case/edge_dom: He=>_ ->; exact: implybT. - case/and3P: Ha=>_ Hch Ha. - apply/allP=>z; rewrite domPtUn inE validPtUn !negb_and negb_or /= negbK Vm (negbTE Npm) /=. - case/orP=>[/eqP {z}->| Hz]; first by rewrite Npch. - move/allP: Ha=>/(_ _ Hz) ->; rewrite andbT. - by apply/negP=>/eqP E; rewrite E Hz in Nxs2. -(* there's a non-trivial path before the loop *) -move=>_ xs1 [_ Exs]; rewrite {xs}Exs in Hp Nxs Ha Hzp *. -move: Hp; rewrite cat_path last_cat !last_rcons rcons_path -andbA /=. -case/and4P=>_ _ He Hp2. -exists ch; first by move: He; rewrite (find_edge _ Hc); case/andP. -apply/connectP; exists xs2; split=>//. -- by case/edge_dom: He=>_ ->; exact: implybT. -move: Ha; rewrite -cat_cons all_cat; case/andP=>_ Ha. -apply/allP=>z Hz; rewrite domPtUn inE validPtUn !negb_and negb_or /= negbK Vm (negbTE Npm) /=. -move/allP: Ha=>/(_ _ Hz) ->; rewrite andbT. -move: Hz; rewrite inE; case/orP=>[/eqP->|Hz] //. -by apply/negP=>/eqP E; rewrite E Hz in Nxs2. -Qed. +Lemma npregraphUnL n g1 g2 : + valid (g1 \+ g2) -> + n_pregraph_axiom n (g1 \+ g2) -> + n_pregraph_axiom n g1. +Proof. by move=>V H x R; apply: H; rewrite rangeUn inE V R. Qed. -Lemma connectMUnSub m1 m2 g x : - valid (m1 \+ m2) -> - forall z, z \in connect (m1 \+ m2) g x -> - z \in connect m2 g x. -Proof. -move=>Vm z; case/connectP=>xs [Hp Ez Nxs] Ha. -apply/connectP; exists xs; split=>//. -apply/sub_all/Ha=>q; apply: contra. -rewrite domUn inE Vm => ->. -by rewrite orbT. -Qed. +Lemma npregraphUnR n g1 g2 : + valid (g1 \+ g2) -> + n_pregraph_axiom n (g1 \+ g2) -> + n_pregraph_axiom n g2. +Proof. by rewrite joinC; apply: npregraphUnL. Qed. -Corollary connectMPtUnHas p m g cs : - valid m -> - p \notin dom m -> - find p g = Some cs -> - forall z, z \in connect m g p = (z == p) || has (fun x => z \in connect (#p \+ m) g x) cs. -Proof. -move=>Vm Npm Hc z. -rewrite (connect_eq_links Hc) //; case: eqP=>/= // /eqP Hzp. -have Vpm : valid (#p \+ m) by rewrite validPtUn Vm. -apply/iffE; split; case/hasP=>x Hx. -- move=>Hz; apply/hasP. - by apply: (connectMPtUn (x:=x)). -move/(connectMUnSub Vpm)=>Hz. -by apply/hasP; exists x. -Qed. +Lemma npregraphF n g x : + n_pregraph_axiom n g -> + n_pregraph_axiom n (free g x). +Proof. by move=>H z /rangeF; apply: H. Qed. -Lemma connectMUn p m1 m2 g x c (cs : seq ptr) : - valid (m2 \+ m1) -> - find p g = Some cs -> - c \in cs -> good_ptr g c -> - dom m2 =i connect m1 g c -> - forall z, x \in cs -> z \in connect m1 g x -> - z \in connect m1 g c - \/ - exists2 y, y \in filter (predC1 c) cs & z \in connect (m2 \+ m1) g y. -Proof. -move=>V21 Hf Hc Hcd Hm z Hxc. case/connectP=>xs [Hp Ez Nxs] Ha. -case/boolP: (all (fun z => z \notin dom m2) (x::xs))=>Hpxs. -- (* path doesn't go through the marked component *) - right; exists x. - - rewrite mem_filter Hxc andbT /=. - case/orP: Hcd=>[/eqP->|Hcd]. - - suff: x \in dom g by move/dom_cond. - by case: {Ez Ha Hpxs}xs Hp Nxs=>//= h t; case/andP; case/edge_dom. - apply/negP=>/eqP Exc; move: Hpxs=>/=; case/andP=>+ _; rewrite Exc Hm. - move/negP; apply; apply: connect0=>//. - by move: Ha=>/=; case/andP=>+ _; rewrite Exc. - apply/connectP; exists xs; split=>//. - apply/allP=>q Hq; rewrite domUn inE negb_and negb_or V21 /=. - by move/allP: Ha=>/(_ _ Hq)->; move/allP: Hpxs=>/(_ _ Hq)->. -(* path goes through the marked component, so z is connected to c *) -left; rewrite -has_predC (eq_has (a2:=fun z=> z \in dom m2)) in Hpxs; last first. -- by move=>q /=; rewrite negbK. -(* q is the last vertex in marked component, xs2 is the free path *) -case: {-1}_ _ _ / (split_findlast Hpxs) (erefl (x::xs))=>{Hpxs} q xs1 xs2 Hq. -rewrite -all_predC (eq_all (a2:=fun z=> z \notin dom m2)) // => Hxs2 Heq. -apply: (connect_trans (y:=q)); rewrite app_predE; first by rewrite -Hm. -case: xs1 Heq=>/=. -- case=>Eq Exs2; rewrite -{q}Eq in Hq *; rewrite -{xs2}Exs2 in Hxs2. - by apply/connectP; exists xs. -move=>_ t [_ Exs]. -apply/connectP; exists xs2; split=>//. -- by move: Hp; rewrite Exs cat_path last_rcons; case/andP. -- by move: Ez; rewrite Exs last_cat last_rcons. -- suff: (edge g (last x t) q) by case/edge_dom=>_ ->; apply: implybT. - rewrite Exs cat_path rcons_path in Hp. - by case/andP: Hp; case/andP. -rewrite Exs -cats1 -catA cat1s -cat_cons all_cat in Ha. -by case/andP: Ha. -Qed. +Definition get_nth (g : pregraph) x := nth null (links g x). -Corollary connectMUnHas p m1 m2 g c (cs : seq ptr) : - valid (m2 \+ m1) -> - find p g = Some cs -> - c \in cs -> good_ptr g c -> - dom m2 =i connect m1 g c -> - forall z, - has (fun x => z \in connect m1 g x) cs = - (z \in connect m1 g c) || - has (fun x => z \in connect (m2 \+ m1) g x) (filter (predC1 c) cs). +Lemma in_node_nth g x n : + graph_axiom g -> + in_node g (get_nth g x n). Proof. -move=>V Hf Hc Hcd Hm z; apply/iffE; split. -- case/hasP=>x Hx Hz. - case: (connectMUn V Hf Hc Hcd Hm Hx Hz); first by move=>->. - by case=>y Hy1 Hy2; apply/orP; right; apply/hasP; exists y. -case/orP; first by move=>Hz; apply/hasP; exists c. -case/hasP=>q; rewrite mem_filter /=; case/andP=>_ Hq /(connectMUnSub V) Hz. -by apply/hasP; exists q. +rewrite /graph_axiom => Hg. +case: (ltnP n (size (links g x)))=>Hm; last first. +- by rewrite /in_node/get_nth nth_default. +have Dx : x \in dom g by apply/size_links/gt0/Hm. +by apply: Hg (range_links Dx) _; rewrite mem_nth. Qed. -End Marking. -Section NGraph. -Definition n_graph (n : nat) (g : pregraph) : bool := - all (fun ns => size ns == n) (range g). -Lemma n_graphUnL n g0 g : - valid (g \+ g0) -> - n_graph n (g \+ g0) -> n_graph n g. -Proof. -by move=>V; apply/subset_all=>z; rewrite rangeUn inE V=>->. -Qed. -(* corollary? *) -Lemma n_graphF n g p : n_graph n g -> n_graph n (free g p). -Proof. by apply/subset_all/rangeF. Qed. -Definition get_nth (ps : seq ptr) (n : nat) : ptr := - nth null ps n. -Lemma all_good_get p g ps m : - good_graph g -> - find p g = Some ps -> - good_ptr g (get_nth ps m). -Proof. -case/andP=>_ Hg Hf. -case: (ltnP m (size ps))=>Hm; last first. -- by apply/orP; left; rewrite /get_nth; rewrite nth_default. -have /allP : all (good_ptr g) ps. -- move/allP: Hg; apply. - apply/mem_rangeX; exists p. - by move/In_find: Hf. -by apply; apply: mem_nth. -Qed. +(*******************) + +Section NGraph. + + + Lemma all_nth n g : n_graph n g -> From 1058fcf06830bd87504c07357b56ad862e6b00fa Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Wed, 7 Aug 2024 21:06:37 +0200 Subject: [PATCH 16/93] blah --- examples/graph.v | 145 +++++++++++++++++++++++++++++------------------ 1 file changed, 89 insertions(+), 56 deletions(-) diff --git a/examples/graph.v b/examples/graph.v index f14d604..6268992 100644 --- a/examples/graph.v +++ b/examples/graph.v @@ -210,11 +210,61 @@ Coercion Pred_of_history (x : pregraph) : Pred_Class := Notation "x &-> v" := (ptsT pregraph x v) (at level 30). -(* internal definition *) +(* links of x includes all outgoing edges, dangling or not *) Definition links (g : pregraph) x := oapp id [::] (find x g). -(* children of node x are those obtained by *) -(* non-dangling edges *) +Lemma links_undef x : links undef x = [::]. +Proof. by []. Qed. + +Lemma links_unit x : links Unit x = [::]. +Proof. by []. Qed. + +Lemma linksND g x : + x \notin dom g -> + links g x = [::]. +Proof. by rewrite /links/oapp; case: dom_find. Qed. + +Lemma linksUnL g1 g2 x : + valid (g1 \+ g2) -> + links (g1 \+ g2) x = + if x \in dom g1 then links g1 x else links g2 x. +Proof. by move=>V; rewrite /links/oapp findUnL //; case: dom_find. Qed. + +Lemma linksUnR g1 g2 x : + valid (g1 \+ g2) -> + links (g1 \+ g2) x = + if x \in dom g2 then links g2 x else links g1 x. +Proof. by rewrite joinC=>/linksUnL; apply. Qed. + +Lemma size_links g x : + size (links g x) > 0 -> + x \in dom g. +Proof. by rewrite /links/oapp; case: dom_find. Qed. + +Lemma linksD (g : pregraph) x y : + y \in links g x -> + x \in dom g. +Proof. by move=>X; apply: size_links; case: (links g x) X. Qed. + +Lemma In_graph (g : pregraph) x xs : + (x, xs) \In g -> + xs = links g x. +Proof. by rewrite /links/oapp=>/In_find ->. Qed. + +Lemma In_graphX (g : pregraph) x : + x \in dom g -> + (x, links g x) \In g. +Proof. +by move=>Dx; apply/In_find; rewrite /links/oapp; case: dom_find Dx. +Qed. + +Lemma range_links (g : pregraph) x : + x \in dom g -> + links g x \in range g. +Proof. by move/In_graphX/mem_range. Qed. + +(* children x are links obtained by non-dangling edges only *) + Definition children (g : pregraph) x : seq node := filter (mem (dom g)) (links g x). @@ -227,24 +277,18 @@ Proof. by []. Qed. Lemma childrenND g x : x \notin dom g -> children g x = [::]. -Proof. by rewrite /children/links/oapp; case: dom_find. Qed. +Proof. by rewrite /children=>/linksND ->. Qed. Lemma childrenD g x : {subset children g x <= dom g}. -Proof. -move=>y; rewrite /children/links/oapp/=. -case D : (find x g)=>[a|//]. -by rewrite mem_filter; case/andP. -Qed. +Proof. by move=>y; rewrite /children mem_filter; case/andP. Qed. Lemma childrenUnL g1 g2 x : valid (g1 \+ g2) -> {subset children g1 x <= children (g1 \+ g2) x}. Proof. -move=>V y; rewrite /children/links/oapp findUnL //. -case: dom_find=>// ys /In_find/In_dom /= D _. -rewrite !mem_filter /= domUn inE V /=. -by case/andP=>->->. +move=>V y; rewrite /children !mem_filter /= =>/andP [Dy Ly]. +by rewrite domUn inE V Dy linksUnL //= (linksD Ly). Qed. Lemma childrenUnR g1 g2 x : @@ -252,19 +296,6 @@ Lemma childrenUnR g1 g2 x : {subset children g2 x <= children (g1 \+ g2) x}. Proof. by rewrite joinC; apply: childrenUnL. Qed. -Lemma size_links g x : - size (links g x) > 0 -> - x \in dom g. -Proof. by rewrite /links/oapp; case: dom_find. Qed. - -Lemma range_links (g : pregraph) x : - x \in dom g -> - links g x \in range g. -Proof. -rewrite /links/oapp. -by case: dom_find=>// xs /In_find/mem_range. -Qed. - Lemma children_links (g : pregraph) x : {subset children g x <= links g x}. Proof. by move=>z; rewrite /children mem_filter=>/andP []. Qed. @@ -280,8 +311,9 @@ move=>Dx; exists (links g x); first by apply: range_links. by apply: children_links. Qed. -(* edge relation is just the applicative variant of children *) -(* thus, consider removing one of them *) +(* edge is applicative variant of children *) +(* NOTE: do we want applicative variant of links *) +(* that computes external edges? *) Definition edge g : rel node := mem \o children g. Arguments edge g x y : simpl never. @@ -334,8 +366,10 @@ Qed. (* lifting the theory of finite graphs to unbounded pregraphs *) -(* list of nodes traversed by depth-first search of g *) +(* list of nodes traversed by depth-first search of g *) (* at depth n, starting from x, and avoiding v *) +(* NOTE: uses children, not links, to traverse internal edges only; *) +(* traversing external edges doesn't seem very sensible *) Fixpoint dfs (g : pregraph) (n : nat) (v : seq node) x := if (x \notin dom g) || (x \in v) then v else if n is n'.+1 then foldl (dfs g n') (x :: v) (children g x) else v. @@ -884,17 +918,27 @@ Qed. (* N-pregraphs *) (***************) -(* pregraphs where there is global bound n *) -(* on the number of links that each node may have *) +(* pregraphs with global bound n *) +(* for the number of links of a node *) Definition n_pregraph_axiom (n : nat) (g : pregraph) := - {in range g, forall xs, size xs = n}. + {in dom g, forall x, size (links g x) = n}. + +Lemma npregraphP n g : + n_pregraph_axiom n g <-> + {in range g, forall xs, size xs = n}. +Proof. +split=>H x; last by move=>Dx; apply: H (range_links Dx). +by case/mem_rangeX=>k X; rewrite (In_graph X); apply: H (In_dom X). +Qed. Lemma npregraphUnL n g1 g2 : valid (g1 \+ g2) -> n_pregraph_axiom n (g1 \+ g2) -> n_pregraph_axiom n g1. -Proof. by move=>V H x R; apply: H; rewrite rangeUn inE V R. Qed. +Proof. +by rewrite !npregraphP=>V H x X; apply: H; rewrite rangeUn inE V X. +Qed. Lemma npregraphUnR n g1 g2 : valid (g1 \+ g2) -> @@ -905,11 +949,11 @@ Proof. by rewrite joinC; apply: npregraphUnL. Qed. Lemma npregraphF n g x : n_pregraph_axiom n g -> n_pregraph_axiom n (free g x). -Proof. by move=>H z /rangeF; apply: H. Qed. +Proof. by rewrite !npregraphP=>H z /rangeF; apply: H. Qed. -Definition get_nth (g : pregraph) x := nth null (links g x). +Definition get_nth g x := nth null (links g x). -Lemma in_node_nth g x n : +Lemma innode_nth g x n : graph_axiom g -> in_node g (get_nth g x n). Proof. @@ -920,32 +964,21 @@ have Dx : x \in dom g by apply/size_links/gt0/Hm. by apply: Hg (range_links Dx) _; rewrite mem_nth. Qed. - - +Lemma links_nth n g x : + n_pregraph_axiom n g -> + x \in dom g -> + links g x = map (get_nth g x) (iota 0 n). +Proof. +move=>H Dx; apply/(eq_from_nth (x0:=null))=>[|i Hi]. +- by rewrite size_map size_iota H. +by rewrite map_nth_iota0 ?H // -(H _ Dx) take_size. +Qed. (*******************) -Section NGraph. - - - - -Lemma all_nth n g : - n_graph n g -> - all (fun ns => ns == map (get_nth ns) (iota 0 n)) (range g). -Proof. -move=>H; apply/sub_all/H=>ns /eqP Hns. -apply/eqP/(eq_from_nth (x0:=null)). -- by rewrite size_map size_iota. -by rewrite Hns=>i Hi; rewrite map_nth_iota0 -Hns // take_size. -Qed. - -End NGraph. - - Section Acyclic. Definition symconnect p g x y := connect p g x y && connect p g y x. From 2394b9f48a896a1434bf82a194d241c24b323000 Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Wed, 7 Aug 2024 21:27:30 +0200 Subject: [PATCH 17/93] blah --- examples/graph.v | 13 ++++++++----- 1 file changed, 8 insertions(+), 5 deletions(-) diff --git a/examples/graph.v b/examples/graph.v index 6268992..63f1d78 100644 --- a/examples/graph.v +++ b/examples/graph.v @@ -368,8 +368,11 @@ Qed. (* list of nodes traversed by depth-first search of g *) (* at depth n, starting from x, and avoiding v *) -(* NOTE: uses children, not links, to traverse internal edges only; *) -(* traversing external edges doesn't seem very sensible *) +(* DEVCOMMENT: *) +(* uses children, not links, to traverse internal edges only *) +(* traversing an external edge may be sensible, *) +(* as then one can reason about connectivity to external nodes *) +(* but it will require a completely different reflection theorem *) Fixpoint dfs (g : pregraph) (n : nat) (v : seq node) x := if (x \notin dom g) || (x \in v) then v else if n is n'.+1 then foldl (dfs g n') (x :: v) (children g x) else v. @@ -452,9 +455,9 @@ Lemma dfs_pathP g n x y v : reflect (dfs_path g v x y) (y \in dfs g n v x). Proof. elim: n=>[|n IHn] /= in x y v * => Hv Uv Sv Ny Dx. -- rewrite addn0 in Hv; case: (uniq_min_size Uv Sv Hv) Ny=>_ Ev /negbTE Ny. - suff: ~ dfs_path g v x y by rewrite Dx if_same Ny; apply: ReflectF. - by case=>xs E _; rewrite disjoint_consR Ev Dx. +- rewrite addn0 in Hv; rewrite Dx if_same (negbTE Ny). + apply: ReflectF; case=>xs E _; rewrite disjoint_consR. + by rewrite (uniq_min_size Uv Sv Hv) Dx. rewrite Dx /=; have [Vx|Vx] := ifPn. - by rewrite (negbTE Ny); apply: ReflectF=>[[xs]]; rewrite disjoint_consR Vx. set v1 := x :: v; set a := children g x; have [->|/eqP Nyx] := eqVneq y x. From 79baaaf11615659b84fd740d0e033bad82d8f33f Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Wed, 7 Aug 2024 22:03:21 +0200 Subject: [PATCH 18/93] comments about do we want to follow dangling edges to their dangling target node or not. i think not. --- examples/graph.v | 37 +++++++++++++++++++++++++------------ 1 file changed, 25 insertions(+), 12 deletions(-) diff --git a/examples/graph.v b/examples/graph.v index 63f1d78..292f2d5 100644 --- a/examples/graph.v +++ b/examples/graph.v @@ -164,10 +164,10 @@ Proof. by case/allPn=>x; rewrite negbK; exists x. Qed. (* Pregraphs are natmaps, mapping nodes into *) (* seq node listing the children of the node (adjacency list) *) -(* Technically pregraph's a "prequiver", because this representation *) -(* allows loops and parallel edges *) (* Pregraph differs from graph, in that edges can *dangle*; that is *) -(* originate or terminate (or both) with a node that isn't in the graph *) +(* terminate with a node that isn't in the graph. *) +(* Here, dangling edges are treated as non-existent; *) +(* that is, target of a dangling edge is treated as null pointer. *) Notation node := nat. Record pregraph := Pregraph {pregraph_base : @UM.base node nat_pred (seq node)}. @@ -210,7 +210,16 @@ Coercion Pred_of_history (x : pregraph) : Pred_Class := Notation "x &-> v" := (ptsT pregraph x v) (at level 30). -(* links of x includes all outgoing edges, dangling or not *) +(* Links of x includes all edges outgoing from x *) +(* and may explicitly include nodes that aren't in dom g. *) +(* These are so called "dangling edges". *) +(* Dangling edges are allowed to enable positional information *) +(* about nodes, as usual in pointer structures. *) +(* For example, if there are 3 links for x, and null is among *) +(* the links, that encodes that the second child of x doesn't exist. *) +(* Non-null dangling links are technially possible, *) +(* but have the same meaning as null. *) + Definition links (g : pregraph) x := oapp id [::] (find x g). Lemma links_undef x : links undef x = [::]. @@ -263,7 +272,7 @@ Lemma range_links (g : pregraph) x : links g x \in range g. Proof. by move/In_graphX/mem_range. Qed. -(* children x are links obtained by non-dangling edges only *) +(* children x removes dangling edges (null or non-null) from links *) Definition children (g : pregraph) x : seq node := filter (mem (dom g)) (links g x). @@ -312,8 +321,11 @@ by apply: children_links. Qed. (* edge is applicative variant of children *) -(* NOTE: do we want applicative variant of links *) -(* that computes external edges? *) +(* NOTE: links to null or to other nodes that aren't *) +(* in graph's domain, are explicitly *not* considered *) +(* to be edges. If it's desired to treat them as edges, *) +(* then the external nodes should be made internal, *) +(* by adding them to the graph's domain *) Definition edge g : rel node := mem \o children g. Arguments edge g x y : simpl never. @@ -368,11 +380,12 @@ Qed. (* list of nodes traversed by depth-first search of g *) (* at depth n, starting from x, and avoiding v *) -(* DEVCOMMENT: *) -(* uses children, not links, to traverse internal edges only *) -(* traversing an external edge may be sensible, *) -(* as then one can reason about connectivity to external nodes *) -(* but it will require a completely different reflection theorem *) +(* NOTE: Definition uses children, not links, to avoid *) +(* following dangling edges. *) +(* To allow following dangling edges to dangling nodes, *) +(* make such dangling (non-null) nodes internal, *) +(* by adding them to the graph's domain *) + Fixpoint dfs (g : pregraph) (n : nat) (v : seq node) x := if (x \notin dom g) || (x \in v) then v else if n is n'.+1 then foldl (dfs g n') (x :: v) (children g x) else v. From a99d04d9d68b5db4f79fe43c58b1af1f99fe981f Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Wed, 7 Aug 2024 22:19:11 +0200 Subject: [PATCH 19/93] comments --- examples/graph.v | 35 ++++++++++++++++------------------- 1 file changed, 16 insertions(+), 19 deletions(-) diff --git a/examples/graph.v b/examples/graph.v index 292f2d5..e6c6a12 100644 --- a/examples/graph.v +++ b/examples/graph.v @@ -166,8 +166,14 @@ Proof. by case/allPn=>x; rewrite negbK; exists x. Qed. (* seq node listing the children of the node (adjacency list) *) (* Pregraph differs from graph, in that edges can *dangle*; that is *) (* terminate with a node that isn't in the graph. *) -(* Here, dangling edges are treated as non-existent; *) -(* that is, target of a dangling edge is treated as null pointer. *) +(* Dangling edges are allowed because they enable encoding positional *) +(* information about nodes, as usual in pointer structures. *) +(* For example, if there are 3 links for x, and null is the 2nd link, *) +(* that encodes that the second child of x doesn't exist. *) +(* Non-null dangling links are technically possible, *) +(* but are treated the same as null. *) +(* If it's desired to treat a non-null dangling link differently *) +(* from null, add that link to the graph to make it non-dangling. *) Notation node := nat. Record pregraph := Pregraph {pregraph_base : @UM.base node nat_pred (seq node)}. @@ -211,15 +217,8 @@ Coercion Pred_of_history (x : pregraph) : Pred_Class := Notation "x &-> v" := (ptsT pregraph x v) (at level 30). (* Links of x includes all edges outgoing from x *) -(* and may explicitly include nodes that aren't in dom g. *) -(* These are so called "dangling edges". *) -(* Dangling edges are allowed to enable positional information *) -(* about nodes, as usual in pointer structures. *) -(* For example, if there are 3 links for x, and null is among *) -(* the links, that encodes that the second child of x doesn't exist. *) -(* Non-null dangling links are technially possible, *) -(* but have the same meaning as null. *) - +(* and may explicitly include nodes that aren't in dom g *) +(* (i.e., are dangling, null or non-null) *) Definition links (g : pregraph) x := oapp id [::] (find x g). Lemma links_undef x : links undef x = [::]. @@ -321,11 +320,10 @@ by apply: children_links. Qed. (* edge is applicative variant of children *) -(* NOTE: links to null or to other nodes that aren't *) -(* in graph's domain, are explicitly *not* considered *) -(* to be edges. If it's desired to treat them as edges, *) -(* then the external nodes should be made internal, *) -(* by adding them to the graph's domain *) +(* NOTE: Links to dangling edges (null or non-null) *) +(* are explicitly *not* considered edges. *) +(* If it's desired to treat a non-null dangling node differently *) +(* from null, add that node to the graph to make it non-dangling. *) Definition edge g : rel node := mem \o children g. Arguments edge g x y : simpl never. @@ -382,9 +380,8 @@ Qed. (* at depth n, starting from x, and avoiding v *) (* NOTE: Definition uses children, not links, to avoid *) (* following dangling edges. *) -(* To allow following dangling edges to dangling nodes, *) -(* make such dangling (non-null) nodes internal, *) -(* by adding them to the graph's domain *) +(* If it's desired to treat a non-null dangling node differently *) +(* from null, add that node to the graph to make it non-dangling. *) Fixpoint dfs (g : pregraph) (n : nat) (v : seq node) x := if (x \notin dom g) || (x \in v) then v else From ecdc4de55c07fcbcec87610959773f5decf9500e Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Wed, 7 Aug 2024 22:20:58 +0200 Subject: [PATCH 20/93] minor --- examples/graph.v | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/examples/graph.v b/examples/graph.v index e6c6a12..266f6f4 100644 --- a/examples/graph.v +++ b/examples/graph.v @@ -165,7 +165,7 @@ Proof. by case/allPn=>x; rewrite negbK; exists x. Qed. (* Pregraphs are natmaps, mapping nodes into *) (* seq node listing the children of the node (adjacency list) *) (* Pregraph differs from graph, in that edges can *dangle*; that is *) -(* terminate with a node that isn't in the graph. *) +(* terminate with a node that isn't in the graph's domain. *) (* Dangling edges are allowed because they enable encoding positional *) (* information about nodes, as usual in pointer structures. *) (* For example, if there are 3 links for x, and null is the 2nd link, *) From d24ddf057806ced8663d6e030532fac85a0751b8 Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Wed, 7 Aug 2024 22:23:05 +0200 Subject: [PATCH 21/93] minor --- examples/graph.v | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/examples/graph.v b/examples/graph.v index 266f6f4..b9500db 100644 --- a/examples/graph.v +++ b/examples/graph.v @@ -171,7 +171,7 @@ Proof. by case/allPn=>x; rewrite negbK; exists x. Qed. (* For example, if there are 3 links for x, and null is the 2nd link, *) (* that encodes that the second child of x doesn't exist. *) (* Non-null dangling links are technically possible, *) -(* but are treated the same as null. *) +(* but are treated same as null. *) (* If it's desired to treat a non-null dangling link differently *) (* from null, add that link to the graph to make it non-dangling. *) From 443bd2e65d900bf8c5001618218a8f35a504dcf0 Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Thu, 8 Aug 2024 13:04:40 +0200 Subject: [PATCH 22/93] blah --- core/prelude.v | 6 +- examples/graph.v | 168 +++++++++++++++++++++++++++++++---------------- pcm/unionmap.v | 60 ++++++++++------- 3 files changed, 149 insertions(+), 85 deletions(-) diff --git a/core/prelude.v b/core/prelude.v index 46fac1b..c85cc3a 100644 --- a/core/prelude.v +++ b/core/prelude.v @@ -17,9 +17,9 @@ limitations under the License. (******************************************************************************) From HB Require Import structures. -From Coq Require Import ssreflect ssrbool ssrfun Eqdep. -From mathcomp Require Import ssrnat seq eqtype choice fintype finfun tuple. -From mathcomp Require Import perm fingroup. +From Coq Require Import ssreflect ssrfun Eqdep. +From mathcomp Require Import ssrbool ssrnat seq eqtype choice. +From mathcomp Require Import fintype finfun tuple perm fingroup. From pcm Require Import options axioms. (***********) diff --git a/examples/graph.v b/examples/graph.v index b9500db..8f4662b 100644 --- a/examples/graph.v +++ b/examples/graph.v @@ -12,8 +12,8 @@ limitations under the License. *) From HB Require Import structures. -From Coq Require Import ssreflect ssrbool ssrfun. -From mathcomp Require Import eqtype ssrnat seq path. +From Coq Require Import ssreflect ssrfun. +From mathcomp Require Import ssrbool eqtype ssrnat seq path. From pcm Require Import options axioms pred prelude. From pcm Require Import pcm unionmap natmap autopcm automap. From pcm Require Import seqext. @@ -527,16 +527,16 @@ Definition component' (p : pred node) g x : seq node := (* y is connected from x if y is visited during dfs *) (* under assumption that x is in dom g *) -Definition connect p (g : pregraph) := - [rel x y | (x \in dom g) && (y \in component' p g x)]. +Definition connect p (g : pregraph) x : pred node := + fun y => (x \in dom g) && (y \in component' p g x). -Lemma connectP p (g : pregraph) x y : - reflect [/\ x \in dom g, x \notin p & exists xs, +Lemma connectP (p : pred node) (g : pregraph) x y : + reflect [/\ x \in dom g, ~~ p x & exists xs, [/\ path (edge g) x xs, y = last x xs & - {in xs, forall z, z \notin p}]] + {in xs, forall z, ~~ p z}]] (y \in connect p g x). Proof. -rewrite /connect/component' inE mem_filter /= andbA. +rewrite /connect/component'/= {2}/in_mem /= mem_filter /= andbA. case: (boolP (x \in dom g))=>Dx; last by constructor; case. case: (boolP (p y))=>Py. - constructor; case=>_ Px [xs][P E Pxs]. @@ -559,11 +559,11 @@ Qed. (* there's path from x to y iff *) (* there's path that doesn't loop through x *) -Lemma connectX p (g : pregraph) x y : - reflect [/\ x \in dom g, x \notin p & exists xs, +Lemma connectX (p : pred node) (g : pregraph) x y : + reflect [/\ x \in dom g, ~~ p x & exists xs, [/\ path (edge g) x xs, y = last x xs, x \notin xs & - {in xs, forall z, z \notin p}]] + {in xs, forall z, ~~ p z}]] (y \in connect p g x). Proof. case: connectP=>H; constructor; last first. @@ -592,12 +592,12 @@ Proof. by move=>y; apply/connectP; case; rewrite dom0. Qed. Lemma connectDx p g x y : y \in connect p g x -> - (x \in dom g) * (x \notin p). + (x \in dom g) * ~~ p x. Proof. by case/connectP. Qed. Lemma connectDy p g x y : y \in connect p g x -> - (y \in dom g) * (y \notin p). + (y \in dom g) * ~~ p y. Proof. case/connectP=>Dx Px [xs][P E Pxs]. have : y \in x :: xs by rewrite E mem_last. @@ -613,7 +613,7 @@ Proof. by move=>C; rewrite (connectDx C) (connectDy C). Qed. Lemma connectDp p g x y : y \in connect p g x -> - (x \notin p) * (y \notin p). + (~~ p x) * (~~ p y). Proof. by move=>C; rewrite (connectDx C) (connectDy C). Qed. Lemma connectDN p g x : @@ -625,7 +625,7 @@ by rewrite (negbTE Dx). Qed. Lemma connectDNp (p : pred node) g x : - x \in p -> + p x -> connect p g x =i pred0. Proof. move=>Hx y; apply/negP=>/connectX []. @@ -633,18 +633,18 @@ by rewrite Hx. Qed. Lemma connect0 p g x : - x \in connect p g x = (x \in dom g) && (x \notin p). + x \in connect p g x = (x \in dom g) && ~~ p x. Proof. by apply/connectP/andP; case=>// H1 H2; split=>//; exists [::]. Qed. -Lemma connect0I p g x : +Lemma connect0I (p : pred node) (g : pregraph) x : x \in dom g -> - x \notin p -> + ~~ p x -> x \in connect p g x. Proof. by rewrite connect0=>->->. Qed. -Lemma connect0N p g x y : +Lemma connect0N (p : pred node) (g : pregraph) x y : x \in dom g -> - x \notin p -> + ~~ p x -> y \notin connect p g x -> x != y. Proof. by move=>Gx Px; apply: contra=>/eqP <-; rewrite connect0 Gx. Qed. @@ -677,8 +677,8 @@ Lemma connectUnR p g1 g2 x : connect p (g1 \+ g2) x}. Proof. by rewrite joinC; apply: connectUnL. Qed. -Lemma connect_sub g p1 p2 x : - {subset p2 <= p1} -> +Lemma connect_sub g (p1 p2 : pred node) x : + subpred p2 p1 -> {subset connect p1 g x <= connect p2 g x}. Proof. move=>S y /connectP [Dx Hx][xs][Px Ey Hxs]. @@ -687,7 +687,7 @@ by exists xs; split=>// z /Hxs; apply/contra/S. Qed. Lemma connect_eq g p1 p2 x : - p1 =i p2 -> + p1 =1 p2 -> connect p1 g x =i connect p2 g x. Proof. by move=>S y; apply/idP/idP; apply/connect_sub=>z; rewrite S. Qed. @@ -791,7 +791,7 @@ Qed. (* part of g reachable from x (avoiding nothing) *) Definition subconnect g x : pregraph := - um_filterk (connect pred0 g x) g. + um_filterk (mem (connect pred0 g x)) g. (* reachable component of a graph is a graph *) Lemma graph_subconnect g x : @@ -817,7 +817,7 @@ Lemma edge_subconnect g x y z : z \in connect pred0 g x & edge g y z]. Proof. -rewrite /edge/children/links/oapp/= find_umfiltk topredE. +rewrite /edge/children/links/oapp/= find_umfiltk /=. case: ifP=>// _; case: (find y g)=>[ys|]; last by rewrite andbF. by rewrite !mem_filter /= dom_umfiltk inE -andbA. Qed. @@ -834,28 +834,28 @@ move=>w; apply/iffE; split; case/connectP. by apply/sub_path/Px=>y z; rewrite edge_subconnect=>/and3P []. move=>Dx _ [xs][Px Ew _]; apply/connectP; split=>//. - by rewrite dom_umfiltk inE topredE connect0 Dx. -exists xs; split=>//; apply/(sub_in_path (P:=connect pred0 g x))/Px. +exists xs; split=>//. +apply/(sub_in_path (P:=[in connect pred0 g x]))/Px. - by move=>y z Cy Cz E; rewrite edge_subconnect Cy Cz E. apply/allP=>z; rewrite inE; case/orP=>[/eqP ->|Hz]. -- by rewrite topredE connect0 Dx. +- by rewrite connect0 Dx. case/splitPr: Hz Px=>xs1 xs2. rewrite -cat1s catA cats1 cat_path=>/andP [Px _]. apply/connectP; split=>//. by exists (rcons xs1 z); rewrite last_rcons. Qed. - (****************************) (* Avoidance sets (marking) *) (****************************) (* deconstructing connecting path from left *) (* strenghtening avoidance *) -Lemma edge_connectX g p (x y : node) : +Lemma edge_connectX p g (x y : node) : y != x -> y \in connect p g x -> exists2 z, - edge g x z & y \in connect [predU pred1 x & p] g z. + edge g x z & y \in connect (predU (pred1 x) p) g z. Proof. move/eqP=>Nxy /connectX [Dx Hx][[|a xs]] /= [Px Ey Nx Hxs]. - by move/Nxy: Ey. @@ -870,11 +870,11 @@ Qed. (* extending connecting path from left *) (* weakening avoidance *) -Lemma edge_connectXI p g x y z : +Lemma edge_connectXI (p : pred node) (g : pregraph) x y z : x \in dom g -> - x \notin p -> + ~~ p x -> edge g x y -> - z \in connect [predU pred1 x & p] g y -> + z \in connect (predU (pred1 x) p) g y -> z \in connect p g x. Proof. move=>Dx Px Ex /(connect_sub (p2:=p)) H. @@ -887,8 +887,8 @@ Qed. (* equivalence variant *) Lemma edge_connectXE p g x y : y \in connect p g x <-> - [/\ x \in dom g, x \notin p & y = x \/ exists2 z, - edge g x z & y \in connect [predU pred1 x & p] g z]. + [/\ x \in dom g, ~~ p x & y = x \/ exists2 z, + edge g x z & y \in connect (predU (pred1 x) p) g z]. Proof. split=>[C|[Dx Hx]]; last first. - case=>[->|[z E C]]; first by rewrite connect0 Dx Hx. @@ -904,10 +904,10 @@ Qed. Lemma connect_avoid p g x y x' : y \in connect p g x -> y \notin connect p g x' -> - y \in connect [predU p & connect p g x'] g x. + y \in connect (predU p [in connect p g x']) g x. Proof. case/connectP=>Dx Hx [xs][Px Ey Hxs] Ny. -have [Nx|Mx] := boolP (has (connect p g x') (x :: xs)); last first. +have [Nx|Mx] := boolP (has [in connect p g x'] (x :: xs)); last first. (* if path contains no nodes reachable from x', nothing to do *) - move/hasPn: Mx=>Mx; apply/connectP; split=>//. - by rewrite inE negb_or Hx Mx // inE eqxx. @@ -915,7 +915,7 @@ have [Nx|Mx] := boolP (has (connect p g x') (x :: xs)); last first. (* if path contains node reachable from x' *) (* let a be the last such node in the path *) case: {-1} _ _ _ / (split_findlast Nx) (erefl (x::xs)). -move=>{Nx} a p1 p2; rewrite topredE=>Ca /hasPn Nx' X. +move=>{Nx} a p1 p2=>Ca /hasPn/= Nx' X. (* suffices to show that y is reachable from a *) (* because then y is also reachable from x'; contradiction *) suff Cy : y \in connect p g a. @@ -987,37 +987,89 @@ move=>H Dx; apply/(eq_from_nth (x0:=null))=>[|i Hi]. by rewrite map_nth_iota0 ?H // -(H _ Dx) take_size. Qed. +(**************) +(* Acyclicity *) +(**************) +Definition biconnect p g x : pred node := + [pred y | (x \in connect p g y) && (y \in connect p g x)]. +Lemma biconnect0 p g x : + x \in dom g -> + ~~ p x -> + biconnect p g x x. +Proof. by move=>Dx Hp; rewrite /biconnect/= connect0 Dx Hp. Qed. -(*******************) +Lemma biconnectUnL p g1 g2 x y : + valid (g1 \+ g2) -> + biconnect p g1 x y -> + biconnect p (g1 \+ g2) x y. +Proof. by move=>V /andP [Cx Cy]; apply/andP; split; apply/connectUnL. Qed. -Section Acyclic. +Lemma biconnectUnR p g1 g2 x y : + valid (g1 \+ g2) -> + biconnect p g2 x y -> + biconnect p (g1 \+ g2) x y. +Proof. by rewrite joinC; apply: biconnectUnL. Qed. -Definition symconnect p g x y := connect p g x y && connect p g y x. +Lemma biconnect_sub g (p1 p2 : pred node) (x : node) : + subpred p2 p1 -> + {subset biconnect p1 g x <= biconnect p2 g x}. +Proof. +by move=>S y; rewrite !inE=>/andP [Hx Hy]; rewrite !(connect_sub S). +Qed. -Lemma symconnect0 p g x : x \in dom g -> x \notin dom p -> symconnect p g x x. -Proof. by move=>Hx Hp; apply/andP; split; apply/connect0. Qed. +Lemma biconnect_eq g (p1 p2 : pred node) (x : node) : + p1 =1 p2 -> + biconnect p1 g x =i biconnect p2 g x. +Proof. by move=>S y; rewrite !inE !(connect_eq g _ S). Qed. -Lemma symconnectUn p g0 g x y : - valid (g \+ g0) -> - symconnect p g x y -> symconnect p (g \+ g0) x y. -Proof. by move=>V; case/andP=>Hxy Hyx; apply/andP; split; apply: connectUn. Qed. +(* elements in a cycle are interconnected *) +(* avoiding everything outside the cycle *) -(* TODO should probably generalize all of this to arbitrary boundary, not Unit *) -Lemma connect_cycle g p : cycle (edge g) p -> {in p &, forall x y, connect Unit g x y}. +(* NOTE: NOTE: NOTE: this is connect lemma, not biconnect *) +Lemma connect_cycle g xs : + cycle (edge g) xs -> + {in xs &, forall x y, y \in connect [predC xs] g x}. Proof. -move=>Hp x y /rot_to[i q Hr]; rewrite -(mem_rot i) Hr => Hy. -have /= Hp1: cycle (edge g) (x :: q) by rewrite -Hr rot_cycle. -have Hd: x \in dom g by move: Hp1; rewrite rcons_path; case/andP=>_ /edge_dom; case. -move: Hp1; case/splitPl: Hy =>r s Hl; rewrite rcons_cat cat_path => /andP[Hxr Hlx]. -apply/connectP; exists r; split=>//; first by rewrite Hd implybT. -by rewrite dom0 all_predT. +move=>C x y /rot_to [i q Hr]; rewrite -(mem_rot i) Hr => Hy. +have Hx : x \in xs by rewrite -(mem_rot i) Hr inE eqxx. +have /= Hp1: cycle (edge g) (x :: q) by rewrite -Hr rot_cycle. +have Dx : x \in dom g. +- by move: Hp1; rewrite rcons_path=>/andP [_ /edge_dom][]. +case/splitPl: Hy Hp1 Hr=>r s Ey. +rewrite rcons_cat cat_path=>/andP [Hxr]. +rewrite Ey rcons_path; case/andP=>Hlx /= Ex Hr. +apply/connectP; split=>//=; first by rewrite negbK. +exists r; split=>// z Z. +by rewrite negbK -(mem_rot i) Hr inE mem_cat Z orbT. Qed. -Lemma symconnect_cycle g p : cycle (edge g) p -> - {in p &, forall x y, symconnect Unit g x y}. -Proof. by move=>Hp x y Hx Hy; rewrite /symconnect !(connect_cycle Hp). Qed. +Lemma connect_cycle0 g xs : + cycle (edge g) xs -> + {in xs &, forall x y, y \in connect pred0 g x}. +Proof. by move=>C x y Hx /(connect_cycle C Hx); apply: connect_sub. Qed. + +(* elements in a cycle are bi-interconnected *) +(* avoiding everything outside the cycle *) + +Lemma biconnect_cycle g xs : + cycle (edge g) xs -> + {in xs &, forall x y, biconnect [predC xs] g x y}. +Proof. by move=>C x y Hx Hy; rewrite /biconnect /= !(connect_cycle C). Qed. + + +Lemma biconnect_cycle0 g xs : + cycle (edge g) xs -> + {in xs &, forall x y, biconnect pred0 g x y}. +Proof. move=>C x y Hx /(biconnect_cycle C Hx); apply: biconnect_sub. + +UP TO HERE + +(*******************) + +Section Acyclic. + Lemma symconnect_cycle2P g x y : x != y -> reflect (exists2 p, y \in p & cycle (edge g) (x :: p)) (symconnect Unit g x y). diff --git a/pcm/unionmap.v b/pcm/unionmap.v index e02fcd0..22049bc 100644 --- a/pcm/unionmap.v +++ b/pcm/unionmap.v @@ -1243,7 +1243,7 @@ End DomEq2Lemmas. (**********) Section UpdateLemmas. -Variable (K : ordType) (C : pred K) (V : Type) (U : union_map C V). +Variable (K : ordType) (C : pred K) (V : Type) (U : union_map C V). Implicit Types (k : K) (v : V) (f : U). Lemma upd_undef k v : upd k v undef = undef :> U. @@ -1632,11 +1632,11 @@ Qed. (* a weaker legacy version of domPtK *) Lemma domPt k v : dom (pts k v : U) =i [pred x | C k & k == x]. -Proof. +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) = C k. +Lemma validPt k v : valid (pts k v : U) = C k. Proof. by rewrite ptsU validU valid_unit andbT. Qed. Lemma domPtV k v : valid (pts k v : U) -> k \in dom (pts k v : U). @@ -2091,7 +2091,7 @@ Lemma umPtPtE (k1 k2 : K) (v1 v2 : V) : if C k1 then [&& C k2, k1 == k2 & v1 == v2] else ~~ C k2. Proof. case D1 : (C k1); last first. -- case D2 : (C k2); rewrite pts_condN ?D1 //. +- case D2 : (C k2); rewrite pts_condN ?D1 //. - by apply/eqP/nesym; rewrite pts_undef D2. by rewrite pts_condN ?D2 // eq_refl. rewrite -(validPt U k1 v1) -(validPt U k2 v2) in D1 *. @@ -2302,21 +2302,25 @@ 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. +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 ->-> [->]. +case: x=>a b /In_find; rewrite findPt2. +by 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 InPtE k v f : - C k -> f = pts k v -> (k, v) \In f. + C k -> + f = pts k v -> (k, v) \In f. Proof. by move=>W ->; apply: In_condPt W. Qed. Lemma In_dom f x : x \In f -> x.1 \in dom f. @@ -3639,7 +3643,7 @@ Lemma omfUE f k (v : V) (x : U) : if C k then if omf f (k, v) is Some v' then upd k v' (f x) else free (f x) k - else undef. + else undef. Proof. case: ifP=>// D; last by rewrite (upd_condN v x (negbT D)) pfundef. case: (normalP x)=>[->|H]. @@ -3756,7 +3760,7 @@ Arguments In_omf {K C V V' U U'} _ {k v w x}. (* omap_fun's with different ranges *) Section OmapFun2. -Variables (K : ordType) (C : pred K) (V V1 V2 : Type). +Variables (K : ordType) (C : pred K) (V V1 V2 : Type). Variables (U : union_map C V) (U1 : union_map C V1) (U2 : union_map C V2). Variables (f1 : omap_fun U U1) (f2 : omap_fun U U2). @@ -4317,7 +4321,7 @@ Lemma umfiltUE p k v f : if p (k, v) then upd k v (um_filter p f) else free (um_filter p f) k else undef. -Proof. by rewrite omfUE omf_omap; case: (p _). Qed. +Proof. by rewrite omfUE omf_omap /=; case: (p _). Qed. Lemma umfiltU p k v f : C k -> @@ -4430,7 +4434,7 @@ Qed. Lemma valid_umfiltkUn p1 p2 f : valid f -> - {in dom f, forall x, x \in p1 -> x \in p2 -> false} -> + {in dom f, forall x, p1 x -> p2 x -> false} -> valid (um_filterk p1 f \+ um_filterk p2 f). Proof. move=>W H; rewrite validUnAE !pfVE W /=. @@ -4439,7 +4443,7 @@ apply/negP; case/In_dom_umfilt=>v2 /= [H1 F2]. by move: (H x (In_dom F1) H1 H2). Qed. -Lemma dom_umfiltk p f : dom (um_filterk p f) =i predI p (mem (dom f)). +Lemma dom_umfiltk p f : dom (um_filterk p f) =i predI p (mem (dom f)). Proof. by move=>k; rewrite dom_umfiltkE mem_filter. Qed. (* DEVCOMMENT *) @@ -4558,7 +4562,7 @@ Proof. by apply: In_umfilt. Qed. End FilterKLemmas. -Arguments In_umfiltk [K C V U] p [x f] _ _. +Arguments In_umfiltk {K C V U} p {x f} _ _. Section FilterVLemmas. Variables (K : ordType) (C : pred K) (V : Type) (U : union_map C V). @@ -4580,7 +4584,7 @@ Proof. by apply: In_umfilt. Qed. End FilterVLemmas. -Arguments In_umfiltv [K C V U] p [x f] _ _. +Arguments In_umfiltv {K C V U} p {x f} _ _. Section OmapMembershipLemmas. Variables (K : ordType) (C : pred K) (V V' : Type). @@ -4626,7 +4630,9 @@ 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]. + {subset dom a <= p} -> + [pcm a <= f] -> + [pcm a <= um_filterk p f]. Proof. move=>D /(umfilt_pleq_mono (p \o fst)). by rewrite (eq_in_umfiltk D) umfilt_predT. @@ -4640,7 +4646,10 @@ Proof. by move=>W; split=>[|<-] // H; case: H W=>b ->; 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]. + valid (a \+ b) -> + [pcm a <= f] -> + [pcm b <= f] -> + [pcm a \+ b <= f]. Proof. case: (normalP f)=>[->???|Df Dab]; first by apply: pleq_undef. move/(umpleqk_pleqE _ Df)=> <- /(umpleqk_pleqE _ Df) <-. @@ -4996,7 +5005,7 @@ Qed. End OrdEvalDefLemmas. -Arguments oev_sub_filter [K C V R U a ks p] _. +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. @@ -5253,7 +5262,8 @@ End EvalRelationalInduction1. Section EvalRelationalInduction2. -Variables (K1 K2 : ordType) (C1 : pred K1) (C2 : pred K2) (V1 V2 R1 R2 : Type). +Variables (K1 K2 : ordType) (C1 : pred K1) (C2 : pred K2). +Variables (V1 V2 R1 R2 : Type). Variables (U1 : union_map C1 V1) (U2 : union_map C2 V2). Variables (U : union_map C1 K2) (P : R1 -> R2 -> Prop). @@ -6113,7 +6123,7 @@ Qed. End MapInvert. -Arguments invert [K V C C' U U']. +Arguments invert {K V C C' U U'}. Section InvertLaws. @@ -6139,7 +6149,7 @@ Qed. End InvertLaws. -Arguments In_invert [K V C C' U U' k v f]. +Arguments In_invert {K V C C' U U' k v f}. (***************************) @@ -6381,11 +6391,11 @@ Hint Resolve umall_undef umall0 : core. Section MapAllDecidable. Variables (K : ordType) (C : pred K) (V : Type) (U : union_map C V). -Implicit Types (f : U) (p : pred (K*V)). +Implicit Types (f : U) (p : pred (K*V)). Definition um_allb p f := um_count p f == size (dom f). -Lemma umallbP p f : reflect (forall x, x \In f -> p x) (um_allb p f). +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. @@ -6451,10 +6461,12 @@ End MapAllDecidable. (* values of equal keys in both maps. *) Section BinaryMapAll. -Variables (K : ordType) (C : pred K) (V : Type) (U : union_map C V) (P : V -> V -> Prop). +Variables (K : ordType) (C : pred K) (V : Type) (U : union_map C V). +Variables (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). +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. From f9a4eb537b3d83b795e95b3558dc615fc638321e Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Thu, 8 Aug 2024 18:13:37 +0200 Subject: [PATCH 23/93] wip --- pcm/autopcm.v | 7 ------- pcm/natmap.v | 8 ++++++-- pcm/unionmap.v | 10 +++++----- 3 files changed, 11 insertions(+), 14 deletions(-) diff --git a/pcm/autopcm.v b/pcm/autopcm.v index e3b5f0f..8c4dea9 100644 --- a/pcm/autopcm.v +++ b/pcm/autopcm.v @@ -354,10 +354,3 @@ End Exports. End PullX. Export PullX.Exports. - - - - - - - diff --git a/pcm/natmap.v b/pcm/natmap.v index 7676532..918ab83 100644 --- a/pcm/natmap.v +++ b/pcm/natmap.v @@ -124,6 +124,9 @@ Lemma In_neq A (U : natmap A) (h : U) k : k.1 > 0. Proof. by move/In_cond; rewrite lt0n. Qed. +Lemma mem_domN0 A (U : natmap A) {h : U} : 0 \notin dom h. +Proof. by rewrite cond_dom. Qed. + Lemma mem_dom0 A (U : natmap A) {h : U} : 0 \in dom h = false. Proof. by rewrite cond_dom. Qed. @@ -136,10 +139,11 @@ Proof. by apply/allP=>x /dom_cond; case: ltngtP. Qed. Lemma all_ltn0 A (U : natmap A) {h : U} : all (ltn 0) (dom h). Proof. by apply/allP=>x /dom_cond; case: ltngtP. Qed. -#[export] Hint Resolve uniq_dom0 all_leq0 all_ltn0 : core. +#[export] Hint Resolve uniq_dom0 all_leq0 all_ltn0 mem_domN0 : core. Lemma In_dom0 {A} {U : natmap A} (h : U) k e : - (k, e) \In h -> k \in 0 :: dom h. + (k, e) \In h -> + k \in 0 :: dom h. Proof. by move=>H; rewrite inE (In_dom H) orbT. Qed. Lemma sorted_leq_dom A {U : natmap A} {h : U} : sorted leq (dom h). diff --git a/pcm/unionmap.v b/pcm/unionmap.v index 22049bc..64ad854 100644 --- a/pcm/unionmap.v +++ b/pcm/unionmap.v @@ -1243,7 +1243,7 @@ End DomEq2Lemmas. (**********) Section UpdateLemmas. -Variable (K : ordType) (C : pred K) (V : Type) (U : union_map C V). +Variable (K : ordType) (C : pred K) (V : Type) (U : union_map C V). Implicit Types (k : K) (v : V) (f : U). Lemma upd_undef k v : upd k v undef = undef :> U. @@ -1636,7 +1636,7 @@ 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) = C k. +Lemma validPt k v : valid (pts k v : U) = C k. Proof. by rewrite ptsU validU valid_unit andbT. Qed. Lemma domPtV k v : valid (pts k v : U) -> k \in dom (pts k v : U). @@ -3643,7 +3643,7 @@ Lemma omfUE f k (v : V) (x : U) : if C k then if omf f (k, v) is Some v' then upd k v' (f x) else free (f x) k - else undef. + else undef. Proof. case: ifP=>// D; last by rewrite (upd_condN v x (negbT D)) pfundef. case: (normalP x)=>[->|H]. @@ -4321,7 +4321,7 @@ Lemma umfiltUE p k v f : if p (k, v) then upd k v (um_filter p f) else free (um_filter p f) k else undef. -Proof. by rewrite omfUE omf_omap /=; case: (p _). Qed. +Proof. by rewrite omfUE omf_omap; case: (p _). Qed. Lemma umfiltU p k v f : C k -> @@ -4443,7 +4443,7 @@ apply/negP; case/In_dom_umfilt=>v2 /= [H1 F2]. by move: (H x (In_dom F1) H1 H2). Qed. -Lemma dom_umfiltk p f : dom (um_filterk p f) =i predI p (mem (dom f)). +Lemma dom_umfiltk p f : dom (um_filterk p f) =i predI p (mem (dom f)). Proof. by move=>k; rewrite dom_umfiltkE mem_filter. Qed. (* DEVCOMMENT *) From 8944fd136ae6ee09145fb39a024754c2a53a2a20 Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Thu, 8 Aug 2024 18:13:49 +0200 Subject: [PATCH 24/93] wip --- examples/graph.v | 186 +++++++++++++++++++++++++++++------------------ 1 file changed, 114 insertions(+), 72 deletions(-) diff --git a/examples/graph.v b/examples/graph.v index 8f4662b..091dcab 100644 --- a/examples/graph.v +++ b/examples/graph.v @@ -358,12 +358,11 @@ rewrite /edge/= => H; split; last by apply: childrenD H. by apply: contraLR H=>/childrenND ->. Qed. -Lemma find_edge g x y ys : - find x g = Some ys -> - edge g x y = (y \in dom g) && (y \in ys). +Lemma edge_links g x y : + edge g x y = (y \in dom g) && (y \in links g x). Proof. -rewrite /edge/children/links/oapp/= => ->. -by rewrite mem_filter. +rewrite /edge/children/links/oapp/= mem_filter /=. +by case: dom_find. Qed. Lemma path_dom g x xs : @@ -803,12 +802,12 @@ have E : g = subconnect g x \+ um_filterk (negb \o connect pred0 g x) g by apply: umfilt_predC. move=>V Ha xs /= n Hxs Hn; have {}Dn : in_node g n. - by apply: Ha Hn; rewrite E rangeUn inE -E V Hxs. -case/mem_rangeX: Hxs=>k /In_umfiltX [/= Ck] /In_find Hf. +case/mem_rangeX: Hxs=>k /In_umfiltX [/= Ck] /In_graph Hf. rewrite /in_node in Dn *; case/boolP: (n == null) Dn=>//= Hnn Dn. case: (connectD Ck)=>_ Dk. rewrite /subconnect dom_umfiltk inE /= Dn andbT. apply: connect_trans Ck _; apply/connectP; split=>//. -by exists [:: n]; split=>//=; rewrite (find_edge _ Hf) Dn Hn. +by exists [:: n]; split=>//=; rewrite edge_links Dn -Hf Hn. Qed. Lemma edge_subconnect g x y z : @@ -987,9 +986,9 @@ move=>H Dx; apply/(eq_from_nth (x0:=null))=>[|i Hi]. by rewrite map_nth_iota0 ?H // -(H _ Dx) take_size. Qed. -(**************) -(* Acyclicity *) -(**************) +(******************) +(* Biconnectivity *) +(******************) Definition biconnect p g x : pred node := [pred y | (x \in connect p g y) && (y \in connect p g x)]. @@ -1055,90 +1054,133 @@ Proof. by move=>C x y Hx /(connect_cycle C Hx); apply: connect_sub. Qed. Lemma biconnect_cycle g xs : cycle (edge g) xs -> - {in xs &, forall x y, biconnect [predC xs] g x y}. -Proof. by move=>C x y Hx Hy; rewrite /biconnect /= !(connect_cycle C). Qed. - + {in xs &, forall x y, y \in biconnect [predC xs] g x}. +Proof. by move=>C x y Hx Hy; rewrite /biconnect inE !(connect_cycle C). Qed. Lemma biconnect_cycle0 g xs : cycle (edge g) xs -> - {in xs &, forall x y, biconnect pred0 g x y}. -Proof. move=>C x y Hx /(biconnect_cycle C Hx); apply: biconnect_sub. + {in xs &, forall x y, y \in biconnect pred0 g x}. +Proof. by move=>C x y Hx /(biconnect_cycle C Hx); apply: biconnect_sub. Qed. + +Lemma biconnect_cycle2P p g x y : + x != y -> + reflect (exists xs : seq node, + [/\ y \in xs, cycle (edge g) (x :: xs) & + {in x :: xs, forall z : node, ~~ p z}]) + (y \in biconnect p g x). +Proof. +move=>Nxy; apply/(iffP idP)=>[|[ys][Py Cy Hys]]; last first. +- apply: biconnect_sub (biconnect_cycle Cy _ _). + - by move=>z; apply/contraL/Hys. + - by rewrite inE eqxx. + by rewrite inE Py orbT. +case/andP=>/connectP [Dy Hy][ys]; elim/last_ind: ys Nxy=>[|ys a IH] Nxy. +- by move/eqP: Nxy=>Nxy [_ /Nxy]. +rewrite rcons_path last_rcons; case=>/[swap] <-{a} /andP [Py Ex Hys]. +case/connectP=>Dx Hx [xs][Px Ey Hxs]; exists (xs ++ ys); split=>/=. +- have := mem_last x xs. + by rewrite -Ey inE eq_sym (negbTE Nxy) /= mem_cat=>->. +- by rewrite rcons_path cat_path last_cat -Ey Px Py Ex. +by move=>z; rewrite -mem_rcons rcons_cat mem_cat=>/orP []; +[apply: Hxs|apply: Hys]. +Qed. -UP TO HERE +Lemma biconnect_cycle2P0 g x y : + x != y -> + reflect (exists2 xs, y \in xs & cycle (edge g) (x :: xs)) + (y \in biconnect pred0 g x). +Proof. +move=>Nxy; apply/(iffP (biconnect_cycle2P pred0 g Nxy)). +- by case=>xs [H1 H2 H3]; exists xs. +by case=>xs H1 H2; exists xs. +Qed. -(*******************) +(**************) +(* Acyclicity *) +(**************) -Section Acyclic. +(* graph is preacyclic if only self-loops are biconnected *) +Definition preacyclic g := + all2rel (fun x y => (y \in biconnect pred0 g x) ==> (x == y)) (dom g). +(* graph is acyclic if it doesn't even have self-loops *) +Definition acyclic g := + preacyclic g && all (fun x => ~~ edge g x x) (dom g). -Lemma symconnect_cycle2P g x y : x != y -> - reflect (exists2 p, y \in p & cycle (edge g) (x :: p)) (symconnect Unit g x y). +Lemma preacyclicP g : + reflect {in dom g &, forall x y, y \in biconnect pred0 g x -> x = y} + (preacyclic g). Proof. -move=> Nxy; apply: (iffP idP) => [|[p yp]]; last first. - by move=> /symconnect_cycle->//; rewrite inE ?eqxx ?yp ?orbT. -move: Nxy =>/[swap]/andP [/connectP[p][Hp {y}-> Np _] /connectP[]]. -elim/last_ind => /= [[] _ <-|q z _]; first by rewrite eqxx. -case; rewrite rcons_path last_rcons => /[swap]{z}<- /andP[gq gzq] _ _ Nxy. -have := mem_last x p; rewrite in_cons eq_sym (negPf Nxy)/= => yp. -exists (p ++ q); first by rewrite mem_cat yp. -by rewrite rcons_path cat_path Hp gq last_cat gzq. +apply: (iffP idP)=>[|H]. +- by move/allrelP=>H x y Dx Dy C; apply/eqP/(implyP _ C)/H. +by apply/allrelP=>x y Dx Dy; apply/implyP=>K; apply/eqP/H. Qed. -Definition preacyclic g := all2rel (fun x y => symconnect Unit g x y ==> (x == y)) (dom g). - -Lemma preacyclicUn g0 g : - valid (g \+ g0) -> - preacyclic (g \+ g0) -> preacyclic g. +Lemma preacyclicL g1 g2 : + valid (g1 \+ g2) -> + preacyclic (g1 \+ g2) -> + preacyclic g1. Proof. -move=>V /allrelP H; apply/allrelP=>x y Hx Hy; apply/implyP=>Hs. -have Hx1: x \in dom (g \+ g0) by rewrite domUn inE V Hx. -have Hy1: y \in dom (g \+ g0) by rewrite domUn inE V Hy. -move/implyP: (H _ _ Hx1 Hy1); apply. -by apply: symconnectUn. +move=>V /preacyclicP H. +apply/preacyclicP=>x y Dx Dy C; apply: H. +- by rewrite domUn inE V Dx. +- by rewrite domUn inE V Dy. +by apply: biconnectUnL V C. Qed. -Definition acyclic g := all (fun x => ~~ edge g x x) (dom g) && preacyclic g. +Lemma preacyclicR g1 g2 : + valid (g1 \+ g2) -> + preacyclic (g1 \+ g2) -> + preacyclic g2. +Proof. by rewrite joinC; apply: preacyclicL. Qed. -Lemma acyclicUn g0 g : - valid (g \+ g0) -> - acyclic (g \+ g0) -> acyclic g. +Lemma acyclicL g1 g2 : + valid (g1 \+ g2) -> + acyclic (g1 \+ g2) -> + acyclic g1. Proof. -move=>V; case/andP=>Ha Hp; apply/andP; split; last by apply: (preacyclicUn V Hp). -apply/allP=>x Hx. -suff: ~~ edge (g \+ g0) x x by apply/contra/edgeUn. -by move/allP: Ha; apply; rewrite domUn inE V Hx. +move=>V /andP [Hp /allP Ha]. +apply/andP; split; first by apply: preacyclicL Hp. +apply/allP=>x Dx; apply: contra (edgeUnL V) (Ha x _). +by rewrite domUn inE V Dx. Qed. -(* TODO all is overkill here, we only need the head of the path in dom g (?) *) +Lemma acyclicR g1 g2 : + valid (g1 \+ g2) -> + acyclic (g1 \+ g2) -> + acyclic g2. +Proof. by rewrite joinC; apply: acyclicL. Qed. + +(* graph is acyclic = graph has no cycles *) Lemma acyclic_cycleP g : - reflect (forall p, ~~ nilp p -> sorted (edge g) p -> all (fun x => x \in dom g) p -> ~~ cycle (edge g) p) - (acyclic g). + reflect (forall x xs, x \in dom g -> ~~ cycle (edge g) (x :: xs)) + (acyclic g). Proof. -rewrite /acyclic; apply: (iffP andP)=>[|Hn]; last first. -- split; first by apply/allP=>x Hx; move: (Hn [::x])=>/=; rewrite !andbT; apply. - apply/allrelP=>x y Hx _; apply/implyP/contraLR => neqxy. - apply/symconnect_cycle2P => // -[p Hp] /[dup]. - rewrite [X in X -> _]/= rcons_path => /andP[/[dup] /path_dom Hd Hg Ha]. - by apply/negP/Hn=>//=; rewrite Hx. -case=>/allP Ne /allrelP Ng. -case=>//= x p _; rewrite rcons_path =>/[dup] E ->/=; case/andP=>Hx. -case: p E=>/=; first by move=>_ _; apply: Ne. -move=>y p; case/andP=>He Hp; case/andP=>Hy Ha. -have: ~~ symconnect Unit g x y. -- apply/negP=>Hs; move: (Ng _ _ Hx Hy); rewrite Hs /= =>/eqP Exy. - by rewrite Exy in He; move: (Ne _ Hy); rewrite He. -apply: contra=>He1; apply: (symconnect_cycle (p:=x::y::p))=>/=; try by rewrite ?(in_cons, eqxx, orbT). -by rewrite rcons_path He Hp. +apply: (iffP idP)=>[|H]; last first. +- apply/andP; split; last first. + - by apply/allP=>x Dx; apply: contra (H _ [::] Dx); rewrite /= =>->. + apply/preacyclicP=>x y Dx Dy By; apply/eqP/(contraLR _ By)=>{By} Nxy. + by apply/(biconnect_cycle2P0 _ Nxy); case=>/= xs Hx; apply/negP/H/Dx. +case/andP=>/preacyclicP Ng /allP Ne x xs Dx /=. +rewrite rcons_path; apply/negP=>/andP []. +case: xs=>[_|y xs /= /andP [Exy Px]]; first by apply/negP/Ne. +have Dy : y \in dom g by rewrite (edge_dom Exy). +have : y \notin biconnect pred0 g x. +- by apply: contraL Exy=>/(Ng x y Dx Dy) <-; apply: Ne Dx. +apply: contraNnot=>Ex; apply: (biconnect_cycle0 (xs:=x::y::xs))=>/=. +- by rewrite Exy rcons_path Px Ex. +- by rewrite inE eqxx. +by rewrite !inE eqxx orbT. Qed. -Lemma acyclic_find g n ns : - acyclic g -> - find n g = Some ns -> - n \notin ns. +Lemma acyclic_links g x : + acyclic g -> + x \notin links g x. Proof. -case/andP=>Ha _ /[dup]Hf /find_some Hn. -move/allP: Ha=>/(_ _ Hn). -by rewrite (find_edge _ Hf) negb_and Hn. +case/andP=>_ /allP H. +have [Dx|Nx] := boolP (x \in dom g); last by rewrite linksND. +by apply: contra (H _ Dx)=>Lx; rewrite edge_links Dx Lx. Qed. -End Acyclic. + + From 83aa06db2a3950cb22c15424839c3709ca09ae02 Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Fri, 9 Aug 2024 17:48:05 +0200 Subject: [PATCH 25/93] refactored schorr.v --- _CoqProject | 1 + core/prelude.v | 17 +- core/seqext.v | 136 +++++++++- examples/graph.v | 573 ++++++++++++++++++++++--------------------- examples/quicksort.v | 2 +- pcm/natmap.v | 45 ++++ 6 files changed, 477 insertions(+), 297 deletions(-) diff --git a/_CoqProject b/_CoqProject index f94a95a..95176bc 100644 --- a/_CoqProject +++ b/_CoqProject @@ -52,3 +52,4 @@ examples/counter.v examples/congmath.v examples/tree.v examples/union_find.v +examples/graph.v \ No newline at end of file diff --git a/core/prelude.v b/core/prelude.v index c85cc3a..f286e8f 100644 --- a/core/prelude.v +++ b/core/prelude.v @@ -530,9 +530,17 @@ case: b1 b2 N; case=>//= _. by case=>_ /(_ erefl). Qed. +Lemma implyb_trans a b c : + a ==> b -> + b ==> c -> + a ==> c. +Proof. by case: a=>//=->. Qed. + Lemma iffE (b1 b2 : bool) : b1 = b2 <-> (b1 <-> b2). Proof. by split=>[->|] //; move/iffPb/eqP. Qed. +(* subsets *) + Lemma subsetC T (p q : mem_pred T) : {subset p <= q} -> {subset predC q <= predC p}. @@ -569,12 +577,6 @@ move=>H1 H2 D1; apply: subset_disj H1 _ => x H /H2. by apply: D1 H. Qed. -Lemma implyb_trans a b c : - a ==> b -> - b ==> c -> - a ==> c. -Proof. by case: a=>//=->. Qed. - (**************) (* empty type *) (**************) @@ -893,9 +895,6 @@ Qed. End SeqPrefixEq. - - - (******************************) (* Some commuting conversions *) (******************************) diff --git a/core/seqext.v b/core/seqext.v index 6a1eb77..174b363 100644 --- a/core/seqext.v +++ b/core/seqext.v @@ -137,6 +137,15 @@ Lemma all_notin (p : pred A) xs y : y \notin xs. Proof. by move/allP=>Ha; apply/contra/Ha. Qed. +Lemma in_consE P x xs : + {in x :: xs, forall z, P z} <-> + P x /\ {in xs, forall z, P z}. +Proof. +split=>[H|[H1 H2]]; last by move=>z /orP [/eqP ->|/H2]. +split=>[|z Z]; first by apply: H; rewrite inE eqxx. +by apply: H; rewrite inE Z orbT. +Qed. + Lemma subset_all a (s1 s2 : seq A) : {subset s1 <= s2} -> all a s2 -> @@ -154,6 +163,33 @@ Lemma subset_nil xs ys : xs = [::]. Proof. by move=>X E; move: E X=>->; apply: subset_nilR. Qed. +Lemma subset_consL x (s1 s2 : seq A) : + {subset x :: s1 <= s2} <-> + x \in s2 /\ {subset s1 <= s2}. +Proof. +split=>[S|[X S]]. +- by split=>[|z Z]; apply: S; rewrite inE ?eqxx ?Z ?orbT. +by move=>z; rewrite inE; case/orP=>[/eqP ->|/S//]. +Qed. + +Lemma subset_consLI x (s1 s2 : seq A) : + x \in s2 -> + {subset s1 <= s2} -> + {subset x :: s1 <= s2}. +Proof. by move=>H1 H2; rewrite subset_consL. Qed. + +Lemma subset_consR x (s : seq A) : + {subset s <= x :: s}. +Proof. by move=>z E; rewrite inE E orbT. Qed. + +Lemma subset_consLR x (s1 s2 : seq A) : + {subset s1 <= s2} -> + {subset x :: s1 <= x :: s2}. +Proof. +move=>X z; rewrite !inE; case/orP=>[|/X] -> //. +by rewrite orbT. +Qed. + Lemma all_mem xs ys : reflect {subset ys <= xs} (all [mem xs] ys). Proof. @@ -341,7 +377,105 @@ Qed. End LemmasEq. -(* finding last occurrence *) +(* decidable sequence disjointness *) + +Definition disjoint {A : eqType} (s1 s2 : seq A) := + all (fun x => x \notin s1) s2. + +Arguments disjoint {A} : simpl never. + +Lemma disjointC (A : eqType) (s1 s2 : seq A) : + disjoint s1 s2 = disjoint s2 s1. +Proof. +apply/idP/idP=>/allP S; apply/allP=>x X; +by apply/negP=>/S; rewrite X. +Qed. + +Lemma disjoint_catR (A : eqType) (s s1 s2 : seq A) : + disjoint s (s1 ++ s2) = + disjoint s s1 && disjoint s s2. +Proof. by rewrite /disjoint all_cat. Qed. + +Lemma disjoint_catL (A : eqType) (s s1 s2 : seq A) : + disjoint (s1 ++ s2) s = + disjoint s1 s && disjoint s2 s. +Proof. by rewrite -!(disjointC s) disjoint_catR. Qed. + +Lemma disjoint1L (A : eqType) x (s : seq A) : + disjoint [:: x] s = (x \notin s). +Proof. +apply/idP/idP. +- by apply: contraL=>X; apply/allPn; exists x=>//; rewrite negbK inE. +by apply: contraR=>/allPn [y H]; rewrite inE negbK =>/eqP <-. +Qed. + +Lemma disjoint1R (A : eqType) x (s : seq A) : + disjoint s [:: x] = (x \notin s). +Proof. by rewrite disjointC disjoint1L. Qed. + +Lemma disjoint_consL (A : eqType) x (s1 s2 : seq A) : + disjoint (x :: s1) s2 = + (x \notin s2) && disjoint s1 s2. +Proof. by rewrite -cat1s disjoint_catL disjoint1L. Qed. + +Lemma disjoint_consR (A : eqType) x (s1 s2 : seq A) : + disjoint s1 (x :: s2) = + (x \notin s1) && disjoint s1 s2. +Proof. by rewrite -cat1s disjoint_catR disjoint1R. Qed. + +Lemma disjoint_consLI (A : eqType) x (s1 s2 : seq A) : + x \notin s2 -> + disjoint s1 s2 -> + disjoint (x :: s1) s2. +Proof. by rewrite disjoint_consL=>->->. Qed. + +Lemma disjoint_consRI (A : eqType) x (s1 s2 : seq A) : + x \notin s1 -> + disjoint s1 s2 -> + disjoint s1 (x :: s2). +Proof. by rewrite disjoint_consR=>->->. Qed. + +Lemma disjoint_consLE (A : eqType) x (s1 s2 : seq A) : + disjoint (x :: s1) s2 -> + (x \notin s2) * (disjoint s1 s2). +Proof. by rewrite disjoint_consL=>/andX. Qed. + +Lemma disjoint_consRE (A : eqType) x (s1 s2 : seq A) : + disjoint s1 (x :: s2) -> + (x \notin s1) * (disjoint s1 s2). +Proof. by rewrite disjoint_consR=>/andX. Qed. + +Lemma disjoint_subL (A : eqType) (s s1 s2 : seq A) : + {subset s2 <= s1} -> + disjoint s s1 -> + disjoint s s2. +Proof. by move=>X /allP H; apply/allP=>z /X /H. Qed. + +Lemma disjoint_subR (A : eqType) (s s1 s2 : seq A) : + {subset s2 <= s1} -> + disjoint s1 s -> + disjoint s2 s. +Proof. +move=>X; rewrite disjointC=>/(disjoint_subL X). +by rewrite disjointC. +Qed. + +Lemma disjoint_eqL {A : eqType} {s s1 s2 : seq A} : + s1 =i s2 -> + disjoint s1 s = disjoint s2 s. +Proof. by move=>X; apply/idP/idP; apply: disjoint_subR=>z; rewrite X. Qed. + +Lemma disjoint_eqR {A : eqType} {s s1 s2 : seq A} : + s1 =i s2 -> + disjoint s s1 = disjoint s s2. +Proof. by move=>X; apply/idP/idP; apply: disjoint_subL=>z; rewrite X. Qed. + +Lemma disjointN (A : eqType) (s1 s2 : seq A) : + ~~ disjoint s1 s2 -> + exists2 x, x \in s1 & x \in s2. +Proof. by case/allPn=>x; rewrite negbK; exists x. Qed. + +(* finding last occurrence of element in a sequence *) Section FindLast. Variables (T : Type). diff --git a/examples/graph.v b/examples/graph.v index 091dcab..aa7965b 100644 --- a/examples/graph.v +++ b/examples/graph.v @@ -14,147 +14,8 @@ limitations under the License. From HB Require Import structures. From Coq Require Import ssreflect ssrfun. From mathcomp Require Import ssrbool eqtype ssrnat seq path. -From pcm Require Import options axioms pred prelude. +From pcm Require Import options axioms pred prelude seqext. From pcm Require Import pcm unionmap natmap autopcm automap. -From pcm Require Import seqext. - -(* subset *) - -Lemma subset_refl (A : Type) (s : pred A) : - {subset s <= s}. -Proof. by []. Qed. - -Lemma subset_trans (A : eqType) (s1 s2 s3 : pred A) : - {subset s1 <= s2} -> - {subset s2 <= s3} -> - {subset s1 <= s3}. -Proof. by move=>S1 S2 z /S1/S2. Qed. - -Lemma subset_consL (A : eqType) x (s1 s2 : seq A) : - {subset x :: s1 <= s2} <-> - x \in s2 /\ {subset s1 <= s2}. -Proof. -split=>[S|[X S]]. -- by split=>[|z Z]; apply: S; rewrite inE ?eqxx ?Z ?orbT. -by move=>z; rewrite inE; case/orP=>[/eqP ->|/S//]. -Qed. - -Lemma subset_consLI (A : eqType) x (s1 s2 : seq A) : - x \in s2 -> - {subset s1 <= s2} -> - {subset x :: s1 <= s2}. -Proof. by move=>H1 H2; rewrite subset_consL. Qed. - -Lemma subset_consR (A : eqType) x (s : seq A) : - {subset s <= x :: s}. -Proof. by move=>z E; rewrite inE E orbT. Qed. - -Lemma subset_consLR (A : eqType) x (s1 s2 : seq A) : - {subset s1 <= s2} -> - {subset x :: s1 <= x :: s2}. -Proof. -move=>X z; rewrite !inE; case/orP=>[|/X] -> //. -by rewrite orbT. -Qed. - -(* disjointness *) - -Definition disjoint {A : eqType} (s1 s2 : seq A) := - all (fun x => x \notin s1) s2. - -Arguments disjoint {A} : simpl never. - -Lemma disjointC (A : eqType) (s1 s2 : seq A) : - disjoint s1 s2 = disjoint s2 s1. -Proof. -apply/idP/idP=>/allP S; apply/allP=>x X; -by apply/negP=>/S; rewrite X. -Qed. - -Lemma disjoint_catR (A : eqType) (s s1 s2 : seq A) : - disjoint s (s1 ++ s2) = - disjoint s s1 && disjoint s s2. -Proof. by rewrite /disjoint all_cat. Qed. - -Lemma disjoint_catL (A : eqType) (s s1 s2 : seq A) : - disjoint (s1 ++ s2) s = - disjoint s1 s && disjoint s2 s. -Proof. by rewrite -!(disjointC s) disjoint_catR. Qed. - -Lemma disjoint1L (A : eqType) x (s : seq A) : - disjoint [:: x] s = (x \notin s). -Proof. -apply/idP/idP. -- by apply: contraL=>X; apply/allPn; exists x=>//; rewrite negbK inE. -by apply: contraR=>/allPn [y H]; rewrite inE negbK =>/eqP <-. -Qed. - -Lemma disjoint1R (A : eqType) x (s : seq A) : - disjoint s [:: x] = (x \notin s). -Proof. by rewrite disjointC disjoint1L. Qed. - -Lemma disjoint_consL (A : eqType) x (s1 s2 : seq A) : - disjoint (x :: s1) s2 = - (x \notin s2) && disjoint s1 s2. -Proof. by rewrite -cat1s disjoint_catL disjoint1L. Qed. - -Lemma disjoint_consR (A : eqType) x (s1 s2 : seq A) : - disjoint s1 (x :: s2) = - (x \notin s1) && disjoint s1 s2. -Proof. by rewrite -cat1s disjoint_catR disjoint1R. Qed. - -Lemma disjoint_consLI (A : eqType) x (s1 s2 : seq A) : - x \notin s2 -> - disjoint s1 s2 -> - disjoint (x :: s1) s2. -Proof. by rewrite disjoint_consL=>->->. Qed. - -Lemma disjoint_consRI (A : eqType) x (s1 s2 : seq A) : - x \notin s1 -> - disjoint s1 s2 -> - disjoint s1 (x :: s2). -Proof. by rewrite disjoint_consR=>->->. Qed. - -Lemma disjoint_consLE (A : eqType) x (s1 s2 : seq A) : - disjoint (x :: s1) s2 -> - (x \notin s2) * (disjoint s1 s2). -Proof. by rewrite disjoint_consL=>/andX. Qed. - -Lemma disjoint_consRE (A : eqType) x (s1 s2 : seq A) : - disjoint s1 (x :: s2) -> - (x \notin s1) * (disjoint s1 s2). -Proof. by rewrite disjoint_consR=>/andX. Qed. - -Lemma disjoint_subL (A : eqType) (s s1 s2 : seq A) : - {subset s2 <= s1} -> - disjoint s s1 -> - disjoint s s2. -Proof. by move=>X /allP H; apply/allP=>z /X /H. Qed. - -Lemma disjoint_subR (A : eqType) (s s1 s2 : seq A) : - {subset s2 <= s1} -> - disjoint s1 s -> - disjoint s2 s. -Proof. -move=>X; rewrite disjointC=>/(disjoint_subL X). -by rewrite disjointC. -Qed. - -Lemma disjoint_eqL {A : eqType} {s s1 s2 : seq A} : - s1 =i s2 -> - disjoint s1 s = disjoint s2 s. -Proof. by move=>X; apply/idP/idP; apply: disjoint_subR=>z; rewrite X. Qed. - -Lemma disjoint_eqR {A : eqType} {s s1 s2 : seq A} : - s1 =i s2 -> - disjoint s s1 = disjoint s s2. -Proof. by move=>X; apply/idP/idP; apply: disjoint_subL=>z; rewrite X. Qed. - -Lemma disjointN (A : eqType) (s1 s2 : seq A) : - ~~ disjoint s1 s2 -> - exists2 x, x \in s1 & x \in s2. -Proof. by case/allPn=>x; rewrite negbK; exists x. Qed. - (*************) (*************) @@ -162,8 +23,8 @@ Proof. by case/allPn=>x; rewrite negbK; exists x. Qed. (*************) (*************) -(* Pregraphs are natmaps, mapping nodes into *) -(* seq node listing the children of the node (adjacency list) *) +(* Pregraphs are natmaps, mapping each graph node x into a node sequence *) +(* that enumerates the children of x (x's adjacency list). *) (* Pregraph differs from graph, in that edges can *dangle*; that is *) (* terminate with a node that isn't in the graph's domain. *) (* Dangling edges are allowed because they enable encoding positional *) @@ -175,6 +36,10 @@ Proof. by case/allPn=>x; rewrite negbK; exists x. Qed. (* If it's desired to treat a non-null dangling link differently *) (* from null, add that link to the graph to make it non-dangling. *) +(* Pregraph differs from fingraph (of mathcomp) *) +(* in that the set of nodes is drawn from an infinite set *) +(* not from a fixed finite set. *) + Notation node := nat. Record pregraph := Pregraph {pregraph_base : @UM.base node nat_pred (seq node)}. @@ -216,11 +81,28 @@ Coercion Pred_of_history (x : pregraph) : Pred_Class := Notation "x &-> v" := (ptsT pregraph x v) (at level 30). +(**************************) +(* Links, children, edges *) +(**************************) + (* Links of x includes all edges outgoing from x *) (* and may explicitly include nodes that aren't in dom g *) (* (i.e., are dangling, null or non-null) *) Definition links (g : pregraph) x := oapp id [::] (find x g). +(* children x removes dangling edges (null or non-null) from links *) +Definition children (g : pregraph) x : seq node := + filter (mem (dom g)) (links g x). + +(* edge is applicative variant of children *) +(* thus, dangling edges (null or non-null) are *not* edges. *) +(* If it's desired to treat a non-null dangling node differently *) +(* from null, add that node to the graph to make it non-dangling. *) +Definition edge g : rel node := mem \o children g. +Arguments edge g x y : simpl never. + +(* links lemmas *) + Lemma links_undef x : links undef x = [::]. Proof. by []. Qed. @@ -271,10 +153,12 @@ Lemma range_links (g : pregraph) x : links g x \in range g. Proof. by move/In_graphX/mem_range. Qed. -(* children x removes dangling edges (null or non-null) from links *) +Lemma links_umfiltk g p x : + links (um_filterk p g) x =i + if p x then links g x else [::]. +Proof. by move=>i; rewrite /links find_umfiltk; case: (p x). Qed. -Definition children (g : pregraph) x : seq node := - filter (mem (dom g)) (links g x). +(* children lemmas *) Lemma children_undef x : children undef x = [::]. Proof. by []. Qed. @@ -318,15 +202,16 @@ Proof. move=>Dx; exists (links g x); first by apply: range_links. by apply: children_links. Qed. - -(* edge is applicative variant of children *) -(* NOTE: Links to dangling edges (null or non-null) *) -(* are explicitly *not* considered edges. *) -(* If it's desired to treat a non-null dangling node differently *) -(* from null, add that node to the graph to make it non-dangling. *) -Definition edge g : rel node := mem \o children g. -Arguments edge g x y : simpl never. +Lemma children_umfiltk g p x y : + y \in children (um_filterk p g) x = + [&& p x, p y & y \in children g x]. +Proof. +rewrite /children !mem_filter /= links_umfiltk dom_umfiltk inE -andbA. +by case: (p x)=>//=; rewrite !andbF. +Qed. + +(* edge lemmas *) Lemma edge_undef x y : edge undef x y = false. Proof. by rewrite /edge/= children_undef. Qed. @@ -350,7 +235,7 @@ Lemma edgeUnR g1 g2 x y : edge (g1 \+ g2) x y. Proof. by move=>V; apply: childrenUnR. Qed. -Lemma edge_dom g x y : +Lemma edgeD g x y : edge g x y -> (x \in dom g) * (y \in dom g). Proof. @@ -369,19 +254,29 @@ Lemma path_dom g x xs : path (edge g) x xs -> {subset xs <= dom g}. Proof. -elim: xs x=>[|a xs IH] x //= /andP [/edge_dom [_ He] /IH]. +elim: xs x=>[|a xs IH] x //= /andP [/edgeD [_ He] /IH]. by apply: subset_consLI He. Qed. -(* lifting the theory of finite graphs to unbounded pregraphs *) +Lemma edge_umfiltk g p x y : + edge (um_filterk p g) x y = + [&& p x, p y & edge g x y]. +Proof. by rewrite /edge /= children_umfiltk. Qed. -(* list of nodes traversed by depth-first search of g *) -(* at depth n, starting from x, and avoiding v *) -(* NOTE: Definition uses children, not links, to avoid *) -(* following dangling edges. *) -(* If it's desired to treat a non-null dangling node differently *) -(* from null, add that node to the graph to make it non-dangling. *) +Lemma edge_umfiltkE g p : + {in predC p &, edge (um_filterk (predC p) g) =2 edge g}. +Proof. by move=>x y; rewrite !inE => X Y; rewrite edge_umfiltk /= X Y. Qed. + +(***********************) +(* Depth-first search *) +(***********************) +(* lifts dfs from mathcomp fingraph to pregraphs *) + +(* list of nodes traversed by depth-first search of g *) +(* at depth n, starting from x, and avoiding v. *) +(* Definition uses children, not links; *) +(* thus, it doesn't follow dangling edges *) Fixpoint dfs (g : pregraph) (n : nat) (v : seq node) x := if (x \notin dom g) || (x \in v) then v else if n is n'.+1 then foldl (dfs g n') (x :: v) (children g x) else v. @@ -487,7 +382,7 @@ have : {subset a <= dom g} by apply: childrenD. elim: {x v}a (x) v1=>[|x xs IHa] x' v /= Dxs Hv U Sv Nv Dx'. - by rewrite (negbTE Nv); apply: ReflectF; case. have Dx : x \in dom g by apply: Dxs; rewrite inE eqxx. -have Da : {subset xs <= dom g} by apply/subset_trans/Dxs/subset_consR. +have Da : {subset xs <= dom g} by move=>z Z; apply/Dxs/subset_consR. set v2 := dfs g n v x. have Sv2 : {subset v <= v2} := @subset_dfs g n v [:: x]. have [Hy2|Ny2] := boolP (y \in v2). @@ -518,6 +413,11 @@ apply/IHn=>//; exists (p ++ pr). by rewrite -cat_cons disjoint_catR D (disjoint_consRE D2). Qed. +(******************) +(* Connectivity *) +(* (reachability) *) +(******************) + (* start dfs with p as avoidance set, then filter out p. *) (* this computes only the nodes visited during dfs. *) (* not for client use, hence primed name *) @@ -525,10 +425,12 @@ Definition component' (p : pred node) g x : seq node := filter (predC p) (dfs g (size (dom g)) (filter p (dom g)) x). (* y is connected from x if y is visited during dfs *) -(* under assumption that x is in dom g *) +(* avoiding nodes from p; assuming x in dom g *) Definition connect p (g : pregraph) x : pred node := fun y => (x \in dom g) && (y \in component' p g x). +(* connect lemmas *) + Lemma connectP (p : pred node) (g : pregraph) x y : reflect [/\ x \in dom g, ~~ p x & exists xs, [/\ path (edge g) x xs, y = last x xs & @@ -580,15 +482,12 @@ case/and3P: Px=>Px Ex Pp2; exists p2; split=>//. by move=>z Z; rewrite Hxs // mem_cat Z orbT. Qed. - - Lemma connect_undef p x : connect p undef x =i pred0. Proof. by move=>y; apply/connectP; case; rewrite dom_undef. Qed. Lemma connect_unit p x : connect p Unit x =i pred0. Proof. by move=>y; apply/connectP; case; rewrite dom0. Qed. - Lemma connectDx p g x y : y \in connect p g x -> (x \in dom g) * ~~ p x. @@ -676,6 +575,39 @@ Lemma connectUnR p g1 g2 x : connect p (g1 \+ g2) x}. Proof. by rewrite joinC; apply: connectUnL. Qed. +Lemma connectUn p g1 g2 x : + [pcm g1 <= g2] -> + valid g2 -> + {subset connect p g1 x <= connect p g2 x}. +Proof. by case=>g ->; apply: connectUnL. Qed. + +Lemma connect_avoid_graph g p x : + connect p g x =i + connect (predU p (predC (mem (dom g)))) g x. +Proof. +move=>y; apply/connectP/connectP; last first. +- case=>Dx; rewrite !inE negb_or Dx andbT=>Hx [xs][Px Ey Hxs]. + split=>//; exists xs; split=>// z /Hxs. + by rewrite !inE negb_or; case/andP. +case=>Dx Hx [xs][Px Ey Hxs]; split=>//. +- by rewrite !inE negb_or Hx Dx. +exists xs; split=>// z Z; rewrite !inE negb_or Hxs //= negbK. +by apply: path_dom Px _ Z. +Qed. + +Lemma connect_umfiltk g p x : + connect p (um_filterk (predC p) g) x =i connect p g x. +Proof. +case: (normalP g)=>[->|V y]; first by rewrite pfundef. +apply/connectP/connectP; case; rewrite dom_umfiltk inE /=. +- case/andP=>Hx Dx _ [xs][Px Ey Hxs]; split=>//. + exists xs; rewrite -(eq_in_path (@edge_umfiltkE g p)) //. + by apply/allP/in_consE. +move=>Dx Hx [xs][Px Ey Hxs]; rewrite Dx Hx; split=>//. +exists xs; rewrite (eq_in_path (@edge_umfiltkE g p)) //. +by apply/allP/in_consE. +Qed. + Lemma connect_sub g (p1 p2 : pred node) x : subpred p2 p1 -> {subset connect p1 g x <= connect p2 g x}. @@ -690,8 +622,27 @@ Lemma connect_eq g p1 p2 x : connect p1 g x =i connect p2 g x. Proof. by move=>S y; apply/idP/idP; apply/connect_sub=>z; rewrite S. Qed. +(* relativized versions of connect_sub and connect_eq *) +(* it suffices to only prove the precondition for nodes in g *) +Lemma connect_in_sub g p1 p2 x : + {in dom g, subpred p2 p1} -> + {subset connect p1 g x <= connect p2 g x}. +Proof. +move=>S y; rewrite connect_avoid_graph=>C. +rewrite connect_avoid_graph. +apply: connect_sub C=>z; rewrite !inE -!(orbC (z \notin _)). +by case Dz: (z \in dom g)=>//=; apply: S Dz. +Qed. +Lemma connect_in_eq g p1 p2 x : + {in dom g, p1 =1 p2} -> + connect p1 g x =i connect p2 g x. +Proof. +move=>S y; rewrite [LHS]connect_avoid_graph [RHS]connect_avoid_graph. +apply: connect_eq=>z; rewrite !inE -!(orbC (z \notin _)). +by case Dz: (z \in dom g)=>//=; rewrite S. +Qed. (* deconstructing connecting path from left *) Lemma edge_connect p g x y : @@ -701,7 +652,7 @@ Lemma edge_connect p g x y : Proof. move/eqP=>Nxy /connectP [Dx H][[|a xs]][/= Px Ey Hxs]; first by move/Nxy: Ey. case/andP: Px=>Px Pxs; exists a=>//; apply/connectP; split. -- by rewrite (edge_dom Px). +- by rewrite (edgeD Px). - by apply: Hxs; rewrite inE eqxx. exists xs; split=>// z Z; apply: Hxs. by rewrite inE Z orbT. @@ -751,65 +702,14 @@ exists (rcons xs z); split=>/=. by move=>w; rewrite mem_rcons inE; case/orP=>[/eqP ->|/Hxs]. Qed. -(**********) -(**********) -(* Graphs *) -(**********) -(**********) - -(* x is in_node if it's in the graph or is null *) -Definition in_node (g : pregraph) (x : node) := - (x == null) || (x \in dom g). - -(* x is out-node if no edge goes into it *) -Definition out_node (g : pregraph) (x : node) := - {in range g, forall xs, x \notin xs}. - -(* pregraph is graph if *) -(* all nodes in all adjacency lists are in-nodes *) -Definition graph_axiom (g : pregraph) := - forall (xs : seq node) (x : node), xs \in range g -> x \in xs -> in_node g x. - -HB.mixin Record isGraph (g : pregraph) := { - graph_subproof : graph_axiom g}. -#[short(type=graph)] -HB.structure Definition Graph := {g of isGraph g }. - -(* removing out-node causes no dangling edges; *) -(* hence preserves graph axiom *) -Lemma graphF g x : - out_node g x -> - graph_axiom g -> - graph_axiom (free g x). -Proof. -move=>Hna Ha xs q /rangeF Hs Hq. -move: (Ha xs q Hs Hq) (Hna _ Hs)=>{}Hs. -rewrite /in_node domF inE in Hs *. -by case: (x =P q) Hq=>//= ->->. -Qed. +(***********************) +(* Connected component *) +(***********************) (* part of g reachable from x (avoiding nothing) *) Definition subconnect g x : pregraph := um_filterk (mem (connect pred0 g x)) g. -(* reachable component of a graph is a graph *) -Lemma graph_subconnect g x : - valid g -> - graph_axiom g -> - graph_axiom (subconnect g x). -Proof. -have E : g = subconnect g x \+ um_filterk - (negb \o connect pred0 g x) g by apply: umfilt_predC. -move=>V Ha xs /= n Hxs Hn; have {}Dn : in_node g n. -- by apply: Ha Hn; rewrite E rangeUn inE -E V Hxs. -case/mem_rangeX: Hxs=>k /In_umfiltX [/= Ck] /In_graph Hf. -rewrite /in_node in Dn *; case/boolP: (n == null) Dn=>//= Hnn Dn. -case: (connectD Ck)=>_ Dk. -rewrite /subconnect dom_umfiltk inE /= Dn andbT. -apply: connect_trans Ck _; apply/connectP; split=>//. -by exists [:: n]; split=>//=; rewrite edge_links Dn -Hf Hn. -Qed. - Lemma edge_subconnect g x y z : edge (subconnect g x) y z = [&& y \in connect pred0 g x, @@ -859,7 +759,7 @@ Proof. move/eqP=>Nxy /connectX [Dx Hx][[|a xs]] /= [Px Ey Nx Hxs]. - by move/Nxy: Ey. case/andP: Px=>Exa Pa; exists a=>//; apply/connectP; split. -- by rewrite (edge_dom Exa). +- by rewrite (edgeD Exa). - rewrite !inE negb_or Hxs ?inE ?eqxx // andbT. by case: (a =P x) Nx=>//= ->; rewrite inE eqxx. exists xs; split=>// z Z. @@ -901,11 +801,13 @@ Qed. (* y is reachable from x by a path that avoids whole *) (* subcomponent of x' *) Lemma connect_avoid p g x y x' : - y \in connect p g x -> y \notin connect p g x' -> - y \in connect (predU p [in connect p g x']) g x. + (y \in connect p g x) = + (y \in connect (predU p [in connect p g x']) g x). Proof. -case/connectP=>Dx Hx [xs][Px Ey Hxs] Ny. +move=>Ny; apply/idP/idP; last first. +- by apply: connect_sub=>z Z; rewrite inE Z. +case/connectP=>Dx Hx [xs][Px Ey Hxs]. have [Nx|Mx] := boolP (has [in connect p g x'] (x :: xs)); last first. (* if path contains no nodes reachable from x', nothing to do *) - move/hasPn: Mx=>Mx; apply/connectP; split=>//. @@ -926,70 +828,24 @@ case/and3P: Px=>Px Ea P2; apply/connectP; rewrite !(connectDy Ca); split=>//. by exists p2; split=>// z Z; rewrite Hxs // mem_cat Z orbT. Qed. -(***************) -(* N-pregraphs *) -(***************) - -(* pregraphs with global bound n *) -(* for the number of links of a node *) - -Definition n_pregraph_axiom (n : nat) (g : pregraph) := - {in dom g, forall x, size (links g x) = n}. - -Lemma npregraphP n g : - n_pregraph_axiom n g <-> - {in range g, forall xs, size xs = n}. -Proof. -split=>H x; last by move=>Dx; apply: H (range_links Dx). -by case/mem_rangeX=>k X; rewrite (In_graph X); apply: H (In_dom X). -Qed. - -Lemma npregraphUnL n g1 g2 : - valid (g1 \+ g2) -> - n_pregraph_axiom n (g1 \+ g2) -> - n_pregraph_axiom n g1. +Lemma connect_avoid1 p g x y x' : + y \notin connect p g x' -> + (y \in connect p g x) = (y \in connect (predU (pred1 x') p) g x). Proof. -by rewrite !npregraphP=>V H x X; apply: H; rewrite rangeUn inE V X. -Qed. - -Lemma npregraphUnR n g1 g2 : - valid (g1 \+ g2) -> - n_pregraph_axiom n (g1 \+ g2) -> - n_pregraph_axiom n g2. -Proof. by rewrite joinC; apply: npregraphUnL. Qed. - -Lemma npregraphF n g x : - n_pregraph_axiom n g -> - n_pregraph_axiom n (free g x). -Proof. by rewrite !npregraphP=>H z /rangeF; apply: H. Qed. - -Definition get_nth g x := nth null (links g x). - -Lemma innode_nth g x n : - graph_axiom g -> - in_node g (get_nth g x n). -Proof. -rewrite /graph_axiom => Hg. -case: (ltnP n (size (links g x)))=>Hm; last first. -- by rewrite /in_node/get_nth nth_default. -have Dx : x \in dom g by apply/size_links/gt0/Hm. -by apply: Hg (range_links Dx) _; rewrite mem_nth. -Qed. - -Lemma links_nth n g x : - n_pregraph_axiom n g -> - x \in dom g -> - links g x = map (get_nth g x) (iota 0 n). -Proof. -move=>H Dx; apply/(eq_from_nth (x0:=null))=>[|i Hi]. -- by rewrite size_map size_iota H. -by rewrite map_nth_iota0 ?H // -(H _ Dx) take_size. +move=>C; apply/idP/idP; last first. +- by apply: connect_sub=>z Z; rewrite inE Z orbT. +rewrite (connect_avoid _ C); apply/connect_in_sub=>z Dz. +by rewrite !inE=>/orP [/eqP <-|->//]; rewrite connect0 Dz orbN. Qed. (******************) (* Biconnectivity *) (******************) +(* symmetric closure of connectivity *) + +(* y is biconnected from x iff *) +(* x and y are mutually connected *) Definition biconnect p g x : pred node := [pred y | (x \in connect p g y) && (y \in connect p g x)]. @@ -1023,10 +879,13 @@ Lemma biconnect_eq g (p1 p2 : pred node) (x : node) : biconnect p1 g x =i biconnect p2 g x. Proof. by move=>S y; rewrite !inE !(connect_eq g _ S). Qed. +(**********) +(* Cycles *) +(**********) + (* elements in a cycle are interconnected *) (* avoiding everything outside the cycle *) -(* NOTE: NOTE: NOTE: this is connect lemma, not biconnect *) Lemma connect_cycle g xs : cycle (edge g) xs -> {in xs &, forall x y, y \in connect [predC xs] g x}. @@ -1035,7 +894,7 @@ move=>C x y /rot_to [i q Hr]; rewrite -(mem_rot i) Hr => Hy. have Hx : x \in xs by rewrite -(mem_rot i) Hr inE eqxx. have /= Hp1: cycle (edge g) (x :: q) by rewrite -Hr rot_cycle. have Dx : x \in dom g. -- by move: Hp1; rewrite rcons_path=>/andP [_ /edge_dom][]. +- by move: Hp1; rewrite rcons_path=>/andP [_ /edgeD][]. case/splitPl: Hy Hp1 Hr=>r s Ey. rewrite rcons_cat cat_path=>/andP [Hxr]. rewrite Ey rcons_path; case/andP=>Hlx /= Ex Hr. @@ -1164,7 +1023,7 @@ apply: (iffP idP)=>[|H]; last first. case/andP=>/preacyclicP Ng /allP Ne x xs Dx /=. rewrite rcons_path; apply/negP=>/andP []. case: xs=>[_|y xs /= /andP [Exy Px]]; first by apply/negP/Ne. -have Dy : y \in dom g by rewrite (edge_dom Exy). +have Dy : y \in dom g by rewrite (edgeD Exy). have : y \notin biconnect pred0 g x. - by apply: contraL Exy=>/(Ng x y Dx Dy) <-; apply: Ne Dx. apply: contraNnot=>Ex; apply: (biconnect_cycle0 (xs:=x::y::xs))=>/=. @@ -1182,5 +1041,147 @@ have [Dx|Nx] := boolP (x \in dom g); last by rewrite linksND. by apply: contra (H _ Dx)=>Lx; rewrite edge_links Dx Lx. Qed. +(***************) +(* N-pregraphs *) +(***************) + +(* getting the n-th link *) +Definition get_nth g x := nth null (links g x). + +Lemma getnth_mem0 g x n : + (get_nth g x n == null) || + (get_nth g x n \in links g x). +Proof. +case: (ltnP n (size (links g x)))=>Hm; last first. +- by rewrite /get_nth nth_default. +by rewrite mem_nth // orbT. +Qed. + +Lemma getnth_mem g x n : + get_nth g x n != null -> + get_nth g x n \in links g x. +Proof. by move=>H; move: (getnth_mem0 g x n); rewrite (negbTE H). Qed. + +(* n-pregraph is pregraph with global bound n *) +(* for the number of links of a node *) + +Definition n_pregraph_axiom (n : nat) (g : pregraph) := + {in dom g, forall x, size (links g x) = n}. + +Lemma npregraphP n g : + n_pregraph_axiom n g <-> + {in range g, forall xs, size xs = n}. +Proof. +split=>H x; last by move=>Dx; apply: H (range_links Dx). +by case/mem_rangeX=>k X; rewrite (In_graph X); apply: H (In_dom X). +Qed. + +Lemma npregraphUnL n g1 g2 : + valid (g1 \+ g2) -> + n_pregraph_axiom n (g1 \+ g2) -> + n_pregraph_axiom n g1. +Proof. +by rewrite !npregraphP=>V H x X; apply: H; rewrite rangeUn inE V X. +Qed. + +Lemma npregraphUnR n g1 g2 : + valid (g1 \+ g2) -> + n_pregraph_axiom n (g1 \+ g2) -> + n_pregraph_axiom n g2. +Proof. by rewrite joinC; apply: npregraphUnL. Qed. + +Lemma npregraphF n g x : + n_pregraph_axiom n g -> + n_pregraph_axiom n (free g x). +Proof. by rewrite !npregraphP=>H z /rangeF; apply: H. Qed. + +Lemma links_nth n g x : + n_pregraph_axiom n g -> + x \in dom g -> + links g x = map (get_nth g x) (iota 0 n). +Proof. +move=>H Dx; apply/(eq_from_nth (x0:=null))=>[|i Hi]. +- by rewrite size_map size_iota H. +by rewrite map_nth_iota0 ?H // -(H _ Dx) take_size. +Qed. + +(**********) +(**********) +(* Graphs *) +(**********) +(**********) + +(* x is in_node if it's in the graph or is null *) +Definition in_node (g : pregraph) (x : node) := + (x == null) || (x \in dom g). + +(* x is out-node if no edge goes into it *) +Definition out_node (g : pregraph) (x : node) := + {in range g, forall xs, x \notin xs}. + +(* pregraph is graph if *) +(* all nodes in all adjacency lists are in-nodes *) +Definition graph_axiom (g : pregraph) := + forall xs x, xs \in range g -> x \in xs -> in_node g x. + +HB.mixin Record isGraph (g : pregraph) := { + graph_subproof : graph_axiom g}. +#[short(type=graph)] +HB.structure Definition Graph := {g of isGraph g }. + +(* removing out-node causes no dangling edges; *) +(* hence preserves graph axiom *) +Lemma graphF g x : + out_node g x -> + graph_axiom g -> + graph_axiom (free g x). +Proof. +move=>Hna Ha xs q /rangeF Hs Hq. +move: (Ha xs q Hs Hq) (Hna _ Hs)=>{}Hs. +rewrite /in_node domF inE in Hs *. +by case: (x =P q) Hq=>//= ->->. +Qed. + +(* reachable component of a graph is a graph *) +Lemma graph_subconnect g x : + valid g -> + graph_axiom g -> + graph_axiom (subconnect g x). +Proof. +have E : g = subconnect g x \+ um_filterk + (negb \o connect pred0 g x) g by apply: umfilt_predC. +move=>V Ha xs /= n Hxs Hn; have {}Dn : in_node g n. +- by apply: Ha Hn; rewrite E rangeUn inE -E V Hxs. +case/mem_rangeX: Hxs=>k /In_umfiltX [/= Ck] /In_graph Hf. +rewrite /in_node in Dn *; case/boolP: (n == null) Dn=>//= Hnn Dn. +case: (connectD Ck)=>_ Dk. +rewrite /subconnect dom_umfiltk inE /= Dn andbT. +apply: connect_trans Ck _; apply/connectP; split=>//. +by exists [:: n]; split=>//=; rewrite edge_links Dn -Hf Hn. +Qed. + +Lemma graph_links g x y : + graph_axiom g -> + y \in links g x -> + in_node g y. +Proof. by move=>H /[dup]/linksD Dx; apply/H/range_links. Qed. + +Lemma graph_children g x y : + graph_axiom g -> + (y \in children g x) = (y != null) && (y \in links g x). +Proof. +move=>H; rewrite mem_filter -!(andbC (y \in links _ _)). +case Ly : (y \in links g x)=>//=. +move/(graph_links H): Ly; rewrite /in_node. +by case: (y =P null)=>//= ->; rewrite cond_dom. +Qed. +Lemma graph_getnth g x n : + graph_axiom g -> + in_node g (get_nth g x n). +Proof. +case: (ltnP n (size (links g x)))=>Hm; last first. +- by rewrite /get_nth nth_default. +by move/(graph_links (x:=x)); apply; rewrite mem_nth. +Qed. diff --git a/examples/quicksort.v b/examples/quicksort.v index 586a164..f9157e7 100644 --- a/examples/quicksort.v +++ b/examples/quicksort.v @@ -634,4 +634,4 @@ Qed. End LomutoQsort. -(* TODO Hoare variant *) + diff --git a/pcm/natmap.v b/pcm/natmap.v index 918ab83..1e732f4 100644 --- a/pcm/natmap.v +++ b/pcm/natmap.v @@ -106,6 +106,51 @@ Lemma xx : (1, 3) \In (1 \-> 3). Abort. (* /DEVCOMMENT *) +(* Sometimes it's useful to not think of natmaps as recording *) +(* times (i.e., being histories), but just as ordinary maps *) +(* over nats. For that occassion, use the nmap type *) + +Record nmap A := NMap {nmap_base : @UM.base nat nat_pred A}. + +Section NMapUMC. +Variables A : Type. +Implicit Type f : nmap A. +Local Coercion nmap_base : nmap >-> UM.base. +Let nm_valid f := @UM.valid nat nat_pred A f. +Let nm_empty := NMap (@UM.empty nat nat_pred A). +Let nm_undef := NMap (@UM.Undef nat nat_pred A). +Let nm_upd k v f := NMap (@UM.upd nat nat_pred A k v f). +Let nm_dom f := @UM.dom nat nat_pred A f. +Let nm_assocs f := @UM.assocs nat nat_pred A f. +Let nm_free f k := NMap (@UM.free nat nat_pred A f k). +Let nm_find k f := @UM.find nat nat_pred A k f. +Let nm_union f1 f2 := NMap (@UM.union nat nat_pred A f1 f2). +Let nm_empb f := @UM.empb nat nat_pred A f. +Let nm_undefb f := @UM.undefb nat nat_pred A f. +Let nm_from (f : nmap A) : UM.base _ _ := f. +Let nm_to (b : @UM.base nat nat_pred A) : nmap A := NMap b. +Let nm_pts k v := NMap (@UM.pts nat nat_pred A k v). + +Lemma nmap_is_umc : + union_map_axiom nm_valid nm_empty nm_undef nm_upd nm_dom + nm_assocs nm_free nm_find nm_union nm_empb + nm_undefb nm_pts nm_from nm_to. +Proof. by split=>//; split=>[|[]]. Qed. + +HB.instance Definition _ := + isUnion_map.Build nat nat_pred A (nmap A) nmap_is_umc. +End NMapUMC. + +HB.instance Definition _ A := isNatMap.Build A (nmap A). +HB.instance Definition _ (A : eqType) := + hasDecEq.Build (nmap A) (@union_map_eqP nat _ A (nmap A)). +Canonical nmap_PredType A : PredType (nat * A) := + Mem_UmMap_PredType (nmap A). +Coercion Pred_of_nmap A (x : nmap A) : Pred_Class := + [eta Mem_UmMap x]. + +(* no points-to notation for nmaps *) + (*************) (* Main defs *) (*************) From 7eb1ccf5216aac8e684b84ad22052f3bcb236bd4 Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Fri, 9 Aug 2024 17:48:31 +0200 Subject: [PATCH 26/93] committed schorr --- examples/schorr.v | 144 ++++++++++++++++++++++++++++++++++++++++++++++ 1 file changed, 144 insertions(+) create mode 100644 examples/schorr.v diff --git a/examples/schorr.v b/examples/schorr.v new file mode 100644 index 0000000..b7570e7 --- /dev/null +++ b/examples/schorr.v @@ -0,0 +1,144 @@ +(* +Copyright 2024 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 HB Require Import structures. +From Coq Require Import ssreflect ssrfun. +From mathcomp Require Import eqtype ssrnat ssrbool seq path bigop. +From pcm Require Import options axioms pred prelude seqperm seqext. +From pcm Require Import pcm unionmap natmap heap autopcm automap. +From htt Require Import options model heapauto graph. + +(*****************) +(* helper lemmas *) +(*****************) + +Lemma eq_connect_aux (g1 g2 : pregraph) p: + valid (g1 \+ g2) -> + {subset dom g2 <= predC p} -> + um_filterk p (g1 \+ g2) = um_filterk p g1. +Proof. +move=>V /(umfiltk_subdom0 p (validR V)) S. +by rewrite umfiltUn // S unitR. +Qed. + +Lemma connectUn_eq (g g1 g2: pregraph) (p : pred node) x : + valid (g \+ g1) -> + valid (g \+ g2) -> + {subset dom g1 <= p} -> + {subset dom g2 <= p} -> + connect p (g \+ g1) x =i connect p (g \+ g2) x. +Proof. +move=>V1 V2 S1 S2 z. +rewrite -connect_umfiltk eq_connect_aux //=; last first. +- by move=>y /S1; rewrite inE negbK. +rewrite -[RHS]connect_umfiltk eq_connect_aux //. +by move=>y /S2; rewrite !inE negbK. +Qed. + +(****************) +(* Schorr-Waite *) +(****************) + +(* short notation for left/right child *) +Notation gl g x := (get_nth g x 0). +Notation gr g x := (get_nth g x 1). + +(* type of markings *) +Inductive mark := U | L | R | LR. +(* decidable equality on mark *) +Definition eq_mark l1 l2 : bool := + if (l1, l2) is ((U, U)|(L, L)|(R, R)|(LR, LR)) then true else false. +Lemma eq_markP : Equality.axiom eq_mark. +Proof. by case; case=>//=; constructor. Qed. +HB.instance Definition _ := hasDecEq.Build mark eq_markP. + +(* marking of children *) + +(* given marking map m, x is m-child of y iff: *) +(* - x is left child of y and m y = L *) +(* - x is right child of y and m y = R *) + +Definition ch (m : nmap mark) (g : pregraph) (x y : nat) := + [|| (find y m == Some L) && (gl g y == x) | + (find y m == Some R) && (gr g y == x)]. + +Lemma chP (m : nmap mark) g x y : + reflect [\/ (y, L) \In m /\ gl g y = x | + (y, R) \In m /\ gr g y = x] + (ch m g x y). +Proof. +rewrite /ch; case: orP=>H; constructor. +- by case: H=>H; [left|right]; case/andP: H=>/eqP/In_find H /eqP. +by case=>[][/In_find/eqP M /eqP G]; apply: H; [left|right]; apply/andP. +Qed. + +Lemma chN0 m g x y : + ch m g x y -> + y != null. +Proof. by case/chP=>[][/In_dom/dom_cond]. Qed. + +Lemma ch_fun m g x y z : + ch m g x y -> + ch m g z y -> + x = z. +Proof. +case/chP=>[][M1 <-] /chP [][M2 <-] //; +by move/(In_fun M1): M2. +Qed. + +Lemma ch_path m g s x : + path (ch m g) null s -> + x \in s -> + exists y, y \in belast null s /\ ch m g y x. +Proof. +elim/last_ind: s=>//= s z IH. +rewrite rcons_path mem_rcons belast_rcons !inE. +case/andP=>H1 H2 /orP [/eqP E|]. +- by exists (last null s); rewrite mem_last E. +by case/(IH H1)=>y [H3 H4]; exists y; rewrite mem_belast. +Qed. + +Lemma ch_path_uniq m g s : + path (ch m g) null s -> + uniq (null :: s). +Proof. +elim/last_ind: s=>[|s x IH] //=. +rewrite rcons_path mem_rcons=>/andP [Px Cx]. +move: {IH}(IH Px) (IH Px); rewrite {1}lastI /=. +rewrite !rcons_uniq inE negb_or=>/andP [N _]/andP [->->]. +rewrite !andbT eq_sym (chN0 Cx) /=. +apply/negP=>/(ch_path Px) [/= y][/[swap]/(ch_fun Cx) <-]. +by rewrite (negbTE N). +Qed. + +Definition upd_edge (m : nmap mark) g x y : seq node := + if find x m is Some L then [:: y; gr g x] else [:: gl g x; y]. + +Fixpoint rev_graph' m g ps t : pregraph := + if ps is p::ps then + rev_graph' m (free g p) ps p \+ pts p (upd_edge m g p t) + else g. + +Definition rev_graph m g S t := rev_graph' m g (rev S) t. + +(* turning (pre)graph into a heap *) +Definition hp (g : pregraph) : heap := + \big[join/Unit]_(x <- dom g) (x :-> (gl g x, gr g x)). + +(* Aleks: not sure what this definition intends to say *) +(* but at least it typechecks now *) +Definition reach (m : nmap mark) (g : pregraph) (S : seq node) t := + forall x, x \notin dom m -> + x \in connect (mem (dom m)) g t \/ + exists y, y \in S /\ x \in connect (mem (dom m)) g y. + From bcd41587b076b39ff41d9eae624c74351dedcd66 Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Fri, 9 Aug 2024 17:51:36 +0200 Subject: [PATCH 27/93] minor --- examples/schorr.v | 5 +---- 1 file changed, 1 insertion(+), 4 deletions(-) diff --git a/examples/schorr.v b/examples/schorr.v index b7570e7..7ea0626 100644 --- a/examples/schorr.v +++ b/examples/schorr.v @@ -91,10 +91,7 @@ Lemma ch_fun m g x y z : ch m g x y -> ch m g z y -> x = z. -Proof. -case/chP=>[][M1 <-] /chP [][M2 <-] //; -by move/(In_fun M1): M2. -Qed. +Proof. by case/chP=>[][H <-] /chP [][/(In_fun H) {}H <-]. Qed. Lemma ch_path m g s x : path (ch m g) null s -> From 75722a2e1fd363a136ac6f82ba7e4664edfe1c9c Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Fri, 9 Aug 2024 17:57:27 +0200 Subject: [PATCH 28/93] minor --- examples/schorr.v | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/examples/schorr.v b/examples/schorr.v index 7ea0626..1cd912c 100644 --- a/examples/schorr.v +++ b/examples/schorr.v @@ -113,7 +113,7 @@ elim/last_ind: s=>[|s x IH] //=. rewrite rcons_path mem_rcons=>/andP [Px Cx]. move: {IH}(IH Px) (IH Px); rewrite {1}lastI /=. rewrite !rcons_uniq inE negb_or=>/andP [N _]/andP [->->]. -rewrite !andbT eq_sym (chN0 Cx) /=. +rewrite eq_sym (chN0 Cx) !andbT /=. apply/negP=>/(ch_path Px) [/= y][/[swap]/(ch_fun Cx) <-]. by rewrite (negbTE N). Qed. From 6b3c09dab7466d5ed58d0325007c348503a33422 Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Fri, 9 Aug 2024 18:01:14 +0200 Subject: [PATCH 29/93] minor --- examples/schorr.v | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) diff --git a/examples/schorr.v b/examples/schorr.v index 1cd912c..be7cb1a 100644 --- a/examples/schorr.v +++ b/examples/schorr.v @@ -55,7 +55,7 @@ Notation gr g x := (get_nth g x 1). (* type of markings *) Inductive mark := U | L | R | LR. -(* decidable equality on mark *) +(* decidable equality on marks *) Definition eq_mark l1 l2 : bool := if (l1, l2) is ((U, U)|(L, L)|(R, R)|(LR, LR)) then true else false. Lemma eq_markP : Equality.axiom eq_mark. @@ -65,8 +65,8 @@ HB.instance Definition _ := hasDecEq.Build mark eq_markP. (* marking of children *) (* given marking map m, x is m-child of y iff: *) -(* - x is left child of y and m y = L *) -(* - x is right child of y and m y = R *) +(* - m y = L and x is left child of y *) +(* - m y = R and x is right child of y *) Definition ch (m : nmap mark) (g : pregraph) (x y : nat) := [|| (find y m == Some L) && (gl g y == x) | @@ -122,7 +122,7 @@ Definition upd_edge (m : nmap mark) g x y : seq node := if find x m is Some L then [:: y; gr g x] else [:: gl g x; y]. Fixpoint rev_graph' m g ps t : pregraph := - if ps is p::ps then + if ps is p :: ps then rev_graph' m (free g p) ps p \+ pts p (upd_edge m g p t) else g. From 40ef923418f76a6d122fd698cc7faf0bfd3baeab Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Fri, 9 Aug 2024 18:06:26 +0200 Subject: [PATCH 30/93] minor --- examples/schorr.v | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/examples/schorr.v b/examples/schorr.v index be7cb1a..c9ba4a9 100644 --- a/examples/schorr.v +++ b/examples/schorr.v @@ -126,7 +126,7 @@ Fixpoint rev_graph' m g ps t : pregraph := rev_graph' m (free g p) ps p \+ pts p (upd_edge m g p t) else g. -Definition rev_graph m g S t := rev_graph' m g (rev S) t. +Definition rev_graph m g s t := rev_graph' m g (rev s) t. (* turning (pre)graph into a heap *) Definition hp (g : pregraph) : heap := @@ -134,8 +134,8 @@ Definition hp (g : pregraph) : heap := (* Aleks: not sure what this definition intends to say *) (* but at least it typechecks now *) -Definition reach (m : nmap mark) (g : pregraph) (S : seq node) t := +Definition reach (m : nmap mark) (g : pregraph) (s : seq node) t := forall x, x \notin dom m -> x \in connect (mem (dom m)) g t \/ - exists y, y \in S /\ x \in connect (mem (dom m)) g y. + exists y, y \in s /\ x \in connect (mem (dom m)) g y. From e0c6c22a514e4d7fe0aa0a54a50e7b48fc4831fe Mon Sep 17 00:00:00 2001 From: Aleks Nanevski Date: Fri, 9 Aug 2024 19:32:36 +0200 Subject: [PATCH 31/93] moved some helper lemmas to seqext.v --- core/seqext.v | 36 ++++++++++++++++++++++++++++++++++++ examples/schorr.v | 30 ++++++++---------------------- 2 files changed, 44 insertions(+), 22 deletions(-) diff --git a/core/seqext.v b/core/seqext.v index 174b363..df10cf5 100644 --- a/core/seqext.v +++ b/core/seqext.v @@ -1480,6 +1480,42 @@ Lemma path_lastR (s : seq A) leT x : leT x (last x s). Proof. by move=>R T P; case: eqP (path_last T P)=>// <- _; apply: R. Qed. +Lemma path_prev leT s a x : + x \in s -> + path leT a s -> + exists y, y \in belast a s /\ leT y x. +Proof. +case/splitPr=>p1 p2; rewrite cat_path /= =>/and3P []. +by exists (last a p1); rewrite belast_cat /= mem_cat inE eqxx orbT. +Qed. + +Lemma path_next leT s a x b : + x \in a :: s -> + path leT a (rcons s b) -> + exists y, y \in rcons s b /\ leT x y. +Proof. +case/splitPl=>p1 p2 Ex; rewrite rcons_cat cat_path rcons_path Ex. +case/and3P=>H1 H2 H3; case: p2 H2 H3=>[|y p2] /= H2 H3. +- by exists b; rewrite mem_cat inE eqxx orbT. +case/andP: H2=>H2 _; exists y; split=>//. +by rewrite mem_cat inE eqxx orbT. +Qed. + +Lemma path_uniq leT a s : + (forall x y, leT x y -> y != a) -> + (forall a b x, leT a x -> leT b x -> a = b) -> + path leT a s -> + uniq (a :: s). +Proof. +move=>Na Fp; elim/last_ind: s=>[|s x IH] //=. +rewrite rcons_path mem_rcons=>/andP [Px Cx]. +move: {IH}(IH Px) (IH Px); rewrite {1}lastI /=. +rewrite !rcons_uniq inE negb_or=>/andP [N _]/andP [->->]. +rewrite eq_sym (Na _ _ Cx) !andbT /=; apply/negP. +move/path_prev=>/(_ _ _ Px) [/= y][/[swap]] /(Fp _ _ _ Cx) <-. +by rewrite (negbTE N). +Qed. + (* in a sorted list, the last element is maximal *) (* and the maximal element is last *) diff --git a/examples/schorr.v b/examples/schorr.v index c9ba4a9..9ee4468 100644 --- a/examples/schorr.v +++ b/examples/schorr.v @@ -31,7 +31,7 @@ move=>V /(umfiltk_subdom0 p (validR V)) S. by rewrite umfiltUn // S unitR. Qed. -Lemma connectUn_eq (g g1 g2: pregraph) (p : pred node) x : +Lemma connectUn_eq (g g1 g2 : pregraph) (p : pred node) x : valid (g \+ g1) -> valid (g \+ g2) -> {subset dom g1 <= p} -> @@ -87,36 +87,22 @@ Lemma chN0 m g x y : y != null. Proof. by case/chP=>[][/In_dom/dom_cond]. Qed. -Lemma ch_fun m g x y z : - ch m g x y -> - ch m g z y -> - x = z. +Lemma ch_fun m g a b x : + ch m g a x -> + ch m g b x -> + a = b. Proof. by case/chP=>[][H <-] /chP [][/(In_fun H) {}H <-]. Qed. Lemma ch_path m g s x : - path (ch m g) null s -> x \in s -> + path (ch m g) null s -> exists y, y \in belast null s /\ ch m g y x. -Proof. -elim/last_ind: s=>//= s z IH. -rewrite rcons_path mem_rcons belast_rcons !inE. -case/andP=>H1 H2 /orP [/eqP E|]. -- by exists (last null s); rewrite mem_last E. -by case/(IH H1)=>y [H3 H4]; exists y; rewrite mem_belast. -Qed. +Proof. exact: path_prev. Qed. Lemma ch_path_uniq m g s : path (ch m g) null s -> uniq (null :: s). -Proof. -elim/last_ind: s=>[|s x IH] //=. -rewrite rcons_path mem_rcons=>/andP [Px Cx]. -move: {IH}(IH Px) (IH Px); rewrite {1}lastI /=. -rewrite !rcons_uniq inE negb_or=>/andP [N _]/andP [->->]. -rewrite eq_sym (chN0 Cx) !andbT /=. -apply/negP=>/(ch_path Px) [/= y][/[swap]/(ch_fun Cx) <-]. -by rewrite (negbTE N). -Qed. +Proof. by apply: path_uniq; [apply: chN0|apply: ch_fun]. Qed. Definition upd_edge (m : nmap mark) g x y : seq node := if find x m is Some L then [:: y; gr g x] else [:: gl g x; y]. From 6ef5717f4c16f2e02567f24b1687d91afc642da3 Mon Sep 17 00:00:00 2001 From: Aleks Nanevski Date: Fri, 9 Aug 2024 19:36:06 +0200 Subject: [PATCH 32/93] blah --- examples/schorr.v | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/examples/schorr.v b/examples/schorr.v index 9ee4468..a5b49b0 100644 --- a/examples/schorr.v +++ b/examples/schorr.v @@ -49,7 +49,7 @@ Qed. (* Schorr-Waite *) (****************) -(* short notation for left/right child *) +(* short notation for left/right child of x *) Notation gl g x := (get_nth g x 0). Notation gr g x := (get_nth g x 1). From f030b03b564add0a3d8a952a1850453a51f95d37 Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Mon, 12 Aug 2024 12:04:14 +0200 Subject: [PATCH 33/93] blah --- examples/array.v | 41 ++++++++++--------- examples/bintree.v | 52 ++++++++++++------------ examples/cyclic.v | 99 ++++++++++++++++++++++++---------------------- examples/dlist.v | 44 ++++++++++----------- examples/gcd.v | 4 +- examples/graph.v | 8 ++-- examples/llist.v | 63 +++++++++++++---------------- examples/queue.v | 26 +++++------- examples/schorr.v | 27 ++++++++----- examples/stack.v | 21 +++++----- 10 files changed, 189 insertions(+), 196 deletions(-) diff --git a/examples/array.v b/examples/array.v index 61c7de4..96aaad7 100644 --- a/examples/array.v +++ b/examples/array.v @@ -19,11 +19,12 @@ From htt Require Import options model heapauto. Definition indx {I : finType} (x : I) := index x (enum I). -(***********************************) -(* Arrays indexed by a finite type *) -(***********************************) +(*********************************) +(* Arrays indexed by finite type *) +(*********************************) -(* an array is a pointer to a contiguous memory region holding its values *) +(* array is (pointer to) a contiguous memory region *) +(* holding the array values *) Record array (I : finType) (T : Type) : Type := Array {orig :> ptr}. Arguments Array {I T}. @@ -38,16 +39,14 @@ Section Array. Variable (I : finType) (T : Type). Notation array := {array I -> T}. -(* an array is specified by a finite function *) - +(* array is specified by finite function *) Definition shape (a : array) (f : {ffun I -> T}) := [Pred h | h = updi a (fgraph f) /\ valid h]. (* main methods *) -(* a new empty array preallocates all cells for all possible index values *) +(* new empty array preallocates all cells for all possible index values *) (* initializing all of them to `x` *) - Program Definition new (x : T) : STsep (emp, [vfun a => shape a [ffun => x]]) := Do (x <-- allocb x #|I|; @@ -62,16 +61,17 @@ Qed. Opaque new. -(* a new array corresponding to a domain of a finite function f *) +(* new array corresponding to a domain of a finite function f *) -(* the loop invariant: *) -(* a partially filled array corresponds to some finite function g acting as a prefix of f *) +(* loop invariant: *) +(* partially filled array corresponds to finite function g *) +(* acting as prefix of f *) Definition fill_loop a (f : {ffun I -> T}) : Type := forall s : seq I, STsep (fun i => exists g s', [/\ i \In shape a g, s' ++ s = enum I & forall x, x \in s' -> g x = f x], - [vfun a => shape a f]). + [vfun a => shape a f]). Program Definition newf (f : {ffun I -> T}) : STsep (emp, [vfun a => shape a f]) := @@ -82,10 +82,10 @@ Program Definition newf (f : {ffun I -> T}) : x <-- new (f v); (* fill it with values of f on I *) let fill := ffix (fun (loop : fill_loop x f) s => - Do (if s is k::t return _ then - x .+ (indx k) ::= f k;; - loop t - else ret (Array x))) + Do (if s is k::t return _ then + x .+ (indx k) ::= f k;; + loop t + else ret (Array x))) in fill (enum I) else ret (Array null)). (* first the loop *) @@ -130,11 +130,11 @@ Definition free_loop (a : array) : Type := [/\ i = updi (a .+ k) xs, valid i & size xs + k = #|I|], - [vfun _ : unit => emp]). + [vfun _ : unit => emp]). Program Definition free (a : array) : STsep (fun i => exists f, i \In shape a f, - [vfun _ : unit => emp]) := + [vfun _ : unit => emp]) := Do (let go := ffix (fun (loop : free_loop a) k => Do (if k == #|I| then ret tt else dealloc a.+k;; @@ -160,10 +160,9 @@ by rewrite ptr0 V {3}codomE size_map -cardE addn0. Qed. (* reading from an array, doesn't modify the heap *) - Program Definition read (a : array) (k : I) : STsep {f h} (fun i => i = h /\ i \In shape a f, - [vfun (y : T) m => m = h /\ y = f k]) := + [vfun (y : T) m => m = h /\ y = f k]) := Do (!a .+ (indx k)). Next Obligation. (* pull out ghost vars *) @@ -176,7 +175,7 @@ Qed. Program Definition write (a : array) (k : I) (x : T) : STsep {f} (shape a f, - [vfun _ : unit => shape a (finfun [eta f with k |-> x])]) := + [vfun _ : unit => shape a (finfun [eta f with k |-> x])]) := Do (a .+ (indx k) ::= x). Next Obligation. (* pull out ghost vars, split the heap *) diff --git a/examples/bintree.v b/examples/bintree.v index 0cbe427..b05e332 100644 --- a/examples/bintree.v +++ b/examples/bintree.v @@ -35,14 +35,13 @@ Fixpoint inorder {A} (t : tree A) : seq A := else [::]. (* Add an element to the rightmost branch *) - Fixpoint addr {A} (y : A) (t : tree A) : tree A := if t is Node l x r then Node l x (addr y r) else Node leaf y leaf. -(* Tree heap predicate: [left branch pointer, value, right branch pointer] *) - +(* Tree heap predicate: *) +(* [left branch pointer, value, right branch pointer] *) Fixpoint shape {A} (p : ptr) (t : tree A) := if t is Node l a r then [Pred h | exists l' r' h', @@ -51,7 +50,6 @@ Fixpoint shape {A} (p : ptr) (t : tree A) := else [Pred h | p = null /\ h = Unit]. (* Null pointer represents a leaf *) - Lemma shape_null {A} (t : tree A) h : valid h -> h \In shape null t -> @@ -62,7 +60,6 @@ by case=>?[?][?][E]; rewrite E validPtUn in V. Qed. (* Non-null pointer represents a node *) - Lemma shape_cont {A} (t : tree A) p h : p != null -> h \In shape p t -> @@ -73,15 +70,14 @@ Lemma shape_cont {A} (t : tree A) p h : Proof. move=>E; case: t=>/= [|l a r]. - by case=>E0; rewrite E0 in E. -case=>l'[r'][h'][E1 E2]. -by exists l,a,r,l',r',h'. +case=>l' [r'][h'][E1 E2]. +by exists l, a, r, l', r', h'. Qed. Section Tree. Variable A : Type. (* Node creation *) - Program Definition mknode (x : A) : STsep (emp, [vfun p => shape p (Node leaf x leaf)]) := @@ -99,7 +95,7 @@ Qed. (* Recursive tree disposal *) -(* We start from a well-formed tree and arrive at an empty heap *) +(* Start from a well-formed tree, and arrive at empty heap *) Definition disposetreeT : Type := forall p, STsep {t : tree A} (shape p t, [vfun _ : unit => emp]). @@ -129,17 +125,17 @@ apply: [stepX r]@hr=>//= _ _ ->; rewrite unitR. by do 3!step; rewrite !unitL. Qed. -(* Calculation of tree size *) +(* Calculate tree size *) (* loop invariant: *) (* the subtree size is added to the accumulator *) Definition treesizeT : Type := forall (ps : ptr * nat), STsep {t : tree A} (shape ps.1 t, - [vfun s h => s == ps.2 + size_tree t /\ shape ps.1 t h]). + [vfun s h => s == ps.2 + size_tree t /\ shape ps.1 t h]). Program Definition treesize p : STsep {t : tree A} (shape p t, - [vfun s h => s == size_tree t /\ shape p t h]) := + [vfun s h => s == size_tree t /\ shape p t h]) := Do (let len := ffix (fun (go : treesizeT) '(p, s) => Do (if p == null then ret s else l <-- !p; @@ -150,10 +146,10 @@ Program Definition treesize p : Next Obligation. (* pull out ghost var + precondition, branch on null check *) move=>_ go _ p s /= [t][] i /= H; case: eqP H=>[{p}->|/eqP Ep] H. -- (* empty tree has size 0 *) - by step=>V; case: (shape_null V H)=>->->/=; rewrite addn0. +(* empty tree has size 0 *) +- by step=>V; case: (shape_null V H)=>->->/=; rewrite addn0. (* non-null pointer is a node, deconstruct it, read branch pointers *) - case: (shape_cont Ep H)=>l[a][r][l'][r'][_][{t H}-> {i}-> [hl][hr][-> Hl Hr]] /=. +case: (shape_cont Ep H)=>l[a][r][l'][r'][_][{t H}-> {i}-> [hl][hr][-> Hl Hr]]. do 2!step. (* calculate left branch size, update the accumulator *) apply: [stepX l]@hl=>//= _ hl' [/eqP -> Hl']. @@ -163,7 +159,7 @@ apply: [gX r]@hr=>//= _ hr' [/eqP -> Hr'] _. by split; [rewrite addnS addSn addnA | vauto]. Qed. Next Obligation. -(* pull out ghost var + precondition, start the loop with empty accumulator *) +(* pull out ghost var + precondition, start loop with empty accumulator *) by move=>/= p [t][] i /= H; apply: [gE t]. Qed. @@ -176,11 +172,11 @@ Opaque insert new. Definition inordertravT : Type := forall (ps : ptr * ptr), STsep {(t : tree A) (l : seq A)} (shape ps.1 t # lseq ps.2 l, - [vfun s h => h \In shape ps.1 t # lseq s (inorder t ++ l)]). + [vfun s h => h \In shape ps.1 t # lseq s (inorder t ++ l)]). Program Definition inordertrav p : STsep {t : tree A} (shape p t, - [vfun s h => h \In shape p t # lseq s (inorder t)]) := + [vfun s h => h \In shape p t # lseq s (inorder t)]) := Do (let loop := ffix (fun (go : inordertravT) '(p, s) => Do (if p == null then ret s else l <-- !p; @@ -192,17 +188,19 @@ Program Definition inordertrav p : in n <-- new A; loop (p, n)). Next Obligation. -(* pull out ghosts + precondition, deconstruct the heap, branch on null check *) +(* pull out ghosts + precondition, destruct heap, branch on null check *) move=>_ go _ p s /= [t][xs][] _ /= [h1][h2][-> H1 H2]. case: eqP H1=>[{p}->|/eqP Ep] H1. -- (* return the accumulated list - empty tree has no values *) - by step=>V; case: (shape_null (validL V) H1)=>->->/=; vauto. -(* non-empty tree is a node, deconstruct the node, read the pointers and the value *) -case: (shape_cont Ep H1)=>l[a][r][l'][r'][_][{t H1}-> {h1}-> [hl][hr][-> Hl Hr]] /=. -do 3!step. -(* run traversal on the right branch first (it's cheaper to grow a linked list to the left) *) +(* return the accumulated list - empty tree has no values *) +- by step=>V; case: (shape_null (validL V) H1)=>->->/=; vauto. +(* non-empty tree is a node *) +(* deconstruct the node, read the pointers and the value *) +case: (shape_cont Ep H1)=>l[a][r][l'][r'][_][{t H1}-> {h1}-> +[hl][hr][-> Hl Hr]] /=; do 3!step. +(* run traversal on the right branch first *) +(* (it's cheaper to grow a linked list to the left) *) apply: [stepX r, xs]@(hr \+ h2)=>//=; first by vauto. -(* we have subheaps corresponding to the left+right branches and the updated list *) +(* subheaps exist corresponding to left/right branches and updated list *) move=>s1 _ [hr'][hs][-> Hr' Hs]. (* prepend the node value to the list *) apply: [stepX (inorder r ++ xs)]@hs=>//= pa _ [s2][h'][-> H']. @@ -236,7 +234,7 @@ Opaque mknode. (* loop invariant: the value is added to the rightmost branch *) Definition expandrightT x : Type := forall (p : ptr), STsep {t : tree A} (shape p t, - [vfun p' => shape p' (addr x t)]). + [vfun p' => shape p' (addr x t)]). Program Definition expandright x : expandrightT x := ffix (fun (go : expandrightT x) p => diff --git a/examples/cyclic.v b/examples/cyclic.v index 76c22bf..968ba47 100644 --- a/examples/cyclic.v +++ b/examples/cyclic.v @@ -18,9 +18,11 @@ From pcm Require Import pcm unionmap heap auto automap autopcm. From htt Require Import options model heapauto. From htt Require Import llist. -(* a queue variant that has a fixed capacity and can overwrite data in a circular way *) +(* queue variant with fixed capacity *) +(* that overwrites data in a circular way *) -(* the structure is a pair of pointers, a mutable length and immutable capacity *) +(* the structure is a pair of pointers: *) +(* a mutable length and immutable capacity *) Record buffer (T : Type) : Type := Buf {active: ptr; inactive: ptr; len: ptr; capacity: nat}. @@ -42,15 +44,14 @@ Definition is_buffer (a i : ptr) (m n : nat) (xs : seq T) := n = size xs + size ys, m = size xs, ha \In lseg a i xs & - hi \In lseg i a ys] - ]. + hi \In lseg i a ys]]. (* the structure itself requires three extra memory cells *) Definition shape (b : buffer) (xs : seq T) := [Pred h | exists a i m h', - [/\ valid (active b :-> a \+ (inactive b :-> i \+ (len b :-> m \+ h'))), - h = active b :-> a \+ (inactive b :-> i \+ (len b :-> m \+ h')) & - h' \In is_buffer a i m (capacity b) xs]]. + [/\ valid (active b :-> a \+ (inactive b :-> i \+ (len b :-> m \+ h'))), + h = active b :-> a \+ (inactive b :-> i \+ (len b :-> m \+ h')) & + h' \In is_buffer a i m (capacity b) xs]]. (* main methods *) @@ -170,12 +171,13 @@ exists a, i, (size xs), h; split=>//. by exists [::], ha, hi; rewrite Eys in Ec Hi. Qed. -(* a version that overwrites data in a cyclic fashion *) -(* we make checking for capacity != 0 the client's problem so it can be dealt with globally *) +(* version that overwrites data in a cyclic fashion *) +(* checking that capacity != 0 is the client's problem *) +(* so it can be dealt with globally *) Program Definition overwrite (x : T) (b : buffer) : STsep {xs} (fun h => 0 < capacity b /\ h \In shape b xs, - [vfun _ => shape b (drop ((size xs).+1 - capacity b) - (rcons xs x))]) := + [vfun _ => shape b (drop ((size xs).+1 - capacity b) + (rcons xs x))]) := Do (i <-- !inactive b; i ::= x;; r <-- !i.+1; @@ -191,26 +193,29 @@ move=>x b [xs []] _ /= [Hc][a][i][_][_][_ -> [ys][ha][hi][-> Ec -> Ha Hi]]. (* read the inactive pointer, case split on inactive part *) (* we do this early on because we need to unroll the heap to proceed *) rewrite Ec in Hc *; step; case: ys Ec Hi Hc=>/=. -- (* inactive part is empty, so the buffer is full & the active part is the whole cycle *) - rewrite addn0=>Ec [Ei {hi}->] Hxs; rewrite unitR; rewrite {i}Ei in Ha *. - (* unroll the (active) heap *) +(* inactive part is empty, so the buffer is full *) +(* and the active part is the whole cycle *) +- rewrite addn0=>Ec [Ei {hi}->] Hxs; rewrite unitR; rewrite {i}Ei in Ha *. +(* unroll the (active) heap *) case/(lseg_lt0n Hxs): Ha=>z[r][h'][Exs {ha}-> H']. - (* run the rest of the program, the postcondition simplifies to beheading the extended xs *) +(* run the rest of the program *) +(* the postcondition simplifies to beheading the extended xs *) do 4!step; rewrite ltnn subSnn drop1; step=>V. - (* the new structure is well-formed if the buffer is *) +(* the new structure is well-formed if the buffer is *) exists r, r, (size xs), (a :-> x \+ (a.+1 :-> r \+ h')); split=>//. - (* most of the conditions hold trivially or by simple arithmetics *) +(* most of the conditions hold trivially or by simple arithmetics *) exists [::], (a :-> x \+ (a.+1 :-> r \+ h')), Unit; split=>//=. - by rewrite unitR. - by rewrite addn0 size_behead size_rcons. - by rewrite size_behead size_rcons. - (* xs is non-empty so behead commutes with rcons *) +(* xs is non-empty so behead commutes with rcons *) rewrite behead_rcons //; apply/lseg_rcons. - (* the segment is well-formed *) +(* the segment is well-formed *) by exists a, h'; split=>//; rewrite [in RHS]joinC joinA. (* inactive part is non-empty, unroll the inactive heap *) move=>y ys Ec [r][h'][{hi}-> H'] _. -(* run the rest of the program, the postcondition simplifies to just the extended xs *) +(* run the rest of the program *) +(* the postcondition simplifies to just the extended xs *) do 4![step]=>{y}; rewrite -{1}(addn0 (size xs)) ltn_add2l /=; step=>V. rewrite addnS subSS subnDA subnn sub0n drop0. (* the new structure is well-formed if the buffer is *) @@ -223,13 +228,13 @@ exists ys, (ha \+ (i :-> x \+ i.+1 :-> r)), h'; split=>//. by apply/lseg_rcons; exists i, ha. Qed. -(* reading (or rather popping) a value from the buffer *) +(* reading (popping) a value from the buffer *) Program Definition read (b : buffer) : STsep {xs} (shape b xs, - fun y h => match y with - | Val x => h \In shape b (behead xs) /\ ohead xs = Some x - | Exn e => e = BufferEmpty /\ xs = [::] - end) := + fun y h => match y with + | Val x => h \In shape b (behead xs) /\ ohead xs = Some x + | Exn e => e = BufferEmpty /\ xs = [::] + end) := Do (m <-- !len b; if 0 < m then a <-- !active b; @@ -240,19 +245,19 @@ Program Definition read (b : buffer) : ret x else throw BufferEmpty). Next Obligation. -(* pull out ghost, destructure the precondition *) +(* pull out ghost, destructure precondition *) move=>b [xs []] _ /= [a][i][_][_][_ -> [ys][ha][hi][-> Hc -> Ha Hi]]. (* read the length, match on it *) step; case: ltnP. -- (* buffer is non-empty, read the active pointer *) - move=>Hxs; step. - (* unroll the active heap to proceed *) +(* buffer is non-empty, read the active pointer *) +- move=>Hxs; step. +(* unroll the active heap to proceed *) case/(lseg_lt0n Hxs): Ha=>x[r][h'][Exs {ha}-> H']. - (* run the rest of the program *) +(* run the rest of the program *) do 5![step]=>V; split; last by rewrite Exs. - (* the new structure is well-formed if the buffer is *) +(* the new structure is well-formed if the buffer is *) exists r, i, (size xs).-1, (a :-> x \+ (a.+1 :-> r \+ h') \+ hi); split=>//. - (* the buffer is well-formed by simple arithmetics *) +(* the buffer is well-formed by simple arithmetics *) exists (rcons ys x), h', (hi \+ (a :-> x \+ a.+1 :-> r)); split=>//. - by rewrite [LHS](pullX (h' \+ hi)) !joinA. - by rewrite size_rcons Hc {1}Exs /= addnS addSn. @@ -263,18 +268,18 @@ by rewrite leqn0=>/nilP->; step. Qed. (* deallocating the buffer *) -(* we do check for capacity = 0 here since this should be a rare operation *) +(* check that capacity = 0 here, since this should be a rare operation *) (* of course, it could also be moved into the precondition as for overwrite *) (* loop invariant for deallocating the inner structure *) Definition free_loopT (n : nat) : Type := forall (pk : ptr * nat), STsep {q (xs : seq T)} (fun h => h \In lseg pk.1 q xs /\ size xs = n - pk.2, - [vfun _ : unit => emp]). + [vfun _ : unit => emp]). Program Definition free (b : buffer) : STsep {xs} (fun h => h \In shape b xs, - [vfun _ => emp]) := + [vfun _ => emp]) := Do (let run := ffix (fun (go : free_loopT (capacity b)) '(r,k) => Do (if k < capacity b then p' <-- !r.+1; @@ -294,13 +299,13 @@ Program Definition free (b : buffer) : Next Obligation. (* pull out ghosts, destructure the preconditions, match on k *) move=>b go _ r k [q][xs][] h /= [H Hs]; case: ltnP. -- (* k < capacity, so the spec is still non-empty *) - move=>Hk; have {Hk}Hxs: 0 < size xs by rewrite Hs subn_gt0. - (* unroll the heap *) +(* k < capacity, so the spec is still non-empty *) +- move=>Hk; have {Hk}Hxs: 0 < size xs by rewrite Hs subn_gt0. +(* unroll the heap *) case/(lseg_lt0n Hxs): H=>x[p][h'][E {h}-> H']. - (* save the next node pointer and deallocate the node *) +(* save the next node pointer and deallocate the node *) do 3!step; rewrite !unitL. - (* run the recursive call, simplify *) +(* run the recursive call, simplify *) apply: [gE q, behead xs]=>//=; split=>//. by rewrite subnS; move: Hs; rewrite {1}E /= =><-. (* k = capacity, so the spec is empty *) @@ -312,18 +317,18 @@ Qed. Next Obligation. (* pull out ghost, destructure the precondition, match on capacity *) move=>b [xs][] _ /= [a][i][m][_][_ -> [ys][ha][hi][-> Hc -> Ha Hi]]; case: ltnP. -- (* 0 < capacity, first apply the vrf_bind primitive to "step into" the `if` block *) - (* run the loop starting from the active part, frame on the corresponding heaps *) - move=>_; apply/vrf_bnd; step; apply: [stepX a, xs++ys]@(ha \+ hi)=>//=. - - (* concatenate the active and inactive parts to satisfy the loop precondition *) - move=>_; split; last by rewrite size_cat subn0. +(* 0 < capacity, first apply vrf_bind to "step into" the `if` block *) +(* run the loop starting from the active part, frame the corresponding heaps *) +- move=>_; apply/vrf_bnd; step; apply: [stepX a, xs++ys]@(ha \+ hi)=>//=. +(* concatenate the active and inactive parts to satisfy the loop precondition *) + - move=>_; split; last by rewrite size_cat subn0. by apply/lseg_cat; exists i, ha, hi. - (* by deallocating the structure we empty the heap fully *) +(* by deallocating the structure, the heap is emptied *) by move=>_ _ ->; rewrite unitR; step=>V; do 3![step]=>_; rewrite !unitR. (* capacity = 0, so just deallocate the structure *) rewrite leqn0=>/eqP E; do 4!step; rewrite !unitL=>_. (* both parts of the spec are empty and so are the corresponding heaps *) -move: Hc; rewrite E =>/eqP; rewrite eq_sym addn_eq0; case/andP=>/nilP Ex /nilP Ey. +move: Hc; rewrite E =>/eqP; rewrite eq_sym addn_eq0=>/andP [/nilP Ex /nilP Ey]. move: Ha; rewrite Ex /=; case=>_->; move: Hi; rewrite Ey /=; case=>_->. by rewrite unitR. Qed. diff --git a/examples/dlist.v b/examples/dlist.v index 4657590..72e551d 100644 --- a/examples/dlist.v +++ b/examples/dlist.v @@ -17,8 +17,8 @@ From pcm Require Import options axioms pred. From pcm Require Import pcm unionmap heap autopcm. From htt Require Import options model heapauto. -(* Doubly linked lists, follows the same structure as singly-linked *) -(* ones, adding a second pointer pointing backwards. *) +(* Doubly-linked lists, follows the same structure as singly-linked *) +(* lists, adding a second pointer pointing backwards. *) (* The arguments are: preceding node, first node, last node, succeeding node *) (* Internally we use three cells for a node: value, next node, last node *) @@ -35,7 +35,7 @@ Variable A : Type. (* structure of a list segment with appended node *) (* adding the predecessor pointer breaks the symmetry with inductive lists *) -(* so we rely on this lemma a lot to manually restructure the spec *) +(* so we rely on this lemma to manually restructure the spec *) Lemma dseg_rcons (xs : seq A) x pp p q qn h : h \In dseg pp p q qn (rcons xs x) <-> exists r h', @@ -57,7 +57,7 @@ rewrite !joinA; split=>//; apply/IH. by exists r, h'; rewrite !joinA. Qed. -(* first node being null means the list is empty *) +(* if first node is null, then list is empty *) Lemma dseg_nullL (xs : seq A) pp q qn h : valid h -> h \In dseg pp null q qn xs -> @@ -67,7 +67,7 @@ case: xs=>[|x xs] /= D H; first by case: H. by case: H D=>r[h'][-> _]; rewrite validPtUn eq_refl. Qed. -(* last node being null also means the list is empty *) +(* if last node is null, then list is empty *) Lemma dseg_nullR (xs : seq A) pp p qn h : valid h -> h \In dseg pp p null qn xs -> @@ -78,7 +78,7 @@ case/dseg_rcons=>r[h'][]; move: D=>/[swap]->/validR. by rewrite validPtUn. Qed. -(* deconstruct a non-trivial segment from the left *) +(* deconstruct non-trivial segment from the left *) Lemma dseg_neqL (xs : seq A) (pp p q qn : ptr) h : p != qn -> h \In dseg pp p q qn xs -> @@ -92,7 +92,7 @@ case: xs=>[|x xs] /= H; last first. by case=>E; rewrite E eq_refl in H. Qed. -(* deconstruct a non-trivial segment from the right *) +(* deconstruct non-trivial segment from the right *) Lemma dseg_neqR (xs : seq A) (pp p q qn : ptr) h : pp != q -> h \In dseg pp p q qn xs -> @@ -107,9 +107,9 @@ case/dseg_rcons=>r[h'][{h}-> H']. by exists xs, x, r, h'. Qed. -(* splitting the list corresponds to splitting the heap *) +(* concatenating/splitting lists = concatenating/splitting heaps *) Lemma dseg_cat (xs ys : seq A) pp p q qn h : - h \In dseg pp p q qn (xs++ys) <-> + h \In dseg pp p q qn (xs ++ ys) <-> exists jn j, h \In dseg pp p j jn xs # dseg j jn q qn ys. Proof. elim: xs pp p h=>/=. @@ -167,8 +167,7 @@ Proof. by rewrite eq_sym=>Hq /(dseg_neqR Hq). Qed. (* main methods *) -(* prepend a value, return pointers to new first and last nodes *) - +(* prepend value x, return pointers to new first and last nodes *) Program Definition insertL p q (x : A) : STsep {l} (dseq p q l, [vfun pq => dseq pq.1 pq.2 (x :: l)]) := Do (r <-- allocb x 3; @@ -192,16 +191,15 @@ exists p, (p :-> y \+ (p.+1 :-> z \+ (p.+2 :-> r \+ h'))). by rewrite !joinA; split=>//; rewrite E /=; exists z, h'; rewrite !joinA. Qed. -(* append a value (in constant), return pointers to new first and last nodes *) - +(* append value x, return pointers to new first and last nodes *) Program Definition insertR p q (x : A) : STsep {l} (dseq p q l, [vfun pq => dseq pq.1 pq.2 (rcons l x)]) := Do (r <-- allocb x 3; r.+1 ::= null;; r.+2 ::= q;; if q == null then ret (r,r) - else q.+1 ::= r;; - ret (p,r)). + else q.+1 ::= r;; + ret (p,r)). Next Obligation. (* pull out ghost + precondition *) move=>p q x [l []] i /= H. @@ -219,22 +217,22 @@ split; first by rewrite joinC. by apply/dseg_rcons; vauto. Qed. -(* traversing the dlist backwards and consing all elements *) +(* travers the dlist backwards and cons all elements *) (* reifies the specification *) -(* the loop invariant: *) -(* 1. the heap is unchanged *) -(* 2. the result is remainder + accumulator *) -(* we carry the pointer to the accumulator as a ghost var *) +(* loop invariant: *) +(* 1. heap is unchanged *) +(* 2. result = remainder + accumulator *) +(* carry the pointer to the accumulator as a logical var *) Definition traverse_backT (p : ptr) : Type := forall (qs : ptr * seq A), STsep {l nx} (dseg null p qs.1 nx l, - [vfun r h => h \In dseg null p qs.1 nx l - /\ r = l ++ qs.2]). + [vfun r h => h \In dseg null p qs.1 nx l /\ + r = l ++ qs.2]). Program Definition traverse_back p q : STsep {l} (dseq p q l, - [vfun r h => h \In dseq p q l /\ r = l]) := + [vfun r h => h \In dseq p q l /\ r = l]) := Do (let tb := ffix (fun (go : traverse_backT p) '(r, acc) => Do (if r == null then ret acc diff --git a/examples/gcd.v b/examples/gcd.v index 9d718c9..40c39b0 100644 --- a/examples/gcd.v +++ b/examples/gcd.v @@ -27,8 +27,8 @@ Definition shape (p q : ptr) (x y : nat) := Definition gcd_loopT (p q : ptr) : Type := unit -> STsep {x y : nat} (shape p q x y, - [vfun (_ : unit) h => - h \In shape p q (gcdn x y) (gcdn x y)]). + [vfun (_ : unit) h => + h \In shape p q (gcdn x y) (gcdn x y)]). Program Definition gcd_loop (p q : ptr) := ffix (fun (go : gcd_loopT p q) _ => diff --git a/examples/graph.v b/examples/graph.v index aa7965b..9a6f2a8 100644 --- a/examples/graph.v +++ b/examples/graph.v @@ -33,8 +33,6 @@ From pcm Require Import pcm unionmap natmap autopcm automap. (* that encodes that the second child of x doesn't exist. *) (* Non-null dangling links are technically possible, *) (* but are treated same as null. *) -(* If it's desired to treat a non-null dangling link differently *) -(* from null, add that link to the graph to make it non-dangling. *) (* Pregraph differs from fingraph (of mathcomp) *) (* in that the set of nodes is drawn from an infinite set *) @@ -96,8 +94,6 @@ Definition children (g : pregraph) x : seq node := (* edge is applicative variant of children *) (* thus, dangling edges (null or non-null) are *not* edges. *) -(* If it's desired to treat a non-null dangling node differently *) -(* from null, add that node to the graph to make it non-dangling. *) Definition edge g : rel node := mem \o children g. Arguments edge g x y : simpl never. @@ -277,6 +273,10 @@ Proof. by move=>x y; rewrite !inE => X Y; rewrite edge_umfiltk /= X Y. Qed. (* at depth n, starting from x, and avoiding v. *) (* Definition uses children, not links; *) (* thus, it doesn't follow dangling edges *) +(* and dfs can't express reachability to an outside node. *) +(* If the latter is desired, it can be separately defined *) +(* as a conjunct of dfs and links properties. *) + Fixpoint dfs (g : pregraph) (n : nat) (v : seq node) x := if (x \notin dom g) || (x \in v) then v else if n is n'.+1 then foldl (dfs g n') (x :: v) (children g x) else v. diff --git a/examples/llist.v b/examples/llist.v index 6c6f0ab..c4bf28d 100644 --- a/examples/llist.v +++ b/examples/llist.v @@ -31,8 +31,7 @@ Definition EmptyList : exn := exn_from_nat 15. Section LSeg. Variable A : Type. -(* appending a value to the list segment *) - +(* appending value to list segment *) Lemma lseg_rcons (xs : seq A) x p r h : h \In lseg p r (rcons xs x) <-> exists q h', h = h' \+ (q :-> x \+ q.+1 :-> r) /\ @@ -48,8 +47,7 @@ exists z, (h2 \+ q :-> y \+ q.+1 :-> r). by rewrite -!joinA; split=>//; apply/IH; eauto. Qed. -(* null pointer represents an empty segment *) - +(* null pointer represents empty segment *) Lemma lseg_null (xs : seq A) q h : valid h -> h \In lseg null q xs -> @@ -59,8 +57,7 @@ case: xs=>[|x xs] D /= H; first by case: H=><- ->. case: H D=>r [h'][->] _; rewrite validPtUn; hhauto. Qed. -(* empty heap represents an empty segment *) - +(* empty heap represents empty segment *) Lemma lseg_empty (xs : seq A) p q : Unit \In lseg p q xs -> p = q /\ xs = [::]. @@ -69,7 +66,6 @@ by case: xs=>[|x xs][] //= r [h][/esym/umap0E][/unitbP]; rewrite um_unitbU. Qed. (* reformulation of the specification *) - Lemma lseg_case (xs : seq A) p q h : h \In lseg p q xs -> [/\ p = q, xs = [::] & h = Unit] \/ @@ -83,8 +79,7 @@ by case=>r [h'][->] H; right; hhauto. Qed. (* non-trivial segment represents a non-empty list *) - -Corollary lseg_neq (xs : seq A) p q h : +Lemma lseg_neq (xs : seq A) p q h : p != q -> h \In lseg p q xs -> exists x r h', [/\ xs = x :: behead xs, @@ -96,7 +91,6 @@ by rewrite E eq_refl in H. Qed. (* non-empty list is represented by a non-trivial segment *) - Lemma lseg_lt0n (xs : seq A) p q h : 0 < size xs -> h \In lseg p q xs -> exists x r h', @@ -109,7 +103,6 @@ by rewrite E in H. Qed. (* splitting the list corresponds to splitting the heap *) - Lemma lseg_cat (xs ys : seq A) p q h : h \In lseg p q (xs++ys) <-> exists j, h \In lseg p j xs # lseg j q ys. @@ -151,12 +144,13 @@ Lemma lseq_pos (xs : seq A) p h : h' \In lseq r (behead xs)]. Proof. by apply: lseg_neq. Qed. -(* a valid heap cannot match two different specs *) +(* valid heap cannot match two different specs *) Lemma lseq_func (l1 l2 : seq A) p h : valid h -> h \In lseq p l1 -> - h \In lseq p l2 -> l1 = l2. + h \In lseq p l2 -> + l1 = l2. Proof. elim: l1 l2 p h => [|x1 xt IH] /= l2 p h V. - by case=>->->; case/lseq_null. @@ -175,7 +169,6 @@ Program Definition new : STsep (emp, [vfun x => lseq x (@nil A)]) := Next Obligation. by move=>[] /= _ ->; step. Qed. (* prepending a value *) - Program Definition insert p (x : A) : STsep {l} (lseq p l, [vfun p' => lseq p' (x::l)]) := Do (q <-- allocb p 2; @@ -190,7 +183,6 @@ Qed. (* getting the head element *) (* an example of a partial program, doesn't modify the heap *) - Program Definition head p : STsep {l} (lseq p l, fun (y : ans A) h => h \In lseq p l /\ @@ -212,7 +204,6 @@ by do 2![step]=>_; split=>//; rewrite E; hhauto. Qed. (* removing the head element, no-op for an empty list *) - Program Definition remove p : STsep {xs : seq A} (lseq p xs, [vfun p' => lseq p' (behead xs)]) := Do (if p == null then ret p @@ -232,16 +223,16 @@ Qed. (* calculating the list length *) -(* the loop invariant: *) +(* loop invariant: *) (* 1. heap is unchanged *) (* 2. total length is accumulator + the length of unprocessed list *) Definition lenT : Type := forall (pl : ptr * nat), STsep {xs : seq A} (lseq pl.1 xs, - [vfun l h => l == pl.2 + length xs /\ lseq pl.1 xs h]). + [vfun l h => l == pl.2 + length xs /\ lseq pl.1 xs h]). Program Definition len p : STsep {xs : seq A} (lseq p xs, - [vfun l h => l == length xs /\ lseq p xs h]) := + [vfun l h => l == length xs /\ lseq p xs h]) := Do (let len := ffix (fun (go : lenT) '(p, l) => Do (if p == null then ret l else pnext <-- !p.+1; @@ -265,22 +256,22 @@ Qed. (* concatenation: modifies the first list, returning nothing *) -(* the loop invariant: *) -(* the first list should not be empty and not overlap the second *) +(* loop invariant: *) +(* first list isn't empty and doesn't overlap with the second *) Definition catT (p2 : ptr) : Type := forall (p1 : ptr), STsep {xs1 xs2 : seq A} (fun h => p1 != null /\ (lseq p1 xs1 # lseq p2 xs2) h, - [vfun _ : unit => lseq p1 (xs1 ++ xs2)]). + [vfun _ : unit => lseq p1 (xs1 ++ xs2)]). Program Definition concat p1 p2 : STsep {xs1 xs2 : seq A} (lseq p1 xs1 # lseq p2 xs2, - [vfun a => lseq a (xs1 ++ xs2)]) := + [vfun a => lseq a (xs1 ++ xs2)]) := Do (let cat := ffix (fun (go : catT p2) q => - Do (next <-- !q.+1; - if (next : ptr) == null - then q.+1 ::= p2;; - ret tt - else go next)) + Do (next <-- !q.+1; + if (next : ptr) == null + then q.+1 ::= p2;; + ret tt + else go next)) in if p1 == null then ret p2 else cat p1;; @@ -320,20 +311,20 @@ Qed. (* in-place reversal by pointer swinging *) -(* the loop invariant: *) -(* 1. the processed and remaining parts should not overlap *) -(* 2. the result is processed part + a reversal of remainder *) +(* loop invariant: *) +(* 1. processed and remaining parts don't overlap *) +(* 2. result = processed part + reversal of remainder *) Definition revT : Type := forall (p : ptr * ptr), STsep {i done : seq A} (lseq p.1 i # lseq p.2 done, - [vfun y => lseq y (catrev i done)]). + [vfun y => lseq y (catrev i done)]). Program Definition reverse p : STsep {xs : seq A} (lseq p xs, [vfun p' => lseq p' (rev xs)]) := Do (let reverse := ffix (fun (go : revT) '(i, done) => - Do (if i == null then ret done - else next <-- !i.+1; - i.+1 ::= done;; - go (next, i))) + Do (if i == null then ret done + else next <-- !i.+1; + i.+1 ::= done;; + go (next, i))) in reverse (p, null)). (* first, the loop *) Next Obligation. diff --git a/examples/queue.v b/examples/queue.v index 0e4a15b..2df546c 100644 --- a/examples/queue.v +++ b/examples/queue.v @@ -26,7 +26,8 @@ Section Queue. Variable T : Type. Notation queue := (queue T). -(* a queue is specified as a singly-linked list split into an initial segment and a last node *) +(* queue is singly-linked list split into *) +(* the initial segment and the last node *) Definition is_queue (fr bq : ptr) (xs : seq T) := if fr == null then [Pred h | [/\ bq = null, xs = [::] & h = Unit]] else [Pred h | exists xt x h', @@ -35,8 +36,8 @@ Definition is_queue (fr bq : ptr) (xs : seq T) := h = h' \+ (bq :-> x \+ bq.+1 :-> null) & h' \In lseg fr bq xt]]. -(* the structure itself is a pair of pointers to body + last node *) -(* insertion happens at the last node, and removal at the head *) +(* queue structure is a pair of pointers to body + last node *) +(* insertion happens at the last node, removal at the head *) Definition shape (q : queue) (xs : seq T) := [Pred h | exists fr bq h', [/\ valid (front q :-> fr \+ (back q :-> bq \+ h')), @@ -44,12 +45,10 @@ Definition shape (q : queue) (xs : seq T) := h' \In is_queue fr bq xs]]. (* well-formed queue is a valid heap *) - Lemma shapeD q xs h : h \In shape q xs -> valid h. Proof. by case=>h1[bq][h'] [] D ->. Qed. (* empty queue is a pair of null pointers *) - Lemma is_queue_nil fr bq h : h \In is_queue fr bq [::] -> [/\ fr = null, bq = null & h = Unit]. @@ -58,7 +57,6 @@ by rewrite /is_queue; case: eqP=>[->[-> _ ->] | _ [[|y xt][x][h'][]]]. Qed. (* restructuring the specification for combined list *) - Lemma is_queue_rcons fr bq xt x h : h \In is_queue fr bq (rcons xt x) <-> (exists h', [/\ valid (h' \+ (bq :-> x \+ bq.+1 :-> null)), @@ -73,8 +71,7 @@ move: (D)=>/[swap]; case/(lseg_null (validL D))=>->->->. by rewrite unitL validPtUn. Qed. -(* pointers should agree in a well-formed queue *) - +(* pointers agree in a well-formed queue *) Lemma backfront fr bq xs h : h \In is_queue fr bq xs -> (fr == null) = (bq == null). @@ -86,7 +83,6 @@ Qed. (* main methods *) (* new queue is a pair of pointers to an empty segment *) - Program Definition new : STsep (emp, [vfun v => shape v [::]]) := Do (x <-- alloc null; @@ -100,7 +96,6 @@ by exists null, null, Unit; rewrite !unitR /= in V *; rewrite joinC. Qed. (* freeing a queue, possible only when it's empty *) - Program Definition free (q : queue) : STsep (shape q [::], [vfun _ h => h = Unit]) := Do (dealloc (front q);; @@ -114,10 +109,9 @@ case/is_queue_nil=>->->->; rewrite unitR=>V. by do 2![step]=>_; rewrite unitR. Qed. -(* for enqueue/dequeue we manipulate the underlying segment directly *) - -(* enqueuing is adding a node at the end *) +(* enqueue/dequeue manipulate the underlying segment directly *) +(* enqueuing = adding a node at the end *) Program Definition enq (q : queue) (x : T) : STsep {xs} (shape q xs, [vfun _ => shape q (rcons xs x)]) := @@ -144,7 +138,8 @@ rewrite -(backfront H) unitR; case: ifP H=>Ef; rewrite /is_queue ?Ef. (* the queue wasn't empty, link the new node to the last one *) case=>s2[x2][i2][->] {}D -> H2; step=>V. (* massage the heap and simplify the goal *) -exists fr, next, (i2 \+ bq :-> x2 \+ bq.+1 :-> next \+ next :-> x \+ next.+1 :-> null). +exists fr, next, (i2 \+ bq :-> x2 \+ bq.+1 :-> next \+ + next :-> x \+ next.+1 :-> null). split; first by apply: (validX V). - by rewrite joinC !joinA. (* the new node conforms to the queue spec *) @@ -154,8 +149,7 @@ rewrite joinA; split=>//; first by apply: (validX V). by apply/lseg_rcons; exists bq, i2; rewrite joinA. Qed. -(* dequeuing is removing the head node and adjusting pointers *) - +(* dequeuing = removing the head node and adjusting pointers *) Program Definition deq (q : queue) : STsep {xs} (shape q xs, fun y h => shape q (behead xs) h /\ diff --git a/examples/schorr.v b/examples/schorr.v index c9ba4a9..67476e6 100644 --- a/examples/schorr.v +++ b/examples/schorr.v @@ -128,14 +128,23 @@ Fixpoint rev_graph' m g ps t : pregraph := Definition rev_graph m g s t := rev_graph' m g (rev s) t. -(* turning (pre)graph into a heap *) -Definition hp (g : pregraph) : heap := - \big[join/Unit]_(x <- dom g) (x :-> (gl g x, gr g x)). - -(* Aleks: not sure what this definition intends to say *) -(* but at least it typechecks now *) +(* layout of graph+marking as heap *) +Definition lay (m : nmap mark) (g : pregraph) : heap := + \big[join/Unit]_(x <- dom g) (x :-> (gl g x, gr g x, odflt L (find x m))). + +(* reach m g s t holds iff *) +(* each unmarked node x in g is reachable through unmarked nodes *) +(* - either starting from t, or *) +(* - starting from a right child of some node in s *) Definition reach (m : nmap mark) (g : pregraph) (s : seq node) t := - forall x, x \notin dom m -> - x \in connect (mem (dom m)) g t \/ - exists y, y \in s /\ x \in connect (mem (dom m)) g y. + forall x, + (* x is node in g *) + x \in dom g -> + (* x is unmarked *) + x \notin dom m -> + (* x is reachable from t avoiding marked nodes *) + x \in connect (mem (dom m)) g t \/ + (* or x is reachable from some right child of a node y in s *) + (* avoiding marked nodes *) + exists y, y \in s /\ x \in connect (mem (dom m)) g (gr g y). diff --git a/examples/stack.v b/examples/stack.v index 931e088..f52dae6 100644 --- a/examples/stack.v +++ b/examples/stack.v @@ -28,15 +28,17 @@ Notation stack := (stack T). Opaque insert head remove. -(* a stack is just a pointer to a singly-linked list *) +(* stack is a pointer to a singly-linked list *) Definition shape s (xs : seq T) := [Pred h | exists p h', [ /\ valid (s :-> p \+ h'), h = s :-> p \+ h' & h' \In lseq p xs]]. -(* a heap cannot match two different specs *) +(* heap cannot match two different specs *) Lemma shape_inv : forall s xs1 xs2 h, - h \In shape s xs1 -> h \In shape s xs2 -> xs1 = xs2. + h \In shape s xs1 -> + h \In shape s xs2 -> + xs1 = xs2. Proof. move=>s xs1 xs2 _ [p][h1][D -> S][p2][h2][D2]. case/(cancelO _ D)=><- _; rewrite !unitL=><-. @@ -44,20 +46,19 @@ by apply: lseq_func=>//; move/validR: D. Qed. (* well-formed stack is a valid heap *) - -Lemma shapeD s xs : forall h, h \In shape s xs -> valid h. -Proof. by move=>h [p][h'][D ->]. Qed. +Lemma shapeD s xs h : + h \In shape s xs -> + valid h. +Proof. by case=>p [h'][D ->]. Qed. (* main methods *) (* new stack is a pointer to an empty heap/list *) - Program Definition new : STsep (emp, [vfun y => shape y [::]]) := Do (alloc null). Next Obligation. by move=>/= [] i ?; step=>??; vauto. Qed. (* freeing a stack, possible only when it's empty *) - Program Definition free s : STsep (shape s [::], [vfun _ h => h = Unit]) := Do (dealloc s). @@ -66,9 +67,8 @@ by case=>i [?][?][V ->][_ ->]; step=>_; rewrite unitR. Qed. (* pushing to the stack is inserting into the list and updating the pointer *) - Program Definition push s x : STsep {xs} (shape s xs, - [vfun _ => shape s (x :: xs)]) := + [vfun _ => shape s (x :: xs)]) := Do (l <-- !s; l' <-- insert l x; s ::= l'). @@ -84,7 +84,6 @@ Qed. (* popping from the stack is: *) (* 1. trying to get the head *) (* 2. removing it from the list and updating the pointer on success *) - Program Definition pop s : STsep {xs} (shape s xs, fun y h => shape s (behead xs) h /\ From ab837f6205b99d81e462c66433ddbaf9c94408a1 Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Mon, 12 Aug 2024 12:13:06 +0200 Subject: [PATCH 34/93] blah --- _CoqProject | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) diff --git a/_CoqProject b/_CoqProject index 95176bc..03f6175 100644 --- a/_CoqProject +++ b/_CoqProject @@ -37,18 +37,18 @@ htt/heapauto.v examples/exploit.v examples/gcd.v examples/llist.v -examples/array.v -examples/bubblesort.v -examples/quicksort.v -examples/stack.v examples/dlist.v +examples/array.v examples/queue.v examples/cyclic.v +examples/stack.v examples/bintree.v examples/bst.v examples/kvmaps.v examples/hashtab.v examples/counter.v +examples/bubblesort.v +examples/quicksort.v examples/congmath.v examples/tree.v examples/union_find.v From 4db8d48af576e6db839d19bbb02863e1d3fcd3c9 Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Mon, 12 Aug 2024 12:27:22 +0200 Subject: [PATCH 35/93] blah --- examples/schorr.v | 7 ++++++- 1 file changed, 6 insertions(+), 1 deletion(-) diff --git a/examples/schorr.v b/examples/schorr.v index 8c87948..9275d34 100644 --- a/examples/schorr.v +++ b/examples/schorr.v @@ -54,6 +54,9 @@ Notation gl g x := (get_nth g x 0). Notation gr g x := (get_nth g x 1). (* type of markings *) +(* Aleks NOTE: Should we need U? My impression was that *) +(* being unmarked is represented by not being in *) +(* the appropriate map? *) Inductive mark := U | L | R | LR. (* decidable equality on marks *) Definition eq_mark l1 l2 : bool := @@ -67,7 +70,6 @@ HB.instance Definition _ := hasDecEq.Build mark eq_markP. (* given marking map m, x is m-child of y iff: *) (* - m y = L and x is left child of y *) (* - m y = R and x is right child of y *) - Definition ch (m : nmap mark) (g : pregraph) (x y : nat) := [|| (find y m == Some L) && (gl g y == x) | (find y m == Some R) && (gr g y == x)]. @@ -129,6 +131,9 @@ Definition reach (m : nmap mark) (g : pregraph) (s : seq node) t := (* x is unmarked *) x \notin dom m -> (* x is reachable from t avoiding marked nodes *) + (* Aleks note: if unmarked nodes are represented by a marking U *) + (* then this conjunct will have to change, as it currently assumes *) + (* that unmarked nodes are just those that aren't in m *) x \in connect (mem (dom m)) g t \/ (* or x is reachable from some right child of a node y in s *) (* avoiding marked nodes *) From 3cc4ffa73da629dc12b94160f290cf9a35e738ce Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Mon, 12 Aug 2024 20:08:36 +0200 Subject: [PATCH 36/93] blah --- _CoqProject | 2 +- core/finmap.v | 75 +++--- core/ordtype.v | 15 +- core/seqext.v | 8 - examples/array.v | 5 +- examples/bintree.v | 8 +- examples/bst.v | 122 +++++---- examples/bubblesort.v | 235 +++++++++-------- examples/congmath.v | 571 ++++++++++++++++++++++++++---------------- examples/counter.v | 10 +- examples/hashtab.v | 112 +++++---- examples/kvmaps.v | 235 ++++++++--------- examples/llist.v | 2 + examples/quicksort.v | 172 ++++++------- examples/stack.v | 2 - examples/union_find.v | 7 - htt/model.v | 7 +- 17 files changed, 852 insertions(+), 736 deletions(-) diff --git a/_CoqProject b/_CoqProject index 03f6175..666e241 100644 --- a/_CoqProject +++ b/_CoqProject @@ -36,6 +36,7 @@ htt/model.v htt/heapauto.v examples/exploit.v examples/gcd.v +examples/counter.v examples/llist.v examples/dlist.v examples/array.v @@ -46,7 +47,6 @@ examples/bintree.v examples/bst.v examples/kvmaps.v examples/hashtab.v -examples/counter.v examples/bubblesort.v examples/quicksort.v examples/congmath.v diff --git a/core/finmap.v b/core/finmap.v index 86f371f..0d06816 100644 --- a/core/finmap.v +++ b/core/finmap.v @@ -1,5 +1,5 @@ (* -Copyright 2010 IMDEA Software Institute +Copyright 2009 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 @@ -11,10 +11,10 @@ See the License for the specific language governing permissions and limitations under the License. *) -(******************************************************************************) -(* This file defines finitely supported maps with keys drawn from *) -(* an ordered type and values from an arbitrary type. *) -(******************************************************************************) +(******************************************************************) +(* This file defines finitely supported maps with keys drawn from *) +(* an ordered type and values from an arbitrary type. *) +(******************************************************************) From HB Require Import structures. From Coq Require Import ssreflect ssrbool ssrfun. @@ -53,7 +53,8 @@ Notation value := (@value K V). Notation predk := (@predk K V). Notation predCk := (@predCk K V). -Lemma fmapE (s1 s2 : fmap) : s1 = s2 <-> seq_of s1 = seq_of s2. +Lemma fmapE (s1 s2 : fmap) : + s1 = s2 <-> seq_of s1 = seq_of s2. Proof. split=>[->|] //. move: s1 s2 => [s1 H1] [s2 H2] /= H. @@ -74,8 +75,9 @@ Fixpoint ins' (k : K) (v : V) (s : seq (K * V)) {struct s} : seq (K * V) := else [:: (k, v)]. Lemma path_ins' s k1 k2 v : - ord k1 k2 -> path ord k1 (map key s) -> - path ord k1 (map key (ins' k2 v s)). + ord k1 k2 -> + path ord k1 (map key s) -> + path ord k1 (map key (ins' k2 v s)). 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. @@ -84,7 +86,8 @@ by rewrite H2 IH //; move: (ord_total k2 k'); rewrite H4 H5. Qed. Lemma sorted_ins' s k v : - sorted ord (map key s) -> sorted ord (map key (ins' k v s)). + sorted ord (map key s) -> + sorted ord (map key (ins' k v s)). Proof. case: s=>// [[k' v']] s /= H. case: ifP=>//= H1; first by rewrite H H1. @@ -95,18 +98,19 @@ Qed. Definition ins k v s := let: FinMap s' p' := s in FinMap (sorted_ins' k v p'). Lemma sorted_filter' k s : - sorted ord (map key s) -> sorted ord (map key (filter (predCk k) s)). + sorted ord (map key s) -> + sorted ord (map key (filter (predCk k) s)). Proof. by move=>H; rewrite -filter_map sorted_filter //; apply: trans. Qed. Definition rem k s := let: FinMap s' p' := s in FinMap (sorted_filter' k p'). -Lemma sorted_behd s : sorted ord (map key s) -> sorted ord (map key (behead s)). +Lemma sorted_behd s : + sorted ord (map key s) -> + sorted ord (map key (behead s)). Proof. by case: s=>//= [[??]] ?; apply: path_sorted. Qed. Definition behd s := let: FinMap s' p' := s in FinMap (sorted_behd p'). - Definition supp s := map key (seq_of s). - Definition cosupp s := map value (seq_of s). End Ops. @@ -119,16 +123,20 @@ Notation fmap := (finMap K V). Notation nil := (nil K V). (* `path_le` specialized to `transitive ord` *) -Lemma ord_path (x y : K) s : ord x y -> path ord y s -> path ord x s. +Lemma ord_path (x y : K) s : + ord x y -> + path ord y s -> + path ord x s. Proof. by apply: path_le. Qed. Lemma last_ins' (x : K) (v : V) s : - path ord x (map key s) -> ins' x v s = (x, v) :: s. + path ord x (map key s) -> + ins' x v s = (x, v) :: s. Proof. by elim: s=>[|[k1 v1] s IH] //=; case: ifP. Qed. Lemma first_ins' (k : K) (v : V) s : (forall x, x \in map key s -> ord x k) -> - ins' k v s = rcons s (k, v). + ins' k v s = rcons s (k, v). Proof. elim: s=>[|[k1 v1] s IH] H //=. move: (H k1); rewrite inE eq_refl; move/(_ (erefl _)). @@ -136,7 +144,9 @@ case: ordP=>// O _; rewrite IH //. by move=>x H'; apply: H; rewrite inE /= H' orbT. Qed. -Lemma notin_path (x : K) s : path ord x s -> x \notin s. +Lemma notin_path (x : K) s : + path ord x s -> + x \notin s. Proof. elim: s=>[|k s IH] //=. rewrite inE negb_or; case/andP=>T1 T2; case: eqP=>H /=. @@ -145,13 +155,19 @@ by apply: IH; apply: ord_path T2. Qed. Lemma path_supp_ord (s : fmap) k : - path ord k (supp s) -> forall m, m \in supp s -> ord k m. + 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: 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. +Lemma all_path_supp (s : fmap) k : + all (ord k) (supp s) -> + path ord k (supp s). +Proof. by rewrite path_sortedE // =>->/=; case: s. Qed. + Lemma notin_filter (x : K) s : x \notin (map key s) -> filter (predk V x) s = [::]. Proof. @@ -1048,12 +1064,12 @@ Definition mapk (m : finMap A V) : finMap B V := (* mapK preserves sorted *) -Lemma sorted_mapk m: - sorted ord (supp (mapk m)). +Lemma sorted_mapk m : sorted ord (supp (mapk m)). Proof. case: (mapk m)=>[s]I //=. Qed. - -Lemma path_mapk m k: path ord k (supp m) -> path ord (f k) (supp (mapk m)). +Lemma path_mapk m k: + path ord k (supp m) -> + path ord (f k) (supp (mapk m)). Proof. elim/fmap_ind': m k =>// k1 v1 s P IH k. rewrite {1}/supp //= {1}seqof_ins // /= => /andP [H]; move/IH=>H1. @@ -1063,7 +1079,6 @@ Qed. Lemma mapk_nil : mapk (nil A V) = nil B V. Proof. by rewrite /mapk //=. Qed. - Lemma mapk_ins k v s : path ord k (supp s) -> mapk (ins k v s) = ins (f k) v (mapk s). @@ -1253,7 +1268,8 @@ by rewrite eq_refl E1 E2; case: eqP. Qed. Lemma zip_rem f1 f2 f x : - zip f1 f2 = Some f -> zip (rem x f1) (rem x f2) = Some (rem x f). + zip f1 f2 = Some f -> + zip (rem x f1) (rem x f2) = Some (rem x f). Proof. case: f1 f2 f=>s1 H1 [s2 H2] [s3 H3] /=; do 2![move: (zip_sorted' _)]. case E1: (zip' s1 s2)=>[t|//]; case E2 : (zip' _ _)=>[q|]; @@ -1262,7 +1278,8 @@ by case: E2=><-{q} pf1 pf2 [E]; congr Some; apply/fmapE; rewrite /= E. Qed. Lemma zip_fnd f1 f2 f x (v : V) : - zip f1 f2 = Some f -> fnd x f = Some v -> + zip f1 f2 = Some f -> + fnd x f = Some v -> exists v1 v2, [/\ zip_f v1 v2 = Some v, fnd x f1 = Some v1 & fnd x f2 = Some v2]. Proof. @@ -1274,7 +1291,7 @@ elim: s1 s2 s E1=>[|[k1 v1] s1 IH]; case=>[|[k2 v2] s2] //= s. case: eqP=>// <-{k2}; case E1: (zip_f v1 v2)=>[w|//]. case E2: (zip' s1 s2)=>[t|//][<-{s}] /=. case: eqP=>[_ [<-]|_]. - (* DEVCOMMENT: exists v1, v2 errors! *) +(* DEVCOMMENT: exists v1, v2 errors! *) - by exists v1; exists v2. by case: (filter (predk V x) t) (IH _ _ E2). Qed. @@ -1298,7 +1315,8 @@ 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 k 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. @@ -1321,7 +1339,8 @@ case: f=>s H; rewrite /supp /= {H}. by elim: s=>[|[k v] s IH] //=; rewrite IH. Qed. -Lemma zunit_disj f1 f2 : disj f1 f2 = disj (zip_unit f1) (zip_unit f2). +Lemma zunit_disj f1 f2 : + disj f1 f2 = disj (zip_unit f1) (zip_unit f2). Proof. case: disjP; case: disjP=>//; rewrite !zunit_supp. - by move=>x H1 H2; move/(_ _ H1); rewrite H2. diff --git a/core/ordtype.v b/core/ordtype.v index 1a83f23..417f0f1 100644 --- a/core/ordtype.v +++ b/core/ordtype.v @@ -67,7 +67,7 @@ Global Arguments ordtype_isOrdered_mixin [_] _. Section type. Local Unset Implicit Arguments. -Polymorphic Cumulative Record +(* Polymorphic Cumulative *) Record type : Type := Pack { sort : Type; _ : axioms_ sort; }. End type. @@ -89,16 +89,17 @@ Global Arguments class : clear implicits. (* Polymorphic Universe annotations added *) Section PolymorphicClonePack. -Polymorphic Universe ou. -Polymorphic Variables (T : Type@{ou}) (cT : type@{ou}). +(* Polymorphic Universe ou. *) +(* Polymorphic Variables (T : Type@{ou}) (cT : type@{ou}). *) +Variables (T : Type) (cT : type). -Polymorphic Definition phant_clone : forall (c : axioms_ T), +(* Polymorphic *) Definition phant_clone : forall (c : axioms_ T), unify Type Type T (sort cT) nomsg -> unify type type cT (Pack (sort:=T) c) nomsg -> type := fun (c : axioms_ T) => fun=> (fun=> (Pack (sort:=T) c)). -Polymorphic Definition pack_ := fun (m : Equality.mixin_of T) +(* Polymorphic *) Definition pack_ := fun (m : Equality.mixin_of T) (m0 : isOrdered.axioms_ T m) => Pack (sort:=T) {|eqtype_hasDecEq_mixin := m; ordtype_isOrdered_mixin := m0 |}. @@ -114,7 +115,7 @@ Notation ordType := Ordered.type. #[reversible] Coercion sort : Ordered.type >-> Sortclass. (* Polymorphic annotation added *) -Polymorphic Definition ordtype_Ordered_class__to__eqtype_Equality_class : +(* Polymorphic *) Definition ordtype_Ordered_class__to__eqtype_Equality_class : forall T : Type, axioms_ T -> Equality.axioms_ T := fun (T : Type) (c : axioms_ T) => {| Equality.eqtype_hasDecEq_mixin := eqtype_hasDecEq_mixin c |}. @@ -125,7 +126,7 @@ Local Arguments ordtype_Ordered_class__to__eqtype_Equality_class : #[reversible] Coercion ordtype_Ordered_class__to__eqtype_Equality_class : ordtype.Ordered.axioms_ >-> eqtype.Equality.axioms_. -Polymorphic Definition ordtype_Ordered__to__eqtype_Equality : +(* Polymorphic *) Definition ordtype_Ordered__to__eqtype_Equality : ordType -> eqType := fun s : ordType => {| Equality.sort := s; Equality.class := class s |}. diff --git a/core/seqext.v b/core/seqext.v index df10cf5..ee757ee 100644 --- a/core/seqext.v +++ b/core/seqext.v @@ -1442,14 +1442,6 @@ Implicit Type ltT leT : rel A. (* ordering with path, seq and last *) -(* renaming *) -#[deprecated(since="fcsl-pcm 1.4.0", note="Use order_path_min instead.")] -Lemma path_all (xs : seq A) x r : - transitive r -> - path r x xs -> - all (r x) xs. -Proof. exact: order_path_min. Qed. - Lemma eq_last (s : seq A) x y : x \in s -> last y s = last x s. diff --git a/examples/array.v b/examples/array.v index 96aaad7..b6612c3 100644 --- a/examples/array.v +++ b/examples/array.v @@ -50,7 +50,7 @@ Definition shape (a : array) (f : {ffun I -> T}) := Program Definition new (x : T) : STsep (emp, [vfun a => shape a [ffun => x]]) := Do (x <-- allocb x #|I|; - ret (Array x)). + ret (Array x)). Next Obligation. (* pull out ghost vars, run the program *) move=>x [] _ /= ->; step=>y; step. @@ -59,8 +59,6 @@ rewrite unitR=>H; split=>//; congr updi; rewrite codomE cardE. by elim: (enum I)=>[|t ts] //= ->; rewrite (ffunE _ _). Qed. -Opaque new. - (* new array corresponding to a domain of a finite function f *) (* loop invariant: *) @@ -187,6 +185,7 @@ Qed. End Array. End Array. + (* Tabulating a finite function over another one *) (* Useful in building linked data structures that are pointed to by *) (* array elements, such as hashtables *) diff --git a/examples/bintree.v b/examples/bintree.v index b05e332..81552af 100644 --- a/examples/bintree.v +++ b/examples/bintree.v @@ -165,8 +165,6 @@ Qed. (* Tree in-order traversal, storing the values in a linked list *) -Opaque insert new. - (* loop invariant: *) (* the subtree is unchanged, its values are prepended to the accumulator list *) Definition inordertravT : Type := forall (ps : ptr * ptr), @@ -206,8 +204,8 @@ move=>s1 _ [hr'][hs][-> Hr' Hs]. apply: [stepX (inorder r ++ xs)]@hs=>//= pa _ [s2][h'][-> H']. (* run the traversal on the left branch with updated list *) apply: [gX l, (a::inorder r ++ xs)]@(hl \+ pa :-> a \+ pa.+1 :-> s2 \+ h')=>//=. -- (* the precondition is satisfied by simple heap manipulation *) - move=>_; exists hl, (pa :-> a \+ (pa.+1 :-> s2 \+ h')). +(* the precondition is satisfied by simple heap manipulation *) +- exists hl, (pa :-> a \+ (pa.+1 :-> s2 \+ h')). by split=>//=; [rewrite !joinA | vauto]. (* the postcondition is also satisfied by simple massaging *) move=>s3 _ [hl''][hs'][-> Hl'' Hs'] _. @@ -229,8 +227,6 @@ Qed. (* Expanding the tree to the right *) -Opaque mknode. - (* loop invariant: the value is added to the rightmost branch *) Definition expandrightT x : Type := forall (p : ptr), STsep {t : tree A} (shape p t, diff --git a/examples/bst.v b/examples/bst.v index eae0dfe..3deedab 100644 --- a/examples/bst.v +++ b/examples/bst.v @@ -20,41 +20,40 @@ From htt Require Import options model heapauto bintree. Section BST. Context {T : ordType}. -(* A binary _search_ tree remains a binary tree structurally, plus: *) -(* 1. its elements must have an ordering defined on them *) -(* 2. a recursive invariant should be satisfied: *) -(* all left branch elements are smaller than the node value & *) -(* all right branch elements are larger than the node value *) +(* Binary _search_ tree is a binary tree structurally, plus: *) +(* 1. its elements have an ordering defined on them *) +(* 2. the following recursive invariant is satisfied: *) +(* - all left branch elements are smaller than the node value & *) +(* - all right branch elements are larger than the node value *) (* Search tree operations *) Fixpoint insert x (t : tree T) : tree T := - if t is Node l a r - then - if x == a then Node l a r - else if ord x a then Node (insert x l) a r - else Node l a (insert x r) + if t is Node l a r then + if x == a then Node l a r + else if ord x a then Node (insert x l) a r + else Node l a (insert x r) else Node leaf x leaf. Fixpoint search (t : tree T) x : bool := - if t is Node l a r - then - if x == a then true - else if ord x a then search l x - else search r x + if t is Node l a r then + if x == a then true + else if ord x a then search l x + else search r x else false. (* Search tree invariant & its interaction with operations *) Fixpoint bst (t : tree T) : bool := - if t is Node l a r - then [&& all (ord^~ a) (inorder l), all (ord a) (inorder r), bst l & bst r] - else true. + if t is Node l a r then + [&& all (ord^~ a) (inorder l), all (ord a) (inorder r), bst l & bst r] + else true. -(* Both BSTs and sorted lists can be used to implement lookup structures, but trees *) -(* are more efficient computationally, while lists are simpler to reason about, *) -(* since the operations on them are associative and commute in many cases. We can *) -(* switch between two specification styles via the in-order traversal. *) +(* BSTs and sorted lists can both be used to implement lookup structures *) +(* but trees are more efficient computationally, while lists are simpler *) +(* to reason about, since the operations on them are associative and *) +(* commute in many cases. One can switch between two specification styles *) +(* via in-order traversal. *) Lemma bst_to_sorted (t : tree T) : bst t = sorted ord (inorder t). @@ -64,7 +63,7 @@ by rewrite sorted_cat_cons_cat /= cats1 sorted_rconsE //= (path_sortedE (@trans T)) andbACA -andbA. Qed. -(* An in-order specification for insertion *) +(* In-order specification for insertion *) Lemma inorder_insert x (t : tree T) : bst t -> perm_eq (inorder (insert x t)) @@ -73,25 +72,22 @@ Proof. elim: t=>//=l IHl a r IHr /and4P [Hal Har /IHl Hl /IHr Hr] {IHl IHr}. rewrite mem_cat inE; case: (ifP [|| _, _ | _]). - case/or3P=>H. - - rewrite H in Hl. - move/allP: Hal=>/(_ x H) /[dup] Hxa /ord_neq/negbTE ->; rewrite Hxa /=. - by apply/permEl/perm_catr. + - rewrite H in Hl; move/allP: Hal=>/(_ x H) /[dup] Hxa /ord_neq/negbTE ->. + by rewrite Hxa; apply/permEl/perm_catr. - by rewrite H. - rewrite H in Hr. - move/allP: Har=>/(_ x H) /[dup] /nsym/negP/negbTE ->. + rewrite H in Hr; move/allP: Har=>/(_ x H) /[dup] /nsym/negP/negbTE ->. move/ord_neq; rewrite eq_sym =>/negbTE -> /=. by apply/permEl/perm_catl; rewrite perm_cons. -move/negbT; rewrite !negb_or; case/and3P=>/negbTE Hxl /negbTE -> /negbTE Hxr; +move/negbT; rewrite !negb_or; case/and3P=>/negbTE Hxl /negbTE -> /negbTE Hxr. rewrite {}Hxl in Hl; rewrite {}Hxr in Hr. -case: ifP=>/= H. -- by rewrite -cat_cons; apply/permEl/perm_catr. +case: ifP=>/= H; first by rewrite -cat_cons; apply/permEl/perm_catr. rewrite -(cat1s x) -(cat1s a) -(cat1s a (inorder r)). rewrite perm_sym perm_catC -!catA catA perm_sym catA. apply/permEl/perm_catl; apply: (perm_trans Hr). by rewrite cats1 -perm_rcons. Qed. -(* As a corollary, we can show that insertion preserves the tree invariant *) +(* Corollary: insertion preserves the tree invariant *) Lemma bst_insert x (t : tree T) : bst t -> bst (insert x t). Proof. elim: t=>//=l IHl a r IHr /and4P [Hal Har Hl Hr]. @@ -138,7 +134,8 @@ Qed. (* Moreover, such representations are equal *) Lemma insert_insert x1 x2 (t : tree T) : bst t -> - inorder (insert x1 (insert x2 t)) = inorder (insert x2 (insert x1 t)). + inorder (insert x1 (insert x2 t)) = + inorder (insert x2 (insert x1 t)). Proof. move=>H; apply: ord_sorted_eq. - by rewrite -bst_to_sorted; do 2!apply: bst_insert. @@ -164,8 +161,6 @@ Qed. (* Pointer-based procedures *) -Opaque mknode. - (* Inserting into the BST *) Definition inserttreeT x : Type := @@ -180,20 +175,22 @@ Program Definition inserttree x : inserttreeT x := ret n else a <-- !p.+1; if x == a then ret p - else if ord x a - then l <-- !p; - l' <-- go l; - p ::= l';; - ret p - else r <-- !p.+2; - r' <-- go r; - p.+2 ::= r';; - ret p)). + else + if ord x a then + l <-- !p; + l' <-- go l; + p ::= l';; + ret p + else + r <-- !p.+2; + r' <-- go r; + p.+2 ::= r';; + ret p)). Next Obligation. (* pull out ghost + precondition, branch on null check *) move=>x go p [t][] i /= H; case: eqP H=>[{p}->|/eqP E] H. -- (* the tree is empty, make a new node *) - apply: vrfV=>V; case: (shape_null V H)=>{t H}->{i V}->. +(* the tree is empty, make a new node *) +- apply: vrfV=>V; case: (shape_null V H)=>{t H}->{i V}->. by apply: [stepE]=>// n m H; step. (* the tree is a node, deconstruct it *) case: (shape_cont E H)=>l[z][r][pl][pr][_][{t H}->{i}->][hl][hr][-> Hl Hr]. @@ -202,8 +199,8 @@ case: (shape_cont E H)=>l[z][r][pl][pr][_][{t H}->{i}->][hl][hr][-> Hl Hr]. step; case: eqP=>_; first by step; vauto. (* branch on comparison, read corresponding pointer *) case: ifP=>Ho; step. -- (* insert in the left branch, update the left pointer *) - apply: [stepX l]@hl=>//= p' h' H'. +(* insert in the left branch, update the left pointer *) +- apply: [stepX l]@hl=>//= p' h' H'. by do 2!step; vauto. (* insert in the right branch, update the right pointer *) apply: [stepX r]@hr=>//= p' h' H'. @@ -220,28 +217,28 @@ Definition searchtreeT x : Type := Program Definition searchtree x : searchtreeT x := ffix (fun (go : searchtreeT x) p => - Do (if p == null - then ret false + Do (if p == null then ret false else a <-- !p.+1; if x == a then ret true - else if ord x a - then l <-- !p; - go l - else r <-- !p.+2; - go r)). + else if ord x a then + l <-- !p; + go l + else + r <-- !p.+2; + go r)). Next Obligation. (* pull out ghost + precondition, branch on null check *) move=>x go p [t][] i /= H; case: eqP H=>[{p}->|/eqP E] H. -- (* tree is empty, it can't contain anything *) - by apply: vrfV=>V; case: (shape_null V H)=>{t H}->{i V}->; step. +(* tree is empty, it can't contain anything *) +- by apply: vrfV=>V; case: (shape_null V H)=>{t H}->{i V}->; step. (* the tree is a node, deconstruct it *) case: (shape_cont E H)=>l[z][r][pl][pr][_][{t H}->{i}->][hl][hr][-> Hl Hr]. (* read the value, compare it to the element, return true if it matches *) step; case: eqP=>_; first by step; vauto. (* branch on comparison, read corresponding pointer *) case: ifP=>Ho; step. -- (* loop on the left branch *) - by apply: [gX l]@hl=>//= b h' [H' E'] _; vauto. +(* loop on the left branch *) +- by apply: [gX l]@hl=>//= b h' [H' E'] _; vauto. (* loop on the right branch *) by apply: [gX r]@hr=>//= b h' [H' E'] _; vauto. Qed. @@ -250,7 +247,7 @@ Qed. (* and lookup finds previously inserted elements *) Program Definition test p x1 x2 : STsep {t : tree T} (fun h => shape p t h /\ bst t, - [vfun (pb : ptr * bool) h => + [vfun (pb : ptr * bool) h => let t' := insert x2 (insert x1 t) in [/\ shape pb.1 t' h, bst t' & pb.2]]) := Do (p1 <-- inserttree x1 p; @@ -268,8 +265,9 @@ step=>_. (* insertions preserve the invariant *) have Hi2: bst (insert x2 t) by apply: bst_insert. have Hi21: bst (insert x2 (insert x1 t)) by do 2!apply: bst_insert. -(* at this point we're done with the separation logic part, i.e. the mutable state *) -(* the only non-trivial goal remaining is search being consistent with insert *) +(* separation logic part (i.e. part about mutable state) is done *) +(* the only remaining non-trivial goal is showing that *) +(* search is consistent with insert *) split=>//{p2 h3 H3}; rewrite {b}Hb. (* switch to in-order lookup *) move: (inorder_search Hi21 x1); rewrite -topredE /= =>->. diff --git a/examples/bubblesort.v b/examples/bubblesort.v index 501aa51..89f570b 100644 --- a/examples/bubblesort.v +++ b/examples/bubblesort.v @@ -23,17 +23,16 @@ From htt Require Import array. From mathcomp Require order. Import order.Order.NatOrder order.Order.TTheory. -(* Mathematics of (bubble) array sorting: *) -(* A theory of permutations built out of (adjacent-element) swaps acting on *) -(* finite functions from bounded nats to ordered values. *) +(* Brief mathematics of (bubble) array sorting: *) +(* Theory of permutations built out of (adjacent-element) swaps acting on *) +(* finite functions from bounded nats to ordered values. *) (* Shorthands for previous/next indices *) (* TODO use fintype.lift instead ? *) Section OrdArith. (* widen by 1 *) -Definition Wo n : 'I_n -> 'I_n.+1 := - widen_ord (leqnSn _). +Definition Wo n : 'I_n -> 'I_n.+1 := widen_ord (leqnSn _). (* increase by 1 *) Definition So n : 'I_n -> 'I_n.+1 := @@ -47,8 +46,8 @@ Proof. by case: i. Qed. Lemma So_WoN n (i : 'I_n) : So i != Wo i. Proof. -case: i=>/= m prf; rewrite /Wo /widen_ord /=; apply/negP=>/eqP; case=>/eqP. -by rewrite eqn_leq leqnSn andbT ltnn. +case: i=>/= m prf; rewrite /Wo /widen_ord /=. +by apply/negP=>/eqP [/eqP]; rewrite eqn_leq leqnSn andbT ltnn. Qed. (* decrease by 1 *) @@ -71,34 +70,31 @@ Proof. by case: i prf=>/= m prf0 prf1; apply/ord_inj. Qed. (* subtract 2 and decrease bound by 1 *) Definition M2lo n (i : 'I_n) (pi : 1%N < i) : 'I_n.-1. Proof. -case: i pi=>m prf /= Hm. -apply: (Ordinal (n:=n.-1) (m:=m.-2)). +case: i pi=>m prf /= Hm; apply: (Ordinal (n:=n.-1) (m:=m.-2)). rewrite ltn_predRL; case: m prf Hm=>//=; case=>//= m Hm _. by apply: ltnW. Defined. Lemma M2lo_eq n (i : 'I_n) (prf : 1%N < i) : nat_of_ord (M2lo prf) = i.-2. Proof. by case: i prf. Qed. - End OrdArith. (* Taking current/next element *) Section OnthCodom. -Variable (A : Type). +Variable A : Type. -Corollary onth_codom1 {n} (i : 'I_n) (f: {ffun 'I_n.+1 -> A}) : - onth (fgraph f) i = Some (f (Wo i)). +Lemma onth_codom1 {n} (i : 'I_n) (f: {ffun 'I_n.+1 -> A}) : + onth (fgraph f) i = Some (f (Wo i)). Proof. by rewrite (onth_codom (Wo i)). Qed. -Corollary onth_codom1S {n} (i : 'I_n) (f: {ffun 'I_n.+1 -> A}) : - onth (fgraph f) i.+1 = Some (f (So i)). +Lemma onth_codom1S {n} (i : 'I_n) (f: {ffun 'I_n.+1 -> A}) : + onth (fgraph f) i.+1 = Some (f (So i)). Proof. by rewrite -(onth_codom (So i)) /= So_eq. Qed. - End OnthCodom. (* Slicing on finfun codomains *) Section CodomSlice. -Variable (n : nat) (A : Type). +Variables (n : nat) (A : Type). Lemma codom1_split (f : {ffun 'I_n.+1 -> A}) (i : 'I_n) : codom f = &:(fgraph f) `]-oo, i : nat[ ++ @@ -107,8 +103,8 @@ Lemma codom1_split (f : {ffun 'I_n.+1 -> A}) (i : 'I_n) : Proof. set i0 := Wo i; set i1 := So i. rewrite {1}codomE (enum_split i0) /= {2}(enum_split i1) /= /heap.indx - (index_enum_ord i0) (index_enum_ord i1) drop_cat size_take - size_enum_ord ltn_ord So_eq ltnn subnn /= map_cat /= map_take map_drop -codomE. + (index_enum_ord i0) (index_enum_ord i1) drop_cat size_take size_enum_ord + ltn_ord So_eq ltnn subnn /= map_cat /= map_take map_drop -codomE. by rewrite /slice.slice/= addn0 addn1 /= drop0 take_size. Qed. @@ -120,14 +116,15 @@ Proof. by rewrite slice_xL // onth_codom1S. Qed. (* TODO notation fails *) Lemma codom1_ax_rcons (f : {ffun 'I_n.+1 -> A}) (a : itv_bound nat) (i : 'I_n) : order.Order.le a (BLeft (i : nat)) -> - &:(fgraph f) (Interval a (BRight (i : nat))) = + &:(fgraph f) (Interval a (BRight (i : nat))) = rcons (&:(fgraph f) (Interval a (BLeft (i : nat)))) (f (Wo i)). Proof. by move=>H; rewrite slice_xR //= onth_codom1 /=. Qed. Lemma codom1_ax_rcons2 (f : {ffun 'I_n.+1 -> A}) (a : itv_bound nat) (i : 'I_n) : order.Order.le a (BLeft (i : nat)) -> &:(fgraph f) (Interval a (BRight i.+1)) = - rcons (rcons (&:(fgraph f) (Interval a (BLeft (i : nat)))) (f (Wo i))) (f (So i)). + rcons (rcons (&:(fgraph f) (Interval a (BLeft (i : nat)))) + (f (Wo i))) (f (So i)). Proof. move=>H; rewrite slice_xR; last first. - by apply: ltW; case: a H=>[[] ax|ab] //; rewrite !bnd_simp ltEnat /= =>/ltnW. @@ -146,17 +143,15 @@ Qed. Lemma codom1_Skk (f : {ffun 'I_n.+1 -> A}) (i : 'I_n) : &:(fgraph f) `[i.+1, i.+1] = [:: f (So i)]. Proof. -rewrite slice_kk /=. -pose i' := cast_ord (esym (card_ord n.+1)) (So i). +rewrite slice_kk /=; pose i' := cast_ord (esym (card_ord n.+1)) (So i). move: (onth_tnth (codom_tuple f) i')=>/=; rewrite So_eq =>->. by move: (@tnth_fgraph _ _ f i'); rewrite enum_val_ord {2}/i' cast_ordKV=>->. Qed. - End CodomSlice. (* Sortedness on codomain slices *) Section CodomSliceOrd. -Variable (n : nat) (A : ordType). +Variables (n : nat) (A : ordType). Lemma codom1_sortedS (f: {ffun 'I_n.+1 -> A}) (i : 'I_n) : oleq (f (Wo i)) (f (So i)) -> @@ -173,14 +168,15 @@ End CodomSliceOrd. (* Swapping adjacent elements and its interaction with slicing *) Section SwapNext. -Variable (n : nat) (A : Type). +Variables (n : nat) (A : Type). Definition swnx (i : 'I_n) : 'S_n.+1 := tperm (Wo i) (So i). Lemma swnxE1 (f : {ffun 'I_n.+1 -> A}) (i : 'I_n) : pffun (swnx i) f (Wo i) = f (So i). Proof. -by rewrite /swnx ffunE; case: tpermP=>// /eqP; rewrite eq_sym (negbTE (So_WoN i)). +rewrite /swnx ffunE; case: tpermP=>// /eqP. +by rewrite eq_sym (negbTE (So_WoN i)). Qed. Lemma swnxE2 (f : {ffun 'I_n.+1 -> A}) (i : 'I_n) : @@ -218,24 +214,22 @@ move=>_ Hx; case: j Hx Hij=>[[] jx|[]] //=. by rewrite addn1 ltEnat /= ltnS leEnat =>->. Qed. -Corollary swnx_ao (f : {ffun 'I_n.+1 -> A}) (i : 'I_n) x k : - (k <= i)%N -> - &:(fgraph (pffun (swnx i) f)) (Interval x (BLeft k)) = - &:(fgraph f) (Interval x (BLeft k)). +Lemma swnx_ao (f : {ffun 'I_n.+1 -> A}) (i : 'I_n) x k : + (k <= i)%N -> + &:(fgraph (pffun (swnx i) f)) (Interval x (BLeft k)) = + &:(fgraph f) (Interval x (BLeft k)). Proof. -move=>Hk. -apply: swnx_notin; rewrite in_itv /= negb_and ltEnat /= -leqNgt; apply/orP; right=>//. -by apply: ltnW. +move=>Hk; apply: swnx_notin; rewrite in_itv /= negb_and ltEnat /= -leqNgt; +by apply/orP; right=>//; apply: ltnW. Qed. -Corollary swnx_oa (f : {ffun 'I_n.+1 -> A}) (i : 'I_n) k y : - (i.+1 <= k)%N -> - &:(fgraph (pffun (swnx i) f)) (Interval (BRight k) y) = - &:(fgraph f) (Interval (BRight k) y). +Lemma swnx_oa (f : {ffun 'I_n.+1 -> A}) (i : 'I_n) k y : + (i.+1 <= k)%N -> + &:(fgraph (pffun (swnx i) f)) (Interval (BRight k) y) = + &:(fgraph f) (Interval (BRight k) y). Proof. -move=>Hk. -apply: swnx_notin; rewrite in_itv /= negb_and ltEnat /= -leqNgt; apply/orP; left=>//. -by apply: ltnW. +move=>Hk; apply: swnx_notin; rewrite in_itv /= negb_and ltEnat /= -leqNgt; +by apply/orP; left=>//; apply: ltnW. Qed. Lemma swnx_split (f : {ffun 'I_n.+1 -> A}) (i : 'I_n) : @@ -265,12 +259,11 @@ Proof. by rewrite codom1_xu_cons swnx_oa // swnxE2. Qed. Lemma swnx_Skk (f : {ffun 'I_n.+1 -> A}) (i : 'I_n) : &:(fgraph (pffun (swnx i) f)) `[i.+1, i.+1] = [:: f (Wo i)]. Proof. by rewrite codom1_Skk swnxE2. Qed. - End SwapNext. (* Swaps/slices with decidable equality *) Section SwapNextEq. -Variable (n : nat) (A : eqType). +Variables (n : nat) (A : eqType). Lemma perm_swnx (f : {ffun 'I_n.+1 -> A}) (i : 'I_n) : perm_eq (fgraph f) (fgraph (pffun (swnx i) f)). @@ -281,22 +274,26 @@ Qed. Lemma perm_swnx_ux (f : {ffun 'I_n.+1 -> A}) (i : 'I_n) k : i <= k -> - perm_eq (&:(fgraph f) `]-oo, k.+1]) (&:(fgraph (pffun (swnx i) f)) `]-oo, k.+1]). + perm_eq &:(fgraph f) `]-oo, k.+1] + &:(fgraph (pffun (swnx i) f)) `]-oo, k.+1]. Proof. move=>H; move: (perm_swnx f i). set f' : {ffun 'I_n.+1 -> A} := pffun (swnx i) f. rewrite {1}(slice_uxou (fgraph f) k.+1) {1}(slice_uxou (fgraph f') k.+1). by rewrite swnx_oa /=; [rewrite perm_cat2r | rewrite ltnS]. Qed. - End SwapNextEq. Section Bubble. -Variable (n : nat) (A : ordType). +Variables (n : nat) (A : ordType). -Opaque Array.write Array.read. +(*****************) +(*****************) +(* Verifications *) +(*****************) +(*****************) -(* We verify 3 versions of the algorithm. *) +(* Verified 3 versions of the algorithm. *) (* *) (* The flag/swapped index is threaded through as a functional argument *) (* instead of storing and reading it from the heap. This would be *) @@ -350,12 +347,13 @@ Program Definition cas_next (a : {array 'I_n.+2 -> A}) (i : 'I_n.+1) : ret true else ret false). Next Obligation. -move=>a i /= [f][] h /= [E V]. -set i0 := Wo i; set i1 := So i. +move=>a i /= [f][] h /= [E V]; set i0 := Wo i; set i1 := So i. do 2!apply: [stepE f, h]=>//= _ _ [->->]. case: oleqP=>H; first by step=>_; exists 1%g; split=>//; rewrite pffunE1. -apply: [stepE f ]=>//= _ _ [-> Vs ]; set fs := finfun [eta f with i0 |-> f i1]. -apply: [stepE fs]=>//= _ _ [-> Vsw]; set fsw := finfun [eta fs with i1 |-> f i0]. +apply: [stepE f ]=>//= _ _ [-> Vs ]. +set fs := finfun [eta f with i0 |-> f i1]. +apply: [stepE fs]=>//= _ _ [-> Vsw]. +set fsw := finfun [eta fs with i1 |-> f i0]. step=>_; exists (swnx i); do!split=>//=. suff: fsw = pffun (swnx i) f by move=>->. rewrite /fsw /fs /swnx; apply/ffunP=>/= x; rewrite !ffunE /= ffunE /=. @@ -363,8 +361,6 @@ case: tpermP=>[->|->|/eqP/negbTE->/eqP/negbTE->] //; rewrite eqxx //. by rewrite /i1 eq_sym (negbTE (So_WoN i)). Qed. -Opaque cas_next. - Definition bubble_pass_loop (a : {array 'I_n.+2 -> A}) := forall isw : 'I_n.+1 * bool, STsep {f : {ffun 'I_n.+2 -> A}} @@ -395,10 +391,10 @@ Next Obligation. move=>a loop _ i sw [f][] h /= [H Hs]. set i0 := Wo i; set i1 := So i. apply: [stepE f]=>//= sw0 m [p][{}H Hsw]; case: decP=>Hin. -- (* i < n, recursive call *) - apply: [gE (pffun p f)]=>//=. - - (* precondition *) - split=>//; rewrite negb_or andbC Sbo_eq. +(* i < n, recursive call *) +- apply: [gE (pffun p f)]=>//=. + (* precondition *) + - split=>//; rewrite negb_or andbC Sbo_eq. case: sw0 Hsw=>//=; case=>{p H}-> Hf; rewrite pffunE1. (* swap didn't happen in this iteration *) apply/(implyb_trans Hs)/implyP=>{Hs}. @@ -416,8 +412,10 @@ step=>Vm; exists p; split=>//; first by apply/implyP=>->. (* swap didn't happen in last iteration, check initial flag *) case: sw0 Hsw=>/=; first by rewrite orbT. (* flag wasn't set, so the pass didn't swap anything - the array is sorted *) -case=>-> Hf; rewrite orbF eqxx /=; apply/(implyb_trans Hs)/implyP=>{}Hs. -rewrite (slice_usize (codom f)) size_codom /= card_ord slice_oSR -[X in `]-oo, X.+1]]Hin. +case=>-> Hf; rewrite orbF eqxx /=. +apply/(implyb_trans Hs)/implyP=>{}Hs. +rewrite (slice_usize (codom f)) size_codom /= card_ord +slice_oSR -[X in `]-oo, X.+1]]Hin. by apply: codom1_sortedS Hs. Qed. Next Obligation. @@ -427,8 +425,6 @@ split=>//; apply: sorted1. by rewrite slice_size /= add0n subn0 size_codom /= card_ord. Qed. -Opaque bubble_pass. - Definition bubble_sort_loop (a : {array 'I_n.+2 -> A}) : Type := unit -> STsep {f : {ffun 'I_n.+2 -> A}} @@ -438,7 +434,8 @@ Definition bubble_sort_loop (a : {array 'I_n.+2 -> A}) : Type := h \In Array.shape a (pffun p f) /\ sorted oleq (fgraph (pffun p f))]). -Program Definition bubble_sort (a : {array 'I_n.+2 -> A}) : bubble_sort_loop a := +Program Definition bubble_sort (a : {array 'I_n.+2 -> A}) : + bubble_sort_loop a := ffix (fun (go : bubble_sort_loop a) _ => Do (sw <-- bubble_pass a; if sw @@ -461,8 +458,6 @@ End Bubble. Section BubbleOpt. Variable (n : nat) (A : ordType). -Opaque cas_next. - (*******************) (* optimization #1 *) (*******************) @@ -488,26 +483,26 @@ Opaque cas_next. Definition bubble_pass_opt_loop (a : {array 'I_n.+2 -> A}) (k : 'I_n.+1) := forall isw : 'I_n.+1 * bool, STsep {f : {ffun 'I_n.+2 -> A}} - (fun h => [/\ h \In Array.shape a f, - (* all elements before f(k+2) are smaller than it *) - allrel oleq (&:(fgraph f) `]-oo, k.+1]) - (&:(fgraph f) `]k.+1, +oo[), - (* and the suffix is sorted *) - sorted oleq (&:(fgraph f) `]k.+1, +oo[), - (* we don't touch it *) - isw.1 <= k, - (* all elements before f(i) are smaller than it *) - all (oleq^~ (f (Wo isw.1))) (&:(fgraph f) `]-oo, isw.1 : nat[) & - ~~ isw.2 ==> sorted oleq (&:(fgraph f) `]-oo, isw.1 : nat[)], - [vfun (b : bool) h => - exists (p : 'S_n.+2), let f' := pffun p f in - [/\ h \In Array.shape a f', - isw.2 ==> b & - if b then - allrel oleq (&:(fgraph f') `]-oo, k : nat]) - (&:(fgraph f') `]k : nat, +oo[) && - sorted oleq (&:(fgraph f') `]k : nat, +oo[) - else (p == 1%g) && sorted oleq (fgraph f)]]). + (fun h => [/\ h \In Array.shape a f, + (* all elements before f(k+2) are smaller than it *) + allrel oleq (&:(fgraph f) `]-oo, k.+1]) + (&:(fgraph f) `]k.+1, +oo[), + (* and the suffix is sorted *) + sorted oleq (&:(fgraph f) `]k.+1, +oo[), + (* we don't touch it *) + isw.1 <= k, + (* all elements before f(i) are smaller than it *) + all (oleq^~ (f (Wo isw.1))) (&:(fgraph f) `]-oo, isw.1 : nat[) & + ~~ isw.2 ==> sorted oleq (&:(fgraph f) `]-oo, isw.1 : nat[)], + [vfun (b : bool) h => + exists (p : 'S_n.+2), let f' := pffun p f in + [/\ h \In Array.shape a f', + isw.2 ==> b & + if b then + allrel oleq (&:(fgraph f') `]-oo, k : nat]) + (&:(fgraph f') `]k : nat, +oo[) && + sorted oleq (&:(fgraph f') `]k : nat, +oo[) + else (p == 1%g) && sorted oleq (fgraph f)]]). Program Definition bubble_pass_opt (a : {array 'I_n.+2 -> A}) (k : 'I_n.+1) : STsep {f : {ffun 'I_n.+2 -> A}} @@ -538,13 +533,13 @@ Next Obligation. move=>a k loop _ i sw [f][] h /= [H Hak Hsk Hik Hai Hsi]. set i0 := Wo i; set i1 := So i. apply: [stepE f]=>//= sw0 m [p][Hm Hsw]; case: decP=>{}H. -- (* i < k, recursive call *) - apply: [gE (pffun p f)]=>//=. - - (* precondition *) - rewrite Sbo_eq Wo_Sbo_eq negb_or slice_oSR. +(* i < k, recursive call *) +- apply: [gE (pffun p f)]=>//=. + (* precondition *) + - rewrite Sbo_eq Wo_Sbo_eq negb_or slice_oSR. case: sw0 Hsw=>/=; case=>Ep Hf; rewrite {p}Ep ?pffunE1 in Hm *. - - (* swap happened before the call *) - rewrite andbF /= swnx_oa //; split=>//. + (* swap happened before the call *) + - rewrite andbF /= swnx_oa //; split=>//. - rewrite (@allrel_in_l _ _ _ _ &:(codom f) `]-oo, k.+1] ) //. by apply/perm_mem; rewrite perm_sym; apply: perm_swnx_ux. by rewrite swnx_ux_rcons all_rcons swnxE2 Hai andbT; apply: ordW. @@ -565,8 +560,8 @@ have {H}Hik : nat_of_ord i = k. step=>Vm; exists p; split=>//; first by apply/implyP=>->. rewrite -Hik in Hak Hsk *; rewrite -slice_oSL. case: sw0 Hsw=>/=; case=>-> Hf. -- (* swap happened on last iteration *) - rewrite orbT swnx_xu_cons /=; apply/andP; split. +(* swap happened on last iteration *) +- rewrite orbT swnx_xu_cons /=; apply/andP; split. - rewrite swnx_ux_rcons allrel_rconsl allrel_consr /=. move: Hak; rewrite codom1_ax_rcons2 // !allrel_rconsl -/i0 -/i1 -andbA. by case/and3P=>-> _ ->; rewrite (ordW Hf) /= !andbT. @@ -575,8 +570,8 @@ case: sw0 Hsw=>/=; case=>-> Hf. by rewrite codom1_ax_rcons2 // 4!(mem_rcons, inE) eqxx /= orbT. (* swap didn't happen on last iteration *) rewrite orbF pffunE1 eqxx /=; case: sw Hsi=>/=[_|]. -- (* flag was set *) - apply/andP; split. +(* flag was set *) +- apply/andP; split. - rewrite codom1_ax_rcons // codom1_xu_cons allrel_rconsl allrel_consr /=. move: Hak; rewrite codom1_ax_rcons2 // !allrel_rconsl -andbA -/i0 -/i1. case/and3P=>->-> _; rewrite Hf /= !andbT. @@ -584,10 +579,13 @@ rewrite orbF pffunE1 eqxx /=; case: sw Hsi=>/=[_|]. rewrite codom1_xu_cons /= (path_sortedE (@otrans A)) Hsk andbT. by move: Hak; rewrite codom1_ax_rcons2 // !allrel_rconsl -andbA; case/and3P. (* flag wasn't set, so the pass didn't swap anything - the array is sorted *) -move=>Hs2; rewrite (codom1_split f i) -/i0 -/i1 /=; rewrite Hik in Hai Hs2 Hsk *. -rewrite (sorted_pairwise (@otrans A)) pairwise_cat /= -!(sorted_pairwise (@otrans A)). +move=>Hs2; rewrite (codom1_split f i) -/i0 -/i1 /=. +rewrite Hik in Hai Hs2 Hsk *. +rewrite (sorted_pairwise (@otrans A)) pairwise_cat /=. +rewrite -!(sorted_pairwise (@otrans A)). rewrite Hf Hsk Hs2 /= andbT !allrel_consr -!andbA Hai /=. -move: Hak; rewrite -Hik codom1_ax_rcons2 // !allrel_rconsl -andbA; case/and3P=>->->->/=. +move: Hak; rewrite -Hik codom1_ax_rcons2 // !allrel_rconsl -andbA. +case/and3P=>->->->/=. by rewrite andbT Hik; apply/sub_all/Hai=>z /otrans; apply. Qed. Next Obligation. @@ -596,8 +594,6 @@ apply: [gE f]=>//= [|v m [p][Hm _ Hv] _]; last by exists p. by rewrite itv_under0R. Qed. -Opaque bubble_pass_opt. - Definition bubble_sort_opt_loop (a : {array 'I_n.+2 -> A}) : Type := forall (k : 'I_n.+1), STsep {f : {ffun 'I_n.+2 -> A}} @@ -625,9 +621,10 @@ Program Definition bubble_sort_opt (a : {array 'I_n.+2 -> A}) : in go ord_max). Next Obligation. move=>a go k [f][] h /= [H Ha Hs]. -apply: [stepE f]=>{H Ha Hs}//=; case=>m; case=>p [Hm /andP [H Hs]]; last first. -- (* id permutation *) - rewrite {p H}(eqP H) in Hm. +apply: [stepE f]=>{H Ha Hs}//=; case=>m; case=>p [Hm /andP [H Hs]]; +last first. +(* id permutation *) +- rewrite {p H}(eqP H) in Hm. by step=>_; exists 1%g; split=>//; rewrite pffunE1. apply: [gE (pffun p f)]=>//=; last first. - by move=> _ m2 [p'][H2 Hs'] _; exists (p' * p)%g; rewrite pffunEM. @@ -654,8 +651,6 @@ End BubbleOpt. Section BubbleOpt2. Variable (n : nat) (A : ordType). -Opaque cas_next. - (*******************) (* optimization #2 *) (*******************) @@ -725,18 +720,19 @@ Next Obligation. move=>a k loop _ i ls [f][] h /= [Hs Hak Hsk /andP [Hils Hik] Hai Hsi]. set i0 := Wo i; set i1 := So i. apply: [stepE f]=>//= sw m [p][Hm Hsw]; case: decP=>H. -- (* i < k, recursive call *) - apply: [gE (pffun p f)]=>//=. - - (* precondition *) - rewrite Sbo_eq. +(* i < k, recursive call *) +- apply: [gE (pffun p f)]=>//=. + (* precondition *) + - rewrite Sbo_eq. case: sw Hsw=>/=; case=>Ep Hf; rewrite {p}Ep ?pffunE1 in Hm *. - - (* swap happened before the call *) - rewrite So_eq swnx_oa; last by rewrite ltnS; apply: ltnW. + (* swap happened before the call *) + - rewrite So_eq swnx_oa; last by rewrite ltnS; apply: ltnW. rewrite swnx_Skk /=; split=>//. - rewrite (@allrel_in_l _ _ _ _ &:(codom f) `]-oo, k.+1]) //. by apply/perm_mem; rewrite perm_sym; apply: perm_swnx_ux. - by rewrite ltnS leqnn. - rewrite allrel1r slice_oSR (slice_split (x:=ls) _ true); last by rewrite in_itv. + rewrite allrel1r slice_oSR (slice_split (x:=ls) _ true); + last by rewrite in_itv. rewrite /= all_cat; apply/andP; split. - rewrite swnx_ao //; move: Hai. by rewrite codom1_ax_rcons // allrel_rconsr; case/andP. @@ -769,9 +765,10 @@ have {H}Hik : nat_of_ord i = k. by move/negP: H; rewrite -leqNgt. rewrite -{loop k}Hik in Hak Hsk *. step=>Vm; exists p; case: sw Hsw=>/=; case=>Ep Hf. -- (* swap happened on last iteration *) - rewrite So_eq; split=>//=; first by apply: ltnW. - - rewrite Ep slice_oSR swnx_ux_rcons swnx_xu_cons allrel_rconsl allrel_consr /=. +(* swap happened on last iteration *) +- rewrite So_eq; split=>//=; first by apply: ltnW. + - rewrite Ep slice_oSR swnx_ux_rcons swnx_xu_cons + allrel_rconsl allrel_consr /=. move: Hak; rewrite codom1_ax_rcons2 // !allrel_rconsl -/i0 -/i1 -andbA. case/and3P=>-> _ ->; rewrite (ordW Hf) /= !andbT. move: Hai; rewrite codom1_ax_rcons //= allrel_rconsr=>/andP [_ Hals]. @@ -814,8 +811,6 @@ split=>//; first by rewrite itv_underR. by apply: sorted1; rewrite slice_size size_codom card_ord. Qed. -Opaque bubble_pass_opt2. - Definition bubble_sort_opt2_loop (a : {array 'I_n.+2 -> A}) : Type := forall (k : 'I_n.+1), STsep {f : {ffun 'I_n.+2 -> A}} @@ -850,8 +845,10 @@ case: decP=>Hx; last first. - by move=>Hx _; rewrite Hx -slice_0L slice_uu in Hs. case: (nat_of_ord x) Ha Hs=>//; case=>//= Ha1 Hs1 _ _. rewrite (slice_uoxu (codom (pffun p f)) 1). - rewrite (sorted_pairwise (@otrans A)) pairwise_cat -!(sorted_pairwise (@otrans A)). - by apply/and3P; split=>//; apply: sorted1; rewrite slice_size size_codom card_ord. + rewrite (sorted_pairwise (@otrans A)) pairwise_cat. + rewrite -!(sorted_pairwise (@otrans A)). + apply/and3P; split=>//; apply: sorted1. + by rewrite slice_size size_codom card_ord. apply: [gE (pffun p f)]=>//=; last first. - move=> _ m2 [p'][H2 Hs'] _; exists (p' * p)%g. by rewrite pffunEM. diff --git a/examples/congmath.v b/examples/congmath.v index 46cdc88..b235213 100644 --- a/examples/congmath.v +++ b/examples/congmath.v @@ -11,10 +11,6 @@ See the License for the specific language governing permissions and limitations under the License. *) -(**********************) -(* Congruence closure *) -(**********************) - From HB Require Import structures. From Coq Require Import Recdef ssreflect ssrbool ssrfun. From mathcomp Require Import eqtype choice ssrnat seq bigop fintype finfun. @@ -22,11 +18,21 @@ From pcm Require Import options prelude ordtype finmap pred seqext. Ltac add_morphism_tactic := SetoidTactics.add_morphism_tactic. +(**********************) +(* Congruence closure *) +(**********************) + +(* Background mathematics for the verification of the *) +(* Barcelogic Congruence Closure algorithm. *) +(* Described in the POPL10 paper *) +(* Structure the verification of heap-manipulating programs *) +(* by Nanevski, Vafeiadis and Berdine *) + (**************************) (* Constants with arities *) (**************************) -Inductive constant : Type := const_with_arity of nat & nat. +Inductive constant : Set := const_with_arity of nat & nat. (* constants are an equality type *) Definition eqcnst (c1 c2 : constant) : bool := @@ -55,43 +61,76 @@ Definition consttag (u : {i : nat & nat}) := Lemma tagconstK : cancel tagconst consttag. Proof. by case. Qed. -(* Lemma consttagK : cancel consttag tagconst. Proof. by case. Qed. -*) HB.instance Definition _ := Countable.copy constant (can_type tagconstK). (***********************************************************) (* The symbols are a predetermined finite set of constants *) (***********************************************************) -Variable symbs : seq constant. -Definition symb := seq_sub symbs. -HB.instance Definition _ := Equality.on symb. -HB.instance Definition _ := Finite.on symb. - -HB.instance Definition symb_ord_mix := - isOrdered.Build symb (fintype_is_ordtype symb). +(* Given sequences s, symb s is the type of elements of s. *) +(* One could use seq_sub from fintype.v here, and just set *) +(* Definition symb (s : seq constant) := seq_sub s. *) +(* However, this is expensive in terms of universes, *) +(* as it forces symb to be Type, not Set. *) +(* Thus, it's best here to roll out a bespoke fintype structure *) +(* copying relevant code from seq_sub from fintype.v *) +(* but setting universe to Set from the beginning. *) + +Record symb (s : seq constant) : Set := + SeqSymb {symb_val : constant; + symb_valP : in_mem symb_val (@mem constant _ s)}. + +#[warning="-projection-no-head-constant"] +HB.instance Definition _ s := [isSub for (@symb_val s)]. +HB.instance Definition _ s := [Equality of symb s by <:]. +Definition symb_enum s : seq (symb s) := undup (pmap insub s). +Lemma mem_symb_enum s x : x \in symb_enum s. +Proof. by rewrite mem_undup mem_pmap -valK map_f ?symb_valP. Qed. +Lemma val_symb_enum s : uniq s -> map val (symb_enum s) = s. +Proof. +move=> Us; rewrite /symb_enum undup_id ?pmap_sub_uniq //. +rewrite (pmap_filter (insubK _)); apply/all_filterP. +by apply/allP => x; rewrite isSome_insub. +Qed. +Definition symb_pickle s x := index x (symb_enum s). +Definition symb_unpickle s n := nth None (map some (symb_enum s)) n. +Lemma symb_pickleK s : pcancel (@symb_pickle s) (@symb_unpickle s). +Proof. +rewrite /symb_unpickle => x. +by rewrite (nth_map x) ?nth_index ?index_mem ?mem_symb_enum. +Qed. +Definition symb_isCountable s := isCountable.Build (symb s) (@symb_pickleK s). +Fact symb_axiom s : Finite.axiom (symb_enum s). +Proof. exact: Finite.uniq_enumP (undup_uniq _) (@mem_symb_enum s). Qed. +Definition symb_isFinite s := isFinite.Build (symb s) (@symb_axiom s). +HB.instance Definition _ s := [Choice of symb s by <:]. +HB.instance Definition _ s : isCountable (symb s) := symb_isCountable s. +HB.instance Definition _ s : isFinite (symb s) := symb_isFinite s. +HB.instance Definition symb_ord_mix s := + isOrdered.Build (symb s) (fintype_is_ordtype (symb s)). (* manual canonical declaration, as HB fails to declare it *) -Canonical symb_ordType : ordType := - Ordered.Pack (sort:=symb) (Ordered.Class symb_ord_mix). +Canonical symb_ordType s : ordType := + Ordered.Pack (sort:=symb s) (Ordered.Class (symb_ord_mix s)). (****************************) (* Expressions over symbols *) (****************************) -Inductive exp : Type := const of symb | app of exp & exp. +(* Consequently, expressions must be Type, not Set *) +Inductive exp s : Set := const of symb s | app of exp s & exp s. (* expressions are an equality type *) -Fixpoint eqexp (e1 e2 : exp) {struct e1} : bool := +Fixpoint eqexp s (e1 e2 : exp s) {struct e1} : bool := match e1, e2 with | const c1, const c2 => c1 == c2 | app e1 e2, app e3 e4 => (eqexp e1 e3) && (eqexp e2 e4) | _, _ => false end. -Lemma eqexpP : Equality.axiom eqexp. +Lemma eqexpP s : Equality.axiom (@eqexp s). Proof. elim=>[c1|e1 IH1 e2 IH2][c2|e3 e4] /=; try by apply: ReflectF. - case: eqP=>[->|H]; last by apply: ReflectF; case. @@ -101,23 +140,34 @@ case: IH2=>[->|H]; last by apply: ReflectF; case. by apply: ReflectT. Qed. -HB.instance Definition _ := hasDecEq.Build exp eqexpP. +HB.instance Definition _ s := hasDecEq.Build (@exp s) (@eqexpP s). (****************************************) (* Congruence relations over expression *) (****************************************) (* R is monotone if it is preserved by application *) +Section Congruence. +Variable s : seq constant. +Notation symb := (symb s). +Notation exp := (exp s). + Notation rel_exp := (Pred (exp*exp)). -Definition Reflexive (r : rel_exp) := forall x, (x, x) \In r. -Definition Symmetric (r : rel_exp) := forall x y, (x, y) \In r -> (y, x) \In r. -Definition Transitive (r : rel_exp) := forall y x z, (x, y) \In r -> (y, z) \In r -> (x, z) \In r. -Definition Equivalence (r : rel_exp) := Reflexive r /\ Symmetric r /\ Transitive r. -Definition Antisymmetric (r : rel_exp) := forall x y, (x, y) \In r -> (y, x) \In r -> x = y. +Definition Reflexive (r : rel_exp) := + forall x, (x, x) \In r. +Definition Symmetric (r : rel_exp) := + forall x y, (x, y) \In r -> (y, x) \In r. +Definition Transitive (r : rel_exp) := + forall y x z, (x, y) \In r -> (y, z) \In r -> (x, z) \In r. +Definition Equivalence (r : rel_exp) := + Reflexive r /\ Symmetric r /\ Transitive r. +Definition Antisymmetric (r : rel_exp) := + forall x y, (x, y) \In r -> (y, x) \In r -> x = y. Definition monotone (R : rel_exp) : Prop := - forall f1 f2 e1 e2, (f1, f2) \In R -> (e1, e2) \In R -> (app f1 e1, app f2 e2) \In R. + forall f1 f2 e1 e2, (f1, f2) \In R -> + (e1, e2) \In R -> (app f1 e1, app f2 e2) \In R. Definition congruence (R : rel_exp) := Equivalence R /\ monotone R. @@ -147,7 +197,6 @@ move=>f1 f2 e1 e2 H1 H2 C H3 H4. by apply: mono_cong (H1 C H3 H4) (H2 C H3 H4). Qed. -#[export] Hint Resolve cong_clos : core. Lemma reflC (R : rel_exp) : Reflexive (closure R). @@ -162,7 +211,6 @@ Proof. by apply: trans_cong. Qed. Lemma monoC (R : rel_exp) : monotone (closure R). Proof. by apply: mono_cong. Qed. -#[export] Hint Resolve reflC symC : core. (* lemmas about closure *) @@ -176,10 +224,12 @@ by rewrite H. Qed. Add Morphism closure with - signature Morphisms.respectful (fun e1 e2 => e1 <~> e2) (fun e1 e2 => e1 <~> e2) as closure_morph. + signature Morphisms.respectful (fun e1 e2 => e1 <~> e2) + (fun e1 e2 => e1 <~> e2) as closure_morph. Proof. by apply: closE. Qed. -Lemma clos_clos (R Q : rel_exp) : closure (closure R +p Q) <~> closure (R +p Q). +Lemma clos_clos (R Q : rel_exp) : + closure (closure R +p Q) <~> closure (R +p Q). Proof. case=>e1 e2; split=>H1 C H2 H3; apply: H1 =>// [[x y]] H4. - by case: H4=>H4; [apply: H4=>// [[x1 y1]] H4|]; apply: H2; [left | right]. @@ -187,18 +237,21 @@ case: H4=>H4; apply: H2; [left | right]=>//. by move=>C0 H5 _; apply: H5. Qed. -Lemma clos_idemp (R : rel_exp) : closure (closure R) <~> closure R. +Lemma clos_idemp (R : rel_exp) : + closure (closure R) <~> closure R. Proof. by rewrite -(orr0 (closure R)) clos_clos !orr0. Qed. Lemma closKR (R1 R2 K : rel_exp) : - closure R1 <~> closure R2 -> closure (K +p R1) <~> closure (K +p R2). + closure R1 <~> closure R2 -> + closure (K +p R1) <~> closure (K +p R2). Proof. move=>H. by rewrite 2!(orrC K) -(clos_clos R1) -(clos_clos R2) H. Qed. Lemma closRK (R1 R2 K : rel_exp) : - closure R1 <~> closure R2 -> closure (R1 +p K) <~> closure (R2 +p K). + closure R1 <~> closure R2 -> + closure (R1 +p K) <~> closure (R2 +p K). Proof. by rewrite 2!(orrC _ K); apply: closKR. Qed. Lemma closI (R : rel_exp) : R ~> closure R. @@ -226,7 +279,8 @@ by rewrite H1 -clos_clos H2 clos_clos orrAC orrA. Qed. Lemma sub_closKR (R1 R2 K : rel_exp) : - (closure R1 ~> closure R2) -> closure (K +p R1) ~> closure (K +p R2). + (closure R1 ~> closure R2) -> + closure (K +p R1) ~> closure (K +p R2). Proof. rewrite -!closAK=>H. rewrite -(orrC (closure _)) clos_clos orrAC -orrA orrI. @@ -236,7 +290,8 @@ by rewrite orrC. Qed. Lemma sub_closRK (R1 R2 K : rel_exp) : - (closure R1 ~> closure R2) -> closure (R1 +p K) ~> closure (R2 +p K). + (closure R1 ~> closure R2) -> + closure (R1 +p K) ~> closure (R2 +p K). Proof. by move=>H; rewrite -!(orrC K); apply: sub_closKR. Qed. Lemma sub_closI (R1 R2 : rel_exp) : @@ -246,7 +301,8 @@ rewrite -orrAb -closAK -(orrC (closure _)) clos_clos => ->. by rewrite (orrC R2) -orrA orrI. Qed. -(* relation that only relates the constant symbols with their image under a function *) +(* relation that only relates the constant symbols *) +(* with their image under a function *) Definition graph (f : symb -> exp) : rel_exp := [Pred e1 e2 | (e1, e2) \in image (fun s => (const s, f s)) predT]. @@ -277,8 +333,10 @@ HB.instance Definition _ := hasDecEq.Build Eq eqEqP. (* interpreting equations as relations *) Definition eq2rel (eq : Eq) : rel_exp := match eq with - | simp_eq c1 c2 => [Pred x y | x = const c1 /\ y = const c2] - | comp_eq c c1 c2 => [Pred x y | x = const c /\ y = app (const c1) (const c2)] + | simp_eq c1 c2 => + [Pred x y | x = const c1 /\ y = const c2] + | comp_eq c c1 c2 => + [Pred x y | x = const c /\ y = app (const c1) (const c2)] end. Fixpoint eqs2rel (eqs : seq Eq) : rel_exp := @@ -289,24 +347,25 @@ Definition symb2eq (t : symb*symb*symb) := let: (c, c1, c2) := t in comp_eq c c1 c2. (* inverting relations equations represented as triples *) -Definition invert (x y : exp) (s : seq (symb*symb*symb)) : - (x, y) \In eqs2rel (map symb2eq s) <-> +Definition invert (x y : exp) (ss : seq (symb*symb*symb)) : + (x, y) \In eqs2rel (map symb2eq ss) <-> exists c c1 c2, - [/\ x = const c, y = app (const c1) (const c2) & (c, c1, c2) \in s]. + [/\ x = const c, y = app (const c1) (const c2) & + (c, c1, c2) \in ss]. Proof. split. -- elim: s=>[|[[c c1] c2] s IH] /=; first by case. +- elim: ss=>[|[[c c1] c2] ss IH] /=; first by case. case; first by case=>->->; exists c, c1, c2; rewrite inE eq_refl. move/IH=>[d][d1][d2][->->H]. by exists d, d1, d2; rewrite inE H orbT. -elim: s=>[|[[c c1] c2] s IH] [d][d1][d2][H1 H2] //. +elim: ss=>[|[[c c1] c2] ss IH] [d][d1][d2][H1 H2] //. rewrite inE; case/orP; first by rewrite H1 H2; case/eqP=>->->->; left. move=>T; apply: sub_orr; apply: IH. by exists d, d1, d2; rewrite H1 H2 T. Qed. -(* We also keep track of pending equations. These are equations that have *) -(* been placed into the input list, but have not yet been processed and *) +(* One must also keep track of pending equations. These are equations *) +(* placed into the input list, but have not yet been processed and *) (* inserted into the data structure for computing the relation. The *) (* pending equations are either simple equations of the form c1 = c2 or *) (* are composite ones of the form (c = c1 c2, d = d1 d2), where c1, d1 *) @@ -347,7 +406,6 @@ Definition pend2eq (p : pend) := (* The definition of the data structures involved in the algorithm. *) (********************************************************************) - Record data : Type := Data {(* An array, indexed by symbs, containing for each symb its *) (* current representative (for the relation obtained by the *) @@ -373,13 +431,11 @@ Definition reps D : seq symb := undup (map (rep D) (enum predT)). Lemma uniq_reps D : uniq (reps D). Proof. by rewrite undup_uniq. Qed. -#[export] Hint Resolve uniq_reps : core. Lemma rep_in_reps D a : rep D a \in reps D. Proof. by move=>*; rewrite mem_undup; apply: map_f; rewrite mem_enum. Qed. -#[export] Hint Resolve rep_in_reps : core. (* The symbols and their representatives are a base case in defining the closure *) @@ -391,7 +447,6 @@ apply: (@closI _ (const a, const (rep D a))); apply: sub_orl; rewrite /= /rep2re by rewrite mem_map /= ?mem_enum // => x1 x2 [->]. Qed. -#[export] Hint Resolve clos_rep : core. (* relation defined by the lookup table *) @@ -423,12 +478,12 @@ Definition CRel D := closure (rep2rel D +p Lemma cong_rel D : congruence (CRel D). Proof. by move=>*; apply: cong_clos. Qed. -#[export] Hint Resolve cong_rel : core. Lemma clos_rel D R a : (const a, const (rep D a)) \In closure (CRel D +p R). Proof. by rewrite /CRel clos_clos orrA; apply: clos_rep. Qed. + (*******************************************) (* Congruence Closure Code and Termination *) (*******************************************) @@ -458,10 +513,11 @@ Lemma join_class_perm (D : data) (a' b' : symb) : (filter (predC1 a') (reps D)). Proof. have ffun_comp: forall (A B C : finType) (f1 : B -> A) (f2 : C -> B), - finfun (f1 \o f2) =1 f1 \o f2 :> (C -> A) by move=>A B C f1 f2 x; rewrite ffunE. + finfun (f1 \o f2) =1 f1 \o f2 :> (C -> A). +- by move=>A B C f1 f2 x; rewrite ffunE. have maps_ffun_comp : forall (A B C : finType) (f1 : B -> C) (f2 : A -> B) s, map (finfun (comp f1 f2)) s = map f1 (map f2 s). -- by move=>A B C f1 f2 s; elim: s=>[|x s IH] //=; rewrite IH /= ffunE. +- by move=>A B C f1 f2 ss; elim: ss=>[|x ss IH] //=; rewrite IH /= ffunE. move=>H H2. apply: uniq_perm; [|apply: filter_uniq|]; try by rewrite undup_uniq. move: H2. @@ -473,17 +529,17 @@ have L1 : forall x, a' != f x. have L2 : forall x, x != a' -> x = f x. - by rewrite /f=>x; case: eqP. rewrite eq_sym in H. -move: (map _ _)=>s H1 x. +move: (map _ _)=>ss H1 x. rewrite !mem_undup. case E1 : (x == a'). - rewrite (eqP E1) mem_filter /= eq_refl /=. - elim: s {H1} =>// t s H1. + elim: ss {H1} =>// t ss H1. by rewrite inE /= (negbTE (L1 t)) H1. case E2 : (x == b'). - rewrite mem_filter /= E1 (eqP E2) H1 (L2 b') //=. by apply: map_f. rewrite mem_filter /= E1 /=. -elim: s {H1}=>// t s; rewrite /= 2!inE=>->. +elim: ss {H1}=>// t ss; rewrite /= 2!inE=>->. case: (t =P a')=>[->|H2]. - by rewrite E1 /f eq_refl E2. by rewrite -(L2 t) //; case: eqP H2. @@ -509,25 +565,27 @@ Qed. (* joining the use lists *) Fixpoint join_use' (D : data) (a' b' : symb) - (old_use : seq (symb * symb * symb)) {struct old_use} : data := + (old_use : seq (symb * symb * symb)) {struct old_use} : data := match old_use with | [::] => D | (c, c1, c2)::p' => if fnd (rep D c1, rep D c2) (lookup D) is Some (d, d1, d2) then - let: newD := Data (rep D) (class D) - (finfun (fun r => if r == a' then behead (use D r) else - use D r)) - (lookup D) - (comp_pend (c, c1, c2) (d, d1, d2) :: pending D) - in join_use' newD a' b' p' + let: newD := + Data (rep D) (class D) + (finfun (fun r => if r == a' then behead (use D r) + else use D r)) + (lookup D) + (comp_pend (c, c1, c2) (d, d1, d2) :: pending D) + in join_use' newD a' b' p' else - let: newD := Data (rep D) (class D) - (finfun (fun r => if r == a' then behead (use D r) else - if r == b' then (c, c1, c2) :: use D r else - use D r)) - (ins (rep D c1, rep D c2) (c, c1, c2) (lookup D)) - (pending D) - in join_use' newD a' b' p' + let: newD := + Data (rep D) (class D) + (finfun (fun r => if r == a' then behead (use D r) else + if r == b' then (c, c1, c2) :: use D r + else use D r)) + (ins (rep D c1, rep D c2) (c, c1, c2) (lookup D)) + (pending D) + in join_use' newD a' b' p' end. Definition join_use D a' b' := join_use' D a' b' (use D a'). @@ -538,7 +596,8 @@ Lemma join_use_metric (D : data) (a' b' : symb) : metric (join_use D a' b') = (metric D) + size (use D a'). Proof. rewrite /join_use; move E: (use D a')=>old_use. -elim: old_use D E=>[|[[c c1] c2] old_use IH] D E H1 H2 H3 /=; first by rewrite addn0. +elim: old_use D E=>[|[[c c1] c2] old_use IH] D E H1 H2 H3 /=. +- by rewrite addn0. move: (mem_split_uniq H2 (uniq_reps D))=>[s1][s2][L1 L2 L3]. have L4 : undup (map (rep D) (enum predT)) = reps D by []. case: (fnd _ _)=>[[[d d1] d2]|]; last first. @@ -626,27 +685,34 @@ Definition class_inv D := forall x a, (rep D x == a) = (x \in class D a). (* use lists only store equations with appropriate representatives *) Definition use_inv D := - forall a c c1 c2, a \in reps D -> (c, c1, c2) \in use D a -> rep D c1 = a \/ rep D c2 = a. + forall a c c1 c2, a \in reps D -> (c, c1, c2) \in use D a -> + rep D c1 = a \/ rep D c2 = a. -(* intermediary invariant which holds after the class of a' has been joined to b' *) +(* intermediary invariant which holds after *) +(* the class of a' has been joined to b' *) Definition use_inv1 D a' b' := forall c c1 c2, - (forall a, a \in reps D -> (c, c1, c2) \in use D a -> rep D c1 = a \/ rep D c2 = a) /\ + (forall a, a \in reps D -> (c, c1, c2) \in use D a -> + rep D c1 = a \/ rep D c2 = a) /\ ((c, c1, c2) \in use D a' -> rep D c1 = b' \/ rep D c2 = b'). -(* lookup table only stores equations with appropriate representatives *) +(* lookup table only stores equations with *) +(* appropriate representatives *) Definition lookup_inv D := forall a b c c1 c2, a \in reps D -> b \in reps D -> - fnd (a, b) (lookup D) = Some (c, c1, c2) -> rep D c1 = a /\ rep D c2 = b. + fnd (a, b) (lookup D) = Some (c, c1, c2) -> + rep D c1 = a /\ rep D c2 = b. (* two symbols are similar if they are either related by representatives *) (* or are to be related after the pending list is processed *) Definition similar D a b := - (const a, const b) \In closure (rep2rel D +p eqs2rel (map pend2eq (pending D))). + (const a, const b) \In closure (rep2rel D +p eqs2rel + (map pend2eq (pending D))). Definition similar1 D a' b' a b := - (const a, const b) \In closure (rep2rel D +p [Pred x y | x = const a' /\ y = const b'] +p - eqs2rel (map pend2eq (pending D))). + (const a, const b) \In + closure (rep2rel D +p [Pred x y | x = const a' /\ y = const b'] +p + eqs2rel (map pend2eq (pending D))). (* invariants tying use lists with the lookup table *) Definition use_lookup_inv D := @@ -657,11 +723,14 @@ Definition use_lookup_inv D := Definition lookup_use_inv D := forall a b d d1 d2, - a \in reps D -> b \in reps D -> fnd (a, b) (lookup D) = Some (d, d1, d2) -> + a \in reps D -> b \in reps D -> + fnd (a, b) (lookup D) = Some (d, d1, d2) -> [/\ exists c c1 c2, - (c, c1, c2) \in use D a /\ rep D c1 = a /\ rep D c2 = b /\ similar D c d & + (c, c1, c2) \in use D a /\ rep D c1 = a /\ + rep D c2 = b /\ similar D c d & exists c c1 c2, - (c, c1, c2) \in use D b /\ rep D c1 = a /\ rep D c2 = b /\ similar D c d]. + (c, c1, c2) \in use D b /\ rep D c1 = a /\ + rep D c2 = b /\ similar D c d]. (* intermediary invariants after removing an equation from the pending list *) Definition use_lookup_inv1 D a' b' := @@ -674,9 +743,11 @@ Definition lookup_use_inv1 D a' b' := forall a b d d1 d2, a \in reps D -> b \in reps D -> fnd (a, b) (lookup D) = Some (d, d1, d2) -> [/\ exists c c1 c2, - (c, c1, c2) \in use D a /\ rep D c1 = a /\ rep D c2 = b /\ similar1 D a' b' c d & + (c, c1, c2) \in use D a /\ rep D c1 = a /\ + rep D c2 = b /\ similar1 D a' b' c d & exists c c1 c2, - (c, c1, c2) \in use D b /\ rep D c1 = a /\ rep D c2 = b /\ similar1 D a' b' c d]. + (c, c1, c2) \in use D b /\ rep D c1 = a /\ + rep D c2 = b /\ similar1 D a' b' c d]. (* intermediary invariants after join_class and during join_use *) Definition use_lookup_inv2 D a' b' := @@ -719,24 +790,31 @@ Definition use_lookup_inv2 D a' b' := Definition lookup_use_inv2 D a' b' := forall d d1 d2, - [/\ forall b, b \in reps D -> fnd (a', b) (lookup D) = Some (d, d1, d2) -> + [/\ forall b, b \in reps D -> + fnd (a', b) (lookup D) = Some (d, d1, d2) -> rep D d1 = b' /\ exists c c1 c2, - (c, c1, c2) \in use D b /\ rep D c1 = b' /\ rep D c2 = b /\ similar D c d, - forall a, a \in reps D -> fnd (a, a') (lookup D) = Some (d, d1, d2) -> + (c, c1, c2) \in use D b /\ rep D c1 = b' /\ + rep D c2 = b /\ similar D c d, + forall a, a \in reps D -> + fnd (a, a') (lookup D) = Some (d, d1, d2) -> rep D d2 = b' /\ exists c c1 c2, - (c, c1, c2) \in use D a /\ rep D c1 = a /\ rep D c2 = b' /\ similar D c d & + (c, c1, c2) \in use D a /\ rep D c1 = a /\ + rep D c2 = b' /\ similar D c d & forall a b, a \in reps D -> b \in reps D -> fnd (a, b) (lookup D) = Some (d, d1, d2) -> [/\ exists c c1 c2, - (c, c1, c2) \in use D a /\ rep D c1 = a /\ rep D c2 = b /\ similar D c d & + (c, c1, c2) \in use D a /\ rep D c1 = a /\ + rep D c2 = b /\ similar D c d & exists c c1 c2, - (c, c1, c2) \in use D b /\ rep D c1 = a /\ rep D c2 = b /\ similar D c d]]. + (c, c1, c2) \in use D b /\ rep D c1 = a /\ + rep D c2 = b /\ similar D c d]]. (* the invariant of the propagate routine *) Definition propagate_inv D := - rep_idemp D /\ use_inv D /\ lookup_inv D /\ use_lookup_inv D /\ lookup_use_inv D. + rep_idemp D /\ use_inv D /\ lookup_inv D /\ + use_lookup_inv D /\ lookup_use_inv D. (****************) (* Verification *) @@ -744,7 +822,8 @@ Definition propagate_inv D := (* first some basic rewrite rules *) -Lemma reps_rep (D : data) (a : symb) : rep_idemp D -> a \in reps D -> rep D a = a. +Lemma reps_rep (D : data) (a : symb) : + rep_idemp D -> a \in reps D -> rep D a = a. Proof. by move=>H; rewrite mem_undup; case/mapP=>x _ ->; apply: H. Qed. Lemma symR R x y : (x, y) \In closure R <-> (y, x) \In closure R. @@ -752,7 +831,8 @@ Proof. by split=>T; apply: symC. Qed. Lemma app_rep D R a b x : (x, app (const a) (const b)) \In closure (rep2rel D +p R) <-> - (x, app (const (rep D a)) (const (rep D b))) \In closure (rep2rel D +p R). + (x, app (const (rep D a)) (const (rep D b))) + \In closure (rep2rel D +p R). Proof. split=>T. - apply: (transC (y:= app (const a) (const (rep D b)))); last first. @@ -789,11 +869,11 @@ Proof. by rewrite /CRel !clos_clos !orrA; apply: const_rep. Qed. Definition init : data := Data [ffun x => x] [ffun c => [:: c]] [ffun c => [::]] (nil _ _) [::]. -Lemma initP : propagate_inv init /\ CRel init <~> closure (graph const). +Lemma initP : propagate_inv init /\ CRel init <~> closure (graph (@const s)). Proof. have S: lookup_rel init <~> Pred0. - by split=>//; case: x=> ??[?][?][?][?][?][]. -have E: graph const <~> graph (fun x => const ([ffun x => x] x)). +have E: graph (@const s) <~> graph (fun x => const ([ffun x => x] x)). - by split; case: x =>x1 x2 /=; case/imageP=>[x] _ [->->]; apply/imageP; exists x; rewrite ?ffunE. rewrite /CRel /= S 2!orr0 /=; split; last by rewrite E. @@ -825,9 +905,11 @@ case; last first. by rewrite /= !ffunE /= reps_rep // eq_refl. case/imageP=>z _ /= [->->] {x y}. case E: (rep D z == a'); last first. -- by apply: (@closI _ (const z, const (rep D z))); apply/imageP; exists z; rewrite // ffunE /= E. +- apply: (@closI _ (const z, const (rep D z))); apply/imageP. + by exists z; rewrite // ffunE /= E. apply: (transC (y:=const b')). -- by apply: (@closI _ (const z, const b')); apply/imageP; exists z; rewrite // ffunE /= E. +- apply: (@closI _ (const z, const b')); apply/imageP. + by exists z; rewrite // ffunE /= E. apply: symC; apply: (@closI _ (const (rep D z), const b')); apply/imageP. by exists (rep D z); rewrite // ffunE /= H1 // E. Qed. @@ -862,10 +944,10 @@ move=>H1 H2 H a b. by rewrite /similar /similar1 -orrA -clos_clos -join_class_repE // clos_clos. Qed. -Module Dummy321. End Dummy321. - Lemma join_class_classP (D : data) (a' b' : symb) : - a' != b' -> class_inv D -> class_inv (join_class D a' b'). + a' != b' -> + class_inv D -> + class_inv (join_class D a' b'). Proof. rewrite /join_class /class_inv=>E /= => H x a; rewrite !ffunE /=. case: ifP=>H1. @@ -878,36 +960,40 @@ by rewrite H. Qed. Lemma join_classP (D : data) (a' b' : symb) : - a' \in reps D -> b' \in reps D -> a' != b' -> - rep_idemp D -> use_inv D -> lookup_inv D -> - use_lookup_inv1 D a' b' -> lookup_use_inv1 D a' b' -> + a' \in reps D -> + b' \in reps D -> + a' != b' -> + rep_idemp D -> + use_inv D -> + lookup_inv D -> + use_lookup_inv1 D a' b' -> + lookup_use_inv1 D a' b' -> use_lookup_inv2 (join_class D a' b') a' b' /\ lookup_use_inv2 (join_class D a' b') a' b'. Proof. move=>H1 H2 H3 L0 L1 L2 T1 T2. split; last first. -- split=>[b|a|a b]; rewrite !join_class_eq // !mem_filter /=; case/andP=>H4 H5. - - move=>F. - rewrite ffunE /=. +- split=>[b|a|a b]; rewrite !join_class_eq // !mem_filter /=; + case/andP=>H4 H5. + - move=>F; rewrite ffunE /=. case: (L2 a' b d d1 d2 H1 H5 F)=>R1 R2; rewrite R1 eq_refl; do !split=>//. case: (T2 a' b d d1 d2 H1 H5 F)=>_ [c][c1][c2][Q1][Q2][Q3]Q4. - exists c, c1, c2; rewrite !ffunE /= Q2 Q3 eq_refl (negbTE H4); do !split=>//. - by rewrite -join_class_simE. - - move=>F. - rewrite ffunE /=. + exists c, c1, c2; rewrite !ffunE /= Q2 Q3 eq_refl (negbTE H4). + by do !split=>//; rewrite -join_class_simE. + - move=>F; rewrite ffunE /=. case: (L2 a a' d d1 d2 H5 H1 F)=>R1 R2; rewrite R2 eq_refl; do !split=>//. case: (T2 a a' d d1 d2 H5 H1 F)=>[[c][c1][c2][Q1][Q2][Q3]Q4 _]. - exists c, c1, c2; rewrite !ffunE /= Q2 Q3 eq_refl (negbTE H4); do !split=>//. - by rewrite -join_class_simE. + exists c, c1, c2; rewrite !ffunE /= Q2 Q3 eq_refl (negbTE H4). + by do !split=>//; rewrite -join_class_simE. case/andP=>H6 H7 F. case: (T2 a b d d1 d2 H5 H7 F). move=>[c][c1][c2][Q1][Q2][Q3] Q4. move=>[e][e1][e2][F1][F2][F3] F4. split. - - exists c, c1, c2; rewrite !ffunE /= Q2 Q3 (negbTE H4) (negbTE H6); do !split=>//. - by rewrite -join_class_simE. - exists e, e1, e2; rewrite !ffunE /= F2 F3 (negbTE H4) (negbTE H6); do !split=>//. - by rewrite -join_class_simE. + - exists c, c1, c2; rewrite !ffunE /= Q2 Q3 (negbTE H4) (negbTE H6). + by do !split=>//; rewrite -join_class_simE. + exists e, e1, e2; rewrite !ffunE /= F2 F3 (negbTE H4) (negbTE H6). + by do !split=>//; rewrite -join_class_simE. split=>[c c1 c2 /= H4 | a c c1 c2]; last first. - rewrite join_class_eq // mem_filter /=. case/andP=>H4 H5 H6. @@ -941,15 +1027,17 @@ rewrite -join_class_simE=>//. by rewrite (eqP H6). Qed. -Module Dummy123. End Dummy123. - -Definition pull {T : Type} (r : Pred T) := (orrC r, orrCA r). - +Definition rpull {T : Type} (r : Pred T) := (orrC r, orrCA r). Lemma join_classE (D : data) (a' b' : symb) : - a' \in reps D -> b' \in reps D -> a' != b' -> - rep_idemp D -> use_lookup_inv1 D a' b' -> lookup_use_inv1 D a' b' -> - closure (CRel (join_class D a' b') +p eqs2rel (map symb2eq (use D a'))) <~> + a' \in reps D -> + b' \in reps D -> + a' != b' -> + rep_idemp D -> + use_lookup_inv1 D a' b' -> + lookup_use_inv1 D a' b' -> + closure (CRel (join_class D a' b') +p + eqs2rel (map symb2eq (use D a'))) <~> closure (CRel D +p [Pred x y | x = const a' /\ y = const b']). Proof. pose ty := [Pred x0 y0 | x0 = const a' /\ y0 = const b']. @@ -962,18 +1050,19 @@ move=>L1 L2 L3 H1 H4 H5 [x y]; split. apply/app_rel. apply: (transC (y := const d)). - rewrite /CRel clos_clos !orrA. - rewrite -!(pull (eqs2rel _)) -3!(pull ty) -!(pull (rep2rel _)) -!orrA. + rewrite -!(rpull (eqs2rel _)) -3!(rpull ty) -!(rpull (rep2rel _)) -!orrA. move: Q4; apply: sub_closI=>{x y} [[x y]] Q4; apply: sub_orl. by rewrite !toPredE in Q4 *; rewrite orrA. rewrite const_rel /CRel clos_clos !orrA symR. - apply: (@closI _ (app (const (rep D c1)) (const (rep D c2)), const (rep D d))). + apply: (@closI _ (app (const (rep D c1)) + (const (rep D c2)), const (rep D d))). apply: sub_orr; apply: sub_orl. rewrite toPredE invert_look. exists (rep D d1), (rep D d2), d, d1, d2. by rewrite -Q2 -Q3 !rep_in_reps. rewrite /CRel !toPredE clos_idemp clos_clos !orrA /=. rewrite -clos_clos join_class_repE // clos_clos orrA. - rewrite -!(pull (eqs2rel (map pend2eq _))). + rewrite -!(rpull (eqs2rel (map pend2eq _))). apply: sub_closKR=>{x y}; apply: sub_closKR=>[[x y]] T1. rewrite toPredE -clos_idemp. apply: sub_closI T1=>{x y} [[x y]]. @@ -988,8 +1077,8 @@ move=>L1 L2 L3 H1 H4 H5 [x y]; split. exists a, b, c, c1, c2=>//. by rewrite H6. rewrite /CRel !toPredE !clos_clos !orrA /= => {H4} T. -rewrite -clos_clos join_class_repE // clos_clos !orrA -(pull ty). -rewrite -3!(pull ty) -!(pull (rep2rel _)) -!(pull (lookup_rel _)) in T *. +rewrite -clos_clos join_class_repE // clos_clos !orrA -(rpull ty). +rewrite -3!(rpull ty) -!(rpull (rep2rel _)) -!(rpull (lookup_rel _)) in T *. rewrite -clos_idemp; apply: sub_closI T=>{x y} [[x y]]. case=>[|T]; last by apply closI; apply: sub_orr; rewrite toPredE -!orrA; apply: sub_orl; rewrite toPredE !orrA. @@ -997,14 +1086,17 @@ case=>a[b][c][c1][c2][-> T1 T2 T3 ->]. case: (H5 a b c c1 c2 T1 T2 T3). move=>[d][d1][d2][Q1][Q2][Q3] Q4 [f][f1][f2][F1][F2][F3] F4. case E1: (a == a'). -- rewrite toPredE !(pull (lookup_rel _)) !orrA -Q2 -Q3 symR -app_rep -const_rep symR. +- rewrite toPredE !(rpull (lookup_rel _)) !orrA -Q2 -Q3 symR + -app_rep -const_rep symR. apply: (transC (y:= const d)); last first. - move: Q4; apply: sub_closI=>{x y} [[x y]]. by rewrite !toPredE -!orrA=>Q4; do 2!apply: sub_orl. - apply: symC; apply closI; do 3!apply: sub_orr; apply: sub_orl; rewrite toPredE invert -(eqP E1). + apply: symC; apply closI; do 3!apply: sub_orr; apply: sub_orl. + rewrite toPredE invert -(eqP E1). by exists d, d1, d2. case E2: (b == a'). -- rewrite toPredE !(pull (lookup_rel _)) !orrA -F2 -F3 symR -app_rep -const_rep symR. +- rewrite toPredE !(rpull (lookup_rel _)) !orrA. + rewrite -F2 -F3 symR -app_rep -const_rep symR. apply: (transC (y:= const f)); last first. - move: F4; apply: sub_closI=>{x y} [[x y]]. by rewrite !toPredE -!orrA=>F4; do 2!apply: sub_orl. @@ -1023,15 +1115,15 @@ exists a, b, c, c1, c2. by rewrite !join_class_eq // !mem_filter /= E1 E2 !ffunE /= E3. Qed. -Module Dummy1. End Dummy1. - (* Lemmas about join_use *) Lemma join_use_classP (D : data) (a' b' : symb) : - a' != b' -> class_inv D -> class_inv (join_use D a' b'). + a' != b' -> + class_inv D -> + class_inv (join_use D a' b'). Proof. rewrite /class_inv /join_use; move E: (use _ _)=>x. -elim: x D E=>[|[[c c1] c2] s IH] D E H1 H2 x a //=. +elim: x D E=>[|[[c c1] c2] ss IH] D E H1 H2 x a //=. by case F : (fnd _ _)=>[[[d d1] d2]|]; rewrite IH //= ffunE eq_refl E. Qed. @@ -1045,18 +1137,21 @@ by case F: (fnd _ _)=>[[[d d1] d2]|] /=; rewrite IH //= ffunE eq_refl H /=; eaut Qed. Lemma join_use_useE (D : data) (a' b' : symb) (l1 l2 : seq (symb*symb*symb)) : - use D a' = l1 ++ l2 -> use (join_use' D a' b' l1) a' = l2. + use D a' = l1 ++ l2 -> + use (join_use' D a' b' l1) a' = l2. Proof. elim: l1 D a' b'=>[|[[c1 c2] c] l1 IH] D a' b' //= H. by case F: (fnd _ _)=>[[[d d1] d2]|] //=; apply: IH; rewrite /= ffunE eq_refl H. Qed. Lemma join_use_useP (D : data) (a' b' : symb) : - a' \notin reps D -> b' \in reps D -> - use_inv1 D a' b' -> use_inv1 (join_use D a' b') a' b'. + a' \notin reps D -> + b' \in reps D -> + use_inv1 D a' b' -> + use_inv1 (join_use D a' b') a' b'. Proof. rewrite /join_use; move E: (use _ _)=>x. -elim: x D E=>[|[[c c1] c2] s IH] D E H1 H2 H3 //=. +elim: x D E=>[|[[c c1] c2] ss IH] D E H1 H2 H3 //=. case F: (fnd _ _)=>[[[d d1] d2]|]; (apply: IH=>//=; first by [rewrite ffunE eq_refl E]; move=>e e1 e2; rewrite /reps /=; split=>[a|]; rewrite ffunE; @@ -1072,11 +1167,13 @@ by rewrite E inE eq_refl. Qed. Lemma join_use_lookupP (D : data) (a' b' : symb) : - a' \notin reps D -> b' \in reps D -> - lookup_inv D -> lookup_inv (join_use D a' b'). + a' \notin reps D -> + b' \in reps D -> + lookup_inv D -> + lookup_inv (join_use D a' b'). Proof. rewrite /join_use; move E: (use _ _)=>x. -elim: x D E=>[|[[c c1] c2] s IH] D E H1 H2 H3 //=. +elim: x D E=>[|[[c c1] c2] ss IH] D E H1 H2 H3 //=. case F: (fnd _ _)=>[[[d d1] d2]|]. - by apply: IH=>//=; first by rewrite ffunE eq_refl E. apply: IH=>//=; first by rewrite ffunE eq_refl E. @@ -1085,9 +1182,13 @@ by case: eqP=>[[-> -> _ _] [_ -> ->]|_]; last by apply: H3. Qed. Lemma join_usePE (D : data) (a' b' : symb) : - a' \notin reps D -> b' \in reps D -> - rep_idemp D -> use_inv1 D a' b' -> lookup_inv D -> - use_lookup_inv2 D a' b' -> lookup_use_inv2 D a' b' -> + a' \notin reps D -> + b' \in reps D -> + rep_idemp D -> + use_inv1 D a' b' -> + lookup_inv D -> + use_lookup_inv2 D a' b' -> + lookup_use_inv2 D a' b' -> rep_idemp (join_use D a' b') /\ use_inv1 (join_use D a' b') a' b' /\ lookup_inv (join_use D a' b') /\ @@ -1097,7 +1198,7 @@ Lemma join_usePE (D : data) (a' b' : symb) : CRel (join_use D a' b'). Proof. rewrite /join_use; move E: (use _ _)=>x. -elim: x D E=>[|[[c c1] c2] s IH] D E L1 L2 H1 H2 H3 H4 H5 /=. +elim: x D E=>[|[[c c1] c2] ss IH] D E L1 L2 H1 H2 H3 H4 H5 /=. - by rewrite orr0 /CRel clos_idemp. case F: (fnd _ _)=>[[[d d1] d2]|]; set D' := Data _ _ _ _ _. - have S1: closure (rep2rel D' +p eqs2rel (map pend2eq (pending D'))) <~> @@ -1105,19 +1206,23 @@ case F: (fnd _ _)=>[[[d d1] d2]|]; set D' := Data _ _ _ _ _. eqs2rel (map pend2eq (pending D))) by []. have S2: forall e1 e2, similar D e1 e2 -> similar D' e1 e2. - rewrite /similar=>e1 e2; move: (const e1) (const e2)=>{e1 e2} x y. - by rewrite S1; apply: sub_closI=>{x y} [[x y]]; case; [left | right; right]. - have T2 : closure (CRel D' +p eqs2rel (map symb2eq s)) <~> - closure (CRel D +p [Pred x y | x = const c /\ y = app (const c1) (const c2)] +p - eqs2rel (map symb2eq s)). + by rewrite S1; apply: sub_closI=>{x y} [[x y]]; + case; [left | right; right]. + have T2 : closure (CRel D' +p eqs2rel (map symb2eq ss)) <~> + closure (CRel D +p [Pred x y | x = const c /\ + y = app (const c1) (const c2)] +p eqs2rel (map symb2eq ss)). - rewrite /CRel {2 3}/D' /= !clos_clos !orrA. - rewrite -!(pull (eqs2rel (map pend2eq _))) -!(pull (eqs2rel (map symb2eq s))). + rewrite -!(rpull (eqs2rel (map pend2eq _))). + rewrite -!(rpull (eqs2rel (map symb2eq ss))). do 2!apply: closKR. rewrite -!orrA=>[[x y]]; split=>T; rewrite toPredE -clos_idemp; apply: sub_closI T=>{x y} [[x y]]. - case; first by move=>H7; apply: closI; apply: sub_orl. case=>->->; rewrite toPredE !orrA symR const_rep symR. - apply: (transC (y:= app (const (rep D c1)) (const (rep D c2)))); last first. - - apply closI; apply: sub_orr; apply: sub_orl; rewrite toPredE invert_look. + apply: (transC (y:= app (const (rep D c1)) (const (rep D c2)))); + last first. + - apply closI; apply: sub_orr; apply: sub_orl. + rewrite toPredE invert_look. by exists (rep D c1), (rep D c2), d, d1, d2. by rewrite -app_rep; apply closI; do 2!right. case; first by move=>H7; apply closI; apply: sub_orl. @@ -1143,8 +1248,10 @@ case F: (fnd _ _)=>[[[d d1] d2]|]; set D' := Data _ _ _ _ _. by exists f, f1, f2; rewrite Q2 Q3; do ![split=>//]; apply: S2. case: eqP L3 L1=>[->-> //|_] L3 L1 H6. case: (H4.2 a e e1 e2 L3 H6)=>[[R1][R2]|[R]|[R1][R2]|[R]]; - [move=>[[h][h1][h2][F1][F2][F3] F4]| idtac | move=>[[h][h1][h2][F1][F2][F3] F4]| idtac]; - move=>[f][f1][f2][Q1][Q2][Q3] Q4; [idtac | apply: Or42 | idtac | apply: Or44]; + [move=>[[h][h1][h2][F1][F2][F3] F4]| idtac | + move=>[[h][h1][h2][F1][F2][F3] F4]| idtac]; + move=>[f][f1][f2][Q1][Q2][Q3] Q4; + [idtac | apply: Or42 | idtac | apply: Or44]; try by [do !split=>//; exists f, f1, f2; do !split=>//; apply: S2]; rewrite E inE in F1; case/orP: F1=>F1; [apply: Or42 | apply: Or41 | apply: Or44 | apply: Or43]; @@ -1174,8 +1281,8 @@ case F: (fnd _ _)=>[[[d d1] d2]|]; set D' := Data _ _ _ _ _. exists h, h1, h2; rewrite {1}/D' /= !ffunE. by case: eqP H7 L1=>[->-> //| _] H7 L1; do !split=>//; apply: S2. have T1: lookup_rel D' <~> lookup_rel D +p - [Pred x y | x = app (const (rep D c1)) (const (rep D c2)) /\ - y = const (rep D c)]. + [Pred x y | x = app (const (rep D c1)) (const (rep D c2)) /\ + y = const (rep D c)]. - rewrite /lookup_rel /D' /= /reps /= =>[[x y]]; split. - move=>[a][b][d][d1][d2][-> T2 T3]. rewrite fnd_ins. @@ -1191,12 +1298,14 @@ have T1: lookup_rel D' <~> lookup_rel D +p exists (rep D c1), (rep D c2), c, c1, c2. rewrite fnd_ins !rep_in_reps. by case: eqP. -have T2: closure (CRel D' +p eqs2rel (map symb2eq s)) <~> - closure (CRel D +p [Pred x y | x = const c /\ y = app (const c1) (const c2)] +p - eqs2rel (map symb2eq s)). -- pose ty1 := [Pred x y | x = app (const (rep D c1)) (const (rep D c2)) /\ y = const (rep D c)]. +have T2: closure (CRel D' +p eqs2rel (map symb2eq ss)) <~> + closure (CRel D +p [Pred x y | x = const c /\ + y = app (const c1) (const c2)] +p eqs2rel (map symb2eq ss)). +- pose ty1 := [Pred x y | x = app (const (rep D c1)) + (const (rep D c2)) /\ y = const (rep D c)]. pose ty2 := [Pred x y | x = const c /\ y = app (const c1) (const c2)]. - rewrite /CRel {3}/D' /= T1 !clos_clos !orrA -2!(pull ty1) -3!(pull ty2) -!orrA -!(pull (rep2rel _)). + rewrite /CRel {3}/D' /= T1 !clos_clos !orrA -2!(rpull ty1) + -3!(rpull ty2) -!orrA -!(rpull (rep2rel _)). do 3!apply: closRK; move=>[x y]; split=>T; rewrite toPredE -clos_idemp; apply: sub_closI T=>{x y} [[x y]]. - case; first by move=>T; apply: closI; left. @@ -1229,12 +1338,14 @@ apply: IH=>//=; first by rewrite ffunE eq_refl E. case: eqP=>[->|H7 H8]; [rewrite inE; case/orP=>[|H8] | idtac]. - case/eqP=>->->->. have H8: (c, c1, c2) \in use D a' by [rewrite E inE eq_refl]. - case: ((H2 c c1 c2).2 H8)=>H9; [apply: Or42 | apply: Or44]; do ![split=>//]; - exists c, c1, c2; rewrite fnd_ins -H9 eq_refl; + case: ((H2 c c1 c2).2 H8)=>H9; [apply: Or42 | apply: Or44]; + do ![split=>//]; exists c, c1, c2; rewrite fnd_ins -H9 eq_refl; by do !split=>//; apply: reflC. - case: (H4.2 b' d d1 d2 L2 H8)=>[[R1][R2]|[R]|[R1][R2]|[R]]; - [move=>[[f][f1][f2][F1][F2][F3] F4]|idtac|move=>[[f][f1][f2][F1][F2][F3] F4]| idtac]; - move=>[e][e1][e2][Q1][Q2][Q3] Q4; [idtac | apply: Or42 | idtac | apply: Or44]; + [move=>[[f][f1][f2][F1][F2][F3] F4]|idtac| + move=>[[f][f1][f2][F1][F2][F3] F4]| idtac]; + move=>[e][e1][e2][Q1][Q2][Q3] Q4; + [idtac | apply: Or42 | idtac | apply: Or44]; try by [do !split=>//; exists e, e1, e2; rewrite fnd_ins; case: eqP Q1=>[[->->]|]; first by rewrite F]; rewrite E inE in F1; case/orP: F1=>F1; @@ -1247,8 +1358,10 @@ apply: IH=>//=; first by rewrite ffunE eq_refl E. case/eqP: F1 F4=><-<-<- F4; by rewrite F3 F2 ?R1 ?R2 eq_refl. case: (H4.2 a d d1 d2 H6 H8)=>[[R1][R2]|[R]|[R1][R2]|[R]]; - [move=>[[f][f1][f2][F1][F2][F3] F4]|idtac|move=>[[f][f1][f2][F1][F2][F3] F4]| idtac]; - move=>[e][e1][e2][Q1][Q2][Q3] Q4; [idtac | apply: Or42 | idtac | apply: Or44]; + [move=>[[f][f1][f2][F1][F2][F3] F4]|idtac| + move=>[[f][f1][f2][F1][F2][F3] F4]| idtac]; + move=>[e][e1][e2][Q1][Q2][Q3] Q4; + [idtac | apply: Or42 | idtac | apply: Or44]; try by [do !split=>//; exists e, e1, e2; rewrite fnd_ins; case: eqP Q1=>[[->->]|]; first by rewrite F]; rewrite E inE in F1; case/orP: F1=>F1; @@ -1323,38 +1436,40 @@ exists f, f1, f2; do !split=>//. by apply: (transC (y:= const e))=>//; apply: symC. Qed. -Module Dummy3. End Dummy3. - (* Lemmas about propagate *) Lemma propagatePE (D : data) : - propagate_inv D -> propagate_inv (propagate D) /\ - pending (propagate D) = [::] /\ - CRel D <~> CRel (propagate D). + propagate_inv D -> + propagate_inv (propagate D) /\ + pending (propagate D) = [::] /\ + CRel D <~> CRel (propagate D). Proof. move: D. -pose ty x0 y0 := [Pred x y | x = const x0 /\ y = const y0]. +pose ty x0 y0 := [Pred x y | x = @const s x0 /\ y = @const s y0]. have L: forall D a b, - closure (rep2rel D +p ty a b) <~> - closure (rep2rel D +p ty (rep D a) (rep D b)). -- move=>D a b [x y]; split=>T; rewrite toPredE -clos_idemp; apply: sub_closI T=>{x y} [[x y]]; + closure (rep2rel D +p ty a b) <~> + closure (rep2rel D +p ty (rep D a) (rep D b)). +- move=>D a b [x y]; split=>T; rewrite toPredE -clos_idemp; + apply: sub_closI T=>{x y} [[x y]]; (case; first by move=>H; apply: closI; left); case=>->->; - [rewrite toPredE const_rep symR const_rep symR | rewrite toPredE -const_rep symR -const_rep symR]; + [rewrite toPredE const_rep symR const_rep symR | + rewrite toPredE -const_rep symR -const_rep symR]; by apply closI; right. pose inv D := propagate_inv D. have L1: forall D pend_eq p' a b a' b' D', - pending D = pend_eq :: p' -> pend2eq pend_eq = simp_eq a b -> - rep D a = a' -> rep D b = b' -> Data (rep D) (class D) (use D) (lookup D) p' = D' -> - (a' == b') -> - (inv D' -> inv (propagate D') /\ pending (propagate D') = [::] /\ - CRel D' <~> CRel (propagate D')) -> - inv D -> inv (propagate D') /\ pending (propagate D') = [::] /\ - CRel D <~> CRel (propagate D'). + pending D = pend_eq :: p' -> pend2eq pend_eq = simp_eq a b -> + rep D a = a' -> rep D b = b' -> + Data (rep D) (class D) (use D) (lookup D) p' = D' -> + (a' == b') -> + (inv D' -> inv (propagate D') /\ pending (propagate D') = [::] /\ + CRel D' <~> CRel (propagate D')) -> + inv D -> inv (propagate D') /\ pending (propagate D') = [::] /\ + CRel D <~> CRel (propagate D'). - move=>D pend_eq p' a b a' b' D' H1 H H2 H3 H4 H5 IH [H6][H7][H8][H9] H10. have L2: forall a1 b1, similar D a1 b1 -> similar D' a1 b1. - - move=>a1 b1; rewrite /similar -{2}H4 H1 /= H /= -(pull (ty a b)) => T. - rewrite -clos_idemp; move: T; apply: sub_closI=>{a1 b1} [[x y]]; rewrite -H4. - case=>[[->->]|T]; last by apply: closI. + - move=>a1 b1; rewrite /similar -{2}H4 H1 /= H /= -(rpull (ty a b)) => T. + rewrite -clos_idemp; move: T; apply: sub_closI=>{a1 b1} [[x y]]; + rewrite -H4; case=>[[->->]|T]; last by apply: closI. rewrite toPredE const_rep symR const_rep symR /=. by rewrite H2 H3 (eqP H5); apply: reflC. suff T: CRel D <~> CRel D'. @@ -1368,29 +1483,32 @@ have L1: forall D pend_eq p' a b a' b' D', split; [exists d, d1, d2 | exists e, e1, e2]; by do !split=>//; rewrite H4; apply: L2. rewrite /CRel H1 /= -{-1}H4 /=. - rewrite -!(pull (lookup_rel _)) -!(pull (eqs2rel (map pend2eq _))). + rewrite -!(rpull (lookup_rel _)) -!(rpull (eqs2rel (map pend2eq _))). do 2![apply: closKR]; rewrite H /= L -H4. symmetry; rewrite closAK=>[[x y]]; case=>->->. by rewrite H2 H3 (eqP H5); apply: reflC. have L2: forall D pend_eq p' a b a' b' D' D'', - pending D = pend_eq :: p' -> pend2eq pend_eq = simp_eq a b -> - rep D a = a' -> rep D b = b' -> - Data (rep D) (class D) (use D) (lookup D) p' = D' -> a' != b' -> - join_class D' a' b' = D'' -> - (inv (join_use D'' a' b') -> inv (propagate (join_use D'' a' b')) /\ - pending (propagate (join_use D'' a' b')) = [::] /\ - CRel (join_use D'' a' b') <~> CRel (propagate (join_use D'' a' b'))) -> - inv D -> inv (propagate (join_use D'' a' b')) /\ - pending (propagate (join_use D'' a' b')) = [::] /\ - CRel D <~> CRel (propagate (join_use D'' a' b')). -- move=>D pend_eq p' a b a' b' D' D'' H1 H H2 H3 H4 H5 H6 IH [H7][H8][H9][H10] H11. - have T: CRel D <~> closure (CRel D' +p [Pred x y | x = const a' /\ y = const b']). + pending D = pend_eq :: p' -> pend2eq pend_eq = simp_eq a b -> + rep D a = a' -> rep D b = b' -> + Data (rep D) (class D) (use D) (lookup D) p' = D' -> a' != b' -> + join_class D' a' b' = D'' -> + (inv (join_use D'' a' b') -> inv (propagate (join_use D'' a' b')) /\ + pending (propagate (join_use D'' a' b')) = [::] /\ + CRel (join_use D'' a' b') <~> CRel (propagate (join_use D'' a' b'))) -> + inv D -> inv (propagate (join_use D'' a' b')) /\ + pending (propagate (join_use D'' a' b')) = [::] /\ + CRel D <~> CRel (propagate (join_use D'' a' b')). +- move=>D pend_eq p' a b a' b' D' D'' H1 H H2 H3 H4 H5 H6 IH + [H7][H8][H9][H10] H11. + have T: CRel D <~> closure (CRel D' +p + [Pred x y | x = const a' /\ y = const b']). - rewrite /CRel H1 -{2 3}H4 /= clos_clos !orrA H /=. - rewrite -!(pull (lookup_rel _)) -!(pull (eqs2rel (map pend2eq _))). + rewrite -!(rpull (lookup_rel _)) -!(rpull (eqs2rel (map pend2eq _))). by do 2![apply: closKR]; rewrite L H2 H3 -H4 /= /rep2rel. have S2: forall a1 b1, similar D a1 b1 -> similar1 D' a' b' a1 b1. - move=>a1 b1. - rewrite /similar /similar1 H1 -H4 /= H /= (pull (ty a b)) (pull (ty a' b')) -!orrA => Q4. + rewrite /similar /similar1 H1 -H4 /= H /=. + rewrite (rpull (ty a b)) (rpull (ty a' b')) -!orrA => Q4. rewrite -clos_idemp. move: Q4; apply: sub_closI=>[[x y]]. case; first by move=>Q4; apply: closI; left. @@ -1438,27 +1556,29 @@ have L2: forall D pend_eq p' a b a' b' D' D'', have L13'' : use_lookup_inv (join_use D'' a' b'). - have U: forall D, use (join_use D a' b') a' = [::]. - rewrite /join_use=>D1; move E: (use D1 _)=>x. - elim: x D1 E=>[|[[f f1] f2] s IH1] D1 E //=. - by case F: (fnd _ _)=>[[[d d1] d2]|]; apply: IH1; rewrite /= !ffunE eq_refl E. + elim: x D1 E=>[|[[f f1] f2] ss IH1] D1 E //=. + case F: (fnd _ _)=>[[[d d1] d2]|]; apply: IH1; + by rewrite /= !ffunE eq_refl E. move=>a1 c c1 c2 H12 H13. case: (T9.2 a1 c c1 c2 H12 H13); first 1 last; try by [case=><- [d][d1][d2][Q1][Q2][Q3] Q4; exists d, d1, d2]; by case=>[_][_][[d][d1][d2]][]; rewrite U. rewrite -H6 /= join_classE // T H6=>->. by apply: IH. -apply/(@propagate_ind (fun d d' => inv d -> inv d' /\ pending d' = [::] /\ CRel d <~> CRel d')); first by []. +apply/(@propagate_ind (fun d d' => inv d -> inv d' /\ + pending d' = [::] /\ CRel d <~> CRel d')); first by []. - by move=>D e p' H1 a' H2 b' H3 D'; apply: L1 H1 _ H2 H3; case: e=>// [[[c c1]] c2] [[d d1] d2]. -move=>D e p' H1 a' H2 b' H3 D' E [] // H _ D''; apply: L2 H1 _ H2 H3 E (negbT H). +move=>D e p' H1 a' H2 b' H3 D' E [] // H _ D''; +apply: L2 H1 _ H2 H3 E (negbT H). by case: e=>// [[[c c1]] c2] [[d d1] d2]. Qed. -Module Dummy4. End Dummy4. - (* Lemmas about interaction of propagate with pending and closure *) Lemma propagate_pendP d eq : propagate_inv d -> - propagate_inv (Data (rep d) (class d) (use d) (lookup d) (eq :: pending d)). + propagate_inv (Data (rep d) (class d) (use d) + (lookup d) (eq :: pending d)). Proof. move=>H; set d' := Data _ _ _ _ _. have L: forall a b, similar d a b -> similar d' a b. @@ -1484,7 +1604,7 @@ move=>PI H. have [R1 R2]: rep d e1 = rep d c1 /\ rep d e2 = rep d c2. - by move: PI=>[_][_][L1] _; apply: (L1 _ _ e). rewrite /CRel /= /eq2rel /=. -rewrite clos_clos -!(pull (eqs2rel _)) !orrA; apply: closKR. +rewrite clos_clos -!(rpull (eqs2rel _)) !orrA; apply: closKR. move=>[z y]; split=>O; rewrite toPredE -clos_idemp; move: O; apply: sub_closI; move=>{z y} [z y]. - case=>[O|]; first by apply: closI; left. @@ -1502,8 +1622,6 @@ exists (rep d c1), (rep d c2), e, e1, e2. by rewrite !rep_in_reps. Qed. -Module DummyT. End DummyT. - Section NoPend. Variables (d : data) (c c1 c2 : symb). Hypotheses (PI : propagate_inv d) @@ -1515,10 +1633,12 @@ Notation u2' := [ffun z => if z == rep d c2 then (c, c1, c2) :: u1' z Lemma propagate_nopendP : propagate_inv (Data (rep d) (class d) u2' - (ins (rep d c1, rep d c2) (c, c1, c2) (lookup d)) (pending d)). + (ins (rep d c1, rep d c2) (c, c1, c2) + (lookup d)) (pending d)). Proof. have E : forall (a c c1 c2 : symb) l, - (c, c1, c2) :: (if a == rep d c1 then (c, c1, c2) :: l else l) =i (c, c1, c2) :: l. + (c, c1, c2) :: (if a == rep d c1 then (c, c1, c2) :: l else l) =i + (c, c1, c2) :: l. - move=>a f f1 f2 l z; rewrite !inE; case: ifP=>H //. by rewrite inE orbA orbb. move: PI=>[R1][U1][L1][UL1] LU1; do 2!split=>//. @@ -1565,8 +1685,10 @@ Qed. Lemma propagate_clos_nopendP : CRel (Data (rep d) (class d) u2' - (ins (rep d c1, rep d c2) (c, c1, c2) (lookup d)) (pending d)) <~> - closure (CRel d +p [Pred x y | x = const c /\ y = app (const c1) (const c2)]). + (ins (rep d c1, rep d c2) (c, c1, c2) (lookup d)) + (pending d)) <~> + closure (CRel d +p + [Pred x y | x = const c /\ y = app (const c1) (const c2)]). Proof. rewrite /CRel clos_clos orrA orrAC -!orrA; apply: closRK; rewrite !orrA. move=>[xx yy]; split=>O; rewrite toPredE -clos_idemp; move: O; apply: sub_closI; @@ -1589,12 +1711,11 @@ Qed. End NoPend. -Module DummyQ. End DummyQ. - (* Lemma about normalization *) -Lemma norm_const (D : data) (t : exp) (s : symb) : - norm D t = const s -> s \in reps D. +Lemma norm_const (D : data) (t : exp) (ss : symb) : + norm D t = const ss -> + ss \in reps D. Proof. elim: t=>[t|t1 IH1 t2 IH2] /=; first by case=><-. do 2![case: (norm D _)=>//] => q2 q1. @@ -1612,7 +1733,8 @@ by rewrite (norm_const E1) (norm_const E2). Qed. Lemma norm_rel (D : data) (t : exp) : - rep_idemp D -> (t, norm D t) \In CRel D. + rep_idemp D -> + (t, norm D t) \In CRel D. Proof. move=>H; elim: t=>[t|t1 IH1 t2 IH2] /=; first by apply: clos_rep. case E1: (norm D t1) IH1=>[q1|] IH1; last by apply: monoC. @@ -1625,7 +1747,8 @@ by rewrite (norm_const E1) (norm_const E2). Qed. Lemma normP (D : data) (x y : exp) : - rep_idemp D -> pending D = [::] -> + rep_idemp D -> + pending D = [::] -> reflect ((x, y) \In CRel D) (norm D x == norm D y). Proof. move=> H1 H2; case: eqP=>H; constructor. @@ -1642,3 +1765,7 @@ move=>[x y]; case. move=>[a][b][c][c1][c2][-> Q1 Q2 Q3 ->] /=; do 2!rewrite reps_rep //. by rewrite Q3 H1. Qed. + +End Congruence. + +Notation rel_exp s := (Pred (exp s * exp s)). diff --git a/examples/counter.v b/examples/counter.v index f85522c..aa80a10 100644 --- a/examples/counter.v +++ b/examples/counter.v @@ -66,8 +66,9 @@ Definition interp (r : rep) : Pred heap := [Pred h | h = rptr r :-> rval r]. Record scnt : Type := SCount {sinv : nat -> rep; - smethod : STsep {v : nat} (interp (sinv v), - [vfun y m => y = v /\ m \In interp (sinv v.+1)])}. + smethod : STsep {v : nat} + (interp (sinv v), + [vfun y m => y = v /\ m \In interp (sinv v.+1)])}. Program Definition snew : STsep (emp, [vfun y m => m \In interp (sinv y 0)]) := @@ -106,8 +107,9 @@ Definition pinterp (R : rep -> Prop) : Pred heap := Record pcnt : Type := PCount {pinv : nat -> rep -> Prop; - pmethod : STsep {v : nat} (pinterp (pinv v), - [vfun y m => y = v /\ m \In pinterp (pinv v.+1)])}. + pmethod : STsep {v : nat} + (pinterp (pinv v), + [vfun y m => y = v /\ m \In pinterp (pinv v.+1)])}. Program Definition pnew : STsep (emp, [vfun y m => m \In pinterp (pinv y 0)]) := diff --git a/examples/hashtab.v b/examples/hashtab.v index 6689359..c9391cf 100644 --- a/examples/hashtab.v +++ b/examples/hashtab.v @@ -22,7 +22,7 @@ From htt Require Import array kvmaps. Module HashTab. Section HashTab. -(* a hash table is an array of buckets, i.e. KV maps *) +(* hash table is array of buckets, i.e. KV maps *) (* bucket indices are provided by the hash function *) Variables (K : ordType) (V : Type) (buckets : KVmap.Sig K V) @@ -32,10 +32,8 @@ Notation KVshape := (@KVmap.shape _ _ buckets). Notation table := (table KVshape). Notation nil := (nil K V). -Opaque Array.write Array.new Array.free Array.read. - (* hash table is specified by a single finMap *) -(* which is morally a "flattening" of all buckets *) +(* which is the "flattening" of all buckets *) Definition shape x (s : finMap K V) := [Pred h | exists (tab : {ffun 'I_n -> KVmap.tp buckets}) (* array spec *) (bucket : 'I_n -> finMap K V), (* buckets spec *) @@ -51,7 +49,7 @@ Definition new_loopinv x := forall k, STsep (fun h => k <= n /\ exists tab, h \In Array.shape x tab # sepit [set x:'I_n | x < k] (table tab (fun=> nil)), - [vfun y => shape y nil]). + [vfun y => shape y nil]). Program Definition new : STsep (emp, [vfun y => shape y nil]) := Do (t <-- Array.new _ (KVmap.default buckets); @@ -65,9 +63,10 @@ Program Definition new : STsep (emp, [vfun y => shape y nil]) := (* first the bucket initialization loop *) Next Obligation. (* pull out preconditions, branch on k comparison *) -move=>/= arr loop k [] _ /= [Eleq][tab][h1][h2][-> H1]; case: decP=>[{Eleq}pf H2|]; last first. -- (* k = n, return the array pointer *) - case: ltnP Eleq (eqn_leq k n)=>// _ -> /= /eqP Ek _ H; step=>_. +move=>/= arr loop k [] _ /= [Eleq][tab][h1][h2][-> H1]. +case: decP=>[{Eleq}pf H2|]; last first. +(* k = n, return the array pointer *) +- case: ltnP Eleq (eqn_leq k n)=>// _ -> /= /eqP Ek _ H; step=>_. (* pass through the constructed table, aggregated finmap is empty *) exists tab, (fun => nil); split=>//; exists h1, h2; split=>//{h1 H1}. (* h2 holds the table *) @@ -79,16 +78,17 @@ apply: [stepX tab] @ h1=>{h1 H1}//= _ m2 [E2 V2]. (* invoke the loop *) apply: [gE]=>//=; split=>//; rewrite joinCA. (* extend the table by the new index/bucket pair, simplify *) -exists [ffun z => if z == Ordinal pf then b else tab z], m2, (m \+ h2); split=>//{m2 E2 V2}. +exists [ffun z => if z == Ordinal pf then b else tab z], m2, (m \+ h2). +split=>//{m2 E2 V2}. (* remove the new bucket from the heap *) -rewrite (sepitS (Ordinal pf)) in_set leqnn {1}/table ffunE eq_refl; exists m, h2; do!split=>{m Hm}//. -apply: tableP2 H2=>{h2}//. +rewrite (sepitS (Ordinal pf)) in_set leqnn {1}/table ffunE eq_refl. +exists m, h2; do!split=>{m Hm}//; apply: tableP2 H2=>{h2}//. - case=>x Hx; rewrite !in_set in_set1 -val_eqE /= ltnS (leq_eqVlt x). by case: ltngtP. (* removing k from the domain of the new table gives the old table back *) by move=>x _; rewrite in_set ffunE; case: eqP=>//->; rewrite ltnn. Qed. -(* now the outer function *) +(* the outer function *) Next Obligation. (* simplify *) move=>/= [] _ ->. @@ -97,12 +97,13 @@ apply: [stepE]=>//= y m [Hm Vm]. (* invoke the loop with index 0 *) apply: [gE]=>//=; split=>//. (* the table is empty *) -exists [ffun => KVmap.default buckets], m, Unit; split=>//=; first by rewrite unitR. +exists [ffun => KVmap.default buckets], m, Unit; split=>//=. +- by rewrite unitR. (* there are no buckets in the heap yet *) by rewrite (eq_sepit (s2 := set0)) // sepit0. Qed. -(* freeing a hashtable is freeing the array + buckets *) +(* freeing hashtable = freeing the array + buckets *) (* loop invariant: *) (* at iteration k, the heap still holds the array and n-k buckets *) @@ -110,11 +111,11 @@ Definition free_loopinv x := forall k, STsep (fun h => k <= n /\ exists t b, h \In Array.shape x t # sepit [set x:'I_n | x >= k] (table t b), - [vfun _ : unit => emp]). + [vfun _ : unit => emp]). Program Definition free x : STsep {s} (shape x s, - [vfun _ : unit => emp]) := - (* we add an extra Do here so we can derive the precondition from the loop *) + [vfun _ : unit => emp]) := + (* the extra Do here enables deriving precondition from the loop *) Do (ffix (fun (loop : free_loopinv x) k => Do (if decP (b := k < n) idP is left pf then b <-- Array.read x (Ordinal pf); @@ -124,18 +125,21 @@ Program Definition free x : STsep {s} (shape x s, (* first the loop *) Next Obligation. (* pull out the ghost + preconditions, branch on k comparison *) -move=>/= x loop k [] _ /= [Eleq][tf][bf][h1][h2][-> [H1 V1]]; case: decP=>[pf H|]; last first. -- (* k = n *) - case: ltnP Eleq (eqn_leq k n)=>// _ -> /= /eqP Ek _ H. +move=>/= x loop k [] _ /= [Eleq][tf][bf][h1][h2][-> [H1 V1]]. +case: decP=>[pf H|]; last first. +(* k = n *) +- case: ltnP Eleq (eqn_leq k n)=>// _ -> /= /eqP Ek _ H. (* free the array *) apply: [gE]=>//=; exists tf. (* h2 is empty *) - move: H; rewrite (eq_sepit (s2 := set0)); first by rewrite sepit0=>->; rewrite unitR. + move: H; rewrite (eq_sepit (s2 := set0)). + - by rewrite sepit0=>->; rewrite unitR. by move=>y; rewrite Ek in_set in_set0 leqNgt ltn_ord. (* k < n, read from array *) apply: [stepX tf, h1] @ h1=>//= _ _ [->->]. (* split out the the k-th bucket *) -move: H; rewrite (sepitS (Ordinal pf)) in_set leqnn; case=>h3[h4][{h2}-> H3 H4]. +move: H; rewrite (sepitS (Ordinal pf)) in_set leqnn. +case=>h3[h4][{h2}-> H3 H4]. (* free it *) apply: [stepX (bf (Ordinal pf))] @ h3=>{h3 H3}//= _ _ ->; rewrite unitL. (* invoke loop, simplify *) @@ -153,12 +157,13 @@ move=>/= x [fm][] h /= [tf][bf][_ _ H]. by apply: [gE]=>//=; eauto. Qed. -(* inserting into a hashmap is inserting into corresponding bucket + updating the array *) -(* returning the pointer is technically not needed, as the array is not moved *) -(* but we need to fit the KV map API *) - -Program Definition insert x k v : STsep {s} (shape x s, - [vfun y => shape y (ins k v s)]) := +(* inserting into hashmap is inserting into *) +(* corresponding bucket + updating the array *) +(* returning the pointer is technically not needed *) +(* as the array is not moved *) +(* but one needs to fit the KV map API *) +Program Definition insert x k v : + STsep {s} (shape x s, [vfun y => shape y (ins k v s)]) := Do (let hk := hash k in b <-- Array.read x hk; b' <-- KVmap.insert b k v; @@ -178,26 +183,29 @@ apply/[stepX tf] @ h1=>{h1 H1}//= _ m3 [E3 V3]; step=>_. (* update the array and buckets' specs *) exists [ffun z => if z == hash k then b' else tf z], (fun b => if b == hash k then ins k v (bf b) else bf b); split=>/=. -- (* the new buckets spec is still a flattening *) - move=>k0; rewrite fnd_ins; case: eqP=>[->|/eqP Ek]. - - (* if the bucket was touched, we get the same value *) - by rewrite eq_refl fnd_ins eq_refl. +(* the new buckets spec is still a flattening *) +- move=>k0; rewrite fnd_ins; case: eqP=>[->|/eqP Ek]. + (* if the bucket was touched, we get the same value *) + - by rewrite eq_refl fnd_ins eq_refl. (* if not, we get the old spec back *) by case: ifP=>_ //; rewrite fnd_ins (negbTE Ek). -- (* the new buckets spec respects the hash *) - move=>i0 k0; case: eqP; last by move=>_; apply: Hh. +(* the new buckets spec respects the hash *) +- move=>i0 k0; case: eqP; last by move=>_; apply: Hh. by move=>->; rewrite supp_ins inE=>/orP; case; [move/eqP=>->|apply: Hh]. (* the shape is respected: first, the array fits *) exists m3, (m2 \+ h4); split=>{Hf Hh m3 E3 V3}//. (* split out the modified bucket *) -rewrite (sepitT1 (hash k)) /table /= ffunE eq_refl; exists m2, h4; split=>{m2 H2}//. +rewrite (sepitT1 (hash k)) /table /= ffunE eq_refl. +exists m2, h4; split=>{m2 H2} //. (* the table fits too *) -by apply/tableP/H4=>/= x0; rewrite !in_set in_set1 andbT ?ffunE =>/negbTE->. +apply/tableP/H4=>/= x0; +by rewrite !in_set in_set1 andbT ?ffunE =>/negbTE->. Qed. -(* removing from a hashmap is removing from corresponding bucket + updating the array *) -(* returning the pointer is again not needed except for the API fit *) - +(* removing from hashmap is removing from *) +(* corresponding bucket + updating the array *) +(* returning the pointer is again not needed *) +(* except for the API fit *) Program Definition remove x k : STsep {s} (shape x s, [vfun y => shape y (rem k s)]) := @@ -220,25 +228,26 @@ apply/[stepX tf] @ h1=>{H1}//= _ m3 [E3 V3]; step=>_. (* update the array and buckets' specs *) exists [ffun z => if z == hash k then b' else tf z], (fun b => if b == hash k then rem k (bf b) else bf b); split=>/=. -- (* the new buckets spec is still a flattening *) - move=>k0; rewrite fnd_rem; case: eqP. - - (* if the bucket was touched, the value is gone *) - by move=>->; rewrite eq_refl fnd_rem eq_refl. +(* the new buckets spec is still a flattening *) +- move=>k0; rewrite fnd_rem; case: eqP. + (* if the bucket was touched, the value is gone *) + - by move=>->; rewrite eq_refl fnd_rem eq_refl. (* if not, we get the old spec back *) by move/eqP=>Ek; case: ifP=>_ //; rewrite fnd_rem (negbTE Ek). -- (* the new buckets spec respects the hash *) - move=>i0 k0; case: eqP; last by move=>_; apply: Hh. +(* the new buckets spec respects the hash *) +- move=>i0 k0; case: eqP; last by move=>_; apply: Hh. by move=>->; rewrite supp_rem !inE=>/andP [] _; apply: Hh. (* the shape is respected: first, the array fits *) exists m3, (m2\+ h4); split=>{m3 E3 V3 Hf Hh}//. (* split out the modified bucket *) -rewrite (sepitT1 (hash k)) /table /= ffunE eq_refl; exists m2, h4; split=>{m2 H2}//. +rewrite (sepitT1 (hash k)) /table /= ffunE eq_refl. +exists m2, h4; split=>{m2 H2} //. (* the table fits too *) -by apply/tableP/H4=>/= x0; rewrite !in_set in_set1 andbT ?ffunE =>/negbTE->. +apply/tableP/H4=>/= x0; +by rewrite !in_set in_set1 andbT ?ffunE => /negbTE->. Qed. (* looking up in a hashmap is looking up in the corresponging bucket *) - Program Definition lookup x k : STsep {s} (shape x s, [vfun y m => m \In shape x s /\ y = fnd k s]) := @@ -252,15 +261,14 @@ apply/[stepX tf, h1] @ h1=>//= _ _ [->->]. (* split out the bucket in the heap *) move: H2; rewrite (sepitT1 (hash k)); case=>h3[h4][{h2}-> H3 H4]. (* look up in the bucket, simplify *) -apply/[gX (bf (hash k))] @ h3=>{h3 H3}//= r m2 [H2 Hr] _; split; last by rewrite Hf. +apply/[gX (bf (hash k))] @ h3=>{h3 H3}//= r m2 [H2 Hr] _. +split; last by rewrite Hf. (* the shape is preserved, as nothing was modified *) -exists tf, bf; split=>//=; exists h1, (m2 \+ h4); split=>{h1 H1}//. +exists tf, bf; split=>//=; exists h1, (m2 \+ h4); split=>{h1 H1} //. by rewrite (sepitT1 (hash k)); vauto. Qed. - (* hash table is a KV map *) - Definition HashTab := KVmap.make (Array null) new free insert remove lookup. End HashTab. diff --git a/examples/kvmaps.v b/examples/kvmaps.v index 5ad60b7..caff2af 100644 --- a/examples/kvmaps.v +++ b/examples/kvmaps.v @@ -32,16 +32,16 @@ Record Sig (K : ordType) (V : Type) : Type := shape : tp -> {finMap K -> V} -> Pred heap; new : STsep (emp, [vfun x => shape x (nil K V)]); free : forall x, STsep {s} (shape x s, - [vfun _ : unit => emp]); + [vfun _ : unit => emp]); insert : forall x k v, STsep {s} (shape x s, - [vfun y => shape y (ins k v s)]); + [vfun y => shape y (ins k v s)]); remove : forall x k, STsep {s} (shape x s, - [vfun y => shape y (rem k s)]); + [vfun y => shape y (rem k s)]); lookup : forall x k, STsep {s} (shape x s, - [vfun y m => m \In shape x s /\ y = fnd k s])}. + [vfun y m => m \In shape x s /\ y = fnd k s])}. End KVmap. (**********************************************************) @@ -58,9 +58,10 @@ Notation nil := (nil K V). Definition entry (p q : ptr) (k : K) (v : V) : heap := p :-> k \+ p.+1 :-> v \+ p.+2 :-> q. -(* similarly to plain linked list, we with a generic definition of a segment and *) -(* then specialize it to self-contained lists specified by a finMap structure *) - +(* similarly to plain linked list *) +(* development starts with generic definition of list segment *) +(* and then specializes to self-contained lists specified *) +(* by finMap structure *) Fixpoint shape_seg' (x y : ptr) (xs : seq (K * V)) : Pred heap := if xs is (k,v) :: tl then [Pred h | exists q h', @@ -75,7 +76,6 @@ Definition shape (x : ptr) (s : fmap) : Pred heap := shape_seg x null s. (* null pointer represents an empty map *) - Lemma shape_null (s : fmap) h : valid h -> h \In shape null s -> @@ -89,45 +89,35 @@ Qed. (* non-empty well-formed region represents a non-empty map such that *) (* a head entry respecting the key order can be pulled out of the map *) - Lemma shape_cont (s : fmap) p h : p != null -> h \In shape p s -> exists k v q h', - [/\ s = ins k v (behd s), (* k:->v is the head entry *) + [/\ s = ins k v (behd s), (* k:->v is the head entry *) all (ord k) (supp (behd s)), h = entry p q k v \+ h' & h' \In shape q (behd s)]. Proof. move=>E; case: s=>xs srt /=. -elim: xs srt=>/=. -- by move=>? [E0]; rewrite E0 in E. +elim: xs srt=>/=; first by move=>? [E0]; rewrite E0 in E. move=>[k v] /= l IH srt [q][h'][-> H]. -exists k,v,q,h'; split=>//; last by apply: order_path_min. +exists k, v, q, h'; split=>//; last by apply: order_path_min. by rewrite fmapE /= last_ins'. Qed. -(* TODO move to finmap? *) -Lemma all_path (s : fmap) k : - all (ord k) (supp s) -> - path ord k (supp s). -Proof. by rewrite path_sortedE // =>->/=; case: s. Qed. - (* inserting an entry with minimal key is prepending to the list *) - Lemma shape_cons (s : fmap) p q h k v : all (ord k) (supp s) -> h \In shape q s -> entry p q k v \+ h \In shape p (ins k v s). Proof. -move/all_path=>S H. +move/all_path_supp=>S H. suff: ins k v s = @FinMap _ _ ((k,v)::seq_of s) S by move=>->; exists q,h. rewrite fmapE /=; case: s {H}S=>/= xs ??. by rewrite last_ins'. Qed. (* inserting an entry with maximal key is appending to the list *) - Lemma shape_seg_rcons (s : fmap) x p q h k v : (* conceptually last s < k *) all (ord^~ k) (supp s) -> @@ -145,7 +135,6 @@ by apply: IH=>//; apply: (path_sorted S). Qed. (* disjoint maps can be concatenated if the order is respected *) - Lemma shape_fcat s1 s2 h1 h2 x y : (* conceptually last s1 < head s2 *) allrel ord (supp s1) (supp s2) -> @@ -169,30 +158,28 @@ Qed. (* main procedures *) -(* a new map is just a null pointer *) - +(* new map is just a null pointer *) Program Definition new : STsep (emp, [vfun x => shape x nil]) := Do (ret null). Next Obligation. by move=>[] /= _ ->; step. Qed. (* freeing a map is deallocating all its elements *) - Definition freeT : Type := forall p, STsep {fm} (shape p fm, [vfun _ : unit => emp]). Program Definition free : freeT := ffix (fun (loop : freeT) p => - Do (if p == null then ret tt - else pnext <-- !p.+2; - dealloc p;; - dealloc p.+1;; - dealloc p.+2;; - loop pnext)). + Do (if p == null then ret tt + else pnext <-- !p.+2; + dealloc p;; + dealloc p.+1;; + dealloc p.+2;; + loop pnext)). Next Obligation. (* pull out ghost var, precondition, branch *) move=>loop p [fm][] i /= H; case: eqP=>[|/eqP] E. -- (* reached the end, heap must be empty *) - by apply: vrfV=>D; step=>_; rewrite E in H; case: (shape_null D H). +(* reached the end, heap must be empty *) +- by apply: vrfV=>D; step=>_; rewrite E in H; case: (shape_null D H). (* pull out the head entry *) case: (shape_cont E H)=>k[v][q][h][_ _ {i H}-> H]. (* deallocate it *) @@ -215,44 +202,44 @@ Program Definition lookup x (k : K) : else k' <-- !cur; if k == k' - then v <-- !cur.+1; - ret (Some v) - else if ord k' k - then next <-- !cur.+2; - loop next - else ret None)) x. + then v <-- !cur.+1; + ret (Some v) + else if ord k' k + then next <-- !cur.+2; + loop next + else ret None)) x. Next Obligation. (* pull out ghost var+precondition, branch on cur being null *) move=>_ k go cur [fm][] h /= H; case: eqP=>[|/eqP] Ec. -- (* reached the end, the heap and the spec must be empty, element not found *) - by apply: vrfV=>Vh; step=>_; rewrite Ec in H; case: (shape_null Vh H)=>->. -(* pull out the head entry *) +(* reached the end, heap and spec must be empty, element not found *) +- by apply: vrfV=>Vh; step=>_; rewrite Ec in H; case: (shape_null Vh H)=>->. +(* pull out head entry *) case: (shape_cont Ec H)=>k'[v][next][h'][Ef O' Ei H']; rewrite {h}Ei in H *. (* read the head key, branch on equality comparison *) step; case: eqP=>[|/eqP] Ek. -- (* the key matches, return the head value *) - by do 2![step]=>_; split=>//; rewrite Ef fnd_ins Ek eq_refl. +(* the key matches, return the head value *) +- by do 2![step]=>_; split=>//; rewrite Ef fnd_ins Ek eq_refl. (* branch on comparison *) case: ifP=>Ho'. -- (* head key is less than the needed one, loop on tail *) - (* (we fall back to gR to preserve associativity) *) - step; apply/[gR (behd fm)] @ h'=>//= v0 h0' [H0 E0] _. +(* head key is less than the needed one, loop on tail *) +(* (fall back to gR to preserve associativity) *) +- step; apply/[gR (behd fm)] @ h'=>//= v0 h0' [H0 E0] _. by rewrite Ef fnd_ins (negbTE Ek); split=>//; exact: shape_cons. (* head key is bigger than the needed one, abort *) move: (connex Ek); rewrite {}Ho' orbC /= =>Ho'; step=>_; split=>//. (* k is not in the head entry *) apply/esym/fnd_supp; rewrite Ef supp_ins inE negb_or; apply/andP; split=>//. (* nor it is in the tail *) -by apply/notin_path/path_le/all_path/O'. +by apply/notin_path/path_le/all_path_supp/O'. Qed. -(* removing an element by key from the map, return the pointer to the new map *) +(* removing element by key from the map, return the pointer to the new map *) (* loop invariant: *) -(* 1. we split the list into a zipper-like structure `fml ++ [k'->v'] ++ fmr` *) -(* 2. the ordering is respected *) +(* 1. split the list into a zipper-like structure `fml ++ [k'->v'] ++ fmr` *) +(* 2. ordering is respected *) (* 3. k is not in fml nor in the focus entry k'->v' *) -(* the focus is needed so that we can connect the remainder of the map to it after removal *) +(* the focus is needed to connect the remainder of the map to it after removal *) Definition removeT p k : Type := forall (prevcur : ptr * ptr), STsep {fm} (fun h => exists fml fmr k' v', @@ -263,7 +250,7 @@ Definition removeT p k : Type := (shape_seg p prevcur.1 fml # (fun h => h = entry prevcur.1 prevcur.2 k' v') # shape prevcur.2 fmr)], - [vfun _ : unit => shape p (rem k fm)]). + [vfun _ : unit => shape p (rem k fm)]). Program Definition remove x (k : K) : STsep {fm} (shape x fm, @@ -297,29 +284,35 @@ Program Definition remove x (k : K) : ret x else ret x). (* first the loop *) -(* we don't return the pointer because it cannot change, as the head is fixed by fml *) +(* don't return the pointer because it cannot change *) +(* as the head is fixed by fml *) Next Obligation. (* pull out ghost var, preconditions and heap validity *) -move=>x k0 go _ prev cur [_][] _ /= [fml][fmr][k][v][-> [Ol Or][El E]][hl][_][-> Hl [_][hr][->->Hr]]. +move=>x k0 go _ prev cur [_][] _ /= [fml][fmr][k][v][-> [Ol Or] +[El E]][hl][_][-> Hl [_][hr][->->Hr]]. (* branch on cur *) case: eqP=>[|/eqP] Ec. -(* cur = null - nothing to left to process, i.e., fmr = [] *) + (* cur = null - nothing to left to process, i.e., fmr = [] *) - apply: vrfV=>Vh; step=>_; rewrite {}Ec in Hr *. (* k is not in fml ++ (k->v) *) - case: (shape_null (validX Vh) Hr)=>/=->->; rewrite fcats0 unitR rem_ins (negbTE E) (rem_supp El). - by exact: shape_seg_rcons. + case: (shape_null (validX Vh) Hr)=>/=->->. + rewrite fcats0 unitR rem_ins (negbTE E) (rem_supp El). + exact: shape_seg_rcons. (* cur <> null, pull out the head entry from fmr *) -case: (shape_cont Ec Hr)=>k'[v'][next][hr'][Efr /all_path Or' {hr Hr Ec}-> Hr']; rewrite joinA joinC. +case: (shape_cont Ec Hr)=>k'[v'][next][hr'] +[Efr /all_path_supp Or' {hr Hr Ec}-> Hr']; rewrite joinA joinC. (* derive ordering facts *) -move/all_path: (Or); rewrite {1}Efr; case/(path_supp_ins_inv Or')/andP=>Ho' Or''. +move/all_path_supp: (Or); rewrite {1}Efr; +case/(path_supp_ins_inv Or')/andP=>Ho' Or''. (* get head key, branch on comparing it to needed one *) step; case: eqP=>[|/eqP] Ek. -- (* k = k' - element is found, run the deallocations *) - do 4!step; rewrite !unitL; do 2![step]=>_. + (* k = k' - element is found, run the deallocations *) +- do 4!step; rewrite !unitL; do 2![step]=>_. (* pull out fml ++ (k->v) *) rewrite Efr -fcat_srem; last by rewrite supp_ins inE negb_or E. (* drop the element in the spec *) - rewrite rem_ins {1}Ek eq_refl rem_supp; last by rewrite Ek; apply: notin_path. + rewrite rem_ins {1}Ek eq_refl rem_supp; + last by rewrite Ek; apply: notin_path. (* heap shape is respected *) rewrite joinC; apply/shape_fcat/Hr'; last by apply: shape_seg_rcons. (* the ordering is respected as well *) @@ -328,22 +321,23 @@ step; case: eqP=>[|/eqP] Ek. by apply/(allrel_trans (z:=k))/order_path_min=>//=. (* k <> k', now branch on order comparison *) case: ifP=>Ho0. -- (* k' is less than k, invoke the loop, shifting the focus to the right *) - step; apply: [gE (fcat (ins k' v' (ins k v fml)) (behd fmr))]=>//=. - - (* prove that all conditions are respected *) - exists (ins k v fml), (behd fmr), k', v'; do!split=>//. - - (* new focus comes after fml ++ old focus *) - rewrite (eq_all_r (s2:=k::supp fml)) /= ?Ho' /=; last by apply: supp_ins. +(* k' is less than k, invoke the loop, shifting the focus to the right *) +- step; apply: [gE (fcat (ins k' v' (ins k v fml)) (behd fmr))]=>//=. + (* prove that all conditions are respected *) + - exists (ins k v fml), (behd fmr), k', v'; do!split=>//. + (* new focus comes after fml ++ old focus *) + - rewrite (eq_all_r (s2:=k::supp fml)) /= ?Ho' /=; + last by apply: supp_ins. by apply/sub_all/Ol=>? /trans; apply. - - (* new focus comes before the new suffix *) - by apply: order_path_min. - - (* the needed key is not in fml ++ old focus *) - by rewrite supp_ins inE negb_or E. + (* new focus comes before the new suffix *) + - by apply: order_path_min. + (* the needed key is not in fml ++ old focus *) + - by rewrite supp_ins inE negb_or E. (* heap shape is respected *) exists (hl \+ entry prev cur k v), (entry cur next k' v' \+ hr'). rewrite joinC; split=>//; last by vauto. by apply: shape_seg_rcons. - (* we can reassemble the spec because insertions of old and new foci commute *) + (* reassemble the spec, as insertions of old and new foci commute *) move=>_ m Hm Vm; rewrite Efr. by rewrite fcat_inss // -?fcat_sins // in Hm; apply: notin_path. (* k' is bigger than k, abort *) @@ -376,29 +370,31 @@ case: (shape_cont Ex H)=>{Ex}k[v][next][h'][Ef Of Eh H']; rewrite Eh. step; case: eqP=>[->|/eqP Ek]. - (* k = k', element found in head position, run deallocations, return new head *) do 5![step]=>_; rewrite !unitL Ef rem_ins eq_refl rem_supp //. - by apply/notin_path/all_path. + by apply/notin_path/all_path_supp. (* k <> k', now branch on order comparison *) case: ifP=>Ho0. (* k' is less than k, start the loop with the head entry as the focus *) - step; apply: [stepE fm]=>//=; last by move=>_ ??; step. (* invariants and shape are satisfied *) exists nil, (behd fm), k, v; do!split=>//. - - by rewrite fcat_inss; [rewrite fcat0s|apply/notin_path/all_path]. + - by rewrite fcat_inss; [rewrite fcat0s|apply/notin_path/all_path_supp]. by exists Unit, (entry x next k v \+ h'); split=>//; [rewrite unitL | vauto]. (* k' is bigger than k, abort *) move: (connex Ek); rewrite {}Ho0 orbC /= =>Ho0. (* return the old head, invariants are preserved *) step=>_; rewrite -Eh rem_supp // Ef supp_ins !inE negb_or Ek /=. -by apply/notin_path/path_le/all_path/Of. +by apply/notin_path/path_le/all_path_supp/Of. Qed. -(* upserting (inserting or updating if the key is found) an entry into the map, return the pointer to the new map *) +(* upserting (inserting or updating if the key is found) *) +(* an entry into the map, return the pointer to the new map *) (* loop invariant, essentially identical to remove: *) -(* 1. as for remove, we split the list into a zipper-like structure `fml ++ [k'->v'] ++ fmr` *) +(* 1. as in remove, split the list into a zipper-like structure *) +(* `fml ++ [k'->v'] ++ fmr` *) (* 2. the ordering is respected *) (* 3. k is not in fml and is less than the key k' of focus entry *) -(* the focus is needed so that we can connect the new element to it after insertion *) +(* the focus is needed to connect the new element after insertion *) Definition insertT p k v : Type := forall (prevcur : ptr * ptr), STsep {fm} (fun h => exists fml fmr k' v', @@ -409,11 +405,11 @@ Definition insertT p k v : Type := (shape_seg p prevcur.1 fml # (fun h => h = entry prevcur.1 prevcur.2 k' v') # shape prevcur.2 fmr)], - [vfun _ : unit => shape p (ins k v fm)]). + [vfun _ : unit => shape p (ins k v fm)]). Program Definition insert x (k : K) (v : V) : STsep {fm} (shape x fm, - [vfun y => shape y (ins k v fm)]) := + [vfun y => shape y (ins k v fm)]) := Do (let go := ffix (fun (loop : insertT x k v) '(prev, cur) => Do (if cur == null then new <-- allocb k 3; @@ -452,14 +448,16 @@ Program Definition insert x (k : K) (v : V) : new.+2 ::= x;; ret new). (* first the loop *) -(* we don't return the pointer because it cannot change, as the head is fixed by fml *) +(* don't return the pointer because it cannot change *) +(* as the head is fixed by fml *) Next Obligation. (* pull out ghost var+preconditions *) -move=>x k0 v0 loop _ prev cur [_][] _ /= [fml][fmr][k][v][-> [Ol Or][El Ho0]][hl][_][-> Hl [_][hr][->-> Hr]]. +move=>x k0 v0 loop _ prev cur [_][] _ /= [fml][fmr][k][v] +[-> [Ol Or][El Ho0]][hl][_][-> Hl [_][hr][->-> Hr]]. (* branch on cur *) case: eqP=>[|/eqP] Ec. -- (* cur = null, insert as the last element *) - rewrite {}Ec in Hr; apply: vrfV=>Vh. +(* cur = null, insert as the last element *) +- rewrite {}Ec in Hr; apply: vrfV=>Vh. step=>p; rewrite unitR; do 2!step; rewrite joinC; do 2![step]=>_. (* fmr is empty *) case: (shape_null (validX Vh) Hr)=>/=->->. @@ -469,14 +467,16 @@ case: eqP=>[|/eqP] Ec. rewrite (eq_all_r (s2:=k::supp fml)) /= ?Ho0 /=; last by apply: supp_ins. by apply/sub_all/Ol=>? /trans; apply. (* cur <> null, pull out the head entry from fmr *) -case: (shape_cont Ec Hr)=>k'[v'][next][hr'][Efr Or' {hr Hr Ec}-> Hr']; rewrite joinA joinC. +case: (shape_cont Ec Hr)=>k'[v'][next][hr'][Efr Or' {hr Hr Ec}-> Hr']. +rewrite joinA joinC. (* derive ordering facts *) -move/all_path: (Or); rewrite {1}Efr; case/(path_supp_ins_inv (all_path Or'))/andP=>Ho'. +move/all_path_supp: (Or); rewrite {1}Efr. +case/(path_supp_ins_inv (all_path_supp Or'))/andP=>Ho'. move/(order_path_min (@trans _))=>Or''. (* get new key, branch on equality comparison *) step; case: eqP=>[|/eqP] Ek. -- (* k = k', update the value at this position *) - do 2![step]=>_; rewrite Efr -fcat_sins ins_ins -Ek eq_refl joinC. +(* k = k', update the value at this position *) +- do 2![step]=>_; rewrite Efr -fcat_sins ins_ins -Ek eq_refl joinC. (* invariants are preserved as the key didn't change *) apply: shape_fcat; first 1 last. - by apply: shape_seg_rcons. @@ -490,31 +490,33 @@ step; case: eqP=>[|/eqP] Ek. case: ifP=>Ho'0. (* k' is less than k, invoke the loop, shifting the focus to the right *) - step; apply/[gE (fcat (ins k' v' (ins k v fml)) (behd fmr))]=>//=. - - (* prove that all conditions are respected *) - exists (ins k v fml), (behd fmr), k', v'; do!split=>//. - - (* new focus comes after fml ++ old focus *) - rewrite (eq_all_r (s2:=k::supp fml)) /= ?Ho' /=; last by apply: supp_ins. + (* prove that all conditions are respected *) + - exists (ins k v fml), (behd fmr), k', v'; do!split=>//. + (* new focus comes after fml ++ old focus *) + - rewrite (eq_all_r (s2:=k::supp fml)) /= ?Ho' /=; last by apply: supp_ins. by apply/sub_all/Ol=>? /trans; apply. - - (* the needed key is not in fml ++ old focus *) - rewrite supp_ins inE negb_or andbC El /=. + (* the needed key is not in fml ++ old focus *) + - rewrite supp_ins inE negb_or andbC El /=. by case: ordP Ho0. (* heap shape is respected *) exists (hl \+ entry prev cur k v), (entry cur next k' v' \+ hr'). by rewrite joinC; split=>//; [apply: shape_seg_rcons | vauto]. - (* we can reassemble the spec because insertions of old and new foci commute *) + (* reassemble the spec, as insertions of old and new foci commute *) move=>_ m Hm _; rewrite Efr. rewrite fcat_inss // in Hm; first by rewrite -fcat_sins in Hm. - by apply/notin_path/all_path. + by apply/notin_path/all_path_supp. (* k' is bigger than k, insert at this position *) move: (connex Ek); rewrite {}Ho'0 orbC /= =>Ho0'. (* run the allocation and assignments *) step=>new; rewrite unitR; do 2!step; rewrite joinA joinC; do 2![step]=>_. rewrite Efr -fcat_sins; apply: shape_fcat; first 1 last. -- (* shape is respected for the prefix fml ++ old focus *) - by apply: shape_seg_rcons. -- (* shape is satisfed for new element + suffix *) - rewrite [X in X \+ (entry _ _ _ _ \+ _)]joinA; apply/shape_cons/shape_cons=>//. - by apply/order_path_min=>//; apply/path_supp_ins=>//; apply/path_le/all_path/Or'. +(* shape is respected for the prefix fml ++ old focus *) +- by apply: shape_seg_rcons. +(* shape is satisfed for new element + suffix *) +- rewrite [X in X \+ (entry _ _ _ _ \+ _)]joinA. + apply/shape_cons/shape_cons=>//. + apply/order_path_min=>//; apply/path_supp_ins=>//. + by apply/path_le/all_path_supp/Or'. (* ordering is respected *) rewrite (allrel_in_l (xs':=k::supp fml) _); last by apply: supp_ins. rewrite (allrel_in_r (ys':=k0::k'::supp (behd fmr)) _ _); last first. @@ -522,7 +524,7 @@ rewrite (allrel_in_r (ys':=k0::k'::supp (behd fmr)) _ _); last first. rewrite allrel_consl !allrel_consr /= Ho0 Ho' Or'' /=; apply/and3P; split. - by apply/sub_all/Ol=>? /trans; apply. - by apply/sub_all/Ol=>? /trans; apply. -by apply: (allrel_trans (z:=k))=>//; exact: trans. +by apply: (allrel_trans (z:=k))=>//; apply: trans. Qed. (* now the outer function, which mostly repeats the loop *) (* the first iteration is special because we don't yet have a left prefix *) @@ -530,8 +532,8 @@ Qed. Next Obligation. (* pull out ghost+precondition, branch on x *) move=>/= x k0 v0 [fm][]h /= H; case: eqP=>[|/eqP] Ex. -- (* x = null, insert as the only element, heap and spec are empty *) - apply: vrfV=>Vh; move: H; rewrite Ex=>/(shape_null Vh) [->->]. +(* x = null, insert as the only element, heap and spec are empty *) +- apply: vrfV=>Vh; move: H; rewrite Ex=>/(shape_null Vh) [->->]. (* run *) step=>p; rewrite !unitR; do 3![step]=>_. by exists null, Unit; rewrite unitR joinA. @@ -539,24 +541,27 @@ move=>/= x k0 v0 [fm][]h /= H; case: eqP=>[|/eqP] Ex. case: (shape_cont Ex H)=>{Ex}k[v][next][h'][Ef Of {h H}-> H]. (* read the head key, branch on equality comparison *) step; case: eqP=>[->|/eqP Ek]. -- (* k = k', exact key found in head position, update the head value *) - do 2![step]=>_; rewrite Ef ins_ins eq_refl. +(* k = k', exact key found in head position, update the head value *) +- do 2![step]=>_; rewrite Ef ins_ins eq_refl. by apply: shape_cons. (* k <> k', now branch on order comparison *) case: ifP=>Ho0. -- (* k' is less than k, run the loop with the head entry as the focus *) - step; apply: [stepE fm]=>//=; last by move=>_ ??; step. +(* k' is less than k, run the loop with the head entry as the focus *) +- step; apply: [stepE fm]=>//=; last by move=>_ ??; step. (* invariants are respected *) exists nil, (behd fm), k, v; do!split=>//. - - by rewrite fcat_inss; [rewrite fcat0s|apply/notin_path/all_path]. - by exists Unit, (entry x next k v \+ h'); split=>//; [rewrite unitL | vauto]. + - by rewrite fcat_inss; [rewrite fcat0s|apply/notin_path/all_path_supp]. + by exists Unit, (entry x next k v \+ h'); split=>//; [rewrite unitL|vauto]. (* k' is bigger than k, insert after the head *) move: (connex Ek); rewrite {}Ho0 orbC /= =>Ho0. (* run allocation and assignments, return the old head *) step=>q; rewrite unitR; do 3![step]=>_. (* invariants are preserved *) -rewrite Ef [X in X \+ (entry _ _ _ _ \+ _)]joinA; apply/shape_cons/shape_cons=>//. -by apply/order_path_min=>//; apply/path_supp_ins=>//; apply/path_le/all_path/Of. +rewrite Ef [X in X \+ (entry _ _ _ _ \+ _)]joinA. +apply/shape_cons/shape_cons=>//. +apply/order_path_min=>//. +apply/path_supp_ins=>//. +by apply/path_le/all_path_supp/Of. Qed. (* ordered association list is a KV map *) diff --git a/examples/llist.v b/examples/llist.v index c4bf28d..95682de 100644 --- a/examples/llist.v +++ b/examples/llist.v @@ -19,6 +19,8 @@ From htt Require Import options model heapauto. (* Linked lists, storing a value and next pointer in consecutive locations. *) (* We start with a more general "segment" definition, where the last node's *) (* next pointer is an arbitrary q *) +(* NOTE: already defined in heap.v under the name llist, but repeated here *) +(* to make the file self-contained. *) Fixpoint lseg {A} (p q : ptr) (xs : seq A) := if xs is hd::tl then diff --git a/examples/quicksort.v b/examples/quicksort.v index f9157e7..661aea3 100644 --- a/examples/quicksort.v +++ b/examples/quicksort.v @@ -12,8 +12,8 @@ limitations under the License. *) From mathcomp Require Import ssreflect ssrbool ssrfun. -From mathcomp Require Import ssrnat eqtype seq path fintype tuple finfun finset. -From mathcomp Require Import fingroup perm. +From mathcomp Require Import ssrnat eqtype seq path fintype tuple finfun. +From mathcomp Require Import finset fingroup perm. From mathcomp Require Import interval. From pcm Require Import options axioms prelude seqext pred ordtype slice. From pcm Require Import pcm unionmap heap. @@ -24,19 +24,10 @@ From htt Require Import array. From mathcomp Require order. Import order.Order.NatOrder order.Order.TTheory. -(* TODO added to mathcomp > 1.16 *) -Lemma perm_on_id {T : finType} (u : {perm T}) (S : {set T}) : - perm_on S u -> - #|S| <= 1 -> - u = 1%g. -Proof. -rewrite leq_eqVlt ltnS leqn0 => pSu S10; apply/permP => t; rewrite perm1. -case/orP : S10; last first. - by move/eqP/cards0_eq => S0; apply: (out_perm pSu); rewrite S0 inE. -move=> /cards1P[x Sx]. -have [-> | ntx] := eqVneq t x; last by apply: (out_perm pSu); rewrite Sx inE. -by apply/eqP; have := perm_closed x pSu; rewrite Sx !inE => ->. -Qed. +(* Brief mathematics of quickorting *) +(* There is some overlap with mathematics developed for bubblesort *) +(* but the development is repeated here to make the files *) +(* self-contained *) Lemma perm_onC {T : finType} (S1 S2 : {set T}) (u1 u2 : {perm T}) : perm_on S1 u1 -> @@ -46,10 +37,10 @@ Lemma perm_onC {T : finType} (S1 S2 : {set T}) (u1 u2 : {perm T}) : Proof. move=> pS1 pS2 S12; apply/permP => t; rewrite !permM. case/boolP : (t \in S1) => tS1. - have /[!disjoint_subset] /subsetP {}S12 := S12. +- have /[!disjoint_subset] /subsetP {}S12 := S12. by rewrite !(out_perm pS2) //; apply: S12; rewrite // perm_closed. case/boolP : (t \in S2) => tS2. - have /[1!disjoint_sym] /[!disjoint_subset] /subsetP {}S12 := S12. +- have /[1!disjoint_sym] /[!disjoint_subset] /subsetP {}S12 := S12. by rewrite !(out_perm pS1) //; apply: S12; rewrite // perm_closed. by rewrite (out_perm pS1) // (out_perm pS2) // (out_perm pS1). Qed. @@ -57,43 +48,15 @@ Qed. Lemma tperm_on {T : finType} (x y : T) : perm_on [set x; y] (tperm x y). Proof. -by apply/subsetP => z /[!inE]; case: tpermP => [->|->|]; rewrite eqxx // orbT. -Qed. - -(* TODO added to fcsl-pcm *) -Lemma slice_oPR {A : Type} a x (s : seq A) : - 0 < x -> - &:s (Interval a (BRight x.-1)) = &:s (Interval a (BLeft x)). -Proof. by move=>Hx; rewrite -{2}(prednK Hx) slice_oSR. Qed. - -Lemma slice_FR {A : Type} (s : seq A) x : - &:s (Interval x +oo) = &:s (Interval x (BLeft (size s))). -Proof. by rewrite /slice.slice /= addn0. Qed. - -Lemma slice_extrude {A : Type} (s : seq A) (i : interval nat) : - order.Order.lt i.1 i.2 -> - s = &:s (Interval -oo i.1) ++ &:s i ++ &:s (Interval i.2 +oo). -Proof. -case: i=>[[[] i|[]][[] j|[]]] //=; rewrite bnd_simp=>H; - rewrite ?itv_minfR ?itv_pinfL /= ?cats0. -- by rewrite -{1}(slice_uu s) (slice_split _ true (x:=i)) //= - (slice_split _ true (x:=j) (i:=`[i, +oo[)) //= in_itv /= andbT; apply/ltnW. -- by rewrite -{1}(slice_uu s) (slice_split _ true (x:=i)) //= - (slice_split _ false (x:=j) (i:=`[i, +oo[)) //= in_itv /= andbT. -- by rewrite -{1}(slice_uu s) (slice_split _ true (x:=i)). -- by rewrite -{1}(slice_uu s) (slice_split _ false (x:=i)) //= - (slice_split _ true (x:=j) (i:=`]i, +oo[)) //= in_itv /= andbT. -- by rewrite -{1}(slice_uu s) (slice_split _ false (x:=i)) //= - (slice_split _ false (x:=j) (i:=`]i, +oo[)) //= in_itv /= andbT. -- by rewrite -{1}(slice_uu s) (slice_split _ false (x:=i)). -- by rewrite -{1}(slice_uu s) (slice_split _ true (x:=j)). -- by rewrite -{1}(slice_uu s) (slice_split _ false (x:=j)). -by rewrite slice_uu. +apply/subsetP => z /[!inE]; case: tpermP => [->|->|]; +by rewrite eqxx // orbT. Qed. (****) -Lemma leq_choose a b : a < b -> sumbool (a.+1 == b) (a.+1 < b). +Lemma leq_choose a b : + a < b -> + sumbool (a.+1 == b) (a.+1 < b). Proof. move=>H. case: (decP (b:=a.+1 < b) idP)=>[H2|/negP H2]; first by right. @@ -101,7 +64,6 @@ by left; rewrite eqn_leq H /= leqNgt. Qed. (* TODO copied from bubble *) -Section OrdArith. Fact ord_trans {n} (j : 'I_n) (i : 'I_n) (Hi : i < j) : (i.+1 < n)%N. Proof. @@ -109,20 +71,27 @@ case: j Hi=>j Hj /= Hi. by apply/leq_trans/Hj. Qed. +Section OrdArith. + (* increase by 1 within the bound *) Definition Sbo n (i : 'I_n) (prf : (i.+1 < n)%N) : 'I_n. Proof. case: i prf=>/= m Hm; apply: Ordinal. Defined. -Lemma Sbo_eq n (i : 'I_n) (prf : (i.+1 < n)%N) : nat_of_ord (Sbo prf) = i.+1. +Lemma Sbo_eq n (i : 'I_n) (prf : (i.+1 < n)%N) : + nat_of_ord (Sbo prf) = i.+1. Proof. by case: i prf. Qed. -Lemma Sbo_lift n (i j : 'I_n) (H1 : i < j) : i.+1 < j -> Sbo (ord_trans H1) < j. +Lemma Sbo_lift n (i j : 'I_n) (H1 : i < j) : + i.+1 < j -> + Sbo (ord_trans H1) < j. Proof. by case: i H1. Qed. -Lemma Sbo_leq n (i j k : 'I_n) (H1 : i <= j) (H2 : j < k) : Sbo (ord_trans (leq_ltn_trans H1 H2)) <= Sbo (ord_trans H2). +Lemma Sbo_leq n (i j k : 'I_n) (H1 : i <= j) (H2 : j < k) : + Sbo (ord_trans (leq_ltn_trans H1 H2)) <= Sbo (ord_trans H2). Proof. by case: j H1 H2; case: i. Qed. -Lemma Sbo_lt n (i j k : 'I_n) (H1 : i <= j) (H2 : j < k) : i <= Sbo (ord_trans H2). +Lemma Sbo_lt n (i j k : 'I_n) (H1 : i <= j) (H2 : j < k) : + i <= Sbo (ord_trans H2). Proof. by case: j H1 H2; case: i=>/=x Hx y Hy Hxy Hyk; apply/ltnW. Qed. (* increase by 1 with saturation *) @@ -133,7 +102,8 @@ case: i=>/= m Hm; case: (ltnP m.+1 n)=>[H|_]. by apply: (@Ordinal _ m Hm). Defined. -Lemma Sso_eq n (i : 'I_n) : nat_of_ord (Sso i) = if (i.+1 < n)%N then i.+1 else i. +Lemma Sso_eq n (i : 'I_n) : + nat_of_ord (Sso i) = if (i.+1 < n)%N then i.+1 else i. Proof. by case: i=>/= m prf; case: ltnP. Qed. (* decrease by 1 *) @@ -146,9 +116,10 @@ Proof. by case: i. Qed. End OrdArith. Section PermFgraph. -Variable (n : nat) (A : Type). +Variables (n : nat) (A : Type). -Lemma perm_on_notin (f : {ffun 'I_n -> A}) (p : 'S_n) (s : {set 'I_n}) (i : interval nat) : +Lemma perm_on_notin (f : {ffun 'I_n -> A}) (p : 'S_n) + (s : {set 'I_n}) (i : interval nat) : perm_on s p -> [disjoint s & [set x : 'I_n | (x : nat) \in i]] -> &:(fgraph (pffun p f)) i = &:(fgraph f) i. @@ -172,9 +143,9 @@ case: j=>[[] jx|[]]; case: i=>[[] ix|[]]; by rewrite leqNgt (ltn_ord y). Qed. -Corollary tperm_notin (f : {ffun 'I_n -> A}) (x y : 'I_n) (i : interval nat) : - (x : nat) \notin i -> (y : nat) \notin i -> - &:(fgraph (pffun (tperm x y) f)) i = &:(fgraph f) i. +Lemma tperm_notin (f : {ffun 'I_n -> A}) (x y : 'I_n) (i : interval nat) : + (x : nat) \notin i -> (y : nat) \notin i -> + &:(fgraph (pffun (tperm x y) f)) i = &:(fgraph f) i. Proof. move=>Hx0 Hx1. apply: perm_on_notin; first by exact: tperm_on. @@ -185,11 +156,11 @@ Qed. End PermFgraph. Section PermFgraphEq. -Variable (n : nat) (A : eqType). +Variables (n : nat) (A : eqType). Lemma perm_fgraph (p : 'S_n) (f : {ffun 'I_n -> A}) : - perm_eq (fgraph (pffun p f)) - (fgraph f). + perm_eq (fgraph (pffun p f)) + (fgraph f). Proof. apply/tuple_permP. exists (cast_perm (esym (card_ord n)) p). @@ -199,9 +170,9 @@ by rewrite tnth_fgraph tnth_map tnth_fgraph ffunE /= tnth_ord_tuple Qed. Lemma perm_on_fgraph (i : interval nat) (p : 'S_n) (f : {ffun 'I_n -> A}) : - perm_on [set x : 'I_n | (x : nat) \in i] p -> - perm_eq &:(fgraph (pffun p f)) i - &:(fgraph f ) i. + perm_on [set x : 'I_n | (x : nat) \in i] p -> + perm_eq &:(fgraph (pffun p f)) i + &:(fgraph f ) i. Proof. case: i=>i j H. case/boolP: (order.Order.lt i j)=>[Hij|]; last first. @@ -231,6 +202,13 @@ Qed. End PermFgraphEq. + +(*****************) +(*****************) +(* Verifications *) +(*****************) +(*****************) + Section Lomuto. Variable (n : nat) (A : ordType). @@ -273,8 +251,6 @@ Variable (n : nat) (A : ordType). (* go 0 (size a)-1 *) (***************************************************) -Opaque Array.write Array.read. - Program Definition swap (a : {array 'I_n -> A}) (i j : 'I_n) : STsep {f : {ffun 'I_n -> A}} (Array.shape a f, @@ -290,15 +266,15 @@ move=>a i j /= [f][] h /= H. case: ifP=>[/eqP->|Hij]. - by step=>_; rewrite tperm1 pffunE1. do 2!apply: [stepE f, h]=>//= _ _ [->->]. -apply: [stepE f]=>//= _ _ [-> V1]; set f1 := finfun [eta f with i |-> f j]. -apply: [gE f1]=>//= _ _ [-> V2]; set f2 := finfun [eta f1 with j |-> f i]. +apply: [stepE f]=>//= _ _ [-> V1]. +set f1 := finfun [eta f with i |-> f j]. +apply: [gE f1]=>//= _ _ [-> V2]. +set f2 := finfun [eta f1 with j |-> f i]. move=>{h H}_; split=>//=; suff {V1 V2}: f2 = pffun (tperm i j) f by move=>->. rewrite {}/f2 {}/f1; apply/ffunP=>/= x; rewrite !ffunE /= ffunE /=. by case: tpermP=>[->|->|/eqP/negbTE->/eqP/negbTE->] {x}//; rewrite eqxx // Hij. Qed. -Opaque swap. - Definition partition_lm_loop (a : {array 'I_n -> A}) (pivot : A) (lo hi : 'I_n) := forall ij : sigT (fun i : 'I_n => sig (fun j : 'I_n => i <= j /\ j < hi)), @@ -329,7 +305,8 @@ Program Definition partition_lm_pass (a : {array 'I_n -> A}) (pivot : A) all (oleq^~ pivot) (&:(fgraph f') `[lo : nat, v : nat[) & all (ord pivot) (&:(fgraph f') `[v : nat, hi : nat[)]]) := Do (let go := - ffix (fun (loop : partition_lm_loop a pivot lo hi) '(existT i (exist j (conj Hi Hj))) => + ffix (fun (loop : partition_lm_loop a pivot lo hi) + '(existT i (exist j (conj Hi Hj))) => Do (x <-- Array.read a j; if oleq x pivot then swap a i j;; @@ -349,11 +326,11 @@ Next Obligation. move=>a pivot lo hi Hlo loop _ i _ j _ Hi Hj [f][] h /= [H Oli Ai Aj]. apply: [stepE f, h]=>//= _ _ [->->]. case: oleqP=>Hfp. -- (* a[j] <= pivot, make swap *) - apply: [stepE f]=>//= _ m Hm. +(* a[j] <= pivot, make swap *) +- apply: [stepE f]=>//= _ m Hm. case: (leq_choose Hj)=>Hj1. - - (* j+1 = hi, exit *) - step=>_; split. + (* j+1 = hi, exit *) + - step=>_; split. - rewrite Sbo_eq; apply/andP; split=>//. by apply/leq_ltn_trans/Hj. exists (tperm i j); rewrite Sbo_eq; split=>//. @@ -374,7 +351,7 @@ case: oleqP=>Hfp. rewrite all_rcons; apply/andP; split. - by rewrite ffunE tpermR. by rewrite tperm_notin ?slice_oSL // - in_itv /= negb_and leEnat ltEnat /= -!leqNgt leqnn // orbT. + in_itv /= negb_and leEnat ltEnat /= -!leqNgt leqnn // orbT. (* j+1 < hi, loop *) apply: [gE (pffun (tperm i j) f)]=>//=. - split=>//; rewrite !Sbo_eq; first by apply/ltnW. @@ -420,8 +397,6 @@ move=>/= a pivot lo hi Hlo [f][] i /= H. by apply: [gE f]=>//=; split=>//; rewrite slice_kk. Qed. -Opaque partition_lm_pass. - Program Definition partition_lm (a : {array 'I_n -> A}) (lo hi : 'I_n) (Hlo : lo < hi): STsep {f : {ffun 'I_n -> A}} @@ -472,8 +447,6 @@ End Lomuto. Section LomutoQsort. Variable (n : nat) (A : ordType). -Opaque partition_lm. - Definition quicksort_lm_loop (a : {array 'I_n.+1 -> A}) := forall (lohi : 'I_n.+1 * 'I_n.+1), let lo := lohi.1 in @@ -500,7 +473,8 @@ Program Definition quicksort_lm (a : {array 'I_n.+1 -> A}) : loop (l, Po v);; (* we use saturating increment to stay under n+1 *) (* and keep the classical form of the algorithm *) - (* the overflow case will exit on next call because v = h = n-1 *) + (* the overflow case will exit on next call *) + (* because v = h = n-1 *) loop (Sso v, h))) in go (ord0, ord_max)). Next Obligation. @@ -528,10 +502,10 @@ exists (pr * (pl * p))%g; split=>//. case/andP=>->/= Hz. apply/leq_trans/Hvh/(leq_trans Hz). by exact: leq_pred. -(* we need to handle two edge cases: v=0 and v=n *) +(* need to handle two edge cases: v=0 and v=n *) case: (eqVneq v ord0)=>[Ev|Nv0]. -- (* if v=0, then l=0 and left partition is empty *) - have El: l = ord0. +(* if v=0, then l=0 and left partition is empty *) +- have El: l = ord0. - by move: Hvl; rewrite Ev leqn0 => /eqP El; apply/ord_inj. rewrite Ev El /= in Hpl. have Epl: pl = 1%g. @@ -550,8 +524,8 @@ case: (eqVneq v ord0)=>[Ev|Nv0]. rewrite pffunEM perm_sym -!slice_oSL (_ : 0 = (ord0 : 'I_n.+1)) //. by apply: perm_on_fgraph. move: (ltn_ord v); rewrite ltnS leq_eqVlt; case/orP=>[/eqP Ev|Nv]. -- (* if v=n, then h=n and right partition is empty *) - have Eh: (h : nat) = n. +(* if v=n, then h=n and right partition is empty *) +- have Eh: (h : nat) = n. - apply/eqP; rewrite eqn_leq; move: Hvh; rewrite Ev=>->; rewrite andbT. by move: (ltn_ord h); rewrite ltnS. rewrite Ev Eh /= ltnn in Hpr. @@ -563,16 +537,18 @@ move: (ltn_ord v); rewrite ltnS leq_eqVlt; case/orP=>[/eqP Ev|Nv]. move: Sl Hpl; rewrite Eh Ev Epr mul1g => Sl Hpl. rewrite slice_xR; last by rewrite bnd_simp leEnat; move: Hvl; rewrite Ev. rewrite {22}(_ : n = (ord_max : 'I_n.+1)) // onth_codom /= sorted_rconsE //=. - move: Sl; rewrite slice_oPR; last by rewrite lt0n -Ev. + move: Sl; rewrite slice_oPR /order.Order.lt/= lt0n -{1}Ev Nv0. move=>->; rewrite andbT; move: Al; rewrite Ev. have ->: pffun (pl * p) f ord_max = pffun p f ord_max. - rewrite !ffunE permM (out_perm Hpl) // inE negb_and -!ltnNge /=. by rewrite ltn_predL lt0n -{3}Ev Nv0 orbT. have ->: v = ord_max by apply/ord_inj. - rewrite [in X in X -> _](perm_all (s2:=&:(codom (pffun (pl * p) f)) `[l: nat, n[)) //. + rewrite [in X in X -> _] + (perm_all (s2:=&:(codom (pffun (pl * p) f)) `[l: nat, n[)) //. rewrite pffunEM perm_sym. rewrite {8 15}(_ : n = (ord_max : 'I_n.+1)) //; apply: perm_on_fgraph. - apply/(subset_trans Hpl)/subsetP=>/= z; rewrite 2!inE in_itv /= leEnat ltEnat /=. + apply/(subset_trans Hpl)/subsetP=>/= z. + rewrite 2!inE in_itv /= leEnat ltEnat /=. by case/andP=>->/= Hz; apply: (leq_ltn_trans Hz); rewrite ltn_predL lt0n -Ev. (* the general case *) rewrite Nv in Hpr Sr. @@ -614,22 +590,22 @@ apply/and5P; split=>//. - move/allP: Al=>/(_ x); apply. move: Hx; congr (_ = _); move: x; apply: perm_mem. rewrite pffunEM; apply: perm_on_fgraph. - apply/(subset_trans Hpl)/subsetP=>/= z; rewrite 2!inE in_itv /= leEnat ltEnat /=. + apply/(subset_trans Hpl)/subsetP=>/= z. + rewrite 2!inE in_itv /= leEnat ltEnat /=. by case/andP=>->/= Hz; apply: (leq_ltn_trans Hz); rewrite ltn_predL lt0n. apply/ordW; move/allP: Ah=>/(_ y); apply. move: Hy; congr (_ = _); move: y; apply: perm_mem. by rewrite pffunEM; apply: perm_on_fgraph. -- by rewrite slice_oPR // in Sl; rewrite lt0n. +- by rewrite slice_oPR /order.Order.lt/= lt0n Nv0 in Sl. have HS: subpred (ord (pffun p f v)) (oleq (pffun p f v)). - by move=>z /ordW. move/(sub_all HS): Ah; congr (_ = _); apply/esym/perm_all. by rewrite pffunEM; apply/perm_on_fgraph. Qed. Next Obligation. -move=>a [f][] i /= H. -apply: [gE f]=>//= _ m [p][Hp Hm Hs] _. -exists p; split=>//. -by rewrite -(slice_uu (codom _)) slice_0L slice_FR size_codom card_ord slice_oSR. +move=>a [f][] i /= H; apply: [gE f]=>//= _ m [p][Hp Hm Hs] _. +exists p; split=>//; rewrite -(slice_uu (codom _)) slice_0L. +by rewrite slice_FR size_codom card_ord slice_oSR. Qed. End LomutoQsort. diff --git a/examples/stack.v b/examples/stack.v index f52dae6..dee4aed 100644 --- a/examples/stack.v +++ b/examples/stack.v @@ -26,8 +26,6 @@ Section Stack. Variable T : Type. Notation stack := (stack T). -Opaque insert head remove. - (* stack is a pointer to a singly-linked list *) Definition shape s (xs : seq T) := [Pred h | exists p h', [ /\ valid (s :-> p \+ h'), diff --git a/examples/union_find.v b/examples/union_find.v index e8eae48..b86f834 100644 --- a/examples/union_find.v +++ b/examples/union_find.v @@ -593,8 +593,6 @@ case=>m [] h [ts] [->-> V]. step=>p. do !step. move=>V2. by exists (TNode p nil :: ts); split. Qed. -Opaque newT. - (********) (* FIND *) (********) @@ -625,8 +623,6 @@ apply: [gE fset ts, r]=>//=. by do !split=>//=; exists ts; split. Qed. -Opaque find1. - (*********) (* UNION *) (*********) @@ -702,9 +698,6 @@ rewrite inE => /eqP X4; move/negP: X2; apply; apply/eqP. by apply: flay_mem_eq X1 B J=>//; rewrite X4 rt_in. Qed. -Opaque union. - - (*********) (* Tests *) (*********) diff --git a/htt/model.v b/htt/model.v index c7f4ff8..7f96384 100644 --- a/htt/model.v +++ b/htt/model.v @@ -141,7 +141,6 @@ Structure STspec G A (s : spec G A) := STprog { Arguments STspec G [A] s. (* some notation *) -Notation "'Do' e" := (@STprog _ _ _ e _) (at level 80). Notation "x '<--' c1 ';' c2" := (bnd c1 (fun x => c2)) (at level 81, right associativity, format "'[v' x '<--' c1 ';' '//' c2 ']'"). @@ -152,6 +151,11 @@ Notation "c1 ';;' c2" := (bnd c1 (fun _ => c2)) (at level 81, right associativity). Notation "'!' x" := (read x) (at level 50). Notation "x '::=' e" := (write x e) (at level 60). +(* using locked for Do, to make the Do definitions opaque *) +(* otherwise, clients will inline the code of Do, preventing *) +(* the use of the program with the spec that has been verified *) +(* with Do. *) +Notation "'Do' e" := (locked (@STprog _ _ _ e _)) (at level 80). (* Fixed point constructor *) (* We shall make fix work over *monotone closure* of argument function. *) @@ -167,7 +171,6 @@ Parameter ffix : forall G A (B : A -> Type) (s : forall x : A, spec G (B x)), forall x : A, STspec G (s x). End VrfSig. - (********************************) (* Definition of the Hoare type *) (********************************) From e666696da6e0d77d3fae14ab8a08498b5855e6e2 Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Mon, 12 Aug 2024 20:18:01 +0200 Subject: [PATCH 37/93] blah --- core/ordtype.v | 15 +++++++-------- examples/bst.v | 4 ++++ 2 files changed, 11 insertions(+), 8 deletions(-) diff --git a/core/ordtype.v b/core/ordtype.v index 417f0f1..1a83f23 100644 --- a/core/ordtype.v +++ b/core/ordtype.v @@ -67,7 +67,7 @@ Global Arguments ordtype_isOrdered_mixin [_] _. Section type. Local Unset Implicit Arguments. -(* Polymorphic Cumulative *) Record +Polymorphic Cumulative Record type : Type := Pack { sort : Type; _ : axioms_ sort; }. End type. @@ -89,17 +89,16 @@ Global Arguments class : clear implicits. (* Polymorphic Universe annotations added *) Section PolymorphicClonePack. -(* Polymorphic Universe ou. *) -(* Polymorphic Variables (T : Type@{ou}) (cT : type@{ou}). *) -Variables (T : Type) (cT : type). +Polymorphic Universe ou. +Polymorphic Variables (T : Type@{ou}) (cT : type@{ou}). -(* Polymorphic *) Definition phant_clone : forall (c : axioms_ T), +Polymorphic Definition phant_clone : forall (c : axioms_ T), unify Type Type T (sort cT) nomsg -> unify type type cT (Pack (sort:=T) c) nomsg -> type := fun (c : axioms_ T) => fun=> (fun=> (Pack (sort:=T) c)). -(* Polymorphic *) Definition pack_ := fun (m : Equality.mixin_of T) +Polymorphic Definition pack_ := fun (m : Equality.mixin_of T) (m0 : isOrdered.axioms_ T m) => Pack (sort:=T) {|eqtype_hasDecEq_mixin := m; ordtype_isOrdered_mixin := m0 |}. @@ -115,7 +114,7 @@ Notation ordType := Ordered.type. #[reversible] Coercion sort : Ordered.type >-> Sortclass. (* Polymorphic annotation added *) -(* Polymorphic *) Definition ordtype_Ordered_class__to__eqtype_Equality_class : +Polymorphic Definition ordtype_Ordered_class__to__eqtype_Equality_class : forall T : Type, axioms_ T -> Equality.axioms_ T := fun (T : Type) (c : axioms_ T) => {| Equality.eqtype_hasDecEq_mixin := eqtype_hasDecEq_mixin c |}. @@ -126,7 +125,7 @@ Local Arguments ordtype_Ordered_class__to__eqtype_Equality_class : #[reversible] Coercion ordtype_Ordered_class__to__eqtype_Equality_class : ordtype.Ordered.axioms_ >-> eqtype.Equality.axioms_. -(* Polymorphic *) Definition ordtype_Ordered__to__eqtype_Equality : +Polymorphic Definition ordtype_Ordered__to__eqtype_Equality : ordType -> eqType := fun s : ordType => {| Equality.sort := s; Equality.class := class s |}. diff --git a/examples/bst.v b/examples/bst.v index 3deedab..01cab5d 100644 --- a/examples/bst.v +++ b/examples/bst.v @@ -17,6 +17,10 @@ From pcm Require Import options axioms pred ordtype seqext. From pcm Require Import pcm unionmap heap autopcm automap. From htt Require Import options model heapauto bintree. +(* hack to avoid "_ *p _" notation clash *) +From mathcomp Require order. +Import order.Order.NatOrder order.Order.TTheory. + Section BST. Context {T : ordType}. From b7bb9b9c722e20bf1882959d6fe1609d99f38ee6 Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Mon, 12 Aug 2024 20:50:25 +0200 Subject: [PATCH 38/93] blah --- core/pred.v | 36 +++++++++---------- examples/bst.v | 4 --- examples/congmath.v | 84 ++++++++++++++++++++++----------------------- 3 files changed, 60 insertions(+), 64 deletions(-) diff --git a/core/pred.v b/core/pred.v index fc711ee..55ed284 100644 --- a/core/pred.v +++ b/core/pred.v @@ -122,9 +122,9 @@ Arguments Pred0 {T}. Arguments PredT {T}. Prenex Implicits Pred0 PredT PredI PredU PredC PredD Preim. -Notation "r1 +p r2" := (PredU r1 r2 : Pred _) +Notation "r1 \+p r2" := (PredU r1 r2 : Pred _) (at level 55, right associativity) : rel_scope. -Notation "r1 *p r2" := (xPredI r1 r2 : Pred _) +Notation "r1 \*p r2" := (xPredI r1 r2 : Pred _) (at level 45, right associativity) : rel_scope. Notation "[ 'Pred' : T | E ]" := (SimplPred (fun _ : T => E)) @@ -517,37 +517,37 @@ Section RelLaws. Variable T : Type. Implicit Type r : Pred T. -Lemma orrI r : r +p r <~> r. +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. +Lemma orrC r1 r2 : r1 \+p r2 <~> r2 \+p r1. Proof. move=>x; split=>/=; tauto. Qed. -Lemma orr0 r : r +p Pred0 <~> r. +Lemma orr0 r : r \+p Pred0 <~> r. Proof. by move=>x; split; [case | left]. Qed. -Lemma or0r r : Pred0 +p r <~> r. +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. +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. +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. +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. +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. +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. @@ -563,28 +563,28 @@ 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. +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. +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. +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. +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. +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. +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. diff --git a/examples/bst.v b/examples/bst.v index 01cab5d..3deedab 100644 --- a/examples/bst.v +++ b/examples/bst.v @@ -17,10 +17,6 @@ From pcm Require Import options axioms pred ordtype seqext. From pcm Require Import pcm unionmap heap autopcm automap. From htt Require Import options model heapauto bintree. -(* hack to avoid "_ *p _" notation clash *) -From mathcomp Require order. -Import order.Order.NatOrder order.Order.TTheory. - Section BST. Context {T : ordType}. diff --git a/examples/congmath.v b/examples/congmath.v index b235213..834b256 100644 --- a/examples/congmath.v +++ b/examples/congmath.v @@ -229,7 +229,7 @@ Add Morphism closure with Proof. by apply: closE. Qed. Lemma clos_clos (R Q : rel_exp) : - closure (closure R +p Q) <~> closure (R +p Q). + closure (closure R \+p Q) <~> closure (R \+p Q). Proof. case=>e1 e2; split=>H1 C H2 H3; apply: H1 =>// [[x y]] H4. - by case: H4=>H4; [apply: H4=>// [[x1 y1]] H4|]; apply: H2; [left | right]. @@ -243,7 +243,7 @@ Proof. by rewrite -(orr0 (closure R)) clos_clos !orr0. Qed. Lemma closKR (R1 R2 K : rel_exp) : closure R1 <~> closure R2 -> - closure (K +p R1) <~> closure (K +p R2). + closure (K \+p R1) <~> closure (K \+p R2). Proof. move=>H. by rewrite 2!(orrC K) -(clos_clos R1) -(clos_clos R2) H. @@ -251,7 +251,7 @@ Qed. Lemma closRK (R1 R2 K : rel_exp) : closure R1 <~> closure R2 -> - closure (R1 +p K) <~> closure (R2 +p K). + closure (R1 \+p K) <~> closure (R2 \+p K). Proof. by rewrite 2!(orrC _ K); apply: closKR. Qed. Lemma closI (R : rel_exp) : R ~> closure R. @@ -259,7 +259,7 @@ Proof. by move=>[x y] H1 T H2 H3; apply: H2. Qed. (* absorption and closure *) Lemma closKA (K R : rel_exp) : - closure R <~> closure (K +p R) <-> K ~> closure R. + closure R <~> closure (K \+p R) <-> K ~> closure R. Proof. rewrite -orrAb; split=>[->|]. - by rewrite orrAb; move=>[x y] H; apply: closI; left. @@ -267,12 +267,12 @@ by rewrite (orrC _ R) -clos_clos orrC=><-; rewrite clos_idemp. Qed. Lemma closAK (K R : rel_exp) : - closure R <~> closure (R +p K) <-> K ~> closure R. + closure R <~> closure (R \+p K) <-> K ~> closure R. Proof. by rewrite orrC closKA. Qed. Lemma clos_or (R1 R2 R : rel_exp) : closure R1 ~> closure R -> closure R2 ~> closure R -> - closure (R1 +p R2) ~> closure R. + closure (R1 \+p R2) ~> closure R. Proof. rewrite -!closAK !(orrC R) !clos_clos -!(orrC R) => H1 H2. by rewrite H1 -clos_clos H2 clos_clos orrAC orrA. @@ -280,7 +280,7 @@ Qed. Lemma sub_closKR (R1 R2 K : rel_exp) : (closure R1 ~> closure R2) -> - closure (K +p R1) ~> closure (K +p R2). + closure (K \+p R1) ~> closure (K \+p R2). Proof. rewrite -!closAK=>H. rewrite -(orrC (closure _)) clos_clos orrAC -orrA orrI. @@ -291,7 +291,7 @@ Qed. Lemma sub_closRK (R1 R2 K : rel_exp) : (closure R1 ~> closure R2) -> - closure (R1 +p K) ~> closure (R2 +p K). + closure (R1 \+p K) ~> closure (R2 \+p K). Proof. by move=>H; rewrite -!(orrC K); apply: sub_closKR. Qed. Lemma sub_closI (R1 R2 : rel_exp) : @@ -340,7 +340,7 @@ Definition eq2rel (eq : Eq) : rel_exp := end. Fixpoint eqs2rel (eqs : seq Eq) : rel_exp := - if eqs is hd::tl then eq2rel hd +p eqs2rel tl else Pred0. + if eqs is hd::tl then eq2rel hd \+p eqs2rel tl else Pred0. (* Coercing triples of symbols into equations *) Definition symb2eq (t : symb*symb*symb) := @@ -441,7 +441,7 @@ Hint Resolve rep_in_reps : core. (* The symbols and their representatives are a base case in defining the closure *) Definition rep2rel D := graph (fun x => const (rep D x)). -Lemma clos_rep D R a : (const a, const (rep D a)) \In closure (rep2rel D +p R). +Lemma clos_rep D R a : (const a, const (rep D a)) \In closure (rep2rel D \+p R). Proof. apply: (@closI _ (const a, const (rep D a))); apply: sub_orl; rewrite /= /rep2rel /graph /=. by rewrite mem_map /= ?mem_enum // => x1 x2 [->]. @@ -471,8 +471,8 @@ by exists a, b, c, c1, c2. Qed. (* The relation defined by the data structure is the following *) -Definition CRel D := closure (rep2rel D +p - lookup_rel D +p +Definition CRel D := closure (rep2rel D \+p + lookup_rel D \+p eqs2rel (map pend2eq (pending D))). Lemma cong_rel D : congruence (CRel D). @@ -480,7 +480,7 @@ Proof. by move=>*; apply: cong_clos. Qed. Hint Resolve cong_rel : core. -Lemma clos_rel D R a : (const a, const (rep D a)) \In closure (CRel D +p R). +Lemma clos_rel D R a : (const a, const (rep D a)) \In closure (CRel D \+p R). Proof. by rewrite /CRel clos_clos orrA; apply: clos_rep. Qed. @@ -706,12 +706,12 @@ Definition lookup_inv D := (* two symbols are similar if they are either related by representatives *) (* or are to be related after the pending list is processed *) Definition similar D a b := - (const a, const b) \In closure (rep2rel D +p eqs2rel + (const a, const b) \In closure (rep2rel D \+p eqs2rel (map pend2eq (pending D))). Definition similar1 D a' b' a b := (const a, const b) \In - closure (rep2rel D +p [Pred x y | x = const a' /\ y = const b'] +p + closure (rep2rel D \+p [Pred x y | x = const a' /\ y = const b'] \+p eqs2rel (map pend2eq (pending D))). (* invariants tying use lists with the lookup table *) @@ -830,9 +830,9 @@ Lemma symR R x y : (x, y) \In closure R <-> (y, x) \In closure R. Proof. by split=>T; apply: symC. Qed. Lemma app_rep D R a b x : - (x, app (const a) (const b)) \In closure (rep2rel D +p R) <-> + (x, app (const a) (const b)) \In closure (rep2rel D \+p R) <-> (x, app (const (rep D a)) (const (rep D b))) - \In closure (rep2rel D +p R). + \In closure (rep2rel D \+p R). Proof. split=>T. - apply: (transC (y:= app (const a) (const (rep D b)))); last first. @@ -846,18 +846,18 @@ by apply: monoC; [apply: symC; apply: clos_rep | apply: reflC]. Qed. Lemma app_rel D R a b x : - (x, app (const a) (const b)) \In closure (CRel D +p R) <-> - (x, app (const (rep D a)) (const (rep D b))) \In closure (CRel D +p R). + (x, app (const a) (const b)) \In closure (CRel D \+p R) <-> + (x, app (const (rep D a)) (const (rep D b))) \In closure (CRel D \+p R). Proof. by rewrite /CRel !clos_clos !orrA; apply: app_rep. Qed. Lemma const_rep D R a x : - (const a, x) \In closure (rep2rel D +p R) <-> - (const (rep D a), x) \In closure (rep2rel D +p R). + (const a, x) \In closure (rep2rel D \+p R) <-> + (const (rep D a), x) \In closure (rep2rel D \+p R). Proof. by split=>T; apply: transC T; [apply: symC|]; apply: clos_rep. Qed. Lemma const_rel D R a x : - (const a, x) \In closure (CRel D +p R) <-> - (const (rep D a), x) \In closure (CRel D +p R). + (const a, x) \In closure (CRel D \+p R) <-> + (const (rep D a), x) \In closure (CRel D \+p R). Proof. by rewrite /CRel !clos_clos !orrA; apply: const_rep. Qed. (* Now the lemmas *) @@ -888,7 +888,7 @@ Qed. Lemma join_class_repE (D : data) (a' b' : symb) : a' \in reps D -> b' \in reps D -> rep_idemp D -> closure (rep2rel (join_class D a' b')) <~> - closure (rep2rel D +p [Pred x y | x = const a' /\ y = const b']). + closure (rep2rel D \+p [Pred x y | x = const a' /\ y = const b']). Proof. move=>H2 H3 H1 [x y]; split=>/= T. - pose ty z := PredU (rep2rel D) [Pred x0 y0 | x0 = const z /\ y0 = const b']. @@ -1036,14 +1036,14 @@ Lemma join_classE (D : data) (a' b' : symb) : rep_idemp D -> use_lookup_inv1 D a' b' -> lookup_use_inv1 D a' b' -> - closure (CRel (join_class D a' b') +p + closure (CRel (join_class D a' b') \+p eqs2rel (map symb2eq (use D a'))) <~> - closure (CRel D +p [Pred x y | x = const a' /\ y = const b']). + closure (CRel D \+p [Pred x y | x = const a' /\ y = const b']). Proof. pose ty := [Pred x0 y0 | x0 = const a' /\ y0 = const b']. move=>L1 L2 L3 H1 H4 H5 [x y]; split. - apply: clos_or; move=>{x y}[x y]; last first. - - move: (clos_idemp (CRel D +p ty) (x,y))=><-. + - move: (clos_idemp (CRel D \+p ty) (x,y))=><-. apply: sub_closI=>{x y}[[x y]] /=. move/invert=>[c][c1][c2][->->H6]. move: (H4 a' c c1 c2 L1 H6)=>[d][d1][d2][Q1][Q2][Q3] Q4. @@ -1194,23 +1194,23 @@ Lemma join_usePE (D : data) (a' b' : symb) : lookup_inv (join_use D a' b') /\ use_lookup_inv2 (join_use D a' b') a' b' /\ lookup_use_inv2 (join_use D a' b') a' b' /\ - closure (CRel D +p eqs2rel (map symb2eq (use D a'))) <~> + closure (CRel D \+p eqs2rel (map symb2eq (use D a'))) <~> CRel (join_use D a' b'). Proof. rewrite /join_use; move E: (use _ _)=>x. elim: x D E=>[|[[c c1] c2] ss IH] D E L1 L2 H1 H2 H3 H4 H5 /=. - by rewrite orr0 /CRel clos_idemp. case F: (fnd _ _)=>[[[d d1] d2]|]; set D' := Data _ _ _ _ _. -- have S1: closure (rep2rel D' +p eqs2rel (map pend2eq (pending D'))) <~> - closure (rep2rel D +p [Pred x y | x = const c /\ y = const d] +p +- have S1: closure (rep2rel D' \+p eqs2rel (map pend2eq (pending D'))) <~> + closure (rep2rel D \+p [Pred x y | x = const c /\ y = const d] \+p eqs2rel (map pend2eq (pending D))) by []. have S2: forall e1 e2, similar D e1 e2 -> similar D' e1 e2. - rewrite /similar=>e1 e2; move: (const e1) (const e2)=>{e1 e2} x y. by rewrite S1; apply: sub_closI=>{x y} [[x y]]; case; [left | right; right]. - have T2 : closure (CRel D' +p eqs2rel (map symb2eq ss)) <~> - closure (CRel D +p [Pred x y | x = const c /\ - y = app (const c1) (const c2)] +p eqs2rel (map symb2eq ss)). + have T2 : closure (CRel D' \+p eqs2rel (map symb2eq ss)) <~> + closure (CRel D \+p [Pred x y | x = const c /\ + y = app (const c1) (const c2)] \+p eqs2rel (map symb2eq ss)). - rewrite /CRel {2 3}/D' /= !clos_clos !orrA. rewrite -!(rpull (eqs2rel (map pend2eq _))). rewrite -!(rpull (eqs2rel (map symb2eq ss))). @@ -1280,7 +1280,7 @@ case F: (fnd _ _)=>[[[d d1] d2]|]; set D' := Data _ _ _ _ _. by case: eqP H6 L1=>[->-> //| _] H6 L1; do !split=>//; apply: S2. exists h, h1, h2; rewrite {1}/D' /= !ffunE. by case: eqP H7 L1=>[->-> //| _] H7 L1; do !split=>//; apply: S2. -have T1: lookup_rel D' <~> lookup_rel D +p +have T1: lookup_rel D' <~> lookup_rel D \+p [Pred x y | x = app (const (rep D c1)) (const (rep D c2)) /\ y = const (rep D c)]. - rewrite /lookup_rel /D' /= /reps /= =>[[x y]]; split. @@ -1298,9 +1298,9 @@ have T1: lookup_rel D' <~> lookup_rel D +p exists (rep D c1), (rep D c2), c, c1, c2. rewrite fnd_ins !rep_in_reps. by case: eqP. -have T2: closure (CRel D' +p eqs2rel (map symb2eq ss)) <~> - closure (CRel D +p [Pred x y | x = const c /\ - y = app (const c1) (const c2)] +p eqs2rel (map symb2eq ss)). +have T2: closure (CRel D' \+p eqs2rel (map symb2eq ss)) <~> + closure (CRel D \+p [Pred x y | x = const c /\ + y = app (const c1) (const c2)] \+p eqs2rel (map symb2eq ss)). - pose ty1 := [Pred x y | x = app (const (rep D c1)) (const (rep D c2)) /\ y = const (rep D c)]. pose ty2 := [Pred x y | x = const c /\ y = app (const c1) (const c2)]. @@ -1447,8 +1447,8 @@ Proof. move: D. pose ty x0 y0 := [Pred x y | x = @const s x0 /\ y = @const s y0]. have L: forall D a b, - closure (rep2rel D +p ty a b) <~> - closure (rep2rel D +p ty (rep D a) (rep D b)). + closure (rep2rel D \+p ty a b) <~> + closure (rep2rel D \+p ty (rep D a) (rep D b)). - move=>D a b [x y]; split=>T; rewrite toPredE -clos_idemp; apply: sub_closI T=>{x y} [[x y]]; (case; first by move=>H; apply: closI; left); case=>->->; @@ -1500,7 +1500,7 @@ have L2: forall D pend_eq p' a b a' b' D' D'', CRel D <~> CRel (propagate (join_use D'' a' b')). - move=>D pend_eq p' a b a' b' D' D'' H1 H H2 H3 H4 H5 H6 IH [H7][H8][H9][H10] H11. - have T: CRel D <~> closure (CRel D' +p + have T: CRel D <~> closure (CRel D' \+p [Pred x y | x = const a' /\ y = const b']). - rewrite /CRel H1 -{2 3}H4 /= clos_clos !orrA H /=. rewrite -!(rpull (lookup_rel _)) -!(rpull (eqs2rel (map pend2eq _))). @@ -1597,7 +1597,7 @@ Lemma propagate_clos_pendP d c c1 c2 e e1 e2 : fnd (rep d c1, rep d c2) (lookup d) = Some (e, e1, e2) -> CRel (Data (rep d) (class d) (use d) (lookup d) (comp_pend (c, c1, c2) (e, e1, e2) :: pending d)) - <~> closure (CRel d +p + <~> closure (CRel d \+p [Pred x y | x = const c /\ y = app (const c1) (const c2)]). Proof. move=>PI H. @@ -1687,7 +1687,7 @@ Lemma propagate_clos_nopendP : CRel (Data (rep d) (class d) u2' (ins (rep d c1, rep d c2) (c, c1, c2) (lookup d)) (pending d)) <~> - closure (CRel d +p + closure (CRel d \+p [Pred x y | x = const c /\ y = app (const c1) (const c2)]). Proof. rewrite /CRel clos_clos orrA orrAC -!orrA; apply: closRK; rewrite !orrA. From bb322d1cbf1711a1c3911a01b4d173f6594a5624 Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Mon, 12 Aug 2024 22:17:56 +0200 Subject: [PATCH 39/93] blah --- examples/congprog.v | 801 ++++++++++++++++++++++++++++++++++++++++++++ 1 file changed, 801 insertions(+) create mode 100644 examples/congprog.v diff --git a/examples/congprog.v b/examples/congprog.v new file mode 100644 index 0000000..d0c6ecd --- /dev/null +++ b/examples/congprog.v @@ -0,0 +1,801 @@ +(*******************************) +(* Stateful congruence closure *) +(*******************************) + +From HB Require Import structures. +From Coq Require SetoidTactics. +From mathcomp Require Import ssreflect ssrbool ssrnat. +From mathcomp Require Import eqtype ssrfun div finset seq fintype finfun. +From mathcomp Require Import choice. +From pcm Require Import axioms ordtype finmap pred pcm unionmap heap. + +From htt Require Import options model heapauto llist array. +From htt Require Import kvmaps hashtab congmath. +Import Prenex Implicits. + +Definition xpull := pcm.pull. + +Ltac add_morphism_tactic := SetoidTactics.add_morphism_tactic. +Notation " R ===> R' " := (@Morphisms.respectful _ _ R R') + (right associativity, at level 55) : signature_scope. + +Variable s : seq constant. +Notation symb := (symb s). + +(* the lookup table is represented as a hash table with 10 buckets *) +Local Definition K : Set := symb * symb. +Local Definition V : Set := symb * symb * symb. + +Definition hash (n : nat) (k : K) := index k (enum (@predT K)) %% n. + +Lemma hash_ord n k : + 0 < n -> + hash n k < n. +Proof. exact: ltn_pmod. Qed. + +Definition hash10 k : 'I_10 := Ordinal (@hash_ord 10 k erefl). + +Definition LT := locked (HT V hash10). + +(* the tables relating arrays with the content of the lists *) +(* ctab is for class lists, utab is for use lists *) + +Definition llist (T : Set) : Set := ptr. + +Definition ctab := @table symb ptr (seq symb) (@lseq symb). +Definition utab := + @table symb ptr (seq (symb * symb * symb)) (@lseq (symb * symb * symb)). + +Definition n := #|{: symb}|. + +(* the empty congruence is one that only relates constant symbols to themselves *) +Definition empty_cong := closure (graph (@const s)). + +Section ShapePredicates. + +(* the algorithm starts with root pointers for the data *) + +Inductive ptrs : Type := + Ptrs of {array symb -> symb} & {array symb -> llist symb} & + {array symb -> llist (symb*symb*symb)} & KVmap.tp LT & ptr. + +Variable (r : {array symb -> symb}). +Variable (clist : {array symb -> llist symb}). +Variable (ulist : {array symb -> llist (symb*symb*symb)}). +Variable (htab : KVmap.tp LT). +Variable (p : ptr). + +(* The layout of the data structure *) + +Definition ashape D := + [Pred h | let: (d, ct, ut) := D in + valid h /\ + h \In Array.shape r (rep d : {ffun symb -> symb}) # + (Array.shape clist ct # sepit setT (ctab ct (class d))) # + (Array.shape ulist ut # sepit setT (utab ut (use d))) # + KVmap.shape htab (lookup d) # + [Pred h' | exists l, h' \In Pred1 (p :-> l) # lseq l (pending d)]]. + +Lemma ashapeD D (h : heap) : + h \In ashape D -> + valid h. +Proof. by case: D; case=>d ct ut []. Qed. + +(* +Lemma ashape_invert : invertible ashape. +Proof. +move=>[[[r1]]] c1 u1 t1 p1 ct1 ut1 [[[r2]]] c2 u2 t2 p2 ct2 ut2 h1 h2 /=. +move=>[hr1][w][->][R1][[hc1]][w'][->{w}][C1][[hu1]][w][->{w'}][U1] + [[ht1]][w'][->{w}][T1][[l1]][hl1][hp1][->{w'}][L1][P1] _ _ _ _ _. +move=>[hr2][w][->][R2][[hc2]][w'][->{w}][C2][[hu2]][w][->{w'}][U2] + [[ht2]][w'][->{w}][T2][[l2]][hl2][hp2][->{w'}][L2][P2] _ _ _ _ _ A. +case: (Array.shape_invert R1 R2 (agreeUnK A)) (A)=>/= -> -> {A}; move/agreeKUn=>A. +case: (table_invert (@lseq_invert _) C1 C2 (agreeUnK A)) (A) =>/= ->; +move/ffunP=>->-> {A}; move/agreeKUn => A. +case: (table_invert (@lseq_invert _) U1 U2 (agreeUnK A)) (A)=>/= ->; +move/ffunP=>->-> {A}; move/agreeKUn => A. +case: (KVmap.shape_invert T1 T2 (agreeUnK A)) (A)=>/= -> -> {A}; move/agreeKUn=>A. +case: (ptr_invert L1 L2 (agreeUnK A)) (A)=>/= E -> {A}; move/agreeKUn=>A. +by rewrite -{}E in P2; case: (lseq_invert P1 P2 A)=>/= -> ->. +Qed. +*) + +(* +Lemma ashape_inv D1 D2 h : h \In ashape D1 -> h \In ashape D2 -> D1 = D2. +Proof. by apply: iinv; [apply: ashape_invert | apply: ashapeD]. Qed. +*) + +Definition bshape d := + [Pred h | class_inv d /\ exists ct ut, h \In ashape (d, ct, ut)]. + +Lemma bshapeD d h : + h \In bshape d -> valid h. +Proof. by case=>_ [ct][ut]; apply: ashapeD. Qed. + +(* +Lemma bshape_invert : invertible bshape. +Proof. +move=>d1 d2 h1 h2 [_][ct1][ut1] H1 [_][ct2][ut2] H2. +by case/(ashape_invert H1 H2)=>[[->]] _ _ ->. +Qed. + +Lemma bshape_inv d1 d2 h : h \In bshape d1 -> h \In bshape d2 -> d1 = d2. +Proof. by apply: iinv; [apply: bshape_invert | apply: bshapeD]. Qed. +*) + +Definition shape (R : rel_exp s) := + [Pred h | exists d, h \In bshape d /\ propagate_inv d /\ pending d = [::] /\ + forall x, x \In CRel d <-> x \In R]. + +Lemma shapeD (R : rel_exp s) h : + h \In shape R -> + valid h. +Proof. by case=>d [H] _; apply: bshapeD H. Qed. + +(* +Lemma shape_invert R1 R2 h1 h2 : + h1 \In shape R1 -> h2 \In shape R2 -> agree h1 h2 -> R1 =r R2 /\ h1 = h2. +Proof. +move=>R1 R2 h1 h2 [d1][S1][_][_] <- [d2][S2][_][_] <- A. +by move: (bshape_invert S1 S2 A)=>[<- <-]. +Qed. + +Lemma shape_inv R1 R2 h : h \In shape R1 -> h \In shape R2 -> R1 =r R2. +Proof. +by move=>??? H; case/(shape_invert H)=>//; apply: agree_refl; apply: shapeD H. +Qed. +*) + +End ShapePredicates. + +(* +Add Parametric Morphism r clist ulist htab p : (shape r clist ulist htab p) with + signature @releq _ ===> @releq _ as shape_morph. +Proof. +move=>R1 R2 H. +split=>x [d1][H1][H2][H3] H4; exists d1; +by [rewrite -H | rewrite H]. +Qed. +*) + +(* Initialization procedure to generate the empty structure *) +Section Initialization. + + +Definition iT (clist : {array symb -> llist symb}): Type := forall k, + STsep (fun i => k <= n /\ exists f, i \In Array.shape clist f # + sepit [set c | indx c < k] (ctab f [ffun c => [:: c]]), + fun y m => y = Val tt /\ exists f, m \In Array.shape clist f # + sepit setT (ctab f [ffun c => [:: c]])). + +Definition ith {I : finType} i (pf : i < #|I|) : I. Admitted. + +Lemma indx_ith {I : finType} i (pf : i < #|I|) : indx (ith i pf) = i. +Proof. +admit. +Admitted. +(* by move=>i pf; rewrite /ith /= /indx index_uniq // -?cardE // enum_uniq. *) + +Lemma ith_indx {I : finType} (s : I) (pf : indx s < #|I|) : ith (indx s) pf = s. +Proof. Admitted. +(* by move=>s /= pf; rewrite /ith /= nth_index // mem_enum. Qed. *) + +Lemma indx_inj I : injective (@indx I). +Admitted. + +Lemma indx_injE {I : finType} s i (pf : i < #|I|) : (s == ith i pf) = (indx s == i). +Proof. +Admitted. + + +Lemma sepit_emp (A : eqType) (s : seq A) f : + (forall x, x \in s -> forall k, k \In f x <-> emp k) -> + (forall k : heap, k \In IterStar.sepit s f <-> emp k). +Proof. +Admitted. + + + +Program Definition init : + STsep (emp, fun y m => exists ptr : ptrs, y = Val ptr /\ + let: Ptrs r c u h p := ptr in + m \In shape r c u h p empty_cong) := + Do (r <-- Array.newf [ffun x : symb=> x]; + clist <-- Array.new _ null; + ffix (fun (loop : iT clist) k => + Do (if decP (b:= k < n) idP is left pf then + x <-- allocb (ith k pf) 2; + x.+1 ::= null;; + Array.write (I:=symb) clist (ith k pf) x;; + loop k.+1 + else ret tt)) 0;; + ulist <-- Array.new _ null; + htab <-- KVmap.new LT; + p <-- alloc null; + ret (Ptrs r clist ulist htab p)). +Next Obligation. +move=>r clist loop k H i [pf][/= f][hc][hct][->{i}][Hc V Hct]. +case: decP=>[{}pf|] /=; last first. +- case: leqP pf (eqn_leq k n)=>// _ -> /= pf _. + step=>W; split=>//; exists f, hc, hct; split=>//. + apply: tableP2 Hct=>// ?; rewrite !in_set (eqP pf). + by rewrite /n cardE index_mem /index mem_enum. +apply: hstep=>x; apply: hstep=>/=; rewrite -!joinA !(xpull _ hc). +apply/vrf_bnd/vrf_frame/[gE f]=>//=. +case=>m [Em _ _ Vm]. apply: [gE]=>[||??[]] //=. +split=>//. +exists [ffun z => if z == ith k pf then x else f z]. +eexists m, _. split=>//. +- rewrite /Array.shape InE /=. + split=>//. by rewrite (validL Vm). +rewrite (sepitS (ith k pf)) in_set indx_ith ltnSn. +rewrite /ctab/table !ffunE eq_refl. simpl. +rewrite !joinA. rewrite /lseq /=. + +rewrite InE. simpl. +eexists (x:->ith k pf \+ x.+1 :-> null), hct. +split=>//. +- by rewrite unitR. +- exists null, Unit. rewrite unitR. by []. +apply: tableP2 Hct=>//; last first. +- move=>s _. rewrite !in_set !ffunE indx_injE. + by case: eqP=>// ->; rewrite ltnn. + +move=>s. +rewrite !in_set ltnS. +rewrite in_set1 indx_injE. +by case: ltngtP=>//. +Qed. + +Next Obligation. +case=>_ ->; apply: [stepE]=>//= r hr [Em Vm]. +rewrite -[hr]unitL; apply: vrf_bnd. +apply/vrf_frame. +apply: [gE]=>//=. +move=>clist hc [Ec _] _ V. +apply/vrf_bnd/vrf_frame. +apply: [gE]=>[||??[]//] /=. +- split=>//. + exists [ffun x => null]. + exists hc, Unit. rewrite unitR. + split=>//. + - rewrite /Array.shape InE /=. split=>//. + by rewrite (validL V). + by rewrite (_ : [set c | indx c < 0] = set0) // sepit0. +case=>_ [_][f][hc'][hrest][->] Hc' Hrest _ V'. +rewrite -[hc' \+ _]unitL -joinA. +apply/vrf_bnd/vrf_frame/[gE]=>//=. +move=>ulist hu [Ehu] _ _ Vhu. +rewrite -[hu \+ _]unitL -joinA. +apply/vrf_bnd/vrf_frame/[gE]=>//=. +move=>htab ht /= Ht _ Vht. +apply/vrf_bnd/vrf_alloc=>p D. +apply: vrf_ret=>// D'. +exists (Ptrs r clist ulist htab p). split=>//. clear V. +pose j := (Data [ffun x => x] [ffun c => [:: c]] [ffun c => [::]] (nil K V) [::]). + +exists j. split; last first. +- move: (initP s). case=>PI Cl. by []. + +split=>/=. +- move=>s a. rewrite !ffunE !inE. by []. +exists f, [ffun s => null]. +split=>//. + +rewrite (_ : p :-> null \+ _ = hr \+ ((hc' \+ hrest) \+ (hu \+ Unit \+ + ht \+ (p :-> null \+ Unit)))); last first. +- heap_congr=>//. by rewrite unitR. + +eexists hr, _; split=>//. +eexists (hc' \+ hrest), _. split=>//. +- by exists hc', hrest; split=>//. +eexists (hu \+ Unit), (ht \+ p :-> null \+ Unit); split=>//. +- by heap_congr. +- rewrite /Array.shape /= !InE. + exists hu, Unit. split=>//. + - by split=>//; rewrite (validL Vhu). + - rewrite sepit_emp //. + move=>k ??. rewrite /utab/table !ffunE /=. simpl. + split=>//. case=>_ ->. by []. + +rewrite -joinA. +eexists ht, _. split=>//. +exists null. exists (p :-> null), Unit. +by []. +Qed. + +End Initialization. + +Module Dummy2. End Dummy2. + +Variable (r : {array symb -> symb}). +Variable (clist : {array symb -> llist symb}). +Variable (ulist : {array symb -> llist (symb*symb*symb)}). +Variable (htab : KVmap.tp LT). +Variable (p : loc). + +Notation ashape' := (ashape r clist ulist htab p). +Notation bshape' := (bshape r clist ulist htab p). +Notation shape' := (shape r clist ulist htab p). + +Definition cT (a' b' : symb) : Type := + forall x:unit, STsep unit + (fun i => (exists D, i \In ashape' D) /\ a' != b', + fun y i m => forall D, i \In ashape' D -> y = Val tt /\ exists ct, exists ut, + let: (d, _, _) := D in + m \In ashape' + (Data [ffun s => if s \in (class d a') then b' + else rep d s] + [ffun s => if s == a' then [::] else + if s == b' then rev (class d a') ++ class d b' + else class d s] + (use d) (lookup d) (pending d), ct, ut)). + +Program +Definition join_hclass (a' b' : symb) : + STsep unit (fun i => (exists d, i \In bshape' d) /\ a' != b', + fun y i m => forall d, i \In bshape' d -> + y = Val tt /\ + m \In bshape' (join_class d a' b')) := + Do (Fix (fun (loop : cT a' b') (x : unit) => + Do (a <-- Array.read clist a'; + b <-- Array.read clist b'; + If a == null :> loc then ret tt + else + s <-- !a; + next <-- !(a .+ 1); + a .+ 1 ::= b;; + Array.write clist b' a;; + Array.write clist a' next;; + Array.write r s b';; + loop tt)) tt). +Next Obligation. +apply: (ghost (d, f0, f)) H1. +move: H0=>{loop f0 f d} N [[d]] ct ut H. +move: (H)=>[i1][w][->][R][[w']][i4][->{w}][[i2]][i3][->{w'}][Ct][Cv] _ [R'] _ D. +move: Cv; rewrite (sepitT1 a') (sepitS b') 3!in_set eq_sym N {1 2}/table /= => Cv. +move: Cv D=>[i5][w][->{i3}][Ca][[i3]][i6][->{w}][Cb][Cc] _ _. +rewrite -(unA i2) -(unCA i2); apply: bnd_gh (Ct)=>[a w [[->]] <-{w}|??[] //]. +apply: bnd_gh (Ct)=>[b w [[->]] <-{w}|??[] //]; rewrite (unCA i2) (unA i2). +case: ifP Ca=>E. +- rewrite (eqP E); case/lseq_null=>Ce -> D. + apply: val_ret=>//; split=>//; exists ct; exists ut. + rewrite InE /= (sepitT1 a') (sepitS b') 3!in_set eq_sym N {1 2}/table /= + (eqP E) Ce /= !ffunE !eq_refl eq_sym (negbTE N) /= + (_ : [ffun s => rep d s] = rep d); last first. + - by rewrite -ffunP => ?; rewrite ffunE. + hauto D; apply: tableP Cc=>?; rewrite !in_set // !ffunE. + by case: ifP=>W; [rewrite (eqP W) Ce | case: eqP]. +case/(lseq_pos (negbT E))=>s [q][i7][Ca] _ [<-{i5}] /= Cq. +rewrite -(unA i2) -2!(unA (ct a' :-> _)) -2!(unA (ct a' .+ 1 :-> _)) + -(unA i7) -(unA i3). +rewrite -2!(push (ct a'))=>D; apply: bnd_read=>//; rewrite 2!(push (ct a')) in D *. +rewrite -3!(push (ct a' .+ 1)) in D *; apply: bnd_read=>//. +apply: bnd_write=>{D} //; rewrite 3!(push (ct a' .+ 1)) -(unCA i2). +do 2![apply: {i2} bnd_gh Ct=>[[] i2 [Ct _]|[?]?[]] //]; rewrite (unCA i2). +apply: {i1} bnd_gh R=>[[] i1 [R _]|[?]?[]] //. +rewrite (unA i3) (unA i7) 2!(unA (ct a' .+ 1 :-> _)) 2!(unA (ct a' :-> _)) (unA i2). +set r2 := (finfun _ )in R; set ct2 := (finfun _) in Ct=>D. +set cv2 := [ffun z => if z == a' then (behead (class d a')) + else if z == b' then s :: (class d b') else class d z]. +apply: (cons_gh1 (Data r2 cv2 (use d) (lookup d) (pending d), ct2, ut))=> + [P|[] m [_][ct1][ut1] /= Wm Dm|??[]||] //=; last 1 first. +- hauto D; rewrite /ct2 (sepitT1 a') (sepitS b') 3!in_set eq_sym N /= /table /= + !ffunE !eq_refl !(eq_sym b') (negbTE N) -!unA -!(unCA i7) !unA unC -!unA unC -unA. + by hauto D; apply: tableP Cc=>?; rewrite !in_set !ffunE; do 2![case: eqP=>//]. +- by split=>//; first by eauto. +split=>{ct2 Ct D} //; exists ct1; exists ut1. +set r1 := (finfun _) in Wm; set c1 := (finfun _) in Wm. +rewrite (_ : (finfun _) = r1); last first. +- rewrite /r1 -ffunP=>?; rewrite !ffunE eq_refl. + case: eqP=>[->|]; first by rewrite if_same Ca inE eq_refl. + by rewrite Ca inE; case: eqP. +rewrite (_ : (finfun _) = c1) //. +rewrite /c1 -ffunP=>?; rewrite !ffunE !eq_refl !(eq_sym b') (negbTE N). +case: eqP=>// H1; case: eqP=>// H2. +by rewrite Ca rev_cons cat_rcons. +Qed. +Next Obligation. +apply: (ghost H) H1; move: H0 {H}=> N d [I][ct][ut] H; move: (ashapeD H). +apply: (cons_gh1 (d, ct, ut))=>[||??[]|] //; first by eauto. +move=>[] m [_] [ct1][ut1] W; split=>//; set d' := (Data _ _ _ _ _) in W. +suff E : join_class d a' b' = d' + by split; [apply: join_class_classP | rewrite E; eauto]. +congr Data. +by set v := (finfun _); rewrite -(ffunP v) /v => x; rewrite !ffunE /= I. +Qed. + +Module Dummy23. End Dummy23. + +Let proj0 (e : symb*symb*symb) := let: (c, c1, c2) := e in c. +Let proj1 (e : symb*symb*symb) := let: (c, c1, c2) := e in c1. +Let proj2 (e : symb*symb*symb) := let: (c, c1, c2) := e in c2. +Notation "e ..0" := (proj0 e) (at level 2). +Notation "e ..1" := (proj1 e) (at level 2). +Notation "e ..2" := (proj2 e) (at level 2). + +Definition uT (a' b' : symb) := forall x : unit, + STsep unit (fun i => exists d, exists done, a' != b' /\ + i \In bshape' (join_use' d a' b' done) /\ + use d a' = done ++ use (join_use' d a' b' done) a', + fun y i m => forall d, i \In bshape' d -> y = Val tt /\ + m \In bshape' (join_use d a' b')). + +Program Definition join_huse (a' b' : symb) : + STsep unit (fun i => exists d, a' != b' /\ i \In bshape' d, + fun y i m => forall d, i \In bshape' d -> + y = Val tt /\ m \In bshape' (join_use d a' b')) := + Do (Fix (fun (loop : uT a' b') x => + Do (a <-- Array.read ulist a'; + If a == null :> loc then ret tt + else + eq <-- !a; + next <-- !(a .+ 1); + Array.write ulist a' next;; + c1 <-- Array.read r eq..1; + c2 <-- Array.read r eq..2; + v <-- KVmap.lookup htab (c1, c2); + match_opt v then + KVmap.insert htab (c1, c2) eq;; + b <-- Array.read ulist b'; + a .+ 1 ::= b;; + Array.write ulist b' a;; + loop tt + else [d] + dealloc a;; + dealloc a .+ 1;; + p' <-- !p; + q <-- insert p' (comp_pend eq d); + p ::= q;; + loop tt)) tt). +Next Obligation. +apply: (ghost1 (join_use' H a' b' H0))=>[?|]; first by apply: bshape_inv. +move: H H0 H1 H2 H3=>d done N. +set d1 := join_use' d a' b' done. +move=>[C1][ct][ut][h1][w][->][H1][[h2]][w'][->{w}][H2][[w]][w''][->{w'}]; +move=>[[h3]][u][->{w}][Ut][U] _ [[h7]][w][->{w''}][H7][[l]][w'][h8][->{w}][]; +case/swp=>->{w'} _ [H8] _ _ _ _ D E; move: D. +move: U; rewrite (sepitT1 a') (sepitS b') 3!in_set eq_sym N {1 2} /table /=. +move=>[h4][w][->{u}][Ua][[h5]][h6][->{w}][Ub][Ru] _ _. +rewrite -(unA h3) -2!(unCA h3); apply: bnd_gh (Ut)=>[_ _ [[->]] <-|??[]] //. +case: ifP=>E1. +- rewrite (eqP E1) in Ua; case/lseq_null: Ua=>E2 ->. + rewrite 2!(unCA h3) (unA h3) /join_use E2=>D; apply: val_ret=>//; do !split=>//. + exists ct; exists ut; hauto D. + by rewrite (sepitT1 a') (sepitS b') 3!in_set eq_sym N /table (eqP E1) E2; hauto D. +case/(lseq_pos (negbT E1)): Ua. +move=>[[c1 c2] c][next][h9][E2] _ [<-{h4}] H9. +rewrite -2!(unA (ut a' :-> _)) -3!(push (ut a'))=>D; apply: bnd_read=>//. +rewrite -2!(unA (ut a' .+ 1 :-> _)) -4!(push (ut a' .+ 1)) in D *. +apply: bnd_read=>//; move: D; rewrite -2!(unCA h3). +apply: {h3} bnd_gh Ut=>[[] h3 [Ut _]|??[] //]; rewrite -3!(unCA h1). +apply: bnd_gh (H1)=>[c1' _ [[Ec1]] <-|??[] //]. +apply: bnd_gh (H1)=>[c2' _ [[Ec2]] <-|??[] //]. +rewrite -(unA h9) -(unA h5) -8!(unCA h7). +apply: {h7} bnd_gh H7=>[v h7 [H7][Eqv]|??[] //]. +set d2 := join_use' d a' b' (done ++ [:: (c1, c2, c)]). +have E3: use d2 a' = behead (use d1 a'). +- rewrite /d2 (@join_useT (behead (use d1 a'))); last by rewrite /= -E2. + by apply: join_use_useE; rewrite /= -E2. +have E4: join_use d2 a' b' = join_use d1 a' b'. +- by rewrite /join_use E3 -!(@join_useT [::]) ?cats0 -?catA ?E /= -?E2. +case: v Eqv=>[[[e1 e2] e]|] /= Eqv. +- rewrite -4!(push (ut a')); apply: bnd_dealloc. + rewrite -3!(push (ut a' .+ 1)); apply: bnd_dealloc. + rewrite -7!(push p)=>D; apply: bnd_read=>//; move: D. + rewrite -(unC h8) -7!(unCA h8). + apply: bnd_gh H8=>{h8 x0} [_ h8 [q][H8][->]|??[?][]//]. + rewrite -(push p); apply: bnd_write=>D. + apply: (cons_gh1 d2)=>[?|[] m [_] P _|??[]||] //. + - by exists d; exists (done ++ [:: (c1, c2, c)]); rewrite -/d2 E3 -catA /= -E2. + - by rewrite -E4. + rewrite (_ : _ :+ _ = + h1 :+ (h2 :+ (h3 :+ (h9 :+ (h5 :+ h6)) :+ (h7 :+ (p :-> q :+ h8))))) in D *; + last by heap_congr. + case: H8 D=>x0 [h'] [_] <- H8 D. + rewrite /d2 (@join_useT (behead (use d1 a'))) -/d1 /= -?E2 // -Ec1 -Ec2 -Eqv; split=>//. + set ut2 := (finfun _) in Ut. + exists ct; exists ut2; hauto D. + rewrite (sepitT1 a') (sepitS b') 3!in_set eq_sym N /table /=. + rewrite !ffunE !eq_refl !(eq_sym b') (negbTE N); hauto D. + by apply: tableP Ru=>s; rewrite !in_set /ut2 ffunE; + case: (s == a')=>//; rewrite andbF. +apply: bnd_gh H7=>{h7} [[] h7 [H7 _]|??[] //]; rewrite -2!(unCA h3). +apply: bnd_gh (Ut)=>[_ _ [[->]] <-|??[] //]; rewrite -3!(push (ut a' .+ 1)). +apply: bnd_write; rewrite ffunE (eq_sym b') (negbTE N) -(unCA h3). +apply: {h3} bnd_gh Ut=>[[] h3 [Ut _]|??[] //]=>D. +apply: (cons_gh1 d2)=>[||??[]||] //. +- by exists d; exists (done ++ [:: (c1, c2, c)]); rewrite -/d2 E3 -catA /= -E2. +- by rewrite -E4. +rewrite (_ : _ :+ _ = h1 :+ (h2 :+ (h3 :+ (h9 :+ ((ut a' :-> (c1, c2, c) :+ + (ut a' .+ 1 :-> ut b' :+ h5)) :+ h6)) :+ (h7 :+ (p :-> l :+ h8))))); + last by heap_congr. +rewrite /d2 (@join_useT (behead (use d1 a'))) -/d1 /= -?E2 // -Ec1 -Ec2 -Eqv. +set ut2 := (finfun _) in Ut; split=>//. +exists ct; exists ut2; hauto D. +rewrite (sepitT1 a') (sepitS b') 3!in_set eq_sym N /table /=. +rewrite /ut2 !ffunE !eq_refl (eq_sym b') (negbTE N); hauto D. +by apply: tableP Ru=>s; rewrite !in_set !ffunE; case: (s == b')=>//; case: (s == a'). +Qed. +Next Obligation. +by apply: cons0=>//; [exists H; exists (Nil (symb*symb*symb)) | apply: bshapeD H1]. +Qed. + +Module Dummy56. End Dummy56. + +Let pend0 (e : pend) := + match e with simp_pend a b => a | comp_pend (a,_,_) (b,_,_) => a end. +Let pend1 (e : pend) := + match e with simp_pend a b => b | comp_pend (a,_,_) (b,_,_) => b end. +Notation "e ..0" := (pend0 e) (at level 2). +Notation "e ..1" := (pend1 e) (at level 2). + +Definition pT : Type := forall x:unit, + STsep unit (fun i => exists d, i \In bshape' d, + fun y i m => forall d, i \In bshape' d -> + y = Val tt /\ m \In bshape' (propagate d)). + +Program Definition hpropagate := + Fix (fun (loop : pT) x => + Do (q <-- !p; + If q == null then ret tt (* pending list is empty *) + else + eq <-- !q; + next <-- !(q .+ 1); + p ::= (next : loc);; + dealloc q;; + dealloc q .+ 1;; + a' <-- Array.read r eq..0; + b' <-- Array.read r eq..1; + If (a' == b') then loop tt + else join_hclass a' b';; + join_huse a' b';; + loop tt)) tt. +Next Obligation. +apply: (ghost1 H)=>[?|]; first by apply: bshape_inv. +move: {loop} H H0=>d [CI][ct][ut][hr][w][->][Hr][[hc]][w'][->{w}][Hc]. +move=>[[hu]][w][->{w'}][Hu][[ht]][w'][->{w}][Ht][[q]][w''][hp][->{w'}][]. +case/swp=>-> _ [Hp] _ _ _ _ {w' w''}. +rewrite -4!(push p)=>D; apply: bnd_read=>//. +case: ifP Hp D=>Eq. +- rewrite (eqP Eq); case/lseq_null=>Ep ->{hp} D. + apply: val_ret=>//; rewrite propagate_equation Ep; do 2!split=>//. + exists ct; exists ut; rewrite 4!(push p) in D *; hauto D. + by rewrite Ep. +case/(lseq_pos (negbT Eq))=>eq [next][hnext][Ep] _ <-{hp} Hp. +rewrite -5!(push q)=>D; apply: bnd_read=>//; rewrite -6!(push (q .+ 1)) in D *. +apply: bnd_read=>//; move: D; rewrite -2!(push p); apply: bnd_write; rewrite -2!(push q). +apply: bnd_dealloc; rewrite -(push (q .+ 1)); apply: bnd_dealloc; rewrite -(unCA hr). +apply: bnd_gh (Hr)=>[a' _ [[Ea]] <-|??[] //]. +apply: bnd_gh (Hr)=>[b' _ [[Eb]] <-|??[] //]. +set d1 := Data (rep d) (class d) (use d) (lookup d) (behead (pending d)). +rewrite 3!(push p). +case: ifP=>E D. +- have T1: propagate d = propagate d1. + - rewrite propagate_equation Ep /=. + by case: eq Ep Ea Eb=>[a2 b2 _ <-<-| [[a2 ?]?][[b2 ?]?] _ <-<-]; rewrite E. + apply: (cons_gh d1)=>[[] m|??[] //||//]; first by rewrite T1. + by split=>//; exists ct; exists ut; hauto D. +rewrite -(unh0 (_ :+ _)) in D *. +apply: (bnd_gh1 d1)=>[|[] m [_] Wm|??[]//||//]; last 1 first. +- by split=>//; exists ct; exists ut; hauto D. +- by rewrite E; eauto. +set d2 := (join_class _ _ _) in Wm. +apply: (bnd_gh1 d2)=>[|{m Wm} [] m [_] Wm|??[]|] //; first by rewrite E /=; eauto. +rewrite unh0; apply: cons0=>/= [|y m2]; first by eauto. +case/(_ _ Wm)=>->{Wm} Wm Dm; split=>//. +rewrite (_ : propagate d = propagate (join_use d2 a' b')) // /d2 + propagate_equation Ep. +by case: eq Ep Ea Eb=>[a2 b2 _ <- <-|[[a2 ?]?][[b2 ?]?] _ <- <-]; rewrite E. +Qed. + +Module Dummy57. End Dummy57. + +Program Definition merge (e : Eq) : + STsep unit (fun i => exists R, i \In shape' R, + fun y i m => forall R, i \In shape' R -> y = Val tt /\ + m \In shape' (closure (R +r eq2rel e))) := + match e with + simp_eq a b => + Do (q <-- !p; + x <-- insert q (simp_pend a b); + p ::= x;; + hpropagate) + | comp_eq c c1 c2 => + Do (c1' <-- Array.read r c1; + c2' <-- Array.read r c2; + v <-- KVmap.lookup htab (c1', c2'); + match_opt v then + KVmap.insert htab (c1', c2') (c, c1, c2);; + u1 <-- Array.read ulist c1'; + x <-- insert u1 (c, c1, c2); + (* if c1' == c2' the equation will be added twice *) + (* but this is okay, so we need not check for equality *) + (* this will rarely occur in practice, because an equation *) + (* c = c1' c1' means that a function c1' is applied to itself *) + (* so by avoiding the check, we optimize for the common case *) + (* this complicates the proof somewhat, however *) + Array.write ulist c1' x;; + u2 <-- Array.read ulist c2'; + x <-- insert u2 (c, c1, c2); + Array.write ulist c2' x + else [b] + q <-- !p; + x <-- insert q (comp_pend (c, c1, c2) b); + p ::= x;; + hpropagate) + end. +Next Obligation. +apply: (ghost2 (t:=H)); first by move=>???; move/(shape_inv H0)=>->. +move: H H0=>R [d][[CI]][ct][ut][hr][w][->][Hr][[hc]][w'][->{w}][Hc]. +move=>[[w]][w''][->{w'}][[hu]][hu'][->{w}][Hu][Hu'] _ [[ht]][w']. +move=>[->{w''}][Ht][[q]][w''][hp][->{w'}][]. +case/swp=>-> _ [Hp] _ _ _ _ {w''} D [PI][Ep] ERel; move: D. +set d1 := Data (rep d) (class d) (use d) (lookup d) (simp_pend a b :: pending d). +rewrite -(unA hu) -5!(push p)=>D; apply: bnd_read=>//; move: D. +rewrite -(unC hp) -5!(unCA hp). +apply: {hp} bnd_gh Hp=>[_ hp [q'] [Hp] [->{x}]|??[?][]] //. +rewrite -(push p); apply: bnd_write=>D. +apply: (cons_gh d1)=>[[] m [_] Wm Dm|??[] //||//]; last first. +- rewrite 4!(unCA hp) (unC hp) 5!(push p) (unA hu) in D *. + move: Hp D=>[x][h'][_] <-; rewrite Ep=>[[->]] ->=>D. + split=>//; exists ct; exists ut; hauto D. + by rewrite Ep. +have L: propagate_inv d1 by apply: propagate_pendP PI. +move: (propagatePE L)=>[L1][L2] L3; split=>//; exists (propagate d1). +by rewrite -L3 -ERel {4}/d1 /CRel /= clos_clos -!orrA orrAC. +Qed. +Next Obligation. +apply: (ghost2 (t:=H)); first by move=>???; move/(shape_inv H0)=>->. +move: H H0=>R [d][[CI]][ct][ut][hr][w][->][Hr][[hc]][w'][->{w}][Hc]. +move=>[[w]][w''][->{w'}][[hu]][hu'][->{w}][Hu][Hu'] _ [[ht]][w']. +move=>[->{w''}][Ht][[q]][w''][hp][->{w'}][]. +case/swp=>-> _ [Hp] _ _ _ _ {w''} D [PI][Ep] ERel; move: D. +do 2!apply: bnd_gh (Hr)=>[_ _ [[->]] <-|??[]//]; rewrite -3!(unCA ht). +apply: {ht} bnd_gh Ht=>[v ht [Ht] [Ev]|??[]//]. +case: v Ev=>[[[e e1] e2]|] Ev. +- set d1 := Data (rep d) (class d) (use d) (lookup d) + (comp_pend (c, c1, c2) (e, e1, e2) :: pending d). + have L: propagate_inv d1 by apply: propagate_pendP PI. + rewrite -4!(push p)=>D; apply: bnd_read=>//; move: D. + rewrite -(unC hp) -4!(unCA hp). + apply: {hp} bnd_gh Hp=>[_ hp [x][Hp][->]|??[?][]//]. + rewrite -(push p); apply: bnd_write. + rewrite 3!(unCA hp) (unC hp) 4!(push p) 3!(unCA ht)=>D. + case: Hp D=>[x1][h'][_] <-; rewrite Ep; case=>->-> D. + apply: (cons_gh d1)=>[[] m [_] Wm Dm |??[]//||//]; last first. + - by split=>//; exists ct; exists ut; hauto D; rewrite Ep. + move: (propagatePE L)=>[L1][L2] L3; split=>//; exists (propagate d1). + rewrite -L3 {4}/d1 -ERel; do 3!split=>//. + by apply: propagate_clos_pendP. +apply: bnd_gh Ht=>{ht x} [_ ht [Ht] _|??[]//]; rewrite -(unA hu) -3!(unCA hu). +apply: bnd_gh (Hu)=>[_ _ [[->]] <-|??[]//]; move: Hu'. +rewrite (sepitT1 (rep d c1)) {1}/table=>[[hu'']][hu2][->{hu'}][Hu''][Hu'] _. +rewrite -(unA hu'') -4!(unCA hu''). +apply: {hu''} bnd_gh Hu''=>[_ hu'' [x] [Hu''][->]|??[?][]//]. +rewrite -(unCA hu); apply: {hu} bnd_gh Hu=>[_ hu [Hu] _|??[] //]. +apply: bnd_gh (Hu)=>[_ _ [[->]] <-|??[]//]. +set ut1 := (finfun _) in Hu. +set u1' := [ffun z => if z == rep d c1 then (c, c1, c2) :: use d z else use d z]. +set u' := [ffun z => if z == rep d c2 then (c, c1, c2) :: u1' z else u1' z]. +set l' := (ins _ _ _) in Ht. +set d1 := Data (rep d) (class d) u' l' (pending d). +pose ut2 x' := [ffun z => if z == rep d c2 then x' else ut1 z]. +case E: (rep d c2 == rep d c1). +- rewrite -(unCA hu''). + apply: (bnd_gh ((c, c1, c2) :: use d (rep d c1))); last first. + - by rewrite ffunE E. + - by move=>??[?][]. + move=>_ {hu'' Hu''} hu'' [x'] [Hu''] [->]. + rewrite -(unCA hu); apply: {hu} val_gh Hu=>[[] hu [Hu] _|??[]//]. + rewrite (_ : _ :+ _ = hr :+ (hc :+ (hu :+ (hu'' :+ hu2) :+ (ht :+ (p :-> q :+ hp))))); + last by heap_congr. + move=>D; split=>//; exists d1; split; last first. + - split; first by apply: propagate_nopendP. + split; first by rewrite /d1. + by rewrite -ERel; apply: propagate_clos_nopendP. + split=>//; exists ct; exists (ut2 x'); hauto D. + rewrite (sepitT1 (rep d c1)) {1}/table !ffunE (eq_sym (rep d c1)) E eq_refl. + hauto D; apply: tableP Hu' => x0; rewrite !in_set andbT /ut1 !ffunE; + by case: ifP=>H1; by [rewrite (eqP H1) E | case: eqP]. +move: Hu'; rewrite (sepitS (rep d c2)) !in_set E /=. +case=>[hu1][hu3][->][Hu1'][Hu2'] _; rewrite -(unA hu1) -5!(unCA hu1). +apply: (bnd_gh (use d (rep d c2))); last first. +- by rewrite ffunE E. +- by move=>??[?][]. +move=>_ {hu1 Hu1'} hu1 [x'] [Hu1'] [->]. +rewrite -!(unCA hu); apply: {hu} val_gh Hu=>[[] hu [Hu] _|??[]//]. +rewrite (_ : _ :+ _ = hr :+ (hc :+ (hu :+ (hu'' :+ (hu1 :+ hu3)) :+ (ht :+ (p :-> q :+ hp))))); + last by heap_congr. +move=>D; split=>//; exists d1; split; last first. +- split; first by apply: propagate_nopendP. + split; first by rewrite /d1. + by rewrite -ERel; apply: propagate_clos_nopendP. +split=>//; exists ct; exists (ut2 x'); hauto D. +rewrite (sepitT1 (rep d c1)) /table !ffunE eq_sym E eq_refl; hauto D. +rewrite (sepitS (rep d c2)) !in_set E /= !ffunE eq_refl E; hauto D. +apply: tableP Hu2'=>y; rewrite !in_set /u' !ffunE; +(case: ifP=>H1; first by rewrite (eqP H1) eq_refl); +by case: ifP=>H2 //; rewrite (eqP H2) eq_refl andbF. +Qed. + +Module Dummy5. End Dummy5. + +Program +Definition Match_exp (A : Type) (t : exp) (p1 : symb -> spec A) (p2 : exp -> exp -> spec A) + (c1 : forall s, STsep A (p1 s)) + (c2 : forall t1 t2, STsep A (p2 t1 t2)) : + STsep A (match t with const s => (p1 s) | app t1 t2 => (p2 t1 t2) end) := + match t with const s => st.do (c1 s) _ | app t1 t2 => st.do (c2 t1 t2) _ end. + +Let pend3 (e : symb*symb*symb) := let: (a, _, _) := e in a. +Notation "e ..0" := (pend3 e) (at level 2). + +Definition nT : Type := + forall t, STsep exp (fun i => exists d, i \In bshape' d, + fun y i m => forall d, i \In bshape' d -> + m \In bshape' d /\ y = Val (norm d t)). + +Program Definition hnorm := + Fix (fun (hnorm : nT) (t : exp) => + Do (Match_exp t + (fun a => a' <-- Array.read r a; + ret (const a')) + (fun t1 t2 => + u1 <-- hnorm t1; + u2 <-- hnorm t2; + Match_exp u1 + (fun w1 => + Match_exp u2 + (fun w2 => + v <-- KVmap.lookup htab (w1, w2); + match_opt v then ret (app u1 u2) + else [a] a' <-- Array.read r (a..0); + ret (const a')) + (fun _ _ => ret (app u1 u2))) + (fun _ _ => ret (app u1 u2))))). +Next Obligation. +apply: (ghost2 (t:=H)); first by move=>???; move/(bshape_inv H0)=>->. +move: H H0=>d [CI][ct][ut][hr][hrest][->][Hr][Hrest]. +case: t=>[s|t1 t2]. +- apply: bnd_gh (Hr)=>[_ _ [[->]] <-|??[]//]=>D. + by apply: val_ret=>//; do 2!split=>//; exists ct; exists ut; hauto D. +rewrite -(unh0 (_ :+ _))=>D. +apply: (bnd_gh d)=>[u1 m1 [W1] [E1] D1|??[]//||//]; last first. +- by split=>//; exists ct; exists ut; hauto D. +apply: (bnd_gh d)=>[u2 m2 [W2] [E2]|??[]||] //; rewrite unh0. +move: W2=>{CI ct ut hr hrest Hr Hrest D}[CI][ct][ut][hr][w][->][Hr][[hc]][w']. +move=>[->{w}][Hc][[hu]][w][->{w'}][Hu][[ht]][hrest][->{w}][Ht][Hrest] _ _ _ _ D. +case: u1 E1=>[w1|??] /= E1; last first. +- rewrite -E1 E2; apply: val_ret=>//; do 2!split=>//. + by exists ct; exists ut; hauto D. +case: u2 E2=>[w2|??] /= E2; last first. +- rewrite -E1 -E2; apply: val_ret=>//; do 2!split=>//. + by exists ct; exists ut; hauto D. +move: D; rewrite -3!(unCA ht). +apply: bnd_gh Ht=>{ht x} [v ht [Ht][Ev]|??[]//]; rewrite 3!(unCA ht) -E1 -E2. +case: v Ev=>[[[a0 a1] a2]|] /= <- D; last first. +- by apply: val_ret=>//; do 2!split=>//; exists ct; exists ut; hauto D. +apply: bnd_gh (Hr) (D)=>[a' _ [[->]] <-|??[]//]. +by apply: val_ret; do 2!split=>//; exists ct; exists ut; hauto D. +Qed. + +Module Dummy6. End Dummy6. + +Program Definition check t1 t2 : + STsep bool (fun i => exists R, i \In shape' R, + fun y i m => forall R, i \In shape' R -> m \In shape' R /\ + exists b, y = Val b /\ (b <-> (t1, t2) \In R)) := + Do (u1 <-- hnorm t1; + u2 <-- hnorm t2; + ret (u1 == u2)). +Next Obligation. +apply: (ghost2 (t:=H)). +- by move=>???; move/(shape_inv H0)=>E [L1][b][->] L2; move: L1 L2; rewrite E; eauto. +move: H H0 (shapeD H0)=>R [d][H][[RI]H1][P] H2; rewrite -(unh0 i). +apply: {i} bnd_gh H=>[u1 i [H] [->]|??[]//]. +apply: {i} bnd_gh H=>[u2 i [H] [->]|??[]//]. +apply: val_ret=>//; rewrite unh0; split; first by exists d. +exists (norm d t1 == norm d t2); split=>//. +by case: normP=>//; rewrite H2; move=>H3; split. +Qed. From f1e7cfe2267a62dba9f1de23eb4ba7cb12ee83e1 Mon Sep 17 00:00:00 2001 From: Aleks Nanevski Date: Tue, 13 Aug 2024 01:07:15 +0200 Subject: [PATCH 40/93] wip --- examples/array.v | 20 ++--- examples/bubblesort.v | 12 +-- examples/congprog.v | 203 +++++++++--------------------------------- examples/hashtab.v | 20 ++--- examples/quicksort.v | 8 +- examples/stack.v | 26 +++--- pcm/pcm.v | 37 +++++--- 7 files changed, 104 insertions(+), 222 deletions(-) diff --git a/examples/array.v b/examples/array.v index b6612c3..c7338e8 100644 --- a/examples/array.v +++ b/examples/array.v @@ -41,7 +41,7 @@ Notation array := {array I -> T}. (* array is specified by finite function *) Definition shape (a : array) (f : {ffun I -> T}) := - [Pred h | h = updi a (fgraph f) /\ valid h]. + [Pred h | h = updi a (fgraph f)]. (* main methods *) @@ -55,7 +55,7 @@ Next Obligation. (* pull out ghost vars, run the program *) move=>x [] _ /= ->; step=>y; step. (* simplify *) -rewrite unitR=>H; split=>//; congr updi; rewrite codomE cardE. +rewrite unitR=>H; congr updi; rewrite codomE cardE. by elim: (enum I)=>[|t ts] //= ->; rewrite (ffunE _ _). Qed. @@ -89,11 +89,11 @@ Program Definition newf (f : {ffun I -> T}) : (* first the loop *) Next Obligation. (* pull out preconditions (note that there are no ghost vars), split *) -move=>f v x loop s [] /= _ [g][s'][[-> _]]; case: s=>[|k t] /= H1 H2. +move=>f v x loop s [] /= _ [g][s'][->]; case: s=>[|k t] /= H1 H2. - (* we've reached the end, return the array *) - rewrite cats0 in H1; step=>H; split=>//. (* g spans the whole f *) - by rewrite (_ : g = f) // -ffunP=>y; apply: H2; rewrite H1 mem_enum. + rewrite cats0 in H1; step=>_; rewrite (_ : g = f) // -ffunP=>y. + by apply: H2; rewrite H1 mem_enum. (* run the loop iteration, split the heap and save its validity *) rewrite (updi_split x k); step; apply: vrfV=>V; apply: [gE]=>//=. (* add the new index+value to g *) @@ -109,11 +109,11 @@ Next Obligation. (* pull out params, check if I is empty *) move=>f [] _ ->; case: fintype.pickP=>[v|] H /=. - (* run the `new` subroutine, simplify *) - apply: [stepE]=>//= a _ [-> V]. + apply: [stepE]=>//= a _ ->. (* invoke the loop, construct g from the first value of f *) by apply: [gE]=>//=; exists [ffun => f v], nil. (* I is empty, so should be the resulting heap *) -step=>_; split=>//; rewrite codom_ffun. +step=>_; rewrite codom_ffun. suff L: #|I| = 0 by case: (fgraph f)=>/=; rewrite L; case. by rewrite cardE; case: (enum I)=>[|x s] //; move: (H x). Qed. @@ -151,7 +151,7 @@ Qed. (* now the outer program *) Next Obligation. (* pull out params, invoke the loop *) -move=>a [] /= _ [f][-> V]; apply: [gE]=>//=. +move=>a [] /= _ [f ->]; apply: vrfV=>V; apply: [gE]=>//=. (* the suffix xs is the whole codomain of f *) exists (tval (fgraph f))=>/=. by rewrite ptr0 V {3}codomE size_map -cardE addn0. @@ -164,7 +164,7 @@ Program Definition read (a : array) (k : I) : Do (!a .+ (indx k)). Next Obligation. (* pull out ghost vars *) -move=>a k [f][_][] _ [->][/= -> _]. +move=>a k [f][_][] _ [->->]. (* split the heap and run the program *) by rewrite (updi_split a k); step. Qed. @@ -177,7 +177,7 @@ Program Definition write (a : array) (k : I) (x : T) : Do (a .+ (indx k) ::= x). Next Obligation. (* pull out ghost vars, split the heap *) -move=>a k x [f][] _ [-> _] /=; rewrite /shape !(updi_split a k). +move=>a k x [f][] _ -> /=; rewrite /shape !(updi_split a k). (* run the program, simplify *) by step; rewrite takeord dropord ffunE /= eq_refl. Qed. diff --git a/examples/bubblesort.v b/examples/bubblesort.v index 89f570b..a2956de 100644 --- a/examples/bubblesort.v +++ b/examples/bubblesort.v @@ -347,12 +347,12 @@ Program Definition cas_next (a : {array 'I_n.+2 -> A}) (i : 'I_n.+1) : ret true else ret false). Next Obligation. -move=>a i /= [f][] h /= [E V]; set i0 := Wo i; set i1 := So i. +move=>a i /= [f][] h /= E; set i0 := Wo i; set i1 := So i. do 2!apply: [stepE f, h]=>//= _ _ [->->]. case: oleqP=>H; first by step=>_; exists 1%g; split=>//; rewrite pffunE1. -apply: [stepE f ]=>//= _ _ [-> Vs ]. +apply: [stepE f ]=>//= _ _ ->. set fs := finfun [eta f with i0 |-> f i1]. -apply: [stepE fs]=>//= _ _ [-> Vsw]. +apply: [stepE fs]=>//= _ _ ->. set fsw := finfun [eta fs with i1 |-> f i0]. step=>_; exists (swnx i); do!split=>//=. suff: fsw = pffun (swnx i) f by move=>->. @@ -442,7 +442,7 @@ Program Definition bubble_sort (a : {array 'I_n.+2 -> A}) : then go tt : ST _ else skip)). Next Obligation. -move=>a go _ [f][] h /= [E V]. +move=>a go _ [f][] h /= E. apply: [stepE f]=>//; case=>m /=. - (* the pass did swap something, loop *) case=>p [Hm _]. @@ -639,7 +639,7 @@ rewrite codom_ux_rcons /= allrel_rconsl Ha1 /=. by have -> //: Ordinal P = @Ordinal n.+2 1 (ltn0Sn n); apply/ord_inj. Qed. Next Obligation. -move=>a [f][] h /= [E V]. +move=>a [f][] h /= E. have Hs: size (codom f) <= n.+2 by rewrite size_codom /= card_ord. apply: [gE f]=>//=; split=>//; rewrite (@itv_overL _ _ (BRight n.+1)) //=; try by rewrite addn1. @@ -856,7 +856,7 @@ rewrite -slice_oSL -slice_oSR (M2lo_eq Hx); suff: x.-2.+2 = x by move=>->. by rewrite -subn2 -addn2 addnBAC // addnK. Qed. Next Obligation. -move=>a [f][] h /= [E V]. +move=>a [f][] h /= E. apply: [gE f]=>//=. rewrite (@itv_overL _ _ _ +oo) /=; first by split=>//; exact: allrel0r. by rewrite leEnat addn1 size_codom /= card_ord. diff --git a/examples/congprog.v b/examples/congprog.v index d0c6ecd..f508c29 100644 --- a/examples/congprog.v +++ b/examples/congprog.v @@ -7,8 +7,7 @@ From Coq Require SetoidTactics. From mathcomp Require Import ssreflect ssrbool ssrnat. From mathcomp Require Import eqtype ssrfun div finset seq fintype finfun. From mathcomp Require Import choice. -From pcm Require Import axioms ordtype finmap pred pcm unionmap heap. - +From pcm Require Import axioms ordtype finmap pred pcm unionmap heap autopcm. From htt Require Import options model heapauto llist array. From htt Require Import kvmaps hashtab congmath. Import Prenex Implicits. @@ -51,14 +50,12 @@ Definition n := #|{: symb}|. (* the empty congruence is one that only relates constant symbols to themselves *) Definition empty_cong := closure (graph (@const s)). -Section ShapePredicates. - (* the algorithm starts with root pointers for the data *) - Inductive ptrs : Type := Ptrs of {array symb -> symb} & {array symb -> llist symb} & {array symb -> llist (symb*symb*symb)} & KVmap.tp LT & ptr. +Section ShapePredicates. Variable (r : {array symb -> symb}). Variable (clist : {array symb -> llist symb}). Variable (ulist : {array symb -> llist (symb*symb*symb)}). @@ -69,82 +66,18 @@ Variable (p : ptr). Definition ashape D := [Pred h | let: (d, ct, ut) := D in - valid h /\ h \In Array.shape r (rep d : {ffun symb -> symb}) # (Array.shape clist ct # sepit setT (ctab ct (class d))) # (Array.shape ulist ut # sepit setT (utab ut (use d))) # KVmap.shape htab (lookup d) # [Pred h' | exists l, h' \In Pred1 (p :-> l) # lseq l (pending d)]]. -Lemma ashapeD D (h : heap) : - h \In ashape D -> - valid h. -Proof. by case: D; case=>d ct ut []. Qed. - -(* -Lemma ashape_invert : invertible ashape. -Proof. -move=>[[[r1]]] c1 u1 t1 p1 ct1 ut1 [[[r2]]] c2 u2 t2 p2 ct2 ut2 h1 h2 /=. -move=>[hr1][w][->][R1][[hc1]][w'][->{w}][C1][[hu1]][w][->{w'}][U1] - [[ht1]][w'][->{w}][T1][[l1]][hl1][hp1][->{w'}][L1][P1] _ _ _ _ _. -move=>[hr2][w][->][R2][[hc2]][w'][->{w}][C2][[hu2]][w][->{w'}][U2] - [[ht2]][w'][->{w}][T2][[l2]][hl2][hp2][->{w'}][L2][P2] _ _ _ _ _ A. -case: (Array.shape_invert R1 R2 (agreeUnK A)) (A)=>/= -> -> {A}; move/agreeKUn=>A. -case: (table_invert (@lseq_invert _) C1 C2 (agreeUnK A)) (A) =>/= ->; -move/ffunP=>->-> {A}; move/agreeKUn => A. -case: (table_invert (@lseq_invert _) U1 U2 (agreeUnK A)) (A)=>/= ->; -move/ffunP=>->-> {A}; move/agreeKUn => A. -case: (KVmap.shape_invert T1 T2 (agreeUnK A)) (A)=>/= -> -> {A}; move/agreeKUn=>A. -case: (ptr_invert L1 L2 (agreeUnK A)) (A)=>/= E -> {A}; move/agreeKUn=>A. -by rewrite -{}E in P2; case: (lseq_invert P1 P2 A)=>/= -> ->. -Qed. -*) - -(* -Lemma ashape_inv D1 D2 h : h \In ashape D1 -> h \In ashape D2 -> D1 = D2. -Proof. by apply: iinv; [apply: ashape_invert | apply: ashapeD]. Qed. -*) - Definition bshape d := [Pred h | class_inv d /\ exists ct ut, h \In ashape (d, ct, ut)]. -Lemma bshapeD d h : - h \In bshape d -> valid h. -Proof. by case=>_ [ct][ut]; apply: ashapeD. Qed. - -(* -Lemma bshape_invert : invertible bshape. -Proof. -move=>d1 d2 h1 h2 [_][ct1][ut1] H1 [_][ct2][ut2] H2. -by case/(ashape_invert H1 H2)=>[[->]] _ _ ->. -Qed. - -Lemma bshape_inv d1 d2 h : h \In bshape d1 -> h \In bshape d2 -> d1 = d2. -Proof. by apply: iinv; [apply: bshape_invert | apply: bshapeD]. Qed. -*) - Definition shape (R : rel_exp s) := [Pred h | exists d, h \In bshape d /\ propagate_inv d /\ pending d = [::] /\ - forall x, x \In CRel d <-> x \In R]. - -Lemma shapeD (R : rel_exp s) h : - h \In shape R -> - valid h. -Proof. by case=>d [H] _; apply: bshapeD H. Qed. - -(* -Lemma shape_invert R1 R2 h1 h2 : - h1 \In shape R1 -> h2 \In shape R2 -> agree h1 h2 -> R1 =r R2 /\ h1 = h2. -Proof. -move=>R1 R2 h1 h2 [d1][S1][_][_] <- [d2][S2][_][_] <- A. -by move: (bshape_invert S1 S2 A)=>[<- <-]. -Qed. - -Lemma shape_inv R1 R2 h : h \In shape R1 -> h \In shape R2 -> R1 =r R2. -Proof. -by move=>??? H; case/(shape_invert H)=>//; apply: agree_refl; apply: shapeD H. -Qed. -*) + CRel d =p R]. End ShapePredicates. @@ -158,16 +91,6 @@ by [rewrite -H | rewrite H]. Qed. *) -(* Initialization procedure to generate the empty structure *) -Section Initialization. - - -Definition iT (clist : {array symb -> llist symb}): Type := forall k, - STsep (fun i => k <= n /\ exists f, i \In Array.shape clist f # - sepit [set c | indx c < k] (ctab f [ffun c => [:: c]]), - fun y m => y = Val tt /\ exists f, m \In Array.shape clist f # - sepit setT (ctab f [ffun c => [:: c]])). - Definition ith {I : finType} i (pf : i < #|I|) : I. Admitted. Lemma indx_ith {I : finType} i (pf : i < #|I|) : indx (ith i pf) = i. @@ -189,12 +112,19 @@ Admitted. Lemma sepit_emp (A : eqType) (s : seq A) f : - (forall x, x \in s -> forall k, k \In f x <-> emp k) -> - (forall k : heap, k \In IterStar.sepit s f <-> emp k). + (forall x, x \in s -> f x =p emp (U:=heap)) -> + IterStar.sepit s f =p emp. Proof. Admitted. +(* Initialization procedure to generate the empty structure *) + +Definition iT (clist : {array symb -> llist symb}): Type := forall k, + STsep (fun i => k <= n /\ exists f, i \In Array.shape clist f # + sepit [set c | indx c < k] (ctab f [ffun c => [:: c]]), + fun y m => y = Val tt /\ exists f, m \In Array.shape clist f # + sepit setT (ctab f [ffun c => [:: c]])). Program Definition init : STsep (emp, fun y m => exists ptr : ptrs, y = Val ptr /\ @@ -206,7 +136,7 @@ Program Definition init : Do (if decP (b:= k < n) idP is left pf then x <-- allocb (ith k pf) 2; x.+1 ::= null;; - Array.write (I:=symb) clist (ith k pf) x;; + Array.write clist (ith k pf) x;; loop k.+1 else ret tt)) 0;; ulist <-- Array.new _ null; @@ -214,99 +144,48 @@ Program Definition init : p <-- alloc null; ret (Ptrs r clist ulist htab p)). Next Obligation. -move=>r clist loop k H i [pf][/= f][hc][hct][->{i}][Hc V Hct]. +move=>r clist loop k H i [pf][/= f][hc][hct][->{i} Hc Hct]. case: decP=>[{}pf|] /=; last first. - case: leqP pf (eqn_leq k n)=>// _ -> /= pf _. - step=>W; split=>//; exists f, hc, hct; split=>//. + step=>_; split=>//; exists f, hc, hct; split=>//. apply: tableP2 Hct=>// ?; rewrite !in_set (eqP pf). by rewrite /n cardE index_mem /index mem_enum. apply: hstep=>x; apply: hstep=>/=; rewrite -!joinA !(xpull _ hc). -apply/vrf_bnd/vrf_frame/[gE f]=>//=. -case=>m [Em _ _ Vm]. apply: [gE]=>[||??[]] //=. -split=>//. -exists [ffun z => if z == ith k pf then x else f z]. -eexists m, _. split=>//. -- rewrite /Array.shape InE /=. - split=>//. by rewrite (validL Vm). -rewrite (sepitS (ith k pf)) in_set indx_ith ltnSn. -rewrite /ctab/table !ffunE eq_refl. simpl. -rewrite !joinA. rewrite /lseq /=. - -rewrite InE. simpl. -eexists (x:->ith k pf \+ x.+1 :-> null), hct. -split=>//. -- by rewrite unitR. -- exists null, Unit. rewrite unitR. by []. -apply: tableP2 Hct=>//; last first. -- move=>s _. rewrite !in_set !ffunE indx_injE. - by case: eqP=>// ->; rewrite ltnn. - -move=>s. -rewrite !in_set ltnS. -rewrite in_set1 indx_injE. -by case: ltngtP=>//. +apply/vrf_bnd/vrf_frame/[gE f]=>//= [[m]] Em _ _. +apply: [gE]=>[||??[]] //=; split=>//. +eexists [ffun z => if z == ith k pf then x else f z], m, _. +split=>//; rewrite (sepitS (ith k pf)) in_set indx_ith ltnSn. +rewrite /ctab/table !ffunE eq_refl joinA unitL. +hhauto; apply: tableP2 Hct=>// s. +- by rewrite !in_set ltnS in_set1 indx_injE; case: ltngtP. +by rewrite !in_set !ffunE indx_injE; case: eqP=>// ->; rewrite ltnn. Qed. - Next Obligation. -case=>_ ->; apply: [stepE]=>//= r hr [Em Vm]. -rewrite -[hr]unitL; apply: vrf_bnd. -apply/vrf_frame. -apply: [gE]=>//=. -move=>clist hc [Ec _] _ V. -apply/vrf_bnd/vrf_frame. -apply: [gE]=>[||??[]//] /=. -- split=>//. - exists [ffun x => null]. - exists hc, Unit. rewrite unitR. - split=>//. - - rewrite /Array.shape InE /=. split=>//. - by rewrite (validL V). - by rewrite (_ : [set c | indx c < 0] = set0) // sepit0. -case=>_ [_][f][hc'][hrest][->] Hc' Hrest _ V'. +case=>_ ->; apply: [stepE]=>//= r hr Er; rewrite -[hr]unitL. +apply/vrf_bnd/vrf_frame/[gE]=>//= clist hc Ec _ _. +apply: [stepX]@hc=>[||??[]] //=. +- split=>//; exists [ffun x => null], hc, Unit; rewrite unitR. + by split=>//; rewrite (_ : [set c | indx c < 0] = set0) // sepit0. +case=>n0 [_][f][hc'][hrest][->] Hc' Hrest. rewrite -[hc' \+ _]unitL -joinA. -apply/vrf_bnd/vrf_frame/[gE]=>//=. -move=>ulist hu [Ehu] _ _ Vhu. +apply/vrf_bnd/vrf_frame/[gE]=>//= ulist hu Ehu _ _. rewrite -[hu \+ _]unitL -joinA. -apply/vrf_bnd/vrf_frame/[gE]=>//=. -move=>htab ht /= Ht _ Vht. -apply/vrf_bnd/vrf_alloc=>p D. -apply: vrf_ret=>// D'. -exists (Ptrs r clist ulist htab p). split=>//. clear V. -pose j := (Data [ffun x => x] [ffun c => [:: c]] [ffun c => [::]] (nil K V) [::]). - -exists j. split; last first. -- move: (initP s). case=>PI Cl. by []. - -split=>/=. -- move=>s a. rewrite !ffunE !inE. by []. +apply/vrf_bnd/vrf_frame/[gE]=>//= htab ht Ht _ _. +apply/vrf_bnd/vrf_alloc=>p D; apply: vrf_ret=>// D'. +exists (Ptrs r clist ulist htab p); split=>//. +exists (Data [ffun x => x] [ffun c => [:: c]] [ffun c => [::]] (nil K V) [::]). +split; last by case: (initP s). +split=>[s a|/=]; first by rewrite !ffunE !inE. exists f, [ffun s => null]. -split=>//. - rewrite (_ : p :-> null \+ _ = hr \+ ((hc' \+ hrest) \+ (hu \+ Unit \+ - ht \+ (p :-> null \+ Unit)))); last first. -- heap_congr=>//. by rewrite unitR. - -eexists hr, _; split=>//. -eexists (hc' \+ hrest), _. split=>//. -- by exists hc', hrest; split=>//. -eexists (hu \+ Unit), (ht \+ p :-> null \+ Unit); split=>//. -- by heap_congr. -- rewrite /Array.shape /= !InE. - exists hu, Unit. split=>//. - - by split=>//; rewrite (validL Vhu). - - rewrite sepit_emp //. - move=>k ??. rewrite /utab/table !ffunE /=. simpl. - split=>//. case=>_ ->. by []. - -rewrite -joinA. -eexists ht, _. split=>//. -exists null. exists (p :-> null), Unit. -by []. + (ht \+ (p :-> null \+ Unit))))); last by heap_congr=>//; rewrite unitR. +hhauto; rewrite sepit_emp //= => k. +by rewrite /utab/table !ffunE; split=>//; case=>_ ->. Qed. -End Initialization. -Module Dummy2. End Dummy2. +UP TO HERE + Variable (r : {array symb -> symb}). Variable (clist : {array symb -> llist symb}). diff --git a/examples/hashtab.v b/examples/hashtab.v index c9391cf..890c1ce 100644 --- a/examples/hashtab.v +++ b/examples/hashtab.v @@ -74,12 +74,12 @@ case: decP=>[{Eleq}pf H2|]; last first. (* k < n, allocate an empty bucket by framing on Unit *) apply: [stepU]=>//= b m Hm. (* write its id to the array under index k *) -apply: [stepX tab] @ h1=>{h1 H1}//= _ m2 [E2 V2]. +apply: [stepX tab] @ h1=>{h1 H1}//= _ m2 E2. (* invoke the loop *) apply: [gE]=>//=; split=>//; rewrite joinCA. (* extend the table by the new index/bucket pair, simplify *) exists [ffun z => if z == Ordinal pf then b else tab z], m2, (m \+ h2). -split=>//{m2 E2 V2}. +split=>//{m2 E2}. (* remove the new bucket from the heap *) rewrite (sepitS (Ordinal pf)) in_set leqnn {1}/table ffunE eq_refl. exists m, h2; do!split=>{m Hm}//; apply: tableP2 H2=>{h2}//. @@ -93,7 +93,7 @@ Next Obligation. (* simplify *) move=>/= [] _ ->. (* allocate the array *) -apply: [stepE]=>//= y m [Hm Vm]. +apply: [stepE]=>//= y m Hm. (* invoke the loop with index 0 *) apply: [gE]=>//=; split=>//. (* the table is empty *) @@ -125,7 +125,7 @@ Program Definition free x : STsep {s} (shape x s, (* first the loop *) Next Obligation. (* pull out the ghost + preconditions, branch on k comparison *) -move=>/= x loop k [] _ /= [Eleq][tf][bf][h1][h2][-> [H1 V1]]. +move=>/= x loop k [] _ /= [Eleq][tf][bf][h1][h2][-> H1]. case: decP=>[pf H|]; last first. (* k = n *) - case: ltnP Eleq (eqn_leq k n)=>// _ -> /= /eqP Ek _ H. @@ -171,7 +171,7 @@ Program Definition insert x k v : ret x). Next Obligation. (* pull out ghost + deconstruct precondition *) -move=>/= x k v [fm][] _ /= [tf][bf][Hf Hh [h1][h2][-> [/= H1 _] H2]]. +move=>/= x k v [fm][] _ /= [tf][bf][Hf Hh [h1][h2][-> /= H1 H2]]. (* read the bucket from array *) apply/[stepX tf, h1] @ h1=>//= _ _ [->->]. (* split out the bucket in the heap *) @@ -179,7 +179,7 @@ move: H2; rewrite (sepitT1 (hash k)) /table; case=>h3[h4][{h2}-> H3 H4]. (* insert into the bucket *) apply/[stepX (bf (hash k))] @ h3=>{h3 H3}//= b' m2 H2. (* write the bucket to the array, return the pointer *) -apply/[stepX tf] @ h1=>{h1 H1}//= _ m3 [E3 V3]; step=>_. +apply/[stepX tf] @ h1=>{h1 H1}//= _ m3 E3; step=>_. (* update the array and buckets' specs *) exists [ffun z => if z == hash k then b' else tf z], (fun b => if b == hash k then ins k v (bf b) else bf b); split=>/=. @@ -193,7 +193,7 @@ exists [ffun z => if z == hash k then b' else tf z], - move=>i0 k0; case: eqP; last by move=>_; apply: Hh. by move=>->; rewrite supp_ins inE=>/orP; case; [move/eqP=>->|apply: Hh]. (* the shape is respected: first, the array fits *) -exists m3, (m2 \+ h4); split=>{Hf Hh m3 E3 V3}//. +exists m3, (m2 \+ h4); split=>{Hf Hh m3 E3}//. (* split out the modified bucket *) rewrite (sepitT1 (hash k)) /table /= ffunE eq_refl. exists m2, h4; split=>{m2 H2} //. @@ -216,7 +216,7 @@ Program Definition remove x k : ret x). Next Obligation. (* pull out ghost + destructure precondition *) -move=>/= x k [fm][] _ /= [tf][bf][Hf Hh [h1][h2][-> [/= H1 _] H2]]. +move=>/= x k [fm][] _ /= [tf][bf][Hf Hh [h1][h2][-> /= H1 H2]]. (* read the bucket from array *) apply/[stepX tf, h1] @ h1=>//= _ _ [->->]. (* split out the bucket in the heap *) @@ -224,7 +224,7 @@ move: H2; rewrite (sepitT1 (hash k)); case=>h3[h4][{h2}-> H3 H4]. (* insert into the bucket *) apply/[stepX (bf (hash k))] @ h3=>{h3 H3}//= b' m2 H2. (* write the bucket to the array, return the pointer *) -apply/[stepX tf] @ h1=>{H1}//= _ m3 [E3 V3]; step=>_. +apply/[stepX tf] @ h1=>{H1}//= _ m3 E3; step=>_. (* update the array and buckets' specs *) exists [ffun z => if z == hash k then b' else tf z], (fun b => if b == hash k then rem k (bf b) else bf b); split=>/=. @@ -238,7 +238,7 @@ exists [ffun z => if z == hash k then b' else tf z], - move=>i0 k0; case: eqP; last by move=>_; apply: Hh. by move=>->; rewrite supp_rem !inE=>/andP [] _; apply: Hh. (* the shape is respected: first, the array fits *) -exists m3, (m2\+ h4); split=>{m3 E3 V3 Hf Hh}//. +exists m3, (m2\+ h4); split=>{m3 E3 Hf Hh}//. (* split out the modified bucket *) rewrite (sepitT1 (hash k)) /table /= ffunE eq_refl. exists m2, h4; split=>{m2 H2} //. diff --git a/examples/quicksort.v b/examples/quicksort.v index 661aea3..f12c259 100644 --- a/examples/quicksort.v +++ b/examples/quicksort.v @@ -266,11 +266,11 @@ move=>a i j /= [f][] h /= H. case: ifP=>[/eqP->|Hij]. - by step=>_; rewrite tperm1 pffunE1. do 2!apply: [stepE f, h]=>//= _ _ [->->]. -apply: [stepE f]=>//= _ _ [-> V1]. +apply: [stepE f]=>//= _ _ ->. set f1 := finfun [eta f with i |-> f j]. -apply: [gE f1]=>//= _ _ [-> V2]. +apply: [gE f1]=>//= _ _ ->. set f2 := finfun [eta f1 with j |-> f i]. -move=>{h H}_; split=>//=; suff {V1 V2}: f2 = pffun (tperm i j) f by move=>->. +move=>{h H} _; suff : f2 = pffun (tperm i j) f by move=>->. rewrite {}/f2 {}/f1; apply/ffunP=>/= x; rewrite !ffunE /= ffunE /=. by case: tpermP=>[->|->|/eqP/negbTE->/eqP/negbTE->] {x}//; rewrite eqxx // Hij. Qed. @@ -413,7 +413,7 @@ Program Definition partition_lm (a : {array 'I_n -> A}) swap a v hi;; ret v). Next Obligation. -move=> a lo hi Hlo /= [f][] h /= [E V]. +move=> a lo hi Hlo /= [f][] h /= E. apply: [stepE f, h]=>//= _ _ [->->]. apply: [stepE f]=>//= v m [Hi][p][Pm Hm Al Ah]. apply: [stepE (pffun p f)]=>//= _ k Hj. diff --git a/examples/stack.v b/examples/stack.v index dee4aed..d374dc8 100644 --- a/examples/stack.v +++ b/examples/stack.v @@ -28,27 +28,21 @@ Notation stack := (stack T). (* stack is a pointer to a singly-linked list *) Definition shape s (xs : seq T) := - [Pred h | exists p h', [ /\ valid (s :-> p \+ h'), - h = s :-> p \+ h' & + [Pred h | exists p h', [ /\ h = s :-> p \+ h' & h' \In lseq p xs]]. (* heap cannot match two different specs *) -Lemma shape_inv : forall s xs1 xs2 h, +Lemma shape_inv s xs1 xs2 h : + valid h -> h \In shape s xs1 -> h \In shape s xs2 -> xs1 = xs2. Proof. -move=>s xs1 xs2 _ [p][h1][D -> S][p2][h2][D2]. -case/(cancelO _ D)=><- _; rewrite !unitL=><-. -by apply: lseq_func=>//; move/validR: D. +move=>V [p][h1][E S][x][h'][]; rewrite {h}E in V *. +case/(cancelO _ V)=><- _; rewrite !unitL=><-. +by apply: lseq_func=>//; move/validR: V. Qed. -(* well-formed stack is a valid heap *) -Lemma shapeD s xs h : - h \In shape s xs -> - valid h. -Proof. by case=>p [h'][D ->]. Qed. - (* main methods *) (* new stack is a pointer to an empty heap/list *) @@ -61,7 +55,7 @@ Program Definition free s : STsep (shape s [::], [vfun _ h => h = Unit]) := Do (dealloc s). Next Obligation. -by case=>i [?][?][V ->][_ ->]; step=>_; rewrite unitR. +by case=>i [?][?][->][_ ->]; step=>_; rewrite unitR. Qed. (* pushing to the stack is inserting into the list and updating the pointer *) @@ -72,9 +66,9 @@ Program Definition push s x : STsep {xs} (shape s xs, s ::= l'). Next Obligation. (* pull out ghost + precondition, get the list *) -case=>xs [] _ /= [l][h][V -> H]; step. +case=>xs [] _ /= [l][h][-> H]; step. (* run the insert procedure with the ghost, deconstruct the new list *) -apply: [stepX xs]@h=>//= x0 _ [r][h'][-> H']. +apply: [stepX xs]@h=>//= l' _ [r][h'][-> H']. (* store the new list *) by step=>V'; hhauto. Qed. @@ -96,7 +90,7 @@ Program Definition pop s : (fun _ => throw EmptyStack)). Next Obligation. (* pull out ghost vars and precondition *) -case=>xs [] _ /= [p][h0][V -> H]. +case=>xs [] _ /= [p][h0][-> H]. (* get the list and invoke head on it, deal with exception first *) step; apply/[tryX xs]@h0=>//= [x|ex] m [Hm]; last first. - (* throw the stack exception *) diff --git a/pcm/pcm.v b/pcm/pcm.v index 6cc0717..e0e248d 100644 --- a/pcm/pcm.v +++ b/pcm/pcm.v @@ -1340,21 +1340,26 @@ Variables (U : pcm) (A : Type). Definition seqjoin (s : seq U) : U := \big[join/Unit]_(i <- s) i. -Definition sepit (s : seq A) (f : A -> Pred U) : Pred U := +Definition sepit' (s : seq A) (f : A -> Pred U) : Pred U := [Pred h | exists hs : seq U, [ /\ size hs = size s, h = seqjoin hs & All (fun '(a, h) => h \In f a) (zip s hs)]]. +Definition sepit s f := locked (sepit' s f). + +Lemma sepitE s f : sepit s f = sepit' s f. +Proof. by rewrite /sepit -lock. Qed. + Lemma sepit0 f : sepit [::] f =p emp. Proof. -move=>h; split. +move=>h; rewrite sepitE; split. - by case=>/= hs [/size0nil -> -> _]; rewrite /seqjoin !big_nil. by move=>->; exists [::]; rewrite /seqjoin !big_nil. Qed. Lemma sepit_cons x s f : sepit (x::s) f =p f x # sepit s f. Proof. -move=>h; split. +move=>h; rewrite !sepitE; split. - case=>/=; case=>[|h0 hs]; case=>//= /eqP; rewrite eqSS =>/eqP Hs. rewrite /seqjoin big_cons =>->[H0 H1]. by exists h0, (seqjoin hs); do!split=>//; exists hs. @@ -1364,17 +1369,17 @@ Qed. Lemma sepit_cat s1 s2 f : sepit (s1 ++ s2) f =p sepit s1 f # sepit s2 f. Proof. -elim: s1 s2=>[|x s1 IH] s2 h /=; split. -- case=>hs [E {h}-> H2]. +rewrite !sepitE; elim: s1 s2=>[|x s1 IH] s2 h /=; split. +- case=>hs [E {h}-> H2]; rewrite -sepitE. exists Unit, (seqjoin hs); rewrite unitL. by split=>//; [rewrite sepit0 | exists hs]. -- by case=>h1[h2][{h}->]; rewrite sepit0=>->; rewrite unitL. +- by case=>h1[h2][{h}->]; rewrite -sepitE sepit0=>->; rewrite unitL. - case=>/=; case=>[|h0 hs]; case=>//= /eqP; rewrite eqSS=>/eqP E. rewrite /seqjoin !big_cons /= =>->[H0 HS]. case: (IH s2 (seqjoin hs)).1; first by exists hs. move=>h1[h2][HJ H1 H2]; rewrite /seqjoin in HJ. exists (h0 \+ h1), h2; rewrite HJ joinA; split=>//. - by rewrite sepit_cons; exists h0, h1. + by rewrite -sepitE sepit_cons sepitE; exists h0, h1. case=>h1[h2][{h}->[]]; case=>[|h0 hs1]; case=>//= /eqP; rewrite eqSS=>/eqP E1. rewrite /seqjoin !big_cons /= =>{h1}->[H0 H1]; case=>hs2[E2 {h2}-> H2]. exists (h0 :: hs1 ++ hs2); rewrite /seqjoin big_cons big_cat joinA; split=>//=. @@ -1391,21 +1396,24 @@ move=>H; have Hx: x \In s2 by apply/(pperm_in H)/In_cons; left. case: (In_split Hx)=>s21[s22] E; rewrite {s2 Hx}E in H *. move/pperm_cons_cat_cons: H=>/IH {}IH. rewrite sepit_cons sepit_cat /= =>h; split. -- case=>h1[h2][{h}-> H1]; rewrite IH sepit_cat. +- case=>h1[h2][{h}-> H1]; rewrite IH sepit_cat !sepitE. case=>_[_][{h2}-> [hs3][E3 -> H3] [hs4][E4 -> H4]]; rewrite joinCA. exists (seqjoin hs3), (h1 \+ seqjoin hs4); split=>//; first by exists hs3. - by rewrite sepit_cons; exists h1, (seqjoin hs4); split=>//; exists hs4. -case=>_[h2][{h}-> [hs1][Hs1 -> H1]]; rewrite sepit_cons. + rewrite -sepitE sepit_cons sepitE; exists h1, (seqjoin hs4). + by split=>//; exists hs4. +rewrite sepitE; case=>_[h2][{h}-> [hs1][Hs1 -> H1]]. +rewrite sepit_cons sepitE. case=>h3[_][{h2}-> H3 [hs2][Hs2 -> H2]]; rewrite joinCA. exists h3, (seqjoin hs1 \+ seqjoin hs2); split=>//. -rewrite IH; exists (hs1 ++ hs2); split. +rewrite IH sepitE; exists (hs1 ++ hs2); split. - by rewrite !size_cat Hs1 Hs2. - by rewrite /seqjoin big_cat. by rewrite zip_cat //; apply/All_cat. Qed. Lemma sepitF s (f1 f2 : A -> Pred U) : - (forall x, x \In s -> f1 x =p f2 x) -> sepit s f1 =p sepit s f2. + (forall x, x \In s -> f1 x =p f2 x) -> + sepit s f1 =p sepit s f2. Proof. elim: s=>[|x s IH] H h; first by rewrite !sepit0. have /IH {IH}H': forall x : A, x \In s -> f1 x =p f2 x. @@ -1438,8 +1446,9 @@ case=>hb[h][{h1}-> Hb H]; rewrite joinCA; exists hb, (ha \+ h); split=>//; [rewrite sepit_cons | rewrite (IH Hx Hu)]; exists ha, h. Qed. -Corollary eq_sepitF s (f1 f2 : A -> Pred U) : - (forall x, x \in s -> f1 x =p f2 x) -> sepit s f1 =p sepit s f2. +Lemma eq_sepitF s (f1 f2 : A -> Pred U) : + (forall x, x \in s -> f1 x =p f2 x) -> + sepit s f1 =p sepit s f2. Proof. by move=>H; apply: sepitF=>x Hx; apply/H/mem_seqP. Qed. Corollary perm_sepit s1 s2 (f : A -> Pred U) : From 507f228679afab9fe436ba781efd165b7ef367bc Mon Sep 17 00:00:00 2001 From: Aleks Nanevski Date: Tue, 13 Aug 2024 03:00:36 +0200 Subject: [PATCH 41/93] some cleaning up in congprog.v --- examples/congprog.v | 26 ++++++++++---------------- 1 file changed, 10 insertions(+), 16 deletions(-) diff --git a/examples/congprog.v b/examples/congprog.v index f508c29..f941ccd 100644 --- a/examples/congprog.v +++ b/examples/congprog.v @@ -150,29 +150,23 @@ case: decP=>[{}pf|] /=; last first. step=>_; split=>//; exists f, hc, hct; split=>//. apply: tableP2 Hct=>// ?; rewrite !in_set (eqP pf). by rewrite /n cardE index_mem /index mem_enum. -apply: hstep=>x; apply: hstep=>/=; rewrite -!joinA !(xpull _ hc). -apply/vrf_bnd/vrf_frame/[gE f]=>//= [[m]] Em _ _. -apply: [gE]=>[||??[]] //=; split=>//. -eexists [ffun z => if z == ith k pf then x else f z], m, _. -split=>//; rewrite (sepitS (ith k pf)) in_set indx_ith ltnSn. -rewrite /ctab/table !ffunE eq_refl joinA unitL. -hhauto; apply: tableP2 Hct=>// s. +step=>x; step; rewrite -!joinA !(xpull _ hc). +apply: [stepX f]@hc=>//= [[m]] Em; apply: [gE]=>[||??[]//] //=. +split=>//; eexists [ffun z => if z == ith k pf then x else f z]. +hhauto; rewrite (sepitS (ith k pf)) in_set indx_ith ltnSn. +rewrite /ctab/table !ffunE eq_refl joinA; hhauto. +apply: tableP2 Hct=>// s. - by rewrite !in_set ltnS in_set1 indx_injE; case: ltngtP. by rewrite !in_set !ffunE indx_injE; case: eqP=>// ->; rewrite ltnn. Qed. Next Obligation. -case=>_ ->; apply: [stepE]=>//= r hr Er; rewrite -[hr]unitL. -apply/vrf_bnd/vrf_frame/[gE]=>//= clist hc Ec _ _. -apply: [stepX]@hc=>[||??[]] //=. +case=>_ ->; apply: [stepE]=>//= r hr Er; apply: [stepU]=>//= clist hc Ec. + apply: [stepX]@hc=>[||??[]] //=. - split=>//; exists [ffun x => null], hc, Unit; rewrite unitR. by split=>//; rewrite (_ : [set c | indx c < 0] = set0) // sepit0. case=>n0 [_][f][hc'][hrest][->] Hc' Hrest. -rewrite -[hc' \+ _]unitL -joinA. -apply/vrf_bnd/vrf_frame/[gE]=>//= ulist hu Ehu _ _. -rewrite -[hu \+ _]unitL -joinA. -apply/vrf_bnd/vrf_frame/[gE]=>//= htab ht Ht _ _. -apply/vrf_bnd/vrf_alloc=>p D; apply: vrf_ret=>// D'. -exists (Ptrs r clist ulist htab p); split=>//. +apply: [stepU]=>//= ulist hu Ehu; apply: [stepU]=>//= htab ht Ht. +step=>p; step=>_; exists (Ptrs r clist ulist htab p); split=>//. exists (Data [ffun x => x] [ffun c => [:: c]] [ffun c => [::]] (nil K V) [::]). split; last by case: (initP s). split=>[s a|/=]; first by rewrite !ffunE !inE. From 67088c5e92612fd12380382e2af1ac92ac66f522 Mon Sep 17 00:00:00 2001 From: Aleks Nanevski Date: Tue, 13 Aug 2024 03:07:41 +0200 Subject: [PATCH 42/93] cleanup in congprog.v --- examples/congprog.v | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) diff --git a/examples/congprog.v b/examples/congprog.v index f941ccd..6ced54d 100644 --- a/examples/congprog.v +++ b/examples/congprog.v @@ -146,10 +146,10 @@ Program Definition init : Next Obligation. move=>r clist loop k H i [pf][/= f][hc][hct][->{i} Hc Hct]. case: decP=>[{}pf|] /=; last first. -- case: leqP pf (eqn_leq k n)=>// _ -> /= pf _. - step=>_; split=>//; exists f, hc, hct; split=>//. - apply: tableP2 Hct=>// ?; rewrite !in_set (eqP pf). - by rewrite /n cardE index_mem /index mem_enum. +- case: (ltngtP k n) pf=>// Ekn _ _; step=>_. + split=>//; exists f, hc, hct; split=>//. + apply: tableP2 Hct=>// x; rewrite !in_set Ekn. + by rewrite /n cardE index_mem /index mem_enum. step=>x; step; rewrite -!joinA !(xpull _ hc). apply: [stepX f]@hc=>//= [[m]] Em; apply: [gE]=>[||??[]//] //=. split=>//; eexists [ffun z => if z == ith k pf then x else f z]. From 39d7b493503dc759b588f4d336cac54bc7f0c2c0 Mon Sep 17 00:00:00 2001 From: Aleks Nanevski Date: Tue, 13 Aug 2024 03:09:44 +0200 Subject: [PATCH 43/93] cleanup --- examples/congprog.v | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/examples/congprog.v b/examples/congprog.v index 6ced54d..65b1f37 100644 --- a/examples/congprog.v +++ b/examples/congprog.v @@ -170,9 +170,9 @@ step=>p; step=>_; exists (Ptrs r clist ulist htab p); split=>//. exists (Data [ffun x => x] [ffun c => [:: c]] [ffun c => [::]] (nil K V) [::]). split; last by case: (initP s). split=>[s a|/=]; first by rewrite !ffunE !inE. -exists f, [ffun s => null]. -rewrite (_ : p :-> null \+ _ = hr \+ ((hc' \+ hrest) \+ (hu \+ Unit \+ - (ht \+ (p :-> null \+ Unit))))); last by heap_congr=>//; rewrite unitR. +exists f, [ffun s => null]. +rewrite (_ : p :-> null \+ _ = hr \+ ((hc' \+ hrest) \+ (hu \+ Unit \+ + (ht \+ (p :-> null \+ Unit))))); last by rewrite unitR; heap_congr. hhauto; rewrite sepit_emp //= => k. by rewrite /utab/table !ffunE; split=>//; case=>_ ->. Qed. From cc4be9803ba13013abf90afd58256eb914c8d78a Mon Sep 17 00:00:00 2001 From: Aleks Nanevski Date: Tue, 13 Aug 2024 03:20:00 +0200 Subject: [PATCH 44/93] cleanup --- examples/congprog.v | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/examples/congprog.v b/examples/congprog.v index 65b1f37..cefa1f0 100644 --- a/examples/congprog.v +++ b/examples/congprog.v @@ -150,9 +150,9 @@ case: decP=>[{}pf|] /=; last first. split=>//; exists f, hc, hct; split=>//. apply: tableP2 Hct=>// x; rewrite !in_set Ekn. by rewrite /n cardE index_mem /index mem_enum. -step=>x; step; rewrite -!joinA !(xpull _ hc). -apply: [stepX f]@hc=>//= [[m]] Em; apply: [gE]=>[||??[]//] //=. -split=>//; eexists [ffun z => if z == ith k pf then x else f z]. +step=>x; step; apply: [stepX f]@hc=>//= [[m]] Em. +apply: [gE]=>[||?? []] //=; split=>//; rewrite !(xpull _ m). +eexists [ffun z => if z == ith k pf then x else f z]. hhauto; rewrite (sepitS (ith k pf)) in_set indx_ith ltnSn. rewrite /ctab/table !ffunE eq_refl joinA; hhauto. apply: tableP2 Hct=>// s. From db551c3578e9f7ab2dbc70197a6ef6208c797e52 Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Tue, 13 Aug 2024 19:40:15 +0200 Subject: [PATCH 45/93] blah --- examples/congprog.v | 470 ++++++++++++++++++++++++++++++-------------- examples/kvmaps.v | 4 +- 2 files changed, 319 insertions(+), 155 deletions(-) diff --git a/examples/congprog.v b/examples/congprog.v index cefa1f0..aa65b8b 100644 --- a/examples/congprog.v +++ b/examples/congprog.v @@ -7,7 +7,7 @@ From Coq Require SetoidTactics. From mathcomp Require Import ssreflect ssrbool ssrnat. From mathcomp Require Import eqtype ssrfun div finset seq fintype finfun. From mathcomp Require Import choice. -From pcm Require Import axioms ordtype finmap pred pcm unionmap heap autopcm. +From pcm Require Import axioms prelude ordtype finmap pred pcm unionmap heap autopcm automap. From htt Require Import options model heapauto llist array. From htt Require Import kvmaps hashtab congmath. Import Prenex Implicits. @@ -34,7 +34,10 @@ Proof. exact: ltn_pmod. Qed. Definition hash10 k : 'I_10 := Ordinal (@hash_ord 10 k erefl). -Definition LT := locked (HT V hash10). +Definition LT := HT V hash10. + +Definition hash_shape : KVmap.tp LT -> finMap K V -> Pred heap := + @HashTab.shape K _ (AssocList.AssocList K V) 10 hash10. (* the tables relating arrays with the content of the lists *) (* ctab is for class lists, utab is for use lists *) @@ -59,17 +62,18 @@ Section ShapePredicates. Variable (r : {array symb -> symb}). Variable (clist : {array symb -> llist symb}). Variable (ulist : {array symb -> llist (symb*symb*symb)}). -Variable (htab : KVmap.tp LT). +Variable (htab : KVmap.tp LT). Variable (p : ptr). (* The layout of the data structure *) Definition ashape D := [Pred h | let: (d, ct, ut) := D in - h \In Array.shape r (rep d : {ffun symb -> symb}) # + h \In + Array.shape r (rep d : {ffun symb -> symb}) # (Array.shape clist ct # sepit setT (ctab ct (class d))) # (Array.shape ulist ut # sepit setT (utab ut (use d))) # - KVmap.shape htab (lookup d) # + hash_shape htab (lookup d) # [Pred h' | exists l, h' \In Pred1 (p :-> l) # lseq l (pending d)]]. Definition bshape d := @@ -178,201 +182,361 @@ by rewrite /utab/table !ffunE; split=>//; case=>_ ->. Qed. -UP TO HERE - - Variable (r : {array symb -> symb}). Variable (clist : {array symb -> llist symb}). Variable (ulist : {array symb -> llist (symb*symb*symb)}). Variable (htab : KVmap.tp LT). -Variable (p : loc). +Variable (p : ptr). Notation ashape' := (ashape r clist ulist htab p). Notation bshape' := (bshape r clist ulist htab p). Notation shape' := (shape r clist ulist htab p). Definition cT (a' b' : symb) : Type := - forall x:unit, STsep unit - (fun i => (exists D, i \In ashape' D) /\ a' != b', - fun y i m => forall D, i \In ashape' D -> y = Val tt /\ exists ct, exists ut, + forall x:unit, STsep {D} + (fun i => i \In ashape' D /\ a' != b', + fun y m => y = Val tt /\ exists ct, exists ut, let: (d, _, _) := D in m \In ashape' - (Data [ffun s => if s \in (class d a') then b' + (Data [ffun s => if s \in class d a' then b' else rep d s] [ffun s => if s == a' then [::] else if s == b' then rev (class d a') ++ class d b' else class d s] (use d) (lookup d) (pending d), ct, ut)). -Program -Definition join_hclass (a' b' : symb) : - STsep unit (fun i => (exists d, i \In bshape' d) /\ a' != b', - fun y i m => forall d, i \In bshape' d -> - y = Val tt /\ - m \In bshape' (join_class d a' b')) := - Do (Fix (fun (loop : cT a' b') (x : unit) => +Program Definition join_hclass (a' b' : symb) : + STsep {d} (fun i => i \In bshape' d /\ a' != b', + fun y m => y = Val tt /\ + m \In bshape' (join_class d a' b')) := + Do (ffix (fun (loop : cT a' b') (x : unit) => Do (a <-- Array.read clist a'; b <-- Array.read clist b'; - If a == null :> loc then ret tt + if a == null then ret tt else s <-- !a; - next <-- !(a .+ 1); - a .+ 1 ::= b;; + next <-- !a.+1; + a.+1 ::= b;; Array.write clist b' a;; Array.write clist a' next;; Array.write r s b';; loop tt)) tt). Next Obligation. -apply: (ghost (d, f0, f)) H1. -move: H0=>{loop f0 f d} N [[d]] ct ut H. -move: (H)=>[i1][w][->][R][[w']][i4][->{w}][[i2]][i3][->{w'}][Ct][Cv] _ [R'] _ D. -move: Cv; rewrite (sepitT1 a') (sepitS b') 3!in_set eq_sym N {1 2}/table /= => Cv. -move: Cv D=>[i5][w][->{i3}][Ca][[i3]][i6][->{w}][Cb][Cc] _ _. -rewrite -(unA i2) -(unCA i2); apply: bnd_gh (Ct)=>[a w [[->]] <-{w}|??[] //]. -apply: bnd_gh (Ct)=>[b w [[->]] <-{w}|??[] //]; rewrite (unCA i2) (unA i2). -case: ifP Ca=>E. -- rewrite (eqP E); case/lseq_null=>Ce -> D. - apply: val_ret=>//; split=>//; exists ct; exists ut. - rewrite InE /= (sepitT1 a') (sepitS b') 3!in_set eq_sym N {1 2}/table /= - (eqP E) Ce /= !ffunE !eq_refl eq_sym (negbTE N) /= - (_ : [ffun s => rep d s] = rep d); last first. - - by rewrite -ffunP => ?; rewrite ffunE. - hauto D; apply: tableP Cc=>?; rewrite !in_set // !ffunE. - by case: ifP=>W; [rewrite (eqP W) Ce | case: eqP]. -case/(lseq_pos (negbT E))=>s [q][i7][Ca] _ [<-{i5}] /= Cq. -rewrite -(unA i2) -2!(unA (ct a' :-> _)) -2!(unA (ct a' .+ 1 :-> _)) - -(unA i7) -(unA i3). -rewrite -2!(push (ct a'))=>D; apply: bnd_read=>//; rewrite 2!(push (ct a')) in D *. -rewrite -3!(push (ct a' .+ 1)) in D *; apply: bnd_read=>//. -apply: bnd_write=>{D} //; rewrite 3!(push (ct a' .+ 1)) -(unCA i2). -do 2![apply: {i2} bnd_gh Ct=>[[] i2 [Ct _]|[?]?[]] //]; rewrite (unCA i2). -apply: {i1} bnd_gh R=>[[] i1 [R _]|[?]?[]] //. -rewrite (unA i3) (unA i7) 2!(unA (ct a' .+ 1 :-> _)) 2!(unA (ct a' :-> _)) (unA i2). -set r2 := (finfun _ )in R; set ct2 := (finfun _) in Ct=>D. +move=>a' b' loop [[[[/= d ct] ut]]][i][H N] /=. +case: H=>/= rh [_][->{i} Rh][_][h][->][th][ctx][->] Th Ctx H. +rewrite (sepitT1 a') in Ctx; case: Ctx=>cta [w][->{ctx}Cta]. +rewrite (sepitS b') !in_set in_set1 eq_sym {1}N /=. +case=>ctb [ctx][->{w}Ctb Ctx]; rewrite /ctab/table in Cta Ctb. +apply: [stepX ct, th]@th=>//= _ _ [->->]. +apply: [stepX ct, th]@th=>//= _ _ [->->]. +apply: vrfV=>V; case: (ct a' =P null) Cta=>[/[dup] Ea' ->|/eqP Na']. +- case/(lseq_null (validX V))=>/= ->->{cta V}. + step=>{}V; split=>//; exists ct, ut. + rewrite (_ : rh \+ _ = rh \+ (th \+ (Unit \+ + (ctb \+ ctx)) \+ h)); last by heap_congr. + rewrite -fin_eta; hhauto; rewrite (sepitT1 a'); hhauto. + - by rewrite /ctab/table ffunE eqxx /= /= Ea'. + rewrite (sepitS b') !in_set in_set1 eq_sym N; hhauto. + - by rewrite /ctab/table /= ffunE eqxx eq_sym (negbTE N). + apply: tableP Ctx=>// x; rewrite !in_set !in_set1 andbT ffunE. + by case/andP=>/negbTE -> /negbTE ->. +case/(lseq_pos Na')=>s {V} [next][cta'][Eca ->{cta}Cta']. +do 3!step; apply: [stepX ct]@th=>//= [_] {th Th} th1 Th1. +set ct1 := (finfun _) in Th1. +apply: [stepX ct1]@th1=>//= [_] {th1 Th1} th2 Th2. +set ct2 := (finfun _) in Th2. +apply: [stepX rep d]@rh=>//= _ {rh Rh} rh1 Rh1. +set r1 := (finfun _) in Rh1. set cv2 := [ffun z => if z == a' then (behead (class d a')) - else if z == b' then s :: (class d b') else class d z]. -apply: (cons_gh1 (Data r2 cv2 (use d) (lookup d) (pending d), ct2, ut))=> - [P|[] m [_][ct1][ut1] /= Wm Dm|??[]||] //=; last 1 first. -- hauto D; rewrite /ct2 (sepitT1 a') (sepitS b') 3!in_set eq_sym N /= /table /= - !ffunE !eq_refl !(eq_sym b') (negbTE N) -!unA -!(unCA i7) !unA unC -!unA unC -unA. - by hauto D; apply: tableP Cc=>?; rewrite !in_set !ffunE; do 2![case: eqP=>//]. -- by split=>//; first by eauto. -split=>{ct2 Ct D} //; exists ct1; exists ut1. -set r1 := (finfun _) in Wm; set c1 := (finfun _) in Wm. -rewrite (_ : (finfun _) = r1); last first. -- rewrite /r1 -ffunP=>?; rewrite !ffunE eq_refl. - case: eqP=>[->|]; first by rewrite if_same Ca inE eq_refl. - by rewrite Ca inE; case: eqP. -rewrite (_ : (finfun _) = c1) //. -rewrite /c1 -ffunP=>?; rewrite !ffunE !eq_refl !(eq_sym b') (negbTE N). -case: eqP=>// H1; case: eqP=>// H2. -by rewrite Ca rev_cons cat_rcons. + else if z == b' then s :: class d b' else class d z]. +apply: [gE (Data r1 cv2 (use d) (lookup d) (pending d), ct2, ut)]=>/=; +last by move=>?? []. +- rewrite (_ : rh1 \+ _ = rh1 \+ (th2 \+ (cta' \+ (ct a' :-> s \+ + ((ct a').+1 :-> ct b' \+ ctb) \+ ctx)) \+ h)); last by heap_congr. + hhauto; rewrite (sepitT1 a'); hhauto. + - by rewrite /ctab/table/ct2/cv2 !ffunE /= eqxx. + rewrite (sepitS b') !in_set in_set1 eq_sym N; hhauto. + - rewrite /ctab/table/ct2/cv2/ct1 !ffunE /= ffunE /= eqxx; + by rewrite eq_sym (negbTE N); hhauto. + apply: tableP Ctx=>x; rewrite /ct2/cv2/ct1 !ffunE /= ?ffunE /=; + by rewrite !in_set !in_set1 andbT; case/andP=>/negbTE -> /negbTE ->. +case=>m [_][] {Th2}ct2 [ut2] /= Hm _; split=>//; exists ct2, ut2. +congr (m \In ashape' (Data _ _ _ _ _, ct2, ut2)): Hm; apply/ffunP=>x. +- by rewrite !ffunE eqxx {2}Eca inE /=; case: (x =P s)=>//= _; rewrite if_same. +rewrite !ffunE !eqxx /= (eq_sym b') (negbTE N). +case: (x =P a')=>// _; case: (x =P b')=>// _. +by rewrite Eca rev_cons cat_rcons. Qed. Next Obligation. -apply: (ghost H) H1; move: H0 {H}=> N d [I][ct][ut] H; move: (ashapeD H). -apply: (cons_gh1 (d, ct, ut))=>[||??[]|] //; first by eauto. -move=>[] m [_] [ct1][ut1] W; split=>//; set d' := (Data _ _ _ _ _) in W. -suff E : join_class d a' b' = d' - by split; [apply: join_class_classP | rewrite E; eauto]. -congr Data. -by set v := (finfun _); rewrite -(ffunP v) /v => x; rewrite !ffunE /= I. +move=>a' b' [d][i][[C]][/= ct][ut H] N. +apply: [gE (d, ct, ut)]=>[||?? []] //= [m][_][ct1][ut1] W _. +set d' := (Data _ _ _ _ _) in W; suff E : join_class d a' b' = d'. +- by split=>//; split; [apply: join_class_classP|rewrite E; eauto]. +by congr Data; apply/ffunP=>x; rewrite !ffunE /= C. Qed. -Module Dummy23. End Dummy23. - -Let proj0 (e : symb*symb*symb) := let: (c, c1, c2) := e in c. -Let proj1 (e : symb*symb*symb) := let: (c, c1, c2) := e in c1. -Let proj2 (e : symb*symb*symb) := let: (c, c1, c2) := e in c2. -Notation "e ..0" := (proj0 e) (at level 2). -Notation "e ..1" := (proj1 e) (at level 2). -Notation "e ..2" := (proj2 e) (at level 2). - +Check HashTab.lookup. Definition uT (a' b' : symb) := forall x : unit, - STsep unit (fun i => exists d, exists done, a' != b' /\ - i \In bshape' (join_use' d a' b' done) /\ - use d a' = done ++ use (join_use' d a' b' done) a', - fun y i m => forall d, i \In bshape' d -> y = Val tt /\ - m \In bshape' (join_use d a' b')). + STsep {d} (fun i => exists don, a' != b' /\ + i \In bshape' (join_use' d a' b' don) /\ + use d a' = don ++ use (join_use' d a' b' don) a', + fun y m => y = Val tt /\ + m \In bshape' (join_use d a' b')). Program Definition join_huse (a' b' : symb) : - STsep unit (fun i => exists d, a' != b' /\ i \In bshape' d, - fun y i m => forall d, i \In bshape' d -> - y = Val tt /\ m \In bshape' (join_use d a' b')) := - Do (Fix (fun (loop : uT a' b') x => + STsep {d} (fun i => a' != b' /\ i \In bshape' d, + fun y m => y = Val tt /\ + m \In bshape' (join_use d a' b')) := + Do (ffix (fun (loop : uT a' b') x => Do (a <-- Array.read ulist a'; - If a == null :> loc then ret tt + if a == null then ret tt else eq <-- !a; - next <-- !(a .+ 1); + next <-- !a.+1; Array.write ulist a' next;; - c1 <-- Array.read r eq..1; - c2 <-- Array.read r eq..2; - v <-- KVmap.lookup htab (c1, c2); - match_opt v then - KVmap.insert htab (c1, c2) eq;; - b <-- Array.read ulist b'; - a .+ 1 ::= b;; - Array.write ulist b' a;; - loop tt - else [d] + c1 <-- Array.read r eq.1.2; + c2 <-- Array.read r eq.2; + v <-- HashTab.lookup hash10 htab (c1, c2); + if v is Some d then dealloc a;; - dealloc a .+ 1;; + dealloc a.+1;; p' <-- !p; q <-- insert p' (comp_pend eq d); p ::= q;; + loop tt + else + HashTab.insert hash10 htab (c1, c2) eq;; + b <-- Array.read ulist b'; + a.+1 ::= b;; + Array.write ulist b' a;; loop tt)) tt). Next Obligation. -apply: (ghost1 (join_use' H a' b' H0))=>[?|]; first by apply: bshape_inv. -move: H H0 H1 H2 H3=>d done N. -set d1 := join_use' d a' b' done. -move=>[C1][ct][ut][h1][w][->][H1][[h2]][w'][->{w}][H2][[w]][w''][->{w'}]; -move=>[[h3]][u][->{w}][Ut][U] _ [[h7]][w][->{w''}][H7][[l]][w'][h8][->{w}][]; -case/swp=>->{w'} _ [H8] _ _ _ _ D E; move: D. -move: U; rewrite (sepitT1 a') (sepitS b') 3!in_set eq_sym N {1 2} /table /=. -move=>[h4][w][->{u}][Ua][[h5]][h6][->{w}][Ub][Ru] _ _. -rewrite -(unA h3) -2!(unCA h3); apply: bnd_gh (Ut)=>[_ _ [[->]] <-|??[]] //. -case: ifP=>E1. -- rewrite (eqP E1) in Ua; case/lseq_null: Ua=>E2 ->. - rewrite 2!(unCA h3) (unA h3) /join_use E2=>D; apply: val_ret=>//; do !split=>//. - exists ct; exists ut; hauto D. - by rewrite (sepitT1 a') (sepitS b') 3!in_set eq_sym N /table (eqP E1) E2; hauto D. -case/(lseq_pos (negbT E1)): Ua. -move=>[[c1 c2] c][next][h9][E2] _ [<-{h4}] H9. -rewrite -2!(unA (ut a' :-> _)) -3!(push (ut a'))=>D; apply: bnd_read=>//. -rewrite -2!(unA (ut a' .+ 1 :-> _)) -4!(push (ut a' .+ 1)) in D *. -apply: bnd_read=>//; move: D; rewrite -2!(unCA h3). -apply: {h3} bnd_gh Ut=>[[] h3 [Ut _]|??[] //]; rewrite -3!(unCA h1). -apply: bnd_gh (H1)=>[c1' _ [[Ec1]] <-|??[] //]. -apply: bnd_gh (H1)=>[c2' _ [[Ec2]] <-|??[] //]. -rewrite -(unA h9) -(unA h5) -8!(unCA h7). -apply: {h7} bnd_gh H7=>[v h7 [H7][Eqv]|??[] //]. -set d2 := join_use' d a' b' (done ++ [:: (c1, c2, c)]). -have E3: use d2 a' = behead (use d1 a'). -- rewrite /d2 (@join_useT (behead (use d1 a'))); last by rewrite /= -E2. - by apply: join_use_useE; rewrite /= -E2. +move=>a' b' loop [[/= d]][i][don][N][H Eu]. +case: H=>C [/= ct][ut][h1][w][->{i} H1][h2][w'][->{w} H2][w][w''][->{w'}]. +case=>h3 [u][->{w} Ut U] H. +case: H=>h7 [wx][->{w''} H7 H8]. +case: H8=>p' [h8][h9][->{wx} <-{h8} /= H]. +set d1 := join_use' d a' b' don in Eu C H1 H2 U H7 H. + +move: U. rewrite (sepitT1 a'). +case=>h4 [w][->{u} Ua]. +rewrite (sepitS b') !in_set !in_set1 eq_sym {1}N /=. +case=>h5 [h6][->{w} Ub Ru]. +apply: [stepX ut, h3]@h3=>//= _ _ [->->]. +rewrite /utab/table/= in Ua. +apply: vrfV=>V. +case: (ut a' =P null) Ua=>[/[dup] Ea' ->|/eqP Na']. +- case/(lseq_null (validX V))=>/= {V} E2 ->{h4}. + step=>_. split=>//. + rewrite (_ : h1 \+ _ = + h1 \+ (h2 \+ (h3 \+ (Unit \+ (h5 \+ h6)) \+ (h7 \+ (p:-> p' \+ h9))))); last by heap_congr. + rewrite /join_use /=. rewrite /class_inv /= in C. + rewrite E2 cats0 in Eu. + rewrite Eu -/d1. + split=>//=. exists ct, ut. + hhauto; rewrite (sepitT1 a'). hhauto. + - by rewrite /utab/table Ea' E2. + rewrite (sepitS b') !in_set !in_set1 eq_sym N /=. + by hhauto. + +case/(lseq_pos Na')=>[[[c1 c2 c]]][nxt][h10][E2] ->{V h4} H10. +step. step. +apply: [stepX ut]@h3=>//=. +case=>{Ut}h3 Ut. +apply: [stepX rep d1, h1]@h1=>//= _ _ [->->]. +apply: [stepX rep d1, h1]@h1=>//= _ _ [->->]. +apply: [stepX lookup d1]@h7=>//=. + +move=>v h7' [H7' Eqv]. +set d2 := join_use' d a' b' (don ++ [:: (c1, c2, c)]). + +have E3 : use d2 a' = behead (use d1 a'). +- rewrite /d2. + rewrite (join_useT (t:=behead (use d1 a'))) //; last first. + - by rewrite /= -E2 -Eu. + apply: join_use_useE. simpl. by rewrite E2 /=. + have E4: join_use d2 a' b' = join_use d1 a' b'. -- by rewrite /join_use E3 -!(@join_useT [::]) ?cats0 -?catA ?E /= -?E2. -case: v Eqv=>[[[e1 e2] e]|] /= Eqv. -- rewrite -4!(push (ut a')); apply: bnd_dealloc. - rewrite -3!(push (ut a' .+ 1)); apply: bnd_dealloc. - rewrite -7!(push p)=>D; apply: bnd_read=>//; move: D. - rewrite -(unC h8) -7!(unCA h8). - apply: bnd_gh H8=>{h8 x0} [_ h8 [q][H8][->]|??[?][]//]. - rewrite -(push p); apply: bnd_write=>D. - apply: (cons_gh1 d2)=>[?|[] m [_] P _|??[]||] //. - - by exists d; exists (done ++ [:: (c1, c2, c)]); rewrite -/d2 E3 -catA /= -E2. - - by rewrite -E4. - rewrite (_ : _ :+ _ = - h1 :+ (h2 :+ (h3 :+ (h9 :+ (h5 :+ h6)) :+ (h7 :+ (p :-> q :+ h8))))) in D *; - last by heap_congr. - case: H8 D=>x0 [h'] [_] <- H8 D. - rewrite /d2 (@join_useT (behead (use d1 a'))) -/d1 /= -?E2 // -Ec1 -Ec2 -Eqv; split=>//. +- rewrite /join_use E3. rewrite /d2. + rewrite -!(join_useT (t:=[::])) /=. + - congr join_use'. + rewrite {2}E2 /= -/d1 //. by rewrite -catA //=. + - by rewrite cats0. + by rewrite cats0 -catA /= -E2 -Eu. + +case: v Eqv=>[[[e1 e2 e3]]|] /= Eqv. + +- have E5 : rep d2 = rep d1. + - rewrite /d2 (join_useT (t:=behead (use d1 a'))) -/d1 /=. + - by rewrite -Eqv /=. + by rewrite Eu E2. + have E6 : class d2 = class d1. + - rewrite /d2 (join_useT (t:=behead (use d1 a'))) -/d1 /=. + - by rewrite -Eqv /=. + by rewrite Eu E2. + have E7 x : x != a' -> use d2 x = use d1 x. + - rewrite /d2. move=>Na. + rewrite (join_useT (t:=behead (use d1 a'))) /= -/d1. + - by rewrite -Eqv /= ffunE (negbTE Na). + by rewrite Eu E2. + have E8 : lookup d2 = lookup d1. + - rewrite /d2 (join_useT (t:=behead (use d1 a'))) /= -/d1. + - by rewrite -Eqv //=. + by rewrite Eu E2. + have E9 : pending d2 = comp_pend (c1, c2, c) (e1, e2, e3) :: pending d1. + - rewrite /d2 (join_useT (t:=behead (use d1 a'))) /= -/d1. + - by rewrite -Eqv //=. + by rewrite Eu E2. + + step. step. step. + apply: [stepX pending d1]@h9=>//=. + move=>q h11 [r][{H}h9][->{h11} H]. + step. + apply: [gE d]=>[||??[]] //=. + exists (don ++ [:: (c1, c2, c)]). split=>//. + split; last first. + - by rewrite -catA /= E3 -E2 -Eu. + rewrite -/d2. + econstructor=>//=. + - by rewrite /class_inv /= E5 E6. + set ut2 := (finfun _) in Ut. - exists ct; exists ut2; hauto D. - rewrite (sepitT1 a') (sepitS b') 3!in_set eq_sym N /table /=. - rewrite !ffunE !eq_refl !(eq_sym b') (negbTE N); hauto D. - by apply: tableP Ru=>s; rewrite !in_set /ut2 ffunE; - case: (s == a')=>//; rewrite andbF. + exists ct, ut2. + rewrite (_ : _ \+ _ = h1 \+ (h2 \+ ((h3 \+ (h10 \+ (h5 \+ h6))) \+ (h7' \+ + (p :-> q \+ (q :-> comp_pend (c1, c2, c) (e1, e2, e3) \+ + (q.+1 :-> r \+ h9))))))); last by heap_congr. + econstructor=>//=. + econstructor=>//=. + econstructor=>//=. + - by rewrite E5. + econstructor=>//=. + econstructor=>//=. + econstructor=>//=. + - by rewrite E6. + econstructor=>//=. + econstructor=>//=. + econstructor=>//=. + - econstructor=>//=. + econstructor=>//=. + econstructor=>//=. + - rewrite (sepitT1 a'). exists h10, (h5 \+ h6). split=>//. + - by rewrite /utab/table E3/ut2 ffunE /= eqxx. + rewrite (sepitS b') !in_set !in_set1 eq_sym N /=. + eexists h5, h6. split=>//. + - rewrite /utab/table/ut2 ffunE /= eq_sym (negbTE N) /=. + by rewrite E7 // eq_sym //. + apply: tableP Ru=>x //=; + rewrite !in_set !in_set1 andbT /d2/ut2 ?ffunE /=. + - by case/andP=>_ /negbTE ->. + by case/andP=>_ X; rewrite E7 //. + econstructor=>//=. + econstructor=>//=. + econstructor=>//=. + - by rewrite E8. + econstructor=>//. + econstructor=>//. + econstructor=>//. + econstructor=>//. rewrite E9 /=. + rewrite /lseq /= InE. simpl. + econstructor=>//. + by econstructor=>//. + +apply: [stepX lookup d1]@h7'=>//=. +move=>htab' h7'' H7''. + +set ut2 := (finfun _) in Ut. + +apply: [stepX ut2, h3]@h3=>//= _ _ [->->]. +step. +apply: [stepX ut2]@h3=>//=. +case=>{Ut}h3 Ut. +set ut3 := (finfun _) in Ut. +apply: [gE d]=>[||??[]] //=. +exists (don++[:: (c1, c2, c)]). +split=>//=. +split; last first. +- by rewrite Eu -catA /= E2 E3. +rewrite -/d2. + +have E5 : rep d2 = rep d1. +- rewrite /d2 (join_useT (t:=behead (use d1 a'))) -/d1 /=. + - by rewrite -Eqv /=. + by rewrite Eu E2. +have E6 : class d2 = class d1. + - rewrite /d2 (join_useT (t:=behead (use d1 a'))) -/d1 /=. + - by rewrite -Eqv /=. + by rewrite Eu E2. +have E7 x : x != a' -> use d2 x = + if x == b' then (c1, c2, c) :: use d1 x else use d1 x. +- rewrite /d2. move=>Na. + rewrite (join_useT (t:=behead (use d1 a'))) /= -/d1. + - by rewrite -Eqv /= ffunE (negbTE Na) //. + by rewrite Eu E2. + +have E8 : lookup d2 = ins (rep d1 c2, rep d1 c) (c1, c2, c) (lookup d1). +- rewrite /d2 (join_useT (t:=behead (use d1 a'))) /= -/d1. + - by rewrite -Eqv //=. + by rewrite Eu E2. +have E9 : pending d2 = pending d1. +- rewrite /d2 (join_useT (t:=behead (use d1 a'))) /= -/d1. + - by rewrite -Eqv //=. + by rewrite Eu E2. + +rewrite (_ : _ \+ _ = h1 \+ (h2 \+ ((h3 \+ (h10 \+ + ((ut a' :-> (c1, c2, c) \+ + ((ut a').+1 :-> ut2 b' \+ h5)) \+ h6))) \+ (h7'' \+ (p :-> p' \+ h9))))); last by heap_congr. + +econstructor. +- by rewrite /class_inv /= E5 E6. +econstructor=>//=. +econstructor=>//=. +econstructor=>//=. +econstructor=>//=. +econstructor=>//=. +- by rewrite E5. +econstructor=>//=. +econstructor=>//=. +econstructor=>//=. +- by rewrite E6; exact: H2. +econstructor=>//=. +econstructor=>//=. +econstructor=>//=. +- econstructor=>//=. + econstructor=>//=. + econstructor=>//=. + - exact :Ut. + rewrite (sepitT1 a'). + econstructor=>//=. + econstructor=>//=. + econstructor=>//=. + - rewrite /utab/table/ut3 ffunE /= (negbTE N) /ut2 ffunE /= eqxx. + by rewrite E3. + rewrite (sepitS b') !in_set !in_set1 /= eq_sym N /=. + econstructor=>//=. + econstructor=>//=. + econstructor=>//=. + - rewrite /utab/table/=/ut3 !ffunE /= eqxx. + rewrite eq_sym (negbTE N). + rewrite E7; last first. by rewrite eq_sym (negbTE N). + rewrite eqxx /=. + rewrite InE /=. exists (ut b'). exists h5. split=>//. + apply: tableP Ru=>x /=; + rewrite !in_set !in_set1 andbT /ut3 ?ffunE /= /ut2 ?ffunE /=. + - by case/andP=>/negbTE -> /negbTE ->. + case/andP=>X1 X2. + by rewrite E7 // (negbTE X1). + +econstructor=>//=. +econstructor=>//=. +econstructor=>//=. +- rewrite E8 /=. + + + + + + + apply: bnd_gh H7=>{h7} [[] h7 [H7 _]|??[] //]; rewrite -2!(unCA h3). apply: bnd_gh (Ut)=>[_ _ [[->]] <-|??[] //]; rewrite -3!(push (ut a' .+ 1)). apply: bnd_write; rewrite ffunE (eq_sym b') (negbTE N) -(unCA h3). diff --git a/examples/kvmaps.v b/examples/kvmaps.v index caff2af..47e5a2a 100644 --- a/examples/kvmaps.v +++ b/examples/kvmaps.v @@ -35,10 +35,10 @@ Record Sig (K : ordType) (V : Type) : Type := [vfun _ : unit => emp]); insert : forall x k v, STsep {s} (shape x s, - [vfun y => shape y (ins k v s)]); + [vfun _ : unit => shape x (ins k v s)]); remove : forall x k, STsep {s} (shape x s, - [vfun y => shape y (rem k s)]); + [vfun _ : unit => shape x (rem k s)]); lookup : forall x k, STsep {s} (shape x s, [vfun y m => m \In shape x s /\ y = fnd k s])}. From 2ff85e2490e41e5813dfb5f63e22d94a6db07bbd Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Wed, 14 Aug 2024 13:16:19 +0200 Subject: [PATCH 46/93] ported proofs of congr_prog --- core/pred.v | 2 +- core/prelude.v | 49 ++- examples/array.v | 14 +- examples/congmath.v | 15 +- examples/congprog.v | 731 +++++++++++++++----------------------------- examples/hashtab.v | 48 +-- examples/kvmaps.v | 137 ++++++++- pcm/heap.v | 35 ++- pcm/pcm.v | 33 +- 9 files changed, 540 insertions(+), 524 deletions(-) diff --git a/core/pred.v b/core/pred.v index 55ed284..d989dc5 100644 --- a/core/pred.v +++ b/core/pred.v @@ -469,7 +469,7 @@ Add Parametric Morphism T (pT : PredType T) : (@Mem T pT) with signature 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. + @EqPredType _ _ ==> @EqPredType _ _ ==> @EqPredType _ _ as predU_morph. Proof. move=>r1 s1 H1 r2 h2 H2 x; split; by case; [move/H1 | move/H2]=>/=; auto. diff --git a/core/prelude.v b/core/prelude.v index f286e8f..b605d0d 100644 --- a/core/prelude.v +++ b/core/prelude.v @@ -19,7 +19,7 @@ limitations under the License. From HB Require Import structures. From Coq Require Import ssreflect ssrfun Eqdep. From mathcomp Require Import ssrbool ssrnat seq eqtype choice. -From mathcomp Require Import fintype finfun tuple perm fingroup. +From mathcomp Require Import fintype finset finfun tuple perm fingroup. From pcm Require Import options axioms. (***********) @@ -1128,6 +1128,53 @@ Proof. by apply/ffunP => i; rewrite !ffunE permM. Qed. End PermFfun. +(* Finite sets *) + +Lemma enum_T (I : finType) : enum setT = enum I. +Proof. by rewrite enum_setT enumT. Qed. + +Lemma enum_0 (I : finType) (s : {set I}) : + s =i set0 -> + enum s = [::]. +Proof. by move/eq_enum=>->; rewrite enum_set0. Qed. + +Lemma setTE (I : finType) (p : pred I) : + p =1 xpredT -> + [set x | p x] = setT :> {set I}. +Proof. by move=>H; apply/setP=>x; rewrite !inE H. Qed. + +Lemma set0E (I : finType) (p : pred I) : + p =1 xpred0 -> + [set x | p x] = set0 :> {set I}. +Proof. by move=>H; apply/setP=>x; rewrite !inE H. Qed. + +(* streamlining for ordinals *) + +Lemma set_ord0N m (p : pred nat) : + (forall x, x < m -> ~~ p x) -> + [set x | p (\val x)] = set0 :> {set 'I_m}. +Proof. by move=>H; apply: set0E; case=>z pf; rewrite (negbTE (H z pf)). Qed. + +Lemma set_ord0 m (p : pred nat) : + (forall x, p x -> m <= x) -> + [set x | p (\val x)] = set0 :> {set 'I_m}. +Proof. +move=>H; apply: set_ord0N=>x; rewrite ltnNge. +by apply/contra/H. +Qed. + +Lemma set_ordT m (p : pred nat) : + (forall x, x < m -> p x) -> + [set x | p (\val x)] = setT :> {set 'I_m}. +Proof. by move=>H; apply: setTE; case=>z pf; rewrite H. Qed. + +Lemma set_ordTN m (p : pred nat) : + (forall x, ~~ p x -> m <= x) -> + [set x | p (\val x)] = setT :> {set 'I_m}. +Proof. +move=>H; apply/set_ordT=>x; rewrite ltnNge. +by apply/contraR/H. +Qed. (* Tagging *) diff --git a/examples/array.v b/examples/array.v index c7338e8..466e7ac 100644 --- a/examples/array.v +++ b/examples/array.v @@ -50,7 +50,7 @@ Definition shape (a : array) (f : {ffun I -> T}) := Program Definition new (x : T) : STsep (emp, [vfun a => shape a [ffun => x]]) := Do (x <-- allocb x #|I|; - ret (Array x)). + ret (Array x)). Next Obligation. (* pull out ghost vars, run the program *) move=>x [] _ /= ->; step=>y; step. @@ -83,7 +83,7 @@ Program Definition newf (f : {ffun I -> T}) : Do (if s is k::t return _ then x .+ (indx k) ::= f k;; loop t - else ret (Array x))) + else ret x)) in fill (enum I) else ret (Array null)). (* first the loop *) @@ -123,7 +123,7 @@ Qed. (* the loop invariant: *) (* a partially freed array still contains valid #|I| - k cells *) (* corresponding to some suffix xs of the original array's spec *) -Definition free_loop (a : array) : Type := +Definition free_loop (a : ptr) : Type := forall k, STsep (fun i => exists xs: seq T, [/\ i = updi (a .+ k) xs, valid i & @@ -192,7 +192,7 @@ End Array. Section Table. -Variables (I : finType) (T S : Type) (x : {array I -> T}) +Variables (I : finType) (T S : Type) (x : ptr) (Ps : T -> S -> Pred heap). Definition table (t : I -> T) (b : I -> S) (i : I) : Pred heap := @@ -201,7 +201,8 @@ Definition table (t : I -> T) (b : I -> S) (i : I) : Pred heap := Lemma tableP (s : {set I}) t1 t2 b1 b2 h : (forall x, x \in s -> t1 x = t2 x) -> (forall x, x \in s -> b1 x = b2 x) -> - h \In sepit s (table t1 b1) -> h \In sepit s (table t2 b2). + h \In sepit s (table t1 b1) -> + h \In sepit s (table t2 b2). Proof. move=>H1 H2 H. apply/(sepitF (f1 := table t1 b1))=>//. @@ -212,7 +213,8 @@ Lemma tableP2 (s1 s2 : {set I}) t1 t2 b1 b2 h : s1 =i s2 -> (forall x, s1 =i s2 -> x \in s1 -> t1 x = t2 x) -> (forall x, s1 =i s2 -> x \in s1 -> b1 x = b2 x) -> - h \In sepit s1 (table t1 b1) -> h \In sepit s2 (table t2 b2). + h \In sepit s1 (table t1 b1) -> + h \In sepit s2 (table t2 b2). Proof. move=>H1 H2 H3; rewrite (eq_sepit _ H1). by apply: tableP=>y; rewrite -H1=>R; [apply: H2 | apply: H3]. diff --git a/examples/congmath.v b/examples/congmath.v index 834b256..51cf72b 100644 --- a/examples/congmath.v +++ b/examples/congmath.v @@ -12,12 +12,10 @@ limitations under the License. *) From HB Require Import structures. -From Coq Require Import Recdef ssreflect ssrbool ssrfun. +From Coq Require Import Recdef Setoid ssreflect ssrbool ssrfun. From mathcomp Require Import eqtype choice ssrnat seq bigop fintype finfun. From pcm Require Import options prelude ordtype finmap pred seqext. -Ltac add_morphism_tactic := SetoidTactics.add_morphism_tactic. - (**********************) (* Congruence closure *) (**********************) @@ -215,7 +213,9 @@ Hint Resolve reflC symC : core. (* lemmas about closure *) -Lemma closE (R1 R2 : rel_exp) : R1 <~> R2 -> closure R1 <~> closure R2. +Lemma closE (R1 R2 : rel_exp) : + R1 <~> R2 -> + closure R1 <~> closure R2. Proof. suff H: forall R1 R2, R1 <~> R2 -> closure R1 ~> closure R2. - by split; apply: H; [| symmetry]. @@ -1769,3 +1769,10 @@ Qed. End Congruence. Notation rel_exp s := (Pred (exp s * exp s)). + +(* repeat the morphism declaration outside the section *) +Add Parametric Morphism s : (@closure s) with + signature Morphisms.respectful (fun e1 e2 => e1 <~> e2) + (fun e1 e2 => e1 <~> e2) as closure_morph'. +Proof. by apply: closE. Qed. + diff --git a/examples/congprog.v b/examples/congprog.v index aa65b8b..a64f9ec 100644 --- a/examples/congprog.v +++ b/examples/congprog.v @@ -3,16 +3,16 @@ (*******************************) From HB Require Import structures. -From Coq Require SetoidTactics. -From mathcomp Require Import ssreflect ssrbool ssrnat. -From mathcomp Require Import eqtype ssrfun div finset seq fintype finfun. -From mathcomp Require Import choice. -From pcm Require Import axioms prelude ordtype finmap pred pcm unionmap heap autopcm automap. +From mathcomp Require Import ssreflect ssrbool ssrnat eqtype ssrfun. +From mathcomp Require Import div finset seq fintype finfun choice. +From pcm Require Import axioms prelude ordtype finmap pred pcm. +From pcm Require Import unionmap heap autopcm automap. From htt Require Import options model heapauto llist array. -From htt Require Import kvmaps hashtab congmath. +From htt Require Import kvmaps hashtab. +From htt Require Import congmath. Import Prenex Implicits. -Definition xpull := pcm.pull. +Notation finE := finset.inE. Ltac add_morphism_tactic := SetoidTactics.add_morphism_tactic. Notation " R ===> R' " := (@Morphisms.respectful _ _ R R') @@ -36,9 +36,6 @@ Definition hash10 k : 'I_10 := Ordinal (@hash_ord 10 k erefl). Definition LT := HT V hash10. -Definition hash_shape : KVmap.tp LT -> finMap K V -> Pred heap := - @HashTab.shape K _ (AssocList.AssocList K V) 10 hash10. - (* the tables relating arrays with the content of the lists *) (* ctab is for class lists, utab is for use lists *) @@ -73,27 +70,28 @@ Definition ashape D := Array.shape r (rep d : {ffun symb -> symb}) # (Array.shape clist ct # sepit setT (ctab ct (class d))) # (Array.shape ulist ut # sepit setT (utab ut (use d))) # - hash_shape htab (lookup d) # + KVmap.shape htab (lookup d) # [Pred h' | exists l, h' \In Pred1 (p :-> l) # lseq l (pending d)]]. Definition bshape d := [Pred h | class_inv d /\ exists ct ut, h \In ashape (d, ct, ut)]. Definition shape (R : rel_exp s) := - [Pred h | exists d, h \In bshape d /\ propagate_inv d /\ pending d = [::] /\ - CRel d =p R]. + [Pred h | exists d, h \In bshape d /\ propagate_inv d /\ + pending d = [::] /\ CRel d <~> R]. End ShapePredicates. -(* Add Parametric Morphism r clist ulist htab p : (shape r clist ulist htab p) with - signature @releq _ ===> @releq _ as shape_morph. + signature @EqPredType _ _ ===> @EqPredType _ _ as shape_morph. Proof. -move=>R1 R2 H. -split=>x [d1][H1][H2][H3] H4; exists d1; -by [rewrite -H | rewrite H]. +move=>R1 R2 H x /=. +split; case=>d [H1][H2][H3 H4]; exists d. +- by rewrite -H. +by rewrite H. Qed. -*) + + Definition ith {I : finType} i (pf : i < #|I|) : I. Admitted. @@ -152,16 +150,18 @@ move=>r clist loop k H i [pf][/= f][hc][hct][->{i} Hc Hct]. case: decP=>[{}pf|] /=; last first. - case: (ltngtP k n) pf=>// Ekn _ _; step=>_. split=>//; exists f, hc, hct; split=>//. - apply: tableP2 Hct=>// x; rewrite !in_set Ekn. + apply: tableP2 Hct=>// x; rewrite !finE Ekn. by rewrite /n cardE index_mem /index mem_enum. step=>x; step; apply: [stepX f]@hc=>//= [[m]] Em. -apply: [gE]=>[||?? []] //=; split=>//; rewrite !(xpull _ m). +apply: [gE]=>[||?? []] //=; split=>//. eexists [ffun z => if z == ith k pf then x else f z]. -hhauto; rewrite (sepitS (ith k pf)) in_set indx_ith ltnSn. -rewrite /ctab/table !ffunE eq_refl joinA; hhauto. +rewrite (_ : _ \+ _ = m \+ (x :-> ith k pf \+ + x.+1 :-> null \+ hct)); last by heap_congr. +hhauto; rewrite (sepitS (ith k pf)) finE indx_ith ltnSn. +rewrite /ctab/table !ffunE eqxx; hhauto. apply: tableP2 Hct=>// s. -- by rewrite !in_set ltnS in_set1 indx_injE; case: ltngtP. -by rewrite !in_set !ffunE indx_injE; case: eqP=>// ->; rewrite ltnn. +- by rewrite !finE ltnS indx_injE; case: ltngtP. +by rewrite !finE !ffunE indx_injE; case: eqP=>// ->; rewrite ltnn. Qed. Next Obligation. case=>_ ->; apply: [stepE]=>//= r hr Er; apply: [stepU]=>//= clist hc Ec. @@ -225,7 +225,7 @@ Next Obligation. move=>a' b' loop [[[[/= d ct] ut]]][i][H N] /=. case: H=>/= rh [_][->{i} Rh][_][h][->][th][ctx][->] Th Ctx H. rewrite (sepitT1 a') in Ctx; case: Ctx=>cta [w][->{ctx}Cta]. -rewrite (sepitS b') !in_set in_set1 eq_sym {1}N /=. +rewrite (sepitS b') !finE eq_sym {1}N /=. case=>ctb [ctx][->{w}Ctb Ctx]; rewrite /ctab/table in Cta Ctb. apply: [stepX ct, th]@th=>//= _ _ [->->]. apply: [stepX ct, th]@th=>//= _ _ [->->]. @@ -236,9 +236,9 @@ apply: vrfV=>V; case: (ct a' =P null) Cta=>[/[dup] Ea' ->|/eqP Na']. (ctb \+ ctx)) \+ h)); last by heap_congr. rewrite -fin_eta; hhauto; rewrite (sepitT1 a'); hhauto. - by rewrite /ctab/table ffunE eqxx /= /= Ea'. - rewrite (sepitS b') !in_set in_set1 eq_sym N; hhauto. + rewrite (sepitS b') !finE eq_sym N; hhauto. - by rewrite /ctab/table /= ffunE eqxx eq_sym (negbTE N). - apply: tableP Ctx=>// x; rewrite !in_set !in_set1 andbT ffunE. + apply: tableP Ctx=>// x; rewrite !finE andbT ffunE. by case/andP=>/negbTE -> /negbTE ->. case/(lseq_pos Na')=>s {V} [next][cta'][Eca ->{cta}Cta']. do 3!step; apply: [stepX ct]@th=>//= [_] {th Th} th1 Th1. @@ -255,11 +255,11 @@ last by move=>?? []. ((ct a').+1 :-> ct b' \+ ctb) \+ ctx)) \+ h)); last by heap_congr. hhauto; rewrite (sepitT1 a'); hhauto. - by rewrite /ctab/table/ct2/cv2 !ffunE /= eqxx. - rewrite (sepitS b') !in_set in_set1 eq_sym N; hhauto. + rewrite (sepitS b') !finE eq_sym N; hhauto. - rewrite /ctab/table/ct2/cv2/ct1 !ffunE /= ffunE /= eqxx; by rewrite eq_sym (negbTE N); hhauto. apply: tableP Ctx=>x; rewrite /ct2/cv2/ct1 !ffunE /= ?ffunE /=; - by rewrite !in_set !in_set1 andbT; case/andP=>/negbTE -> /negbTE ->. + by rewrite !finE andbT; case/andP=>/negbTE -> /negbTE ->. case=>m [_][] {Th2}ct2 [ut2] /= Hm _; split=>//; exists ct2, ut2. congr (m \In ashape' (Data _ _ _ _ _, ct2, ut2)): Hm; apply/ffunP=>x. - by rewrite !ffunE eqxx {2}Eca inE /=; case: (x =P s)=>//= _; rewrite if_same. @@ -275,7 +275,6 @@ set d' := (Data _ _ _ _ _) in W; suff E : join_class d a' b' = d'. by congr Data; apply/ffunP=>x; rewrite !ffunE /= C. Qed. -Check HashTab.lookup. Definition uT (a' b' : symb) := forall x : unit, STsep {d} (fun i => exists don, a' != b' /\ i \In bshape' (join_use' d a' b' don) /\ @@ -312,328 +311,132 @@ Program Definition join_huse (a' b' : symb) : loop tt)) tt). Next Obligation. move=>a' b' loop [[/= d]][i][don][N][H Eu]. -case: H=>C [/= ct][ut][h1][w][->{i} H1][h2][w'][->{w} H2][w][w''][->{w'}]. -case=>h3 [u][->{w} Ut U] H. -case: H=>h7 [wx][->{w''} H7 H8]. -case: H8=>p' [h8][h9][->{wx} <-{h8} /= H]. -set d1 := join_use' d a' b' don in Eu C H1 H2 U H7 H. - -move: U. rewrite (sepitT1 a'). -case=>h4 [w][->{u} Ua]. -rewrite (sepitS b') !in_set !in_set1 eq_sym {1}N /=. -case=>h5 [h6][->{w} Ub Ru]. -apply: [stepX ut, h3]@h3=>//= _ _ [->->]. -rewrite /utab/table/= in Ua. -apply: vrfV=>V. -case: (ut a' =P null) Ua=>[/[dup] Ea' ->|/eqP Na']. -- case/(lseq_null (validX V))=>/= {V} E2 ->{h4}. - step=>_. split=>//. - rewrite (_ : h1 \+ _ = - h1 \+ (h2 \+ (h3 \+ (Unit \+ (h5 \+ h6)) \+ (h7 \+ (p:-> p' \+ h9))))); last by heap_congr. - rewrite /join_use /=. rewrite /class_inv /= in C. - rewrite E2 cats0 in Eu. - rewrite Eu -/d1. - split=>//=. exists ct, ut. - hhauto; rewrite (sepitT1 a'). hhauto. - - by rewrite /utab/table Ea' E2. - rewrite (sepitS b') !in_set !in_set1 eq_sym N /=. - by hhauto. - -case/(lseq_pos Na')=>[[[c1 c2 c]]][nxt][h10][E2] ->{V h4} H10. -step. step. -apply: [stepX ut]@h3=>//=. -case=>{Ut}h3 Ut. -apply: [stepX rep d1, h1]@h1=>//= _ _ [->->]. -apply: [stepX rep d1, h1]@h1=>//= _ _ [->->]. -apply: [stepX lookup d1]@h7=>//=. - -move=>v h7' [H7' Eqv]. -set d2 := join_use' d a' b' (don ++ [:: (c1, c2, c)]). - -have E3 : use d2 a' = behead (use d1 a'). -- rewrite /d2. - rewrite (join_useT (t:=behead (use d1 a'))) //; last first. - - by rewrite /= -E2 -Eu. - apply: join_use_useE. simpl. by rewrite E2 /=. - -have E4: join_use d2 a' b' = join_use d1 a' b'. -- rewrite /join_use E3. rewrite /d2. - rewrite -!(join_useT (t:=[::])) /=. - - congr join_use'. - rewrite {2}E2 /= -/d1 //. by rewrite -catA //=. - - by rewrite cats0. - by rewrite cats0 -catA /= -E2 -Eu. - -case: v Eqv=>[[[e1 e2 e3]]|] /= Eqv. - -- have E5 : rep d2 = rep d1. - - rewrite /d2 (join_useT (t:=behead (use d1 a'))) -/d1 /=. - - by rewrite -Eqv /=. - by rewrite Eu E2. - have E6 : class d2 = class d1. - - rewrite /d2 (join_useT (t:=behead (use d1 a'))) -/d1 /=. - - by rewrite -Eqv /=. - by rewrite Eu E2. - have E7 x : x != a' -> use d2 x = use d1 x. - - rewrite /d2. move=>Na. - rewrite (join_useT (t:=behead (use d1 a'))) /= -/d1. - - by rewrite -Eqv /= ffunE (negbTE Na). - by rewrite Eu E2. - have E8 : lookup d2 = lookup d1. - - rewrite /d2 (join_useT (t:=behead (use d1 a'))) /= -/d1. - - by rewrite -Eqv //=. - by rewrite Eu E2. - have E9 : pending d2 = comp_pend (c1, c2, c) (e1, e2, e3) :: pending d1. - - rewrite /d2 (join_useT (t:=behead (use d1 a'))) /= -/d1. - - by rewrite -Eqv //=. - by rewrite Eu E2. - - step. step. step. - apply: [stepX pending d1]@h9=>//=. - move=>q h11 [r][{H}h9][->{h11} H]. - step. - apply: [gE d]=>[||??[]] //=. - exists (don ++ [:: (c1, c2, c)]). split=>//. - split; last first. - - by rewrite -catA /= E3 -E2 -Eu. - rewrite -/d2. - econstructor=>//=. - - by rewrite /class_inv /= E5 E6. - - set ut2 := (finfun _) in Ut. - exists ct, ut2. - rewrite (_ : _ \+ _ = h1 \+ (h2 \+ ((h3 \+ (h10 \+ (h5 \+ h6))) \+ (h7' \+ - (p :-> q \+ (q :-> comp_pend (c1, c2, c) (e1, e2, e3) \+ - (q.+1 :-> r \+ h9))))))); last by heap_congr. - econstructor=>//=. - econstructor=>//=. - econstructor=>//=. - - by rewrite E5. - econstructor=>//=. - econstructor=>//=. - econstructor=>//=. - - by rewrite E6. - econstructor=>//=. - econstructor=>//=. - econstructor=>//=. - - econstructor=>//=. - econstructor=>//=. - econstructor=>//=. - - rewrite (sepitT1 a'). exists h10, (h5 \+ h6). split=>//. - - by rewrite /utab/table E3/ut2 ffunE /= eqxx. - rewrite (sepitS b') !in_set !in_set1 eq_sym N /=. - eexists h5, h6. split=>//. - - rewrite /utab/table/ut2 ffunE /= eq_sym (negbTE N) /=. - by rewrite E7 // eq_sym //. - apply: tableP Ru=>x //=; - rewrite !in_set !in_set1 andbT /d2/ut2 ?ffunE /=. - - by case/andP=>_ /negbTE ->. - by case/andP=>_ X; rewrite E7 //. - econstructor=>//=. - econstructor=>//=. - econstructor=>//=. - - by rewrite E8. - econstructor=>//. - econstructor=>//. - econstructor=>//. - econstructor=>//. rewrite E9 /=. - rewrite /lseq /= InE. simpl. - econstructor=>//. - by econstructor=>//. - -apply: [stepX lookup d1]@h7'=>//=. -move=>htab' h7'' H7''. - -set ut2 := (finfun _) in Ut. - -apply: [stepX ut2, h3]@h3=>//= _ _ [->->]. -step. -apply: [stepX ut2]@h3=>//=. -case=>{Ut}h3 Ut. -set ut3 := (finfun _) in Ut. -apply: [gE d]=>[||??[]] //=. -exists (don++[:: (c1, c2, c)]). -split=>//=. -split; last first. -- by rewrite Eu -catA /= E2 E3. -rewrite -/d2. - -have E5 : rep d2 = rep d1. -- rewrite /d2 (join_useT (t:=behead (use d1 a'))) -/d1 /=. - - by rewrite -Eqv /=. - by rewrite Eu E2. -have E6 : class d2 = class d1. - - rewrite /d2 (join_useT (t:=behead (use d1 a'))) -/d1 /=. - - by rewrite -Eqv /=. - by rewrite Eu E2. -have E7 x : x != a' -> use d2 x = - if x == b' then (c1, c2, c) :: use d1 x else use d1 x. -- rewrite /d2. move=>Na. - rewrite (join_useT (t:=behead (use d1 a'))) /= -/d1. - - by rewrite -Eqv /= ffunE (negbTE Na) //. - by rewrite Eu E2. - -have E8 : lookup d2 = ins (rep d1 c2, rep d1 c) (c1, c2, c) (lookup d1). -- rewrite /d2 (join_useT (t:=behead (use d1 a'))) /= -/d1. - - by rewrite -Eqv //=. - by rewrite Eu E2. -have E9 : pending d2 = pending d1. -- rewrite /d2 (join_useT (t:=behead (use d1 a'))) /= -/d1. - - by rewrite -Eqv //=. - by rewrite Eu E2. - -rewrite (_ : _ \+ _ = h1 \+ (h2 \+ ((h3 \+ (h10 \+ - ((ut a' :-> (c1, c2, c) \+ - ((ut a').+1 :-> ut2 b' \+ h5)) \+ h6))) \+ (h7'' \+ (p :-> p' \+ h9))))); last by heap_congr. - -econstructor. -- by rewrite /class_inv /= E5 E6. -econstructor=>//=. -econstructor=>//=. -econstructor=>//=. -econstructor=>//=. -econstructor=>//=. -- by rewrite E5. -econstructor=>//=. -econstructor=>//=. -econstructor=>//=. -- by rewrite E6; exact: H2. -econstructor=>//=. -econstructor=>//=. -econstructor=>//=. -- econstructor=>//=. - econstructor=>//=. - econstructor=>//=. - - exact :Ut. - rewrite (sepitT1 a'). - econstructor=>//=. - econstructor=>//=. - econstructor=>//=. - - rewrite /utab/table/ut3 ffunE /= (negbTE N) /ut2 ffunE /= eqxx. - by rewrite E3. - rewrite (sepitS b') !in_set !in_set1 /= eq_sym N /=. - econstructor=>//=. - econstructor=>//=. - econstructor=>//=. - - rewrite /utab/table/=/ut3 !ffunE /= eqxx. - rewrite eq_sym (negbTE N). - rewrite E7; last first. by rewrite eq_sym (negbTE N). - rewrite eqxx /=. - rewrite InE /=. exists (ut b'). exists h5. split=>//. - apply: tableP Ru=>x /=; - rewrite !in_set !in_set1 andbT /ut3 ?ffunE /= /ut2 ?ffunE /=. - - by case/andP=>/negbTE -> /negbTE ->. - case/andP=>X1 X2. - by rewrite E7 // (negbTE X1). - -econstructor=>//=. -econstructor=>//=. -econstructor=>//=. -- rewrite E8 /=. - - - - - - - -apply: bnd_gh H7=>{h7} [[] h7 [H7 _]|??[] //]; rewrite -2!(unCA h3). -apply: bnd_gh (Ut)=>[_ _ [[->]] <-|??[] //]; rewrite -3!(push (ut a' .+ 1)). -apply: bnd_write; rewrite ffunE (eq_sym b') (negbTE N) -(unCA h3). -apply: {h3} bnd_gh Ut=>[[] h3 [Ut _]|??[] //]=>D. -apply: (cons_gh1 d2)=>[||??[]||] //. -- by exists d; exists (done ++ [:: (c1, c2, c)]); rewrite -/d2 E3 -catA /= -E2. -- by rewrite -E4. -rewrite (_ : _ :+ _ = h1 :+ (h2 :+ (h3 :+ (h9 :+ ((ut a' :-> (c1, c2, c) :+ - (ut a' .+ 1 :-> ut b' :+ h5)) :+ h6)) :+ (h7 :+ (p :-> l :+ h8))))); +set d1 := join_use' d a' b' don in H Eu. +case: H=>C [/= ct][ut][rh][_][->{i} Rh][cth][_][-> Htc][_][_][->] +[ru][hu][-> Ut Hu][ht][_][-> Ht][p'][_][hp][-> <- Hp]. +move: Hu; rewrite (sepitT1 a'); case=>ha' [w][->{hu} Ua]. +rewrite (sepitS b') !finE eq_sym {1}N; case=>hb' [h][->{w} Ub R]. +rewrite /utab/table/= in Ua Ub. +apply: [stepX ut, ru]@ru=>//= _ _ [->->]; apply: vrfV=>V. +case: (ut a' =P null) Ua=>[/[dup] Ea' ->|/eqP Na' {V}]. +- case/(lseq_null (validX V))=>/= {V} Eu1 ->{ha'}; step=>_; split=>//. + rewrite (_ : rh \+ _ = rh \+ (cth \+ (ru \+ (Unit \+ (hb' \+ h)) \+ + (ht \+ (p:-> p' \+ hp))))); last by heap_congr. + rewrite /join_use Eu Eu1 cats0 -/d1; split=>//=; exists ct, ut. + hhauto; rewrite (sepitT1 a'); hhauto; first by rewrite /utab/table Ea' Eu1. + by rewrite (sepitS b') !finE eq_sym N; hhauto. +case/(lseq_pos Na')=>[[[c1 c2 c3]]][nxt][hh][Eu1 ->{ha'} Hh]; step; step. +set c := (c1, c2, c3) in Eu1 *; apply: [stepX ut]@ru=>//= [[]] {Ut}ru Ut2. +set ut2 := (finfun _) in Ut2. +apply: [stepX rep d1, rh]@rh=>//= _ _ [->->]. +apply: [stepX rep d1, rh]@rh=>//= _ _ [->->]. +apply: [stepX lookup d1]@ht=>//= v {Ht}ht [Ht Eqv]. +rewrite Eu1 in Eu; set d2 := join_use' d a' b' (don ++ [:: c]). +set a1' := behead (use d1 a') in Eu. +have Eu2 : use d2 a' = a1'. +- by rewrite /d2 (join_useT (t:=behead (use d1 a'))) //; apply: join_use_useE. +have Ej2: join_use d2 a' b' = join_use d1 a' b'. +- rewrite /join_use/d2 Eu2. + by rewrite -!(join_useT (t:=[::])) ?cats0 -?catA //= ?Eu -?Eu1. +case: v Eqv=>[[[e1 e2 e3]]|] /= /esym Eqv. +- set e := (e1, e2, e3) in Eqv. + do 3!step; apply: [stepX pending d1]@hp=>//= q _ [r][{Hp}hp [-> Hp]]. + step; apply: [gE d]=>[||??[]] //=. + exists (don ++ [:: c]); rewrite -/d2 Eu2 -catA; do 2!split=>//. + rewrite /bshape'/class_inv/ashape/d2 (join_useT (t:=a1')) //= Eqv -/d1 /=. + rewrite (_ : _ \+ _ = rh \+ (cth \+ ((ru \+ (hh \+ (hb' \+ h))) \+ (ht \+ + (p :-> q \+ (q :-> comp_pend c e \+ + (q.+1 :-> r \+ hp))))))); last by heap_congr. + case: Htc=>x [y][->] X1 X2; hhauto; [eauto|eauto|exact: Ut2|]. + rewrite (sepitT1 a'); hhauto. + - by rewrite /utab/table/ut2 !ffunE /= eqxx. + rewrite (sepitS b') !finE eq_sym N; hhauto. + - by rewrite /utab/table/ut2 !ffunE /= eq_sym (negbTE N). + apply: tableP R=>a /=; rewrite !finE andbT /ut2 ?ffunE /=; + by case/andP=>_ /negbTE ->. +apply: [stepX lookup d1]@ht=>//= [[{Ht}ht Ht]]. +apply: [stepX ut2, ru]@ru=>//= _ _ [->->]; step. +apply: [stepX ut2]@ru=>//= [[]] {Ut2}ru Ut3; set ut3 := (finfun _) in Ut3. +apply: [gE d]=>[||??[]] //=; exists (don ++ [:: c]). +rewrite -/d2 Eu2 -catA; do 2!split=>//. +rewrite /bshape'/class_inv/ashape/d2 (join_useT (t:=a1')) //= Eqv -/d1 /=. +rewrite (_ : _ \+ _ = rh \+ (cth \+ ((ru \+ (hh \+ ((ut a' :-> c \+ + ((ut a').+1 :-> ut2 b' \+ hb')) \+ h))) \+ (ht \+ (p :-> p' \+ hp))))); last by heap_congr. -rewrite /d2 (@join_useT (behead (use d1 a'))) -/d1 /= -?E2 // -Ec1 -Ec2 -Eqv. -set ut2 := (finfun _) in Ut; split=>//. -exists ct; exists ut2; hauto D. -rewrite (sepitT1 a') (sepitS b') 3!in_set eq_sym N /table /=. -rewrite /ut2 !ffunE !eq_refl (eq_sym b') (negbTE N); hauto D. -by apply: tableP Ru=>s; rewrite !in_set !ffunE; case: (s == b')=>//; case: (s == a'). +case: Htc=>x [y][->] X1 X2; hhauto; [eauto|eauto|exact: Ut3|]. +rewrite (sepitT1 a'); hhauto. +- by rewrite /utab/table/ut3 ffunE /= (negbTE N) /ut2 !ffunE /= eqxx. +rewrite (sepitS b') !finE eq_sym N /=; hhauto. +- by rewrite /utab/table/=/ut3 !ffunE /= eqxx eq_sym (negbTE N); hhauto. +apply: tableP R=>a /=; rewrite !finE andbT /ut3/ut2 !ffunE /= ?ffunE /=; +by case/andP=>/negbTE -> /negbTE ->. Qed. Next Obligation. -by apply: cons0=>//; [exists H; exists (Nil (symb*symb*symb)) | apply: bshapeD H1]. +by move=>a' b' [/= d][i][N H]; apply: [gE d]=>[||??[]] //; exists [::]. Qed. -Module Dummy56. End Dummy56. -Let pend0 (e : pend) := +Let pend0 (e : pend s) := match e with simp_pend a b => a | comp_pend (a,_,_) (b,_,_) => a end. -Let pend1 (e : pend) := +Let pend1 (e : pend s) := match e with simp_pend a b => b | comp_pend (a,_,_) (b,_,_) => b end. Notation "e ..0" := (pend0 e) (at level 2). Notation "e ..1" := (pend1 e) (at level 2). Definition pT : Type := forall x:unit, - STsep unit (fun i => exists d, i \In bshape' d, - fun y i m => forall d, i \In bshape' d -> - y = Val tt /\ m \In bshape' (propagate d)). + STsep {d} (fun i => i \In bshape' d, + fun y m => y = Val tt /\ m \In bshape' (propagate d)). Program Definition hpropagate := - Fix (fun (loop : pT) x => - Do (q <-- !p; - If q == null then ret tt (* pending list is empty *) + ffix (fun (loop : pT) x => + Do (q <-- @read ptr p; + if q == null then ret tt (* pending list is empty *) else eq <-- !q; - next <-- !(q .+ 1); - p ::= (next : loc);; + next <-- !q.+1; + p ::= (next : ptr);; dealloc q;; - dealloc q .+ 1;; + dealloc q.+1;; a' <-- Array.read r eq..0; b' <-- Array.read r eq..1; - If (a' == b') then loop tt - else join_hclass a' b';; - join_huse a' b';; - loop tt)) tt. + if a' == b' then loop tt : ST _ + else + join_hclass a' b';; + join_huse a' b';; + loop tt)) tt. Next Obligation. -apply: (ghost1 H)=>[?|]; first by apply: bshape_inv. -move: {loop} H H0=>d [CI][ct][ut][hr][w][->][Hr][[hc]][w'][->{w}][Hc]. -move=>[[hu]][w][->{w'}][Hu][[ht]][w'][->{w}][Ht][[q]][w''][hp][->{w'}][]. -case/swp=>-> _ [Hp] _ _ _ _ {w' w''}. -rewrite -4!(push p)=>D; apply: bnd_read=>//. -case: ifP Hp D=>Eq. -- rewrite (eqP Eq); case/lseq_null=>Ep ->{hp} D. - apply: val_ret=>//; rewrite propagate_equation Ep; do 2!split=>//. - exists ct; exists ut; rewrite 4!(push p) in D *; hauto D. - by rewrite Ep. -case/(lseq_pos (negbT Eq))=>eq [next][hnext][Ep] _ <-{hp} Hp. -rewrite -5!(push q)=>D; apply: bnd_read=>//; rewrite -6!(push (q .+ 1)) in D *. -apply: bnd_read=>//; move: D; rewrite -2!(push p); apply: bnd_write; rewrite -2!(push q). -apply: bnd_dealloc; rewrite -(push (q .+ 1)); apply: bnd_dealloc; rewrite -(unCA hr). -apply: bnd_gh (Hr)=>[a' _ [[Ea]] <-|??[] //]. -apply: bnd_gh (Hr)=>[b' _ [[Eb]] <-|??[] //]. +move=>loop [][/= d][i][C][/= ct][ut][hr][_][-> Hr][hc][_][-> Hc] +[hu][_][-> Hu][ht][_][-> Ht][q][_][hp][->] <- Hp. +step; case: (q =P null) Hp=>[->{q}|/eqP N] Hp. +- apply: vrfV=>V; case/(lseq_null (validX V)): Hp=>{V} Ep /= ->{hp}. + step=>_; rewrite propagate_equation Ep; do !split=>//=. + by exists ct, ut; hhauto; rewrite Ep. +case/(lseq_pos N): Hp=>eq [next][hnext][Ep] ->{hp} Hp; do !step. +apply: [stepX rep d, hr]@hr=>//= a' _ [-> Ea']. +apply: [stepX rep d, hr]@hr=>//= b' _ [-> Eb']. set d1 := Data (rep d) (class d) (use d) (lookup d) (behead (pending d)). -rewrite 3!(push p). -case: ifP=>E D. -- have T1: propagate d = propagate d1. - - rewrite propagate_equation Ep /=. - by case: eq Ep Ea Eb=>[a2 b2 _ <-<-| [[a2 ?]?][[b2 ?]?] _ <-<-]; rewrite E. - apply: (cons_gh d1)=>[[] m|??[] //||//]; first by rewrite T1. - by split=>//; exists ct; exists ut; hauto D. -rewrite -(unh0 (_ :+ _)) in D *. -apply: (bnd_gh1 d1)=>[|[] m [_] Wm|??[]//||//]; last 1 first. -- by split=>//; exists ct; exists ut; hauto D. -- by rewrite E; eauto. -set d2 := (join_class _ _ _) in Wm. -apply: (bnd_gh1 d2)=>[|{m Wm} [] m [_] Wm|??[]|] //; first by rewrite E /=; eauto. -rewrite unh0; apply: cons0=>/= [|y m2]; first by eauto. -case/(_ _ Wm)=>->{Wm} Wm Dm; split=>//. -rewrite (_ : propagate d = propagate (join_use d2 a' b')) // /d2 - propagate_equation Ep. -by case: eq Ep Ea Eb=>[a2 b2 _ <- <-|[[a2 ?]?][[b2 ?]?] _ <- <-]; rewrite E. +case: ifPn=>[E|E]. +- have T1 : propagate d = propagate d1. + - rewrite propagate_equation Ep. + by case: eq Ep Ea' Eb' E=>[a2 b2 _|[[a2 ??]][[b2 ??]] _] ->->->. + apply: [gE d1]=>[||??[] //] /=; last by rewrite T1. + by split=>//=; exists ct, ut; hhauto. +apply: [stepE d1]=>[||??[]//] /=. +- by do !split=>//; exists ct, ut; hhauto. +set d2 := (join_class _ _ _); case=>m [_] H. +apply: [stepE d2]=>[//||??[]//] /= [n][_] Hn. +set d3 := (join_use _ _ _) in Hn. +suff -> : propagate d = propagate d3 by apply: [gE d3]. +rewrite /d3/d2 propagate_equation Ep; +by case: eq Ep Ea' Eb' E=>[a2 b2 _|[[a2 ??]][[b2 ??]] _] ->-> /negbTE ->. Qed. -Module Dummy57. End Dummy57. - -Program Definition merge (e : Eq) : - STsep unit (fun i => exists R, i \In shape' R, - fun y i m => forall R, i \In shape' R -> y = Val tt /\ - m \In shape' (closure (R +r eq2rel e))) := +Program Definition merge (e : Eq s) : + STsep {R} (fun i => i \In shape' R, + fun y m => y = Val tt /\ + m \In shape' (closure (R \+p eq2rel e))) := match e with - simp_eq a b => + | simp_eq a b => Do (q <-- !p; x <-- insert q (simp_pend a b); p ::= x;; @@ -642,7 +445,12 @@ Program Definition merge (e : Eq) : Do (c1' <-- Array.read r c1; c2' <-- Array.read r c2; v <-- KVmap.lookup htab (c1', c2'); - match_opt v then + if v is Some b then + q <-- !p; + x <-- insert q (comp_pend (c, c1, c2) b); + p ::= x;; + hpropagate + else KVmap.insert htab (c1', c2') (c, c1, c2);; u1 <-- Array.read ulist c1'; x <-- insert u1 (c, c1, c2); @@ -655,184 +463,151 @@ Program Definition merge (e : Eq) : Array.write ulist c1' x;; u2 <-- Array.read ulist c2'; x <-- insert u2 (c, c1, c2); - Array.write ulist c2' x - else [b] - q <-- !p; - x <-- insert q (comp_pend (c, c1, c2) b); - p ::= x;; - hpropagate) + Array.write ulist c2' x) end. Next Obligation. -apply: (ghost2 (t:=H)); first by move=>???; move/(shape_inv H0)=>->. -move: H H0=>R [d][[CI]][ct][ut][hr][w][->][Hr][[hc]][w'][->{w}][Hc]. -move=>[[w]][w''][->{w'}][[hu]][hu'][->{w}][Hu][Hu'] _ [[ht]][w']. -move=>[->{w''}][Ht][[q]][w''][hp][->{w'}][]. -case/swp=>-> _ [Hp] _ _ _ _ {w''} D [PI][Ep] ERel; move: D. -set d1 := Data (rep d) (class d) (use d) (lookup d) (simp_pend a b :: pending d). -rewrite -(unA hu) -5!(push p)=>D; apply: bnd_read=>//; move: D. -rewrite -(unC hp) -5!(unCA hp). -apply: {hp} bnd_gh Hp=>[_ hp [q'] [Hp] [->{x}]|??[?][]] //. -rewrite -(push p); apply: bnd_write=>D. -apply: (cons_gh d1)=>[[] m [_] Wm Dm|??[] //||//]; last first. -- rewrite 4!(unCA hp) (unC hp) 5!(push p) (unA hu) in D *. - move: Hp D=>[x][h'][_] <-; rewrite Ep=>[[->]] ->=>D. - split=>//; exists ct; exists ut; hauto D. - by rewrite Ep. +move=>e a b [R][_][d][[C]][/= ct][ut][hr][_][-> Hr][hc][_][-> Hc][_][_] +[->][hu][hu'][-> Hu Hu'][ht][_][-> Ht][q][_][hp][->] <- Hp [PI][Ep Erel]. +step; apply: [stepX (pending d)]@hp=>//= x _ [r][{Hp}hp][-> Hp]. +set d1:=Data (rep d) (class d) (use d) (lookup d) (simp_pend a b :: pending d). +step; apply: [gE d1]=>[||??[]] //=. +- rewrite Ep /= in Hp; case: Hp=>->{r} ->{hp}. + rewrite (_ : _ \+ _ = hr \+ (hc \+ (hu \+ hu' \+ (ht \+ + (p :-> x \+ (x :-> simp_pend a b \+ (x.+1 :-> null \+ Unit))))))); + last by heap_congr. + by split=>//=; exists ct, ut; hhauto; rewrite Ep. +case=>m [_ Hm _]; split=>//. have L: propagate_inv d1 by apply: propagate_pendP PI. -move: (propagatePE L)=>[L1][L2] L3; split=>//; exists (propagate d1). -by rewrite -L3 -ERel {4}/d1 /CRel /= clos_clos -!orrA orrAC. +case: (propagatePE L)=>L1 [L2 L3]; exists (propagate d1); do 3!split=>//. +by rewrite -L3 -Erel clos_clos /CRel -!orrA orrAC. Qed. Next Obligation. -apply: (ghost2 (t:=H)); first by move=>???; move/(shape_inv H0)=>->. -move: H H0=>R [d][[CI]][ct][ut][hr][w][->][Hr][[hc]][w'][->{w}][Hc]. -move=>[[w]][w''][->{w'}][[hu]][hu'][->{w}][Hu][Hu'] _ [[ht]][w']. -move=>[->{w''}][Ht][[q]][w''][hp][->{w'}][]. -case/swp=>-> _ [Hp] _ _ _ _ {w''} D [PI][Ep] ERel; move: D. -do 2!apply: bnd_gh (Hr)=>[_ _ [[->]] <-|??[]//]; rewrite -3!(unCA ht). -apply: {ht} bnd_gh Ht=>[v ht [Ht] [Ev]|??[]//]. -case: v Ev=>[[[e e1] e2]|] Ev. -- set d1 := Data (rep d) (class d) (use d) (lookup d) - (comp_pend (c, c1, c2) (e, e1, e2) :: pending d). - have L: propagate_inv d1 by apply: propagate_pendP PI. - rewrite -4!(push p)=>D; apply: bnd_read=>//; move: D. - rewrite -(unC hp) -4!(unCA hp). - apply: {hp} bnd_gh Hp=>[_ hp [x][Hp][->]|??[?][]//]. - rewrite -(push p); apply: bnd_write. - rewrite 3!(unCA hp) (unC hp) 4!(push p) 3!(unCA ht)=>D. - case: Hp D=>[x1][h'][_] <-; rewrite Ep; case=>->-> D. - apply: (cons_gh d1)=>[[] m [_] Wm Dm |??[]//||//]; last first. - - by split=>//; exists ct; exists ut; hauto D; rewrite Ep. - move: (propagatePE L)=>[L1][L2] L3; split=>//; exists (propagate d1). - rewrite -L3 {4}/d1 -ERel; do 3!split=>//. - by apply: propagate_clos_pendP. -apply: bnd_gh Ht=>{ht x} [_ ht [Ht] _|??[]//]; rewrite -(unA hu) -3!(unCA hu). -apply: bnd_gh (Hu)=>[_ _ [[->]] <-|??[]//]; move: Hu'. -rewrite (sepitT1 (rep d c1)) {1}/table=>[[hu'']][hu2][->{hu'}][Hu''][Hu'] _. -rewrite -(unA hu'') -4!(unCA hu''). -apply: {hu''} bnd_gh Hu''=>[_ hu'' [x] [Hu''][->]|??[?][]//]. -rewrite -(unCA hu); apply: {hu} bnd_gh Hu=>[_ hu [Hu] _|??[] //]. -apply: bnd_gh (Hu)=>[_ _ [[->]] <-|??[]//]. -set ut1 := (finfun _) in Hu. -set u1' := [ffun z => if z == rep d c1 then (c, c1, c2) :: use d z else use d z]. -set u' := [ffun z => if z == rep d c2 then (c, c1, c2) :: u1' z else u1' z]. +move=>X c c1 c2 [R][_][d][[C]][/= ct][ut][hr][_][-> Hr][hc][_][-> H] +[_][_][->][hu][hu'][-> Hu Hu'][ht][_][-> Ht][q][_] +[hp][->] <- Hp [PI][Ep Erel]; set cx := (c, c1, c2). +apply: [stepX rep d, hr]@hr=>//= _ _ [->->]. +apply: [stepX rep d, hr]@hr=>//= _ _ [->->]. +apply: [stepX lookup d]@ht=>//= v {Ht}ht [Ht Ev]. +case: v Ev=>[[[e e1 e2]]|] Ev. +- set ex := (e, e1, e2). + set d1 := Data (rep d) (class d) (use d) + (lookup d) (comp_pend cx ex :: pending d). + step; apply: [stepX pending d]@hp=>//= x _ [x1][{}hp][-> {}Hp]. + step; apply: [gE d1]=>[||??[]] //=. + - rewrite (_ : _ \+ _ = hr \+ (hc \+ (hu \+ hu' \+ (ht \+ (p :-> x \+ + (x :-> comp_pend cx ex \+ (x.+1 :-> x1 \+ hp))))))); + last by heap_congr. + by split=>//; exists ct, ut; hhauto. + case=>m [_] Hm _; split=>//. + have L : propagate_inv d1 by apply: propagate_pendP PI. + case: (propagatePE L)=>L1 [L2] L3; exists (propagate d1). + by rewrite -L3 -Erel propagate_clos_pendP. +apply: [stepX lookup d]@ht=>//= _ {Ht}ht Ht. +apply: [stepX ut, hu]@hu=>//= _ _ [->->]. +move: Hu'; rewrite (sepitT1 (rep d c1)). +case=>hu'' [hu2][->{hu'} Hu'' Hu']. +apply: [stepX use d (rep d c1)]@hu''=>//= x _ [r][{Hu''}hu''][-> Hu'']. +apply: [stepX ut]@hu=>//= _ {Hu}hu Hu; set ut1 := (finfun _) in Hu. +apply: [stepX ut1, hu]@hu=>//= _ _ [->->]. +set u1' := [ffun z => if z == rep d c1 then cx :: use d z else use d z]. +set u' := [ffun z => if z == rep d c2 then cx :: u1' z else u1' z]. set l' := (ins _ _ _) in Ht. set d1 := Data (rep d) (class d) u' l' (pending d). pose ut2 x' := [ffun z => if z == rep d c2 then x' else ut1 z]. -case E: (rep d c2 == rep d c1). -- rewrite -(unCA hu''). - apply: (bnd_gh ((c, c1, c2) :: use d (rep d c1))); last first. - - by rewrite ffunE E. - - by move=>??[?][]. - move=>_ {hu'' Hu''} hu'' [x'] [Hu''] [->]. - rewrite -(unCA hu); apply: {hu} val_gh Hu=>[[] hu [Hu] _|??[]//]. - rewrite (_ : _ :+ _ = hr :+ (hc :+ (hu :+ (hu'' :+ hu2) :+ (ht :+ (p :-> q :+ hp))))); - last by heap_congr. - move=>D; split=>//; exists d1; split; last first. +case E : (rep d c2 == rep d c1). +- apply: [stepX cx::use d (rep d c1)]@(x :-> _ \+ x.+1 :-> _ \+ hu'')=>//=. + - by exists r, hu''; rewrite joinA (eqP E) /ut1 !ffunE /= eqxx. + move=>v {Hu''}_ [x'][_][->][r0][{}hu''][-> Hu'']. + apply: [gX ut1]@hu=>//= [[]] {Hu}hu Hu _; split=>//. + rewrite (_ : _ \+ _ = hr \+ (hc \+ (hu \+ (v :-> cx \+ + (v.+1 :-> x' \+ (x' :-> cx \+ (x'.+1 :-> r0 \+ hu''))) \+ + hu2) \+ (ht \+ (p :-> q \+ hp))))); last by heap_congr. + exists d1; split=>//; last first. - split; first by apply: propagate_nopendP. - split; first by rewrite /d1. - by rewrite -ERel; apply: propagate_clos_nopendP. - split=>//; exists ct; exists (ut2 x'); hauto D. - rewrite (sepitT1 (rep d c1)) {1}/table !ffunE (eq_sym (rep d c1)) E eq_refl. - hauto D; apply: tableP Hu' => x0; rewrite !in_set andbT /ut1 !ffunE; - by case: ifP=>H1; by [rewrite (eqP H1) E | case: eqP]. -move: Hu'; rewrite (sepitS (rep d c2)) !in_set E /=. -case=>[hu1][hu3][->][Hu1'][Hu2'] _; rewrite -(unA hu1) -5!(unCA hu1). -apply: (bnd_gh (use d (rep d c2))); last first. -- by rewrite ffunE E. -- by move=>??[?][]. -move=>_ {hu1 Hu1'} hu1 [x'] [Hu1'] [->]. -rewrite -!(unCA hu); apply: {hu} val_gh Hu=>[[] hu [Hu] _|??[]//]. -rewrite (_ : _ :+ _ = hr :+ (hc :+ (hu :+ (hu'' :+ (hu1 :+ hu3)) :+ (ht :+ (p :-> q :+ hp))))); - last by heap_congr. -move=>D; split=>//; exists d1; split; last first. + by rewrite -Erel propagate_clos_nopendP. + split=>//=; exists ct, (ut2 v); hhauto. + rewrite (sepitT1 (rep d c1)) /utab/table/ut2 /= !ffunE eq_sym eqxx E /=. + hhauto; apply: tableP Hu'=>// a; rewrite !finE /u'/u1' !ffunE /= andbT; + by rewrite (eqP E)=>/negbTE ->. +move: Hu'; rewrite (sepitS (rep d c2)) !finE E /=. +case=>hu1 [hu3][->{hu2} Hu1' Hu2']. +apply: [stepX use d (rep d c2)]@hu1=>//=; first by rewrite /ut1 ffunE /= E. +move=>v _ [x'][{Hu1'}hu1][-> Hu1']. +apply: [gX ut1]@hu=>//= [[]] {Hu}hu Hu _; split=>//. +rewrite (_ : _ \+ _ = hr \+ (hc \+ (hu \+ + (x :-> cx \+ (x.+1 :-> r \+ hu'') \+ + ((v :-> cx \+ (v.+1 :-> x' \+ hu1) \+ hu3))) \+ + ((ht \+ (p :-> q \+ hp)))))); last by heap_congr. +exists d1; split; last first. - split; first by apply: propagate_nopendP. split; first by rewrite /d1. - by rewrite -ERel; apply: propagate_clos_nopendP. -split=>//; exists ct; exists (ut2 x'); hauto D. -rewrite (sepitT1 (rep d c1)) /table !ffunE eq_sym E eq_refl; hauto D. -rewrite (sepitS (rep d c2)) !in_set E /= !ffunE eq_refl E; hauto D. -apply: tableP Hu2'=>y; rewrite !in_set /u' !ffunE; -(case: ifP=>H1; first by rewrite (eqP H1) eq_refl); -by case: ifP=>H2 //; rewrite (eqP H2) eq_refl andbF. + by rewrite -Erel; apply: propagate_clos_nopendP. +split=>//; exists ct, (ut2 v); hhauto. +rewrite (sepitT1 (rep d c1))/utab/table/ut2 !ffunE /= eqxx eq_sym E. +hhauto; rewrite (sepitS (rep d c2)) !finE !ffunE /= !eqxx E /=. +hhauto; apply: tableP Hu2'=>/= a; rewrite !finE !ffunE /= andbT; +by case/andP=>/negbTE -> /negbTE ->. Qed. -Module Dummy5. End Dummy5. - -Program -Definition Match_exp (A : Type) (t : exp) (p1 : symb -> spec A) (p2 : exp -> exp -> spec A) - (c1 : forall s, STsep A (p1 s)) - (c2 : forall t1 t2, STsep A (p2 t1 t2)) : - STsep A (match t with const s => (p1 s) | app t1 t2 => (p2 t1 t2) end) := - match t with const s => st.do (c1 s) _ | app t1 t2 => st.do (c2 t1 t2) _ end. Let pend3 (e : symb*symb*symb) := let: (a, _, _) := e in a. Notation "e ..0" := (pend3 e) (at level 2). Definition nT : Type := - forall t, STsep exp (fun i => exists d, i \In bshape' d, - fun y i m => forall d, i \In bshape' d -> - m \In bshape' d /\ y = Val (norm d t)). + forall t, STsep {d} (fun i => i \In bshape' d, + fun y m => m \In bshape' d /\ y = Val (norm d t)). Program Definition hnorm := - Fix (fun (hnorm : nT) (t : exp) => - Do (Match_exp t - (fun a => a' <-- Array.read r a; - ret (const a')) - (fun t1 t2 => - u1 <-- hnorm t1; - u2 <-- hnorm t2; - Match_exp u1 - (fun w1 => - Match_exp u2 - (fun w2 => - v <-- KVmap.lookup htab (w1, w2); - match_opt v then ret (app u1 u2) - else [a] a' <-- Array.read r (a..0); - ret (const a')) - (fun _ _ => ret (app u1 u2))) - (fun _ _ => ret (app u1 u2))))). + ffix (fun (hnorm : nT) (t : exp s) => + Do (match t with + | const a => + a' <-- Array.read r a; + ret (const a') + | app t1 t2 => + u1 <-- hnorm t1; + u2 <-- hnorm t2; + if (u1, u2) is (const w1, const w2) then + v <-- KVmap.lookup htab (w1, w2); + if v is Some a then + a' <-- Array.read r (a..0); + ret (const a') + else ret (app u1 u2) + else ret (app u1 u2) + end)). Next Obligation. -apply: (ghost2 (t:=H)); first by move=>???; move/(bshape_inv H0)=>->. -move: H H0=>d [CI][ct][ut][hr][hrest][->][Hr][Hrest]. +move=>hnorm t [d][_][Ci][/= ct][ut][hr][hrest][-> Hr Hrest]. case: t=>[s|t1 t2]. -- apply: bnd_gh (Hr)=>[_ _ [[->]] <-|??[]//]=>D. - by apply: val_ret=>//; do 2!split=>//; exists ct; exists ut; hauto D. -rewrite -(unh0 (_ :+ _))=>D. -apply: (bnd_gh d)=>[u1 m1 [W1] [E1] D1|??[]//||//]; last first. -- by split=>//; exists ct; exists ut; hauto D. -apply: (bnd_gh d)=>[u2 m2 [W2] [E2]|??[]||] //; rewrite unh0. -move: W2=>{CI ct ut hr hrest Hr Hrest D}[CI][ct][ut][hr][w][->][Hr][[hc]][w']. -move=>[->{w}][Hc][[hu]][w][->{w'}][Hu][[ht]][hrest][->{w}][Ht][Hrest] _ _ _ _ D. +- apply: [stepX rep d, hr]@hr=>//= _ _ [->->]. + by step; do 2!split=>//; exists ct, ut; hhauto. +apply: [stepE d]=>[||??[]] //=. +- by split=>//; exists ct, ut; hhauto. +move=>u1 m [H][E1]; apply: [stepE d]=>[||??[]] //=. +move=>u2 {}m [{}H [E2]] {Ci ct ut hr hrest Hr Hrest}. +case: H=>Ci [/= ct][ut][hr][w][->{m} Hr][hc][w'] +[->{w} Hc][hu][w][->{w'} Hu][ht][hrest][->{w} Ht Hrest]. case: u1 E1=>[w1|??] /= E1; last first. -- rewrite -E1 E2; apply: val_ret=>//; do 2!split=>//. - by exists ct; exists ut; hauto D. +- by rewrite -E1 E2; step; do 2!split=>//; exists ct, ut; hhauto. case: u2 E2=>[w2|??] /= E2; last first. -- rewrite -E1 -E2; apply: val_ret=>//; do 2!split=>//. - by exists ct; exists ut; hauto D. -move: D; rewrite -3!(unCA ht). -apply: bnd_gh Ht=>{ht x} [v ht [Ht][Ev]|??[]//]; rewrite 3!(unCA ht) -E1 -E2. -case: v Ev=>[[[a0 a1] a2]|] /= <- D; last first. -- by apply: val_ret=>//; do 2!split=>//; exists ct; exists ut; hauto D. -apply: bnd_gh (Hr) (D)=>[a' _ [[->]] <-|??[]//]. -by apply: val_ret; do 2!split=>//; exists ct; exists ut; hauto D. +- by rewrite -E1 -E2; step; do 2!split=>//; exists ct, ut; hhauto. +rewrite -E1 -E2; apply: [stepX lookup d]@ht=>//= v {Ht}ht [Ht Ev]. +case: v Ev=>[[[a1 a2 a3]]|] <-; last first. +- by step; do 2!split=>//; exists ct, ut; hhauto. +apply: [stepX rep d, hr]@hr=>//= _ _ [->->]. +by step; do 2!split=>//; exists ct, ut; hhauto. Qed. -Module Dummy6. End Dummy6. - Program Definition check t1 t2 : - STsep bool (fun i => exists R, i \In shape' R, - fun y i m => forall R, i \In shape' R -> m \In shape' R /\ - exists b, y = Val b /\ (b <-> (t1, t2) \In R)) := + STsep {R} (fun i => i \In shape' R, + fun y m => m \In shape' R /\ + exists b:bool, y = Val b /\ (b <-> (t1, t2) \In R)) := Do (u1 <-- hnorm t1; u2 <-- hnorm t2; ret (u1 == u2)). Next Obligation. -apply: (ghost2 (t:=H)). -- by move=>???; move/(shape_inv H0)=>E [L1][b][->] L2; move: L1 L2; rewrite E; eauto. -move: H H0 (shapeD H0)=>R [d][H][[RI]H1][P] H2; rewrite -(unh0 i). -apply: {i} bnd_gh H=>[u1 i [H] [->]|??[]//]. -apply: {i} bnd_gh H=>[u2 i [H] [->]|??[]//]. -apply: val_ret=>//; rewrite unh0; split; first by exists d. +move=>t1 t2 [R][h][d][H][[RI X]][Ep PI]. +apply: [stepX d]@h=>[||??[]] //= _ {H}h [H][->]. +apply: [stepX d]@h=>[||??[]] //= _ {H}h [H][->]. +step=>_. split; first by exists d. exists (norm d t1 == norm d t2); split=>//. -by case: normP=>//; rewrite H2; move=>H3; split. +by case: normP=>//; rewrite PI. Qed. diff --git a/examples/hashtab.v b/examples/hashtab.v index 890c1ce..ac41cfa 100644 --- a/examples/hashtab.v +++ b/examples/hashtab.v @@ -24,18 +24,22 @@ Section HashTab. (* hash table is array of buckets, i.e. KV maps *) (* bucket indices are provided by the hash function *) +(* using dynaming kv-maps for buckets *) +(* DEVCOMMENT: *) +(* it's possible to develop this with static buckers *) +(* /DEVCOMMENT *) -Variables (K : ordType) (V : Type) (buckets : KVmap.Sig K V) +Variables (K : ordType) (V : Type) (buckets : DynKVmap.Sig K V) (n : nat) (hash : K -> 'I_n). -Definition hashtab := {array 'I_n -> KVmap.tp buckets}. -Notation KVshape := (@KVmap.shape _ _ buckets). +Definition hashtab := {array 'I_n -> DynKVmap.tp buckets}. +Notation KVshape := (@DynKVmap.shape _ _ buckets). Notation table := (table KVshape). Notation nil := (nil K V). (* hash table is specified by a single finMap *) (* which is the "flattening" of all buckets *) Definition shape x (s : finMap K V) := - [Pred h | exists (tab : {ffun 'I_n -> KVmap.tp buckets}) (* array spec *) + [Pred h | exists (tab : {ffun 'I_n -> DynKVmap.tp buckets}) (* array spec *) (bucket : 'I_n -> finMap K V), (* buckets spec *) [/\ forall k, fnd k s = fnd k (bucket (hash k)), forall i k, k \in supp (bucket i) -> hash k = i & @@ -52,10 +56,10 @@ Definition new_loopinv x := forall k, [vfun y => shape y nil]). Program Definition new : STsep (emp, [vfun y => shape y nil]) := - Do (t <-- Array.new _ (KVmap.default buckets); + Do (t <-- Array.new 'I_n (DynKVmap.default buckets); let go := ffix (fun (loop : new_loopinv t) k => Do (if decP (b := k < n) idP is left pf then - b <-- KVmap.new buckets; + b <-- DynKVmap.new buckets; Array.write t (Ordinal pf) b;; loop k.+1 else ret t)) @@ -97,7 +101,7 @@ apply: [stepE]=>//= y m Hm. (* invoke the loop with index 0 *) apply: [gE]=>//=; split=>//. (* the table is empty *) -exists [ffun => KVmap.default buckets], m, Unit; split=>//=. +exists [ffun => DynKVmap.default buckets], m, Unit; split=>//=. - by rewrite unitR. (* there are no buckets in the heap yet *) by rewrite (eq_sepit (s2 := set0)) // sepit0. @@ -119,7 +123,7 @@ Program Definition free x : STsep {s} (shape x s, Do (ffix (fun (loop : free_loopinv x) k => Do (if decP (b := k < n) idP is left pf then b <-- Array.read x (Ordinal pf); - KVmap.free b;; + DynKVmap.free b;; loop k.+1 else Array.free x)) 0). (* first the loop *) @@ -159,16 +163,14 @@ Qed. (* inserting into hashmap is inserting into *) (* corresponding bucket + updating the array *) -(* returning the pointer is technically not needed *) +(* returning the pointer is not needed *) (* as the array is not moved *) -(* but one needs to fit the KV map API *) Program Definition insert x k v : - STsep {s} (shape x s, [vfun y => shape y (ins k v s)]) := + STsep {s} (shape x s, [vfun _ : unit => shape x (ins k v s)]) := Do (let hk := hash k in b <-- Array.read x hk; - b' <-- KVmap.insert b k v; - Array.write x hk b';; - ret x). + b' <-- DynKVmap.insert b k v; + Array.write x hk b'). Next Obligation. (* pull out ghost + deconstruct precondition *) move=>/= x k v [fm][] _ /= [tf][bf][Hf Hh [h1][h2][-> /= H1 H2]]. @@ -179,7 +181,7 @@ move: H2; rewrite (sepitT1 (hash k)) /table; case=>h3[h4][{h2}-> H3 H4]. (* insert into the bucket *) apply/[stepX (bf (hash k))] @ h3=>{h3 H3}//= b' m2 H2. (* write the bucket to the array, return the pointer *) -apply/[stepX tf] @ h1=>{h1 H1}//= _ m3 E3; step=>_. +apply: [gX tf]@h1=>{h1 H1} //= _ m3 E3 _. (* update the array and buckets' specs *) exists [ffun z => if z == hash k then b' else tf z], (fun b => if b == hash k then ins k v (bf b) else bf b); split=>/=. @@ -205,15 +207,13 @@ Qed. (* removing from hashmap is removing from *) (* corresponding bucket + updating the array *) (* returning the pointer is again not needed *) -(* except for the API fit *) Program Definition remove x k : STsep {s} (shape x s, - [vfun y => shape y (rem k s)]) := + [vfun _ : unit => shape x (rem k s)]) := Do (let hk := hash k in b <-- Array.read x hk; - b' <-- KVmap.remove b k; - Array.write x hk b';; - ret x). + b' <-- DynKVmap.remove b k; + Array.write x hk b'). Next Obligation. (* pull out ghost + destructure precondition *) move=>/= x k [fm][] _ /= [tf][bf][Hf Hh [h1][h2][-> /= H1 H2]]. @@ -224,7 +224,7 @@ move: H2; rewrite (sepitT1 (hash k)); case=>h3[h4][{h2}-> H3 H4]. (* insert into the bucket *) apply/[stepX (bf (hash k))] @ h3=>{h3 H3}//= b' m2 H2. (* write the bucket to the array, return the pointer *) -apply/[stepX tf] @ h1=>{H1}//= _ m3 E3; step=>_. +apply/[gX tf] @ h1=>{H1}//= _ m3 E3 _. (* update the array and buckets' specs *) exists [ffun z => if z == hash k then b' else tf z], (fun b => if b == hash k then rem k (bf b) else bf b); split=>/=. @@ -252,7 +252,7 @@ Program Definition lookup x k : STsep {s} (shape x s, [vfun y m => m \In shape x s /\ y = fnd k s]) := Do (b <-- Array.read x (hash k); - KVmap.lookup b k). + DynKVmap.lookup b k). Next Obligation. (* pull out ghost + destructure precondition *) move=>/= x k [fm][] _ /= [tf][bf][Hf Hh [h1][h2][-> H1 H2]]. @@ -268,10 +268,10 @@ exists tf, bf; split=>//=; exists h1, (m2 \+ h4); split=>{h1 H1} //. by rewrite (sepitT1 (hash k)); vauto. Qed. -(* hash table is a KV map *) +(* hash table is a *static* KV map *) Definition HashTab := KVmap.make (Array null) new free insert remove lookup. End HashTab. End HashTab. -Definition HT K V := HashTab.HashTab (AssocList.AssocList K V). +Definition HT K V := HashTab.HashTab (DynAssocList.AssocList K V). diff --git a/examples/kvmaps.v b/examples/kvmaps.v index 47e5a2a..a4705ce 100644 --- a/examples/kvmaps.v +++ b/examples/kvmaps.v @@ -25,17 +25,52 @@ From htt Require Import options model heapauto. (* stateful KV map ADT *) (***********************) +(* Dynamic KV map is determined by the root pointer. *) +(* Functions such insert and remove may modify *) +(* the root, and will correspondingly return the new one. *) +(* tp is abstracted to facilitate passing K and V to methods *) +(* (thus, the move reduces annotations) *) +Module DynKVmap. +Record Sig (K : ordType) (V : Type) : Type := + make {tp :> Type; + default : tp; + shape : tp -> {finMap K -> V} -> Pred heap; + new : STsep (emp, [vfun x => shape x (nil K V)]); + free : forall x, STsep {s} (shape x s, + [vfun _ : unit => emp]); + insert : forall x k v, + STsep {s} (shape x s, + [vfun y => shape y (ins k v s)]); + remove : forall x k, + STsep {s} (shape x s, + [vfun y => shape y (rem k s)]); + lookup : forall x k, + STsep {s} (shape x s, + [vfun y m => m \In shape x s /\ y = fnd k s])}. +End DynKVmap. + +(* Static KV map (or just KV map) are those that *) +(* don't need to modify the root pointer. *) +(* Typical example will be hash tables *) +(* Another example is a structure obtained by packaing *) +(* the root pointer along with the dynamic KV map that the *) +(* root points to. *) + Module KVmap. Record Sig (K : ordType) (V : Type) : Type := make {tp :> Type; default : tp; shape : tp -> {finMap K -> V} -> Pred heap; + (* allocate root pointer and empty structure *) new : STsep (emp, [vfun x => shape x (nil K V)]); + (* remove the struture, and the root pointer *) free : forall x, STsep {s} (shape x s, [vfun _ : unit => emp]); + (* insert, keeping the root pointer unchanged *) insert : forall x k v, STsep {s} (shape x s, [vfun _ : unit => shape x (ins k v s)]); + (* remove, keeping the root pointer unchanged *) remove : forall x k, STsep {s} (shape x s, [vfun _ : unit => shape x (rem k s)]); @@ -48,7 +83,7 @@ End KVmap. (* KV map implemented as a sorted association linked list *) (**********************************************************) -Module AssocList. +Module DynAssocList. Section AssocList. Variables (K : ordType) (V : Set). Notation fmap := (finMap K V). @@ -75,7 +110,7 @@ Definition shape_seg (x y : ptr) (s : fmap) : Pred heap := Definition shape (x : ptr) (s : fmap) : Pred heap := shape_seg x null s. -(* null pointer represents an empty map *) +(* null pointer represents empty map *) Lemma shape_null (s : fmap) h : valid h -> h \In shape null s -> @@ -161,9 +196,10 @@ Qed. (* new map is just a null pointer *) Program Definition new : STsep (emp, [vfun x => shape x nil]) := Do (ret null). -Next Obligation. by move=>[] /= _ ->; step. Qed. +Next Obligation. by case=>_ /= ->; step. Qed. (* freeing a map is deallocating all its elements *) +(* and the root pointer *) Definition freeT : Type := forall p, STsep {fm} (shape p fm, [vfun _ : unit => emp]). @@ -190,6 +226,7 @@ Qed. (* looking up an element in the map *) + Definition lookupT k : Type := forall p, STsep {fm} (shape p fm, [vfun y m => m \In shape p fm /\ y = fnd k fm]). @@ -564,9 +601,99 @@ apply/path_supp_ins=>//. by apply/path_le/all_path_supp/Of. Qed. -(* ordered association list is a KV map *) +(* ordered association list is a dynamic KV map *) -Definition AssocList := KVmap.make null new free insert remove lookup. +Definition AssocList := DynKVmap.make null new free insert remove lookup. +End AssocList. +End DynAssocList. + +(* static variant packages the root pointer *) +(* with the dynamic part of the structure *) + +Module AssocList. +Section AssocList. +Variables (K : ordType) (V : Set). +Notation fmap := (finMap K V). +Notation nil := (nil K V). + +Definition shape (x : ptr) (f : fmap) : Pred heap := + [Pred h | exists (a : ptr) h', + h = x :-> a \+ h' /\ @DynAssocList.shape K V a f h']. + +(* new structure, but the root pointer is given *) +Program Definition new0 (x : ptr) : + STsep {a : ptr} (fun h => h = x :-> a, + [vfun _ : unit => shape x nil]) := + Do (a <-- @DynAssocList.new K V; + x ::= a). +Next Obligation. +move=>x [a][/= _ ->]; apply: [stepU]=>//= b h H. +by step; exists b, h; rewrite joinC. +Qed. + +Program Definition new : + STsep (emp, [vfun x => shape x nil]) := + Do (a <-- @DynAssocList.new K V; + alloc a). +Next Obligation. +case=>_ /= ->; apply: [stepU]=>//= a h H. +by step=>x; exists a, h; rewrite unitR. +Qed. + +(* free structure, keep the root pointer *) +Program Definition free0 x : + STsep {s} (shape x s, + [vfun (_ : unit) h => exists a : ptr, h = x :-> a]) := + Do (a <-- !x; + DynAssocList.free K V a). +Next Obligation. +move=>x [fm][/= _ [a][h][-> H]]; step. +by apply: [gX fm]@h=>//= _ _ -> _; rewrite unitR; eauto. +Qed. + +Program Definition free x : + STsep {s} (shape x s, + [vfun _ : unit => emp]) := + Do (a <-- !x; + dealloc x;; + DynAssocList.free K V a). +Next Obligation. +by move=>x [fm][/= _ [a][h][-> H]]; step; step; apply: [gX fm]@h. +Qed. +Program Definition insert x k v : + STsep {s} (shape x s, + [vfun _ : unit => shape x (ins k v s)]) := + Do (a <-- !x; + y <-- DynAssocList.insert a k v; + x ::= y). +Next Obligation. +move=>x k v [fm][/= _ [a][h][-> H]]; step. +by apply: [stepX fm]@h=>//= y {}h {}H; step; hhauto. +Qed. + +Program Definition remove x k : + STsep {s} (shape x s, + [vfun _ : unit => shape x (rem k s)]) := + Do (a <-- !x; + y <-- DynAssocList.remove V a k; + x ::= y). +Next Obligation. +move=>x k [fm][/= _ [a][h][-> H]]; step. +by apply: [stepX fm]@h=>//= y {}h {}H; step; hhauto. +Qed. + +Program Definition lookup x k : + STsep {s} (shape x s, + [vfun y m => m \In shape x s /\ y = fnd k s]) := + Do (a <-- !x; + DynAssocList.lookup V a k). +Next Obligation. +move=>x k [fm][/= _ [a][h][-> H]]; step. +by apply: [gX fm]@h=>//= y {}h {H}[]; hhauto. +Qed. + +(* ordered association list is a KV map *) +Definition AssocList := KVmap.make null new free insert remove lookup. End AssocList. End AssocList. diff --git a/pcm/heap.v b/pcm/heap.v index a1280be..be1ee5e 100644 --- a/pcm/heap.v +++ b/pcm/heap.v @@ -19,7 +19,7 @@ limitations under the License. From HB Require Import structures. From Coq Require Import ssreflect ssrbool ssrfun Eqdep. -From mathcomp Require Import ssrnat eqtype seq path. +From mathcomp Require Import ssrnat eqtype fintype seq path bigop. From pcm Require Import options axioms finmap. From pcm Require Import pcm unionmap natmap. @@ -237,7 +237,7 @@ Prenex Implicits heap_eta heap_eta2. (*******************************************) Section BlockUpdate. -Variable (A : Type). +Variable A : Type. Fixpoint updi (x : ptr) (vs : seq A) {struct vs} : heap := if vs is v'::vs' then x :-> v' \+ updi x.+1 vs' else Unit. @@ -336,10 +336,39 @@ move=>[E]; rewrite -!joinA=>D; case/(hcancelV D)=><-{}D. by case/(IH _ _ _ _ E D)=>->->. Qed. +Lemma updi_iota n (x : ptr) (f : nat -> A) : + updi x (map f (iota 0 n)) = + \big[join/Unit]_(i <- iota 0 n) x.+i :-> f i. +Proof. +elim: n x=>[|n IH] x; first by rewrite big_nil. +rewrite -addn1 iotaD add0n map_cat /= big_cat big_seq1 /=. +by rewrite updi_cat /= size_map size_iota /= unitR -IH. +Qed. + +Lemma updi_ord n (x : ptr) (f : 'I_n -> A) : + updi x (map f (enum 'I_n)) = + \big[join/Unit]_(i in 'I_n) x.+i :-> f i. +Proof. +case: n f=>[|n] f; first by rewrite -big_enum enum_ord0 big_nil. +set F := fun i => + if decP (b:=i < n.+1) idP is left pf then f (Ordinal pf) + else f (Ordinal (n:=n.+1) (m:=0) erefl). +have Ef i : f i = F i. +- rewrite /F; case: decP=>//. + by case: i=>i pf1 pf2; rewrite (pf_irr pf1 pf2). +set J := RHS. +have -> : J = \big[join/Unit]_(i in 'I_n.+1) x.+i :-> F i. +- by apply: eq_bigr=>i _; rewrite Ef. +rewrite -big_enum -(big_map _ predT (fun i=>x.+i :-> F i)). +rewrite val_enum_ord -updi_iota -val_enum_ord -map_comp. +by congr updi; apply: eq_map. +Qed. + End BlockUpdate. Lemma domeqUP A1 A2 x (xs1 : seq A1) (xs2 : seq A2) : - size xs1 = size xs2 -> dom_eq (updi x xs1) (updi x xs2). + size xs1 = size xs2 -> + dom_eq (updi x xs1) (updi x xs2). Proof. move=>E; apply/domeqP; split; first by rewrite !updiD E. apply/domE=>z; case: updiP=>[[H][m][->]|X]; first by rewrite updimV H E. diff --git a/pcm/pcm.v b/pcm/pcm.v index e0e248d..adb5b7a 100644 --- a/pcm/pcm.v +++ b/pcm/pcm.v @@ -1477,13 +1477,16 @@ by move=>x; rewrite !mem_filter /= in_set0. Qed. Lemma sepitF (s : {set I}) f1 f2 : - (forall x, x \in s -> f1 x =p f2 x) -> sepit s f1 =p sepit s f2. + (forall x, x \in s -> f1 x =p f2 x) -> + sepit s f1 =p sepit s f2. Proof. move=>H; apply: IterStar.eq_sepitF=>x H1; apply: H. by rewrite -mem_enum. Qed. -Lemma eq_sepit (s1 s2 : {set I}) f : s1 =i s2 -> sepit s1 f =p sepit s2 f. +Lemma eq_sepit (s1 s2 : {set I}) f : + s1 =i s2 -> + sepit s1 f =p sepit s2 f. Proof. by move/eq_enum=>H; apply: IterStar.perm_sepit; rewrite H. Qed. Lemma sepitS x (s : {set I}) f : @@ -1503,4 +1506,30 @@ Qed. Lemma sepitT1 x f : sepit setT f =p f x # sepit (setT :\ x) f. Proof. by rewrite (sepitS x) in_setT. Qed. +Lemma big_sepit (s : {set I}) (f : I -> U) m : + m = \big[join/Unit]_(i in s) (f i) <-> + m \In sepit s (fun i h => h = f i). +Proof. +rewrite {1}/sepit IterStar.sepitE/IterStar.sepit'/IterStar.seqjoin. +split; last first. +- case=>hs [Es Em] H. + suff {Em}E : hs = map f (enum s) by rewrite Em E big_map big_enum. + elim: hs (enum s) Es H=>[|h hs IH] es; first by move/esym/size0nil=>->. + by case: es=>//= e es [E][H1 H2]; rewrite H1 -IH. +move=>Em; exists (map f (enum s)); split. +- by rewrite size_map. +- by rewrite big_map big_enum. +apply/AllP; case=>a b. +rewrite {1}(_ : enum s = map id (enum s)); last by rewrite map_id. +by rewrite zip_map; case/MapP=>x /= _ [->->]. +Qed. + +Lemma big_sepitT (f : I -> U) m : + m = \big[join/Unit]_i (f i) <-> + m \In sepit setT (fun i h => h = f i). +Proof. +rewrite -big_sepit; apply: iff_sym. +by rewrite -big_enum /= enum_T big_enum. +Qed. + End FinIterStar. From 55dfc75656b7b975928557bfd00b6fc2a49ad85f Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Wed, 14 Aug 2024 13:50:04 +0200 Subject: [PATCH 47/93] modified congprog to use postconditions with vfun. --- _CoqProject | 1 + examples/congprog.v | 132 ++++++++++++++++++-------------------------- 2 files changed, 54 insertions(+), 79 deletions(-) diff --git a/_CoqProject b/_CoqProject index 666e241..27a2862 100644 --- a/_CoqProject +++ b/_CoqProject @@ -50,6 +50,7 @@ examples/hashtab.v examples/bubblesort.v examples/quicksort.v examples/congmath.v +examples/congprog.v examples/tree.v examples/union_find.v examples/graph.v \ No newline at end of file diff --git a/examples/congprog.v b/examples/congprog.v index a64f9ec..2cc680e 100644 --- a/examples/congprog.v +++ b/examples/congprog.v @@ -8,16 +8,10 @@ From mathcomp Require Import div finset seq fintype finfun choice. From pcm Require Import axioms prelude ordtype finmap pred pcm. From pcm Require Import unionmap heap autopcm automap. From htt Require Import options model heapauto llist array. -From htt Require Import kvmaps hashtab. -From htt Require Import congmath. -Import Prenex Implicits. +From htt Require Import kvmaps hashtab congmath. Notation finE := finset.inE. -Ltac add_morphism_tactic := SetoidTactics.add_morphism_tactic. -Notation " R ===> R' " := (@Morphisms.respectful _ _ R R') - (right associativity, at level 55) : signature_scope. - Variable s : seq constant. Notation symb := (symb s). @@ -82,17 +76,6 @@ Definition shape (R : rel_exp s) := End ShapePredicates. -Add Parametric Morphism r clist ulist htab p : (shape r clist ulist htab p) with - signature @EqPredType _ _ ===> @EqPredType _ _ as shape_morph. -Proof. -move=>R1 R2 H x /=. -split; case=>d [H1][H2][H3 H4]; exists d. -- by rewrite -H. -by rewrite H. -Qed. - - - Definition ith {I : finType} i (pf : i < #|I|) : I. Admitted. Lemma indx_ith {I : finType} i (pf : i < #|I|) : indx (ith i pf) = i. @@ -125,14 +108,14 @@ Admitted. Definition iT (clist : {array symb -> llist symb}): Type := forall k, STsep (fun i => k <= n /\ exists f, i \In Array.shape clist f # sepit [set c | indx c < k] (ctab f [ffun c => [:: c]]), - fun y m => y = Val tt /\ exists f, m \In Array.shape clist f # - sepit setT (ctab f [ffun c => [:: c]])). + [vfun (_ : unit) m => exists f, m \In Array.shape clist f # + sepit setT (ctab f [ffun c => [:: c]])]). Program Definition init : - STsep (emp, fun y m => exists ptr : ptrs, y = Val ptr /\ - let: Ptrs r c u h p := ptr in - m \In shape r c u h p empty_cong) := - Do (r <-- Array.newf [ffun x : symb=> x]; + STsep (emp, + [vfun y m => let: Ptrs r c u h p := y in + m \In shape r c u h p empty_cong]) := + Do (r <-- Array.newf [ffun x : symb => x]; clist <-- Array.new _ null; ffix (fun (loop : iT clist) k => Do (if decP (b:= k < n) idP is left pf then @@ -149,11 +132,11 @@ Next Obligation. move=>r clist loop k H i [pf][/= f][hc][hct][->{i} Hc Hct]. case: decP=>[{}pf|] /=; last first. - case: (ltngtP k n) pf=>// Ekn _ _; step=>_. - split=>//; exists f, hc, hct; split=>//. + exists f, hc, hct; split=>//. apply: tableP2 Hct=>// x; rewrite !finE Ekn. by rewrite /n cardE index_mem /index mem_enum. step=>x; step; apply: [stepX f]@hc=>//= [[m]] Em. -apply: [gE]=>[||?? []] //=; split=>//. +apply: [gE]=>//=; split=>//. eexists [ffun z => if z == ith k pf then x else f z]. rewrite (_ : _ \+ _ = m \+ (x :-> ith k pf \+ x.+1 :-> null \+ hct)); last by heap_congr. @@ -165,14 +148,13 @@ by rewrite !finE !ffunE indx_injE; case: eqP=>// ->; rewrite ltnn. Qed. Next Obligation. case=>_ ->; apply: [stepE]=>//= r hr Er; apply: [stepU]=>//= clist hc Ec. - apply: [stepX]@hc=>[||??[]] //=. +apply: [stepX]@hc=>//=. - split=>//; exists [ffun x => null], hc, Unit; rewrite unitR. by split=>//; rewrite (_ : [set c | indx c < 0] = set0) // sepit0. -case=>n0 [_][f][hc'][hrest][->] Hc' Hrest. +case=>_ [f][hc'][hrest][-> Hc' Hrest]. apply: [stepU]=>//= ulist hu Ehu; apply: [stepU]=>//= htab ht Ht. -step=>p; step=>_; exists (Ptrs r clist ulist htab p); split=>//. -exists (Data [ffun x => x] [ffun c => [:: c]] [ffun c => [::]] (nil K V) [::]). -split; last by case: (initP s). +set d := Data [ffun x => x] [ffun c => [:: c]] [ffun c => [::]] (nil K V) [::]. +step=>p; step; exists d; split; last by case: (initP s). split=>[s a|/=]; first by rewrite !ffunE !inE. exists f, [ffun s => null]. rewrite (_ : p :-> null \+ _ = hr \+ ((hc' \+ hrest) \+ (hu \+ Unit \+ @@ -193,22 +175,21 @@ Notation bshape' := (bshape r clist ulist htab p). Notation shape' := (shape r clist ulist htab p). Definition cT (a' b' : symb) : Type := - forall x:unit, STsep {D} + forall x : unit, STsep {D} (fun i => i \In ashape' D /\ a' != b', - fun y m => y = Val tt /\ exists ct, exists ut, - let: (d, _, _) := D in - m \In ashape' - (Data [ffun s => if s \in class d a' then b' - else rep d s] - [ffun s => if s == a' then [::] else - if s == b' then rev (class d a') ++ class d b' - else class d s] - (use d) (lookup d) (pending d), ct, ut)). + [vfun (_ : unit) m => exists ct, exists ut, + let: (d, _, _) := D in + m \In ashape' + (Data [ffun s => if s \in class d a' then b' + else rep d s] + [ffun s => if s == a' then [::] else + if s == b' then rev (class d a') ++ class d b' + else class d s] + (use d) (lookup d) (pending d), ct, ut)]). Program Definition join_hclass (a' b' : symb) : STsep {d} (fun i => i \In bshape' d /\ a' != b', - fun y m => y = Val tt /\ - m \In bshape' (join_class d a' b')) := + [vfun (_ : unit) m => m \In bshape' (join_class d a' b')]) := Do (ffix (fun (loop : cT a' b') (x : unit) => Do (a <-- Array.read clist a'; b <-- Array.read clist b'; @@ -230,8 +211,7 @@ case=>ctb [ctx][->{w}Ctb Ctx]; rewrite /ctab/table in Cta Ctb. apply: [stepX ct, th]@th=>//= _ _ [->->]. apply: [stepX ct, th]@th=>//= _ _ [->->]. apply: vrfV=>V; case: (ct a' =P null) Cta=>[/[dup] Ea' ->|/eqP Na']. -- case/(lseq_null (validX V))=>/= ->->{cta V}. - step=>{}V; split=>//; exists ct, ut. +- case/(lseq_null (validX V))=>/= ->->{cta V}; step=>{}V; exists ct, ut. rewrite (_ : rh \+ _ = rh \+ (th \+ (Unit \+ (ctb \+ ctx)) \+ h)); last by heap_congr. rewrite -fin_eta; hhauto; rewrite (sepitT1 a'); hhauto. @@ -260,7 +240,7 @@ last by move=>?? []. by rewrite eq_sym (negbTE N); hhauto. apply: tableP Ctx=>x; rewrite /ct2/cv2/ct1 !ffunE /= ?ffunE /=; by rewrite !finE andbT; case/andP=>/negbTE -> /negbTE ->. -case=>m [_][] {Th2}ct2 [ut2] /= Hm _; split=>//; exists ct2, ut2. +case=>m [{Th2}ct2][ut2] /= Hm _; exists ct2, ut2. congr (m \In ashape' (Data _ _ _ _ _, ct2, ut2)): Hm; apply/ffunP=>x. - by rewrite !ffunE eqxx {2}Eca inE /=; case: (x =P s)=>//= _; rewrite if_same. rewrite !ffunE !eqxx /= (eq_sym b') (negbTE N). @@ -269,9 +249,9 @@ by rewrite Eca rev_cons cat_rcons. Qed. Next Obligation. move=>a' b' [d][i][[C]][/= ct][ut H] N. -apply: [gE (d, ct, ut)]=>[||?? []] //= [m][_][ct1][ut1] W _. +apply: [gE (d, ct, ut)]=>//= [[m]][ct1][ut1] W _. set d' := (Data _ _ _ _ _) in W; suff E : join_class d a' b' = d'. -- by split=>//; split; [apply: join_class_classP|rewrite E; eauto]. +- by split; [apply: join_class_classP|rewrite E; eauto]. by congr Data; apply/ffunP=>x; rewrite !ffunE /= C. Qed. @@ -279,13 +259,11 @@ Definition uT (a' b' : symb) := forall x : unit, STsep {d} (fun i => exists don, a' != b' /\ i \In bshape' (join_use' d a' b' don) /\ use d a' = don ++ use (join_use' d a' b' don) a', - fun y m => y = Val tt /\ - m \In bshape' (join_use d a' b')). + [vfun (_ : unit) m => m \In bshape' (join_use d a' b')]). Program Definition join_huse (a' b' : symb) : STsep {d} (fun i => a' != b' /\ i \In bshape' d, - fun y m => y = Val tt /\ - m \In bshape' (join_use d a' b')) := + [vfun (_ : unit) m => m \In bshape' (join_use d a' b')]) := Do (ffix (fun (loop : uT a' b') x => Do (a <-- Array.read ulist a'; if a == null then ret tt @@ -319,7 +297,7 @@ rewrite (sepitS b') !finE eq_sym {1}N; case=>hb' [h][->{w} Ub R]. rewrite /utab/table/= in Ua Ub. apply: [stepX ut, ru]@ru=>//= _ _ [->->]; apply: vrfV=>V. case: (ut a' =P null) Ua=>[/[dup] Ea' ->|/eqP Na' {V}]. -- case/(lseq_null (validX V))=>/= {V} Eu1 ->{ha'}; step=>_; split=>//. +- case/(lseq_null (validX V))=>/= {V} Eu1 ->{ha'}; step. rewrite (_ : rh \+ _ = rh \+ (cth \+ (ru \+ (Unit \+ (hb' \+ h)) \+ (ht \+ (p:-> p' \+ hp))))); last by heap_congr. rewrite /join_use Eu Eu1 cats0 -/d1; split=>//=; exists ct, ut. @@ -383,9 +361,9 @@ Let pend1 (e : pend s) := Notation "e ..0" := (pend0 e) (at level 2). Notation "e ..1" := (pend1 e) (at level 2). -Definition pT : Type := forall x:unit, +Definition pT : Type := forall x : unit, STsep {d} (fun i => i \In bshape' d, - fun y m => y = Val tt /\ m \In bshape' (propagate d)). + [vfun (_ : unit) m => m \In bshape' (propagate d)]). Program Definition hpropagate := ffix (fun (loop : pT) x => @@ -409,7 +387,7 @@ move=>loop [][/= d][i][C][/= ct][ut][hr][_][-> Hr][hc][_][-> Hc] [hu][_][-> Hu][ht][_][-> Ht][q][_][hp][->] <- Hp. step; case: (q =P null) Hp=>[->{q}|/eqP N] Hp. - apply: vrfV=>V; case/(lseq_null (validX V)): Hp=>{V} Ep /= ->{hp}. - step=>_; rewrite propagate_equation Ep; do !split=>//=. + step; rewrite propagate_equation Ep; split=>//=. by exists ct, ut; hhauto; rewrite Ep. case/(lseq_pos N): Hp=>eq [next][hnext][Ep] ->{hp} Hp; do !step. apply: [stepX rep d, hr]@hr=>//= a' _ [-> Ea']. @@ -421,10 +399,8 @@ case: ifPn=>[E|E]. by case: eq Ep Ea' Eb' E=>[a2 b2 _|[[a2 ??]][[b2 ??]] _] ->->->. apply: [gE d1]=>[||??[] //] /=; last by rewrite T1. by split=>//=; exists ct, ut; hhauto. -apply: [stepE d1]=>[||??[]//] /=. -- by do !split=>//; exists ct, ut; hhauto. -set d2 := (join_class _ _ _); case=>m [_] H. -apply: [stepE d2]=>[//||??[]//] /= [n][_] Hn. +apply: [stepE d1]=>//=; first by do 2split=>//; exists ct, ut; hhauto. +set d2 := (join_class _ _ _); case=>m H; apply: [stepE d2]=>//= [[n]] Hn. set d3 := (join_use _ _ _) in Hn. suff -> : propagate d = propagate d3 by apply: [gE d3]. rewrite /d3/d2 propagate_equation Ep; @@ -433,8 +409,8 @@ Qed. Program Definition merge (e : Eq s) : STsep {R} (fun i => i \In shape' R, - fun y m => y = Val tt /\ - m \In shape' (closure (R \+p eq2rel e))) := + [vfun (_ : unit) m => + m \In shape' (closure (R \+p eq2rel e))]) := match e with | simp_eq a b => Do (q <-- !p; @@ -470,13 +446,13 @@ move=>e a b [R][_][d][[C]][/= ct][ut][hr][_][-> Hr][hc][_][-> Hc][_][_] [->][hu][hu'][-> Hu Hu'][ht][_][-> Ht][q][_][hp][->] <- Hp [PI][Ep Erel]. step; apply: [stepX (pending d)]@hp=>//= x _ [r][{Hp}hp][-> Hp]. set d1:=Data (rep d) (class d) (use d) (lookup d) (simp_pend a b :: pending d). -step; apply: [gE d1]=>[||??[]] //=. +step; apply: [gE d1]=>//=. - rewrite Ep /= in Hp; case: Hp=>->{r} ->{hp}. rewrite (_ : _ \+ _ = hr \+ (hc \+ (hu \+ hu' \+ (ht \+ (p :-> x \+ (x :-> simp_pend a b \+ (x.+1 :-> null \+ Unit))))))); last by heap_congr. by split=>//=; exists ct, ut; hhauto; rewrite Ep. -case=>m [_ Hm _]; split=>//. +case=>m Hm _. have L: propagate_inv d1 by apply: propagate_pendP PI. case: (propagatePE L)=>L1 [L2 L3]; exists (propagate d1); do 3!split=>//. by rewrite -L3 -Erel clos_clos /CRel -!orrA orrAC. @@ -493,12 +469,12 @@ case: v Ev=>[[[e e1 e2]]|] Ev. set d1 := Data (rep d) (class d) (use d) (lookup d) (comp_pend cx ex :: pending d). step; apply: [stepX pending d]@hp=>//= x _ [x1][{}hp][-> {}Hp]. - step; apply: [gE d1]=>[||??[]] //=. + step; apply: [gE d1]=>//=. - rewrite (_ : _ \+ _ = hr \+ (hc \+ (hu \+ hu' \+ (ht \+ (p :-> x \+ (x :-> comp_pend cx ex \+ (x.+1 :-> x1 \+ hp))))))); last by heap_congr. by split=>//; exists ct, ut; hhauto. - case=>m [_] Hm _; split=>//. + case=>m Hm _. have L : propagate_inv d1 by apply: propagate_pendP PI. case: (propagatePE L)=>L1 [L2] L3; exists (propagate d1). by rewrite -L3 -Erel propagate_clos_pendP. @@ -518,7 +494,7 @@ case E : (rep d c2 == rep d c1). - apply: [stepX cx::use d (rep d c1)]@(x :-> _ \+ x.+1 :-> _ \+ hu'')=>//=. - by exists r, hu''; rewrite joinA (eqP E) /ut1 !ffunE /= eqxx. move=>v {Hu''}_ [x'][_][->][r0][{}hu''][-> Hu'']. - apply: [gX ut1]@hu=>//= [[]] {Hu}hu Hu _; split=>//. + apply: [gX ut1]@hu=>//= [[]] {Hu}hu Hu _. rewrite (_ : _ \+ _ = hr \+ (hc \+ (hu \+ (v :-> cx \+ (v.+1 :-> x' \+ (x' :-> cx \+ (x'.+1 :-> r0 \+ hu''))) \+ hu2) \+ (ht \+ (p :-> q \+ hp))))); last by heap_congr. @@ -533,7 +509,7 @@ move: Hu'; rewrite (sepitS (rep d c2)) !finE E /=. case=>hu1 [hu3][->{hu2} Hu1' Hu2']. apply: [stepX use d (rep d c2)]@hu1=>//=; first by rewrite /ut1 ffunE /= E. move=>v _ [x'][{Hu1'}hu1][-> Hu1']. -apply: [gX ut1]@hu=>//= [[]] {Hu}hu Hu _; split=>//. +apply: [gX ut1]@hu=>//= [[]] {Hu}hu Hu _. rewrite (_ : _ \+ _ = hr \+ (hc \+ (hu \+ (x :-> cx \+ (x.+1 :-> r \+ hu'') \+ ((v :-> cx \+ (v.+1 :-> x' \+ hu1) \+ hu3))) \+ @@ -555,7 +531,7 @@ Notation "e ..0" := (pend3 e) (at level 2). Definition nT : Type := forall t, STsep {d} (fun i => i \In bshape' d, - fun y m => m \In bshape' d /\ y = Val (norm d t)). + [vfun y m => y = norm d t /\ m \In bshape' d]). Program Definition hnorm := ffix (fun (hnorm : nT) (t : exp s) => @@ -579,10 +555,9 @@ move=>hnorm t [d][_][Ci][/= ct][ut][hr][hrest][-> Hr Hrest]. case: t=>[s|t1 t2]. - apply: [stepX rep d, hr]@hr=>//= _ _ [->->]. by step; do 2!split=>//; exists ct, ut; hhauto. -apply: [stepE d]=>[||??[]] //=. -- by split=>//; exists ct, ut; hhauto. -move=>u1 m [H][E1]; apply: [stepE d]=>[||??[]] //=. -move=>u2 {}m [{}H [E2]] {Ci ct ut hr hrest Hr Hrest}. +apply: [stepE d]=>//=; first by split=>//; exists ct, ut; hhauto. +move=>u1 m [E1 H]; apply: [stepE d]=>//=. +move=>u2 {}m [E2 {}H] {Ci ct ut hr hrest Hr Hrest}. case: H=>Ci [/= ct][ut][hr][w][->{m} Hr][hc][w'] [->{w} Hc][hu][w][->{w'} Hu][ht][hrest][->{w} Ht Hrest]. case: u1 E1=>[w1|??] /= E1; last first. @@ -598,16 +573,15 @@ Qed. Program Definition check t1 t2 : STsep {R} (fun i => i \In shape' R, - fun y m => m \In shape' R /\ - exists b:bool, y = Val b /\ (b <-> (t1, t2) \In R)) := + [vfun (b : bool) m => m \In shape' R /\ + (b <-> (t1, t2) \In R)]) := Do (u1 <-- hnorm t1; u2 <-- hnorm t2; ret (u1 == u2)). Next Obligation. move=>t1 t2 [R][h][d][H][[RI X]][Ep PI]. -apply: [stepX d]@h=>[||??[]] //= _ {H}h [H][->]. -apply: [stepX d]@h=>[||??[]] //= _ {H}h [H][->]. -step=>_. split; first by exists d. -exists (norm d t1 == norm d t2); split=>//. +apply: [stepX d]@h=>//= _ {H}h [-> H]. +apply: [stepX d]@h=>//= _ {H}h [-> H]. +step; split; first by exists d. by case: normP=>//; rewrite PI. Qed. From cd96ba584b2ed03d9927d125f4edb11dfc19462d Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Wed, 14 Aug 2024 14:55:16 +0200 Subject: [PATCH 48/93] abstracting the code in congprog.v --- examples/congprog.v | 181 +++++++++++++++++++++++++++----------------- examples/kvmaps.v | 12 +-- 2 files changed, 120 insertions(+), 73 deletions(-) diff --git a/examples/congprog.v b/examples/congprog.v index 2cc680e..f2b8e52 100644 --- a/examples/congprog.v +++ b/examples/congprog.v @@ -12,6 +12,75 @@ From htt Require Import kvmaps hashtab congmath. Notation finE := finset.inE. +Definition ith {I : finType} i (pf : i < #|I|) : I. Admitted. + +Lemma indx_ith {I : finType} i (pf : i < #|I|) : indx (ith i pf) = i. +Proof. +admit. +Admitted. +(* by move=>i pf; rewrite /ith /= /indx index_uniq // -?cardE // enum_uniq. *) + +Lemma ith_indx {I : finType} (s : I) (pf : indx s < #|I|) : ith (indx s) pf = s. +Proof. Admitted. +(* by move=>s /= pf; rewrite /ith /= nth_index // mem_enum. Qed. *) + +Lemma indx_inj I : injective (@indx I). +Admitted. + +Lemma indx_injE {I : finType} s i (pf : i < #|I|) : (s == ith i pf) = (indx s == i). +Proof. +Admitted. + + +Lemma sepit_emp (A : eqType) (s : seq A) f : + (forall x, x \in s -> f x =p emp (U:=heap)) -> + IterStar.sepit s f =p emp. +Proof. +Admitted. + + +(* empty congruence only relates constant symbols to themselves *) +Definition empty_cong s := closure (graph (@const s)). + +(*************) +(* Signature *) +(*************) + +Module Type CongrSig. +(* abstract type for collection of root pointers *) +(* seq const is the list of constants over which *) +(* structure computes *) +(* this is a global argument, hence exposed by tp *) +Parameter tp : seq constant -> Set. + +(* abstract predicate tying roots, (congruence) relation, heap *) +Parameter shape : forall {s}, tp s -> rel_exp s -> Pred heap. + +(* initialize empty congruence structure; return roots *) +Parameter init : forall {s}, + STsep (emp, [vfun rt m => m \In shape rt (empty_cong s)]). + +(* merge new equation into the structure rooted by rt *) +Parameter merge : forall {s} rt (e : Eq s), + STsep {R} (fun i => i \In shape rt R, + [vfun (_ : unit) m => + m \In shape rt (closure (R \+p eq2rel e))]). + +(* check if two expressions are congruent in the structure rooted by rt *) +Parameter check : forall {s} rt (t1 t2 : exp s), + STsep {R} (fun i => i \In shape rt R, + [vfun (b : bool) m => m \In shape rt R /\ + (b <-> (t1, t2) \In R)]). +End CongrSig. + +(******************) +(* Implementation *) +(******************) + +(* faithful to Barcelogic's actual implementation *) + +Module Congr : CongrSig. +Section Congr. Variable s : seq constant. Notation symb := (symb s). @@ -41,20 +110,21 @@ Definition utab := Definition n := #|{: symb}|. -(* the empty congruence is one that only relates constant symbols to themselves *) -Definition empty_cong := closure (graph (@const s)). - (* the algorithm starts with root pointers for the data *) -Inductive ptrs : Type := +Inductive ptrs : Set := Ptrs of {array symb -> symb} & {array symb -> llist symb} & {array symb -> llist (symb*symb*symb)} & KVmap.tp LT & ptr. +(* renaming type to satisfy the signature check *) +Definition tp := ptrs. + Section ShapePredicates. -Variable (r : {array symb -> symb}). -Variable (clist : {array symb -> llist symb}). -Variable (ulist : {array symb -> llist (symb*symb*symb)}). -Variable (htab : KVmap.tp LT). -Variable (p : ptr). +Variable rt : ptrs. +Let r := let: Ptrs r clist ulist htab p := rt in r. +Let clist := let: Ptrs r clist ulist htab p := rt in clist. +Let ulist := let: Ptrs r clist ulist htab p := rt in ulist. +Let htab := let: Ptrs r clist ulist htab p := rt in htab. +Let p := let: Ptrs r clist ulist htab p := rt in p. (* The layout of the data structure *) @@ -70,39 +140,12 @@ Definition ashape D := Definition bshape d := [Pred h | class_inv d /\ exists ct ut, h \In ashape (d, ct, ut)]. -Definition shape (R : rel_exp s) := +Definition shape (R : rel_exp s) : Pred heap := [Pred h | exists d, h \In bshape d /\ propagate_inv d /\ pending d = [::] /\ CRel d <~> R]. End ShapePredicates. -Definition ith {I : finType} i (pf : i < #|I|) : I. Admitted. - -Lemma indx_ith {I : finType} i (pf : i < #|I|) : indx (ith i pf) = i. -Proof. -admit. -Admitted. -(* by move=>i pf; rewrite /ith /= /indx index_uniq // -?cardE // enum_uniq. *) - -Lemma ith_indx {I : finType} (s : I) (pf : indx s < #|I|) : ith (indx s) pf = s. -Proof. Admitted. -(* by move=>s /= pf; rewrite /ith /= nth_index // mem_enum. Qed. *) - -Lemma indx_inj I : injective (@indx I). -Admitted. - -Lemma indx_injE {I : finType} s i (pf : i < #|I|) : (s == ith i pf) = (indx s == i). -Proof. -Admitted. - - -Lemma sepit_emp (A : eqType) (s : seq A) f : - (forall x, x \in s -> f x =p emp (U:=heap)) -> - IterStar.sepit s f =p emp. -Proof. -Admitted. - - (* Initialization procedure to generate the empty structure *) Definition iT (clist : {array symb -> llist symb}): Type := forall k, @@ -112,9 +155,7 @@ Definition iT (clist : {array symb -> llist symb}): Type := forall k, sepit setT (ctab f [ffun c => [:: c]])]). Program Definition init : - STsep (emp, - [vfun y m => let: Ptrs r c u h p := y in - m \In shape r c u h p empty_cong]) := + STsep (emp, [vfun rt m => m \In shape rt (empty_cong s)]) := Do (r <-- Array.newf [ffun x : symb => x]; clist <-- Array.new _ null; ffix (fun (loop : iT clist) k => @@ -128,7 +169,7 @@ Program Definition init : htab <-- KVmap.new LT; p <-- alloc null; ret (Ptrs r clist ulist htab p)). -Next Obligation. +Next Obligation. move=>r clist loop k H i [pf][/= f][hc][hct][->{i} Hc Hct]. case: decP=>[{}pf|] /=; last first. - case: (ltngtP k n) pf=>// Ekn _ _; step=>_. @@ -142,7 +183,7 @@ rewrite (_ : _ \+ _ = m \+ (x :-> ith k pf \+ x.+1 :-> null \+ hct)); last by heap_congr. hhauto; rewrite (sepitS (ith k pf)) finE indx_ith ltnSn. rewrite /ctab/table !ffunE eqxx; hhauto. -apply: tableP2 Hct=>// s. +apply: tableP2 Hct=>// a. - by rewrite !finE ltnS indx_injE; case: ltngtP. by rewrite !finE !ffunE indx_injE; case: eqP=>// ->; rewrite ltnn. Qed. @@ -155,7 +196,7 @@ case=>_ [f][hc'][hrest][-> Hc' Hrest]. apply: [stepU]=>//= ulist hu Ehu; apply: [stepU]=>//= htab ht Ht. set d := Data [ffun x => x] [ffun c => [:: c]] [ffun c => [::]] (nil K V) [::]. step=>p; step; exists d; split; last by case: (initP s). -split=>[s a|/=]; first by rewrite !ffunE !inE. +split=>[a b|/=]; first by rewrite !ffunE !inE. exists f, [ffun s => null]. rewrite (_ : p :-> null \+ _ = hr \+ ((hc' \+ hrest) \+ (hu \+ Unit \+ (ht \+ (p :-> null \+ Unit))))); last by rewrite unitR; heap_congr. @@ -163,16 +204,17 @@ hhauto; rewrite sepit_emp //= => k. by rewrite /utab/table !ffunE; split=>//; case=>_ ->. Qed. +Section Internal. +Variable rt : ptrs. +Notation ashape' := (ashape rt). +Notation bshape' := (bshape rt). +Notation shape' := (shape rt). -Variable (r : {array symb -> symb}). -Variable (clist : {array symb -> llist symb}). -Variable (ulist : {array symb -> llist (symb*symb*symb)}). -Variable (htab : KVmap.tp LT). -Variable (p : ptr). - -Notation ashape' := (ashape r clist ulist htab p). -Notation bshape' := (bshape r clist ulist htab p). -Notation shape' := (shape r clist ulist htab p). +Let r := let: Ptrs r clist ulist htab p := rt in r. +Let clist := let: Ptrs r clist ulist htab p := rt in clist. +Let ulist := let: Ptrs r clist ulist htab p := rt in ulist. +Let htab := let: Ptrs r clist ulist htab p := rt in htab. +Let p := let: Ptrs r clist ulist htab p := rt in p. Definition cT (a' b' : symb) : Type := forall x : unit, STsep {D} @@ -220,7 +262,7 @@ apply: vrfV=>V; case: (ct a' =P null) Cta=>[/[dup] Ea' ->|/eqP Na']. - by rewrite /ctab/table /= ffunE eqxx eq_sym (negbTE N). apply: tableP Ctx=>// x; rewrite !finE andbT ffunE. by case/andP=>/negbTE -> /negbTE ->. -case/(lseq_pos Na')=>s {V} [next][cta'][Eca ->{cta}Cta']. +case/(lseq_pos Na')=>a {V} [next][cta'][Eca ->{cta}Cta']. do 3!step; apply: [stepX ct]@th=>//= [_] {th Th} th1 Th1. set ct1 := (finfun _) in Th1. apply: [stepX ct1]@th1=>//= [_] {th1 Th1} th2 Th2. @@ -228,10 +270,10 @@ set ct2 := (finfun _) in Th2. apply: [stepX rep d]@rh=>//= _ {rh Rh} rh1 Rh1. set r1 := (finfun _) in Rh1. set cv2 := [ffun z => if z == a' then (behead (class d a')) - else if z == b' then s :: class d b' else class d z]. + else if z == b' then a :: class d b' else class d z]. apply: [gE (Data r1 cv2 (use d) (lookup d) (pending d), ct2, ut)]=>/=; last by move=>?? []. -- rewrite (_ : rh1 \+ _ = rh1 \+ (th2 \+ (cta' \+ (ct a' :-> s \+ +- rewrite (_ : rh1 \+ _ = rh1 \+ (th2 \+ (cta' \+ (ct a' :-> a \+ ((ct a').+1 :-> ct b' \+ ctb) \+ ctx)) \+ h)); last by heap_congr. hhauto; rewrite (sepitT1 a'); hhauto. - by rewrite /ctab/table/ct2/cv2 !ffunE /= eqxx. @@ -242,7 +284,7 @@ last by move=>?? []. by rewrite !finE andbT; case/andP=>/negbTE -> /negbTE ->. case=>m [{Th2}ct2][ut2] /= Hm _; exists ct2, ut2. congr (m \In ashape' (Data _ _ _ _ _, ct2, ut2)): Hm; apply/ffunP=>x. -- by rewrite !ffunE eqxx {2}Eca inE /=; case: (x =P s)=>//= _; rewrite if_same. +- by rewrite !ffunE eqxx {2}Eca inE /=; case: (x =P a)=>//= _; rewrite if_same. rewrite !ffunE !eqxx /= (eq_sym b') (negbTE N). case: (x =P a')=>// _; case: (x =P b')=>// _. by rewrite Eca rev_cons cat_rcons. @@ -318,13 +360,13 @@ have Ej2: join_use d2 a' b' = join_use d1 a' b'. by rewrite -!(join_useT (t:=[::])) ?cats0 -?catA //= ?Eu -?Eu1. case: v Eqv=>[[[e1 e2 e3]]|] /= /esym Eqv. - set e := (e1, e2, e3) in Eqv. - do 3!step; apply: [stepX pending d1]@hp=>//= q _ [r][{Hp}hp [-> Hp]]. + do 3!step; apply: [stepX pending d1]@hp=>//= q _ [r0][{Hp}hp [-> Hp]]. step; apply: [gE d]=>[||??[]] //=. exists (don ++ [:: c]); rewrite -/d2 Eu2 -catA; do 2!split=>//. rewrite /bshape'/class_inv/ashape/d2 (join_useT (t:=a1')) //= Eqv -/d1 /=. rewrite (_ : _ \+ _ = rh \+ (cth \+ ((ru \+ (hh \+ (hb' \+ h))) \+ (ht \+ (p :-> q \+ (q :-> comp_pend c e \+ - (q.+1 :-> r \+ hp))))))); last by heap_congr. + (q.+1 :-> r0 \+ hp))))))); last by heap_congr. case: Htc=>x [y][->] X1 X2; hhauto; [eauto|eauto|exact: Ut2|]. rewrite (sepitT1 a'); hhauto. - by rewrite /utab/table/ut2 !ffunE /= eqxx. @@ -353,7 +395,6 @@ Next Obligation. by move=>a' b' [/= d][i][N H]; apply: [gE d]=>[||??[]] //; exists [::]. Qed. - Let pend0 (e : pend s) := match e with simp_pend a b => a | comp_pend (a,_,_) (b,_,_) => a end. Let pend1 (e : pend s) := @@ -444,10 +485,10 @@ Program Definition merge (e : Eq s) : Next Obligation. move=>e a b [R][_][d][[C]][/= ct][ut][hr][_][-> Hr][hc][_][-> Hc][_][_] [->][hu][hu'][-> Hu Hu'][ht][_][-> Ht][q][_][hp][->] <- Hp [PI][Ep Erel]. -step; apply: [stepX (pending d)]@hp=>//= x _ [r][{Hp}hp][-> Hp]. +step; apply: [stepX (pending d)]@hp=>//= x _ [r0][{Hp}hp][-> Hp]. set d1:=Data (rep d) (class d) (use d) (lookup d) (simp_pend a b :: pending d). step; apply: [gE d1]=>//=. -- rewrite Ep /= in Hp; case: Hp=>->{r} ->{hp}. +- rewrite Ep /= in Hp; case: Hp=>->{r0} ->{hp}. rewrite (_ : _ \+ _ = hr \+ (hc \+ (hu \+ hu' \+ (ht \+ (p :-> x \+ (x :-> simp_pend a b \+ (x.+1 :-> null \+ Unit))))))); last by heap_congr. @@ -482,7 +523,7 @@ apply: [stepX lookup d]@ht=>//= _ {Ht}ht Ht. apply: [stepX ut, hu]@hu=>//= _ _ [->->]. move: Hu'; rewrite (sepitT1 (rep d c1)). case=>hu'' [hu2][->{hu'} Hu'' Hu']. -apply: [stepX use d (rep d c1)]@hu''=>//= x _ [r][{Hu''}hu''][-> Hu'']. +apply: [stepX use d (rep d c1)]@hu''=>//= x _ [r0][{Hu''}hu''][-> Hu'']. apply: [stepX ut]@hu=>//= _ {Hu}hu Hu; set ut1 := (finfun _) in Hu. apply: [stepX ut1, hu]@hu=>//= _ _ [->->]. set u1' := [ffun z => if z == rep d c1 then cx :: use d z else use d z]. @@ -492,11 +533,11 @@ set d1 := Data (rep d) (class d) u' l' (pending d). pose ut2 x' := [ffun z => if z == rep d c2 then x' else ut1 z]. case E : (rep d c2 == rep d c1). - apply: [stepX cx::use d (rep d c1)]@(x :-> _ \+ x.+1 :-> _ \+ hu'')=>//=. - - by exists r, hu''; rewrite joinA (eqP E) /ut1 !ffunE /= eqxx. - move=>v {Hu''}_ [x'][_][->][r0][{}hu''][-> Hu'']. + - by exists r0, hu''; rewrite joinA (eqP E) /ut1 !ffunE /= eqxx. + move=>v {Hu''}_ [x'][_][->][r1][{}hu''][-> Hu'']. apply: [gX ut1]@hu=>//= [[]] {Hu}hu Hu _. rewrite (_ : _ \+ _ = hr \+ (hc \+ (hu \+ (v :-> cx \+ - (v.+1 :-> x' \+ (x' :-> cx \+ (x'.+1 :-> r0 \+ hu''))) \+ + (v.+1 :-> x' \+ (x' :-> cx \+ (x'.+1 :-> r1 \+ hu''))) \+ hu2) \+ (ht \+ (p :-> q \+ hp))))); last by heap_congr. exists d1; split=>//; last first. - split; first by apply: propagate_nopendP. @@ -511,7 +552,7 @@ apply: [stepX use d (rep d c2)]@hu1=>//=; first by rewrite /ut1 ffunE /= E. move=>v _ [x'][{Hu1'}hu1][-> Hu1']. apply: [gX ut1]@hu=>//= [[]] {Hu}hu Hu _. rewrite (_ : _ \+ _ = hr \+ (hc \+ (hu \+ - (x :-> cx \+ (x.+1 :-> r \+ hu'') \+ + (x :-> cx \+ (x.+1 :-> r0 \+ hu'') \+ ((v :-> cx \+ (v.+1 :-> x' \+ hu1) \+ hu3))) \+ ((ht \+ (p :-> q \+ hp)))))); last by heap_congr. exists d1; split; last first. @@ -525,7 +566,6 @@ hhauto; apply: tableP Hu2'=>/= a; rewrite !finE !ffunE /= andbT; by case/andP=>/negbTE -> /negbTE ->. Qed. - Let pend3 (e : symb*symb*symb) := let: (a, _, _) := e in a. Notation "e ..0" := (pend3 e) (at level 2). @@ -552,7 +592,7 @@ Program Definition hnorm := end)). Next Obligation. move=>hnorm t [d][_][Ci][/= ct][ut][hr][hrest][-> Hr Hrest]. -case: t=>[s|t1 t2]. +case: t=>[a|t1 t2]. - apply: [stepX rep d, hr]@hr=>//= _ _ [->->]. by step; do 2!split=>//; exists ct, ut; hhauto. apply: [stepE d]=>//=; first by split=>//; exists ct, ut; hhauto. @@ -585,3 +625,8 @@ apply: [stepX d]@h=>//= _ {H}h [-> H]. step; split; first by exists d. by case: normP=>//; rewrite PI. Qed. + +End Internal. +End Congr. +End Congr. + diff --git a/examples/kvmaps.v b/examples/kvmaps.v index a4705ce..b3d515a 100644 --- a/examples/kvmaps.v +++ b/examples/kvmaps.v @@ -25,14 +25,16 @@ From htt Require Import options model heapauto. (* stateful KV map ADT *) (***********************) -(* Dynamic KV map is determined by the root pointer. *) +(* Dynamic KV map is determined by its root pointer(s). *) (* Functions such insert and remove may modify *) (* the root, and will correspondingly return the new one. *) -(* tp is abstracted to facilitate passing K and V to methods *) -(* (thus, the move reduces annotations) *) +(* Tp is abstracted to facilitate structures that may *) +(* have more than one root pointers. Also, it enables *) +(* passing K and V to methods thus reducing annotations *) + Module DynKVmap. Record Sig (K : ordType) (V : Type) : Type := - make {tp :> Type; + make {tp :> Set; default : tp; shape : tp -> {finMap K -> V} -> Pred heap; new : STsep (emp, [vfun x => shape x (nil K V)]); @@ -58,7 +60,7 @@ End DynKVmap. Module KVmap. Record Sig (K : ordType) (V : Type) : Type := - make {tp :> Type; + make {tp :> Set; default : tp; shape : tp -> {finMap K -> V} -> Pred heap; (* allocate root pointer and empty structure *) From 241fac7314d69c6df0bcfac74c5c407bee65415f Mon Sep 17 00:00:00 2001 From: kevinlopez Date: Wed, 14 Aug 2024 14:59:41 +0200 Subject: [PATCH 49/93] defining shape predicate --- examples/schorr.v | 9 +++++++++ 1 file changed, 9 insertions(+) diff --git a/examples/schorr.v b/examples/schorr.v index 9275d34..82d8c60 100644 --- a/examples/schorr.v +++ b/examples/schorr.v @@ -139,3 +139,12 @@ Definition reach (m : nmap mark) (g : pregraph) (s : seq node) t := (* avoiding marked nodes *) exists y, y \in s /\ x \in connect (mem (dom m)) g (gr g y). +Definition shape g0 r p t := + fun h:heap => (exists s g m, + dom g = dom h /\ + h = lay m g /\ + path (ch m g) null s /\ + rev_graph m g s t = g0 /\ + reach m g s t /\ + p = last null s /\ + r = head t s). From 230f0f7f4cb2d2bac946d98e7cd3784046ae2168 Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Wed, 14 Aug 2024 15:04:27 +0200 Subject: [PATCH 50/93] made abstract type of root pointers of arrays and kvmaps be small (ie.. declared as Set) --- examples/array.v | 2 +- examples/kvmaps.v | 3 ++- 2 files changed, 3 insertions(+), 2 deletions(-) diff --git a/examples/array.v b/examples/array.v index 466e7ac..f92b21e 100644 --- a/examples/array.v +++ b/examples/array.v @@ -26,7 +26,7 @@ Definition indx {I : finType} (x : I) := index x (enum I). (* array is (pointer to) a contiguous memory region *) (* holding the array values *) -Record array (I : finType) (T : Type) : Type := Array {orig :> ptr}. +Record array (I : finType) (T : Type) : Set := Array {orig :> ptr}. Arguments Array {I T}. Definition array_for (I : finType) (T : Type) of phant (I -> T) := array I T. diff --git a/examples/kvmaps.v b/examples/kvmaps.v index b3d515a..267de8c 100644 --- a/examples/kvmaps.v +++ b/examples/kvmaps.v @@ -31,7 +31,8 @@ From htt Require Import options model heapauto. (* Tp is abstracted to facilitate structures that may *) (* have more than one root pointers. Also, it enables *) (* passing K and V to methods thus reducing annotations *) - +(* There's no deep reason to make tp : Set, except that *) +(* it should be thought of a collection of pointers, hence small. *) Module DynKVmap. Record Sig (K : ordType) (V : Type) : Type := make {tp :> Set; From 5616c43e922e43dc56c404b516e376579b4fff12 Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Wed, 14 Aug 2024 18:26:01 +0200 Subject: [PATCH 51/93] developed a number of lemmas about star. to be moved to other files later --- examples/congmath.v | 7 +- examples/congprog.v | 265 +++++++++++++++++++++++++++++++++++--------- 2 files changed, 219 insertions(+), 53 deletions(-) diff --git a/examples/congmath.v b/examples/congmath.v index 51cf72b..7375753 100644 --- a/examples/congmath.v +++ b/examples/congmath.v @@ -22,10 +22,15 @@ From pcm Require Import options prelude ordtype finmap pred seqext. (* Background mathematics for the verification of the *) (* Barcelogic Congruence Closure algorithm. *) -(* Described in the POPL10 paper *) +(* Described in the POPL10 paper: *) (* Structure the verification of heap-manipulating programs *) (* by Nanevski, Vafeiadis and Berdine *) +(* This file contains the proofs of the mathematical properties *) +(* needed for the verification. The implementation and the *) +(* proofs of the actual stateful programs is in the *) +(* accompanying file congprog.v *) + (**************************) (* Constants with arities *) (**************************) diff --git a/examples/congprog.v b/examples/congprog.v index f2b8e52..6a67161 100644 --- a/examples/congprog.v +++ b/examples/congprog.v @@ -1,7 +1,3 @@ -(*******************************) -(* Stateful congruence closure *) -(*******************************) - From HB Require Import structures. From mathcomp Require Import ssreflect ssrbool ssrnat eqtype ssrfun. From mathcomp Require Import div finset seq fintype finfun choice. @@ -10,33 +6,184 @@ From pcm Require Import unionmap heap autopcm automap. From htt Require Import options model heapauto llist array. From htt Require Import kvmaps hashtab congmath. +(***********************************) +(* Congruence closure verification *) +(***********************************) + +(* Spec, code and proofs for the verification of the *) +(* Barcelogic Congruence Closure algorithm. *) +(* Described in the POPL10 paper: *) +(* Structure the verification of heap-manipulating programs *) +(* by Nanevski, Vafeiadis and Berdine *) + +(* This file contains the verification of the stateful programs. *) +(* The associated math properties used in the verification are *) +(* proved in the accompanying file congmath.v *) + Notation finE := finset.inE. -Definition ith {I : finType} i (pf : i < #|I|) : I. Admitted. +Lemma inhabF (I : finType) : 0 < #|I| -> ~ @predT I =1 xpred0. +Proof. by case/card_gt0P=>x _ /(_ x). Qed. -Lemma indx_ith {I : finType} i (pf : i < #|I|) : indx (ith i pf) = i. -Proof. -admit. -Admitted. -(* by move=>i pf; rewrite /ith /= /indx index_uniq // -?cardE // enum_uniq. *) +Definition inhab0 {I : finType} (pf : 0 < #|I|) : I := + match pickP predT with + | Pick x _ => x + | Nopick qf => False_rect I (inhabF I pf qf) + end. -Lemma ith_indx {I : finType} (s : I) (pf : indx s < #|I|) : ith (indx s) pf = s. -Proof. Admitted. -(* by move=>s /= pf; rewrite /ith /= nth_index // mem_enum. Qed. *) +Definition inhab {I : ifinType} : I := inhab0 card_inhab. + +(* variant of nth that needs no seed value *) +Definition ith {I : finType} i (pf : i < #|I|) : I := + nth (inhab0 (leq_ltn_trans (leq0n i) pf)) (enum I) i. + +(* dually, variant of index that doesn't overflow *) +Definition indx {I : finType} (x : I) := index x (enum I). + +Lemma indx_ith {I : finType} i (pf : i < #|I|) : + indx (ith i pf) = i. +Proof. by rewrite /indx/ith index_uniq ?enum_uniq -?cardE. Qed. + +Lemma ith_indx {I : finType} (s : I) (pf : indx s < #|I|) : + ith (indx s) pf = s. +Proof. by rewrite /ith/indx nth_index // mem_enum. Qed. Lemma indx_inj I : injective (@indx I). -Admitted. +Proof. +rewrite /indx=>x1 x2. +have [] : x1 \in enum I /\ x2 \in enum I by rewrite !mem_enum. +elim: (enum I)=>[|x xs IH] //=; rewrite !inE !(eq_sym x). +case: (x1 =P x)=>[<-|] _ /=; first by case: (x2 =P x1). +by case: (x2 =P x)=>//= _ X1 X2 []; apply: IH X1 X2. +Qed. + +From mathcomp Require Import bigop. +From Coq Require Import Setoid. +Prenex Implicits star. + +Add Parametric Morphism U : (@star U) with signature + @EqPredType _ _ ==> @EqPredType _ _ ==> @EqPredType _ _ as star_morph. +Proof. +move=>x y E x1 y1 E1 m /=. +split; case=>h1 [h2][-> H1 H2]; exists h1, h2; split=>//. + by apply/E. by apply/E1. + by apply/E. by apply/E1. +Qed. + +Lemma sepit0' (U : pcm) A (f : A -> Pred U) : + IterStar.sepit [::] f <~> emp. +Proof. exact: IterStar.sepit0. Qed. + +Lemma sepit_cons' (U : pcm) A x (s : seq A) (f : A -> Pred U) : + IterStar.sepit (x :: s) f <~> + f x # IterStar.sepit s f. +Proof. exact: IterStar.sepit_cons. Qed. + +Lemma sepitX (U : pcm) A (s : seq A) (f : A -> Pred U) : + IterStar.sepit s f <~> + \big[star/emp]_(i <- s) f i. +Proof. +elim: s=>[|x s IH]. +- rewrite sepit0'. rewrite big_nil. by []. +rewrite sepit_cons'. +rewrite big_cons. +move=>m. +split; case=>h1 [h2][Em H1 H2]; exists h1, h2; split=>//. +- by rewrite -IH. +by rewrite IH. +Qed. + +Lemma sepit0 (U : pcm) A (f : A -> Pred U) : + IterStar.sepit [::] f <~> emp. +Proof. by rewrite sepitX big_nil. Qed. + +Lemma sepit_cons (U : pcm) A x (s : seq A) (f : A -> Pred U) : + IterStar.sepit (x::s) f <~> f x # IterStar.sepit s f. +Proof. by rewrite sepitX big_cons sepitX. Qed. -Lemma indx_injE {I : finType} s i (pf : i < #|I|) : (s == ith i pf) = (indx s == i). +(* U has the laws of commutative monoids from bigop *) +(* but doesn't really, as it's up to extensional equality *) +Lemma starA (U : pcm) (x y z : Pred U) : + x # y # z <~> (x # y) # z. Proof. -Admitted. +move=>m /=; split; case=>h1 [h2][H1 H2]. +- case=>h3 [h4][H3 H4 H5]. + subst h2. + eexists (h1 \+ h3), h4. rewrite -joinA. split=>//. + by exists h1, h3. +case: H2=>h3 [h4][? H2 H3 H4]. subst h1. +exists h3, (h4 \+ h2). rewrite joinA. split=>//. +by exists h4, h2. +Qed. +Lemma starC (U : pcm) (x y : Pred U) : + x # y <~> y # x. +Proof. +move=>m /=. +split. +- case=>h1 [h2][H H1 H2]. + exists h2, h1. rewrite joinC. by []. +case=>h1 [h2][H H1 H2]. +exists h2, h1. by rewrite joinC. +Qed. -Lemma sepit_emp (A : eqType) (s : seq A) f : - (forall x, x \in s -> f x =p emp (U:=heap)) -> - IterStar.sepit s f =p emp. +Lemma starL (U : pcm) (x : Pred U) : + emp # x <~> x. Proof. -Admitted. +move=>m /=. +split. +- case=>h1 [h2][->->]. rewrite unitL. by []. +move=>H. +exists Unit, m. rewrite unitL. by []. +Qed. + +Lemma starR (U : pcm) (x : Pred U) : + x # emp <~> x. +Proof. by rewrite starC starL. Qed. + +(* +HB.instance Definition _ (U : pcm) := + Monoid.isComLaw.Build (Pred U) emp (@star U) (@starA U) (@starC U) (@starL U). +*) + +Lemma big_sepit_seq (U : pcm) A (s : seq A) (f : A -> U) m : + m = \big[join/Unit]_(i <- s) f i <-> + m \In IterStar.sepit s (fun i h => h = f i). +Proof. +rewrite sepitX. +elim: s m=>[|x xs IH] /= m. +- by rewrite !big_nil. +rewrite !big_cons. +split. +- move=>E. rewrite InE. exists (f x), (\big[join/Unit]_(j <- xs) f j). + split=>//. rewrite -IH. by []. +case=>h1 [h2][Em H1 H2]. rewrite -toPredE /= in H1. +rewrite Em H1. +congr (_ \+ _). +by rewrite IH. +Qed. + +Lemma sepit_emp (U : pcm) (A : eqType) (s : seq A) (f : A -> Pred U) : + (forall x, x \in s -> f x <~> emp (U:=U)) -> + IterStar.sepit s f <~> emp. +Proof. +move=>H. +rewrite sepitX. +elim: s H=>[|a xs IH] H. +- by rewrite big_nil. +rewrite big_cons. +rewrite H; last by rewrite inE eqxx. +rewrite starL IH //. +move=>x X. apply: H. +by rewrite inE X orbT. +Qed. + +Lemma big_sepit (U : pcm) (I : finType) (s : {set I}) (f : I -> U) m : + m = \big[join/Unit]_(i in s) (f i) <-> + m \In sepit s (fun i h => h = f i). +Proof. by rewrite /sepit /= -big_sepit_seq -big_enum. Qed. + + (* empty congruence only relates constant symbols to themselves *) @@ -110,10 +257,24 @@ Definition utab := Definition n := #|{: symb}|. -(* the algorithm starts with root pointers for the data *) +(* roots for the structure *) +(* all components are isomorphic to pointers *) Inductive ptrs : Set := - Ptrs of {array symb -> symb} & {array symb -> llist symb} & - {array symb -> llist (symb*symb*symb)} & KVmap.tp LT & ptr. + Ptrs of + (* array of representatives; for each c *) + (* returns symbol that represents c's class *) + {array symb -> symb} & + (* class list; for each c *) + (* singly-linked list storing whole class of c *) + {array symb -> llist symb} & + (* use list; internal structure *) + (* see the paper for description *) + {array symb -> llist (symb*symb*symb)} & + (* hash table; for each pair of representatives *) + (* stores equations; see paper for description *) + KVmap.tp LT & + (* list of pending equations *) + ptr. (* renaming type to satisfy the signature check *) Definition tp := ptrs. @@ -126,7 +287,7 @@ Let ulist := let: Ptrs r clist ulist htab p := rt in ulist. Let htab := let: Ptrs r clist ulist htab p := rt in htab. Let p := let: Ptrs r clist ulist htab p := rt in p. -(* The layout of the data structure *) +(* Data structure's layout in the heap *) Definition ashape D := [Pred h | let: (d, ct, ut) := D in @@ -144,9 +305,10 @@ Definition shape (R : rel_exp s) : Pred heap := [Pred h | exists d, h \In bshape d /\ propagate_inv d /\ pending d = [::] /\ CRel d <~> R]. -End ShapePredicates. +End ShapePredicates. -(* Initialization procedure to generate the empty structure *) +(* Initialization procedure to generate *) +(* the empty structure and its root pointers *) Definition iT (clist : {array symb -> llist symb}): Type := forall k, STsep (fun i => k <= n /\ exists f, i \In Array.shape clist f # @@ -156,21 +318,21 @@ Definition iT (clist : {array symb -> llist symb}): Type := forall k, Program Definition init : STsep (emp, [vfun rt m => m \In shape rt (empty_cong s)]) := - Do (r <-- Array.newf [ffun x : symb => x]; - clist <-- Array.new _ null; - ffix (fun (loop : iT clist) k => + Do (rx <-- Array.newf [ffun x : symb => x]; + cl <-- Array.new _ null; + ffix (fun (loop : iT cl) k => Do (if decP (b:= k < n) idP is left pf then x <-- allocb (ith k pf) 2; x.+1 ::= null;; - Array.write clist (ith k pf) x;; + Array.write cl (ith k pf) x;; loop k.+1 else ret tt)) 0;; - ulist <-- Array.new _ null; - htab <-- KVmap.new LT; + ul <-- Array.new _ null; + htb <-- KVmap.new LT; p <-- alloc null; - ret (Ptrs r clist ulist htab p)). + ret (Ptrs rx cl ul htb p)). Next Obligation. -move=>r clist loop k H i [pf][/= f][hc][hct][->{i} Hc Hct]. +move=>_ cl loop k H i [pf][/= f][hc][hct][->{i} Hc Hct]. case: decP=>[{}pf|] /=; last first. - case: (ltngtP k n) pf=>// Ekn _ _; step=>_. exists f, hc, hct; split=>//. @@ -188,22 +350,25 @@ apply: tableP2 Hct=>// a. by rewrite !finE !ffunE indx_injE; case: eqP=>// ->; rewrite ltnn. Qed. Next Obligation. -case=>_ ->; apply: [stepE]=>//= r hr Er; apply: [stepU]=>//= clist hc Ec. +case=>_ ->; apply: [stepE]=>//= rx hr Er; apply: [stepU]=>//= cl hc Ec. apply: [stepX]@hc=>//=. - split=>//; exists [ffun x => null], hc, Unit; rewrite unitR. by split=>//; rewrite (_ : [set c | indx c < 0] = set0) // sepit0. case=>_ [f][hc'][hrest][-> Hc' Hrest]. -apply: [stepU]=>//= ulist hu Ehu; apply: [stepU]=>//= htab ht Ht. +apply: [stepU]=>//= ul hu Ehu; apply: [stepU]=>//= htb ht Ht. set d := Data [ffun x => x] [ffun c => [:: c]] [ffun c => [::]] (nil K V) [::]. -step=>p; step; exists d; split; last by case: (initP s). +step=>px; step; exists d; split; last by case: (initP s). split=>[a b|/=]; first by rewrite !ffunE !inE. exists f, [ffun s => null]. -rewrite (_ : p :-> null \+ _ = hr \+ ((hc' \+ hrest) \+ (hu \+ Unit \+ - (ht \+ (p :-> null \+ Unit))))); last by rewrite unitR; heap_congr. +rewrite (_ : px :-> null \+ _ = hr \+ ((hc' \+ hrest) \+ (hu \+ Unit \+ + (ht \+ (px :-> null \+ Unit))))); last by rewrite unitR; heap_congr. hhauto; rewrite sepit_emp //= => k. by rewrite /utab/table !ffunE; split=>//; case=>_ ->. Qed. +(* Method implementations use various input root pointers. *) +(* These are given explicit names as projections from rt. *) + Section Internal. Variable rt : ptrs. Notation ashape' := (ashape rt). @@ -395,17 +560,16 @@ Next Obligation. by move=>a' b' [/= d][i][N H]; apply: [gE d]=>[||??[]] //; exists [::]. Qed. -Let pend0 (e : pend s) := - match e with simp_pend a b => a | comp_pend (a,_,_) (b,_,_) => a end. -Let pend1 (e : pend s) := - match e with simp_pend a b => b | comp_pend (a,_,_) (b,_,_) => b end. -Notation "e ..0" := (pend0 e) (at level 2). -Notation "e ..1" := (pend1 e) (at level 2). - Definition pT : Type := forall x : unit, STsep {d} (fun i => i \In bshape' d, [vfun (_ : unit) m => m \In bshape' (propagate d)]). +(* left/right symbol in pending equation *) +Definition pendL (e : pend s) := + let: (simp_pend a _ | comp_pend (a,_,_) _) := e in a. +Definition pendR (e : pend s) := + let: (simp_pend _ b | comp_pend _ (b,_,_)) := e in b. + Program Definition hpropagate := ffix (fun (loop : pT) x => Do (q <-- @read ptr p; @@ -416,8 +580,8 @@ Program Definition hpropagate := p ::= (next : ptr);; dealloc q;; dealloc q.+1;; - a' <-- Array.read r eq..0; - b' <-- Array.read r eq..1; + a' <-- Array.read r (pendL eq); + b' <-- Array.read r (pendR eq); if a' == b' then loop tt : ST _ else join_hclass a' b';; @@ -566,9 +730,6 @@ hhauto; apply: tableP Hu2'=>/= a; rewrite !finE !ffunE /= andbT; by case/andP=>/negbTE -> /negbTE ->. Qed. -Let pend3 (e : symb*symb*symb) := let: (a, _, _) := e in a. -Notation "e ..0" := (pend3 e) (at level 2). - Definition nT : Type := forall t, STsep {d} (fun i => i \In bshape' d, [vfun y m => y = norm d t /\ m \In bshape' d]). @@ -585,7 +746,7 @@ Program Definition hnorm := if (u1, u2) is (const w1, const w2) then v <-- KVmap.lookup htab (w1, w2); if v is Some a then - a' <-- Array.read r (a..0); + a' <-- Array.read r a.1.1; ret (const a') else ret (app u1 u2) else ret (app u1 u2) From ab0b85c28ddac72502f13a059370458a37db5cd3 Mon Sep 17 00:00:00 2001 From: kevinlopez Date: Thu, 15 Aug 2024 17:21:57 +0200 Subject: [PATCH 52/93] defining shape predicate --- examples/schorr.v | 6 ++++-- 1 file changed, 4 insertions(+), 2 deletions(-) diff --git a/examples/schorr.v b/examples/schorr.v index 82d8c60..c985c2f 100644 --- a/examples/schorr.v +++ b/examples/schorr.v @@ -145,6 +145,8 @@ Definition shape g0 r p t := h = lay m g /\ path (ch m g) null s /\ rev_graph m g s t = g0 /\ - reach m g s t /\ + reach m g s t /\ + [pcm um_filterv (fun v => ( eq_mark v L || eq_mark v LR)) g <= dom S ] /\ p = last null s /\ - r = head t s). + r = head t s /\ + graph_axiom g). From 295c20da70c3216f4cc4460817736b879a9f91f0 Mon Sep 17 00:00:00 2001 From: kevinlopez Date: Thu, 15 Aug 2024 17:27:51 +0200 Subject: [PATCH 53/93] defining shape predicate --- Unnamed_coqscript_1.crashcoqide | 0 Unnamed_coqscript_2.crashcoqide | 2 + core/prelude.v.crashcoqide | 1162 +++++++++++++++++++++++++++++++ examples/schorr.v | 2 +- examples/schorr.v.crashcoqide | 142 ++++ htt/domain.v.crashcoqide | 1038 +++++++++++++++++++++++++++ 6 files changed, 2345 insertions(+), 1 deletion(-) create mode 100644 Unnamed_coqscript_1.crashcoqide create mode 100644 Unnamed_coqscript_2.crashcoqide create mode 100644 core/prelude.v.crashcoqide create mode 100644 examples/schorr.v.crashcoqide create mode 100644 htt/domain.v.crashcoqide diff --git a/Unnamed_coqscript_1.crashcoqide b/Unnamed_coqscript_1.crashcoqide new file mode 100644 index 0000000..e69de29 diff --git a/Unnamed_coqscript_2.crashcoqide b/Unnamed_coqscript_2.crashcoqide new file mode 100644 index 0000000..8f0fbb9 --- /dev/null +++ b/Unnamed_coqscript_2.crashcoqide @@ -0,0 +1,2 @@ +From HB Require Import structures. +From mathcomp Require Import ssreflect ssrfun ssrbool. diff --git a/core/prelude.v.crashcoqide b/core/prelude.v.crashcoqide new file mode 100644 index 0000000..f286e8f --- /dev/null +++ b/core/prelude.v.crashcoqide @@ -0,0 +1,1162 @@ +(* +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 HB Require Import structures. +From Coq Require Import ssreflect ssrfun Eqdep. +From mathcomp Require Import ssrbool ssrnat seq eqtype choice. +From mathcomp Require Import fintype finfun tuple perm fingroup. +From pcm Require Import options axioms. + +(***********) +(* Prelude *) +(***********) + +(* 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}. + +(* 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}. + +(* 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. + +Lemma trans_eq A (x y z : A) : + x = y -> + x = z -> + y = z. +Proof. by move/esym; apply: eq_trans. 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. + +(************) +(* Products *) +(************) + +Inductive Prod3 U1 U2 U3 := + mk3 {proj31 : U1; proj32 : U2; proj33 : U3}. +Inductive Prod4 U1 U2 U3 U4 := + mk4 {proj41 : U1; proj42 : U2; proj43 : U3; proj44 : U4}. +Inductive Prod5 U1 U2 U3 U4 U5 := + mk5 {proj51 : U1; proj52 : U2; proj53 : U3; proj54 : U4; proj55 : U5}. +Inductive Prod6 U1 U2 U3 U4 U5 U6 := + mk6 {proj61 : U1; proj62 : U2; proj63 : U3; + proj64 : U4; proj65 : U5; proj66 : U6}. +Inductive Prod7 U1 U2 U3 U4 U5 U6 U7 := + mk7 {proj71 : U1; proj72 : U2; proj73 : U3; + proj74 : U4; proj75 : U5; proj76 : U6; proj77 : U7}. +Prenex Implicits proj31 proj32 proj33. +Prenex Implicits proj41 proj42 proj43 proj44. +Prenex Implicits proj51 proj52 proj53 proj54 proj55. +Prenex Implicits proj61 proj62 proj63 proj64 proj65 proj66. +Prenex Implicits proj71 proj72 proj73 proj74 proj75 proj76 proj77. + +Definition eq3 (U1 U2 U3 : eqType) (x y : Prod3 U1 U2 U3) := + [&& proj31 x == proj31 y, proj32 x == proj32 y & proj33 x == proj33 y]. +Definition eq4 (U1 U2 U3 U4 : eqType) (x y : Prod4 U1 U2 U3 U4) := + [&& proj41 x == proj41 y, proj42 x == proj42 y, proj43 x == proj43 y & + proj44 x == proj44 y]. +Definition eq5 (U1 U2 U3 U4 U5 : eqType) (x y : Prod5 U1 U2 U3 U4 U5) := + [&& proj51 x == proj51 y, proj52 x == proj52 y, proj53 x == proj53 y, + proj54 x == proj54 y & proj55 x == proj55 y]. +Definition eq6 (U1 U2 U3 U4 U5 U6 : eqType) (x y : Prod6 U1 U2 U3 U4 U5 U6) := + [&& proj61 x == proj61 y, proj62 x == proj62 y, proj63 x == proj63 y, + proj64 x == proj64 y, proj65 x == proj65 y & proj66 x == proj66 y]. +Definition eq7 (U1 U2 U3 U4 U5 U6 U7 : eqType) (x y : Prod7 U1 U2 U3 U4 U5 U6 U7) := + [&& proj71 x == proj71 y, proj72 x == proj72 y, proj73 x == proj73 y, + proj74 x == proj74 y, proj75 x == proj75 y, proj76 x == proj76 y & + proj77 x == proj77 y]. + +Lemma eq3P U1 U2 U3 : + Equality.axiom (@eq3 U1 U2 U3). +Proof. +rewrite /eq3; case=>x1 x2 x3 [y1 y2 y3] /=. +by do ![case: eqP=>[->|]]; constructor=>//; case. +Qed. + +Lemma eq4P U1 U2 U3 U4 : + Equality.axiom (@eq4 U1 U2 U3 U4). +Proof. +rewrite /eq4; case=>x1 x2 x3 x4 [y1 y2 y3 y4] /=. +by do ![case: eqP=>[->|]]; constructor=>//; case. +Qed. + +Lemma eq5P U1 U2 U3 U4 U5 : + Equality.axiom (@eq5 U1 U2 U3 U4 U5). +Proof. +rewrite /eq5; case=>x1 x2 x3 x4 x5 [y1 y2 y3 y4 y5] /=. +by do ![case: eqP=>[->|]]; constructor=>//; case. +Qed. + +Lemma eq6P U1 U2 U3 U4 U5 U6 : + Equality.axiom (@eq6 U1 U2 U3 U4 U5 U6). +Proof. +rewrite /eq6; case=>x1 x2 x3 x4 x5 x6 [y1 y2 y3 y4 y5 y6] /=. +by do ![case: eqP=>[->|]]; constructor=>//; case. +Qed. + +Lemma eq7P U1 U2 U3 U4 U5 U6 U7 : + Equality.axiom (@eq7 U1 U2 U3 U4 U5 U6 U7). +Proof. +rewrite /eq7; case=>x1 x2 x3 x4 x5 x6 x7 [y1 y2 y3 y4 y5 y6 y7] /=. +by do ![case: eqP=>[->|]]; constructor=>//; case. +Qed. + +HB.instance Definition _ (U1 U2 U3 : eqType) := + hasDecEq.Build (Prod3 U1 U2 U3) (@eq3P U1 U2 U3). +HB.instance Definition _ (U1 U2 U3 U4 : eqType) := + hasDecEq.Build (Prod4 U1 U2 U3 U4) (@eq4P U1 U2 U3 U4). +HB.instance Definition _ (U1 U2 U3 U4 U5 : eqType) := + hasDecEq.Build (Prod5 U1 U2 U3 U4 U5) (@eq5P U1 U2 U3 U4 U5). +HB.instance Definition _ (U1 U2 U3 U4 U5 U6 : eqType) := + hasDecEq.Build (Prod6 U1 U2 U3 U4 U5 U6) (@eq6P U1 U2 U3 U4 U5 U6). +HB.instance Definition _ (U1 U2 U3 U4 U5 U6 U7 : eqType) := + hasDecEq.Build (Prod7 U1 U2 U3 U4 U5 U6 U7) (@eq7P U1 U2 U3 U4 U5 U6 U7). + +(***************************) +(* 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. + +Lemma compE A B C (g : B -> C) (f : A -> B) x : + g (f x) = (g \o f) x. +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, aka. fork/fanout/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. + +(************************) +(* 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 , P6 , P7 , P8 , P9 , P10 & P11 ]" (at level 0, format + "'[hv' [ /\ '[' P1 , '/' P2 , '/' P3 , '/' P4 , '/' P5 , '/' P6 , '/' P7 , '/' P8 , '/' P9 , '/' P10 ']' '/ ' & P11 ] ']'"). +Reserved Notation "[ /\ P1 , P2 , P3 , P4 , P5 , P6 , P7 , P8 , P9 , P10 , P11 & P12 ]" (at level 0, format + "'[hv' [ /\ '[' P1 , '/' P2 , '/' P3 , '/' P4 , '/' P5 , '/' P6 , '/' P7 , '/' P8 , '/' P9 , '/' P10 , '/' P11 ']' '/ ' & P12 ] ']'"). +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 and11 (P1 P2 P3 P4 P5 P6 P7 P8 P9 P10 P11 : Prop) : Prop := + And11 of P1 & P2 & P3 & P4 & P5 & P6 & P7 & P8 & P9 & P10 & P11. +Inductive and12 (P1 P2 P3 P4 P5 P6 P7 P8 P9 P10 P11 P12 : Prop) : Prop := + And12 of P1 & P2 & P3 & P4 & P5 & P6 & P7 & P8 & P9 & P10 & P11 & P12. + +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 , P6 , P7 , P8 , P9 , P10 & P11 ]" := + (and11 P1 P2 P3 P4 P5 P6 P7 P8 P9 P10 P11) : type_scope. +Notation "[ /\ P1 , P2 , P3 , P4 , P5 , P6 , P7 , P8 , P9 , P10 , P11 & P12 ]" := + (and12 P1 P2 P3 P4 P5 P6 P7 P8 P9 P10 P11 P12) : 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 *) + +(* DEVCOMMENT *) +(* TODO: we should move some of the following to [ssrbool] in Coq *) +(* /DEVCOMMENT *) + +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 *) +#[export] Program Instance and3_impl_morphism : Proper (iter_respectful impl impl 3) and3 | 1. +#[export] Program Instance and3_iff_morphism : Proper (iter_respectful iff iff 3) and3 | 1. + +#[export] Program Instance and4_impl_morphism : Proper (iter_respectful impl impl 4) and4 | 1. +#[export] Program Instance and4_iff_morphism : Proper (iter_respectful iff iff 4) and4 | 1. + +#[export] Program Instance and5_impl_morphism : Proper (iter_respectful impl impl 5) and5 | 1. +#[export] Program Instance and5_iff_morphism : Proper (iter_respectful iff iff 5) and5 | 1. + +#[export] Program Instance and6_impl_morphism : Proper (iter_respectful impl impl 6) and6 | 1. +#[export] Program Instance and6_iff_morphism : Proper (iter_respectful iff iff 6) and6 | 1. + +#[export] Program Instance and7_impl_morphism : Proper (iter_respectful impl impl 7) and7 | 1. +#[export] Program Instance and7_iff_morphism : Proper (iter_respectful iff iff 7) and7 | 1. + +#[export] Program Instance and8_impl_morphism : Proper (iter_respectful impl impl 8) and8 | 1. +#[export] Program Instance and8_iff_morphism : Proper (iter_respectful iff iff 8) and8 | 1. + +#[export] Program Instance and9_impl_morphism : Proper (iter_respectful impl impl 9) and9 | 1. +#[export] Program Instance and9_iff_morphism : Proper (iter_respectful iff iff 9) and9 | 1. + +#[export] Program Instance and10_impl_morphism : Proper (iter_respectful impl impl 10) and10 | 1. +#[export] Program Instance and10_iff_morphism : Proper (iter_respectful iff iff 10) and10 | 1. + +#[export] Program Instance and11_impl_morphism : Proper (iter_respectful impl impl 11) and11 | 1. +#[export] Program Instance and11_iff_morphism : Proper (iter_respectful iff iff 11) and11 | 1. + +#[export] Program Instance and12_impl_morphism : Proper (iter_respectful impl impl 12) and12 | 1. +#[export] Program Instance and12_iff_morphism : Proper (iter_respectful iff iff 12) and12 | 1. + +(** Logical disjunction *) +#[export] Program Instance or3_impl_morphism : Proper (iter_respectful impl impl 3) or3 | 1. +#[export] Program Instance or3_iff_morphism : Proper (iter_respectful iff iff 3) or3 | 1. + +#[export] Program Instance or4_impl_morphism : Proper (iter_respectful impl impl 4) or4 | 1. +#[export] Program Instance or4_iff_morphism : Proper (iter_respectful iff iff 4) or4 | 1. + +#[export] Program Instance or5_impl_morphism : Proper (iter_respectful impl impl 5) or5 | 1. +#[export] Program Instance or5_iff_morphism : Proper (iter_respectful iff iff 5) or5 | 1. + +#[export] Program Instance or6_impl_morphism : Proper (iter_respectful impl impl 6) or6 | 1. +#[export] Program Instance or6_iff_morphism : Proper (iter_respectful iff iff 6) or6 | 1. + +#[export] Program Instance or7_impl_morphism : Proper (iter_respectful impl impl 7) or7 | 1. +#[export] 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 b11 b12 : bool. + +Lemma and6P : reflect [/\ b1, b2, b3, b4, b5 & b6] [&& b1, b2, b3, b4, b5 & b6]. +Proof. +case: b1=>/=; last by constructor; case. +case: b2=>/=; last by constructor; case. +case: b3=>/=; last by constructor; case. +case: b4=>/=; last by constructor; case. +case: b5=>/=; last by constructor; case. +case: b6=>/=; last by constructor; case. +by constructor. +Qed. + +Lemma and7P : reflect [/\ b1, b2, b3, b4, b5, b6 & b7] [&& b1, b2, b3, b4, b5, b6 & b7]. +Proof. +case: b1=>/=; last by constructor; case. +case: b2=>/=; last by constructor; case. +case: b3=>/=; last by constructor; case. +case: b4=>/=; last by constructor; case. +case: b5=>/=; last by constructor; case. +case: b6=>/=; last by constructor; case. +case: b7=>/=; last by constructor; case. +by constructor. +Qed. + +Lemma and8P : reflect [/\ b1, b2, b3, b4, b5, b6, b7 & b8] [&& b1, b2, b3, b4, b5, b6, b7 & b8]. +Proof. +case: b1=>/=; last by constructor; case. +case: b2=>/=; last by constructor; case. +case: b3=>/=; last by constructor; case. +case: b4=>/=; last by constructor; case. +case: b5=>/=; last by constructor; case. +case: b6=>/=; last by constructor; case. +case: b7=>/=; last by constructor; case. +case: b8=>/=; last by constructor; case. +by constructor. +Qed. + +Lemma and9P : reflect [/\ b1, b2, b3, b4, b5, b6, b7, b8 & b9] [&& b1, b2, b3, b4, b5, b6, b7, b8 & b9]. +Proof. +case: b1=>/=; last by constructor; case. +case: b2=>/=; last by constructor; case. +case: b3=>/=; last by constructor; case. +case: b4=>/=; last by constructor; case. +case: b5=>/=; last by constructor; case. +case: b6=>/=; last by constructor; case. +case: b7=>/=; last by constructor; case. +case: b8=>/=; last by constructor; case. +case: b9=>/=; last by constructor; case. +by constructor. +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. +case: b1=>/=; last by constructor; case. +case: b2=>/=; last by constructor; case. +case: b3=>/=; last by constructor; case. +case: b4=>/=; last by constructor; case. +case: b5=>/=; last by constructor; case. +case: b6=>/=; last by constructor; case. +case: b7=>/=; last by constructor; case. +case: b8=>/=; last by constructor; case. +case: b9=>/=; last by constructor; case. +case: b10=>/=; last by constructor; case. +by constructor. +Qed. + +Lemma and11P : reflect [/\ b1, b2, b3, b4, b5, b6, b7, b8, b9, b10 & b11] + [&& b1, b2, b3, b4, b5, b6, b7, b8, b9, b10 & b11]. +Proof. +case: b1=>/=; last by constructor; case. +case: b2=>/=; last by constructor; case. +case: b3=>/=; last by constructor; case. +case: b4=>/=; last by constructor; case. +case: b5=>/=; last by constructor; case. +case: b6=>/=; last by constructor; case. +case: b7=>/=; last by constructor; case. +case: b8=>/=; last by constructor; case. +case: b9=>/=; last by constructor; case. +case: b10=>/=; last by constructor; case. +case: b11=>/=; last by constructor; case. +by constructor. +Qed. + +Lemma and12P : reflect [/\ b1, b2, b3, b4, b5, b6, b7, b8, b9, b10, b11 & b12] + [&& b1, b2, b3, b4, b5, b6, b7, b8, b9, b10, b11 & b12]. +Proof. +case: b1=>/=; last by constructor; case. +case: b2=>/=; last by constructor; case. +case: b3=>/=; last by constructor; case. +case: b4=>/=; last by constructor; case. +case: b5=>/=; last by constructor; case. +case: b6=>/=; last by constructor; case. +case: b7=>/=; last by constructor; case. +case: b8=>/=; last by constructor; case. +case: b9=>/=; last by constructor; case. +case: b10=>/=; last by constructor; case. +case: b11=>/=; last by constructor; case. +case: b12=>/=; last by constructor; case. +by constructor. +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 and11P {b1 b2 b3 b4 b5 b6 b7 b8 b9 b10 b11}. +Arguments and12P {b1 b2 b3 b4 b5 b6 b7 b8 b9 b10 b11 b12}. + +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}. + +Lemma iffPb (b1 b2 : bool) : reflect (b1 <-> b2) (b1 == b2). +Proof. +case: eqP=>[->|N]; constructor=>//. +case: b1 b2 N; case=>//= _. +- by case=>/(_ erefl). +by case=>_ /(_ erefl). +Qed. + +Lemma implyb_trans a b c : + a ==> b -> + b ==> c -> + a ==> c. +Proof. by case: a=>//=->. Qed. + +Lemma iffE (b1 b2 : bool) : b1 = b2 <-> (b1 <-> b2). +Proof. by split=>[->|] //; move/iffPb/eqP. Qed. + +(* subsets *) + +Lemma subsetC T (p q : mem_pred T) : + {subset p <= q} -> + {subset predC q <= predC p}. +Proof. by move=>H x; apply: contra (H x). Qed. + +Lemma subsetR T (p q : mem_pred T) : + {subset p <= predC q} -> + {subset q <= predC p}. +Proof. by move=>H x; apply: contraL (H x). Qed. + +Lemma subsetL T (p q : mem_pred T) : + {subset predC p <= q} -> + {subset predC q <= p}. +Proof. by move=>H x; apply: contraR (H x). Qed. + +Lemma subsetLR T (p q : mem_pred T) : + {subset predC p <= predC q} -> + {subset q <= p}. +Proof. by move=>H x; apply: contraLR (H x). Qed. + +Lemma subset_disj T (p q r : mem_pred T) : + {subset p <= q} -> + (forall x, x \in q -> x \in r -> false) -> + (forall x, x \in p -> x \in r -> false). +Proof. by move=>H D x /H/D. Qed. + +Lemma subset_disj2 T (p p1 q q1 : mem_pred T) : + {subset p1 <= p} -> + {subset q1 <= q} -> + (forall x, x \in p -> x \in q -> false) -> + (forall x, x \in p1 -> x \in q1 -> false). +Proof. +move=>H1 H2 D1; apply: subset_disj H1 _ => x H /H2. +by apply: D1 H. +Qed. + +(**************) +(* empty type *) +(**************) + +Lemma False_eqP : Equality.axiom (fun _ _ : False => true). +Proof. by case. Qed. + +HB.instance Definition _ := hasDecEq.Build False False_eqP. + +(*************) +(* 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. + +#[export] Hint Extern 0 (ocancel _ _) => + (apply: lft_inl_ocanc || apply: rgt_inr_ocanc) : core. + +(****************) +(* option types *) +(****************) + +(* Alternative mechanism for manipulating *) +(* properties of the form isSome X. *) +(* To use a lemma of the form isSome X *) +(* one tyoically has to explicitly put the conclusion *) +(* of that lemma onto the context, and then case on X. *) +(* If such lemma is restated using is_some X *) +(* then it suffices to case on the lemma's name. *) +(* This saves typing X explicitly, which *) +(* may be significant if X is large. *) + +Inductive is_some_spec A x : option A -> Prop := +| is_some_case v of x = Some v : is_some_spec x (Some v). + +Hint Resolve is_some_case : core. + +Notation is_some x := (is_some_spec x x). + +Lemma is_someP A (x : option A) : reflect (is_some x) (isSome x). +Proof. by case: x=>[a|]; constructor=>//; case. Qed. + +(* some simplifications *) +Lemma oapp_some A (x : option A) : oapp [eta Some] None x = x. +Proof. by case: x. Qed. + +Lemma obind_some A (x : option A) : obind [eta Some] x = x. +Proof. exact: oapp_some. Qed. + +(********) +(* nats *) +(********) + +Lemma gt0 m n : m < n -> 0 < n. +Proof. by case: n. Qed. + +Lemma neq0 m n : m < n -> n != 0. +Proof. by move/gt0; rewrite lt0n. Qed. + +Lemma neqSn n : n.+1 != n. +Proof. by elim: n. Qed. + +Lemma neqnS n : n != n.+1. +Proof. by elim: n. Qed. + +(**************************************) +(* Inhabited (non-empty) finite types *) +(**************************************) + +Lemma inhabits (T : finType) (t : T) : 0 < #|T|. +Proof. by apply/card_gt0P; exists t. Qed. + +Lemma inhabits_irr (T : finType) (t1 t2 : T) : + inhabits t1 = inhabits t2. +Proof. by apply: bool_irrelevance. Qed. + +Definition inhabited_axiom (T : finType) := 0 < #|T|. + +HB.mixin Record isInhabited T of Finite T := { + card_inhab : inhabited_axiom T}. + +#[short(type="ifinType")] +HB.structure Definition Inhabited := {T of isInhabited T & Finite T}. + +Lemma inhabF (T : ifinType) : ~ (@predT T =1 xpred0). +Proof. by case/card_gt0P: (@card_inhab T)=>x _ /(_ x). Qed. + +Definition inhab {T : ifinType} : T := + match pickP predT with + | Pick x _ => x + | Nopick pf => False_rect T (inhabF pf) + end. + +HB.instance Definition _ := isInhabited.Build unit (inhabits tt). +HB.instance Definition _ := isInhabited.Build bool (inhabits false). +HB.instance Definition _ n := isInhabited.Build 'I_n.+1 (inhabits ord0). +HB.instance Definition _ (T : finType) := + isInhabited.Build (option T) (inhabits None). + +(*************************************) +(* A copy of booleans with mnemonics *) +(* LL and RR for working with sides *) +(*************************************) + +Inductive Side := LL | RR. +Definition Side_eq x y := + match x, y with LL, LL => true | RR, RR => true | _, _ => false end. +Definition nat2Side x := if odd x then LL else RR. +Definition Side2nat x := if x is RR then 0 else 1. + +Lemma ssrcanc : ssrfun.cancel Side2nat nat2Side. Proof. by case. Qed. +HB.instance Definition _ : isCountable Side := CanIsCountable ssrcanc. + +Lemma Side_enumP : Finite.axiom [:: LL; RR]. Proof. by case. Qed. +HB.instance Definition _ : isFinite Side := isFinite.Build Side Side_enumP. +HB.instance Definition _ := isInhabited.Build Side (inhabits LL). + +(*************) +(* Sequences *) +(*************) + +(* folds *) +(* TODO upstream to mathcomp *) + +Lemma map_foldr {T1 T2} (f : T1 -> T2) xs : + map f xs = foldr (fun x ys => f x :: ys) [::] xs. +Proof. by []. Qed. + +Lemma fusion_foldr {T R Q} (g : R -> Q) f0 f1 z0 z1 (xs : seq T) : + (forall x y, g (f0 x y) = f1 x (g y)) -> + g z0 = z1 -> + g (foldr f0 z0 xs) = foldr f1 z1 xs. +Proof. by move=>Hf Hz; elim: xs=>//= x xs <-. Qed. + +Lemma fusion_foldl {T R Q} (g : R -> Q) f0 f1 z0 z1 (xs : seq T) : + (forall x y, g (f0 x y) = f1 (g x) y) -> + g z0 = z1 -> + g (foldl f0 z0 xs) = foldl f1 z1 xs. +Proof. +move=>Hf Hz; elim: xs z0 z1 Hz =>//= x xs IH z0 z1 Hz. +by apply: IH; rewrite Hf Hz. +Qed. + +Lemma foldl_foldr {T R} (f : R -> T -> R) z xs : + foldl f z xs = foldr (fun b g x => g (f x b)) id xs z. +Proof. by elim: xs z=>/=. Qed. + +Lemma foldr_foldl {T R} (f : T -> R -> R) z xs : + foldr f z xs = foldl (fun g b x => g (f b x)) id xs z. +Proof. +elim/last_ind: xs z=>//= xs x IH z. +by rewrite foldl_rcons -IH foldr_rcons. +Qed. + +(* pmap *) +(* TODO upstream to mathcomp *) +Lemma pmap_pcomp {S T U} (f : T -> option U) (g : S -> option T) s : + pmap (pcomp f g) s = pmap f (pmap g s). +Proof. by elim: s=>//= x s ->; rewrite /pcomp; case: (g x). Qed. + +(* sequence prefixes *) + +(* Two helper concepts for searching in sequences: *) +(* - onth: like nth, but returns None when the element is not found *) +(* - Prefix: a prefix relation on sequences, used for growing *) +(* interpretation contexts *) + +Fixpoint onth A (s : seq A) n : option A := + if s is x::sx then if n is nx.+1 then onth sx nx else Some x else None. + +Definition Prefix A (s1 s2 : seq A) : Prop := + forall n x, onth s1 n = some x -> onth s2 n = some x. + +(* Lemmas *) + +Section SeqPrefix. +Variable A : Type. +Implicit Type s : seq A. + +Variant onth_spec s n : bool -> Type := +| onth_none : onth s n = None -> onth_spec s n false +| onth_some v : onth s n = Some v -> onth_spec s n true. + +Lemma onth_sizeP s n : onth_spec s n (n < size s). +Proof. +elim: s n=>/= [|a s IH] n; first by rewrite ltn0; constructor. +case: n=>[|n] /=; first by apply: (@onth_some _ _ a). +rewrite ltnS; case: (IH n)=>[|v] H. +- by constructor. +by apply: (@onth_some _ _ v). +Qed. + +Lemma size_onth s n : n < size s -> exists x, onth s n = Some x. +Proof. +by case: onth_sizeP=>// v H _; exists v. +Qed. + +Lemma onth_size s n x : onth s n = Some x -> n < size s. +Proof. +by case: onth_sizeP=>//->. +Qed. + +Lemma size_onthPn s n : reflect (onth s n = None) (size s <= n). +Proof. +by rewrite leqNgt; apply: (iffP idP); case: onth_sizeP=>//= v ->. +Qed. + +Lemma onth_sizeN s : onth s (size s) = None. +Proof. by apply/size_onthPn. Qed. + +Lemma nth_onth x0 n s : nth x0 s n = odflt x0 (onth s n). +Proof. +elim: s n=>/= [|a s IH] n /=; first by apply: nth_nil. +by case: n. +Qed. + +Lemma onth_nth x0 n s : + n < size s -> + onth s n = Some (nth x0 s n). +Proof. +elim: s n=>//= a s IH n. +by rewrite ltnS; case: n. +Qed. + +Lemma Prefix_refl s : Prefix s s. +Proof. by move=>n x <-. Qed. + +Lemma Prefix_trans s2 s1 s3 : + Prefix s1 s2 -> Prefix s2 s3 -> Prefix s1 s3. +Proof. by move=>H1 H2 n x E; apply: H2; apply: H1. Qed. + +Lemma Prefix_cons x s1 s2 : Prefix (x :: s1) (x :: s2) <-> Prefix s1 s2. +Proof. by split=>E n; [apply: (E n.+1) | case: n]. Qed. + +Lemma Prefix_cons' x y s1 s2 : + Prefix (x :: s1) (y :: s2) -> x = y /\ Prefix s1 s2. +Proof. by move=>H; case: (H 0 x (erefl _)) (H)=>-> /Prefix_cons. Qed. + +Lemma Prefix_rcons x s : Prefix s (rcons s x). +Proof. by elim: s=>//= y ys IH; apply/Prefix_cons; apply: IH. Qed. + +Lemma Prefix_cat s1 s2 : Prefix s1 (s1 ++ s2). +Proof. +elim: s2 s1=>[|x xs IH] s1; first by rewrite cats0. +rewrite -cat_rcons; apply: Prefix_trans (IH _). +by apply: Prefix_rcons. +Qed. + +Lemma Prefix_size s1 s2 : Prefix s1 s2 -> size s1 <= size s2. +Proof. +elim: s1 s2=>[//|a s1 IH] [|b s2] H; first by move: (H 0 a (erefl _)). +by rewrite ltnS; apply: (IH _ (proj2 (Prefix_cons' H))). +Qed. + +Lemma Prefix_onth s t x : + x < size s -> + Prefix s t -> onth s x = onth t x. +Proof. +elim:s t x =>[//|a s IH] [|b t] x H1 H2; first by move: (H2 0 a (erefl _)). +by case/Prefix_cons': H2=><- H2; case: x H1=>[|n] //= H1; apply: IH. +Qed. + +Lemma PrefixE s1 s2 : Prefix s1 s2 <-> exists s3, s2 = s1 ++ s3. +Proof. +split; last by case=>s3 ->; apply: Prefix_cat. +elim: s1 s2=>[|x xs IH] s2; first by exists s2. +case: s2=>[/(_ 0 x erefl)//|y ys /Prefix_cons' [?]]. +by subst y=>/IH [s3 ->]; exists s3. +Qed. + +End SeqPrefix. + +#[export] Hint Resolve Prefix_refl : core. + +(* when A : eqType *) + +Section SeqPrefixEq. +Variable A : eqType. +Implicit Type s : seq A. + +Lemma onth_mem s n x : + onth s n = Some x -> + x \in s. +Proof. +by elim: s n=>//= a s IH [[->]|n /IH]; rewrite inE ?eq_refl // orbC =>->. +Qed. + +Lemma onth_index (s : seq A) x : + onth s (index x s) = + if x \in s then Some x else None. +Proof. +by elim: s=>//=h s IH; rewrite inE eq_sym; case: eqP=>//= ->. +Qed. + +Lemma PrefixP (s1 s2 : seq A) : + reflect (Prefix s1 s2) (prefix s1 s2). +Proof. +apply/(equivP (prefixP (s1:=s1) (s2:=s2))). +by apply: iff_sym; exact: PrefixE. +Qed. + +End SeqPrefixEq. + +(******************************) +(* Some commuting conversions *) +(******************************) + +Lemma fun_op A B C (b : option A) (vS : A -> B) (vN : B) (f : B -> C) : + f (if b is Some v then vS v else vN) = + if b is Some v then f (vS v) else f vN. +Proof. by case: b=>//. Qed. + +Lemma op_if A B (b : bool) (vS vN : option A) (vS1 : A -> B) (vN1 : B) : + (if (if b then vS else vN) is Some v then vS1 v else vN1) = + if b then if vS is Some v then vS1 v else vN1 + else if vN is Some v then vS1 v else vN1. +Proof. by case: b. Qed. + +Lemma if_triv b : (if b then true else false) = b. +Proof. by case: b. Qed. + +(*************************************************************) +(* quick ways to extract projections and transitivity proofs *) +(* out of iterated inequalities *) +(*************************************************************) + +Lemma ltn13 a b c : a < b < c -> a < c. +Proof. by case/andP; apply: ltn_trans. Qed. + +Lemma ltn12 a b c : a < b < c -> a < b. +Proof. by case/andP. Qed. + +Lemma ltn23 a b c : a < b < c -> b < c. +Proof. by case/andP. Qed. + +Lemma leq13 a b c : a <= b <= c -> a <= c. +Proof. by case/andP; apply: leq_trans. Qed. + +Lemma leq12 a b c : a <= b <= c -> a <= b. +Proof. by case/andP. Qed. + +Lemma leq23 a b c : a <= b <= c -> b <= c. +Proof. by case/andP. Qed. + +Lemma lqt13 a b c : a <= b < c -> a < c. +Proof. by case/andP; apply: leq_ltn_trans. Qed. + +Lemma lqt12 a b c : a <= b < c -> a <= b. +Proof. by case/andP. Qed. + +Lemma lqt23 a b c : a <= b < c -> b < c. +Proof. by case/andP. Qed. + +(********************) +(* Finite functions *) +(********************) + +Section FinFun. +Variables (T : finType) (Us : T -> Type). +Implicit Type f : {dffun forall t, Us t}. + +(* Explicit name for finite function application. *) +(* Will be used to hang canonical projections onto. *) +(* The function/argument order is reversed to facilitate rewriting. *) +Definition sel tg f : Us tg := f tg. + +Lemma ffinP f1 f2 : (forall t, sel t f1 = sel t f2) <-> f1 = f2. +Proof. by rewrite ffunP. Qed. + +(* beta equality *) +Lemma sel_fin t (f : forall t, Us t) : sel t (finfun f) = f t. +Proof. by rewrite /sel ffunE. Qed. + +(* eta equality *) +Lemma fin_eta f : f = finfun (sel^~ f). +Proof. by apply/ffinP=>t; rewrite sel_fin. Qed. + +(* function *) +Definition splice tg f (v : Us tg) : {dffun _} := + finfun (fun x => + if decP (x =P tg) is left pf then cast Us pf v + else sel x f). + +Lemma sel_splice t f x (v : Us x) : + sel t (splice f v) = + if decP (t =P x) is left pf then cast Us pf v + else sel t f. +Proof. by rewrite sel_fin. Qed. + +Lemma sel_spliceE t f v : sel t (splice f v) = v. +Proof. by rewrite sel_fin; case: eqP=>//= pf; rewrite eqd. Qed. + +Lemma sel_spliceN t x f (w : Us x) : + t <> x -> sel t (splice f w) = sel t f. +Proof. by move=>N; rewrite sel_fin; case: eqP. Qed. + +Lemma splice_eta t f : splice f (sel t f) = f. +Proof. +apply/ffinP=>x; rewrite sel_fin; case: eqP=>//=. +by move/[dup]=>-> pf; rewrite eqd. +Qed. + +End FinFun. + +Arguments sel {T Us} tg f. +Arguments splice {T Us tg} f v. + +(* notation for building finfuns *) +(* DEVCOMMENT *) +(* copied from finfun to fix some bad spacing in formatting *) +(* /DEVCOMMENT *) + +Notation "[ 'ffun' x : aT => E ]" := (finfun (fun x : aT => E)) + (at level 0, x name, format "[ 'ffun' x : aT => E ]") : fun_scope. + +Notation "[ 'ffun' x => E ]" := (@finfun _ (fun=> _) (fun x => E)) + (at level 0, x name, format "[ 'ffun' x => E ]") : fun_scope. + +Notation "['ffun' => E ]" := [ffun _ => E] + (at level 0, format "['ffun' => E ]") : fun_scope. + +Section IteratedNotation. +Variables (T : finType) (Us : T -> Type). + +Variant dfun_delta : Type := DFunDelta t of Us t. + +(* for iteration that starts with function ends with function *) +Definition dapp_fdelta df (f : forall t, Us t) z := + let: DFunDelta t v := df in + if decP (z =P t) is left pf then cast Us pf v + else f z. + +(* for iteration that starts with finfun ends with function *) +Definition splice' df (f : {ffun forall t, Us t}) z := + dapp_fdelta df f z. + +End IteratedNotation. + +Delimit Scope fun_delta_scope with FUN_DELTA. +Arguments dapp_fdelta {T Us} df%_FUN_DELTA f. + +Notation "y \\ x" := (@DFunDelta _ _ x y) (at level 1). + +(* notation for simultaneous update of f with d1,..,dn *) +(* rewrite by sel_fin peels all layers *) +Notation "[ 'ext' f 'with' d1 , .. , dn ]" := + (finfun ( + dapp_fdelta d1%_FUN_DELTA .. (dapp_fdelta dn%_FUN_DELTA f) ..)) + (at level 0, format + "'[hv' [ '[' 'ext' '/ ' f ']' '/' 'with' '[' d1 , '/' .. , '/' dn ']' ] ']'" + ) : fun_scope. + +(* notation for iterated update of f with d1, then d2, ... *) +(* rewrite by sel_fin peels top layer only *) +Notation "[ 'splice' F 'with' d1 , .. , dn ]" := + (finfun (splice' + d1%_FUN_DELTA .. (finfun (splice' dn%_FUN_DELTA (finfun F))) ..)) + (at level 0, format + "'[hv' [ '[' 'splice' '/ ' F ']' '/' 'with' '[' d1 , '/' .. , '/' dn ']' ] ']'" + ) : fun_scope. + +Section TestingNotation. +Variables (T : finType) (Us : T -> Type). +Variables (f : {dffun forall t, Us t}) (t1 t2 : T) (v1 : Us t1) (v2 : Us t2). + +(* we have three different options in displaying splices *) +(* splice, [splice], and [ext] *) +Lemma test : + [/\ sel t2 [splice f with v1 \\ t1, v2 \\ t2] = v2, + sel t2 (splice (splice f v1) v2) = v2, + sel t2 [ext f with v1 \\ t1, v2 \\ t2] = v2 & +(* and we can use underscores to elide some info *) + sel t2 [ext f with v1 \\ _, v2 \\ _] = v2]. +Abort. +End TestingNotation. + +(* mapping over simply-typed finite functions *) + +Section FinFunMap. +Variables (T : finType) (A B : Type). +Implicit Types (f : A -> B) (x : {ffun T -> A}). + +Definition fmap f x := [ffun tg => f (sel tg x)]. + +Lemma sel_fmap f x tg : sel tg (fmap f x) = f (sel tg x). +Proof. exact: sel_fin. Qed. + +Lemma fmap_splice f x tg (v : A) : + fmap f (splice (tg:=tg) x v) = splice (tg:=tg) (fmap f x) (f v). +Proof. +apply/ffinP=>t; rewrite !sel_fin; case: eqP=>//= ?; subst t. +by rewrite !eqc. +Qed. + +End FinFunMap. + +(* surgery on tuples and finfuns *) + +Section OnthCodom. +Variable A : Type. + +Lemma onth_tnth {n} (s : n.-tuple A) (i : 'I_n) : + onth s i = Some (tnth s i). +Proof. +elim: n s i =>[|n IH] s i; first by case: i. +case/tupleP: s=>/=x s; case: (unliftP ord0 i)=>[j|]-> /=. +- by rewrite tnthS. +by rewrite tnth0. +Qed. + +Lemma onth_codom {n} (i : 'I_n) (f: {ffun 'I_n -> A}) : + onth (fgraph f) i = Some (f i). +Proof. +pose i' := cast_ord (esym (card_ord n)) i. +move: (@tnth_fgraph _ _ f i'); rewrite (enum_val_ord) {2}/i' cast_ordKV=><-. +by rewrite (onth_tnth (fgraph f) i'). +Qed. + +End OnthCodom. + +(* ffun and permutation *) +Section PermFfun. +Variables (I : finType) (A : Type). + +Definition pffun (p : {perm I}) (f : {ffun I -> A}) := + [ffun i => f (p i)]. + +Lemma pffunE1 (f : {ffun I -> A}) : pffun 1%g f = f. +Proof. by apply/ffunP=>i; rewrite !ffunE permE. Qed. + +Lemma pffunEM (p p' : {perm I}) (f : {ffun I -> A}) : + pffun (p * p') f = pffun p (pffun p' f). +Proof. by apply/ffunP => i; rewrite !ffunE permM. Qed. + +End PermFfun. + + +(* Tagging *) + +Notation Tag := (@existT _ _). + +Lemma Tag_inj T Us (t1 t2 : T) i1 i2 : + Tag t1 i1 = Tag t2 i2 -> + t1 = t2 /\ jmeq Us i1 i2. +Proof. by case=>?; subst t2=>/inj_pair2 ->. Qed. +Arguments Tag_inj {T Us t1 t2 i1 i2}. + +(* tagged union of equality types is equality type *) + +Section TaggedEq. +Variables (T : eqType) (Us : T -> eqType). + +Definition tag_eq : sigT Us -> sigT Us -> bool := + fun '(Tag tx opx) '(Tag ty opy) => + if decP (tx =P ty) is left pf then opx == cast Us pf opy + else false. + +Lemma tag_eqP : Equality.axiom tag_eq. +Proof. +case=>tx opx [ty opy] /=; case: (tx =P ty)=>pf; last first. +- by constructor; case=>/pf. +subst ty; rewrite /= eqc; case: eqP=>pf; constructor; +by [rewrite pf|case=>/inj_pair2/pf]. +Qed. + +HB.instance Definition _ := hasDecEq.Build (sigT Us) tag_eqP. +End TaggedEq. + diff --git a/examples/schorr.v b/examples/schorr.v index c985c2f..c76afa1 100644 --- a/examples/schorr.v +++ b/examples/schorr.v @@ -146,7 +146,7 @@ Definition shape g0 r p t := path (ch m g) null s /\ rev_graph m g s t = g0 /\ reach m g s t /\ - [pcm um_filterv (fun v => ( eq_mark v L || eq_mark v LR)) g <= dom S ] /\ + [pcm um_filterv (fun v => ( eq_mark v L || eq_mark v LR)) m <= s ] /\ p = last null s /\ r = head t s /\ graph_axiom g). diff --git a/examples/schorr.v.crashcoqide b/examples/schorr.v.crashcoqide new file mode 100644 index 0000000..3207c6f --- /dev/null +++ b/examples/schorr.v.crashcoqide @@ -0,0 +1,142 @@ +(* +Copyright 2024 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 HB Require Import structures. +From Coq Require Import ssreflect ssrfun. +From mathcomp Require Import eqtype ssrnat ssrbool seq path bigop. +From htt Require Import options model heapauto graph. +From pcm Require Import options axioms pred prelude seqperm seqext. +From pcm Require Import pcm unionmap natmap heap autopcm automap. + + +(*****************) +(* helper lemmas *) +(*****************) + +Lemma eq_connect_aux (g1 g2 : pregraph) p: + valid (g1 \+ g2) -> + {subset dom g2 <= predC p} -> + um_filterk p (g1 \+ g2) = um_filterk p g1. +Proof. +move=>V /(umfiltk_subdom0 p (validR V)) S. +by rewrite umfiltUn // S unitR. +Qed. + +Lemma connectUn_eq (g g1 g2 : pregraph) (p : pred node) x : + valid (g \+ g1) -> + valid (g \+ g2) -> + {subset dom g1 <= p} -> + {subset dom g2 <= p} -> + connect p (g \+ g1) x =i connect p (g \+ g2) x. +Proof. +move=>V1 V2 S1 S2 z. +rewrite -connect_umfiltk eq_connect_aux //=; last first. +- by move=>y /S1; rewrite inE negbK. +rewrite -[RHS]connect_umfiltk eq_connect_aux //. +by move=>y /S2; rewrite !inE negbK. +Qed. + +(****************) +(* Schorr-Waite *) +(****************) + +(* short notation for left/right child of x *) +Notation gl g x := (get_nth g x 0). +Notation gr g x := (get_nth g x 1). + +(* type of markings *) +(* Aleks NOTE: Should we need U? My impression was that *) +(* being unmarked is represented by not being in *) +(* the appropriate map? *) +Inductive mark := U | L | R | LR. +(* decidable equality on marks *) +Definition eq_mark l1 l2 : bool := + if (l1, l2) is ((U, U)|(L, L)|(R, R)|(LR, LR)) then true else false. +Lemma eq_markP : Equality.axiom eq_mark. +Proof. by case; case=>//=; constructor. Qed. +HB.instance Definition _ := hasDecEq.Build mark eq_markP. + +(* marking of children *) + +(* given marking map m, x is m-child of y iff: *) +(* - m y = L and x is left child of y *) +(* - m y = R and x is right child of y *) +Definition ch (m : nmap mark) (g : pregraph) (x y : nat) := + [|| (find y m == Some L) && (gl g y == x) | + (find y m == Some R) && (gr g y == x)]. + +Lemma chP (m : nmap mark) g x y : + reflect [\/ (y, L) \In m /\ gl g y = x | + (y, R) \In m /\ gr g y = x] + (ch m g x y). +Proof. +rewrite /ch; case: orP=>H; constructor. +- by case: H=>H; [left|right]; case/andP: H=>/eqP/In_find H /eqP. +by case=>[][/In_find/eqP M /eqP G]; apply: H; [left|right]; apply/andP. +Qed. + +Lemma chN0 m g x y : + ch m g x y -> + y != null. +Proof. by case/chP=>[][/In_dom/dom_cond]. Qed. + +Lemma ch_fun m g a b x : + ch m g a x -> + ch m g b x -> + a = b. +Proof. by case/chP=>[][H <-] /chP [][/(In_fun H) {}H <-]. Qed. + +Lemma ch_path m g s x : + x \in s -> + path (ch m g) null s -> + exists y, y \in belast null s /\ ch m g y x. +Proof. exact: path_prev. Qed. + +Lemma ch_path_uniq m g s : + path (ch m g) null s -> + uniq (null :: s). +Proof. by apply: path_uniq; [apply: chN0|apply: ch_fun]. Qed. + +Definition upd_edge (m : nmap mark) g x y : seq node := + if find x m is Some L then [:: y; gr g x] else [:: gl g x; y]. + +Fixpoint rev_graph' m g ps t : pregraph := + if ps is p :: ps then + rev_graph' m (free g p) ps p \+ pts p (upd_edge m g p t) + else g. + +Definition rev_graph m g s t := rev_graph' m g (rev s) t. + +(* layout of graph+marking as heap *) +Definition lay (m : nmap mark) (g : pregraph) : heap := + \big[join/Unit]_(x <- dom g) (x :-> (gl g x, gr g x, odflt L (find x m))). + +(* reach m g s t holds iff *) +(* each unmarked node x in g is reachable through unmarked nodes *) +(* - either starting from t, or *) +(* - starting from a right child of some node in s *) +Definition reach (m : nmap mark) (g : pregraph) (s : seq node) t := + forall x, + (* x is node in g *) + x \in dom g -> + (* x is unmarked *) + x \notin dom m -> + (* x is reachable from t avoiding marked nodes *) + (* Aleks note: if unmarked nodes are represented by a marking U *) + (* then this conjunct will have to change, as it currently assumes *) + (* that unmarked nodes are just those that aren't in m *) + x \in connect (mem (dom m)) g t \/ + (* or x is reachable from some right child of a node y in s *) + (* avoiding marked nodes *) + exists y, y \in s /\ x \in connect (mem (dom m)) g (gr g y). + diff --git a/htt/domain.v.crashcoqide b/htt/domain.v.crashcoqide new file mode 100644 index 0000000..8a14f09 --- /dev/null +++ b/htt/domain.v.crashcoqide @@ -0,0 +1,1038 @@ +(* +Copyright 2019 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 HB Require Import structures. +From Coq Require Import ssreflect ssrbool ssrfun. +From mathcomp Require Import ssrnat eqtype. +From pcm Require Import options axioms pred prelude. + +(**************************************************) +(* This file develops some basic domain theory of *) +(* posets, complete lattices, complete partial *) +(* orders (CPOs). It developes the standard fixed *) +(* point theorems: Knaster-Tarski's for complete *) +(* lattices, and Kleene's for CPOs. *) +(**************************************************) + +(**********) +(* Posets *) +(**********) + +Definition poset_axiom T (leq : T -> T -> Prop) := + [/\ forall x, leq x x, + forall x y, leq x y -> leq y x -> x = y & + forall x y z, leq x y -> leq y z -> leq x z]. + +HB.mixin Record isPoset T := { + poset_leq : T -> T -> Prop; + poset_subproof : poset_axiom poset_leq}. + +#[short(type="poset")] +HB.structure Definition Poset := {T of isPoset T}. + +Notation "x <== y" := (poset_leq x y) (at level 70). + +Section Repack. +Variable T : poset. +Lemma poset_refl (x : T) : x <== x. +Proof. by case: (@poset_subproof T). Qed. +Lemma poset_asym (x y : T) : x <== y -> y <== x -> x = y. +Proof. by case: (@poset_subproof T)=>_ H _; apply: H. Qed. +Lemma poset_trans (y x z : T) : x <== y -> y <== z -> x <== z. +Proof. by case: (@poset_subproof T)=>_ _; apply. Qed. +End Repack. + +#[export] Hint Resolve poset_refl : core. + +Add Parametric Relation (T : poset) : T (@poset_leq T) + reflexivity proved by (@poset_refl _) + transitivity proved by (fun x y => @poset_trans _ y x) as poset_rel. + +(**********************) +(* monotone functions *) +(**********************) + +Definition monofun_axiom (T1 T2 : poset) (f : T1 -> T2) := + forall x y, x <== y -> f x <== f y. + +HB.mixin Record isMonoFun (T1 T2 : poset) (f : T1 -> T2) := { + monofun_subproof : monofun_axiom f}. + +#[short(type="mono_fun")] +HB.structure Definition MonoFun (T1 T2 : poset) := + {f of isMonoFun T1 T2 f}. + +Lemma monofunP (T1 T2 : poset) (f : mono_fun T1 T2) x y : + x <== y -> f x <== f y. +Proof. exact: monofun_subproof. Qed. + +(**************************) +(* some basic definitions *) +(**************************) + +Section IdealDef. +Variable T : poset. + +Structure ideal (P : T) := Ideal {id_val : T; id_pf : id_val <== P}. + +(* Changing the type of the ideal *) + +Lemma relaxP (P1 P2 : T) : P1 <== P2 -> forall p, p <== P1 -> p <== P2. +Proof. by move=>H1 p H2; apply: poset_trans H1. Qed. + +Definition relax (P1 P2 : T) (x : ideal P1) (pf : P1 <== P2) := + Ideal (relaxP pf (id_pf x)). + +End IdealDef. + +(***********************) +(* poset constructions *) +(***********************) + +Section SubPoset. +Variables (T : poset) (s : Pred T). +Local Notation tp := {x : T | x \In s}. + +Definition sub_leq (p1 p2 : tp) := sval p1 <== sval p2. + +Lemma sub_is_poset : poset_axiom sub_leq. +Proof. +split=>[x|[x Hx][y Hy]|[x Hx][y Hy][z Hz]]; rewrite /sub_leq //=. +- move=>H1 H2; rewrite -(poset_asym H1 H2) in Hy *. + by rewrite (pf_irr Hx). +by apply: poset_trans. +Qed. +HB.instance Definition _ := isPoset.Build tp sub_is_poset. +End SubPoset. + +(* pairing of posets *) + +Section ProdPoset. +Variables A B : poset. +Local Notation tp := (A * B)%type. + +Definition poset_prod_leq := + [fun p1 p2 : tp => p1.1 <== p2.1 /\ p1.2 <== p2.2]. + +Lemma prod_is_poset : poset_axiom poset_prod_leq. +Proof. +split=>[//|[x1 x2][y1 y2]|[x1 x2][y1 y2][z1 z2]] /=. +- by case=>H1 H2 [H3 H4]; congr (_, _); apply: poset_asym. +case=>H1 H2 [H3 H4]; split; +by [apply: poset_trans H3|apply: poset_trans H4]. +Qed. +HB.instance Definition _ := isPoset.Build tp prod_is_poset. +End ProdPoset. + +(* functions into poset form a poset *) + +Section FunPoset. +Variable (A : Type) (B : poset). +Local Notation tp := (A -> B). + +Definition poset_fun_leq := [fun p1 p2 : tp => forall x, p1 x <== p2 x]. + +Lemma fun_is_poset : poset_axiom poset_fun_leq. +Proof. +split=>[x|x y H1 H2|x y z H1 H2 t] //=. +- by apply: fext=>z; apply: poset_asym; [apply: H1|apply: H2]. +by apply: poset_trans (H2 t). +Qed. +HB.instance Definition _ := isPoset.Build tp fun_is_poset. +End FunPoset. + +(* dependent functions into a poset form a poset *) + +Section DFunPoset. +Variables (A : Type) (B : A -> poset). +Local Notation tp := (forall x, B x). + +Definition poset_dfun_leq := [fun p1 p2 : tp => forall x, p1 x <== p2 x]. + +Lemma dfun_is_poset : poset_axiom poset_dfun_leq. +Proof. +split=>[x|x y H1 H2|x y z H1 H2 t] //. +- by apply: fext=>z; apply: poset_asym; [apply: H1|apply: H2]. +by apply: poset_trans (H2 t). +Qed. + +(* no instance declaration, since it's keyed on forall *) +(* and forall is not a symbol *) +Definition dfun_poset_mix := isPoset.Build tp dfun_is_poset. +Definition dfun_poset : poset := Poset.pack_ dfun_poset_mix. +End DFunPoset. + +(* ideal of poset is poset *) + +Section IdealPoset. +Variable (T : poset) (P : T). + +Definition poset_ideal_leq := + [fun p1 p2 : ideal P => id_val p1 <== id_val p2]. + +Lemma ideal_is_poset : poset_axiom poset_ideal_leq. +Proof. +split=>[[x]|[x1 H1][x2 H2]|[x1 H1][x2 H2][x3 H3]] //=. +- move=>H3 H4; rewrite (poset_asym H3 H4) in H1 H2 *. + by rewrite (pf_irr H1). +by apply: poset_trans. +Qed. +HB.instance Definition _ := isPoset.Build (ideal P) ideal_is_poset. +End IdealPoset. + +(* Prop is poset *) +Definition poset_prop_leq := [fun p1 p2 : Prop => p1 -> p2]. +Lemma prop_is_poset : poset_axiom poset_prop_leq. +Proof. by split=>[x|x y H1 H2|x y z H1 H2 /H1] //; apply: pext. Qed. +HB.instance Definition _ := isPoset.Build Prop prop_is_poset. + + +(*********************) +(* Complete lattices *) +(*********************) + +Definition lattice_axiom (T : poset) (sup : Pred T -> T) := + [/\ (* sup is upper bound *) + forall (s : Pred T) x, x \In s -> x <== sup s & + (* sup is least upper bound *) + forall (s : Pred T) x, + (forall y, y \In s -> y <== x) -> sup s <== x]. + +HB.mixin Record isLattice T of Poset T := { + sup : Pred T -> T; + lattice_subproof : lattice_axiom sup}. + +#[short(type="lattice")] +HB.structure Definition Lattice := {T of Poset T & isLattice T}. + +Section Repack. +Variable T : lattice. +Lemma supP (s : Pred T) x : x \In s -> x <== sup s. +Proof. by case: (@lattice_subproof T)=>H _; apply: H. Qed. +Lemma supM (s : Pred T) x : (forall y, y \In s -> y <== x) -> sup s <== x. +Proof. by case: (@lattice_subproof T)=>_; apply. Qed. +End Repack. + +(* infima follow *) +Section Infimum. +Variable T : lattice. + +Definition inf (s : Pred T) := + sup [Pred x : T | forall y, y \In s -> x <== y]. + +Lemma infP s x : x \In s -> inf s <== x. +Proof. by move=>H; apply: supM=>y; apply. Qed. + +Lemma infM s x : (forall y, y \In s -> x <== y) -> x <== inf s. +Proof. by apply: supP. Qed. + +End Infimum. + +(* least and greatest fixed points by Knaster-Tarski construction *) +Section KnasterTarskiDefs. +Variable T : lattice. + +Definition lbot : T := sup Pred0. + +Definition tarski_lfp (f : T -> T) := inf [Pred x : T | f x <== x]. +Definition tarski_gfp (f : T -> T) := sup [Pred x : T | x <== f x]. + +Definition sup_closed (T : lattice) := + [Pred s : Pred T | forall d, d <=p s -> sup d \In s]. + +Definition sup_closure (T : lattice) (s : Pred T) := + [Pred p : T | forall t : Pred T, s <=p t -> t \In sup_closed T -> p \In t]. + +End KnasterTarskiDefs. + +Arguments lbot {T}. +Arguments sup_closed {T}. +Arguments sup_closure [T] s. +Prenex Implicits sup_closed sup_closure. + +Section BasicProperties. +Variable T : lattice. + +Lemma sup_mono (s1 s2 : Pred T) : s1 <=p s2 -> sup s1 <== sup s2. +Proof. by move=>H; apply: supM=>y /H/supP. Qed. + +Lemma supE (s1 s2 : Pred T) : s1 =p s2 -> sup s1 = sup s2. +Proof. by move=>H; apply: poset_asym; apply: supM=>y /H/supP. Qed. + +(* Knaster-Tarski theorem *) +Lemma tarski_lfp_fixed (f : mono_fun T T) : + f (tarski_lfp f) = tarski_lfp f. +Proof. +suff L : f (tarski_lfp f) <== tarski_lfp f. +- by apply: poset_asym=>//; apply/infP/monofunP. +by apply: infM=>x L; apply: poset_trans (L); apply/monofunP/infP. +Qed. + +Lemma tarski_lfp_least f (x : T) : f x <== x -> tarski_lfp f <== x. +Proof. by move=>H; apply: infP. Qed. + +Lemma tarski_gfp_fixed (f : mono_fun T T) : + f (tarski_gfp f) = tarski_gfp f. +Proof. +suff L: tarski_gfp f <== f (tarski_gfp f). +- by apply: poset_asym=>//; apply/supP/monofunP. +by apply: supM=>x L; apply: poset_trans (L) _; apply/monofunP/supP. +Qed. + +Lemma tarski_gfp_greatest f (x : T) : x <== f x -> x <== tarski_gfp f. +Proof. by move=>H; apply: supP. Qed. + +Lemma tarski_lfp_mono (f1 : T -> T) (f2 : mono_fun T T) : + f1 <== f2 -> tarski_lfp f1 <== tarski_lfp f2. +Proof. +move=>H; apply: infP; apply: poset_trans (H (tarski_lfp f2)) _. +by rewrite tarski_lfp_fixed. +Qed. + +Lemma tarski_gfp_mono (f1 : mono_fun T T) (f2 : T -> T) : + (f1 : T -> T) <== f2 -> tarski_gfp f1 <== tarski_gfp f2. +Proof. +move=>H; apply/supP/poset_trans/(H (tarski_gfp f1)). +by rewrite tarski_gfp_fixed. +Qed. + +(* closure contains s *) +Lemma sup_clos_sub (s : Pred T) : s <=p sup_closure s. +Proof. by move=>p H1 t H2 H3; apply: H2 H1. Qed. + +(* closure is smallest *) +Lemma sup_clos_min (s : Pred T) : + forall t, s <=p t -> sup_closed t -> sup_closure s <=p t. +Proof. by move=>t H1 H2 p; move/(_ _ H1 H2). Qed. + +(* closure is closed *) +Lemma sup_closP (s : Pred T) : sup_closed (sup_closure s). +Proof. +move=>d H1 t H3 H4; move: (sup_clos_min H3 H4)=>{}H3. +by apply: H4=>x; move/H1; move/H3. +Qed. + +Lemma sup_clos_idemp (s : Pred T) : sup_closed s -> sup_closure s =p s. +Proof. by move=>p; split; [apply: sup_clos_min | apply: sup_clos_sub]. Qed. + +Lemma sup_clos_mono (s1 s2 : Pred T) : + s1 <=p s2 -> sup_closure s1 <=p sup_closure s2. +Proof. +move=>H1; apply: sup_clos_min (@sup_closP s2). +by move=>t /H1; apply: sup_clos_sub. +Qed. + +End BasicProperties. + +(* lattice constructions *) + +(* subset lattice *) +Section SubLattice. +Variables (T : lattice) (s : Pred T) (C : sup_closed s). +Local Notation tp := {x : T | x \In s}. + +Definition lattice_sub_sup' (u : Pred tp) : T := + sup [Pred x : T | exists y, y \In u /\ x = sval y]. +Lemma lattice_sub_supX (u : Pred tp) : lattice_sub_sup' u \In s. +Proof. by apply: C=>x [][y H][_] ->. Qed. +Definition lattice_sub_sup (u : Pred tp) : tp := + exist _ (lattice_sub_sup' u) (lattice_sub_supX u). + +Lemma sub_is_lattice : lattice_axiom lattice_sub_sup. +Proof. +split=>u x H; first by apply: supP; exists x. +by apply: supM=>y [z][H1] ->; apply: H H1. +Qed. +HB.instance Definition _ := isLattice.Build tp sub_is_lattice. +End SubLattice. + +(* product *) +Section ProdLattice. +Variables (A B : lattice). +Local Notation tp := (A * B)%type. + +Definition lattice_prod_sup (s : Pred tp) : tp := + (sup [Pred p | exists f, p = f.1 /\ f \In s], + sup [Pred p | exists f, p = f.2 /\ f \In s]). + +Lemma prod_is_lattice : lattice_axiom lattice_prod_sup. +Proof. +split=>s p H; first by split; apply: supP; exists p. +by split; apply: supM=>y [f][->]; case/H. +Qed. +HB.instance Definition _ := isLattice.Build tp prod_is_lattice. +End ProdLattice. + +(* functions into latice form lattice *) +Section FunLattice. +Variables (A : Type) (B : lattice). +Local Notation tp := (A -> B). + +Definition lattice_fun_sup (s : Pred tp) : tp := + fun x => sup [Pred p | exists f, f \In s /\ p = f x]. + +Lemma fun_is_lattice : lattice_axiom lattice_fun_sup. +Proof. +split=>s p H x; first by apply: supP; exists p. +by apply: supM=>y [f][H1] ->; apply: H. +Qed. +HB.instance Definition _ := isLattice.Build tp fun_is_lattice. +End FunLattice. + +(* dependent functions into a lattice form a lattice *) +Section DFunLattice. +Variables (A : Type) (B : A -> lattice). +Local Notation tp := (dfun_poset B). + +Definition lattice_dfun_sup (s : Pred tp) : tp := + fun x => sup [Pred p | exists f, f \In s /\ p = f x]. + +Lemma dfun_is_lattice : lattice_axiom lattice_dfun_sup. +Proof. +split=>s p H x; first by apply: supP; exists p. +by apply: supM=>y [f][H1] ->; apply: H. +Qed. + +Definition dfun_lattice_mix := isLattice.Build tp dfun_is_lattice. +Definition dfun_lattice : lattice := Lattice.pack_ dfun_lattice_mix. +End DFunLattice. + +(* applied sup equals the sup of applications *) +Lemma sup_appE A (B : lattice) (s : Pred (A -> B)) (x : A) : + sup s x = sup [Pred y : B | exists f, f \In s /\ y = f x]. +Proof. by []. Qed. + +Lemma sup_dappE A (B : A -> lattice) (s : Pred (dfun_lattice B)) (x : A) : + sup s x = sup [Pred y : B x | exists f, f \In s /\ y = f x]. +Proof. by []. Qed. + +(* ideal of a lattice forms a lattice *) +Section IdealLattice. +Variables (T : lattice) (P : T). + +Definition lattice_ideal_sup' (s : Pred (ideal P)) := + sup [Pred x | exists p, p \In s /\ id_val p = x]. +Lemma lattice_ideal_supP' (s : Pred (ideal P)) : lattice_ideal_sup' s <== P. +Proof. by apply: supM=>y [[x]] H /= [_] <-. Qed. +Definition lattice_ideal_sup (s : Pred (ideal P)) := + Ideal (lattice_ideal_supP' s). + +Lemma ideal_is_lattice : lattice_axiom lattice_ideal_sup. +Proof. +split=>s p H; first by apply: supP; exists p. +by apply: supM=>y [q][H1] <-; apply: H. +Qed. +HB.instance Definition _ := isLattice.Build (ideal P) ideal_is_lattice. +End IdealLattice. + +(* Prop is a lattice *) +Definition lattice_prop_sup (s : Pred Prop) : Prop := + exists p, p \In s /\ p. +Lemma prop_is_lattice : lattice_axiom lattice_prop_sup. +Proof. by split=>s p; [exists p|move=>H [r][]; move/H]. Qed. +HB.instance Definition _ := isLattice.Build Prop prop_is_lattice. + +(****************) +(* Monotonicity *) +(****************) + +Section MonoCanonicals. +Variables T T1 T2 T3 S1 S2 : poset. + +Lemma idfun_is_mono : monofun_axiom (@idfun T). Proof. by []. Qed. +HB.instance Definition _ := isMonoFun.Build T T idfun idfun_is_mono. + +Definition const_fun (y : T2) (_ : T1) := y. +Lemma const_is_mono (y : T2) : monofun_axiom (const_fun y). Proof. by []. Qed. +HB.instance Definition _ y := + isMonoFun.Build T1 T2 (const_fun y) (const_is_mono y). + +Lemma fst_is_mono : monofun_axiom (@fst T1 T2). +Proof. by case=>x1 x2 [y1 y2][]. Qed. +HB.instance Definition _ := isMonoFun.Build (T1*T2)%type T1 fst fst_is_mono. + +Lemma snd_is_mono : monofun_axiom (@snd T1 T2). +Proof. by case=>x1 x2 [y1 y2][]. Qed. +HB.instance Definition _ := isMonoFun.Build (T1*T2)%type T2 snd snd_is_mono. + +Definition diag (x : T) := (x, x). +Lemma diag_is_mono : monofun_axiom diag. Proof. by []. Qed. +HB.instance Definition _ := isMonoFun.Build T (T*T)%type diag diag_is_mono. + +Lemma sval_is_mono (P : Pred T) : monofun_axiom (@sval T P). +Proof. by []. Qed. +HB.instance Definition _ P := + isMonoFun.Build {x : T | P x} T sval (@sval_is_mono P). + +Definition app A B (x : A) := fun f : A -> B => f x. +Lemma app_is_mono A (x : A) : monofun_axiom (@app A T x). +Proof. by move=>f1 f2; apply. Qed. +HB.instance Definition _ A x := + isMonoFun.Build (A -> T) T (@app A T x) (app_is_mono x). + +Definition dapp A (B : A -> poset) (x : A) := fun f : dfun_poset B => f x. +Lemma dapp_is_mono A (B : A -> poset) (x : A) : monofun_axiom (@dapp A B x). +Proof. by move=>f1 f2; apply. Qed. +HB.instance Definition _ A B x := + isMonoFun.Build (dfun_poset B) (B x) (@dapp A B x) (dapp_is_mono x). + +Section Comp. +Variables (f1 : mono_fun T2 T1) (f2 : mono_fun T3 T2). +Lemma comp_is_mono : monofun_axiom (f1 \o f2). +Proof. by move=>x y H; apply/monofunP/monofunP. Qed. +HB.instance Definition _ := isMonoFun.Build T3 T1 (f1 \o f2) comp_is_mono. +End Comp. + +Section Prod. +Variables (f1 : mono_fun S1 T1) (f2 : mono_fun S2 T2). +Lemma funprod_is_mono : monofun_axiom (f1 \* f2). +Proof. by case=>x1 x2 [y1 y2][/= H1 H2]; split=>/=; apply: monofunP. Qed. +HB.instance Definition _ := + isMonoFun.Build (S1 * S2)%type (T1 * T2)%type (f1 \* f2) funprod_is_mono. +End Prod. +End MonoCanonicals. + +Prenex Implicits diag app. + +(**********) +(* Chains *) +(**********) + +Definition chain_axiom (T : poset) (s : Pred T) := + (exists d, d \In s) /\ + (forall x y, x \In s -> y \In s -> x <== y \/ y <== x). + +HB.mixin Record isChain (T : poset) (s : Pred T) := { + chain_subproof : chain_axiom s}. + +#[short(type="chain")] +HB.structure Definition Chain (T : poset) := {s of isChain T s}. + +Arguments chain T%_type. + +(* \In notation *) +Definition chain_pred (T : poset) (x : chain T) : Pred T := x. +Canonical chainPredType (T : poset) := + Eval hnf in @mkPredType T (@chain T) (@chain_pred T). + +Lemma chainE (T : poset) (s1 s2 : chain T) : + s1 = s2 <-> chain_pred s1 =p chain_pred s2. +Proof. +split=>[->//|]; case: s1 s2=>s1 H1 [s2 H2] /= E; move: H1 H2. +suff: s1 = s2. +- move=>-> [[H1]][[H2]]; congr (Chain.Pack (Chain.Class _)). + by rewrite (pf_irr H1). +by apply: fext=>x; apply: pext; split=> /E. +Qed. + +Section ImageChain. +Variables (T1 T2 : poset) (f : mono_fun T1 T2) (s : chain T1). + +Lemma image_is_chain : chain_axiom (Image f s). +Proof. +case: s=>p [[[[d H1] H2]]] /=. +split=>[|x y]; first by exists (f d); exists d. +case=>x1 -> H3 [y1 -> H4]; rewrite -!toPredE /= in H3 H4. +by case: (H2 x1 y1 H3 H4)=>L; [left | right]; apply: monofunP L. +Qed. +HB.instance Definition _ := isChain.Build T2 (Image f s) image_is_chain. +End ImageChain. + +Lemma id_chainE T (s : chain T) f : f =1 idfun -> Image f s =p s. +Proof. +move=>H x; split; first by case=>y ->; rewrite H. +by exists x=>//; rewrite H. +Qed. + +Section ChainConst. +Variables (T1 T2 : poset) (y : T2). + +Definition const_chain : Pred _ := Pred1 y. +Lemma const_is_chain : chain_axiom const_chain. +Proof. by split; [exists y | move=>x1 x2 <-<-; left]. Qed. +HB.instance Definition _ := isChain.Build T2 const_chain const_is_chain. + +Lemma const_chainE (s : chain T1) : Image (const_fun y) s =p const_chain. +Proof. +move=>z1; split; first by case=>z2 ->. +by case: s=>s [[[[d H1]] H2]] /= <-; exists d. +Qed. +End ChainConst. + +Section ChainCompose. +Variables T1 T2 T3 : poset. +Variables (f1 : mono_fun T2 T1) (f2 : mono_fun T3 T2). +Variable s : chain T3. +Lemma comp_chainE : Image f1 (Image f2 s) =p Image (f1 \o f2) s. +Proof. by move=>x; apply: ImageIm. Qed. +End ChainCompose. + +Section ProjChain. +Variables T1 T2 : poset. +Variables s : chain (T1 * T2). +Definition proj1_chain : Pred T1 := Image fst s. +HB.instance Definition _ := Chain.copy proj1_chain proj1_chain. +Definition proj2_chain : Pred T2 := Image snd s. +HB.instance Definition _ := Chain.copy proj2_chain proj2_chain. +End ProjChain. + +Section DiagChain. +Variables (T : poset) (s : chain T). +Definition diag_chain : Pred (T * T) := Image diag s. +HB.instance Definition _ := Chain.copy diag_chain diag_chain. +End DiagChain. + + +(*********) +(* CPO's *) +(*********) + +Definition cpo_axiom (T : poset) (bot : T) (lim : chain T -> T) := + [/\ forall x, bot <== x, + forall (s : chain T) x, x \In s -> x <== lim s & + forall (s : chain T) x, + (forall y, y \In s -> y <== x) -> lim s <== x]. + +HB.mixin Record isCPO T of Poset T := { + bot : T; + lim_op : chain T -> T; + cpo_subproof: cpo_axiom bot lim_op}. + +#[short(type="cpo")] +HB.structure Definition CPO := {T of Poset T & isCPO T}. + +Definition limx (D : cpo) (s : chain D) of phantom (Pred _) s := lim_op s. +Notation lim s := (limx (Phantom (Pred _) s)). + +Section Repack. +Variable D : cpo. +Lemma botP (x : D) : bot <== x. +Proof. by case: (@cpo_subproof D). Qed. +Lemma limP (s : chain D) x : x \In s -> x <== lim s. +Proof. by case: (@cpo_subproof D)=>_ H _; apply: H. Qed. +Lemma limM (s : chain D) x : (forall y, y \In s -> y <== x) -> lim s <== x. +Proof. by case: (@cpo_subproof D)=>_ _; apply. Qed. +End Repack. + +#[export] Hint Resolve botP : core. + +Lemma limE (D : cpo) (s1 s2 : chain D) : s1 =p s2 -> lim s1 = lim s2. +Proof. by move/chainE=>->. Qed. + +(* common chain constructions *) + +(* adding bot to a chain *) + +Definition lift (D : cpo) (s : Pred D) : Pred D := + fun x => x = bot \/ x \In s. + +Lemma lift_is_chain (D : cpo) (s : chain D) : chain_axiom (lift s). +Proof. +case: s=>p [[[[d H1]] H2]] /=. +split=>[|x y]; first by exists d; right. +by case=>[->|H3][->|H4]; auto. +Qed. + +HB.instance Definition _ (D : cpo) (s : chain D) := + isChain.Build D (lift s) (@lift_is_chain D s). + + +(****************************) +(* common cpo constructions *) +(****************************) + +(* products *) +Section ProdCPO. +Variables (A B : cpo). +Definition cpo_prod_bot := (@bot A, @bot B). +Definition cpo_prod_lim (s : chain (A * B)) := + (lim (Image fst s), lim (Image snd s)). + +Lemma prod_is_cpo : cpo_axiom cpo_prod_bot cpo_prod_lim. +Proof. +split=>[x|s x|s x]. +- by split; apply: botP. +- by split; apply: limP; exists x. +by split; apply: limM=>y [z ->]; case/H. +Qed. +HB.instance Definition _ := isCPO.Build (A * B)%type prod_is_cpo. +End ProdCPO. + +(* functions *) +Section FunCPO. +Variable (A : Type) (B : cpo). +Definition cpo_fun_bot (x : A) := @bot B. +Definition cpo_fun_lim (s : chain (A -> B)) x := + lim (Image (app x) s). + +Lemma fun_is_cpo : cpo_axiom cpo_fun_bot cpo_fun_lim. +Proof. +split=>[f|s f|s f]. +- by move=>y; apply: botP. +- by move=>H t; apply: limP; exists f. +by move=>H1 t; apply: limM=>y [g ->] H2; apply: H1. +Qed. +HB.instance Definition _ := isCPO.Build (A -> B) fun_is_cpo. +End FunCPO. + +(* dependent functions *) +Section DFunCPO. +Variables (A : Type) (B : A -> cpo). + +Definition cpo_dfun_bot : dfun_poset B := fun x => bot. +Definition cpo_dfun_lim (s : chain (dfun_poset B)) : dfun_poset B := + fun x => lim (Image (dapp x) s). + +Lemma dfun_is_cpo : cpo_axiom cpo_dfun_bot cpo_dfun_lim. +Proof. +split=>[x|s x|s x]. +- by move=>y; apply: botP. +- by move=>H t; apply: limP; exists x. +by move=>H1 t; apply: limM=>y [f ->] H2; apply: H1. +Qed. + +Definition dfun_cpo_mix := isCPO.Build (dfun_poset B) dfun_is_cpo. +Definition dfun_cpo : cpo := CPO.pack_ dfun_cpo_mix. +End DFunCPO. + +(* Prop *) +Definition cpo_prop_bot : Prop := False. +Definition cpo_prop_lim (s : chain Prop) : Prop := + exists p, p \In s /\ p. +Lemma prop_is_cpo : cpo_axiom cpo_prop_bot cpo_prop_lim. +Proof. by split=>[x|s p|s p H] //=; [exists p|case=>r []; move/H]. Qed. +HB.instance Definition _ := isCPO.Build Prop prop_is_cpo. + +(* Pred *) +HB.instance Definition _ A := CPO.copy (Pred A) _. + +(* every complete lattice is cpo *) +Section LatticeCPO. +Variable A : lattice. + +Definition lat_bot : A := lbot. +Definition lat_lim (s : Pred A) : A := sup s. + +Lemma lat_is_cpo : cpo_axiom lat_bot lat_lim. +Proof. by split=>[x|s x|s x]; [apply: supM|apply: supP|apply: supM]. Qed. + +(* no instance declaration as nothing to *) +(* hang the new structure onto *) +Definition lat_cpo_mix := isCPO.Build A lat_is_cpo. +Definition lat_cpo := CPO.pack_ lat_cpo_mix. +End LatticeCPO. + +(* sub-CPO's *) + +(* every chain-closed subset of a cpo is a cpo *) + +Section AdmissibleClosure. +Variable T : cpo. + +Definition chain_closed := + [Pred s : Pred T | + bot \In s /\ forall d : chain T, d <=p s -> lim d \In s]. + +(* admissible closure of s is the smallest closed set containing s *) +(* basically extends s to include the sups of chains *) +Definition chain_closure (s : Pred T) := + [Pred p : T | forall t : Pred T, s <=p t -> chain_closed t -> p \In t]. + +(* admissible closure contains s *) +Lemma chain_clos_sub (s : Pred T) : s <=p chain_closure s. +Proof. by move=>p H1 t H2 H3; apply: H2 H1. Qed. + +(* admissible closure is smallest *) +Lemma chain_clos_min (s : Pred T) t : + s <=p t -> chain_closed t -> chain_closure s <=p t. +Proof. by move=>H1 H2 p; move/(_ _ H1 H2). Qed. + +(* admissible closure is closed *) +Lemma chain_closP (s : Pred T) : chain_closed (chain_closure s). +Proof. +split; first by move=>t _ []. +move=>d H1 t H3 H4; move: (chain_clos_min H3 H4)=>{}H3. +by case: H4=>_; apply=>// x; move/H1; move/H3. +Qed. + +Lemma chain_clos_idemp (s : Pred T) : + chain_closed s -> chain_closure s =p s. +Proof. +move=>p; split; last by apply: chain_clos_sub. +by apply: chain_clos_min=>// /chain_closP. +Qed. + +Lemma chain_clos_mono (s1 s2 : Pred T) : + s1 <=p s2 -> chain_closure s1 <=p chain_closure s2. +Proof. +move=>H1; apply: chain_clos_min (chain_closP s2)=>p H2. +by apply: chain_clos_sub; apply: H1. +Qed. + +(* intersection of admissible sets is admissible *) +Lemma chain_closI (s1 s2 : Pred T) : + chain_closed s1 -> chain_closed s2 -> chain_closed (PredI s1 s2). +Proof. +move=>[H1 S1][H2 S2]; split=>// d H. +by split; [apply: S1 | apply: S2]=>// x; case/H. +Qed. + +End AdmissibleClosure. + +Arguments chain_closed {T}. +Prenex Implicits chain_closed. + +(* diagonal of an admissible set of pairs is admissible *) +Lemma chain_clos_diag (T : cpo) (s : Pred (T * T)) : + chain_closed s -> chain_closed [Pred t : T | (t, t) \In s]. +Proof. +case=>B H1; split=>// d H2; move: (H1 (Image diag d)). +rewrite /limx/=/lim_op /= /cpo_prod_lim /=. +rewrite !(limE (comp_chainE _ _ _)) !(limE (id_chainE _ _)) //=. +by apply; case=>x1 x2 [z ->]; apply: H2. +Qed. + +Section SubCPO. +Variables (D : cpo) (s : Pred D) (C : chain_closed s). +Local Notation tp := {x : D | x \In s}. + +Definition subcpo_bot := exist _ (@bot D) (proj1 C). +Lemma subcpo_limX (u : chain tp) : lim (Image sval u) \In s. +Proof. by case: C u=>P H u; apply: (H)=>t [[y H1 ->]]. Qed. +Definition subcpo_lim (u : chain tp) : tp := + exist _ (lim (Image sval u)) (subcpo_limX u). + +Lemma subcpo_is_cpo : cpo_axiom subcpo_bot subcpo_lim. +Proof. +split=>[x|u x H|u x H] //=. +- by apply: botP. +- by apply: limP; exists x. +by apply: limM=>y [z ->]; apply: H. +Qed. +HB.instance Definition _ := isCPO.Build tp subcpo_is_cpo. +End SubCPO. + +(***********************************************) +(* Continuity and Kleene's fixed point theorem *) +(***********************************************) + +Lemma lim_mono (D : cpo) (s1 s2 : chain D) : + s1 <=p s2 -> lim s1 <== lim s2. +Proof. by move=>H; apply: limM=>y /H/limP. Qed. + +Lemma lim_liftE (D : cpo) (s : chain D) : lim s = lim (lift s). +Proof. +apply: poset_asym; apply: limM=>y H; first by apply: limP; right. +by case: H; [move=>-> | apply: limP]. +Qed. + +(* applied lim equals the lim of applications *) +(* ie., part of continuity of application *) +(* but is so often used, I give it a name *) +Lemma lim_appE A (D : cpo) (s : chain (A -> D)) (x : A) : + lim s x = lim (Image (app x) s). +Proof. by []. Qed. + +Lemma lim_dappE A (D : A -> cpo) (s : chain (dfun_cpo D)) (x : A) : + lim s x = lim (Image (dapp x) s). +Proof. by []. Qed. + +(************************) +(* continuous functions *) +(************************) + +Definition contfun_axiom (D1 D2 : cpo) (f : mono_fun D1 D2) := + forall s : chain D1, f (lim s) <== lim (Image f s). + +HB.mixin Record isContFun (D1 D2 : cpo) (f : D1 -> D2) + of @MonoFun D1 D2 f := {contfun_subproof : contfun_axiom f}. + +#[short(type="cont_fun")] +HB.structure Definition ContFun (D1 D2 : cpo) := + {f of @MonoFun D1 D2 f & isContFun D1 D2 f}. + +Lemma contE (D1 D2 : cpo) (f : cont_fun D1 D2) (s : chain D1) : + f (lim s) = lim (Image f s). +Proof. +apply: poset_asym; first by apply: contfun_subproof. +by apply: limM=>y [x1 ->] /limP/monofunP. +Qed. + +(* just for this proof, we want that nat is a poset under <= *) +(* in other scopes, we'll use the equality as ordering *) +(* Thus, we encapsulate under module. *) + +Module Kleene. + +Lemma nat_is_poset : poset_axiom leq. +Proof. +split=>[|x y H1 H2|x y z] //; last by apply: leq_trans. +by apply: anti_leq; rewrite H1 H2. +Qed. +HB.instance Definition _ := isPoset.Build nat nat_is_poset. + +(* nats form a chain *) +Definition nats : Pred nat := PredT. +Lemma nats_chain_axiom : chain_axiom nats. +Proof. +split=>[|x y _ _]; first by exists 0. +rewrite /poset_leq /= [y <= x]leq_eqVlt. +by case: leqP; [left | rewrite orbT; right]. +Qed. +HB.instance Definition _ := isChain.Build nat nats nats_chain_axiom. + +Section Kleene. +Variables (D : cpo) (f : cont_fun D D). + +Fixpoint pow m := if m is n.+1 then f (pow n) else bot. + +Lemma pow_is_mono : monofun_axiom pow. +Proof. +move=>m n; elim: n m=>[|n IH] m /=; first by case: m. +rewrite {1}/poset_leq /= leq_eqVlt ltnS. +case/orP; first by move/eqP=>->. +move/IH=>{IH} H; apply: poset_trans H _. +by elim: n=>[|n IH] //=; apply: monofunP IH. +Qed. +HB.instance Definition _ := isMonoFun.Build nat D pow pow_is_mono. + +Lemma reindex : Image pow nats =p lift (Image f (Image pow nats)). +Proof. +move=>x; split. +- case; case=>[|n] -> _; first by left. + by right; exists (pow n)=>//; exists n. +case=>/=; first by move=>->; exists 0. +by case=>y -> [n ->]; exists n.+1. +Qed. + +Definition kleene_lfp := lim (Image pow nats). + +Lemma kleene_lfp_fixed : f kleene_lfp = kleene_lfp. +Proof. by rewrite /kleene_lfp contE /= (limE reindex) /= lim_liftE. Qed. + +Lemma kleene_lfp_least x : f x = x -> kleene_lfp <== x. +Proof. +move=>H; apply: limM=>y [n ->] _. +by elim: n=>[|n IH] //=; rewrite -H; apply: monofunP IH. +Qed. + +End Kleene. + +Lemma kleene_lfp_mono (D : cpo) (f1 f2 : cont_fun D D) : + (f1 : D -> D) <== f2 -> + kleene_lfp f1 <== kleene_lfp f2. +Proof. +move=>H; apply: limM=>x [n ->] X. +have N : forall n, pow f1 n <== pow f2 n. +- by elim=>//= m /(monofunP f2); apply: poset_trans (H _). +by apply: poset_trans (N n) _; apply: limP; exists n. +Qed. + +Module Exports. +Notation kleene_lfp := kleene_lfp. +Notation kleene_lfp_fixed := kleene_lfp_fixed. +Notation kleene_lfp_least := kleene_lfp_least. +Notation kleene_lfp_mono := kleene_lfp_mono. +End Exports. + +End Kleene. + +Export Kleene.Exports. + + +(**********************************) +(* Continuity of common functions *) +(**********************************) + +Lemma idfun_is_cont (D : cpo) : contfun_axiom (@idfun D). +Proof. by move=>d; rewrite (limE (id_chainE _ _)). Qed. +HB.instance Definition _ (D : cpo) := + isContFun.Build D D idfun (@idfun_is_cont D). + +Lemma const_is_cont (D1 D2 : cpo) (y : D2) : + contfun_axiom (@const_fun D1 D2 y). +Proof. +by move=>s; apply: limP; case: s=>[p][[[[d H1] H2]]]; exists d. +Qed. +HB.instance Definition _ (D1 D2 : cpo) (y : D2) := + isContFun.Build D1 D2 (const_fun y) (const_is_cont y). + +Section ContCompose. +Variables D1 D2 D3 : cpo. +Variable (f1 : cont_fun D2 D1) (f2 : cont_fun D3 D2). +Lemma comp_is_cont : contfun_axiom (f1 \o f2). +Proof. by move=>s; rewrite /= !contE (limE (comp_chainE _ _ _)). Qed. +HB.instance Definition _ := isContFun.Build D3 D1 (f1 \o f2) comp_is_cont. +End ContCompose. + +Lemma fst_is_cont (D1 D2 : cpo) : contfun_axiom (@fst D1 D2). +Proof. by []. Qed. +HB.instance Definition _ (D1 D2 : cpo) := + isContFun.Build (D1*D2)%type D1 fst (@fst_is_cont D1 D2). + +Lemma snd_is_cont (D1 D2 : cpo) : contfun_axiom (@snd D1 D2). +Proof. by []. Qed. +HB.instance Definition _ (D1 D2 : cpo) := + isContFun.Build (D1*D2)%type D2 snd (@snd_is_cont D1 D2). + +Lemma diag_is_cont (D : cpo) : contfun_axiom (@diag D). +Proof. +move=>d; split=>/=; +by rewrite (limE (comp_chainE _ _ _)) (limE (id_chainE _ _)). +Qed. +HB.instance Definition _ (D : cpo) := + isContFun.Build D (D*D)%type diag (@diag_is_cont D). + +Lemma app_is_cont A (D : cpo) x : contfun_axiom (@app A D x). +Proof. by []. Qed. +HB.instance Definition _ A (D : cpo) x := + isContFun.Build (A -> D) D (@app A D x) (app_is_cont x). + +(* can't reuse dapp and it's proof of monotonicity *) +(* because inheritance doesn't propagate properly *) +(* for dependent functions; so we duplicate *) +(* See mathcomp/finfun.v for solution (dfun vs. dffun). *) +Definition dapp2 A (B : A -> cpo) (x : A) := + fun f : dfun_cpo B => f x. +Lemma dapp2_is_mono A (B : A -> cpo) x : monofun_axiom (@dapp2 A B x). +Proof. by move=>f1 f2; apply. Qed. +HB.instance Definition _ A (B : A -> cpo) x := + isMonoFun.Build (dfun_poset B) (B x) (@dapp2 A B x) (dapp2_is_mono x). +Lemma dapp2_is_cont A (B : A -> cpo) (x : A) : contfun_axiom (@dapp2 A B x). +Proof. +move=>s; have <- // : dapp2 x (lim s) = lim [Image dapp2 x i | i <- s]. +by apply: limE=>/=. +Qed. +HB.instance Definition _ A (B : A -> cpo) x := + isContFun.Build (dfun_cpo B) (B x) (@dapp2 A B x) (dapp2_is_cont x). + +Section ProdCont. +Variables S1 S2 T1 T2 : cpo. +Variables (f1 : cont_fun S1 T1) (f2 : cont_fun S2 T2). + +Lemma prod_is_cont : contfun_axiom (f1 \* f2). +Proof. +move=>d; split=>//=; +by rewrite contE !(limE (comp_chainE _ _ _)); apply: lim_mono. +Qed. +HB.instance Definition _ := + isContFun.Build (S1*S2)%type (T1*T2)%type (f1 \* f2) prod_is_cont. +End ProdCont. + +(* DEVCOMMENT *) +(* TODO: *) +(* 1. limit of a chain of continuous functions is a continuous function *) +(* 2. CPO of continuous functions should be the subCPO of functions *) +(* 3. kleene_lfp is a continuous operation *) +(* /DEVCOMMENT *) + From 4a624786be69d245cfec694fc9d1f429dd164c6d Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Thu, 15 Aug 2024 18:14:12 +0200 Subject: [PATCH 54/93] blah --- _CoqProject | 4 ++-- examples/schorr.v | 9 +++------ 2 files changed, 5 insertions(+), 8 deletions(-) diff --git a/_CoqProject b/_CoqProject index 27a2862..ff398d9 100644 --- a/_CoqProject +++ b/_CoqProject @@ -50,7 +50,7 @@ examples/hashtab.v examples/bubblesort.v examples/quicksort.v examples/congmath.v -examples/congprog.v +#examples/congprog.v examples/tree.v examples/union_find.v -examples/graph.v \ No newline at end of file +examples/graph.v diff --git a/examples/schorr.v b/examples/schorr.v index c76afa1..ff49d98 100644 --- a/examples/schorr.v +++ b/examples/schorr.v @@ -54,13 +54,10 @@ Notation gl g x := (get_nth g x 0). Notation gr g x := (get_nth g x 1). (* type of markings *) -(* Aleks NOTE: Should we need U? My impression was that *) -(* being unmarked is represented by not being in *) -(* the appropriate map? *) -Inductive mark := U | L | R | LR. +Inductive mark := L | R | LR. (* decidable equality on marks *) Definition eq_mark l1 l2 : bool := - if (l1, l2) is ((U, U)|(L, L)|(R, R)|(LR, LR)) then true else false. + if (l1, l2) is ((L, L)|(R, R)|(LR, LR)) then true else false. Lemma eq_markP : Equality.axiom eq_mark. Proof. by case; case=>//=; constructor. Qed. HB.instance Definition _ := hasDecEq.Build mark eq_markP. @@ -146,7 +143,7 @@ Definition shape g0 r p t := path (ch m g) null s /\ rev_graph m g s t = g0 /\ reach m g s t /\ - [pcm um_filterv (fun v => ( eq_mark v L || eq_mark v LR)) m <= s ] /\ + {subset dom (um_filterv (predU (pred1 L) (pred1 LR)) m) <= s} /\ p = last null s /\ r = head t s /\ graph_axiom g). From 4f64d25222ce31d34bb582494ca7d23548b2aa20 Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Fri, 16 Aug 2024 07:14:30 +0200 Subject: [PATCH 55/93] reorgainzation of pred.v prelude.v and consequences. also, largely cleaned up kvmaps.v and hashtable.v to get the inheritance and hiding to behave properly --- core/pred.v | 858 ++++++++++++++++++++++++++---------------- core/prelude.v | 200 +++++++++- examples/array.v | 42 ++- examples/bubblesort.v | 11 +- examples/congmath.v | 185 ++++----- examples/congprog.v | 142 +++---- examples/graph.v | 4 +- examples/hashtab.v | 79 ++-- examples/kvmaps.v | 203 ++++++---- examples/quicksort.v | 9 +- examples/tree.v | 6 +- htt/domain.v | 7 +- pcm/heap.v | 86 +---- pcm/natmap.v | 12 +- pcm/unionmap.v | 20 +- 15 files changed, 1119 insertions(+), 745 deletions(-) diff --git a/core/pred.v b/core/pred.v index d989dc5..e1c6a3e 100644 --- a/core/pred.v +++ b/core/pred.v @@ -58,10 +58,12 @@ Open Scope fun_scope. (**************************************************************************) Definition Pred T := T -> Prop. -Identity Coercion fun_of_Pred : Pred >-> Funclass. +Identity Coercion Fun_Of_Pred : Pred >-> Funclass. + +Definition SubPred T (p1 p2 : Pred T) := forall x : T, p1 x -> p2 x. Notation xPred0 := (fun _ => False). -Notation xPred1 := (fun x y => x = y). +Notation xPred1 := (fun x => eq^~ x). 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). @@ -69,260 +71,401 @@ 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. +(* The packed class interface for pred-like types. *) -(* simple predicates *) +Structure PredType T := + PropPredType {Pred_Sort :> Type; toPred : Pred_Sort -> Pred T}. -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. +Definition Clone_Pred T U := + fun pT & @Pred_Sort T pT -> U => + fun toP (pT' := @PropPredType T U toP) & phant_id pT' pT => pT'. +Notation "[ 'PredType' 'of' T ]" := (@Clone_Pred _ T _ id _ id) : form_scope. -(* it's useful to declare the operations as simple predicates, so that *) -(* complex expressions automatically reduce when used in applicative *) -(* positions *) +Canonical PredPredType T := PropPredType (@id (Pred T)). +Canonical PropFunPredType T := PropPredType (@id (T -> Prop)). -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). +Notation "{ 'Pred' T }" := (Pred_Sort (PredPredType T)) + (format "{ 'Pred' T }") : type_scope. -(* membership predicates *) +Definition Eq_Pred {T} (pT : PredType T) (mp1 mp2 : pT) := + forall x : T, toPred mp1 x <-> toPred mp2 x. +Definition Sub_Pred {T} (pT : PredType T) (mp1 mp2 : pT) := + forall x : T, toPred mp1 x -> toPred mp2 x. +Definition Eq_PredFun {T1 T2} (pT2 : PredType T2) p1 p2 := + forall x : T1, @Eq_Pred T2 pT2 (p1 x) (p2 x). +Definition Sub_PredFun T1 T2 (pT2 : PredType T2) p1 p2 := + forall x : T1, @Sub_Pred T2 pT2 (p1 x) (p2 x). -Variant Mem_Pred : Type := MemProp of Pred T. -Definition isMem pT toPred mem := mem = (fun p : pT => MemProp [eta toPred p]). +Notation "A <~> B" := (@Eq_Pred _ _ A B) + (at level 70, no associativity) : rel_scope. +Notation "A ~> B" := (@Sub_Pred _ _ A B) + (at level 70, no associativity) : rel_scope. +Notation "A <~1> B" := (@Eq_PredFun _ _ _ A B) + (at level 70, no associativity) : rel_scope. +Notation "A ~1> B" := (@Sub_PredFun _ _ _ A B) + (at level 70, no associativity) : rel_scope. -(* the general structure for predicates *) +(* The type of self-simplifying collective predicates. *) + +Definition Simpl_Pred T := simpl_fun T Prop. +Definition PropSimplPred {T} (p : Pred T) : Simpl_Pred T := SimplFun p. + +(* Some Simple_Pred constructors *) + +Definition Pred0 {T} := @PropSimplPred T xPred0. +Definition PredT {T} := @PropSimplPred T xPredT. +Definition Pred1 {T} (x : T) := PropSimplPred (xPred1 x). +Definition PredI {T} (p1 p2 : Pred T) := PropSimplPred (xPredI p1 p2). +Definition PredU {T} (p1 p2 : Pred T) := PropSimplPred (xPredU p1 p2). +Definition PredC {T} (p : Pred T) := PropSimplPred (xPredC p). +Definition PredD {T} (p1 p2 : Pred T) := PropSimplPred (xPredD p1 p2). +Definition Preim {aT rT} (f : aT -> rT) (d : Pred rT) := + PropSimplPred (xPreim f d). + +Notation "[ 'Pred' : T | E ]" := (PropSimplPred (fun _ : T => E)) + (format "[ 'Pred' : T | E ]") : fun_scope. +Notation "[ 'Pred' x | E ]" := (PropSimplPred (fun x => E)) + (x name, format "[ 'Pred' x | E ]") : fun_scope. +Notation "[ 'Pred' x : T | E ]" := (PropSimplPred (fun x : T => E)) + (x name, only parsing) : fun_scope. +Notation "[ 'Pred' x | E1 & E2 ]" := [Pred x | E1 /\ E2 ] + (x name) : fun_scope. +Notation "[ 'Pred' x : T | E1 & E2 ]" := [Pred x : T | E1 /\ E2 ] + (x name, only parsing) : fun_scope. + +(* coercion for Simpl_Pred *) + +Module PredOfSimpl. +Definition Coerce T (sp : Simpl_Pred T) : Pred T := fun_of_simpl sp. +End PredOfSimpl. +Notation Pred_Of_Simpl := PredOfSimpl.Coerce. +Coercion Pred_Of_Simpl : Simpl_Pred >-> Pred. +Canonical SimplPredType T := PropPredType (@Pred_Of_Simpl T). + +Module Type PropPredSortOfSimplSignature. +Parameter Coerce : forall T, Simpl_Pred T -> {Pred T}. +End PropPredSortOfSimplSignature. +Module PropDeclarePredSortOfSimpl + (PropPredSortOfSimpl : PropPredSortOfSimplSignature). +Coercion PropPredSortOfSimpl.Coerce : Simpl_Pred >-> Pred_Sort. +End PropDeclarePredSortOfSimpl. +Module Export PropPredSortOfSimplCoercion := + PropDeclarePredSortOfSimpl PredOfSimpl. -Structure PredType : Type := PropPredType { - Pred_Sort :> Type; - toPred : Pred_Sort -> Pred T; - _ : {mem | isMem toPred mem}}. +Definition PredArgType := Type. +Bind Scope type_scope with PredArgType. +Identity Coercion Sort_Of_PredArgType : PredArgType >-> Sortclass. +Coercion Pred_Of_ArgType (T : PredArgType) : Simpl_Pred T := PredT. +Notation "{ :: T }" := (T%type : PredArgType) : type_scope. -Definition mkPredType pT toP := PropPredType (exist (@isMem pT toP) _ (erefl _)). +(* prop relations *) -(* 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). +Definition Rel T := T -> Pred T. +Identity Coercion Fun_Of_Rel : Rel >-> Funclass. -End Predicates. +Definition eqRel {T} (R1 R2 : Rel T) := forall x y, R1 x y <-> R2 x y. +Definition subRel {T} (R1 R2 : Rel T) := forall x y : T, R1 x y -> R2 x y. -Arguments Pred0 {T}. -Arguments PredT {T}. -Prenex Implicits Pred0 PredT PredI PredU PredC PredD Preim. +Notation "A <~2> B" := (eqRel A B) + (at level 70, no associativity) : rel_scope. +Notation "A ~2> B" := (subRel A B) + (at level 70, no associativity) : rel_scope. -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. +(* composing relation and function *) -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 name, format "[ 'Pred' x | E ]") : fun_scope. -Notation "[ 'Pred' x : T | E ]" := (SimplPred (fun x : T => E)) - (at level 0, x name, only parsing) : fun_scope. -Notation "[ 'Pred' x y | E ]" := (SimplPred (fun t => let: (x, y) := t in E)) - (at level 0, x name, y name, 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 name, y name, only parsing) : fun_scope. - -Definition clone_Pred T U := - fun pT & @Pred_Sort T pT -> U => - fun toP (pT' := @PropPredType T U toP) & phant_id pT' pT => pT'. -Notation "[ 'PredType' 'of' T ]" := (@clone_Pred _ T _ id _ id) : form_scope. +Definition Rel_app A B (D : Rel A) (f : B -> A) : Rel B := + fun x y => D (f x) (f y). -(* -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. -*) +Arguments Rel_app {A B} D f _ _ /. +Notation "D \\o f" := (@Rel_app _ _ D f) + (at level 42, left associativity). -Notation Pred_Class := (Pred_Sort (PredPredType _)). -Coercion Sort_of_Simpl_Pred T (p : Simpl_Pred T) : Pred_Class := p : Pred T. +Lemma rel_app1 T (g : T -> T -> Prop) : g \\o id <~2> g. +Proof. by []. Qed. -Definition PredArgType := Type. -Coercion Pred_of_argType (T : PredArgType) : Simpl_Pred T := PredT. +(* simplifying relations *) -Notation "{ :: T }" := (T%type : PredArgType) - (at level 0, format "{ :: T }") : type_scope. +Definition Simpl_Rel T := T -> Simpl_Pred T. -(* 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. +Coercion Rel_Of_Simpl T (sr : Simpl_Rel T) : Rel T := fun x : T => sr x. +Arguments Rel_Of_Simpl {T} sr x /. -Prenex Implicits Mem. +Notation xRelU := (fun (r1 r2 : Rel _) x y => r1 x y \/ r2 x y). +Notation xRelPre := (fun f (r : Rel _) x y => r (f x) (f y)). -(* Membership Predicates can be used as simple ones *) -Coercion Pred_of_Mem_Pred T mp := [Pred x : T | InMem x mp]. +Definition PropSimplRel {T} (r : Rel T) : Simpl_Rel T := + fun x => PropSimplPred (r x). +Definition RelU {T} (r1 r2 : Rel T) := PropSimplRel (xRelU r1 r2). +Definition RelPre {aT rT} (f : aT -> rT) (r : Rel rT) := PropSimplRel (xRelPre f r). -(* equality and subset *) +Notation "[ 'Rel' x y | E ]" := (PropSimplRel (fun x y => E)) + (x name, y name, only parsing) : fun_scope. +Notation "[ 'Rel' x y : T | E ]" := (PropSimplRel (fun x y : T => E)) + (x name, y name, only parsing) : fun_scope. -Definition EqPredType T (pT : PredType T) (p1 p2 : pT) := - forall x : T, toPred p1 x <-> toPred p2 x. +Lemma subRelUl T (r1 r2 : Rel T) : subRel r1 (RelU r1 r2). +Proof. by left. Qed. -Definition SubPredType T (pT : PredType T) (p1 p2 : pT) := - forall x : T, toPred p1 x -> toPred p2 x. +Lemma subRelUr T (r1 r2 : Rel T) : subRel r2 (RelU r1 r2). +Proof. by right. Qed. -(* DEVCOMMENT *) -(* are these needed? *) -(* Definition EqPred T (p1 p2 : Pred T) := EqPredType p1 p2. *) -(* Definition SubPred T (p1 p2 : Pred T) := SubPredType p1 p2. *) -(* -Definition EqMem T (p1 p2 : Mem_Pred T) := EqPredType p1 p2. -Definition SubMem T (p1 p2 : Mem_Pred T) := SubPredType p1 p2. -*) +(* Variant of Simpl_Pred specialised to the membership operator. *) -(* /DEVCOMMENT *) +Variant Mem_Pred T := PropMem of Pred T. -Definition EqSimplPred T (p1 p2 : Simpl_Pred T) := EqPredType p1 p2. -Definition SubSimplPred T (p1 p2 : Simpl_Pred T) := SubPredType p1 p2. +Coercion Pred_Of_Mem {T} mp : {Pred T} := let: PropMem p := mp in [eta p]. +Canonical MemPredType T := PropPredType (@Pred_Of_Mem T). -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 In_Mem {T} (x : T) mp := Pred_Of_Mem mp x. +Definition Eq_Mem {T} mp1 mp2 := forall x : T, In_Mem x mp1 <-> In_Mem x mp2. +Definition Sub_Mem {T} mp1 mp2 := forall x : T, In_Mem x mp1 -> In_Mem x mp2. -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. +Arguments In_Mem {T} x mp : simpl never. +Global Typeclasses Opaque Eq_Mem Sub_Mem. -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. +(* +The Simpl_Of_Mem; Pred_Of_Simpl path provides a new Mem_Pred >-> Pred + coercion, but does not override the Pred_Of_Mem : Mem_Pred >-> Pred_sort + explicit coercion declaration above. +*) + +Coercion Simpl_Of_Mem {T} mp := PropSimplPred (fun x : T => In_Mem x mp). + +Lemma Sub_Refl T (mp : Mem_Pred T) : Sub_Mem mp mp. +Proof. by []. Qed. +Arguments Sub_Refl {T mp} [x] mp_x. -Notation "x \In A" := (InMem x (Mem A)) +Definition Mem T (pT : PredType T) : pT -> Mem_Pred T := + let: PropPredType _ toP := pT in fun A => PropMem [eta toP A]. +Arguments Mem {T pT} A : rename, simpl never. + +Notation "x \In A" := (In_Mem 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)) +Notation "A =p B" := (Eq_Mem (Mem A) (Mem B)) (at level 70, no associativity) : type_scope. -Notation "A <=p B" := (SubMem (Mem A) (Mem B)) +Notation "A <=p B" := (Sub_Mem (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. +(* conversions *) +Lemma eq_conv T (A B : Pred T) : (A <~> B) = (A =p B). +Proof. by []. Qed. +Lemma sub_conv T (A B : Pred T) : (A ~> B) = (A <=p B). +Proof. by []. Qed. + +Notation "[ 'Mem' A ]" := + (Pred_Of_Simpl (Simpl_Of_Mem (Mem A))) (only parsing) : fun_scope. + Notation "[ 'PredI' A & B ]" := (PredI [Mem A] [Mem B]) - (at level 0, format "[ 'PredI' A & B ]") : fun_scope. + (format "[ 'PredI' A & B ]") : fun_scope. Notation "[ 'PredU' A & B ]" := (PredU [Mem A] [Mem B]) - (at level 0, format "[ 'PredU' A & B ]") : fun_scope. + (format "[ 'PredU' A & B ]") : fun_scope. Notation "[ 'PredD' A & B ]" := (PredD [Mem A] [Mem B]) - (at level 0, format "[ 'PredD' A & B ]") : fun_scope. + (format "[ 'PredD' A & B ]") : fun_scope. Notation "[ 'PredC' A ]" := (PredC [Mem A]) - (at level 0, format "[ 'PredC' A ]") : fun_scope. + (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 name, format "[ 'Pred' x \In A ]") : fun_scope. -Notation "[ 'Pred' x \In A | E ]" := [Pred x | (x \In A) /\ E] - (at level 0, x name, 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] + (format "[ 'Preim' f 'of' A ]") : fun_scope. + +Notation "[ 'Pred' x 'In' A ]" := [Pred x | x \In A] + (x name) : fun_scope. +Notation "[ 'Pred' x 'In' A | E ]" := [Pred x | (x \In A) /\ E] + (x name) : fun_scope. +Notation "[ 'Pred' x 'In' A | E1 & E2 ]" := + [Pred x | x \In A & E1 /\ E2 ] + (x name) : fun_scope. + +(* infix notation for PredU and PredI *) +(* is sometimes more readable *) + +Notation "A \+p B" := (PredOfSimpl.Coerce (PredU A B)) + (at level 55, right associativity) : rel_scope. +Notation "A \*p B" := (PredOfSimpl.Coerce (PredI A B)) + (at level 45, right associativity) : rel_scope. + +Notation "[ 'Rel' x y 'In' A & B | E ]" := + [Rel x y | (x \In A) /\ (y \In B) /\ E] (at level 0, x name, y name, - 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)] + format "[ 'Rel' x y 'In' A & B | E ]") : fun_scope. +Notation "[ 'Rel' x y 'In' A & B ]" := [Rel x y | (x \In A) /\ (y \In B)] (at level 0, x name, y name, - format "[ 'Pred' x y \In A & B ]") : fun_scope. -Notation "[ 'Pred' x y \In A | E ]" := [Pred x y \In A & A | E] + format "[ 'Rel' x y 'In' A & B ]") : fun_scope. +Notation "[ 'Rel' x y 'In' A | E ]" := [Rel x y In A & A | E] (at level 0, x name, y name, - format "[ 'Pred' x y \In A | E ]") : fun_scope. -Notation "[ 'Pred' x y \In A ]" := [Pred x y \In A & A] + format "[ 'Rel' x y 'In' A | E ]") : fun_scope. +Notation "[ 'Rel' x y 'In' A ]" := [Rel x y In A & A] (at level 0, x name, y name, - format "[ 'Pred' x y \In A ]") : fun_scope. + format "[ 'Rel' x y 'In' A ]") : fun_scope. + +Definition Applicative_Pred T := Pred T. +Definition Collective_Pred T := Pred T. +Coercion Applicative_Pred_Of_Simpl T (sp : Simpl_Pred T) : Applicative_Pred T := + fun_of_simpl sp. +Coercion Collective_Pred_Of_Simpl T (sp : Simpl_Pred T) : Collective_Pred T := + let: SimplFun p := sp in p. + +(* Explicit simplification rules for predicate application and membership. *) +Section PredicateSimplification. +Variables T : Type. +Implicit Types (p : Pred T) (pT : PredType T) (sp : Simpl_Pred T). +Implicit Types (mp : Mem_Pred T). + +Structure Registered_Applicative_Pred p := PropRegisteredApplicativePred { + Applicative_Pred_Value :> Pred T; + _ : Applicative_Pred_Value = p +}. +Definition PropApplicativePred p := PropRegisteredApplicativePred (erefl p). +Canonical Applicative_Pred_Applicative sp := + PropApplicativePred (Applicative_Pred_Of_Simpl sp). + +Structure Manifest_Simpl_Pred p := PropManifestSimplPred { + Simpl_Pred_Value :> Simpl_Pred T; + _ : Simpl_Pred_Value = PropSimplPred p +}. +Canonical Expose_Simpl_Pred p := PropManifestSimplPred (erefl (PropSimplPred p)). + +Structure Manifest_Mem_Pred p := PropManifestMemPred { + Mem_Pred_Value :> Mem_Pred T; + _ : Mem_Pred_Value = PropMem [eta p] +}. +Canonical Expose_Mem_Pred p := PropManifestMemPred (erefl (PropMem [eta p])). + +Structure Applicative_Mem_Pred p := + PropApplicativeMemPred {Applicative_Mem_Pred_Value :> Manifest_Mem_Pred p}. +Canonical Check_Applicative_Mem_Pred p (ap : Registered_Applicative_Pred p) := + [eta @PropApplicativeMemPred ap]. + +Lemma Mem_toPred pT (pp : pT) : Mem (toPred pp) = Mem pp. +Proof. by case: pT pp. Qed. + +Lemma toPredE pT x (pp : pT) : toPred pp x = (x \In pp). +Proof. by case: pT pp. Qed. + +Lemma App_PredE x p (ap : Registered_Applicative_Pred p) : ap x = (x \In p). +Proof. by case: ap=>ap /= ->. Qed. + +Lemma In_Applicative x p (amp : Applicative_Mem_Pred p) : In_Mem x amp = p x. +Proof. by case: amp; case=>amp /= ->. Qed. + +Lemma In_Collective x p (msp : Manifest_Simpl_Pred p) : + (x \In Collective_Pred_Of_Simpl msp) = p x. +Proof. by case: msp=>msp pf; rewrite -toPredE /= pf. Qed. + +Lemma In_Simpl x p (msp : Manifest_Simpl_Pred p) : + In_Mem x (PropMem [eta Pred_Of_Simpl msp]) = p x. +Proof. by case: msp=>msp /= ->. Qed. + +Lemma Unfold_In x p : (x \In ([eta p] : Pred T)) = p x. +Proof. by []. Qed. -Section Simplifications. -Variables (T : Type) (pT : PredType T). +Lemma Simpl_PredE p : PropSimplPred p =1 p. +Proof. by []. Qed. -Lemma Mem_toPred : forall (p : pT), Mem (toPred p) = Mem p. -Proof. by rewrite /Mem; case: pT => T1 app1 [mem1 /= ->]. Qed. +Lemma Mem_Simpl sp : Mem sp = sp :> Pred T. +Proof. by []. Qed. -Lemma toPredE x (p : pT) : toPred p x = (x \In p). -Proof. by rewrite -Mem_toPred. Qed. +Definition MemE := Mem_Simpl. +Lemma Mem_Mem mp : + (Mem mp = mp) * (Mem (mp : Simpl_Pred T) = mp) * (Mem (mp : Pred T) = mp). +Proof. by case: mp. Qed. -Lemma In_Simpl x (p : Simpl_Pred T) : (x \In p) = p x. -Proof. by []. Qed. +End PredicateSimplification. -Lemma Simpl_PredE (p : Pred T) : p <~> [Pred x | p x]. -Proof. by []. Qed. +(* Qualifiers and keyed predicates. *) -(* DEVCOMMENT *) -(* is this needed? *) -(* Definition InE := (In_Simpl, Simpl_PredE). (* to be extended *) *) -(* /DEVCOMMENT *) +Variant Qualifier (q : nat) T := PropQualifier of {Pred T}. -Lemma Mem_Simpl (p : Simpl_Pred T) : Mem p = p :> Pred T. +Coercion Has_Quality n T (q : Qualifier n T) : {Pred T} := + fun x => let: PropQualifier p := q in p x. +Arguments Has_Quality n {T}. + +Lemma QualifE n T p x : (x \In @PropQualifier n T p) = p x. Proof. by []. Qed. -Definition MemE := Mem_Simpl. (* could be extended *) +Notation "x \Is A" := (x \In Has_Quality 0 A) + (at level 0, only parsing) : fun_scope. +Notation "x \Is A" := (x \In Has_Quality 0 A) + (at level 0, only printing) : fun_scope. +Notation "x \Is 'a' A" := (x \In Has_Quality 1 A) + (at level 0, only parsing) : fun_scope. +Notation "x \Is 'a' A" := (x \In Has_Quality 1 A) + (at level 0, only printing) : fun_scope. +Notation "x \Is 'an' A" := (x \In Has_Quality 2 A) + (at level 0, only parsing) : fun_scope. +Notation "x \Is 'an' A" := (x \In Has_Quality 2 A) + (at level 0, only printing) : fun_scope. +Notation "x \Isn't A" := (x \Notin Has_Quality 0 A) + (at level 0) : fun_scope. +Notation "x \Isn't 'a' A" := (x \Notin Has_Quality 1 A) + (at level 0) : fun_scope. +Notation "x \Isn't 'an' A" := (x \Notin Has_Quality 2 A) + (at level 0) : fun_scope. +Notation "[ 'Qualify' x | P ]" := (PropQualifier 0 (fun x => P)) : form_scope. +Notation "[ 'Qualify' x : T | P ]" := + (PropQualifier 0 (fun x : T => P)) (only parsing) : form_scope. +Notation "[ 'Qualify' 'a' x | P ]" := + (PropQualifier 1 (fun x => P)) : form_scope. +Notation "[ 'Qualify' 'a' x : T | P ]" := + (PropQualifier 1 (fun x : T => P)) (only parsing) : form_scope. +Notation "[ 'Qualify' 'an' x | P ]" := + (PropQualifier 2 (fun x => P)) : form_scope. +Notation "[ 'Qualify' 'an' x : T | P ]" := + (PropQualifier 2 (fun x : T => P)) (only parsing) : form_scope. + +(* Keyed predicates: support for property-bearing predicate interfaces. *) + +Section KeyPred. +Variable T : Type. +Variant Pred_Key (p : {Pred T}) : Prop := PropDefaultPredKey. -Lemma Mem_Mem (p : pT) : (Mem (Mem p) = Mem p) * (Mem [Mem p] = Mem p). -Proof. by rewrite -Mem_toPred. Qed. +Variable p : {Pred T}. +Structure Keyed_Pred (k : Pred_Key p) := + PropPackKeyedPred {Unkey_Pred :> {Pred T}; _ : Unkey_Pred =p p}. -End Simplifications. +Variable k : Pred_Key p. +Definition PropKeyedPred := @PropPackKeyedPred k p (fun x => iff_refl _). -(* We also add subrelation and relation equality *) -(* for non-collective binary relations *) +Variable k_p : Keyed_Pred k. +Lemma Keyed_PredE : k_p =p p. +Proof. by case: k_p. Qed. -(* Not a definition to avoid universe inconsistencies. *) -(* DEVCOMMENT *) -(* Universe Polymorphism does not quite work yet *) -(* /DEVCOMMENT *) -Local Notation Rel A := (A -> A -> Prop). +Canonical Keyed_Mem := + @PropPackKeyedPred k (Pred_Of_Mem (Mem k_p)) Keyed_PredE. +Canonical Leyed_Mem_Simpl := + @PropPackKeyedPred k (Pred_Of_Simpl (Mem k_p)) Keyed_PredE. -Definition eqrel {A} (R1 R2 : Rel A) := forall x y, R1 x y <-> R2 x y. -Definition subrel {A} (R1 R2 : Rel A) := forall x y, R1 x y -> R2 x y. +End KeyPred. -Notation "A <~2> B" := (eqrel A B) - (at level 70, no associativity) : rel_scope. -Notation "A ~2> B" := (subrel A B) - (at level 70, no associativity) : rel_scope. +Local Notation In_Unkey x S := (x \In @Unkey_Pred _ S _ _) (only parsing). +Notation "x \In S" := (In_Unkey x S) (only printing) : fun_scope. -Section TransferLemmas. -Variables (A : Type) (R1 R2 : Rel A). -Lemma eqrelI : (forall x y, R1 x y <-> R2 x y) -> R1 <~2> R2. Proof. by []. Qed. -Lemma eqrelE : R1 <~2> R2 -> forall x y, R1 x y <-> R2 x y. Proof. by []. Qed. -Lemma subrelI : (forall x y, R1 x y -> R2 x y) -> R1 ~2> R2. Proof. by []. Qed. -Lemma subrelE : R1 ~2> R2 -> forall x y, R1 x y -> R2 x y. Proof. by []. Qed. -End TransferLemmas. +Section KeyedQualifier. +Variables (T : Type) (n : nat) (q : Qualifier n T). -(* composing relation and function *) +Structure Keyed_Qualifier (k : Pred_Key q) := + PropPackKeyedQualifier {Unkey_Qualifier; _ : Unkey_Qualifier = q}. +Definition PropKeyedQualifier k := PropPackKeyedQualifier k (erefl q). +Variables (k : Pred_Key q) (k_q : Keyed_Qualifier k). +Fact Keyed_Qualifier_Subproof : Unkey_Qualifier k_q =p q. +Proof. by case: k_q=>kq pf x /=; rewrite pf. Qed. +Canonical Keyed_Qualifier_Keyed := PropPackKeyedPred k Keyed_Qualifier_Subproof. -Definition Rel_app A B (D : Rel A) (f : B -> A) : Rel B := - fun x y => D (f x) (f y). +End KeyedQualifier. -Arguments Rel_app {A B} D f _ _ /. -Notation "D \\o f" := (@Rel_app _ _ D f) - (at level 42, left associativity). +Notation "x \Is A" := + (In_Unkey x (Has_Quality 0 A)) (only printing) : fun_scope. +Notation "x \Is 'a' A" := + (In_Unkey x (Has_Quality 1 A)) (only printing) : fun_scope. +Notation "x \Is 'an' A" := + (In_Unkey x (Has_Quality 2 A)) (only printing) : fun_scope. -Lemma rel_app1 T (g : T -> T -> Prop) : g \\o id <~2> g. Proof. by []. Qed. +Module DefaultKeying. +Canonical Default_Keyed_Pred T p := PropKeyedPred (@PropDefaultPredKey T p). +Canonical Default_Keyed_Qualifier T n (q : Qualifier n T) := + PropKeyedQualifier (PropDefaultPredKey q). +End DefaultKeying. -(* the same for decidable relations *) - -Definition rel_app U V (S : rel V) f : rel U := - fun x y => S (f x) (f y). -Arguments rel_app {U V} S f _ _ /. (**************************************) (* Definitions and lemmas for setoids *) @@ -330,101 +473,81 @@ Arguments rel_app {U V} S f _ _ /. (* Declaration of relations *) -Section EqPredType. -Variables (T : Type) (pT : PredType T). -Lemma EqPredType_refl (r : pT) : EqPredType r r. Proof. by []. Qed. -Lemma EqPredType_sym (r1 r2 : pT) : EqPredType r1 r2 -> EqPredType r2 r1. +Section EqPred. +Context {T : Type} {pT : PredType T}. +Lemma EqPred_refl (r : pT) : Eq_Pred r r. Proof. by []. Qed. +Lemma EqPred_sym (r1 r2 : pT) : Eq_Pred r1 r2 -> Eq_Pred 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. +Lemma EqPred_trans' (r1 r2 r3 : pT) : + Eq_Pred r1 r2 -> Eq_Pred r2 r3 -> Eq_Pred r1 r3. Proof. by move=>H1 H2 x; split; [move/H1; move/H2 | move/H2; move/H1]. Qed. -Definition EqPredType_trans r2 r1 r3 := @EqPredType_trans' r1 r2 r3. -End EqPredType. - -#[export] Hint Resolve EqPredType_refl : core. +Definition EqPred_trans r2 r1 r3 := @EqPred_trans' r1 r2 r3. +End EqPred. -(* declare relations for all implicit coercions and canonical structures *) +#[export] Hint Resolve EqPred_refl : core. -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. +(* declare relations for all Eq_Pred and Eq_Mem *) -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 (pT : PredType T) : pT (@Eq_Pred _ pT) + reflexivity proved by (@EqPred_refl _ _) + symmetry proved by (@EqPred_sym _ _) + transitivity proved by (@EqPred_trans' _ _) as EqPred_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. - -(* DEVCOMMENT *) -(* are these needed? *) -(* -Add Parametric Relation T : (Pred T) (@EqPred _) - reflexivity proved by (@EqPredType_refl _ _) - symmetry proved by (@EqPredType_sym _ _) - transitivity proved by (@EqPredType_trans' _ _) as EqPred_rel. -*) +(* do we need to add relations for mem_pred as well? *) +Add Parametric Relation T : (Mem_Pred T) (@Eq_Mem T) + reflexivity proved by (@EqPred_refl _ _) + symmetry proved by (@EqPred_sym _ _) + transitivity proved by (@EqPred_trans' _ _) as EqMem_rel. -Section SubPredType. -Variables (T : Type) (pT : PredType T). -Lemma SubPredType_refl (r : pT) : SubPredType r r. Proof. by []. Qed. -Lemma SubPredType_trans' (r1 r2 r3 : pT) : - SubPredType r1 r2 -> SubPredType r2 r3 -> SubPredType r1 r3. +Section SubPred. +Context {T : Type} {pT : PredType T}. +Lemma SubPred_refl (r : pT) : Sub_Pred r r. Proof. by []. Qed. +Lemma SubPred_trans' (r1 r2 r3 : pT) : + Sub_Pred r1 r2 -> Sub_Pred r2 r3 -> Sub_Pred r1 r3. Proof. by move=>H1 H2 x; move/H1; move/H2. Qed. -Definition SubPredType_trans r2 r1 r3 := @SubPredType_trans' r1 r2 r3. -End SubPredType. - -#[export] Hint Resolve SubPredType_refl : core. +Definition SubPred_trans r2 r1 r3 := @SubPred_trans' r1 r2 r3. +End SubPred. -Add Parametric Relation T (pT : PredType T) : pT (@SubPredType _ pT) - reflexivity proved by (@SubPredType_refl _ _) - transitivity proved by (@SubPredType_trans' _ _) as SubPredType_rel. +#[export] Hint Resolve SubPred_refl : core. -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 (pT : PredType T) : pT (@Sub_Pred _ pT) + reflexivity proved by (@SubPred_refl _ _) + transitivity proved by (@SubPred_trans' _ _) as SubPred_rel. -Add Parametric Relation T : (Mem_Pred T) (@SubMem _) - reflexivity proved by (@SubPredType_refl _ _) - transitivity proved by (@SubPredType_trans' _ _) as SubMem_rel. - -(* DEVCOMMENT *) -(* are these needed? *) -(* -Add Parametric Relation T : (Pred T) (@SubPred _) - reflexivity proved by (@SubPredType_refl _ _) - transitivity proved by (@SubPredType_trans' _ _) as SubPred_rel. -*) -(* /DEVCOMMENT *) +Add Parametric Relation T : (Mem_Pred T) (@Sub_Mem _) + reflexivity proved by (@SubPred_refl _ _) + transitivity proved by (@SubPred_trans' _ _) as SubMem_rel. -Lemma eqrel_refl {A R} : @eqrel A R R. Proof. by []. Qed. -Lemma eqrel_sym {A R1 R2} : @eqrel A R1 R2 -> eqrel R2 R1. +Section EqRel. +Context {T : Type}. +Lemma eqRel_refl (R : Rel T) : eqRel R R. Proof. by []. Qed. +Lemma eqRel_sym {R1 R2 : Rel T} : eqRel R1 R2 -> eqRel R2 R1. Proof. by move=>H x y; rewrite H. Qed. -Lemma eqrel_trans {A R1 R2 R3} : - @eqrel A R1 R2 -> eqrel R2 R3 -> eqrel R1 R3. +Lemma eqRel_trans {R1 R2 R3 : Rel T} : + eqRel R1 R2 -> eqRel R2 R3 -> eqRel R1 R3. Proof. by move=>H1 H2 x y; rewrite H1 H2. Qed. +End EqRel. -#[export] Hint Resolve eqrel_refl : core. +#[export] Hint Resolve eqRel_refl : core. -Add Parametric Relation (A : Type) : (Rel A) (@eqrel A) - reflexivity proved by (@eqrel_refl A) - symmetry proved by (@eqrel_sym A) - transitivity proved by (@eqrel_trans A) as eqrel_eq_rel. +Add Parametric Relation (T : Type) : (Rel T) (@eqRel T) + reflexivity proved by (@eqRel_refl T) + symmetry proved by (@eqRel_sym T) + transitivity proved by (@eqRel_trans T) as eqRel_rel. -Lemma subrel_refl {A} R : @subrel A R R. Proof. by []. Qed. -Lemma subrel_trans A R1 R2 R3 : - @subrel A R1 R2 -> subrel R2 R3 -> subrel R1 R3. +Section SubRel. +Context {T : Type}. +Lemma subRel_refl (R : Rel T) : subRel R R. Proof. by []. Qed. +Lemma subRel_trans (R1 R2 R3 : Rel T) : + subRel R1 R2 -> subRel R2 R3 -> subRel R1 R3. Proof. by move=>H1 H2 x y /H1/H2. Qed. +End SubRel. -#[export] Hint Resolve subrel_refl : core. +#[export] Hint Resolve subRel_refl : core. -Add Parametric Relation (A : Type) : (Rel A) (@subrel A) - reflexivity proved by (@subrel_refl A) - transitivity proved by (@subrel_trans A) as subrel_rel. +Add Parametric Relation (T : Type) : (Rel T) (@subRel T) + reflexivity proved by (@subRel_refl T) + transitivity proved by (@subRel_trans T) as subRel_rel. (* Declaring morphisms. *) @@ -432,9 +555,11 @@ Add Parametric Relation (A : Type) : (Rel A) (@subrel A) (* Annoyingly, even the coercions must be declared *) (* /DEVCOMMENT *) +(* Add Parametric Morphism T : (@Pred_of_Simpl T) with signature @EqSimplPred _ ==> @EqPredType T (PredPredType T) as Pred_of_Simpl_morph. Proof. by []. Qed. +*) (* DEVCOMMENT *) (* Do we need other coercions? We'll discover as we go *) @@ -452,50 +577,102 @@ Proof. by move=>r1 s1 H1 r2 s2 H2; rewrite H1 H2. Qed. *) (* /DEVCOMMENT *) -Add Parametric Morphism T (pT : PredType T) : (@EqPredType T pT) with signature - @EqPredType _ _ ==> @EqPredType _ _ ==> iff as EqPredType_morph. +Add Parametric Morphism T (pT : PredType T) : (@Eq_Pred T pT) + with signature + @Eq_Pred _ _ ==> @Eq_Pred _ _ ==> iff as EqPred_morph. +Proof. by move=>r1 s1 H1 r2 s2 H2; rewrite H1 H2. Qed. + +Add Parametric Morphism T (pT : PredType T) : (@Sub_Pred T pT) + with signature + @Eq_Pred _ _ ==> @Eq_Pred _ _ ==> 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 : (@Eq_Mem T) + with signature @Eq_Mem _ ==> @Eq_Mem _ ==> iff as EqMem_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. +Add Parametric Morphism T : (@Sub_Mem T) + with signature @Eq_Mem _ ==> @Eq_Mem _ ==> iff as SubMem_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. +Add Parametric Morphism T : (@Eq_Pred T _) + with signature @Eq_Mem _ ==> @Eq_Mem _ ==> iff as EqPred_Mem_morph. +Proof. +move=>x y H x0 y0 H0; split=>X p /=; move: (H p) (H0 p)=>{}H {}H0; +rewrite /In_Mem /= in H H0. +- by rewrite -H -H0; apply/X. +by rewrite H H0; apply/X. +Qed. + +Add Parametric Morphism T : (@In_Mem T) with signature + @eq _ ==> @Eq_Mem _ ==> 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. +(* to rewrite under [Mem], must enable rewriting under *) +(* Mem, Simpl_Of_Mem and Pred_Of_Simpl *) + +(* rewriting by Eq_Pred under Mem *) +Add Parametric Morphism T (pT : PredType T) : (@Mem T pT) + with signature @Eq_Pred _ _ ==> @Eq_Pred _ _ as Mem_morph. +Proof. by case: pT=>pT F /= x y; apply. Qed. + +(* and under Eq_Mem \o Mem *) +Add Parametric Morphism T (pT : PredType T) : (@Mem T pT) + with signature @Eq_Pred _ _ ==> @Eq_Mem _ as Mem_morph'. +Proof. by move=>x y H p; rewrite /Eq_Pred -!toPredE in H *; apply: H. Qed. + +(* rewriting under Simpl_Of_Mem *) +Add Parametric Morphism T : (@Simpl_Of_Mem T) + with signature @Eq_Pred _ _ ==> @Eq_Pred _ _ as SimplOfMem_morph. +Proof. by []. Qed. + +(* rewriting under Pred_Of_Simpl *) +Add Parametric Morphism T : (@Pred_Of_Simpl T) + with signature @Eq_Pred _ _ ==> @Eq_Pred _ _ as Pred_Of_Simpl_morph. +Proof. by []. Qed. + +(* rewriting by Eq_Mem under Eq_Pred *) +Add Parametric Morphism T : (@Mem T _) + with signature @Eq_Mem _ ==> @Eq_Pred _ _ as Mem_Morph. +Proof. by []. Qed. + +Add Parametric Morphism T : (@Simpl_Of_Mem T) + with signature @Eq_Mem _ ==> @Eq_Pred _ _ as SimplOfMem_morph'. +Proof. by []. Qed. + +(* rewriting under Pred_Of_Simpl *) +Add Parametric Morphism T : (@Pred_Of_Simpl T) + with signature @Eq_Mem _ ==> @Eq_Pred _ _ as Pred_Of_Simpl_morph'. +Proof. by []. Qed. Add Parametric Morphism T : (@PredU T) with signature - @EqPredType _ _ ==> @EqPredType _ _ ==> @EqPredType _ _ as predU_morph. + @Eq_Pred T _ ==> @Eq_Pred T _ ==> @Eq_Pred T _ 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. + @Eq_Pred _ _ ==> @Eq_Pred _ _ ==> @Eq_Pred _ _ 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. + @Eq_Pred _ _ ==> @Eq_Pred _ _ as predC_morph. Proof. by move=>r s H x; split=>H1; apply/H. Qed. -Add Parametric Morphism A : (@subrel A) with signature - eqrel ==> eqrel ==> iff as subrel_eq. +Add Parametric Morphism A : (@subRel A) with signature + eqRel ==> eqRel ==> iff as subrel_eq. Proof. by move=>x y H x1 y1 H1; split=>X a b /H/X/H1. Qed. Add Parametric Morphism A B : (@Rel_app A B) with signature - eqrel ==> eq ==> eqrel as Relapp_eq. + eqRel ==> eq ==> eqRel as Relapp_eq. Proof. by move=>x y H f s s'; split=>/H. Qed. Add Parametric Morphism A B : (@Rel_app A B) with signature - subrel ==> eq ==> subrel as Relapp_sub. + subRel ==> eq ==> subRel as Relapp_sub. Proof. by move=>x y H f s s' /H. Qed. (* Can we declare relation application as morphism? *) @@ -517,11 +694,11 @@ Section RelLaws. Variable T : Type. Implicit Type r : Pred T. -Lemma orrI r : r \+p r <~> r. +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 orrC r1 r2 : r1 \+p r2 <~> r2 \+p r1. +Proof. by split; (case=>/=; [right|left]). Qed. Lemma orr0 r : r \+p Pred0 <~> r. Proof. by move=>x; split; [case | left]. Qed. @@ -529,25 +706,36 @@ 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 orrCA r1 r2 r3 : + r1 \+p r2 \+p r3 <~> r2 \+p r1 \+p r3. simpl. +Proof. by move=>x /=; intuition. Qed. -Lemma orrAC r1 r2 r3 : (r1 \+p r2) \+p r3 <~> (r1 \+p r3) \+p r2. -Proof. by move=>?; split=>/=; intuition. Qed. +Lemma orrAC r1 r2 r3 : + (r1 \+p r2) \+p r3 <~> (r1 \+p r3) \+p r2. +Proof. by move=>x /=; 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. +Lemma orrA r1 r2 r3 : + (r1 \+p r2) \+p r3 <~> r1 \+p r2 \+p r3. +Proof. by rewrite (orrC r2) orrCA (orrC r3). Qed. (* absorption *) -Lemma orrAb r1 r2 : r1 <~> r1 \+p r2 <-> r2 ~> r1. +(* DEVCOMMENT: when lemmas use sub_rel, it's usually better *) +(* to use the collective form sub_mem *) +(* Typically, with sub_rel, there will be lemma application *) +(* that will turn expression of the from x \In A into A x *) +(* and the latter immediately prevents setoid rewriting *) +(* /DEVCOMMENT *) +Lemma orrAb r1 r2 : + r1 <~> r1 \+p r2 <-> + r2 <=p r1. Proof. -split; first by move=>-> x /=; auto. +split; first by move=>-> x; rewrite -!toPredE /=; 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. +Lemma sub_orl r1 r2 : r1 <=p r1 \+p r2. Proof. by left. Qed. +Lemma sub_orr r1 r2 : r2 <=p r1 \+p r2. Proof. by right. Qed. End RelLaws. Section SubMemLaws. @@ -557,34 +745,40 @@ 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. +Lemma subp_asym p1 p2 : p1 <=p p2 -> p2 <=p p1 -> p1 <~> 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. +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. +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. +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. +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. +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. +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. @@ -599,12 +793,11 @@ Fixpoint Mem_Seq (s : seq T) := 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. +Coercion Pred_Of_Seq (s : EqSeq_Class) : {Pred T} := [eta Mem_Seq s]. +Canonical Structure seq_PredType := + PropPredType (Pred_Of_Seq : seq T -> Pred T). +(* The line below makes mem_seq a canonical instance of topred. *) +Canonical Structure Mem_Seq_PredType := PropPredType Mem_Seq. Lemma In_cons y s x : (x \In y :: s) <-> (x = y) \/ (x \In s). Proof. by []. Qed. @@ -615,13 +808,9 @@ 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). - -(* DEVCOMMENT *) -(* I also wanted to add Simpl_PredE, but setoid rewrite returns an error *) -(* and instead of trying the other rules in the tuple, it just stops *) -(* This is ridiculuous *) -(* /DEVCOMMENT *) +Definition InE := + (Mem_Seq1, In_cons, + (In_Applicative, In_Simpl, Simpl_PredE)). Lemma Mem_cat x : forall s1 s2, (x \In s1 ++ s2) <-> x \In s1 \/ x \In s2. Proof. @@ -869,19 +1058,25 @@ Notation "[ 'Image' E | i : T <- s & C ]" := [Image E | i : T <- [PredI s & C]] (at level 0, E at level 99, i name, 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. +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. + 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. + injective f -> + Image f X1 <~> Image f X2 -> + X1 <~> 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]. + Image f (PredU X1 X2) <~> PredU (Image f X1) (Image f X2). Proof. move=>x; split. - by case=>y -> [H|H]; [left | right]; apply: Image_mem. @@ -889,14 +1084,15 @@ 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. + Image f1 (Image f2 X) <~> 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. + f1 =1 f2 -> + Image f1 X <~> Image f2 X. Proof. by move=>H x; split; case=>a ->; exists a. Qed. (********************************) @@ -921,14 +1117,14 @@ by case=>y' E; exists y'; rewrite /Rel_app -E. Qed. Add Parametric Morphism A B : (@rel_on_func A B) with signature - eq ==> @EqPredType _ _ ==> eqrel ==> iff as relon_eq. + eq ==> @Eq_Pred _ _ ==> eqRel ==> iff as relon_eq. Proof. move=>f p1 p2 H r1 r2 E; split=>X; by apply/onE=>x y /H /= P /E /(X _ _ P) [y'][-> /H]; eauto. Qed. Add Parametric Morphism A B : (@rel_on_func A B) with signature - eq ==> @EqPredType _ _ ==> subrel ==> flip impl as relon_sub. + eq ==> @Eq_Pred _ _ ==> subRel ==> flip impl as relon_sub. Proof. move=>f p1 p2 H r1 r2 E X. by apply/onE=>x y /H /= P /E /(X _ _ P) [y'][-> /H]; eauto. @@ -1211,11 +1407,11 @@ Lemma iter_pair {T} {g : unit*T -> unit*T -> Prop} : Proof. by apply: iter_eqf=>[x y []|x [[y]]] //; exists y. Qed. Add Parametric Morphism A : (@iter A) with signature - eqrel ==> eqrel as iter_eq'. + eqRel ==> eqRel as iter_eq'. Proof. by move=>x y H x1 y1; rewrite (iter_eq H _). Qed. Add Parametric Morphism A : (@iter A) with signature - subrel ==> subrel as iter_sub'. + subRel ==> subRel as iter_sub'. Proof. by move=>x y H x1 y1; apply: iter_sub. Qed. Lemma iter_pres' A (g : Rel A) n C : preserved C g -> preserved C (iter' g n). @@ -1474,27 +1670,27 @@ Arguments pre_rel {A} P R _ _ /. Arguments post_rel {A} R Q _ _ /. Lemma union_eq A (x1 y1 x2 y2 : Rel A) : - eqrel x1 y1 -> eqrel x2 y2 -> - eqrel (Union_rel x1 x2) (Union_rel y1 y2). -Proof. by rewrite /eqrel/Union_rel=>E1 E2 x y; rewrite E1 E2. Qed. + eqRel x1 y1 -> eqRel x2 y2 -> + eqRel (Union_rel x1 x2) (Union_rel y1 y2). +Proof. by rewrite /eqRel/Union_rel=>E1 E2 x y; rewrite E1 E2. Qed. Arguments union_eq {A x1 y1 x2 y2}. Prenex Implicits union_eq. Add Parametric Morphism A : (@Union_rel A) with signature - eqrel ==> eqrel ==> eqrel as union_eq'. + eqRel ==> eqRel ==> eqRel as union_eq'. Proof. by move=>*; apply: union_eq. Qed. Lemma union_sub A (x1 y1 x2 y2 : Rel A) : - subrel x1 y1 -> subrel x2 y2 -> - subrel (Union_rel x1 x2) (Union_rel y1 y2). + subRel x1 y1 -> subRel x2 y2 -> + subRel (Union_rel x1 x2) (Union_rel y1 y2). Proof. by move=>E1 E2 x y [/E1|/E2]; [left | right]. Qed. Arguments union_sub {A x1 y1 x2 y2}. Prenex Implicits union_sub. Add Parametric Morphism A : (@Union_rel A) with signature - subrel ==> subrel ==> subrel as union_sub1. + subRel ==> subRel ==> subRel as union_sub1. Proof. by move=>*; apply: union_sub. Qed. Lemma subrelL A (x y : Rel A) : x ~2> Union_rel x y. @@ -1504,7 +1700,7 @@ Lemma subrelR A (x y : Rel A) : y ~2> Union_rel x y. Proof. by right. Qed. #[export] -Instance eqrel_subrel A : subrelation (@eqrel A) (@subrel A). +Instance eqrel_subrel A : subrelation (@eqRel A) (@subRel A). Proof. by move=>x y H ?? /H. Qed. Lemma union_app A B D1 D2 (f : B -> A) : diff --git a/core/prelude.v b/core/prelude.v index b605d0d..8d45ce5 100644 --- a/core/prelude.v +++ b/core/prelude.v @@ -18,7 +18,7 @@ limitations under the License. From HB Require Import structures. From Coq Require Import ssreflect ssrfun Eqdep. -From mathcomp Require Import ssrbool ssrnat seq eqtype choice. +From mathcomp Require Import ssrbool ssrnat seq eqtype choice path. From mathcomp Require Import fintype finset finfun tuple perm fingroup. From pcm Require Import options axioms. @@ -225,6 +225,11 @@ Notation "f1 \** f2 " := (pmorphism f1 f2) Lemma prod_feta (A B : Type) : @idfun (A * B) = fst \** snd. Proof. by apply: fext=>x; rewrite /pmorphism -prod_eta. Qed. +(* composing relation and function *) +Definition rel_app A B (S : rel A) f : rel B := + fun x y => S (f x) (f y). +Arguments rel_app {A B} S f _ _ /. + Reserved Notation "[ ** f1 & f2 ]" (at level 0). Reserved Notation "[ ** f1 , f2 & f3 ]" (at level 0, format "'[hv' [ ** '[' f1 , '/' f2 ']' '/ ' & f3 ] ']'"). @@ -657,10 +662,192 @@ Proof. by elim: n. Qed. Lemma neqnS n : n != n.+1. Proof. by elim: n. Qed. +Lemma subn_eq0P m n : reflect (m - n = 0) (m <= n). +Proof. by rewrite -subn_eq0; apply/eqP. Qed. + (**************************************) (* Inhabited (non-empty) finite types *) (**************************************) +(* if a type is non-empty, one can pick an element *) + +Lemma inhabF {T : finType} : + 0 < #|T| -> ~ @predT T =1 xpred0. +Proof. by case/card_gt0P=>x _ /(_ x). Qed. + +Definition inhab0 {T : finType} (pf : 0 < #|T|) : T := + match pickP predT with + | Pick x _ => x + | Nopick qf => False_rect T (inhabF pf qf) + end. + +(* variant of nth that needs no seed value *) +Definition ith {T : finType} i (pf : i < #|T|) : T := + nth (inhab0 (leq_ltn_trans (leq0n i) pf)) (enum T) i. + +Arguments ith {T}. + +(* dually, variant of index that doesn't overflow *) +Definition indx {T : finType} (x : T) := index x (enum T). + +(* lemmas associated with ith/indx *) + +Lemma indx_card {T : finType} (i : T) : indx i < #|T|. +Proof. by rewrite cardT index_mem mem_enum. Qed. + +Lemma indx_ith {T : finType} i (pf : i < #|T|) : + indx (ith i pf) = i. +Proof. by rewrite /indx/ith index_uniq ?enum_uniq -?cardE. Qed. + +Lemma ith_indx {T : finType} (i : T) (pf : indx i < #|T|) : + ith (indx i) pf = i. +Proof. by rewrite /ith/indx nth_index // mem_enum. Qed. + +Lemma indx_inj {T} : injective (@indx T). +Proof. +rewrite /indx=>x1 x2. +have [] : x1 \in enum T /\ x2 \in enum T by rewrite !mem_enum. +elim: (enum T)=>[|x xs IH] //=; rewrite !inE !(eq_sym x). +case: (x1 =P x)=>[<-|] _ /=; first by case: (x2 =P x1). +by case: (x2 =P x)=>//= _ X1 X2 []; apply: IH X1 X2. +Qed. + +Lemma ith_inj {T : finType} i1 i2 (pf1 : i1 < #|T|) (pf2 : i2 < #|T|) : + ith i1 pf1 = ith i2 pf2 -> + i1 = i2. +Proof. +rewrite /ith; move: (inhab0 _) (inhab0 _)=>o1 o2 H. +rewrite cardE in pf1 pf2. +elim: (enum T) i1 i2 pf1 pf2 o1 o2 H (enum_uniq T)=>[|x xs IH] //=. +case=>[|i1][|i2] //= pf1 pf2 o1 o2. +- by move=>->; rewrite mem_nth. +- by move=><-; rewrite mem_nth. +by move=>H /andP [_ U]; rewrite (IH _ _ pf1 pf2 o1 o2). +Qed. + +Lemma indx_injE {T : finType} s i (pf : i < #|T|) : + (s == ith i pf) = (indx s == i). +Proof. +apply/eqP/eqP=>[->|E]; first by rewrite indx_ith. +by rewrite -E in pf *; rewrite ith_indx. +Qed. + +Lemma map_indx {T : finType} : map indx (enum T) = iota 0 #|T|. +Proof. +rewrite cardE /indx. +elim: (enum T) (enum_uniq T)=>[|x xs IH] //. +set f := index^~(x :: xs)=>/= /andP [H1 H2]. +rewrite {1}/f /= eqxx; congr (0 :: _). +case: (eq_in_map f (fun x=>(index x xs).+1) xs)=>E _. +rewrite E; last first. +- by move=>z R; rewrite /f /=; case: (x =P z) R H1=>//= ->->. +by rewrite -add1n iotaDl -IH // -map_comp. +Qed. + +Lemma take_enum {T : finType} x i : + x \in take i (enum T) = (indx x < i). +Proof. +pose f x := indx x. +rewrite -(mem_map indx_inj) map_take map_indx take_iota. +case: (leqP i #|T|)=>H; rewrite mem_iota /=; first by rewrite add0n. +rewrite add0n cardE index_mem mem_enum inE /=; apply: sym_eq. +apply: (@ltn_trans #|T|) H. +by rewrite cardE index_mem mem_enum. +Qed. + +Lemma drop_enum {T : finType} x i : + x \in drop i (enum T) = (i <= indx x). +Proof. +pose f x := index x (enum T). +rewrite -(mem_map indx_inj) map_drop map_indx drop_iota mem_iota add0n. +case H : (i <= indx x)=>//=; rewrite subnKC ?indx_card //. +by apply/(leq_trans H)/ltnW/indx_card. +Qed. + +Lemma take_enum_filter {T : finType} k : + filter (preim indx [pred x | x < k]) (enum T) = + take k (enum T). +Proof. +apply: (inj_map indx_inj). +rewrite map_take map_indx -filter_map map_indx. +apply: (sorted_eq leq_trans anti_leq). +- by apply/(sorted_filter leq_trans)/iota_sorted. +- by apply/take_sorted/iota_sorted. +apply: uniq_perm. +- by rewrite filter_uniq // iota_uniq. +- by rewrite take_uniq // iota_uniq. +move=>x; rewrite mem_filter take_iota cardE. +case: (leqP k (size _))=>H1; rewrite !mem_iota !add0n /=. +- by case: (ltnP x k)=>H2 //=; apply: leq_trans H1. +case: (ltnP x (size _))=>H2 //=; last by rewrite andbF. +by rewrite andbT; apply: ltn_trans H1. +Qed. + +Lemma drop_enum_filter {T : finType} k : + filter (preim indx [pred x | x >= k]) (enum T) = + drop k (enum T). +Proof. +apply: (inj_map indx_inj). +rewrite map_drop map_indx -filter_map map_indx. +apply: (sorted_eq leq_trans anti_leq). +- by apply/(sorted_filter leq_trans)/iota_sorted. +- by apply/drop_sorted/iota_sorted. +apply: uniq_perm. +- by rewrite filter_uniq // iota_uniq. +- by rewrite drop_uniq // iota_uniq. +move=>x; rewrite mem_filter drop_iota cardE !add0n inE. +case: (leqP k (size (enum T)))=>H1; rewrite !mem_iota add0n. +- by rewrite subnKC. +case: subn_eq0P (ltnW H1)=>// -> _; rewrite addn0 leq0n /=. +by case: (leqP k x)=>//= K1; case: ltngtP (leq_trans H1 K1). +Qed. + +Lemma enum_split {T : finType} k : + enum T = take (indx k) (enum T) ++ k :: drop (indx k).+1 (enum T). +Proof. +rewrite -{2}(@nth_index T k k (enum T)) ?mem_enum //. +by rewrite -drop_nth ?index_mem ?mem_enum // cat_take_drop. +Qed. + +Lemma takeord {I : finType} T k x (f : {ffun I -> T}) : + take (indx k) (fgraph [ffun y => [eta f with k |-> x] y]) = + take (indx k) (fgraph f). +Proof. +set f' := (finfun _). +suff E: {in take (indx k) (enum I), f =1 f'}. +- by rewrite !fgraph_codom /= !codomE -2!map_take; move/eq_in_map: E. +move: (enum_uniq I); rewrite {1}(enum_split k) cat_uniq /= =>H4. +move=>y H5; rewrite /f' /= !ffunE /=; case: eqP H5 H4=>// -> ->. +by rewrite andbF. +Qed. + +Lemma dropord {I : finType} T k x (f : {ffun I -> T}) : + drop (indx k).+1 (fgraph [ffun y => [eta f with k |->x] y]) = + drop (indx k).+1 (fgraph f). +Proof. +set f' := (finfun _). +suff E: {in drop (indx k).+1 (enum I), f =1 f'}. +- by rewrite !fgraph_codom /= !codomE -2!map_drop; move/eq_in_map: E. +move: (enum_uniq I); rewrite {1}(enum_split k) cat_uniq /= => H4. +move=>y H5; rewrite /f' /= !ffunE /=; case: eqP H5 H4=>// -> ->. +by rewrite !andbF. +Qed. + +Lemma size_fgraph {I : finType} T1 T2 + (r1 : {ffun I -> T1}) (r2 : {ffun I -> T2}) : + size (fgraph r1) = size (fgraph r2). +Proof. by rewrite !fgraph_codom /= !codomE !size_map. Qed. + +Lemma fgraphE {I : finType} T (r1 r2 : {ffun I -> T}) : + fgraph r1 = fgraph r2 -> + r1 = r2. +Proof. +move=> eq_r12; apply/ffunP=> x. +by rewrite -[x]enum_rankK -!tnth_fgraph eq_r12. +Qed. + +(* building inhabited finite types *) + Lemma inhabits (T : finType) (t : T) : 0 < #|T|. Proof. by apply/card_gt0P; exists t. Qed. @@ -676,21 +863,14 @@ HB.mixin Record isInhabited T of Finite T := { #[short(type="ifinType")] HB.structure Definition Inhabited := {T of isInhabited T & Finite T}. -Lemma inhabF (T : ifinType) : ~ (@predT T =1 xpred0). -Proof. by case/card_gt0P: (@card_inhab T)=>x _ /(_ x). Qed. - -Definition inhab {T : ifinType} : T := - match pickP predT with - | Pick x _ => x - | Nopick pf => False_rect T (inhabF pf) - end. - HB.instance Definition _ := isInhabited.Build unit (inhabits tt). HB.instance Definition _ := isInhabited.Build bool (inhabits false). HB.instance Definition _ n := isInhabited.Build 'I_n.+1 (inhabits ord0). HB.instance Definition _ (T : finType) := isInhabited.Build (option T) (inhabits None). +Definition inhab {I : ifinType} : I := inhab0 card_inhab. + (*************************************) (* A copy of booleans with mnemonics *) (* LL and RR for working with sides *) diff --git a/examples/array.v b/examples/array.v index f92b21e..d79fb40 100644 --- a/examples/array.v +++ b/examples/array.v @@ -17,8 +17,6 @@ From pcm Require Import options axioms prelude pred. From pcm Require Import pcm unionmap heap. From htt Require Import options model heapauto. -Definition indx {I : finType} (x : I) := index x (enum I). - (*********************************) (* Arrays indexed by finite type *) (*********************************) @@ -34,13 +32,41 @@ Identity Coercion array_for_array : array_for >-> array. Notation "{ 'array' aT }" := (array_for (Phant aT)) (at level 0, format "{ 'array' '[hv' aT ']' }") : type_scope. -Module Array. +Module Type ArraySig. +Parameter shape : forall {I : finType} {T : Type}, + {array I -> T} -> {ffun I -> T} -> Pred heap. + +(* build new array with all cells initialized to x *) +Parameter new : forall {I : finType} {T : Type} (x : T), + STsep (emp, [vfun (a : {array I -> T}) h => h \In shape a [ffun => x]]). + +(* build new array with cells initialized by finite function f *) +Parameter newf : forall {I : finType} {T : Type} (f : {ffun I -> T}), + STsep (emp, [vfun a h => h \In shape a f]). + +(* free the array *) +Parameter free : forall {I : finType} {T : Type} (a : {array I -> T}), + STsep (fun i => exists f, i \In shape a f, [vfun _ : unit => emp]). + +(* read k-th cell *) +Parameter read : forall {I : finType} {T : Type} (a : {array I -> T}) (k : I), + STsep {f h} (fun i => i = h /\ i \In shape a f, + [vfun (y : T) m => m = h /\ y = f k]). + +(* write k-th cell *) +Parameter write : forall {I : finType} {T : Type} (a : {array I -> T}) k x, + STsep {f} (shape a f, + [vfun (_ : unit) h => h \In shape a (finfun [eta f with k |-> x])]). +End ArraySig. + + +Module Array : ArraySig. Section Array. -Variable (I : finType) (T : Type). +Context {I : finType} {T : Type}. Notation array := {array I -> T}. (* array is specified by finite function *) -Definition shape (a : array) (f : {ffun I -> T}) := +Definition shape (a : array) (f : {ffun I -> T}) : Pred heap := [Pred h | h = updi a (fgraph f)]. (* main methods *) @@ -55,7 +81,7 @@ Next Obligation. (* pull out ghost vars, run the program *) move=>x [] _ /= ->; step=>y; step. (* simplify *) -rewrite unitR=>H; congr updi; rewrite codomE cardE. +rewrite unitR=>_; congr updi; rewrite /= codomE cardE. by elim: (enum I)=>[|t ts] //= ->; rewrite (ffunE _ _). Qed. @@ -113,7 +139,7 @@ move=>f [] _ ->; case: fintype.pickP=>[v|] H /=. (* invoke the loop, construct g from the first value of f *) by apply: [gE]=>//=; exists [ffun => f v], nil. (* I is empty, so should be the resulting heap *) -step=>_; rewrite codom_ffun. +step; rewrite /shape /= codom_ffun. suff L: #|I| = 0 by case: (fgraph f)=>/=; rewrite L; case. by rewrite cardE; case: (enum I)=>[|x s] //; move: (H x). Qed. @@ -170,7 +196,6 @@ by rewrite (updi_split a k); step. Qed. (* writing to an array, updates the spec function with a new value *) - Program Definition write (a : array) (k : I) (x : T) : STsep {f} (shape a f, [vfun _ : unit => shape a (finfun [eta f with k |-> x])]) := @@ -191,7 +216,6 @@ End Array. (* array elements, such as hashtables *) Section Table. - Variables (I : finType) (T S : Type) (x : ptr) (Ps : T -> S -> Pred heap). diff --git a/examples/bubblesort.v b/examples/bubblesort.v index a2956de..ab192b1 100644 --- a/examples/bubblesort.v +++ b/examples/bubblesort.v @@ -19,7 +19,6 @@ From pcm Require Import seqext pcm unionmap heap. From htt Require Import options model heapauto. From htt Require Import array. -(* hack to avoid "_ *p _" notation clash *) From mathcomp Require order. Import order.Order.NatOrder order.Order.TTheory. @@ -102,7 +101,7 @@ Lemma codom1_split (f : {ffun 'I_n.+1 -> A}) (i : 'I_n) : &:(fgraph f) `]i.+1, +oo[. Proof. set i0 := Wo i; set i1 := So i. -rewrite {1}codomE (enum_split i0) /= {2}(enum_split i1) /= /heap.indx +rewrite {1}codomE (enum_split i0) /= {2}(enum_split i1) /= /indx (index_enum_ord i0) (index_enum_ord i1) drop_cat size_take size_enum_ord ltn_ord So_eq ltnn subnn /= map_cat /= map_take map_drop -codomE. by rewrite /slice.slice/= addn0 addn1 /= drop0 take_size. @@ -350,12 +349,12 @@ Next Obligation. move=>a i /= [f][] h /= E; set i0 := Wo i; set i1 := So i. do 2!apply: [stepE f, h]=>//= _ _ [->->]. case: oleqP=>H; first by step=>_; exists 1%g; split=>//; rewrite pffunE1. -apply: [stepE f ]=>//= _ _ ->. +apply: [stepE f]=>//= _ {E}h E. set fs := finfun [eta f with i0 |-> f i1]. -apply: [stepE fs]=>//= _ _ ->. -set fsw := finfun [eta fs with i1 |-> f i0]. +apply: [stepE fs]=>//= _ {E}h E. +set fsw := (finfun _) in E. step=>_; exists (swnx i); do!split=>//=. -suff: fsw = pffun (swnx i) f by move=>->. +suff: fsw = pffun (swnx i) f by move=><-. rewrite /fsw /fs /swnx; apply/ffunP=>/= x; rewrite !ffunE /= ffunE /=. case: tpermP=>[->|->|/eqP/negbTE->/eqP/negbTE->] //; rewrite eqxx //. by rewrite /i1 eq_sym (negbTE (So_WoN i)). diff --git a/examples/congmath.v b/examples/congmath.v index 7375753..a2abdd3 100644 --- a/examples/congmath.v +++ b/examples/congmath.v @@ -187,7 +187,7 @@ Definition mono_cong (R : rel_exp) : congruence R -> monotone R. Proof. by move=>[[]]. Qed. Definition closure (R : rel_exp) : rel_exp := - [Pred e1 e2 | forall C, R ~> C -> congruence C -> (e1, e2) \In C]. + [Pred e | forall C, R ~> C -> congruence C -> e \In C]. Lemma cong_clos (R : rel_exp) : congruence (closure R). Proof. @@ -259,12 +259,13 @@ Lemma closRK (R1 R2 K : rel_exp) : closure (R1 \+p K) <~> closure (R2 \+p K). Proof. by rewrite 2!(orrC _ K); apply: closKR. Qed. -Lemma closI (R : rel_exp) : R ~> closure R. +Lemma closI (R : rel_exp) : R <=p closure R. Proof. by move=>[x y] H1 T H2 H3; apply: H2. Qed. (* absorption and closure *) Lemma closKA (K R : rel_exp) : - closure R <~> closure (K \+p R) <-> K ~> closure R. + closure R <~> closure (K \+p R) <-> + K <=p closure R. Proof. rewrite -orrAb; split=>[->|]. - by rewrite orrAb; move=>[x y] H; apply: closI; left. @@ -272,20 +273,22 @@ by rewrite (orrC _ R) -clos_clos orrC=><-; rewrite clos_idemp. Qed. Lemma closAK (K R : rel_exp) : - closure R <~> closure (R \+p K) <-> K ~> closure R. + closure R <~> closure (R \+p K) <-> + K <=p closure R. Proof. by rewrite orrC closKA. Qed. Lemma clos_or (R1 R2 R : rel_exp) : - closure R1 ~> closure R -> closure R2 ~> closure R -> - closure (R1 \+p R2) ~> closure R. + closure R1 <=p closure R -> + closure R2 <=p closure R -> + closure (R1 \+p R2) <=p closure R. Proof. rewrite -!closAK !(orrC R) !clos_clos -!(orrC R) => H1 H2. by rewrite H1 -clos_clos H2 clos_clos orrAC orrA. Qed. Lemma sub_closKR (R1 R2 K : rel_exp) : - (closure R1 ~> closure R2) -> - closure (K \+p R1) ~> closure (K \+p R2). + (closure R1 <=p closure R2) -> + closure (K \+p R1) <=p closure (K \+p R2). Proof. rewrite -!closAK=>H. rewrite -(orrC (closure _)) clos_clos orrAC -orrA orrI. @@ -295,12 +298,13 @@ by rewrite orrC. Qed. Lemma sub_closRK (R1 R2 K : rel_exp) : - (closure R1 ~> closure R2) -> - closure (R1 \+p K) ~> closure (R2 \+p K). + (closure R1 <=p closure R2) -> + closure (R1 \+p K) <=p closure (R2 \+p K). Proof. by move=>H; rewrite -!(orrC K); apply: sub_closKR. Qed. Lemma sub_closI (R1 R2 : rel_exp) : - R1 ~> R2 -> closure R1 ~> closure R2. + R1 <=p R2 -> + closure R1 <=p closure R2. Proof. rewrite -orrAb -closAK -(orrC (closure _)) clos_clos => ->. by rewrite (orrC R2) -orrA orrI. @@ -309,7 +313,7 @@ Qed. (* relation that only relates the constant symbols *) (* with their image under a function *) Definition graph (f : symb -> exp) : rel_exp := - [Pred e1 e2 | (e1, e2) \in image (fun s => (const s, f s)) predT]. + [Pred e | e \in image (fun s => (const s, f s)) predT]. (* Equations in canonical form, as required by the congruence closure *) (* algorithm. An equation is in canonical form if it is an equation *) @@ -338,10 +342,9 @@ HB.instance Definition _ := hasDecEq.Build Eq eqEqP. (* interpreting equations as relations *) Definition eq2rel (eq : Eq) : rel_exp := match eq with - | simp_eq c1 c2 => - [Pred x y | x = const c1 /\ y = const c2] + | simp_eq c1 c2 => Pred1 (const c1, const c2) | comp_eq c c1 c2 => - [Pred x y | x = const c /\ y = app (const c1) (const c2)] + Pred1 (const c, app (const c1) (const c2)) end. Fixpoint eqs2rel (eqs : seq Eq) : rel_exp := @@ -448,19 +451,18 @@ Definition rep2rel D := graph (fun x => const (rep D x)). Lemma clos_rep D R a : (const a, const (rep D a)) \In closure (rep2rel D \+p R). Proof. -apply: (@closI _ (const a, const (rep D a))); apply: sub_orl; rewrite /= /rep2rel /graph /=. -by rewrite mem_map /= ?mem_enum // => x1 x2 [->]. +by apply: closI; apply: sub_orl; rewrite InE mem_map ?mem_enum // =>x1 x2 []. Qed. Hint Resolve clos_rep : core. (* relation defined by the lookup table *) Definition lookup_rel D : rel_exp := - [Pred e1 e2 | + [Pred e | exists a b c c1 c2, - [/\ e1 = app (const a) (const b), + [/\ e.1 = app (const a) (const b), a \in reps D, b \in reps D, - fnd (a, b) (lookup D) = Some (c, c1, c2) & e2 = const (rep D c)]]. + fnd (a, b) (lookup D) = Some (c, c1, c2) & e.2 = const (rep D c)]]. (* inverting relations equations represented as triples *) Lemma invert_look D (x y : exp) : @@ -716,7 +718,7 @@ Definition similar D a b := Definition similar1 D a' b' a b := (const a, const b) \In - closure (rep2rel D \+p [Pred x y | x = const a' /\ y = const b'] \+p + closure (rep2rel D \+p Pred1 (const a', const b') \+p eqs2rel (map pend2eq (pending D))). (* invariants tying use lists with the lookup table *) @@ -891,12 +893,14 @@ Qed. (* Lemmas about join_class *) Lemma join_class_repE (D : data) (a' b' : symb) : - a' \in reps D -> b' \in reps D -> rep_idemp D -> + a' \in reps D -> + b' \in reps D -> + rep_idemp D -> closure (rep2rel (join_class D a' b')) <~> - closure (rep2rel D \+p [Pred x y | x = const a' /\ y = const b']). + closure (rep2rel D \+p Pred1 (const a', const b')). Proof. move=>H2 H3 H1 [x y]; split=>/= T. -- pose ty z := PredU (rep2rel D) [Pred x0 y0 | x0 = const z /\ y0 = const b']. +- pose ty z := PredU (rep2rel D) (Pred1 (const z, const b')). move: (clos_idemp (ty a') (x,y))=>/=<-. apply: sub_closI T=>{x y} [[x y]]. case/imageP=>z _ /= [->->] {x y}; rewrite ffunE /=. @@ -920,8 +924,11 @@ by exists (rep D z); rewrite // ffunE /= H1 // E. Qed. Lemma join_class_useP (D : data) (a' b' : symb) : - a' \in reps D -> b' \in reps D -> a' != b' -> - use_inv D -> use_inv1 (join_class D a' b') a' b'. + a' \in reps D -> + b' \in reps D -> + a' != b' -> + use_inv D -> + use_inv1 (join_class D a' b') a' b'. Proof. move=>H1 H2 H3 H4 c c1 c2; split=>/= [a|H5]; rewrite !ffunE /=; last first. - by case: (H4 _ _ _ _ H1 H5)=>->; [left | right]; rewrite /= eq_refl. @@ -931,8 +938,11 @@ by case: (H4 _ _ _ _ H6 H7)=>->; [left | right]; rewrite (negbTE H5). Qed. Lemma join_class_lookupP (D : data) (a' b' : symb) : - a' \in reps D -> b' \in reps D -> a' != b' -> - lookup_inv D -> lookup_inv (join_class D a' b'). + a' \in reps D -> + b' \in reps D -> + a' != b' -> + lookup_inv D -> + lookup_inv (join_class D a' b'). Proof. move=>H1 H2 H3 H a b c c1 c2. rewrite !join_class_eq // !mem_filter /= !ffunE /=. @@ -942,8 +952,11 @@ by rewrite (negbTE T1) (negbTE Q1). Qed. Lemma join_class_simE (D : data) (a' b' : symb) : - a' \in reps D -> b' \in reps D -> rep_idemp D -> - forall a b, similar1 D a' b' a b <-> similar (join_class D a' b') a b. + a' \in reps D -> + b' \in reps D -> + rep_idemp D -> + forall a b, similar1 D a' b' a b <-> + similar (join_class D a' b') a b. Proof. move=>H1 H2 H a b. by rewrite /similar /similar1 -orrA -clos_clos -join_class_repE // clos_clos. @@ -1043,12 +1056,12 @@ Lemma join_classE (D : data) (a' b' : symb) : lookup_use_inv1 D a' b' -> closure (CRel (join_class D a' b') \+p eqs2rel (map symb2eq (use D a'))) <~> - closure (CRel D \+p [Pred x y | x = const a' /\ y = const b']). + closure (CRel D \+p Pred1 (const a', const b')). Proof. -pose ty := [Pred x0 y0 | x0 = const a' /\ y0 = const b']. -move=>L1 L2 L3 H1 H4 H5 [x y]; split. +pose ty := Pred1 (const a', const b'). +move=>L1 L2 L3 H1 H4 H5 [x y]; rewrite !toPredE; split. - apply: clos_or; move=>{x y}[x y]; last first. - - move: (clos_idemp (CRel D \+p ty) (x,y))=><-. + - rewrite -(clos_idemp (CRel D \+p ty)). apply: sub_closI=>{x y}[[x y]] /=. move/invert=>[c][c1][c2][->->H6]. move: (H4 a' c c1 c2 L1 H6)=>[d][d1][d2][Q1][Q2][Q3] Q4. @@ -1056,66 +1069,65 @@ move=>L1 L2 L3 H1 H4 H5 [x y]; split. apply: (transC (y := const d)). - rewrite /CRel clos_clos !orrA. rewrite -!(rpull (eqs2rel _)) -3!(rpull ty) -!(rpull (rep2rel _)) -!orrA. - move: Q4; apply: sub_closI=>{x y} [[x y]] Q4; apply: sub_orl. - by rewrite !toPredE in Q4 *; rewrite orrA. + move: Q4; apply: sub_closI=>{x y} [[x y]] Q4; apply: sub_orl. + by rewrite orrA. rewrite const_rel /CRel clos_clos !orrA symR. apply: (@closI _ (app (const (rep D c1)) (const (rep D c2)), const (rep D d))). apply: sub_orr; apply: sub_orl. - rewrite toPredE invert_look. + rewrite invert_look. exists (rep D d1), (rep D d2), d, d1, d2. by rewrite -Q2 -Q3 !rep_in_reps. - rewrite /CRel !toPredE clos_idemp clos_clos !orrA /=. + rewrite /CRel clos_idemp clos_clos !orrA /=. rewrite -clos_clos join_class_repE // clos_clos orrA. rewrite -!(rpull (eqs2rel (map pend2eq _))). apply: sub_closKR=>{x y}; apply: sub_closKR=>[[x y]] T1. - rewrite toPredE -clos_idemp. + rewrite -clos_idemp. apply: sub_closI T1=>{x y} [[x y]]. - rewrite toPredE. case=>[T|]; first by apply: closI; right. - case=>a[b][c][c1][c2][-> +++ ->]. + case=>a [b][c][c1][c2][/= -> +++ ->]. rewrite !join_class_eq //= !mem_filter /=. case/andP=>T1 T1'; case/andP=>T2 T2' T3; rewrite ffunE /=. case: eqP=>[H6|_]; [apply: (transC (y:= const a')); last by [apply closI; right] |]; - apply closI; apply: sub_orl; rewrite toPredE invert_look; + apply closI; apply: sub_orl; rewrite invert_look; exists a, b, c, c1, c2=>//. by rewrite H6. -rewrite /CRel !toPredE !clos_clos !orrA /= => {H4} T. +rewrite /CRel !clos_clos !orrA /= => {H4} T. rewrite -clos_clos join_class_repE // clos_clos !orrA -(rpull ty). rewrite -3!(rpull ty) -!(rpull (rep2rel _)) -!(rpull (lookup_rel _)) in T *. rewrite -clos_idemp; apply: sub_closI T=>{x y} [[x y]]. case=>[|T]; last by apply closI; apply: sub_orr; - rewrite toPredE -!orrA; apply: sub_orl; rewrite toPredE !orrA. -case=>a[b][c][c1][c2][-> T1 T2 T3 ->]. + rewrite -!orrA; apply: sub_orl; rewrite !orrA. +case=>a [b][c][c1][c2][/= -> T1 T2 T3 ->]. case: (H5 a b c c1 c2 T1 T2 T3). move=>[d][d1][d2][Q1][Q2][Q3] Q4 [f][f1][f2][F1][F2][F3] F4. -case E1: (a == a'). -- rewrite toPredE !(rpull (lookup_rel _)) !orrA -Q2 -Q3 symR +case E1: (a == a'). +- rewrite !(rpull (lookup_rel _)) !orrA -Q2 -Q3 symR -app_rep -const_rep symR. apply: (transC (y:= const d)); last first. - move: Q4; apply: sub_closI=>{x y} [[x y]]. - by rewrite !toPredE -!orrA=>Q4; do 2!apply: sub_orl. + by rewrite -!orrA=>Q4; do 2!apply: sub_orl. apply: symC; apply closI; do 3!apply: sub_orr; apply: sub_orl. - rewrite toPredE invert -(eqP E1). + rewrite invert -(eqP E1). by exists d, d1, d2. case E2: (b == a'). -- rewrite toPredE !(rpull (lookup_rel _)) !orrA. +- rewrite !(rpull (lookup_rel _)) !orrA. rewrite -F2 -F3 symR -app_rep -const_rep symR. apply: (transC (y:= const f)); last first. - move: F4; apply: sub_closI=>{x y} [[x y]]. - by rewrite !toPredE -!orrA=>F4; do 2!apply: sub_orl. + by rewrite -!orrA=>F4; do 2!apply: sub_orl. apply: symC; apply closI; do 3!apply: sub_orr; apply: sub_orl. - rewrite toPredE invert -(eqP E2). + rewrite invert -(eqP E2). by exists f, f1, f2. case E3: (rep D c == a'); last first. -- apply: closI; apply: sub_orl; rewrite toPredE invert_look. +- apply: closI; apply: sub_orl; rewrite invert_look. exists a, b, c, c1, c2. by rewrite !join_class_eq // !mem_filter /= E1 E2 !ffunE /= E3. rewrite (eqP E3). apply: (transC (y := const b')); last first. - by apply: symC; apply closI; do 2!apply: sub_orr; apply: sub_orl. -apply closI; apply: sub_orl; rewrite toPredE invert_look. +apply closI; apply: sub_orl; rewrite invert_look. exists a, b, c, c1, c2. by rewrite !join_class_eq // !mem_filter /= E1 E2 !ffunE /= E3. Qed. @@ -1138,7 +1150,8 @@ Lemma join_useT t (D : data) (a' b' : symb) (l1 l2 : seq (symb*symb*symb)) : join_use' (join_use' D a' b' l1) a' b' l2. Proof. elim: l1 D a' b'=>[|[[c1 c2] c] l1 IH] D a' b' //= H. -by case F: (fnd _ _)=>[[[d d1] d2]|] /=; rewrite IH //= ffunE eq_refl H /=; eauto. +case F: (fnd _ _)=>[[[d d1] d2]|] /=; +by rewrite IH //= ffunE eq_refl H /=; eauto. Qed. Lemma join_use_useE (D : data) (a' b' : symb) (l1 l2 : seq (symb*symb*symb)) : @@ -1207,15 +1220,15 @@ elim: x D E=>[|[[c c1] c2] ss IH] D E L1 L2 H1 H2 H3 H4 H5 /=. - by rewrite orr0 /CRel clos_idemp. case F: (fnd _ _)=>[[[d d1] d2]|]; set D' := Data _ _ _ _ _. - have S1: closure (rep2rel D' \+p eqs2rel (map pend2eq (pending D'))) <~> - closure (rep2rel D \+p [Pred x y | x = const c /\ y = const d] \+p + closure (rep2rel D \+p Pred1 (const c, const d) \+p eqs2rel (map pend2eq (pending D))) by []. have S2: forall e1 e2, similar D e1 e2 -> similar D' e1 e2. - rewrite /similar=>e1 e2; move: (const e1) (const e2)=>{e1 e2} x y. by rewrite S1; apply: sub_closI=>{x y} [[x y]]; case; [left | right; right]. have T2 : closure (CRel D' \+p eqs2rel (map symb2eq ss)) <~> - closure (CRel D \+p [Pred x y | x = const c /\ - y = app (const c1) (const c2)] \+p eqs2rel (map symb2eq ss)). + closure (CRel D \+p Pred1 (const c, app (const c1) (const c2)) + \+p eqs2rel (map symb2eq ss)). - rewrite /CRel {2 3}/D' /= !clos_clos !orrA. rewrite -!(rpull (eqs2rel (map pend2eq _))). rewrite -!(rpull (eqs2rel (map symb2eq ss))). @@ -1223,19 +1236,19 @@ case F: (fnd _ _)=>[[[d d1] d2]|]; set D' := Data _ _ _ _ _. rewrite -!orrA=>[[x y]]; split=>T; rewrite toPredE -clos_idemp; apply: sub_closI T=>{x y} [[x y]]. - case; first by move=>H7; apply: closI; apply: sub_orl. - case=>->->; rewrite toPredE !orrA symR const_rep symR. + case=>->->; rewrite !orrA symR const_rep symR. apply: (transC (y:= app (const (rep D c1)) (const (rep D c2)))); last first. - apply closI; apply: sub_orr; apply: sub_orl. - rewrite toPredE invert_look. + rewrite invert_look. by exists (rep D c1), (rep D c2), d, d1, d2. by rewrite -app_rep; apply closI; do 2!right. case; first by move=>H7; apply closI; apply: sub_orl. - case=>->->; rewrite toPredE !orrA app_rep. + case=>->->; rewrite !orrA app_rep. apply: (transC (y := const d)); first by apply closI; right; right. apply: (transC (y := const (rep D d))); first by apply: clos_rep. apply: symC; apply closI; apply: sub_orr; apply: sub_orl. - rewrite toPredE invert_look. + rewrite invert_look. exists (rep D c1), (rep D c2), d, d1, d2. by rewrite !rep_in_reps. rewrite -T2. @@ -1286,10 +1299,9 @@ case F: (fnd _ _)=>[[[d d1] d2]|]; set D' := Data _ _ _ _ _. exists h, h1, h2; rewrite {1}/D' /= !ffunE. by case: eqP H7 L1=>[->-> //| _] H7 L1; do !split=>//; apply: S2. have T1: lookup_rel D' <~> lookup_rel D \+p - [Pred x y | x = app (const (rep D c1)) (const (rep D c2)) /\ - y = const (rep D c)]. + Pred1 (app (const (rep D c1)) (const (rep D c2)), const (rep D c)). - rewrite /lookup_rel /D' /= /reps /= =>[[x y]]; split. - - move=>[a][b][d][d1][d2][-> T2 T3]. + - move=>[a][b][d][d1][d2][/= -> T2 T3]. rewrite fnd_ins. case: eqP=>[[->->] [<- _ _] ->|_ T4 T5]; [by right | left]. by exists a, b, d, d1, d2. @@ -1304,20 +1316,19 @@ have T1: lookup_rel D' <~> lookup_rel D \+p rewrite fnd_ins !rep_in_reps. by case: eqP. have T2: closure (CRel D' \+p eqs2rel (map symb2eq ss)) <~> - closure (CRel D \+p [Pred x y | x = const c /\ - y = app (const c1) (const c2)] \+p eqs2rel (map symb2eq ss)). -- pose ty1 := [Pred x y | x = app (const (rep D c1)) - (const (rep D c2)) /\ y = const (rep D c)]. - pose ty2 := [Pred x y | x = const c /\ y = app (const c1) (const c2)]. + closure (CRel D \+p Pred1 (const c, app (const c1) (const c2)) + \+p eqs2rel (map symb2eq ss)). +- pose ty1 := Pred1 (app (const (rep D c1)) (const (rep D c2)), const (rep D c)). + pose ty2 := Pred1 (const c, app (const c1) (const c2)). rewrite /CRel {3}/D' /= T1 !clos_clos !orrA -2!(rpull ty1) -3!(rpull ty2) -!orrA -!(rpull (rep2rel _)). do 3!apply: closRK; move=>[x y]; split=>T; rewrite toPredE -clos_idemp; apply: sub_closI T=>{x y} [[x y]]. - case; first by move=>T; apply: closI; left. - case=>->->; rewrite toPredE symR -const_rep -app_rep. + case=>->->; rewrite symR -const_rep -app_rep. by apply closI; right. case=>[T|[->->]]; first by apply: closI; left. - by rewrite toPredE const_rep app_rep symR; apply closI; right. + by rewrite const_rep app_rep symR; apply closI; right. rewrite -T2. apply: IH=>//=; first by rewrite ffunE eq_refl E. - move=>e e1 e2; split=>/= [a H6|]; rewrite !ffunE; last first. @@ -1450,15 +1461,15 @@ Lemma propagatePE (D : data) : CRel D <~> CRel (propagate D). Proof. move: D. -pose ty x0 y0 := [Pred x y | x = @const s x0 /\ y = @const s y0]. +pose ty x0 y0 := Pred1 (@const s x0, @const s y0). have L: forall D a b, closure (rep2rel D \+p ty a b) <~> closure (rep2rel D \+p ty (rep D a) (rep D b)). - move=>D a b [x y]; split=>T; rewrite toPredE -clos_idemp; apply: sub_closI T=>{x y} [[x y]]; (case; first by move=>H; apply: closI; left); case=>->->; - [rewrite toPredE const_rep symR const_rep symR | - rewrite toPredE -const_rep symR -const_rep symR]; + [rewrite const_rep symR const_rep symR | + rewrite -const_rep symR -const_rep symR]; by apply closI; right. pose inv D := propagate_inv D. have L1: forall D pend_eq p' a b a' b' D', @@ -1475,7 +1486,7 @@ have L1: forall D pend_eq p' a b a' b' D', - move=>a1 b1; rewrite /similar -{2}H4 H1 /= H /= -(rpull (ty a b)) => T. rewrite -clos_idemp; move: T; apply: sub_closI=>{a1 b1} [[x y]]; rewrite -H4; case=>[[->->]|T]; last by apply: closI. - rewrite toPredE const_rep symR const_rep symR /=. + rewrite const_rep symR const_rep symR /=. by rewrite H2 H3 (eqP H5); apply: reflC. suff T: CRel D <~> CRel D'. - rewrite T; apply: IH; rewrite -H4; do 4!split=>//. @@ -1506,7 +1517,7 @@ have L2: forall D pend_eq p' a b a' b' D' D'', - move=>D pend_eq p' a b a' b' D' D'' H1 H H2 H3 H4 H5 H6 IH [H7][H8][H9][H10] H11. have T: CRel D <~> closure (CRel D' \+p - [Pred x y | x = const a' /\ y = const b']). + Pred1 (const a', const b')). - rewrite /CRel H1 -{2 3}H4 /= clos_clos !orrA H /=. rewrite -!(rpull (lookup_rel _)) -!(rpull (eqs2rel (map pend2eq _))). by do 2![apply: closKR]; rewrite L H2 H3 -H4 /= /rep2rel. @@ -1517,7 +1528,7 @@ have L2: forall D pend_eq p' a b a' b' D' D'', rewrite -clos_idemp. move: Q4; apply: sub_closI=>[[x y]]. case; first by move=>Q4; apply: closI; left. - case=>->->; rewrite toPredE !orrA const_rep symR const_rep symR /= H2 H3. + case=>->->; rewrite !orrA const_rep symR const_rep symR /= H2 H3. by apply closI; do 2!right. have [L2' L3']: a' \in reps D' /\ b' \in reps D' by rewrite -H2 -H3 -H4 !rep_in_reps. @@ -1603,7 +1614,7 @@ Lemma propagate_clos_pendP d c c1 c2 e e1 e2 : CRel (Data (rep d) (class d) (use d) (lookup d) (comp_pend (c, c1, c2) (e, e1, e2) :: pending d)) <~> closure (CRel d \+p - [Pred x y | x = const c /\ y = app (const c1) (const c2)]). + Pred1 (const c, app (const c1) (const c2))). Proof. move=>PI H. have [R1 R2]: rep d e1 = rep d c1 /\ rep d e2 = rep d c2. @@ -1614,7 +1625,7 @@ move=>[z y]; split=>O; rewrite toPredE -clos_idemp; move: O; apply: sub_closI; move=>{z y} [z y]. - case=>[O|]; first by apply: closI; left. case=>[O|]; first by apply: closI; right; left. - case=>->->; rewrite toPredE symR const_rep symR. + case=>->->; rewrite symR const_rep symR. apply: (transC (y := app (const (rep d e1)) (const (rep d e2)))). - by rewrite R1 R2 -app_rep; apply closI; right; right. apply closI; right; left; rewrite R1 R2. @@ -1693,23 +1704,23 @@ Lemma propagate_clos_nopendP : (ins (rep d c1, rep d c2) (c, c1, c2) (lookup d)) (pending d)) <~> closure (CRel d \+p - [Pred x y | x = const c /\ y = app (const c1) (const c2)]). + Pred1 (const c, app (const c1) (const c2))). Proof. rewrite /CRel clos_clos orrA orrAC -!orrA; apply: closRK; rewrite !orrA. move=>[xx yy]; split=>O; rewrite toPredE -clos_idemp; move: O; apply: sub_closI; move=>{xx yy} [xx yy]. - case=>[O|]; first by apply: closI; left. - rewrite /lookup_rel =>[[a]][b][e][e1][e2][-> R1 R2 T1 ->]. + rewrite /lookup_rel =>[[a]][b][e][e1][e2][/= -> R1 R2 T1 ->]. rewrite fnd_ins in T1; case: ifP T1=>[|O T1]. - - case/eqP=>->->[->->->]; rewrite toPredE symR -app_rep -const_rep. + - case/eqP=>->->[->->->]; rewrite symR -app_rep -const_rep. by apply closI; right; right. by apply: closI; right; left; exists a, b, e, e1, e2. case=>[O|]; first by apply: closI; left. case=>[O|[->->]]. -- apply: closI; right; move: O=>[a][b][e][e1][e2][-> R1 R2 O ->]. +- apply: closI; right; move: O=>[a][b][e][e1][e2][/= -> R1 R2 O ->]. exists a, b, e, e1, e2; do !split=>//. by rewrite fnd_ins; case: eqP O=>//; case=>->->; rewrite Ev. -rewrite toPredE app_rep const_rep symR; apply closI; right. +rewrite app_rep const_rep symR; apply closI; right. exists (rep d c1), (rep d c2), c, c1, c2. by rewrite !rep_in_reps fnd_ins eq_refl. Qed. @@ -1761,13 +1772,13 @@ move=> H1 H2; case: eqP=>H; constructor. apply: (transC (y := norm D y)); last by apply: symC; apply: norm_rel. by rewrite H; apply: reflC. rewrite /CRel H2 /= orr0 => H3; case: H. -suff: (x, y) \In [Pred e1 e2 | norm D e1 = norm D e2] by []. +suff: (x, y) \In [Pred e | norm D e.1 = norm D e.2] by []. apply: {x y} H3; last first. - split; last by move=>????; rewrite !InE /= => -> ->. by do !split=>//; move=> x y z; rewrite !InE /= => ->. move=>[x y]; case. - by case/imageP=>a _ [->] ->; rewrite /= H1. -move=>[a][b][c][c1][c2][-> Q1 Q2 Q3 ->] /=; do 2!rewrite reps_rep //. +move=>[a][b][c][c1][c2][/= -> Q1 Q2 Q3 ->] /=; do 2!rewrite reps_rep //. by rewrite Q3 H1. Qed. diff --git a/examples/congprog.v b/examples/congprog.v index 6a67161..78dbcc0 100644 --- a/examples/congprog.v +++ b/examples/congprog.v @@ -2,7 +2,7 @@ From HB Require Import structures. From mathcomp Require Import ssreflect ssrbool ssrnat eqtype ssrfun. From mathcomp Require Import div finset seq fintype finfun choice. From pcm Require Import axioms prelude ordtype finmap pred pcm. -From pcm Require Import unionmap heap autopcm automap. +From pcm Require Import unionmap heap autopcm automap. From htt Require Import options model heapauto llist array. From htt Require Import kvmaps hashtab congmath. @@ -22,47 +22,15 @@ From htt Require Import kvmaps hashtab congmath. Notation finE := finset.inE. -Lemma inhabF (I : finType) : 0 < #|I| -> ~ @predT I =1 xpred0. -Proof. by case/card_gt0P=>x _ /(_ x). Qed. - -Definition inhab0 {I : finType} (pf : 0 < #|I|) : I := - match pickP predT with - | Pick x _ => x - | Nopick qf => False_rect I (inhabF I pf qf) - end. - -Definition inhab {I : ifinType} : I := inhab0 card_inhab. - -(* variant of nth that needs no seed value *) -Definition ith {I : finType} i (pf : i < #|I|) : I := - nth (inhab0 (leq_ltn_trans (leq0n i) pf)) (enum I) i. - -(* dually, variant of index that doesn't overflow *) -Definition indx {I : finType} (x : I) := index x (enum I). - -Lemma indx_ith {I : finType} i (pf : i < #|I|) : - indx (ith i pf) = i. -Proof. by rewrite /indx/ith index_uniq ?enum_uniq -?cardE. Qed. - -Lemma ith_indx {I : finType} (s : I) (pf : indx s < #|I|) : - ith (indx s) pf = s. -Proof. by rewrite /ith/indx nth_index // mem_enum. Qed. - -Lemma indx_inj I : injective (@indx I). -Proof. -rewrite /indx=>x1 x2. -have [] : x1 \in enum I /\ x2 \in enum I by rewrite !mem_enum. -elim: (enum I)=>[|x xs IH] //=; rewrite !inE !(eq_sym x). -case: (x1 =P x)=>[<-|] _ /=; first by case: (x2 =P x1). -by case: (x2 =P x)=>//= _ X1 X2 []; apply: IH X1 X2. -Qed. +Prenex Implicits sepit. +Arguments nil {K V}. From mathcomp Require Import bigop. From Coq Require Import Setoid. Prenex Implicits star. Add Parametric Morphism U : (@star U) with signature - @EqPredType _ _ ==> @EqPredType _ _ ==> @EqPredType _ _ as star_morph. + @Eq_Pred _ _ ==> @Eq_Pred _ _ ==> @Eq_Pred _ _ as star_morph. Proof. move=>x y E x1 y1 E1 m /=. split; case=>h1 [h2][-> H1 H2]; exists h1, h2; split=>//. @@ -93,7 +61,7 @@ split; case=>h1 [h2][Em H1 H2]; exists h1, h2; split=>//. by rewrite IH. Qed. -Lemma sepit0 (U : pcm) A (f : A -> Pred U) : +Lemma Iter_sepit0 (U : pcm) A (f : A -> Pred U) : IterStar.sepit [::] f <~> emp. Proof. by rewrite sepitX big_nil. Qed. @@ -163,7 +131,7 @@ congr (_ \+ _). by rewrite IH. Qed. -Lemma sepit_emp (U : pcm) (A : eqType) (s : seq A) (f : A -> Pred U) : +Lemma Iter_sepit_emp (U : pcm) (A : eqType) (s : seq A) (f : A -> Pred U) : (forall x, x \in s -> f x <~> emp (U:=U)) -> IterStar.sepit s f <~> emp. Proof. @@ -178,14 +146,12 @@ move=>x X. apply: H. by rewrite inE X orbT. Qed. + Lemma big_sepit (U : pcm) (I : finType) (s : {set I}) (f : I -> U) m : m = \big[join/Unit]_(i in s) (f i) <-> m \In sepit s (fun i h => h = f i). Proof. by rewrite /sepit /= -big_sepit_seq -big_enum. Qed. - - - (* empty congruence only relates constant symbols to themselves *) Definition empty_cong s := closure (graph (@const s)). @@ -244,8 +210,6 @@ Proof. exact: ltn_pmod. Qed. Definition hash10 k : 'I_10 := Ordinal (@hash_ord 10 k erefl). -Definition LT := HT V hash10. - (* the tables relating arrays with the content of the lists *) (* ctab is for class lists, utab is for use lists *) @@ -272,7 +236,7 @@ Inductive ptrs : Set := {array symb -> llist (symb*symb*symb)} & (* hash table; for each pair of representatives *) (* stores equations; see paper for description *) - KVmap.tp LT & + htab V hash10 & (* list of pending equations *) ptr. @@ -281,11 +245,11 @@ Definition tp := ptrs. Section ShapePredicates. Variable rt : ptrs. -Let r := let: Ptrs r clist ulist htab p := rt in r. -Let clist := let: Ptrs r clist ulist htab p := rt in clist. -Let ulist := let: Ptrs r clist ulist htab p := rt in ulist. -Let htab := let: Ptrs r clist ulist htab p := rt in htab. -Let p := let: Ptrs r clist ulist htab p := rt in p. +Notation r := (let: Ptrs r _ _ _ _ := rt in r). +Notation clist := (let: Ptrs _ clist _ _ _ := rt in clist). +Notation ulist := (let: Ptrs _ _ ulist _ _ := rt in ulist). +Notation htb := (let: Ptrs _ _ _ htb _ := rt in htb). +Notation p := (let: Ptrs _ _ _ _ p := rt in p). (* Data structure's layout in the heap *) @@ -295,7 +259,7 @@ Definition ashape D := Array.shape r (rep d : {ffun symb -> symb}) # (Array.shape clist ct # sepit setT (ctab ct (class d))) # (Array.shape ulist ut # sepit setT (utab ut (use d))) # - KVmap.shape htab (lookup d) # + kvm_shape htb (lookup d) # [Pred h' | exists l, h' \In Pred1 (p :-> l) # lseq l (pending d)]]. Definition bshape d := @@ -319,7 +283,7 @@ Definition iT (clist : {array symb -> llist symb}): Type := forall k, Program Definition init : STsep (emp, [vfun rt m => m \In shape rt (empty_cong s)]) := Do (rx <-- Array.newf [ffun x : symb => x]; - cl <-- Array.new _ null; + cl <-- Array.new null; ffix (fun (loop : iT cl) k => Do (if decP (b:= k < n) idP is left pf then x <-- allocb (ith k pf) 2; @@ -327,12 +291,12 @@ Program Definition init : Array.write cl (ith k pf) x;; loop k.+1 else ret tt)) 0;; - ul <-- Array.new _ null; - htb <-- KVmap.new LT; + ul <-- Array.new null; + htb <-- kvm_new (htab V hash10); p <-- alloc null; ret (Ptrs rx cl ul htb p)). Next Obligation. -move=>_ cl loop k H i [pf][/= f][hc][hct][->{i} Hc Hct]. +case=>i [pf][/= f][hc][hct][->{i} Hc Hct]. case: decP=>[{}pf|] /=; last first. - case: (ltngtP k n) pf=>// Ekn _ _; step=>_. exists f, hc, hct; split=>//. @@ -343,7 +307,7 @@ apply: [gE]=>//=; split=>//. eexists [ffun z => if z == ith k pf then x else f z]. rewrite (_ : _ \+ _ = m \+ (x :-> ith k pf \+ x.+1 :-> null \+ hct)); last by heap_congr. -hhauto; rewrite (sepitS (ith k pf)) finE indx_ith ltnSn. +hhauto; rewrite (sepitS (ith k pf)) finE /= indx_ith ltnSn. rewrite /ctab/table !ffunE eqxx; hhauto. apply: tableP2 Hct=>// a. - by rewrite !finE ltnS indx_injE; case: ltngtP. @@ -356,13 +320,13 @@ apply: [stepX]@hc=>//=. by split=>//; rewrite (_ : [set c | indx c < 0] = set0) // sepit0. case=>_ [f][hc'][hrest][-> Hc' Hrest]. apply: [stepU]=>//= ul hu Ehu; apply: [stepU]=>//= htb ht Ht. -set d := Data [ffun x => x] [ffun c => [:: c]] [ffun c => [::]] (nil K V) [::]. +set d := Data [ffun x => x] [ffun c => [:: c]] [ffun c => [::]] (@nil K V) [::]. step=>px; step; exists d; split; last by case: (initP s). split=>[a b|/=]; first by rewrite !ffunE !inE. exists f, [ffun s => null]. rewrite (_ : px :-> null \+ _ = hr \+ ((hc' \+ hrest) \+ (hu \+ Unit \+ (ht \+ (px :-> null \+ Unit))))); last by rewrite unitR; heap_congr. -hhauto; rewrite sepit_emp //= => k. +hhauto; rewrite /sepit Iter_sepit_emp // => k. by rewrite /utab/table !ffunE; split=>//; case=>_ ->. Qed. @@ -375,11 +339,11 @@ Notation ashape' := (ashape rt). Notation bshape' := (bshape rt). Notation shape' := (shape rt). -Let r := let: Ptrs r clist ulist htab p := rt in r. -Let clist := let: Ptrs r clist ulist htab p := rt in clist. -Let ulist := let: Ptrs r clist ulist htab p := rt in ulist. -Let htab := let: Ptrs r clist ulist htab p := rt in htab. -Let p := let: Ptrs r clist ulist htab p := rt in p. +Notation r := (let: Ptrs r _ _ _ _ := rt in r). +Notation clist := (let: Ptrs _ clist _ _ _ := rt in clist). +Notation ulist := (let: Ptrs _ _ ulist _ _ := rt in ulist). +Notation htb := (let: Ptrs _ _ _ htb _ := rt in htb). +Notation p := (let: Ptrs _ _ _ _ p := rt in p). Definition cT (a' b' : symb) : Type := forall x : unit, STsep {D} @@ -397,7 +361,7 @@ Definition cT (a' b' : symb) : Type := Program Definition join_hclass (a' b' : symb) : STsep {d} (fun i => i \In bshape' d /\ a' != b', [vfun (_ : unit) m => m \In bshape' (join_class d a' b')]) := - Do (ffix (fun (loop : cT a' b') (x : unit) => + Do (ffix (fun (loop : cT a' b') (xx : unit) => Do (a <-- Array.read clist a'; b <-- Array.read clist b'; if a == null then ret tt @@ -410,8 +374,8 @@ Program Definition join_hclass (a' b' : symb) : Array.write r s b';; loop tt)) tt). Next Obligation. -move=>a' b' loop [[[[/= d ct] ut]]][i][H N] /=. -case: H=>/= rh [_][->{i} Rh][_][h][->][th][ctx][->] Th Ctx H. +case=>[[[/= d ct] ut]][i][H N] /=. +case: H=>/= rh [_][->{i} Rh][_][h][->][th][ctx][->] Th Ctx H. rewrite (sepitT1 a') in Ctx; case: Ctx=>cta [w][->{ctx}Cta]. rewrite (sepitS b') !finE eq_sym {1}N /=. case=>ctb [ctx][->{w}Ctb Ctx]; rewrite /ctab/table in Cta Ctb. @@ -455,7 +419,7 @@ case: (x =P a')=>// _; case: (x =P b')=>// _. by rewrite Eca rev_cons cat_rcons. Qed. Next Obligation. -move=>a' b' [d][i][[C]][/= ct][ut H] N. +case=>d [i][[C]][/= ct][ut H] N. apply: [gE (d, ct, ut)]=>//= [[m]][ct1][ut1] W _. set d' := (Data _ _ _ _ _) in W; suff E : join_class d a' b' = d'. - by split; [apply: join_class_classP|rewrite E; eauto]. @@ -466,12 +430,12 @@ Definition uT (a' b' : symb) := forall x : unit, STsep {d} (fun i => exists don, a' != b' /\ i \In bshape' (join_use' d a' b' don) /\ use d a' = don ++ use (join_use' d a' b' don) a', - [vfun (_ : unit) m => m \In bshape' (join_use d a' b')]). + [vfun (_ : unit) m => m \In bshape' (join_use d a' b')]). Program Definition join_huse (a' b' : symb) : STsep {d} (fun i => a' != b' /\ i \In bshape' d, [vfun (_ : unit) m => m \In bshape' (join_use d a' b')]) := - Do (ffix (fun (loop : uT a' b') x => + Do (ffix (fun (loop : uT a' b') xx => Do (a <-- Array.read ulist a'; if a == null then ret tt else @@ -480,7 +444,7 @@ Program Definition join_huse (a' b' : symb) : Array.write ulist a' next;; c1 <-- Array.read r eq.1.2; c2 <-- Array.read r eq.2; - v <-- HashTab.lookup hash10 htab (c1, c2); + v <-- kvm_lookup htb (c1, c2); if v is Some d then dealloc a;; dealloc a.+1;; @@ -489,16 +453,16 @@ Program Definition join_huse (a' b' : symb) : p ::= q;; loop tt else - HashTab.insert hash10 htab (c1, c2) eq;; + kvm_insert htb (c1, c2) eq;; b <-- Array.read ulist b'; a.+1 ::= b;; Array.write ulist b' a;; loop tt)) tt). Next Obligation. -move=>a' b' loop [[/= d]][i][don][N][H Eu]. -set d1 := join_use' d a' b' don in H Eu. +case=>[/= d][i][don][N][H Eu]. +set d1 := join_use' d a' b' don in H Eu. case: H=>C [/= ct][ut][rh][_][->{i} Rh][cth][_][-> Htc][_][_][->] -[ru][hu][-> Ut Hu][ht][_][-> Ht][p'][_][hp][-> <- Hp]. +[ru][hu][-> Ut Hu][ht][_][-> Ht][p'][_][hp][->-> Hp]. move: Hu; rewrite (sepitT1 a'); case=>ha' [w][->{hu} Ua]. rewrite (sepitS b') !finE eq_sym {1}N; case=>hb' [h][->{w} Ub R]. rewrite /utab/table/= in Ua Ub. @@ -551,13 +515,13 @@ rewrite (_ : _ \+ _ = rh \+ (cth \+ ((ru \+ (hh \+ ((ut a' :-> c \+ case: Htc=>x [y][->] X1 X2; hhauto; [eauto|eauto|exact: Ut3|]. rewrite (sepitT1 a'); hhauto. - by rewrite /utab/table/ut3 ffunE /= (negbTE N) /ut2 !ffunE /= eqxx. -rewrite (sepitS b') !finE eq_sym N /=; hhauto. +rewrite (sepitS b') !finE eq_sym N /=; hhauto. - by rewrite /utab/table/=/ut3 !ffunE /= eqxx eq_sym (negbTE N); hhauto. apply: tableP R=>a /=; rewrite !finE andbT /ut3/ut2 !ffunE /= ?ffunE /=; by case/andP=>/negbTE -> /negbTE ->. Qed. Next Obligation. -by move=>a' b' [/= d][i][N H]; apply: [gE d]=>[||??[]] //; exists [::]. +by case=>d [i][N H]; apply: [gE d]=>[||??[]] //; exists [::]. Qed. Definition pT : Type := forall x : unit, @@ -588,8 +552,8 @@ Program Definition hpropagate := join_huse a' b';; loop tt)) tt. Next Obligation. -move=>loop [][/= d][i][C][/= ct][ut][hr][_][-> Hr][hc][_][-> Hc] -[hu][_][-> Hu][ht][_][-> Ht][q][_][hp][->] <- Hp. +case=>/= d [i][C][/= ct][ut][hr][_][-> Hr][hc][_][-> Hc] +[hu][_][-> Hu][ht][_][-> Ht][q][_][hp][->] -> Hp. step; case: (q =P null) Hp=>[->{q}|/eqP N] Hp. - apply: vrfV=>V; case/(lseq_null (validX V)): Hp=>{V} Ep /= ->{hp}. step; rewrite propagate_equation Ep; split=>//=. @@ -607,7 +571,7 @@ case: ifPn=>[E|E]. apply: [stepE d1]=>//=; first by do 2split=>//; exists ct, ut; hhauto. set d2 := (join_class _ _ _); case=>m H; apply: [stepE d2]=>//= [[n]] Hn. set d3 := (join_use _ _ _) in Hn. -suff -> : propagate d = propagate d3 by apply: [gE d3]. +suff -> : propagate d = propagate d3 by apply: [gE d3]. simpl in Hu. rewrite /d3/d2 propagate_equation Ep; by case: eq Ep Ea' Eb' E=>[a2 b2 _|[[a2 ??]][[b2 ??]] _] ->-> /negbTE ->. Qed. @@ -625,14 +589,14 @@ Program Definition merge (e : Eq s) : | comp_eq c c1 c2 => Do (c1' <-- Array.read r c1; c2' <-- Array.read r c2; - v <-- KVmap.lookup htab (c1', c2'); + v <-- kvm_lookup htb (c1', c2'); if v is Some b then q <-- !p; x <-- insert q (comp_pend (c, c1, c2) b); p ::= x;; hpropagate else - KVmap.insert htab (c1', c2') (c, c1, c2);; + kvm_insert htb (c1', c2') (c, c1, c2);; u1 <-- Array.read ulist c1'; x <-- insert u1 (c, c1, c2); (* if c1' == c2' the equation will be added twice *) @@ -647,8 +611,8 @@ Program Definition merge (e : Eq s) : Array.write ulist c2' x) end. Next Obligation. -move=>e a b [R][_][d][[C]][/= ct][ut][hr][_][-> Hr][hc][_][-> Hc][_][_] -[->][hu][hu'][-> Hu Hu'][ht][_][-> Ht][q][_][hp][->] <- Hp [PI][Ep Erel]. +case=>R [_][d][[C]][/= ct][ut][hr][_][-> Hr][hc][_][-> Hc][_][_] +[->][hu][hu'][-> Hu Hu'][ht][_][-> Ht][q][_][hp][->] -> Hp [PI][Ep Erel]. step; apply: [stepX (pending d)]@hp=>//= x _ [r0][{Hp}hp][-> Hp]. set d1:=Data (rep d) (class d) (use d) (lookup d) (simp_pend a b :: pending d). step; apply: [gE d1]=>//=. @@ -663,14 +627,14 @@ case: (propagatePE L)=>L1 [L2 L3]; exists (propagate d1); do 3!split=>//. by rewrite -L3 -Erel clos_clos /CRel -!orrA orrAC. Qed. Next Obligation. -move=>X c c1 c2 [R][_][d][[C]][/= ct][ut][hr][_][-> Hr][hc][_][-> H] +case=>R [_][d][[C]][/= ct][ut][hr][_][-> Hr][hc][_][-> H] [_][_][->][hu][hu'][-> Hu Hu'][ht][_][-> Ht][q][_] -[hp][->] <- Hp [PI][Ep Erel]; set cx := (c, c1, c2). +[hp][->] -> Hp [PI][Ep Erel]; set cx := (c, c1, c2). apply: [stepX rep d, hr]@hr=>//= _ _ [->->]. apply: [stepX rep d, hr]@hr=>//= _ _ [->->]. apply: [stepX lookup d]@ht=>//= v {Ht}ht [Ht Ev]. -case: v Ev=>[[[e e1 e2]]|] Ev. -- set ex := (e, e1, e2). +case: v Ev=>[[[e0 e1 e2]]|] Ev. +- set ex := (e0, e1, e2). set d1 := Data (rep d) (class d) (use d) (lookup d) (comp_pend cx ex :: pending d). step; apply: [stepX pending d]@hp=>//= x _ [x1][{}hp][-> {}Hp]. @@ -744,7 +708,7 @@ Program Definition hnorm := u1 <-- hnorm t1; u2 <-- hnorm t2; if (u1, u2) is (const w1, const w2) then - v <-- KVmap.lookup htab (w1, w2); + v <-- kvm_lookup htb (w1, w2); if v is Some a then a' <-- Array.read r a.1.1; ret (const a') @@ -752,7 +716,7 @@ Program Definition hnorm := else ret (app u1 u2) end)). Next Obligation. -move=>hnorm t [d][_][Ci][/= ct][ut][hr][hrest][-> Hr Hrest]. +case=>d [_][Ci][/= ct][ut][hr][hrest][-> Hr Hrest]. case: t=>[a|t1 t2]. - apply: [stepX rep d, hr]@hr=>//= _ _ [->->]. by step; do 2!split=>//; exists ct, ut; hhauto. @@ -780,7 +744,7 @@ Program Definition check t1 t2 : u2 <-- hnorm t2; ret (u1 == u2)). Next Obligation. -move=>t1 t2 [R][h][d][H][[RI X]][Ep PI]. +case=>R [h][d][H][[RI X]][Ep PI]. apply: [stepX d]@h=>//= _ {H}h [-> H]. apply: [stepX d]@h=>//= _ {H}h [-> H]. step; split; first by exists d. diff --git a/examples/graph.v b/examples/graph.v index 9a6f2a8..84f728c 100644 --- a/examples/graph.v +++ b/examples/graph.v @@ -73,8 +73,8 @@ HB.instance Definition _ := isNatMap.Build (seq node) pregraph. HB.instance Definition _ := hasDecEq.Build pregraph (@union_map_eqP node _ (seq node) pregraph). Canonical pregraph_PredType : PredType (node * (seq node)) := - Mem_UmMap_PredType pregraph. -Coercion Pred_of_history (x : pregraph) : Pred_Class := + um_PredType pregraph. +Coercion Pred_of_history (x : pregraph) : {Pred _} := [eta Mem_UmMap x]. Notation "x &-> v" := (ptsT pregraph x v) (at level 30). diff --git a/examples/hashtab.v b/examples/hashtab.v index ac41cfa..e73e343 100644 --- a/examples/hashtab.v +++ b/examples/hashtab.v @@ -19,9 +19,6 @@ From pcm Require Import pcm unionmap heap autopcm. From htt Require Import options model heapauto. From htt Require Import array kvmaps. -Module HashTab. -Section HashTab. - (* hash table is array of buckets, i.e. KV maps *) (* bucket indices are provided by the hash function *) (* using dynaming kv-maps for buckets *) @@ -29,18 +26,49 @@ Section HashTab. (* it's possible to develop this with static buckers *) (* /DEVCOMMENT *) -Variables (K : ordType) (V : Type) (buckets : DynKVmap.Sig K V) - (n : nat) (hash : K -> 'I_n). -Definition hashtab := {array 'I_n -> DynKVmap.tp buckets}. -Notation KVshape := (@DynKVmap.shape _ _ buckets). +Module Type Hashtab_sig. +Parameter root : forall {K : ordType} {V : Type} (buckets : dkvm K V) + {n : nat} (hash : K -> 'I_n), Set. +Parameter null_root : forall {K : ordType} {V : Type} {buckets : dkvm K V} + {n : nat} {hash : K -> 'I_n}, root buckets hash. +Parameter shape : forall {K : ordType} {V : Type} {buckets : dkvm K V} + {n : nat} {hash : K -> 'I_n}, root buckets hash -> {finMap K -> V} -> + Pred heap. +Parameter new : forall {K : ordType} {V : Type} {buckets : dkvm K V} + {n : nat} {hash : K -> 'I_n}, + STsep (emp, [vfun (x : root buckets hash) h => h \In shape x (nil K V)]). +Parameter free : forall {K : ordType} {V : Type} {buckets : dkvm K V} + {n : nat} {hash : K -> 'I_n} (x : root buckets hash), + STsep {s} (shape x s, [vfun _ : unit => emp]). +Parameter insert : forall {K : ordType} {V : Type} {buckets : dkvm K V} + {n : nat} {hash : K -> 'I_n} (x : root buckets hash) k v, + STsep {s} (shape x s, [vfun (_ : unit) h => h \In shape x (ins k v s)]). +Parameter remove : forall {K : ordType} {V : Type} {buckets : dkvm K V} + {n : nat} {hash : K -> 'I_n} (x : root buckets hash) k, + STsep {s} (shape x s, [vfun (_ : unit) h => h \In shape x (rem k s)]). +Parameter lookup : forall {K : ordType} {V : Type} {buckets : dkvm K V} + {n : nat} {hash : K -> 'I_n} (x : root buckets hash) k, + STsep {s} (shape x s, [vfun y h => h \In shape x s /\ y = fnd k s]). +End Hashtab_sig. + +Module HashTab : Hashtab_sig. +Definition root K V (buckets : dkvm K V) n (hash : K -> 'I_n) : + Set := {array 'I_n -> buckets}. +Definition null_root K V buckets n hash : + @root K V buckets n hash := Array null. +Section HashTab. +Context (K : ordType) (V : Type) {buckets : dkvm K V} + {n : nat} {hash : K -> 'I_n}. +Notation KVshape := (@dkvm_shape _ _ buckets). Notation table := (table KVshape). Notation nil := (nil K V). +Notation root := (root buckets hash). (* hash table is specified by a single finMap *) (* which is the "flattening" of all buckets *) -Definition shape x (s : finMap K V) := - [Pred h | exists (tab : {ffun 'I_n -> DynKVmap.tp buckets}) (* array spec *) - (bucket : 'I_n -> finMap K V), (* buckets spec *) +Definition shape (x : root) (s : finMap K V) : Pred heap := + [Pred h | exists (tab : {ffun 'I_n -> buckets}) (* array spec *) + (bucket : 'I_n -> {finMap K -> V}), (* buckets spec *) [/\ forall k, fnd k s = fnd k (bucket (hash k)), forall i k, k \in supp (bucket i) -> hash k = i & h \In Array.shape x tab # sepit setT (table tab bucket)] ]. @@ -56,16 +84,16 @@ Definition new_loopinv x := forall k, [vfun y => shape y nil]). Program Definition new : STsep (emp, [vfun y => shape y nil]) := - Do (t <-- Array.new 'I_n (DynKVmap.default buckets); + Do (t <-- Array.new dkvm_null; let go := ffix (fun (loop : new_loopinv t) k => Do (if decP (b := k < n) idP is left pf then - b <-- DynKVmap.new buckets; + b <-- dkvm_new buckets; Array.write t (Ordinal pf) b;; loop k.+1 else ret t)) in go 0). (* first the bucket initialization loop *) -Next Obligation. +Next Obligation. (* pull out preconditions, branch on k comparison *) move=>/= arr loop k [] _ /= [Eleq][tab][h1][h2][-> H1]. case: decP=>[{Eleq}pf H2|]; last first. @@ -101,7 +129,7 @@ apply: [stepE]=>//= y m Hm. (* invoke the loop with index 0 *) apply: [gE]=>//=; split=>//. (* the table is empty *) -exists [ffun => DynKVmap.default buckets], m, Unit; split=>//=. +exists [ffun => dkvm_null], m, Unit; split=>//=. - by rewrite unitR. (* there are no buckets in the heap yet *) by rewrite (eq_sepit (s2 := set0)) // sepit0. @@ -123,7 +151,7 @@ Program Definition free x : STsep {s} (shape x s, Do (ffix (fun (loop : free_loopinv x) k => Do (if decP (b := k < n) idP is left pf then b <-- Array.read x (Ordinal pf); - DynKVmap.free b;; + dkvm_free b;; loop k.+1 else Array.free x)) 0). (* first the loop *) @@ -169,7 +197,7 @@ Program Definition insert x k v : STsep {s} (shape x s, [vfun _ : unit => shape x (ins k v s)]) := Do (let hk := hash k in b <-- Array.read x hk; - b' <-- DynKVmap.insert b k v; + b' <-- dkvm_insert b k v; Array.write x hk b'). Next Obligation. (* pull out ghost + deconstruct precondition *) @@ -212,7 +240,7 @@ Program Definition remove x k : [vfun _ : unit => shape x (rem k s)]) := Do (let hk := hash k in b <-- Array.read x hk; - b' <-- DynKVmap.remove b k; + b' <-- dkvm_remove b k; Array.write x hk b'). Next Obligation. (* pull out ghost + destructure precondition *) @@ -252,7 +280,7 @@ Program Definition lookup x k : STsep {s} (shape x s, [vfun y m => m \In shape x s /\ y = fnd k s]) := Do (b <-- Array.read x (hash k); - DynKVmap.lookup b k). + dkvm_lookup b k). Next Obligation. (* pull out ghost + destructure precondition *) move=>/= x k [fm][] _ /= [tf][bf][Hf Hh [h1][h2][-> H1 H2]]. @@ -268,10 +296,17 @@ exists tf, bf; split=>//=; exists h1, (m2 \+ h4); split=>{h1 H1} //. by rewrite (sepitT1 (hash k)); vauto. Qed. -(* hash table is a *static* KV map *) -Definition HashTab := KVmap.make (Array null) new free insert remove lookup. - End HashTab. End HashTab. -Definition HT K V := HashTab.HashTab (DynAssocList.AssocList K V). +(* hash table is (static) KV map *) +Notation hashtab := HashTab.root. +HB.instance Definition _ K V (buckets : dkvm K V) n (hash : K -> 'I_n) := + isKVM.Build K V (hashtab buckets hash) HashTab.null_root + HashTab.new HashTab.free HashTab.insert HashTab.remove HashTab.lookup. + +(* htab is specific simple hash tab where buckets are association lists *) +Definition htab K V n hash := @hashtab K V (dalist K V) n hash. +HB.instance Definition _ K V n hash := KVM.on (@htab K V n hash). + + diff --git a/examples/kvmaps.v b/examples/kvmaps.v index 267de8c..e7f8444 100644 --- a/examples/kvmaps.v +++ b/examples/kvmaps.v @@ -15,16 +15,13 @@ limitations under the License. (* Key-Value maps *) (******************) +From HB Require Import structures. From Coq Require Import ssreflect ssrbool ssrfun. From mathcomp Require Import eqtype seq path ssrnat. From pcm Require Import options axioms pred ordtype finmap seqext. From pcm Require Import pcm unionmap heap autopcm automap. From htt Require Import options model heapauto. -(***********************) -(* stateful KV map ADT *) -(***********************) - (* Dynamic KV map is determined by its root pointer(s). *) (* Functions such insert and remove may modify *) (* the root, and will correspondingly return the new one. *) @@ -33,24 +30,24 @@ From htt Require Import options model heapauto. (* passing K and V to methods thus reducing annotations *) (* There's no deep reason to make tp : Set, except that *) (* it should be thought of a collection of pointers, hence small. *) -Module DynKVmap. -Record Sig (K : ordType) (V : Type) : Type := - make {tp :> Set; - default : tp; - shape : tp -> {finMap K -> V} -> Pred heap; - new : STsep (emp, [vfun x => shape x (nil K V)]); - free : forall x, STsep {s} (shape x s, - [vfun _ : unit => emp]); - insert : forall x k v, - STsep {s} (shape x s, - [vfun y => shape y (ins k v s)]); - remove : forall x k, - STsep {s} (shape x s, - [vfun y => shape y (rem k s)]); - lookup : forall x k, - STsep {s} (shape x s, - [vfun y m => m \In shape x s /\ y = fnd k s])}. -End DynKVmap. + +HB.mixin Record isDKVM (K : ordType) (V : Type) (root : Set) := { + dkvm_null : root; + dkvm_shape : root -> {finMap K -> V} -> Pred heap; + dkvm_new : STsep (emp, [vfun x => dkvm_shape x (nil K V)]); + dkvm_free : forall x, + STsep {s} (dkvm_shape x s, [vfun _ : unit => emp]); + dkvm_insert : forall x k v, + STsep {s} (dkvm_shape x s,[vfun y => dkvm_shape y (ins k v s)]); + dkvm_remove : forall x k, + STsep {s} (dkvm_shape x s,[vfun y => dkvm_shape y (rem k s)]); + dkvm_lookup : forall x k, + STsep {s} (dkvm_shape x s, [vfun y m => m \In dkvm_shape x s /\ y = fnd k s])}. + +#[short(type="dkvm")] +HB.structure Definition DKVM K V := {root of @isDKVM K V root}. + +Arguments dkvm_new {K V}. (* Static KV map (or just KV map) are those that *) (* don't need to modify the root pointer. *) @@ -59,36 +56,50 @@ End DynKVmap. (* the root pointer along with the dynamic KV map that the *) (* root points to. *) -Module KVmap. -Record Sig (K : ordType) (V : Type) : Type := - make {tp :> Set; - default : tp; - shape : tp -> {finMap K -> V} -> Pred heap; - (* allocate root pointer and empty structure *) - new : STsep (emp, [vfun x => shape x (nil K V)]); - (* remove the struture, and the root pointer *) - free : forall x, STsep {s} (shape x s, - [vfun _ : unit => emp]); - (* insert, keeping the root pointer unchanged *) - insert : forall x k v, - STsep {s} (shape x s, - [vfun _ : unit => shape x (ins k v s)]); - (* remove, keeping the root pointer unchanged *) - remove : forall x k, - STsep {s} (shape x s, - [vfun _ : unit => shape x (rem k s)]); - lookup : forall x k, - STsep {s} (shape x s, - [vfun y m => m \In shape x s /\ y = fnd k s])}. -End KVmap. - -(**********************************************************) -(* KV map implemented as a sorted association linked list *) -(**********************************************************) - -Module DynAssocList. +HB.mixin Record isKVM (K : ordType) (V : Type) (root : Set) := { + kvm_null : root; + kvm_shape : root -> {finMap K -> V} -> Pred heap; + kvm_new : STsep (emp, [vfun x => kvm_shape x (nil K V)]); + kvm_free : forall x, + STsep {s} (kvm_shape x s, [vfun _ : unit => emp]); + kvm_insert : forall x k v, + STsep {s} (kvm_shape x s,[vfun _ : unit => kvm_shape x (ins k v s)]); + kvm_remove : forall x k, + STsep {s} (kvm_shape x s,[vfun _ : unit => kvm_shape x (rem k s)]); + kvm_lookup : forall x k, + STsep {s} (kvm_shape x s, [vfun y m => m \In kvm_shape x s /\ y = fnd k s])}. + +#[short(type="kvm")] +HB.structure Definition KVM K V := {root of @isKVM K V root}. + +Arguments kvm_new {K V}. + +(*****************************************) +(* Association lists with DKVM interface *) +(*****************************************) + +Module Type DAList_sig. +Parameter root : ordType -> Type -> Set. +Parameter null_root : forall {K : ordType} {V : Type}, root K V. +Parameter shape : forall {K : ordType} {V : Type}, + root K V -> {finMap K -> V} -> Pred heap. +Parameter new : forall {K : ordType} {V : Type}, + STsep (emp, [vfun (x : root K V) h => h \In shape x (nil K V)]). +Parameter free : forall {K : ordType} {V : Type} (x : root K V), + STsep {s} (shape x s, [vfun _ : unit => emp]). +Parameter insert : forall {K : ordType} {V : Type} (x : root K V) k v, + STsep {s} (shape x s, [vfun (y : root K V) h => h \In shape y (ins k v s)]). +Parameter remove : forall {K : ordType} {V : Type} (x : root K V) k, + STsep {s} (shape x s, [vfun (y : root K V) h => h \In shape y (rem k s)]). +Parameter lookup : forall {K : ordType} {V : Type} (x : root K V) k, + STsep {s} (shape x s, [vfun y m => m \In shape x s /\ y = fnd k s]). +End DAList_sig. + +Module DAList : DAList_sig. +Definition root (K : ordType) (V : Type) := ptr. +Definition null_root K V : @root K V := null. Section AssocList. -Variables (K : ordType) (V : Set). +Variables (K : ordType) (V : Type). Notation fmap := (finMap K V). Notation nil := (nil K V). @@ -110,6 +121,8 @@ Fixpoint shape_seg' (x y : ptr) (xs : seq (K * V)) : Pred heap := Definition shape_seg (x y : ptr) (s : fmap) : Pred heap := shape_seg' x y (seq_of s). +(* definition that adds dymmy dependence on D *) +(* to make the module system pass *) Definition shape (x : ptr) (s : fmap) : Pred heap := shape_seg x null s. @@ -119,10 +132,10 @@ Lemma shape_null (s : fmap) h : h \In shape null s -> s = nil /\ h = Unit. Proof. -move=>D; case: s; case=>/=. +move=>W; case: s; case=>/=. - by move=>? [_ ->] //; rewrite fmapE. move=>[??]??[?][?][H]. -by rewrite H -!joinA validPtUn in D. +by rewrite H -!joinA validPtUn in W. Qed. (* non-empty well-formed region represents a non-empty map such that *) @@ -197,7 +210,8 @@ Qed. (* main procedures *) (* new map is just a null pointer *) -Program Definition new : STsep (emp, [vfun x => shape x nil]) := +Program Definition new : + STsep (emp, [vfun x => shape x nil]) := Do (ret null). Next Obligation. by case=>_ /= ->; step. Qed. @@ -218,7 +232,7 @@ Next Obligation. (* pull out ghost var, precondition, branch *) move=>loop p [fm][] i /= H; case: eqP=>[|/eqP] E. (* reached the end, heap must be empty *) -- by apply: vrfV=>D; step=>_; rewrite E in H; case: (shape_null D H). +- by apply: vrfV=>W; step=>_; rewrite E in H; case: (shape_null W H). (* pull out the head entry *) case: (shape_cont E H)=>k[v][q][h][_ _ {i H}-> H]. (* deallocate it *) @@ -229,7 +243,6 @@ Qed. (* looking up an element in the map *) - Definition lookupT k : Type := forall p, STsep {fm} (shape p fm, [vfun y m => m \In shape p fm /\ y = fnd k fm]). @@ -604,30 +617,56 @@ apply/path_supp_ins=>//. by apply/path_le/all_path_supp/Of. Qed. -(* ordered association list is a dynamic KV map *) - -Definition AssocList := DynKVmap.make null new free insert remove lookup. End AssocList. -End DynAssocList. +End DAList. + +(* association list is dynamic KV map *) +Notation dalist := DAList.root. +HB.instance Definition _ K V := + isDKVM.Build K V (dalist K V) DAList.null_root + DAList.new DAList.free DAList.insert DAList.remove DAList.lookup. + +(***************************************) +(* Association maps with KVM interface *) +(***************************************) (* static variant packages the root pointer *) (* with the dynamic part of the structure *) -Module AssocList. +Module Type AList_sig. +Parameter root : ordType -> Type -> Set. +Parameter null_root : forall {K : ordType} {V : Type}, root K V. +Parameter shape : forall {K : ordType} {V : Type}, + root K V -> {finMap K -> V} -> Pred heap. +Parameter new : forall {K : ordType} {V : Type}, + STsep (emp, [vfun (x : root K V) h => h \In shape x (nil K V)]). +Parameter free : forall {K : ordType} {V : Type} (x : root K V), + STsep {s} (shape x s, [vfun _ : unit => emp]). +Parameter insert : forall {K : ordType} {V : Type} (x : root K V) k v, + STsep {s} (shape x s, [vfun (_ : unit) h => h \In shape x (ins k v s)]). +Parameter remove : forall {K : ordType} {V : Type} (x : root K V) k, + STsep {s} (shape x s, [vfun (_ : unit) h => h \In shape x (rem k s)]). +Parameter lookup : forall {K : ordType} {V : Type} (x : root K V) k, + STsep {s} (shape x s, [vfun y m => m \In shape x s /\ y = fnd k s]). +End AList_sig. + +Module AList : AList_sig. Section AssocList. -Variables (K : ordType) (V : Set). +Definition root : ordType -> Type -> Set := fun _ _ => ptr. +Variables (K : ordType) (V : Type). Notation fmap := (finMap K V). Notation nil := (nil K V). +Definition null_root : root K V := null. Definition shape (x : ptr) (f : fmap) : Pred heap := - [Pred h | exists (a : ptr) h', - h = x :-> a \+ h' /\ @DynAssocList.shape K V a f h']. + [Pred h | exists (a : dalist K V) h', + h = x :-> a \+ h' /\ dkvm_shape a f h']. (* new structure, but the root pointer is given *) Program Definition new0 (x : ptr) : - STsep {a : ptr} (fun h => h = x :-> a, + STsep {v : dalist K V} (fun h => h = x :-> v, [vfun _ : unit => shape x nil]) := - Do (a <-- @DynAssocList.new K V; + Do (a <-- dkvm_new (dalist K V); x ::= a). Next Obligation. move=>x [a][/= _ ->]; apply: [stepU]=>//= b h H. @@ -636,7 +675,7 @@ Qed. Program Definition new : STsep (emp, [vfun x => shape x nil]) := - Do (a <-- @DynAssocList.new K V; + Do (a <-- dkvm_new (dalist K V); alloc a). Next Obligation. case=>_ /= ->; apply: [stepU]=>//= a h H. @@ -645,12 +684,12 @@ Qed. (* free structure, keep the root pointer *) Program Definition free0 x : - STsep {s} (shape x s, - [vfun (_ : unit) h => exists a : ptr, h = x :-> a]) := + STsep {s} (shape x s, [vfun (_ : unit) h => + exists a : dalist K V, h = x :-> a]) := Do (a <-- !x; - DynAssocList.free K V a). + dkvm_free (a : dalist K V)). Next Obligation. -move=>x [fm][/= _ [a][h][-> H]]; step. +move=>x [fm][/= _ [a][h][-> H]]; step. by apply: [gX fm]@h=>//= _ _ -> _; rewrite unitR; eauto. Qed. @@ -659,16 +698,16 @@ Program Definition free x : [vfun _ : unit => emp]) := Do (a <-- !x; dealloc x;; - DynAssocList.free K V a). + dkvm_free (a : dalist K V)). Next Obligation. by move=>x [fm][/= _ [a][h][-> H]]; step; step; apply: [gX fm]@h. Qed. Program Definition insert x k v : STsep {s} (shape x s, - [vfun _ : unit => shape x (ins k v s)]) := + [vfun (_ : unit) h => h \In shape x (ins k v s)]) := Do (a <-- !x; - y <-- DynAssocList.insert a k v; + y <-- dkvm_insert (a : dalist K V) k v; x ::= y). Next Obligation. move=>x k v [fm][/= _ [a][h][-> H]]; step. @@ -679,7 +718,7 @@ Program Definition remove x k : STsep {s} (shape x s, [vfun _ : unit => shape x (rem k s)]) := Do (a <-- !x; - y <-- DynAssocList.remove V a k; + y <-- dkvm_remove (a : dalist K V) k; x ::= y). Next Obligation. move=>x k [fm][/= _ [a][h][-> H]]; step. @@ -690,13 +729,17 @@ Program Definition lookup x k : STsep {s} (shape x s, [vfun y m => m \In shape x s /\ y = fnd k s]) := Do (a <-- !x; - DynAssocList.lookup V a k). + dkvm_lookup (a : dalist K V) k). Next Obligation. move=>x k [fm][/= _ [a][h][-> H]]; step. by apply: [gX fm]@h=>//= y {}h {H}[]; hhauto. Qed. - -(* ordered association list is a KV map *) -Definition AssocList := KVmap.make null new free insert remove lookup. -End AssocList. End AssocList. +End AList. + +(* association list is (static) KV map *) +Notation alist := AList.root. +HB.instance Definition _ K V := + isKVM.Build K V (alist K V) AList.null_root + AList.new AList.free AList.insert AList.remove AList.lookup. + diff --git a/examples/quicksort.v b/examples/quicksort.v index f12c259..5b004d2 100644 --- a/examples/quicksort.v +++ b/examples/quicksort.v @@ -20,7 +20,6 @@ From pcm Require Import pcm unionmap heap. From htt Require Import options model heapauto. From htt Require Import array. -(* hack to avoid "_ *p _" notation clash *) From mathcomp Require order. Import order.Order.NatOrder order.Order.TTheory. @@ -266,11 +265,9 @@ move=>a i j /= [f][] h /= H. case: ifP=>[/eqP->|Hij]. - by step=>_; rewrite tperm1 pffunE1. do 2!apply: [stepE f, h]=>//= _ _ [->->]. -apply: [stepE f]=>//= _ _ ->. -set f1 := finfun [eta f with i |-> f j]. -apply: [gE f1]=>//= _ _ ->. -set f2 := finfun [eta f1 with j |-> f i]. -move=>{h H} _; suff : f2 = pffun (tperm i j) f by move=>->. +apply: [stepE f]=>//= _ {H}h H; set f1 := (finfun _) in H. +apply: [gE f1]=>//= _ {H}h H; set f2 := (finfun _) in H. +suff {H}: f2 = pffun (tperm i j) f by move=><-. rewrite {}/f2 {}/f1; apply/ffunP=>/= x; rewrite !ffunE /= ffunE /=. by case: tpermP=>[->|->|/eqP/negbTE->/eqP/negbTE->] {x}//; rewrite eqxx // Hij. Qed. diff --git a/examples/tree.v b/examples/tree.v index 542437f..9908835 100644 --- a/examples/tree.v +++ b/examples/tree.v @@ -122,12 +122,10 @@ Fixpoint mem_tree (t : tree A) : pred A := Definition tree_eqclass := tree A. Identity Coercion tree_of_eqclass : tree_eqclass >-> tree. Coercion pred_of_tree (t : tree_eqclass) : {pred A} := mem_tree t. - Canonical tree_predType := ssrbool.PredType (pred_of_tree : tree A -> pred A). -(* The line below makes mem_tree a canonical instance of topred. *) -Canonical mem_tree_predType := ssrbool.PredType mem_tree. -Lemma in_tnode a t ts : (t \in TNode a ts) = (t == a) || has (fun q => t \in q) ts. +Lemma in_tnode a t ts : + (t \in TNode a ts) = (t == a) || has (fun q => t \in q) ts. Proof. by []. Qed. (* frequently used facts about membership in a tree *) diff --git a/htt/domain.v b/htt/domain.v index 8a14f09..cbb13ef 100644 --- a/htt/domain.v +++ b/htt/domain.v @@ -522,8 +522,7 @@ Arguments chain T%_type. (* \In notation *) Definition chain_pred (T : poset) (x : chain T) : Pred T := x. -Canonical chainPredType (T : poset) := - Eval hnf in @mkPredType T (@chain T) (@chain_pred T). +Canonical chainPredType (T : poset) := PropPredType (@chain_pred T). Lemma chainE (T : poset) (s1 s2 : chain T) : s1 = s2 <-> chain_pred s1 =p chain_pred s2. @@ -557,9 +556,9 @@ Qed. Section ChainConst. Variables (T1 T2 : poset) (y : T2). -Definition const_chain : Pred _ := Pred1 y. +Definition const_chain : Pred _ := xPred1 y. Lemma const_is_chain : chain_axiom const_chain. -Proof. by split; [exists y | move=>x1 x2 <-<-; left]. Qed. +Proof. by split; [exists y | move=>x1 x2 ->->; left]. Qed. HB.instance Definition _ := isChain.Build T2 const_chain const_is_chain. Lemma const_chainE (s : chain T1) : Image (const_fun y) s =p const_chain. diff --git a/pcm/heap.v b/pcm/heap.v index be1ee5e..0364474 100644 --- a/pcm/heap.v +++ b/pcm/heap.v @@ -19,8 +19,8 @@ limitations under the License. From HB Require Import structures. From Coq Require Import ssreflect ssrbool ssrfun Eqdep. -From mathcomp Require Import ssrnat eqtype fintype seq path bigop. -From pcm Require Import options axioms finmap. +From mathcomp Require Import ssrnat eqtype fintype tuple finfun seq path bigop. +From pcm Require Import options axioms prelude finmap. From pcm Require Import pcm unionmap natmap. (************) @@ -366,6 +366,16 @@ Qed. End BlockUpdate. +Lemma updi_split {I : finType} T p k (f : {ffun I -> T}) : + updi p (fgraph f) = updi p (take (indx k) (fgraph f)) \+ + p.+(indx k) :-> f k \+ + updi (p.+(indx k).+1) (drop (indx k).+1 (fgraph f)). +Proof. +rewrite fgraph_codom /= codomE {1}(enum_split k) map_cat updi_cat /=. +rewrite map_take map_drop size_takel ?joinA; first by rewrite -ptr1 ptrA addn1. +by rewrite size_map index_size. +Qed. + Lemma domeqUP A1 A2 x (xs1 : seq A1) (xs2 : seq A2) : size xs1 = size xs2 -> dom_eq (updi x xs1) (updi x xs2). @@ -536,76 +546,4 @@ case: xs=>[|x xs] /= V H; first by case: H. by case: H V=>p [h'][-> _] /validPtUn_cond. Qed. -(******************************) -(******************************) -(* Custom lemmas about arrays *) -(******************************) -(******************************) - -From mathcomp Require Import fintype tuple finfun. - -Definition indx {I : finType} (x : I) := index x (enum I). - -Prenex Implicits indx. - -(***********************************) -(* Arrays indexed by a finite type *) -(***********************************) - -Section Array. -Variables (p : ptr) (I : finType). - -(* helper lemmas *) - -Lemma enum_split k : - enum I = take (indx k) (enum I) ++ k :: drop (indx k).+1 (enum I). -Proof. -rewrite -{2}(@nth_index I k k (enum I)) ?mem_enum //. -by rewrite -drop_nth ?index_mem ?mem_enum // cat_take_drop. -Qed. - -Lemma updi_split T k (f : {ffun I -> T}) : - updi p (fgraph f) = updi p (take (indx k) (fgraph f)) \+ - p.+(indx k) :-> f k \+ - updi (p.+(indx k).+1) (drop (indx k).+1 (fgraph f)). -Proof. -rewrite fgraph_codom /= codomE {1}(enum_split k) map_cat updi_cat /=. -rewrite map_take map_drop size_takel ?joinA; first by rewrite -ptr1 ptrA addn1. -by rewrite size_map index_size. -Qed. - -Lemma takeord T k x (f : {ffun I -> T}) : - take (indx k) (fgraph [ffun y => [eta f with k |-> x] y]) = - take (indx k) (fgraph f). -Proof. -set f' := (finfun _). -suff E: {in take (indx k) (enum I), f =1 f'}. -- by rewrite !fgraph_codom /= !codomE -2!map_take; move/eq_in_map: E. -move: (enum_uniq I); rewrite {1}(enum_split k) cat_uniq /= =>H4. -move=>y H5; rewrite /f' /= !ffunE /=; case: eqP H5 H4=>// -> ->. -by rewrite andbF. -Qed. - -Lemma dropord T k x (f : {ffun I -> T}) : - drop (indx k).+1 (fgraph [ffun y => [eta f with k |->x] y]) = - drop (indx k).+1 (fgraph f). -Proof. -set f' := (finfun _). -suff E: {in drop (indx k).+1 (enum I), f =1 f'}. -- by rewrite !fgraph_codom /= !codomE -2!map_drop; move/eq_in_map: E. -move: (enum_uniq I); rewrite {1}(enum_split k) cat_uniq /= => H4. -move=>y H5; rewrite /f' /= !ffunE /=; case: eqP H5 H4=>// -> ->. -by rewrite !andbF. -Qed. - -Lemma size_fgraph T1 T2 (r1 : {ffun I -> T1}) (r2 : {ffun I -> T2}) : - size (fgraph r1) = size (fgraph r2). -Proof. by rewrite !fgraph_codom /= !codomE !size_map. Qed. - -Lemma fgraphE T (r1 r2 : {ffun I -> T}) : - fgraph r1 = fgraph r2 -> r1 = r2. -Proof. -by move=> eq_r12; apply/ffunP=> x; rewrite -[x]enum_rankK -!tnth_fgraph eq_r12. -Qed. -End Array. diff --git a/pcm/natmap.v b/pcm/natmap.v index 1e732f4..41e582a 100644 --- a/pcm/natmap.v +++ b/pcm/natmap.v @@ -40,8 +40,8 @@ HB.structure Definition NatMap V := { (* pred structure (to write x \In h) is copied from union_map *) Canonical natmap_PredType A (U : natmap A) : PredType (nat * A) := - Mem_UmMap_PredType U. -Coercion Pred_of_natmap A (U : natmap A) (x : U) : Pred_Class := + um_PredType U. +Coercion Pred_of_natmap A (U : natmap A) (x : U) : {Pred _} := [eta Mem_UmMap x]. (* canonical natmap is histories *) @@ -81,8 +81,8 @@ HB.instance Definition _ A := isNatMap.Build A (history A). HB.instance Definition _ (A : eqType) := hasDecEq.Build (history A) (@union_map_eqP nat _ A (history A)). Canonical history_PredType A : PredType (nat * A) := - Mem_UmMap_PredType (history A). -Coercion Pred_of_history A (x : history A) : Pred_Class := + um_PredType (history A). +Coercion Pred_of_history A (x : history A) : {Pred _} := [eta Mem_UmMap x]. Notation "x \-> v" := (ptsT (history _) x v) (at level 30). @@ -145,8 +145,8 @@ HB.instance Definition _ A := isNatMap.Build A (nmap A). HB.instance Definition _ (A : eqType) := hasDecEq.Build (nmap A) (@union_map_eqP nat _ A (nmap A)). Canonical nmap_PredType A : PredType (nat * A) := - Mem_UmMap_PredType (nmap A). -Coercion Pred_of_nmap A (x : nmap A) : Pred_Class := + um_PredType (nmap A). +Coercion Pred_of_nmap A (x : nmap A) : {Pred _} := [eta Mem_UmMap x]. (* no points-to notation for nmaps *) diff --git a/pcm/unionmap.v b/pcm/unionmap.v index 64ad854..9e70998 100644 --- a/pcm/unionmap.v +++ b/pcm/unionmap.v @@ -2262,17 +2262,7 @@ Implicit Type f : U. Definition Mem_UmMap f := fun x : K * V => valid f /\ [pcm pts x.1 x.2 <= f]. - -(* DEVCOMMENT *) -(* is this needed? *) -(* Coercion Pred_of_Mem_UmMap (f : U) : Pred_Class := [eta Mem_UmMap f]. *) -(* /DEVCOMMENT *) -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. +Canonical um_PredType := PropPredType Mem_UmMap. Lemma In_undef x : x \In (undef : U) -> False. Proof. by case; rewrite valid_undef. Qed. @@ -2510,12 +2500,12 @@ Prenex Implicits InPt In_eta InPtUn In_dom InPtUnEL In_findE. (* Their membership structures aren't directly inheritted from that *) (* of union_map, but must be declared separately *) Canonical umap_PredType (K : ordType) V : PredType (K * V) := - Mem_UmMap_PredType (umap K V). -Coercion Pred_of_umap K V (x : umap K V) : Pred_Class := + um_PredType (umap K V). +Coercion Pred_of_umap K V (x : umap K V) : {Pred _} := [eta Mem_UmMap x]. Canonical fset_PredType (K : ordType) : PredType (K * unit) := - Mem_UmMap_PredType (fset K). -Coercion Pred_of_fset K (x : fset K) : Pred_Class := + um_PredType (fset K). +Coercion Pred_of_fset K (x : fset K) : {Pred _} := [eta Mem_UmMap x]. Section MorphMembership. From 0084f0a01220e180ef5a7c2c4d4bbbd6a7011e2f Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Fri, 16 Aug 2024 07:18:38 +0200 Subject: [PATCH 56/93] removed some files that shouldn't be here --- Unnamed_coqscript_1.crashcoqide | 0 Unnamed_coqscript_2.crashcoqide | 2 - core/prelude.v.crashcoqide | 1162 ------------------------------- examples/schorr.v.crashcoqide | 142 ---- htt/domain.v.crashcoqide | 1038 --------------------------- 5 files changed, 2344 deletions(-) delete mode 100644 Unnamed_coqscript_1.crashcoqide delete mode 100644 Unnamed_coqscript_2.crashcoqide delete mode 100644 core/prelude.v.crashcoqide delete mode 100644 examples/schorr.v.crashcoqide delete mode 100644 htt/domain.v.crashcoqide diff --git a/Unnamed_coqscript_1.crashcoqide b/Unnamed_coqscript_1.crashcoqide deleted file mode 100644 index e69de29..0000000 diff --git a/Unnamed_coqscript_2.crashcoqide b/Unnamed_coqscript_2.crashcoqide deleted file mode 100644 index 8f0fbb9..0000000 --- a/Unnamed_coqscript_2.crashcoqide +++ /dev/null @@ -1,2 +0,0 @@ -From HB Require Import structures. -From mathcomp Require Import ssreflect ssrfun ssrbool. diff --git a/core/prelude.v.crashcoqide b/core/prelude.v.crashcoqide deleted file mode 100644 index f286e8f..0000000 --- a/core/prelude.v.crashcoqide +++ /dev/null @@ -1,1162 +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 HB Require Import structures. -From Coq Require Import ssreflect ssrfun Eqdep. -From mathcomp Require Import ssrbool ssrnat seq eqtype choice. -From mathcomp Require Import fintype finfun tuple perm fingroup. -From pcm Require Import options axioms. - -(***********) -(* Prelude *) -(***********) - -(* 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}. - -(* 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}. - -(* 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. - -Lemma trans_eq A (x y z : A) : - x = y -> - x = z -> - y = z. -Proof. by move/esym; apply: eq_trans. 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. - -(************) -(* Products *) -(************) - -Inductive Prod3 U1 U2 U3 := - mk3 {proj31 : U1; proj32 : U2; proj33 : U3}. -Inductive Prod4 U1 U2 U3 U4 := - mk4 {proj41 : U1; proj42 : U2; proj43 : U3; proj44 : U4}. -Inductive Prod5 U1 U2 U3 U4 U5 := - mk5 {proj51 : U1; proj52 : U2; proj53 : U3; proj54 : U4; proj55 : U5}. -Inductive Prod6 U1 U2 U3 U4 U5 U6 := - mk6 {proj61 : U1; proj62 : U2; proj63 : U3; - proj64 : U4; proj65 : U5; proj66 : U6}. -Inductive Prod7 U1 U2 U3 U4 U5 U6 U7 := - mk7 {proj71 : U1; proj72 : U2; proj73 : U3; - proj74 : U4; proj75 : U5; proj76 : U6; proj77 : U7}. -Prenex Implicits proj31 proj32 proj33. -Prenex Implicits proj41 proj42 proj43 proj44. -Prenex Implicits proj51 proj52 proj53 proj54 proj55. -Prenex Implicits proj61 proj62 proj63 proj64 proj65 proj66. -Prenex Implicits proj71 proj72 proj73 proj74 proj75 proj76 proj77. - -Definition eq3 (U1 U2 U3 : eqType) (x y : Prod3 U1 U2 U3) := - [&& proj31 x == proj31 y, proj32 x == proj32 y & proj33 x == proj33 y]. -Definition eq4 (U1 U2 U3 U4 : eqType) (x y : Prod4 U1 U2 U3 U4) := - [&& proj41 x == proj41 y, proj42 x == proj42 y, proj43 x == proj43 y & - proj44 x == proj44 y]. -Definition eq5 (U1 U2 U3 U4 U5 : eqType) (x y : Prod5 U1 U2 U3 U4 U5) := - [&& proj51 x == proj51 y, proj52 x == proj52 y, proj53 x == proj53 y, - proj54 x == proj54 y & proj55 x == proj55 y]. -Definition eq6 (U1 U2 U3 U4 U5 U6 : eqType) (x y : Prod6 U1 U2 U3 U4 U5 U6) := - [&& proj61 x == proj61 y, proj62 x == proj62 y, proj63 x == proj63 y, - proj64 x == proj64 y, proj65 x == proj65 y & proj66 x == proj66 y]. -Definition eq7 (U1 U2 U3 U4 U5 U6 U7 : eqType) (x y : Prod7 U1 U2 U3 U4 U5 U6 U7) := - [&& proj71 x == proj71 y, proj72 x == proj72 y, proj73 x == proj73 y, - proj74 x == proj74 y, proj75 x == proj75 y, proj76 x == proj76 y & - proj77 x == proj77 y]. - -Lemma eq3P U1 U2 U3 : - Equality.axiom (@eq3 U1 U2 U3). -Proof. -rewrite /eq3; case=>x1 x2 x3 [y1 y2 y3] /=. -by do ![case: eqP=>[->|]]; constructor=>//; case. -Qed. - -Lemma eq4P U1 U2 U3 U4 : - Equality.axiom (@eq4 U1 U2 U3 U4). -Proof. -rewrite /eq4; case=>x1 x2 x3 x4 [y1 y2 y3 y4] /=. -by do ![case: eqP=>[->|]]; constructor=>//; case. -Qed. - -Lemma eq5P U1 U2 U3 U4 U5 : - Equality.axiom (@eq5 U1 U2 U3 U4 U5). -Proof. -rewrite /eq5; case=>x1 x2 x3 x4 x5 [y1 y2 y3 y4 y5] /=. -by do ![case: eqP=>[->|]]; constructor=>//; case. -Qed. - -Lemma eq6P U1 U2 U3 U4 U5 U6 : - Equality.axiom (@eq6 U1 U2 U3 U4 U5 U6). -Proof. -rewrite /eq6; case=>x1 x2 x3 x4 x5 x6 [y1 y2 y3 y4 y5 y6] /=. -by do ![case: eqP=>[->|]]; constructor=>//; case. -Qed. - -Lemma eq7P U1 U2 U3 U4 U5 U6 U7 : - Equality.axiom (@eq7 U1 U2 U3 U4 U5 U6 U7). -Proof. -rewrite /eq7; case=>x1 x2 x3 x4 x5 x6 x7 [y1 y2 y3 y4 y5 y6 y7] /=. -by do ![case: eqP=>[->|]]; constructor=>//; case. -Qed. - -HB.instance Definition _ (U1 U2 U3 : eqType) := - hasDecEq.Build (Prod3 U1 U2 U3) (@eq3P U1 U2 U3). -HB.instance Definition _ (U1 U2 U3 U4 : eqType) := - hasDecEq.Build (Prod4 U1 U2 U3 U4) (@eq4P U1 U2 U3 U4). -HB.instance Definition _ (U1 U2 U3 U4 U5 : eqType) := - hasDecEq.Build (Prod5 U1 U2 U3 U4 U5) (@eq5P U1 U2 U3 U4 U5). -HB.instance Definition _ (U1 U2 U3 U4 U5 U6 : eqType) := - hasDecEq.Build (Prod6 U1 U2 U3 U4 U5 U6) (@eq6P U1 U2 U3 U4 U5 U6). -HB.instance Definition _ (U1 U2 U3 U4 U5 U6 U7 : eqType) := - hasDecEq.Build (Prod7 U1 U2 U3 U4 U5 U6 U7) (@eq7P U1 U2 U3 U4 U5 U6 U7). - -(***************************) -(* 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. - -Lemma compE A B C (g : B -> C) (f : A -> B) x : - g (f x) = (g \o f) x. -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, aka. fork/fanout/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. - -(************************) -(* 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 , P6 , P7 , P8 , P9 , P10 & P11 ]" (at level 0, format - "'[hv' [ /\ '[' P1 , '/' P2 , '/' P3 , '/' P4 , '/' P5 , '/' P6 , '/' P7 , '/' P8 , '/' P9 , '/' P10 ']' '/ ' & P11 ] ']'"). -Reserved Notation "[ /\ P1 , P2 , P3 , P4 , P5 , P6 , P7 , P8 , P9 , P10 , P11 & P12 ]" (at level 0, format - "'[hv' [ /\ '[' P1 , '/' P2 , '/' P3 , '/' P4 , '/' P5 , '/' P6 , '/' P7 , '/' P8 , '/' P9 , '/' P10 , '/' P11 ']' '/ ' & P12 ] ']'"). -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 and11 (P1 P2 P3 P4 P5 P6 P7 P8 P9 P10 P11 : Prop) : Prop := - And11 of P1 & P2 & P3 & P4 & P5 & P6 & P7 & P8 & P9 & P10 & P11. -Inductive and12 (P1 P2 P3 P4 P5 P6 P7 P8 P9 P10 P11 P12 : Prop) : Prop := - And12 of P1 & P2 & P3 & P4 & P5 & P6 & P7 & P8 & P9 & P10 & P11 & P12. - -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 , P6 , P7 , P8 , P9 , P10 & P11 ]" := - (and11 P1 P2 P3 P4 P5 P6 P7 P8 P9 P10 P11) : type_scope. -Notation "[ /\ P1 , P2 , P3 , P4 , P5 , P6 , P7 , P8 , P9 , P10 , P11 & P12 ]" := - (and12 P1 P2 P3 P4 P5 P6 P7 P8 P9 P10 P11 P12) : 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 *) - -(* DEVCOMMENT *) -(* TODO: we should move some of the following to [ssrbool] in Coq *) -(* /DEVCOMMENT *) - -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 *) -#[export] Program Instance and3_impl_morphism : Proper (iter_respectful impl impl 3) and3 | 1. -#[export] Program Instance and3_iff_morphism : Proper (iter_respectful iff iff 3) and3 | 1. - -#[export] Program Instance and4_impl_morphism : Proper (iter_respectful impl impl 4) and4 | 1. -#[export] Program Instance and4_iff_morphism : Proper (iter_respectful iff iff 4) and4 | 1. - -#[export] Program Instance and5_impl_morphism : Proper (iter_respectful impl impl 5) and5 | 1. -#[export] Program Instance and5_iff_morphism : Proper (iter_respectful iff iff 5) and5 | 1. - -#[export] Program Instance and6_impl_morphism : Proper (iter_respectful impl impl 6) and6 | 1. -#[export] Program Instance and6_iff_morphism : Proper (iter_respectful iff iff 6) and6 | 1. - -#[export] Program Instance and7_impl_morphism : Proper (iter_respectful impl impl 7) and7 | 1. -#[export] Program Instance and7_iff_morphism : Proper (iter_respectful iff iff 7) and7 | 1. - -#[export] Program Instance and8_impl_morphism : Proper (iter_respectful impl impl 8) and8 | 1. -#[export] Program Instance and8_iff_morphism : Proper (iter_respectful iff iff 8) and8 | 1. - -#[export] Program Instance and9_impl_morphism : Proper (iter_respectful impl impl 9) and9 | 1. -#[export] Program Instance and9_iff_morphism : Proper (iter_respectful iff iff 9) and9 | 1. - -#[export] Program Instance and10_impl_morphism : Proper (iter_respectful impl impl 10) and10 | 1. -#[export] Program Instance and10_iff_morphism : Proper (iter_respectful iff iff 10) and10 | 1. - -#[export] Program Instance and11_impl_morphism : Proper (iter_respectful impl impl 11) and11 | 1. -#[export] Program Instance and11_iff_morphism : Proper (iter_respectful iff iff 11) and11 | 1. - -#[export] Program Instance and12_impl_morphism : Proper (iter_respectful impl impl 12) and12 | 1. -#[export] Program Instance and12_iff_morphism : Proper (iter_respectful iff iff 12) and12 | 1. - -(** Logical disjunction *) -#[export] Program Instance or3_impl_morphism : Proper (iter_respectful impl impl 3) or3 | 1. -#[export] Program Instance or3_iff_morphism : Proper (iter_respectful iff iff 3) or3 | 1. - -#[export] Program Instance or4_impl_morphism : Proper (iter_respectful impl impl 4) or4 | 1. -#[export] Program Instance or4_iff_morphism : Proper (iter_respectful iff iff 4) or4 | 1. - -#[export] Program Instance or5_impl_morphism : Proper (iter_respectful impl impl 5) or5 | 1. -#[export] Program Instance or5_iff_morphism : Proper (iter_respectful iff iff 5) or5 | 1. - -#[export] Program Instance or6_impl_morphism : Proper (iter_respectful impl impl 6) or6 | 1. -#[export] Program Instance or6_iff_morphism : Proper (iter_respectful iff iff 6) or6 | 1. - -#[export] Program Instance or7_impl_morphism : Proper (iter_respectful impl impl 7) or7 | 1. -#[export] 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 b11 b12 : bool. - -Lemma and6P : reflect [/\ b1, b2, b3, b4, b5 & b6] [&& b1, b2, b3, b4, b5 & b6]. -Proof. -case: b1=>/=; last by constructor; case. -case: b2=>/=; last by constructor; case. -case: b3=>/=; last by constructor; case. -case: b4=>/=; last by constructor; case. -case: b5=>/=; last by constructor; case. -case: b6=>/=; last by constructor; case. -by constructor. -Qed. - -Lemma and7P : reflect [/\ b1, b2, b3, b4, b5, b6 & b7] [&& b1, b2, b3, b4, b5, b6 & b7]. -Proof. -case: b1=>/=; last by constructor; case. -case: b2=>/=; last by constructor; case. -case: b3=>/=; last by constructor; case. -case: b4=>/=; last by constructor; case. -case: b5=>/=; last by constructor; case. -case: b6=>/=; last by constructor; case. -case: b7=>/=; last by constructor; case. -by constructor. -Qed. - -Lemma and8P : reflect [/\ b1, b2, b3, b4, b5, b6, b7 & b8] [&& b1, b2, b3, b4, b5, b6, b7 & b8]. -Proof. -case: b1=>/=; last by constructor; case. -case: b2=>/=; last by constructor; case. -case: b3=>/=; last by constructor; case. -case: b4=>/=; last by constructor; case. -case: b5=>/=; last by constructor; case. -case: b6=>/=; last by constructor; case. -case: b7=>/=; last by constructor; case. -case: b8=>/=; last by constructor; case. -by constructor. -Qed. - -Lemma and9P : reflect [/\ b1, b2, b3, b4, b5, b6, b7, b8 & b9] [&& b1, b2, b3, b4, b5, b6, b7, b8 & b9]. -Proof. -case: b1=>/=; last by constructor; case. -case: b2=>/=; last by constructor; case. -case: b3=>/=; last by constructor; case. -case: b4=>/=; last by constructor; case. -case: b5=>/=; last by constructor; case. -case: b6=>/=; last by constructor; case. -case: b7=>/=; last by constructor; case. -case: b8=>/=; last by constructor; case. -case: b9=>/=; last by constructor; case. -by constructor. -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. -case: b1=>/=; last by constructor; case. -case: b2=>/=; last by constructor; case. -case: b3=>/=; last by constructor; case. -case: b4=>/=; last by constructor; case. -case: b5=>/=; last by constructor; case. -case: b6=>/=; last by constructor; case. -case: b7=>/=; last by constructor; case. -case: b8=>/=; last by constructor; case. -case: b9=>/=; last by constructor; case. -case: b10=>/=; last by constructor; case. -by constructor. -Qed. - -Lemma and11P : reflect [/\ b1, b2, b3, b4, b5, b6, b7, b8, b9, b10 & b11] - [&& b1, b2, b3, b4, b5, b6, b7, b8, b9, b10 & b11]. -Proof. -case: b1=>/=; last by constructor; case. -case: b2=>/=; last by constructor; case. -case: b3=>/=; last by constructor; case. -case: b4=>/=; last by constructor; case. -case: b5=>/=; last by constructor; case. -case: b6=>/=; last by constructor; case. -case: b7=>/=; last by constructor; case. -case: b8=>/=; last by constructor; case. -case: b9=>/=; last by constructor; case. -case: b10=>/=; last by constructor; case. -case: b11=>/=; last by constructor; case. -by constructor. -Qed. - -Lemma and12P : reflect [/\ b1, b2, b3, b4, b5, b6, b7, b8, b9, b10, b11 & b12] - [&& b1, b2, b3, b4, b5, b6, b7, b8, b9, b10, b11 & b12]. -Proof. -case: b1=>/=; last by constructor; case. -case: b2=>/=; last by constructor; case. -case: b3=>/=; last by constructor; case. -case: b4=>/=; last by constructor; case. -case: b5=>/=; last by constructor; case. -case: b6=>/=; last by constructor; case. -case: b7=>/=; last by constructor; case. -case: b8=>/=; last by constructor; case. -case: b9=>/=; last by constructor; case. -case: b10=>/=; last by constructor; case. -case: b11=>/=; last by constructor; case. -case: b12=>/=; last by constructor; case. -by constructor. -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 and11P {b1 b2 b3 b4 b5 b6 b7 b8 b9 b10 b11}. -Arguments and12P {b1 b2 b3 b4 b5 b6 b7 b8 b9 b10 b11 b12}. - -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}. - -Lemma iffPb (b1 b2 : bool) : reflect (b1 <-> b2) (b1 == b2). -Proof. -case: eqP=>[->|N]; constructor=>//. -case: b1 b2 N; case=>//= _. -- by case=>/(_ erefl). -by case=>_ /(_ erefl). -Qed. - -Lemma implyb_trans a b c : - a ==> b -> - b ==> c -> - a ==> c. -Proof. by case: a=>//=->. Qed. - -Lemma iffE (b1 b2 : bool) : b1 = b2 <-> (b1 <-> b2). -Proof. by split=>[->|] //; move/iffPb/eqP. Qed. - -(* subsets *) - -Lemma subsetC T (p q : mem_pred T) : - {subset p <= q} -> - {subset predC q <= predC p}. -Proof. by move=>H x; apply: contra (H x). Qed. - -Lemma subsetR T (p q : mem_pred T) : - {subset p <= predC q} -> - {subset q <= predC p}. -Proof. by move=>H x; apply: contraL (H x). Qed. - -Lemma subsetL T (p q : mem_pred T) : - {subset predC p <= q} -> - {subset predC q <= p}. -Proof. by move=>H x; apply: contraR (H x). Qed. - -Lemma subsetLR T (p q : mem_pred T) : - {subset predC p <= predC q} -> - {subset q <= p}. -Proof. by move=>H x; apply: contraLR (H x). Qed. - -Lemma subset_disj T (p q r : mem_pred T) : - {subset p <= q} -> - (forall x, x \in q -> x \in r -> false) -> - (forall x, x \in p -> x \in r -> false). -Proof. by move=>H D x /H/D. Qed. - -Lemma subset_disj2 T (p p1 q q1 : mem_pred T) : - {subset p1 <= p} -> - {subset q1 <= q} -> - (forall x, x \in p -> x \in q -> false) -> - (forall x, x \in p1 -> x \in q1 -> false). -Proof. -move=>H1 H2 D1; apply: subset_disj H1 _ => x H /H2. -by apply: D1 H. -Qed. - -(**************) -(* empty type *) -(**************) - -Lemma False_eqP : Equality.axiom (fun _ _ : False => true). -Proof. by case. Qed. - -HB.instance Definition _ := hasDecEq.Build False False_eqP. - -(*************) -(* 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. - -#[export] Hint Extern 0 (ocancel _ _) => - (apply: lft_inl_ocanc || apply: rgt_inr_ocanc) : core. - -(****************) -(* option types *) -(****************) - -(* Alternative mechanism for manipulating *) -(* properties of the form isSome X. *) -(* To use a lemma of the form isSome X *) -(* one tyoically has to explicitly put the conclusion *) -(* of that lemma onto the context, and then case on X. *) -(* If such lemma is restated using is_some X *) -(* then it suffices to case on the lemma's name. *) -(* This saves typing X explicitly, which *) -(* may be significant if X is large. *) - -Inductive is_some_spec A x : option A -> Prop := -| is_some_case v of x = Some v : is_some_spec x (Some v). - -Hint Resolve is_some_case : core. - -Notation is_some x := (is_some_spec x x). - -Lemma is_someP A (x : option A) : reflect (is_some x) (isSome x). -Proof. by case: x=>[a|]; constructor=>//; case. Qed. - -(* some simplifications *) -Lemma oapp_some A (x : option A) : oapp [eta Some] None x = x. -Proof. by case: x. Qed. - -Lemma obind_some A (x : option A) : obind [eta Some] x = x. -Proof. exact: oapp_some. Qed. - -(********) -(* nats *) -(********) - -Lemma gt0 m n : m < n -> 0 < n. -Proof. by case: n. Qed. - -Lemma neq0 m n : m < n -> n != 0. -Proof. by move/gt0; rewrite lt0n. Qed. - -Lemma neqSn n : n.+1 != n. -Proof. by elim: n. Qed. - -Lemma neqnS n : n != n.+1. -Proof. by elim: n. Qed. - -(**************************************) -(* Inhabited (non-empty) finite types *) -(**************************************) - -Lemma inhabits (T : finType) (t : T) : 0 < #|T|. -Proof. by apply/card_gt0P; exists t. Qed. - -Lemma inhabits_irr (T : finType) (t1 t2 : T) : - inhabits t1 = inhabits t2. -Proof. by apply: bool_irrelevance. Qed. - -Definition inhabited_axiom (T : finType) := 0 < #|T|. - -HB.mixin Record isInhabited T of Finite T := { - card_inhab : inhabited_axiom T}. - -#[short(type="ifinType")] -HB.structure Definition Inhabited := {T of isInhabited T & Finite T}. - -Lemma inhabF (T : ifinType) : ~ (@predT T =1 xpred0). -Proof. by case/card_gt0P: (@card_inhab T)=>x _ /(_ x). Qed. - -Definition inhab {T : ifinType} : T := - match pickP predT with - | Pick x _ => x - | Nopick pf => False_rect T (inhabF pf) - end. - -HB.instance Definition _ := isInhabited.Build unit (inhabits tt). -HB.instance Definition _ := isInhabited.Build bool (inhabits false). -HB.instance Definition _ n := isInhabited.Build 'I_n.+1 (inhabits ord0). -HB.instance Definition _ (T : finType) := - isInhabited.Build (option T) (inhabits None). - -(*************************************) -(* A copy of booleans with mnemonics *) -(* LL and RR for working with sides *) -(*************************************) - -Inductive Side := LL | RR. -Definition Side_eq x y := - match x, y with LL, LL => true | RR, RR => true | _, _ => false end. -Definition nat2Side x := if odd x then LL else RR. -Definition Side2nat x := if x is RR then 0 else 1. - -Lemma ssrcanc : ssrfun.cancel Side2nat nat2Side. Proof. by case. Qed. -HB.instance Definition _ : isCountable Side := CanIsCountable ssrcanc. - -Lemma Side_enumP : Finite.axiom [:: LL; RR]. Proof. by case. Qed. -HB.instance Definition _ : isFinite Side := isFinite.Build Side Side_enumP. -HB.instance Definition _ := isInhabited.Build Side (inhabits LL). - -(*************) -(* Sequences *) -(*************) - -(* folds *) -(* TODO upstream to mathcomp *) - -Lemma map_foldr {T1 T2} (f : T1 -> T2) xs : - map f xs = foldr (fun x ys => f x :: ys) [::] xs. -Proof. by []. Qed. - -Lemma fusion_foldr {T R Q} (g : R -> Q) f0 f1 z0 z1 (xs : seq T) : - (forall x y, g (f0 x y) = f1 x (g y)) -> - g z0 = z1 -> - g (foldr f0 z0 xs) = foldr f1 z1 xs. -Proof. by move=>Hf Hz; elim: xs=>//= x xs <-. Qed. - -Lemma fusion_foldl {T R Q} (g : R -> Q) f0 f1 z0 z1 (xs : seq T) : - (forall x y, g (f0 x y) = f1 (g x) y) -> - g z0 = z1 -> - g (foldl f0 z0 xs) = foldl f1 z1 xs. -Proof. -move=>Hf Hz; elim: xs z0 z1 Hz =>//= x xs IH z0 z1 Hz. -by apply: IH; rewrite Hf Hz. -Qed. - -Lemma foldl_foldr {T R} (f : R -> T -> R) z xs : - foldl f z xs = foldr (fun b g x => g (f x b)) id xs z. -Proof. by elim: xs z=>/=. Qed. - -Lemma foldr_foldl {T R} (f : T -> R -> R) z xs : - foldr f z xs = foldl (fun g b x => g (f b x)) id xs z. -Proof. -elim/last_ind: xs z=>//= xs x IH z. -by rewrite foldl_rcons -IH foldr_rcons. -Qed. - -(* pmap *) -(* TODO upstream to mathcomp *) -Lemma pmap_pcomp {S T U} (f : T -> option U) (g : S -> option T) s : - pmap (pcomp f g) s = pmap f (pmap g s). -Proof. by elim: s=>//= x s ->; rewrite /pcomp; case: (g x). Qed. - -(* sequence prefixes *) - -(* Two helper concepts for searching in sequences: *) -(* - onth: like nth, but returns None when the element is not found *) -(* - Prefix: a prefix relation on sequences, used for growing *) -(* interpretation contexts *) - -Fixpoint onth A (s : seq A) n : option A := - if s is x::sx then if n is nx.+1 then onth sx nx else Some x else None. - -Definition Prefix A (s1 s2 : seq A) : Prop := - forall n x, onth s1 n = some x -> onth s2 n = some x. - -(* Lemmas *) - -Section SeqPrefix. -Variable A : Type. -Implicit Type s : seq A. - -Variant onth_spec s n : bool -> Type := -| onth_none : onth s n = None -> onth_spec s n false -| onth_some v : onth s n = Some v -> onth_spec s n true. - -Lemma onth_sizeP s n : onth_spec s n (n < size s). -Proof. -elim: s n=>/= [|a s IH] n; first by rewrite ltn0; constructor. -case: n=>[|n] /=; first by apply: (@onth_some _ _ a). -rewrite ltnS; case: (IH n)=>[|v] H. -- by constructor. -by apply: (@onth_some _ _ v). -Qed. - -Lemma size_onth s n : n < size s -> exists x, onth s n = Some x. -Proof. -by case: onth_sizeP=>// v H _; exists v. -Qed. - -Lemma onth_size s n x : onth s n = Some x -> n < size s. -Proof. -by case: onth_sizeP=>//->. -Qed. - -Lemma size_onthPn s n : reflect (onth s n = None) (size s <= n). -Proof. -by rewrite leqNgt; apply: (iffP idP); case: onth_sizeP=>//= v ->. -Qed. - -Lemma onth_sizeN s : onth s (size s) = None. -Proof. by apply/size_onthPn. Qed. - -Lemma nth_onth x0 n s : nth x0 s n = odflt x0 (onth s n). -Proof. -elim: s n=>/= [|a s IH] n /=; first by apply: nth_nil. -by case: n. -Qed. - -Lemma onth_nth x0 n s : - n < size s -> - onth s n = Some (nth x0 s n). -Proof. -elim: s n=>//= a s IH n. -by rewrite ltnS; case: n. -Qed. - -Lemma Prefix_refl s : Prefix s s. -Proof. by move=>n x <-. Qed. - -Lemma Prefix_trans s2 s1 s3 : - Prefix s1 s2 -> Prefix s2 s3 -> Prefix s1 s3. -Proof. by move=>H1 H2 n x E; apply: H2; apply: H1. Qed. - -Lemma Prefix_cons x s1 s2 : Prefix (x :: s1) (x :: s2) <-> Prefix s1 s2. -Proof. by split=>E n; [apply: (E n.+1) | case: n]. Qed. - -Lemma Prefix_cons' x y s1 s2 : - Prefix (x :: s1) (y :: s2) -> x = y /\ Prefix s1 s2. -Proof. by move=>H; case: (H 0 x (erefl _)) (H)=>-> /Prefix_cons. Qed. - -Lemma Prefix_rcons x s : Prefix s (rcons s x). -Proof. by elim: s=>//= y ys IH; apply/Prefix_cons; apply: IH. Qed. - -Lemma Prefix_cat s1 s2 : Prefix s1 (s1 ++ s2). -Proof. -elim: s2 s1=>[|x xs IH] s1; first by rewrite cats0. -rewrite -cat_rcons; apply: Prefix_trans (IH _). -by apply: Prefix_rcons. -Qed. - -Lemma Prefix_size s1 s2 : Prefix s1 s2 -> size s1 <= size s2. -Proof. -elim: s1 s2=>[//|a s1 IH] [|b s2] H; first by move: (H 0 a (erefl _)). -by rewrite ltnS; apply: (IH _ (proj2 (Prefix_cons' H))). -Qed. - -Lemma Prefix_onth s t x : - x < size s -> - Prefix s t -> onth s x = onth t x. -Proof. -elim:s t x =>[//|a s IH] [|b t] x H1 H2; first by move: (H2 0 a (erefl _)). -by case/Prefix_cons': H2=><- H2; case: x H1=>[|n] //= H1; apply: IH. -Qed. - -Lemma PrefixE s1 s2 : Prefix s1 s2 <-> exists s3, s2 = s1 ++ s3. -Proof. -split; last by case=>s3 ->; apply: Prefix_cat. -elim: s1 s2=>[|x xs IH] s2; first by exists s2. -case: s2=>[/(_ 0 x erefl)//|y ys /Prefix_cons' [?]]. -by subst y=>/IH [s3 ->]; exists s3. -Qed. - -End SeqPrefix. - -#[export] Hint Resolve Prefix_refl : core. - -(* when A : eqType *) - -Section SeqPrefixEq. -Variable A : eqType. -Implicit Type s : seq A. - -Lemma onth_mem s n x : - onth s n = Some x -> - x \in s. -Proof. -by elim: s n=>//= a s IH [[->]|n /IH]; rewrite inE ?eq_refl // orbC =>->. -Qed. - -Lemma onth_index (s : seq A) x : - onth s (index x s) = - if x \in s then Some x else None. -Proof. -by elim: s=>//=h s IH; rewrite inE eq_sym; case: eqP=>//= ->. -Qed. - -Lemma PrefixP (s1 s2 : seq A) : - reflect (Prefix s1 s2) (prefix s1 s2). -Proof. -apply/(equivP (prefixP (s1:=s1) (s2:=s2))). -by apply: iff_sym; exact: PrefixE. -Qed. - -End SeqPrefixEq. - -(******************************) -(* Some commuting conversions *) -(******************************) - -Lemma fun_op A B C (b : option A) (vS : A -> B) (vN : B) (f : B -> C) : - f (if b is Some v then vS v else vN) = - if b is Some v then f (vS v) else f vN. -Proof. by case: b=>//. Qed. - -Lemma op_if A B (b : bool) (vS vN : option A) (vS1 : A -> B) (vN1 : B) : - (if (if b then vS else vN) is Some v then vS1 v else vN1) = - if b then if vS is Some v then vS1 v else vN1 - else if vN is Some v then vS1 v else vN1. -Proof. by case: b. Qed. - -Lemma if_triv b : (if b then true else false) = b. -Proof. by case: b. Qed. - -(*************************************************************) -(* quick ways to extract projections and transitivity proofs *) -(* out of iterated inequalities *) -(*************************************************************) - -Lemma ltn13 a b c : a < b < c -> a < c. -Proof. by case/andP; apply: ltn_trans. Qed. - -Lemma ltn12 a b c : a < b < c -> a < b. -Proof. by case/andP. Qed. - -Lemma ltn23 a b c : a < b < c -> b < c. -Proof. by case/andP. Qed. - -Lemma leq13 a b c : a <= b <= c -> a <= c. -Proof. by case/andP; apply: leq_trans. Qed. - -Lemma leq12 a b c : a <= b <= c -> a <= b. -Proof. by case/andP. Qed. - -Lemma leq23 a b c : a <= b <= c -> b <= c. -Proof. by case/andP. Qed. - -Lemma lqt13 a b c : a <= b < c -> a < c. -Proof. by case/andP; apply: leq_ltn_trans. Qed. - -Lemma lqt12 a b c : a <= b < c -> a <= b. -Proof. by case/andP. Qed. - -Lemma lqt23 a b c : a <= b < c -> b < c. -Proof. by case/andP. Qed. - -(********************) -(* Finite functions *) -(********************) - -Section FinFun. -Variables (T : finType) (Us : T -> Type). -Implicit Type f : {dffun forall t, Us t}. - -(* Explicit name for finite function application. *) -(* Will be used to hang canonical projections onto. *) -(* The function/argument order is reversed to facilitate rewriting. *) -Definition sel tg f : Us tg := f tg. - -Lemma ffinP f1 f2 : (forall t, sel t f1 = sel t f2) <-> f1 = f2. -Proof. by rewrite ffunP. Qed. - -(* beta equality *) -Lemma sel_fin t (f : forall t, Us t) : sel t (finfun f) = f t. -Proof. by rewrite /sel ffunE. Qed. - -(* eta equality *) -Lemma fin_eta f : f = finfun (sel^~ f). -Proof. by apply/ffinP=>t; rewrite sel_fin. Qed. - -(* function *) -Definition splice tg f (v : Us tg) : {dffun _} := - finfun (fun x => - if decP (x =P tg) is left pf then cast Us pf v - else sel x f). - -Lemma sel_splice t f x (v : Us x) : - sel t (splice f v) = - if decP (t =P x) is left pf then cast Us pf v - else sel t f. -Proof. by rewrite sel_fin. Qed. - -Lemma sel_spliceE t f v : sel t (splice f v) = v. -Proof. by rewrite sel_fin; case: eqP=>//= pf; rewrite eqd. Qed. - -Lemma sel_spliceN t x f (w : Us x) : - t <> x -> sel t (splice f w) = sel t f. -Proof. by move=>N; rewrite sel_fin; case: eqP. Qed. - -Lemma splice_eta t f : splice f (sel t f) = f. -Proof. -apply/ffinP=>x; rewrite sel_fin; case: eqP=>//=. -by move/[dup]=>-> pf; rewrite eqd. -Qed. - -End FinFun. - -Arguments sel {T Us} tg f. -Arguments splice {T Us tg} f v. - -(* notation for building finfuns *) -(* DEVCOMMENT *) -(* copied from finfun to fix some bad spacing in formatting *) -(* /DEVCOMMENT *) - -Notation "[ 'ffun' x : aT => E ]" := (finfun (fun x : aT => E)) - (at level 0, x name, format "[ 'ffun' x : aT => E ]") : fun_scope. - -Notation "[ 'ffun' x => E ]" := (@finfun _ (fun=> _) (fun x => E)) - (at level 0, x name, format "[ 'ffun' x => E ]") : fun_scope. - -Notation "['ffun' => E ]" := [ffun _ => E] - (at level 0, format "['ffun' => E ]") : fun_scope. - -Section IteratedNotation. -Variables (T : finType) (Us : T -> Type). - -Variant dfun_delta : Type := DFunDelta t of Us t. - -(* for iteration that starts with function ends with function *) -Definition dapp_fdelta df (f : forall t, Us t) z := - let: DFunDelta t v := df in - if decP (z =P t) is left pf then cast Us pf v - else f z. - -(* for iteration that starts with finfun ends with function *) -Definition splice' df (f : {ffun forall t, Us t}) z := - dapp_fdelta df f z. - -End IteratedNotation. - -Delimit Scope fun_delta_scope with FUN_DELTA. -Arguments dapp_fdelta {T Us} df%_FUN_DELTA f. - -Notation "y \\ x" := (@DFunDelta _ _ x y) (at level 1). - -(* notation for simultaneous update of f with d1,..,dn *) -(* rewrite by sel_fin peels all layers *) -Notation "[ 'ext' f 'with' d1 , .. , dn ]" := - (finfun ( - dapp_fdelta d1%_FUN_DELTA .. (dapp_fdelta dn%_FUN_DELTA f) ..)) - (at level 0, format - "'[hv' [ '[' 'ext' '/ ' f ']' '/' 'with' '[' d1 , '/' .. , '/' dn ']' ] ']'" - ) : fun_scope. - -(* notation for iterated update of f with d1, then d2, ... *) -(* rewrite by sel_fin peels top layer only *) -Notation "[ 'splice' F 'with' d1 , .. , dn ]" := - (finfun (splice' - d1%_FUN_DELTA .. (finfun (splice' dn%_FUN_DELTA (finfun F))) ..)) - (at level 0, format - "'[hv' [ '[' 'splice' '/ ' F ']' '/' 'with' '[' d1 , '/' .. , '/' dn ']' ] ']'" - ) : fun_scope. - -Section TestingNotation. -Variables (T : finType) (Us : T -> Type). -Variables (f : {dffun forall t, Us t}) (t1 t2 : T) (v1 : Us t1) (v2 : Us t2). - -(* we have three different options in displaying splices *) -(* splice, [splice], and [ext] *) -Lemma test : - [/\ sel t2 [splice f with v1 \\ t1, v2 \\ t2] = v2, - sel t2 (splice (splice f v1) v2) = v2, - sel t2 [ext f with v1 \\ t1, v2 \\ t2] = v2 & -(* and we can use underscores to elide some info *) - sel t2 [ext f with v1 \\ _, v2 \\ _] = v2]. -Abort. -End TestingNotation. - -(* mapping over simply-typed finite functions *) - -Section FinFunMap. -Variables (T : finType) (A B : Type). -Implicit Types (f : A -> B) (x : {ffun T -> A}). - -Definition fmap f x := [ffun tg => f (sel tg x)]. - -Lemma sel_fmap f x tg : sel tg (fmap f x) = f (sel tg x). -Proof. exact: sel_fin. Qed. - -Lemma fmap_splice f x tg (v : A) : - fmap f (splice (tg:=tg) x v) = splice (tg:=tg) (fmap f x) (f v). -Proof. -apply/ffinP=>t; rewrite !sel_fin; case: eqP=>//= ?; subst t. -by rewrite !eqc. -Qed. - -End FinFunMap. - -(* surgery on tuples and finfuns *) - -Section OnthCodom. -Variable A : Type. - -Lemma onth_tnth {n} (s : n.-tuple A) (i : 'I_n) : - onth s i = Some (tnth s i). -Proof. -elim: n s i =>[|n IH] s i; first by case: i. -case/tupleP: s=>/=x s; case: (unliftP ord0 i)=>[j|]-> /=. -- by rewrite tnthS. -by rewrite tnth0. -Qed. - -Lemma onth_codom {n} (i : 'I_n) (f: {ffun 'I_n -> A}) : - onth (fgraph f) i = Some (f i). -Proof. -pose i' := cast_ord (esym (card_ord n)) i. -move: (@tnth_fgraph _ _ f i'); rewrite (enum_val_ord) {2}/i' cast_ordKV=><-. -by rewrite (onth_tnth (fgraph f) i'). -Qed. - -End OnthCodom. - -(* ffun and permutation *) -Section PermFfun. -Variables (I : finType) (A : Type). - -Definition pffun (p : {perm I}) (f : {ffun I -> A}) := - [ffun i => f (p i)]. - -Lemma pffunE1 (f : {ffun I -> A}) : pffun 1%g f = f. -Proof. by apply/ffunP=>i; rewrite !ffunE permE. Qed. - -Lemma pffunEM (p p' : {perm I}) (f : {ffun I -> A}) : - pffun (p * p') f = pffun p (pffun p' f). -Proof. by apply/ffunP => i; rewrite !ffunE permM. Qed. - -End PermFfun. - - -(* Tagging *) - -Notation Tag := (@existT _ _). - -Lemma Tag_inj T Us (t1 t2 : T) i1 i2 : - Tag t1 i1 = Tag t2 i2 -> - t1 = t2 /\ jmeq Us i1 i2. -Proof. by case=>?; subst t2=>/inj_pair2 ->. Qed. -Arguments Tag_inj {T Us t1 t2 i1 i2}. - -(* tagged union of equality types is equality type *) - -Section TaggedEq. -Variables (T : eqType) (Us : T -> eqType). - -Definition tag_eq : sigT Us -> sigT Us -> bool := - fun '(Tag tx opx) '(Tag ty opy) => - if decP (tx =P ty) is left pf then opx == cast Us pf opy - else false. - -Lemma tag_eqP : Equality.axiom tag_eq. -Proof. -case=>tx opx [ty opy] /=; case: (tx =P ty)=>pf; last first. -- by constructor; case=>/pf. -subst ty; rewrite /= eqc; case: eqP=>pf; constructor; -by [rewrite pf|case=>/inj_pair2/pf]. -Qed. - -HB.instance Definition _ := hasDecEq.Build (sigT Us) tag_eqP. -End TaggedEq. - diff --git a/examples/schorr.v.crashcoqide b/examples/schorr.v.crashcoqide deleted file mode 100644 index 3207c6f..0000000 --- a/examples/schorr.v.crashcoqide +++ /dev/null @@ -1,142 +0,0 @@ -(* -Copyright 2024 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 HB Require Import structures. -From Coq Require Import ssreflect ssrfun. -From mathcomp Require Import eqtype ssrnat ssrbool seq path bigop. -From htt Require Import options model heapauto graph. -From pcm Require Import options axioms pred prelude seqperm seqext. -From pcm Require Import pcm unionmap natmap heap autopcm automap. - - -(*****************) -(* helper lemmas *) -(*****************) - -Lemma eq_connect_aux (g1 g2 : pregraph) p: - valid (g1 \+ g2) -> - {subset dom g2 <= predC p} -> - um_filterk p (g1 \+ g2) = um_filterk p g1. -Proof. -move=>V /(umfiltk_subdom0 p (validR V)) S. -by rewrite umfiltUn // S unitR. -Qed. - -Lemma connectUn_eq (g g1 g2 : pregraph) (p : pred node) x : - valid (g \+ g1) -> - valid (g \+ g2) -> - {subset dom g1 <= p} -> - {subset dom g2 <= p} -> - connect p (g \+ g1) x =i connect p (g \+ g2) x. -Proof. -move=>V1 V2 S1 S2 z. -rewrite -connect_umfiltk eq_connect_aux //=; last first. -- by move=>y /S1; rewrite inE negbK. -rewrite -[RHS]connect_umfiltk eq_connect_aux //. -by move=>y /S2; rewrite !inE negbK. -Qed. - -(****************) -(* Schorr-Waite *) -(****************) - -(* short notation for left/right child of x *) -Notation gl g x := (get_nth g x 0). -Notation gr g x := (get_nth g x 1). - -(* type of markings *) -(* Aleks NOTE: Should we need U? My impression was that *) -(* being unmarked is represented by not being in *) -(* the appropriate map? *) -Inductive mark := U | L | R | LR. -(* decidable equality on marks *) -Definition eq_mark l1 l2 : bool := - if (l1, l2) is ((U, U)|(L, L)|(R, R)|(LR, LR)) then true else false. -Lemma eq_markP : Equality.axiom eq_mark. -Proof. by case; case=>//=; constructor. Qed. -HB.instance Definition _ := hasDecEq.Build mark eq_markP. - -(* marking of children *) - -(* given marking map m, x is m-child of y iff: *) -(* - m y = L and x is left child of y *) -(* - m y = R and x is right child of y *) -Definition ch (m : nmap mark) (g : pregraph) (x y : nat) := - [|| (find y m == Some L) && (gl g y == x) | - (find y m == Some R) && (gr g y == x)]. - -Lemma chP (m : nmap mark) g x y : - reflect [\/ (y, L) \In m /\ gl g y = x | - (y, R) \In m /\ gr g y = x] - (ch m g x y). -Proof. -rewrite /ch; case: orP=>H; constructor. -- by case: H=>H; [left|right]; case/andP: H=>/eqP/In_find H /eqP. -by case=>[][/In_find/eqP M /eqP G]; apply: H; [left|right]; apply/andP. -Qed. - -Lemma chN0 m g x y : - ch m g x y -> - y != null. -Proof. by case/chP=>[][/In_dom/dom_cond]. Qed. - -Lemma ch_fun m g a b x : - ch m g a x -> - ch m g b x -> - a = b. -Proof. by case/chP=>[][H <-] /chP [][/(In_fun H) {}H <-]. Qed. - -Lemma ch_path m g s x : - x \in s -> - path (ch m g) null s -> - exists y, y \in belast null s /\ ch m g y x. -Proof. exact: path_prev. Qed. - -Lemma ch_path_uniq m g s : - path (ch m g) null s -> - uniq (null :: s). -Proof. by apply: path_uniq; [apply: chN0|apply: ch_fun]. Qed. - -Definition upd_edge (m : nmap mark) g x y : seq node := - if find x m is Some L then [:: y; gr g x] else [:: gl g x; y]. - -Fixpoint rev_graph' m g ps t : pregraph := - if ps is p :: ps then - rev_graph' m (free g p) ps p \+ pts p (upd_edge m g p t) - else g. - -Definition rev_graph m g s t := rev_graph' m g (rev s) t. - -(* layout of graph+marking as heap *) -Definition lay (m : nmap mark) (g : pregraph) : heap := - \big[join/Unit]_(x <- dom g) (x :-> (gl g x, gr g x, odflt L (find x m))). - -(* reach m g s t holds iff *) -(* each unmarked node x in g is reachable through unmarked nodes *) -(* - either starting from t, or *) -(* - starting from a right child of some node in s *) -Definition reach (m : nmap mark) (g : pregraph) (s : seq node) t := - forall x, - (* x is node in g *) - x \in dom g -> - (* x is unmarked *) - x \notin dom m -> - (* x is reachable from t avoiding marked nodes *) - (* Aleks note: if unmarked nodes are represented by a marking U *) - (* then this conjunct will have to change, as it currently assumes *) - (* that unmarked nodes are just those that aren't in m *) - x \in connect (mem (dom m)) g t \/ - (* or x is reachable from some right child of a node y in s *) - (* avoiding marked nodes *) - exists y, y \in s /\ x \in connect (mem (dom m)) g (gr g y). - diff --git a/htt/domain.v.crashcoqide b/htt/domain.v.crashcoqide deleted file mode 100644 index 8a14f09..0000000 --- a/htt/domain.v.crashcoqide +++ /dev/null @@ -1,1038 +0,0 @@ -(* -Copyright 2019 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 HB Require Import structures. -From Coq Require Import ssreflect ssrbool ssrfun. -From mathcomp Require Import ssrnat eqtype. -From pcm Require Import options axioms pred prelude. - -(**************************************************) -(* This file develops some basic domain theory of *) -(* posets, complete lattices, complete partial *) -(* orders (CPOs). It developes the standard fixed *) -(* point theorems: Knaster-Tarski's for complete *) -(* lattices, and Kleene's for CPOs. *) -(**************************************************) - -(**********) -(* Posets *) -(**********) - -Definition poset_axiom T (leq : T -> T -> Prop) := - [/\ forall x, leq x x, - forall x y, leq x y -> leq y x -> x = y & - forall x y z, leq x y -> leq y z -> leq x z]. - -HB.mixin Record isPoset T := { - poset_leq : T -> T -> Prop; - poset_subproof : poset_axiom poset_leq}. - -#[short(type="poset")] -HB.structure Definition Poset := {T of isPoset T}. - -Notation "x <== y" := (poset_leq x y) (at level 70). - -Section Repack. -Variable T : poset. -Lemma poset_refl (x : T) : x <== x. -Proof. by case: (@poset_subproof T). Qed. -Lemma poset_asym (x y : T) : x <== y -> y <== x -> x = y. -Proof. by case: (@poset_subproof T)=>_ H _; apply: H. Qed. -Lemma poset_trans (y x z : T) : x <== y -> y <== z -> x <== z. -Proof. by case: (@poset_subproof T)=>_ _; apply. Qed. -End Repack. - -#[export] Hint Resolve poset_refl : core. - -Add Parametric Relation (T : poset) : T (@poset_leq T) - reflexivity proved by (@poset_refl _) - transitivity proved by (fun x y => @poset_trans _ y x) as poset_rel. - -(**********************) -(* monotone functions *) -(**********************) - -Definition monofun_axiom (T1 T2 : poset) (f : T1 -> T2) := - forall x y, x <== y -> f x <== f y. - -HB.mixin Record isMonoFun (T1 T2 : poset) (f : T1 -> T2) := { - monofun_subproof : monofun_axiom f}. - -#[short(type="mono_fun")] -HB.structure Definition MonoFun (T1 T2 : poset) := - {f of isMonoFun T1 T2 f}. - -Lemma monofunP (T1 T2 : poset) (f : mono_fun T1 T2) x y : - x <== y -> f x <== f y. -Proof. exact: monofun_subproof. Qed. - -(**************************) -(* some basic definitions *) -(**************************) - -Section IdealDef. -Variable T : poset. - -Structure ideal (P : T) := Ideal {id_val : T; id_pf : id_val <== P}. - -(* Changing the type of the ideal *) - -Lemma relaxP (P1 P2 : T) : P1 <== P2 -> forall p, p <== P1 -> p <== P2. -Proof. by move=>H1 p H2; apply: poset_trans H1. Qed. - -Definition relax (P1 P2 : T) (x : ideal P1) (pf : P1 <== P2) := - Ideal (relaxP pf (id_pf x)). - -End IdealDef. - -(***********************) -(* poset constructions *) -(***********************) - -Section SubPoset. -Variables (T : poset) (s : Pred T). -Local Notation tp := {x : T | x \In s}. - -Definition sub_leq (p1 p2 : tp) := sval p1 <== sval p2. - -Lemma sub_is_poset : poset_axiom sub_leq. -Proof. -split=>[x|[x Hx][y Hy]|[x Hx][y Hy][z Hz]]; rewrite /sub_leq //=. -- move=>H1 H2; rewrite -(poset_asym H1 H2) in Hy *. - by rewrite (pf_irr Hx). -by apply: poset_trans. -Qed. -HB.instance Definition _ := isPoset.Build tp sub_is_poset. -End SubPoset. - -(* pairing of posets *) - -Section ProdPoset. -Variables A B : poset. -Local Notation tp := (A * B)%type. - -Definition poset_prod_leq := - [fun p1 p2 : tp => p1.1 <== p2.1 /\ p1.2 <== p2.2]. - -Lemma prod_is_poset : poset_axiom poset_prod_leq. -Proof. -split=>[//|[x1 x2][y1 y2]|[x1 x2][y1 y2][z1 z2]] /=. -- by case=>H1 H2 [H3 H4]; congr (_, _); apply: poset_asym. -case=>H1 H2 [H3 H4]; split; -by [apply: poset_trans H3|apply: poset_trans H4]. -Qed. -HB.instance Definition _ := isPoset.Build tp prod_is_poset. -End ProdPoset. - -(* functions into poset form a poset *) - -Section FunPoset. -Variable (A : Type) (B : poset). -Local Notation tp := (A -> B). - -Definition poset_fun_leq := [fun p1 p2 : tp => forall x, p1 x <== p2 x]. - -Lemma fun_is_poset : poset_axiom poset_fun_leq. -Proof. -split=>[x|x y H1 H2|x y z H1 H2 t] //=. -- by apply: fext=>z; apply: poset_asym; [apply: H1|apply: H2]. -by apply: poset_trans (H2 t). -Qed. -HB.instance Definition _ := isPoset.Build tp fun_is_poset. -End FunPoset. - -(* dependent functions into a poset form a poset *) - -Section DFunPoset. -Variables (A : Type) (B : A -> poset). -Local Notation tp := (forall x, B x). - -Definition poset_dfun_leq := [fun p1 p2 : tp => forall x, p1 x <== p2 x]. - -Lemma dfun_is_poset : poset_axiom poset_dfun_leq. -Proof. -split=>[x|x y H1 H2|x y z H1 H2 t] //. -- by apply: fext=>z; apply: poset_asym; [apply: H1|apply: H2]. -by apply: poset_trans (H2 t). -Qed. - -(* no instance declaration, since it's keyed on forall *) -(* and forall is not a symbol *) -Definition dfun_poset_mix := isPoset.Build tp dfun_is_poset. -Definition dfun_poset : poset := Poset.pack_ dfun_poset_mix. -End DFunPoset. - -(* ideal of poset is poset *) - -Section IdealPoset. -Variable (T : poset) (P : T). - -Definition poset_ideal_leq := - [fun p1 p2 : ideal P => id_val p1 <== id_val p2]. - -Lemma ideal_is_poset : poset_axiom poset_ideal_leq. -Proof. -split=>[[x]|[x1 H1][x2 H2]|[x1 H1][x2 H2][x3 H3]] //=. -- move=>H3 H4; rewrite (poset_asym H3 H4) in H1 H2 *. - by rewrite (pf_irr H1). -by apply: poset_trans. -Qed. -HB.instance Definition _ := isPoset.Build (ideal P) ideal_is_poset. -End IdealPoset. - -(* Prop is poset *) -Definition poset_prop_leq := [fun p1 p2 : Prop => p1 -> p2]. -Lemma prop_is_poset : poset_axiom poset_prop_leq. -Proof. by split=>[x|x y H1 H2|x y z H1 H2 /H1] //; apply: pext. Qed. -HB.instance Definition _ := isPoset.Build Prop prop_is_poset. - - -(*********************) -(* Complete lattices *) -(*********************) - -Definition lattice_axiom (T : poset) (sup : Pred T -> T) := - [/\ (* sup is upper bound *) - forall (s : Pred T) x, x \In s -> x <== sup s & - (* sup is least upper bound *) - forall (s : Pred T) x, - (forall y, y \In s -> y <== x) -> sup s <== x]. - -HB.mixin Record isLattice T of Poset T := { - sup : Pred T -> T; - lattice_subproof : lattice_axiom sup}. - -#[short(type="lattice")] -HB.structure Definition Lattice := {T of Poset T & isLattice T}. - -Section Repack. -Variable T : lattice. -Lemma supP (s : Pred T) x : x \In s -> x <== sup s. -Proof. by case: (@lattice_subproof T)=>H _; apply: H. Qed. -Lemma supM (s : Pred T) x : (forall y, y \In s -> y <== x) -> sup s <== x. -Proof. by case: (@lattice_subproof T)=>_; apply. Qed. -End Repack. - -(* infima follow *) -Section Infimum. -Variable T : lattice. - -Definition inf (s : Pred T) := - sup [Pred x : T | forall y, y \In s -> x <== y]. - -Lemma infP s x : x \In s -> inf s <== x. -Proof. by move=>H; apply: supM=>y; apply. Qed. - -Lemma infM s x : (forall y, y \In s -> x <== y) -> x <== inf s. -Proof. by apply: supP. Qed. - -End Infimum. - -(* least and greatest fixed points by Knaster-Tarski construction *) -Section KnasterTarskiDefs. -Variable T : lattice. - -Definition lbot : T := sup Pred0. - -Definition tarski_lfp (f : T -> T) := inf [Pred x : T | f x <== x]. -Definition tarski_gfp (f : T -> T) := sup [Pred x : T | x <== f x]. - -Definition sup_closed (T : lattice) := - [Pred s : Pred T | forall d, d <=p s -> sup d \In s]. - -Definition sup_closure (T : lattice) (s : Pred T) := - [Pred p : T | forall t : Pred T, s <=p t -> t \In sup_closed T -> p \In t]. - -End KnasterTarskiDefs. - -Arguments lbot {T}. -Arguments sup_closed {T}. -Arguments sup_closure [T] s. -Prenex Implicits sup_closed sup_closure. - -Section BasicProperties. -Variable T : lattice. - -Lemma sup_mono (s1 s2 : Pred T) : s1 <=p s2 -> sup s1 <== sup s2. -Proof. by move=>H; apply: supM=>y /H/supP. Qed. - -Lemma supE (s1 s2 : Pred T) : s1 =p s2 -> sup s1 = sup s2. -Proof. by move=>H; apply: poset_asym; apply: supM=>y /H/supP. Qed. - -(* Knaster-Tarski theorem *) -Lemma tarski_lfp_fixed (f : mono_fun T T) : - f (tarski_lfp f) = tarski_lfp f. -Proof. -suff L : f (tarski_lfp f) <== tarski_lfp f. -- by apply: poset_asym=>//; apply/infP/monofunP. -by apply: infM=>x L; apply: poset_trans (L); apply/monofunP/infP. -Qed. - -Lemma tarski_lfp_least f (x : T) : f x <== x -> tarski_lfp f <== x. -Proof. by move=>H; apply: infP. Qed. - -Lemma tarski_gfp_fixed (f : mono_fun T T) : - f (tarski_gfp f) = tarski_gfp f. -Proof. -suff L: tarski_gfp f <== f (tarski_gfp f). -- by apply: poset_asym=>//; apply/supP/monofunP. -by apply: supM=>x L; apply: poset_trans (L) _; apply/monofunP/supP. -Qed. - -Lemma tarski_gfp_greatest f (x : T) : x <== f x -> x <== tarski_gfp f. -Proof. by move=>H; apply: supP. Qed. - -Lemma tarski_lfp_mono (f1 : T -> T) (f2 : mono_fun T T) : - f1 <== f2 -> tarski_lfp f1 <== tarski_lfp f2. -Proof. -move=>H; apply: infP; apply: poset_trans (H (tarski_lfp f2)) _. -by rewrite tarski_lfp_fixed. -Qed. - -Lemma tarski_gfp_mono (f1 : mono_fun T T) (f2 : T -> T) : - (f1 : T -> T) <== f2 -> tarski_gfp f1 <== tarski_gfp f2. -Proof. -move=>H; apply/supP/poset_trans/(H (tarski_gfp f1)). -by rewrite tarski_gfp_fixed. -Qed. - -(* closure contains s *) -Lemma sup_clos_sub (s : Pred T) : s <=p sup_closure s. -Proof. by move=>p H1 t H2 H3; apply: H2 H1. Qed. - -(* closure is smallest *) -Lemma sup_clos_min (s : Pred T) : - forall t, s <=p t -> sup_closed t -> sup_closure s <=p t. -Proof. by move=>t H1 H2 p; move/(_ _ H1 H2). Qed. - -(* closure is closed *) -Lemma sup_closP (s : Pred T) : sup_closed (sup_closure s). -Proof. -move=>d H1 t H3 H4; move: (sup_clos_min H3 H4)=>{}H3. -by apply: H4=>x; move/H1; move/H3. -Qed. - -Lemma sup_clos_idemp (s : Pred T) : sup_closed s -> sup_closure s =p s. -Proof. by move=>p; split; [apply: sup_clos_min | apply: sup_clos_sub]. Qed. - -Lemma sup_clos_mono (s1 s2 : Pred T) : - s1 <=p s2 -> sup_closure s1 <=p sup_closure s2. -Proof. -move=>H1; apply: sup_clos_min (@sup_closP s2). -by move=>t /H1; apply: sup_clos_sub. -Qed. - -End BasicProperties. - -(* lattice constructions *) - -(* subset lattice *) -Section SubLattice. -Variables (T : lattice) (s : Pred T) (C : sup_closed s). -Local Notation tp := {x : T | x \In s}. - -Definition lattice_sub_sup' (u : Pred tp) : T := - sup [Pred x : T | exists y, y \In u /\ x = sval y]. -Lemma lattice_sub_supX (u : Pred tp) : lattice_sub_sup' u \In s. -Proof. by apply: C=>x [][y H][_] ->. Qed. -Definition lattice_sub_sup (u : Pred tp) : tp := - exist _ (lattice_sub_sup' u) (lattice_sub_supX u). - -Lemma sub_is_lattice : lattice_axiom lattice_sub_sup. -Proof. -split=>u x H; first by apply: supP; exists x. -by apply: supM=>y [z][H1] ->; apply: H H1. -Qed. -HB.instance Definition _ := isLattice.Build tp sub_is_lattice. -End SubLattice. - -(* product *) -Section ProdLattice. -Variables (A B : lattice). -Local Notation tp := (A * B)%type. - -Definition lattice_prod_sup (s : Pred tp) : tp := - (sup [Pred p | exists f, p = f.1 /\ f \In s], - sup [Pred p | exists f, p = f.2 /\ f \In s]). - -Lemma prod_is_lattice : lattice_axiom lattice_prod_sup. -Proof. -split=>s p H; first by split; apply: supP; exists p. -by split; apply: supM=>y [f][->]; case/H. -Qed. -HB.instance Definition _ := isLattice.Build tp prod_is_lattice. -End ProdLattice. - -(* functions into latice form lattice *) -Section FunLattice. -Variables (A : Type) (B : lattice). -Local Notation tp := (A -> B). - -Definition lattice_fun_sup (s : Pred tp) : tp := - fun x => sup [Pred p | exists f, f \In s /\ p = f x]. - -Lemma fun_is_lattice : lattice_axiom lattice_fun_sup. -Proof. -split=>s p H x; first by apply: supP; exists p. -by apply: supM=>y [f][H1] ->; apply: H. -Qed. -HB.instance Definition _ := isLattice.Build tp fun_is_lattice. -End FunLattice. - -(* dependent functions into a lattice form a lattice *) -Section DFunLattice. -Variables (A : Type) (B : A -> lattice). -Local Notation tp := (dfun_poset B). - -Definition lattice_dfun_sup (s : Pred tp) : tp := - fun x => sup [Pred p | exists f, f \In s /\ p = f x]. - -Lemma dfun_is_lattice : lattice_axiom lattice_dfun_sup. -Proof. -split=>s p H x; first by apply: supP; exists p. -by apply: supM=>y [f][H1] ->; apply: H. -Qed. - -Definition dfun_lattice_mix := isLattice.Build tp dfun_is_lattice. -Definition dfun_lattice : lattice := Lattice.pack_ dfun_lattice_mix. -End DFunLattice. - -(* applied sup equals the sup of applications *) -Lemma sup_appE A (B : lattice) (s : Pred (A -> B)) (x : A) : - sup s x = sup [Pred y : B | exists f, f \In s /\ y = f x]. -Proof. by []. Qed. - -Lemma sup_dappE A (B : A -> lattice) (s : Pred (dfun_lattice B)) (x : A) : - sup s x = sup [Pred y : B x | exists f, f \In s /\ y = f x]. -Proof. by []. Qed. - -(* ideal of a lattice forms a lattice *) -Section IdealLattice. -Variables (T : lattice) (P : T). - -Definition lattice_ideal_sup' (s : Pred (ideal P)) := - sup [Pred x | exists p, p \In s /\ id_val p = x]. -Lemma lattice_ideal_supP' (s : Pred (ideal P)) : lattice_ideal_sup' s <== P. -Proof. by apply: supM=>y [[x]] H /= [_] <-. Qed. -Definition lattice_ideal_sup (s : Pred (ideal P)) := - Ideal (lattice_ideal_supP' s). - -Lemma ideal_is_lattice : lattice_axiom lattice_ideal_sup. -Proof. -split=>s p H; first by apply: supP; exists p. -by apply: supM=>y [q][H1] <-; apply: H. -Qed. -HB.instance Definition _ := isLattice.Build (ideal P) ideal_is_lattice. -End IdealLattice. - -(* Prop is a lattice *) -Definition lattice_prop_sup (s : Pred Prop) : Prop := - exists p, p \In s /\ p. -Lemma prop_is_lattice : lattice_axiom lattice_prop_sup. -Proof. by split=>s p; [exists p|move=>H [r][]; move/H]. Qed. -HB.instance Definition _ := isLattice.Build Prop prop_is_lattice. - -(****************) -(* Monotonicity *) -(****************) - -Section MonoCanonicals. -Variables T T1 T2 T3 S1 S2 : poset. - -Lemma idfun_is_mono : monofun_axiom (@idfun T). Proof. by []. Qed. -HB.instance Definition _ := isMonoFun.Build T T idfun idfun_is_mono. - -Definition const_fun (y : T2) (_ : T1) := y. -Lemma const_is_mono (y : T2) : monofun_axiom (const_fun y). Proof. by []. Qed. -HB.instance Definition _ y := - isMonoFun.Build T1 T2 (const_fun y) (const_is_mono y). - -Lemma fst_is_mono : monofun_axiom (@fst T1 T2). -Proof. by case=>x1 x2 [y1 y2][]. Qed. -HB.instance Definition _ := isMonoFun.Build (T1*T2)%type T1 fst fst_is_mono. - -Lemma snd_is_mono : monofun_axiom (@snd T1 T2). -Proof. by case=>x1 x2 [y1 y2][]. Qed. -HB.instance Definition _ := isMonoFun.Build (T1*T2)%type T2 snd snd_is_mono. - -Definition diag (x : T) := (x, x). -Lemma diag_is_mono : monofun_axiom diag. Proof. by []. Qed. -HB.instance Definition _ := isMonoFun.Build T (T*T)%type diag diag_is_mono. - -Lemma sval_is_mono (P : Pred T) : monofun_axiom (@sval T P). -Proof. by []. Qed. -HB.instance Definition _ P := - isMonoFun.Build {x : T | P x} T sval (@sval_is_mono P). - -Definition app A B (x : A) := fun f : A -> B => f x. -Lemma app_is_mono A (x : A) : monofun_axiom (@app A T x). -Proof. by move=>f1 f2; apply. Qed. -HB.instance Definition _ A x := - isMonoFun.Build (A -> T) T (@app A T x) (app_is_mono x). - -Definition dapp A (B : A -> poset) (x : A) := fun f : dfun_poset B => f x. -Lemma dapp_is_mono A (B : A -> poset) (x : A) : monofun_axiom (@dapp A B x). -Proof. by move=>f1 f2; apply. Qed. -HB.instance Definition _ A B x := - isMonoFun.Build (dfun_poset B) (B x) (@dapp A B x) (dapp_is_mono x). - -Section Comp. -Variables (f1 : mono_fun T2 T1) (f2 : mono_fun T3 T2). -Lemma comp_is_mono : monofun_axiom (f1 \o f2). -Proof. by move=>x y H; apply/monofunP/monofunP. Qed. -HB.instance Definition _ := isMonoFun.Build T3 T1 (f1 \o f2) comp_is_mono. -End Comp. - -Section Prod. -Variables (f1 : mono_fun S1 T1) (f2 : mono_fun S2 T2). -Lemma funprod_is_mono : monofun_axiom (f1 \* f2). -Proof. by case=>x1 x2 [y1 y2][/= H1 H2]; split=>/=; apply: monofunP. Qed. -HB.instance Definition _ := - isMonoFun.Build (S1 * S2)%type (T1 * T2)%type (f1 \* f2) funprod_is_mono. -End Prod. -End MonoCanonicals. - -Prenex Implicits diag app. - -(**********) -(* Chains *) -(**********) - -Definition chain_axiom (T : poset) (s : Pred T) := - (exists d, d \In s) /\ - (forall x y, x \In s -> y \In s -> x <== y \/ y <== x). - -HB.mixin Record isChain (T : poset) (s : Pred T) := { - chain_subproof : chain_axiom s}. - -#[short(type="chain")] -HB.structure Definition Chain (T : poset) := {s of isChain T s}. - -Arguments chain T%_type. - -(* \In notation *) -Definition chain_pred (T : poset) (x : chain T) : Pred T := x. -Canonical chainPredType (T : poset) := - Eval hnf in @mkPredType T (@chain T) (@chain_pred T). - -Lemma chainE (T : poset) (s1 s2 : chain T) : - s1 = s2 <-> chain_pred s1 =p chain_pred s2. -Proof. -split=>[->//|]; case: s1 s2=>s1 H1 [s2 H2] /= E; move: H1 H2. -suff: s1 = s2. -- move=>-> [[H1]][[H2]]; congr (Chain.Pack (Chain.Class _)). - by rewrite (pf_irr H1). -by apply: fext=>x; apply: pext; split=> /E. -Qed. - -Section ImageChain. -Variables (T1 T2 : poset) (f : mono_fun T1 T2) (s : chain T1). - -Lemma image_is_chain : chain_axiom (Image f s). -Proof. -case: s=>p [[[[d H1] H2]]] /=. -split=>[|x y]; first by exists (f d); exists d. -case=>x1 -> H3 [y1 -> H4]; rewrite -!toPredE /= in H3 H4. -by case: (H2 x1 y1 H3 H4)=>L; [left | right]; apply: monofunP L. -Qed. -HB.instance Definition _ := isChain.Build T2 (Image f s) image_is_chain. -End ImageChain. - -Lemma id_chainE T (s : chain T) f : f =1 idfun -> Image f s =p s. -Proof. -move=>H x; split; first by case=>y ->; rewrite H. -by exists x=>//; rewrite H. -Qed. - -Section ChainConst. -Variables (T1 T2 : poset) (y : T2). - -Definition const_chain : Pred _ := Pred1 y. -Lemma const_is_chain : chain_axiom const_chain. -Proof. by split; [exists y | move=>x1 x2 <-<-; left]. Qed. -HB.instance Definition _ := isChain.Build T2 const_chain const_is_chain. - -Lemma const_chainE (s : chain T1) : Image (const_fun y) s =p const_chain. -Proof. -move=>z1; split; first by case=>z2 ->. -by case: s=>s [[[[d H1]] H2]] /= <-; exists d. -Qed. -End ChainConst. - -Section ChainCompose. -Variables T1 T2 T3 : poset. -Variables (f1 : mono_fun T2 T1) (f2 : mono_fun T3 T2). -Variable s : chain T3. -Lemma comp_chainE : Image f1 (Image f2 s) =p Image (f1 \o f2) s. -Proof. by move=>x; apply: ImageIm. Qed. -End ChainCompose. - -Section ProjChain. -Variables T1 T2 : poset. -Variables s : chain (T1 * T2). -Definition proj1_chain : Pred T1 := Image fst s. -HB.instance Definition _ := Chain.copy proj1_chain proj1_chain. -Definition proj2_chain : Pred T2 := Image snd s. -HB.instance Definition _ := Chain.copy proj2_chain proj2_chain. -End ProjChain. - -Section DiagChain. -Variables (T : poset) (s : chain T). -Definition diag_chain : Pred (T * T) := Image diag s. -HB.instance Definition _ := Chain.copy diag_chain diag_chain. -End DiagChain. - - -(*********) -(* CPO's *) -(*********) - -Definition cpo_axiom (T : poset) (bot : T) (lim : chain T -> T) := - [/\ forall x, bot <== x, - forall (s : chain T) x, x \In s -> x <== lim s & - forall (s : chain T) x, - (forall y, y \In s -> y <== x) -> lim s <== x]. - -HB.mixin Record isCPO T of Poset T := { - bot : T; - lim_op : chain T -> T; - cpo_subproof: cpo_axiom bot lim_op}. - -#[short(type="cpo")] -HB.structure Definition CPO := {T of Poset T & isCPO T}. - -Definition limx (D : cpo) (s : chain D) of phantom (Pred _) s := lim_op s. -Notation lim s := (limx (Phantom (Pred _) s)). - -Section Repack. -Variable D : cpo. -Lemma botP (x : D) : bot <== x. -Proof. by case: (@cpo_subproof D). Qed. -Lemma limP (s : chain D) x : x \In s -> x <== lim s. -Proof. by case: (@cpo_subproof D)=>_ H _; apply: H. Qed. -Lemma limM (s : chain D) x : (forall y, y \In s -> y <== x) -> lim s <== x. -Proof. by case: (@cpo_subproof D)=>_ _; apply. Qed. -End Repack. - -#[export] Hint Resolve botP : core. - -Lemma limE (D : cpo) (s1 s2 : chain D) : s1 =p s2 -> lim s1 = lim s2. -Proof. by move/chainE=>->. Qed. - -(* common chain constructions *) - -(* adding bot to a chain *) - -Definition lift (D : cpo) (s : Pred D) : Pred D := - fun x => x = bot \/ x \In s. - -Lemma lift_is_chain (D : cpo) (s : chain D) : chain_axiom (lift s). -Proof. -case: s=>p [[[[d H1]] H2]] /=. -split=>[|x y]; first by exists d; right. -by case=>[->|H3][->|H4]; auto. -Qed. - -HB.instance Definition _ (D : cpo) (s : chain D) := - isChain.Build D (lift s) (@lift_is_chain D s). - - -(****************************) -(* common cpo constructions *) -(****************************) - -(* products *) -Section ProdCPO. -Variables (A B : cpo). -Definition cpo_prod_bot := (@bot A, @bot B). -Definition cpo_prod_lim (s : chain (A * B)) := - (lim (Image fst s), lim (Image snd s)). - -Lemma prod_is_cpo : cpo_axiom cpo_prod_bot cpo_prod_lim. -Proof. -split=>[x|s x|s x]. -- by split; apply: botP. -- by split; apply: limP; exists x. -by split; apply: limM=>y [z ->]; case/H. -Qed. -HB.instance Definition _ := isCPO.Build (A * B)%type prod_is_cpo. -End ProdCPO. - -(* functions *) -Section FunCPO. -Variable (A : Type) (B : cpo). -Definition cpo_fun_bot (x : A) := @bot B. -Definition cpo_fun_lim (s : chain (A -> B)) x := - lim (Image (app x) s). - -Lemma fun_is_cpo : cpo_axiom cpo_fun_bot cpo_fun_lim. -Proof. -split=>[f|s f|s f]. -- by move=>y; apply: botP. -- by move=>H t; apply: limP; exists f. -by move=>H1 t; apply: limM=>y [g ->] H2; apply: H1. -Qed. -HB.instance Definition _ := isCPO.Build (A -> B) fun_is_cpo. -End FunCPO. - -(* dependent functions *) -Section DFunCPO. -Variables (A : Type) (B : A -> cpo). - -Definition cpo_dfun_bot : dfun_poset B := fun x => bot. -Definition cpo_dfun_lim (s : chain (dfun_poset B)) : dfun_poset B := - fun x => lim (Image (dapp x) s). - -Lemma dfun_is_cpo : cpo_axiom cpo_dfun_bot cpo_dfun_lim. -Proof. -split=>[x|s x|s x]. -- by move=>y; apply: botP. -- by move=>H t; apply: limP; exists x. -by move=>H1 t; apply: limM=>y [f ->] H2; apply: H1. -Qed. - -Definition dfun_cpo_mix := isCPO.Build (dfun_poset B) dfun_is_cpo. -Definition dfun_cpo : cpo := CPO.pack_ dfun_cpo_mix. -End DFunCPO. - -(* Prop *) -Definition cpo_prop_bot : Prop := False. -Definition cpo_prop_lim (s : chain Prop) : Prop := - exists p, p \In s /\ p. -Lemma prop_is_cpo : cpo_axiom cpo_prop_bot cpo_prop_lim. -Proof. by split=>[x|s p|s p H] //=; [exists p|case=>r []; move/H]. Qed. -HB.instance Definition _ := isCPO.Build Prop prop_is_cpo. - -(* Pred *) -HB.instance Definition _ A := CPO.copy (Pred A) _. - -(* every complete lattice is cpo *) -Section LatticeCPO. -Variable A : lattice. - -Definition lat_bot : A := lbot. -Definition lat_lim (s : Pred A) : A := sup s. - -Lemma lat_is_cpo : cpo_axiom lat_bot lat_lim. -Proof. by split=>[x|s x|s x]; [apply: supM|apply: supP|apply: supM]. Qed. - -(* no instance declaration as nothing to *) -(* hang the new structure onto *) -Definition lat_cpo_mix := isCPO.Build A lat_is_cpo. -Definition lat_cpo := CPO.pack_ lat_cpo_mix. -End LatticeCPO. - -(* sub-CPO's *) - -(* every chain-closed subset of a cpo is a cpo *) - -Section AdmissibleClosure. -Variable T : cpo. - -Definition chain_closed := - [Pred s : Pred T | - bot \In s /\ forall d : chain T, d <=p s -> lim d \In s]. - -(* admissible closure of s is the smallest closed set containing s *) -(* basically extends s to include the sups of chains *) -Definition chain_closure (s : Pred T) := - [Pred p : T | forall t : Pred T, s <=p t -> chain_closed t -> p \In t]. - -(* admissible closure contains s *) -Lemma chain_clos_sub (s : Pred T) : s <=p chain_closure s. -Proof. by move=>p H1 t H2 H3; apply: H2 H1. Qed. - -(* admissible closure is smallest *) -Lemma chain_clos_min (s : Pred T) t : - s <=p t -> chain_closed t -> chain_closure s <=p t. -Proof. by move=>H1 H2 p; move/(_ _ H1 H2). Qed. - -(* admissible closure is closed *) -Lemma chain_closP (s : Pred T) : chain_closed (chain_closure s). -Proof. -split; first by move=>t _ []. -move=>d H1 t H3 H4; move: (chain_clos_min H3 H4)=>{}H3. -by case: H4=>_; apply=>// x; move/H1; move/H3. -Qed. - -Lemma chain_clos_idemp (s : Pred T) : - chain_closed s -> chain_closure s =p s. -Proof. -move=>p; split; last by apply: chain_clos_sub. -by apply: chain_clos_min=>// /chain_closP. -Qed. - -Lemma chain_clos_mono (s1 s2 : Pred T) : - s1 <=p s2 -> chain_closure s1 <=p chain_closure s2. -Proof. -move=>H1; apply: chain_clos_min (chain_closP s2)=>p H2. -by apply: chain_clos_sub; apply: H1. -Qed. - -(* intersection of admissible sets is admissible *) -Lemma chain_closI (s1 s2 : Pred T) : - chain_closed s1 -> chain_closed s2 -> chain_closed (PredI s1 s2). -Proof. -move=>[H1 S1][H2 S2]; split=>// d H. -by split; [apply: S1 | apply: S2]=>// x; case/H. -Qed. - -End AdmissibleClosure. - -Arguments chain_closed {T}. -Prenex Implicits chain_closed. - -(* diagonal of an admissible set of pairs is admissible *) -Lemma chain_clos_diag (T : cpo) (s : Pred (T * T)) : - chain_closed s -> chain_closed [Pred t : T | (t, t) \In s]. -Proof. -case=>B H1; split=>// d H2; move: (H1 (Image diag d)). -rewrite /limx/=/lim_op /= /cpo_prod_lim /=. -rewrite !(limE (comp_chainE _ _ _)) !(limE (id_chainE _ _)) //=. -by apply; case=>x1 x2 [z ->]; apply: H2. -Qed. - -Section SubCPO. -Variables (D : cpo) (s : Pred D) (C : chain_closed s). -Local Notation tp := {x : D | x \In s}. - -Definition subcpo_bot := exist _ (@bot D) (proj1 C). -Lemma subcpo_limX (u : chain tp) : lim (Image sval u) \In s. -Proof. by case: C u=>P H u; apply: (H)=>t [[y H1 ->]]. Qed. -Definition subcpo_lim (u : chain tp) : tp := - exist _ (lim (Image sval u)) (subcpo_limX u). - -Lemma subcpo_is_cpo : cpo_axiom subcpo_bot subcpo_lim. -Proof. -split=>[x|u x H|u x H] //=. -- by apply: botP. -- by apply: limP; exists x. -by apply: limM=>y [z ->]; apply: H. -Qed. -HB.instance Definition _ := isCPO.Build tp subcpo_is_cpo. -End SubCPO. - -(***********************************************) -(* Continuity and Kleene's fixed point theorem *) -(***********************************************) - -Lemma lim_mono (D : cpo) (s1 s2 : chain D) : - s1 <=p s2 -> lim s1 <== lim s2. -Proof. by move=>H; apply: limM=>y /H/limP. Qed. - -Lemma lim_liftE (D : cpo) (s : chain D) : lim s = lim (lift s). -Proof. -apply: poset_asym; apply: limM=>y H; first by apply: limP; right. -by case: H; [move=>-> | apply: limP]. -Qed. - -(* applied lim equals the lim of applications *) -(* ie., part of continuity of application *) -(* but is so often used, I give it a name *) -Lemma lim_appE A (D : cpo) (s : chain (A -> D)) (x : A) : - lim s x = lim (Image (app x) s). -Proof. by []. Qed. - -Lemma lim_dappE A (D : A -> cpo) (s : chain (dfun_cpo D)) (x : A) : - lim s x = lim (Image (dapp x) s). -Proof. by []. Qed. - -(************************) -(* continuous functions *) -(************************) - -Definition contfun_axiom (D1 D2 : cpo) (f : mono_fun D1 D2) := - forall s : chain D1, f (lim s) <== lim (Image f s). - -HB.mixin Record isContFun (D1 D2 : cpo) (f : D1 -> D2) - of @MonoFun D1 D2 f := {contfun_subproof : contfun_axiom f}. - -#[short(type="cont_fun")] -HB.structure Definition ContFun (D1 D2 : cpo) := - {f of @MonoFun D1 D2 f & isContFun D1 D2 f}. - -Lemma contE (D1 D2 : cpo) (f : cont_fun D1 D2) (s : chain D1) : - f (lim s) = lim (Image f s). -Proof. -apply: poset_asym; first by apply: contfun_subproof. -by apply: limM=>y [x1 ->] /limP/monofunP. -Qed. - -(* just for this proof, we want that nat is a poset under <= *) -(* in other scopes, we'll use the equality as ordering *) -(* Thus, we encapsulate under module. *) - -Module Kleene. - -Lemma nat_is_poset : poset_axiom leq. -Proof. -split=>[|x y H1 H2|x y z] //; last by apply: leq_trans. -by apply: anti_leq; rewrite H1 H2. -Qed. -HB.instance Definition _ := isPoset.Build nat nat_is_poset. - -(* nats form a chain *) -Definition nats : Pred nat := PredT. -Lemma nats_chain_axiom : chain_axiom nats. -Proof. -split=>[|x y _ _]; first by exists 0. -rewrite /poset_leq /= [y <= x]leq_eqVlt. -by case: leqP; [left | rewrite orbT; right]. -Qed. -HB.instance Definition _ := isChain.Build nat nats nats_chain_axiom. - -Section Kleene. -Variables (D : cpo) (f : cont_fun D D). - -Fixpoint pow m := if m is n.+1 then f (pow n) else bot. - -Lemma pow_is_mono : monofun_axiom pow. -Proof. -move=>m n; elim: n m=>[|n IH] m /=; first by case: m. -rewrite {1}/poset_leq /= leq_eqVlt ltnS. -case/orP; first by move/eqP=>->. -move/IH=>{IH} H; apply: poset_trans H _. -by elim: n=>[|n IH] //=; apply: monofunP IH. -Qed. -HB.instance Definition _ := isMonoFun.Build nat D pow pow_is_mono. - -Lemma reindex : Image pow nats =p lift (Image f (Image pow nats)). -Proof. -move=>x; split. -- case; case=>[|n] -> _; first by left. - by right; exists (pow n)=>//; exists n. -case=>/=; first by move=>->; exists 0. -by case=>y -> [n ->]; exists n.+1. -Qed. - -Definition kleene_lfp := lim (Image pow nats). - -Lemma kleene_lfp_fixed : f kleene_lfp = kleene_lfp. -Proof. by rewrite /kleene_lfp contE /= (limE reindex) /= lim_liftE. Qed. - -Lemma kleene_lfp_least x : f x = x -> kleene_lfp <== x. -Proof. -move=>H; apply: limM=>y [n ->] _. -by elim: n=>[|n IH] //=; rewrite -H; apply: monofunP IH. -Qed. - -End Kleene. - -Lemma kleene_lfp_mono (D : cpo) (f1 f2 : cont_fun D D) : - (f1 : D -> D) <== f2 -> - kleene_lfp f1 <== kleene_lfp f2. -Proof. -move=>H; apply: limM=>x [n ->] X. -have N : forall n, pow f1 n <== pow f2 n. -- by elim=>//= m /(monofunP f2); apply: poset_trans (H _). -by apply: poset_trans (N n) _; apply: limP; exists n. -Qed. - -Module Exports. -Notation kleene_lfp := kleene_lfp. -Notation kleene_lfp_fixed := kleene_lfp_fixed. -Notation kleene_lfp_least := kleene_lfp_least. -Notation kleene_lfp_mono := kleene_lfp_mono. -End Exports. - -End Kleene. - -Export Kleene.Exports. - - -(**********************************) -(* Continuity of common functions *) -(**********************************) - -Lemma idfun_is_cont (D : cpo) : contfun_axiom (@idfun D). -Proof. by move=>d; rewrite (limE (id_chainE _ _)). Qed. -HB.instance Definition _ (D : cpo) := - isContFun.Build D D idfun (@idfun_is_cont D). - -Lemma const_is_cont (D1 D2 : cpo) (y : D2) : - contfun_axiom (@const_fun D1 D2 y). -Proof. -by move=>s; apply: limP; case: s=>[p][[[[d H1] H2]]]; exists d. -Qed. -HB.instance Definition _ (D1 D2 : cpo) (y : D2) := - isContFun.Build D1 D2 (const_fun y) (const_is_cont y). - -Section ContCompose. -Variables D1 D2 D3 : cpo. -Variable (f1 : cont_fun D2 D1) (f2 : cont_fun D3 D2). -Lemma comp_is_cont : contfun_axiom (f1 \o f2). -Proof. by move=>s; rewrite /= !contE (limE (comp_chainE _ _ _)). Qed. -HB.instance Definition _ := isContFun.Build D3 D1 (f1 \o f2) comp_is_cont. -End ContCompose. - -Lemma fst_is_cont (D1 D2 : cpo) : contfun_axiom (@fst D1 D2). -Proof. by []. Qed. -HB.instance Definition _ (D1 D2 : cpo) := - isContFun.Build (D1*D2)%type D1 fst (@fst_is_cont D1 D2). - -Lemma snd_is_cont (D1 D2 : cpo) : contfun_axiom (@snd D1 D2). -Proof. by []. Qed. -HB.instance Definition _ (D1 D2 : cpo) := - isContFun.Build (D1*D2)%type D2 snd (@snd_is_cont D1 D2). - -Lemma diag_is_cont (D : cpo) : contfun_axiom (@diag D). -Proof. -move=>d; split=>/=; -by rewrite (limE (comp_chainE _ _ _)) (limE (id_chainE _ _)). -Qed. -HB.instance Definition _ (D : cpo) := - isContFun.Build D (D*D)%type diag (@diag_is_cont D). - -Lemma app_is_cont A (D : cpo) x : contfun_axiom (@app A D x). -Proof. by []. Qed. -HB.instance Definition _ A (D : cpo) x := - isContFun.Build (A -> D) D (@app A D x) (app_is_cont x). - -(* can't reuse dapp and it's proof of monotonicity *) -(* because inheritance doesn't propagate properly *) -(* for dependent functions; so we duplicate *) -(* See mathcomp/finfun.v for solution (dfun vs. dffun). *) -Definition dapp2 A (B : A -> cpo) (x : A) := - fun f : dfun_cpo B => f x. -Lemma dapp2_is_mono A (B : A -> cpo) x : monofun_axiom (@dapp2 A B x). -Proof. by move=>f1 f2; apply. Qed. -HB.instance Definition _ A (B : A -> cpo) x := - isMonoFun.Build (dfun_poset B) (B x) (@dapp2 A B x) (dapp2_is_mono x). -Lemma dapp2_is_cont A (B : A -> cpo) (x : A) : contfun_axiom (@dapp2 A B x). -Proof. -move=>s; have <- // : dapp2 x (lim s) = lim [Image dapp2 x i | i <- s]. -by apply: limE=>/=. -Qed. -HB.instance Definition _ A (B : A -> cpo) x := - isContFun.Build (dfun_cpo B) (B x) (@dapp2 A B x) (dapp2_is_cont x). - -Section ProdCont. -Variables S1 S2 T1 T2 : cpo. -Variables (f1 : cont_fun S1 T1) (f2 : cont_fun S2 T2). - -Lemma prod_is_cont : contfun_axiom (f1 \* f2). -Proof. -move=>d; split=>//=; -by rewrite contE !(limE (comp_chainE _ _ _)); apply: lim_mono. -Qed. -HB.instance Definition _ := - isContFun.Build (S1*S2)%type (T1*T2)%type (f1 \* f2) prod_is_cont. -End ProdCont. - -(* DEVCOMMENT *) -(* TODO: *) -(* 1. limit of a chain of continuous functions is a continuous function *) -(* 2. CPO of continuous functions should be the subCPO of functions *) -(* 3. kleene_lfp is a continuous operation *) -(* /DEVCOMMENT *) - From 323ccb9dca99be9e7b3d36b02d58debead430cca Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Fri, 16 Aug 2024 12:25:52 +0200 Subject: [PATCH 57/93] moved additional lemmas from individual files to fcsl-pcm. also, cleaned up congprog, hashtab, kvmaps, etc. --- _CoqProject | 2 +- core/finmap.v | 49 +++--- core/prelude.v | 4 +- core/seqperm.v | 16 ++ examples/congmath.v | 7 +- examples/congprog.v | 174 +++---------------- examples/hashtab.v | 33 ++-- examples/kvmaps.v | 51 +++--- examples/tree.v | 36 ++-- examples/union_find.v | 68 ++------ pcm/heap.v | 2 +- pcm/pcm.v | 394 ++++++++++++++++++++++++++++-------------- pcm/unionmap.v | 36 +++- 13 files changed, 439 insertions(+), 433 deletions(-) diff --git a/_CoqProject b/_CoqProject index ff398d9..c488066 100644 --- a/_CoqProject +++ b/_CoqProject @@ -50,7 +50,7 @@ examples/hashtab.v examples/bubblesort.v examples/quicksort.v examples/congmath.v -#examples/congprog.v +examples/congprog.v examples/tree.v examples/union_find.v examples/graph.v diff --git a/core/finmap.v b/core/finmap.v index 0d06816..346e367 100644 --- a/core/finmap.v +++ b/core/finmap.v @@ -115,12 +115,13 @@ Definition cosupp s := map value (seq_of s). End Ops. -Prenex Implicits fnd ins rem supp. +Arguments nil {K V}. +Prenex Implicits fnd ins rem supp nil. Section Laws. Variables (K : ordType) (V : Type). Notation fmap := (finMap K V). -Notation nil := (nil K V). +Notation nil := (@nil K V). (* `path_le` specialized to `transitive ord` *) Lemma ord_path (x y : K) s : @@ -472,7 +473,7 @@ End Laws. Section Append. Variable (K : ordType) (V : Type). Notation fmap := (finMap K V). -Notation nil := (nil K V). +Notation nil := (@nil K V). Lemma seqof_ins k v (s : fmap) : path ord k (supp s) -> seq_of (ins k v s) = (k, v) :: seq_of s. @@ -515,7 +516,8 @@ Qed. (* forward induction principle *) Lemma fmap_ind' (P : fmap -> Prop) : - P nil -> (forall k v s, path ord k (supp s) -> P s -> P (ins k v s)) -> + P nil -> + (forall k v s, path ord k (supp s) -> P s -> P (ins k v s)) -> forall s, P s. Proof. move=>H1 H2; case; elim=>[|[k v] s IH] /= H. @@ -547,7 +549,6 @@ rewrite (_ : FinMap _ = ins k v (FinMap S)). by rewrite fmapE /= first_ins'. Qed. - Fixpoint fcat' (s1 : fmap) (s2 : seq (K * V)) {struct s2} : fmap := if s2 is (k, v)::t then fcat' (ins k v s1) t else s1. @@ -684,7 +685,7 @@ End FMapInd. Section Filtering. Variables (K : ordType) (V : Type). Notation fmap := (finMap K V). -Notation nil := (nil K V). +Notation nil := (@nil K V). Definition kfilter' (p : pred K) (s : fmap) := filter (fun kv => p kv.1) (seq_of s). @@ -788,7 +789,7 @@ End Filtering. Section DisjointUnion. Variable (K : ordType) (V : Type). Notation fmap := (finMap K V). -Notation nil := (nil K V). +Notation nil := (@nil K V). Definition disj (s1 s2 : fmap) := all (predC (fun x => x \in supp s2)) (supp s1). @@ -860,7 +861,9 @@ case: (k \in supp s2)=>//=; first by rewrite andbF. by rewrite -!(disjC s) IH. Qed. -Lemma fcatC (s1 s2 : fmap) : disj s1 s2 -> fcat s1 s2 = fcat s2 s1. +Lemma fcatC (s1 s2 : fmap) : + disj s1 s2 -> + fcat s1 s2 = fcat s2 s1. Proof. rewrite /fcat. elim/fmap_ind': s2 s1=>[|k v s2 L IH] s1 /=; first by rewrite fcat_nil'. @@ -869,7 +872,8 @@ by rewrite fcat_ins' // -IH // seqof_ins //= -fcat_ins' ?notin_path. Qed. Lemma fcatA (s1 s2 s3 : fmap) : - disj s2 s3 -> fcat (fcat s1 s2) s3 = fcat s1 (fcat s2 s3). + disj s2 s3 -> + fcat (fcat s1 s2) s3 = fcat s1 (fcat s2 s3). Proof. move=>H. elim/fmap_ind': s3 s1 s2 H=>[|k v s3 L IH] s1 s2 /=; first by rewrite !fcats0. @@ -901,7 +905,8 @@ by apply: cancel_ins H5; rewrite supp_fcat negb_or /= ?H1?H3 H. Qed. Lemma fcatKs (s s1 s2 : fmap) : - disj s s1 && disj s s2 -> fcat s s1 = fcat s s2 -> s1 = s2. + disj s s1 && disj s s2 -> + fcat s s1 = fcat s s2 -> s1 = s2. Proof. case/andP=>H1 H2. rewrite (fcatC H1) (fcatC H2); apply: fcatsK. @@ -1012,7 +1017,7 @@ Lemma mapf_fcat s1 s2 : mapf (fcat s1 s2) = fcat (mapf s1) (mapf s2). Proof. elim/fmap_ind': s2 s1=>[|k v s2 H IH] s1 /=. - rewrite fcats0; set j := FinMap _. - by rewrite (_ : j = nil K V) ?fcat0s //; apply/fmapE. + by rewrite (_ : j = @nil K V) ?fcat0s //; apply/fmapE. by rewrite fcat_sins mapf_ins IH -fcat_sins mapf_ins. Qed. @@ -1041,26 +1046,25 @@ Variables (A B: ordType) (V C: Type). Definition foldfmap g (e: C) (s: finMap A V) := foldr g e (seq_of s). - Lemma foldf_nil g e : foldfmap g e (@nil A V) = e. Proof. by rewrite /foldfmap //=. Qed. Lemma foldf_ins g e k v f: - path ord k (supp f) -> - foldfmap g e (ins k v f) = g (k, v) (foldfmap g e f). + path ord k (supp f) -> + foldfmap g e (ins k v f) = g (k, v) (foldfmap g e f). Proof. by move=> H; rewrite /foldfmap //= seqof_ins //. Qed. End FoldFMap. Section KeyMap. Section MapDef. -Variables (A B: ordType) (V : Type). +Variables (A B : ordType) (V : Type). Variable (f: A -> B). Hypothesis Hf : forall x y, strictly_increasing f x y. Definition mapk (m : finMap A V) : finMap B V := - foldfmap (fun p s => ins (f (key p)) (value p) s) (nil B V) m. + foldfmap (fun p s => ins (f (key p)) (value p) s) nil m. (* mapK preserves sorted *) @@ -1076,17 +1080,18 @@ rewrite {1}/supp //= {1}seqof_ins // /= => /andP [H]; move/IH=>H1. by rewrite /mapk foldf_ins // /supp /= seqof_ins //= H1 andbT (Hf H). Qed. -Lemma mapk_nil : mapk (nil A V) = nil B V. +Lemma mapk_nil : mapk (@nil A V) = @nil B V. Proof. by rewrite /mapk //=. Qed. Lemma mapk_ins k v s : - path ord k (supp s) -> - mapk (ins k v s) = ins (f k) v (mapk s). + path ord k (supp s) -> + mapk (ins k v s) = ins (f k) v (mapk s). Proof. by move=> H; rewrite /mapk foldf_ins =>//. Qed. End MapDef. + Arguments mapk {A B V} f m. -Variables (A B C : ordType)(V : Type)(f : A -> B) (g : B -> C). +Variables (A B C : ordType) (V : Type) (f : A -> B) (g : B -> C). Hypothesis Hf : forall x y, strictly_increasing f x y. Lemma mapk_id m : @mapk A A V id m = m. @@ -1312,7 +1317,7 @@ case: eqP=>/=; last by case: eqP=>// _ _; apply: IH. by move=>->{k1}; rewrite eq_refl; case=><- [<-]. Qed. -Lemma zunit0 : zip_unit (nil K V) = nil K V. +Lemma zunit0 : zip_unit (@nil K V) = @nil K V. Proof. by apply/fmapE. Qed. Lemma zunit_ins f k v : @@ -1329,7 +1334,7 @@ Lemma zunit_fcat f1 f2 : Proof. elim/fmap_ind': f2 f1=>[|k v f2 H IH] f1 /=. - rewrite fcats0; set j := FinMap _. - by rewrite (_ : j = nil K V) ?fcat0s //; apply/fmapE. + by rewrite (_ : j = @nil K V) ?fcat0s //; apply/fmapE. by rewrite fcat_sins zunit_ins IH -fcat_sins zunit_ins. Qed. diff --git a/core/prelude.v b/core/prelude.v index 8d45ce5..1302659 100644 --- a/core/prelude.v +++ b/core/prelude.v @@ -1190,8 +1190,8 @@ Notation "[ 'ffun' x : aT => E ]" := (finfun (fun x : aT => E)) Notation "[ 'ffun' x => E ]" := (@finfun _ (fun=> _) (fun x => E)) (at level 0, x name, format "[ 'ffun' x => E ]") : fun_scope. -Notation "['ffun' => E ]" := [ffun _ => E] - (at level 0, format "['ffun' => E ]") : fun_scope. +Notation "[ 'ffun' => E ]" := [ffun _ => E] + (at level 0, format "[ 'ffun' => E ]") : fun_scope. Section IteratedNotation. Variables (T : finType) (Us : T -> Type). diff --git a/core/seqperm.v b/core/seqperm.v index 012b36f..962d1a7 100644 --- a/core/seqperm.v +++ b/core/seqperm.v @@ -204,6 +204,22 @@ apply: (@pperm_trans (x::s1++t1))=>//; apply: (@pperm_trans (x::s2++t2))=>//. by apply/pperm_cons. Qed. +Lemma pperm_rcons s1 s2 x : + perm s1 s2 <-> + perm (rcons s1 x) (rcons s2 x). +Proof. +rewrite -!cats1; split=>[|H]; first by apply: pperm_cat2rL. +by rewrite -[s1]cats0 -[s2]cats0 -(pperm_cat_cons x). +Qed. + +Lemma pperm_rev s : perm s (rev s). +Proof. +elim: s=>[|x xs IH] //=. +rewrite rev_cons -cats1. +apply: (@pperm_trans ([:: x] ++ rev xs)); last by apply: pperm_catC. +by rewrite pperm_cons. +Qed. + End Permutations. #[export] Hint Resolve pperm_refl pperm_catC pperm_cons_catCA diff --git a/examples/congmath.v b/examples/congmath.v index a2abdd3..a0c8f61 100644 --- a/examples/congmath.v +++ b/examples/congmath.v @@ -1,5 +1,5 @@ (* -Copyright 2010 IMDEA Software Institute +Copyright 2009 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 @@ -874,7 +874,7 @@ Proof. by rewrite /CRel !clos_clos !orrA; apply: const_rep. Qed. (* table and the pending list all start empty. *) Definition init : data := - Data [ffun x => x] [ffun c => [:: c]] [ffun c => [::]] (nil _ _) [::]. + Data [ffun x => x] [ffun c => [:: c]] [ffun c => [::]] nil [::]. Lemma initP : propagate_inv init /\ CRel init <~> closure (graph (@const s)). Proof. @@ -1792,3 +1792,6 @@ Add Parametric Morphism s : (@closure s) with (fun e1 e2 => e1 <~> e2) as closure_morph'. Proof. by apply: closE. Qed. +(* empty congruence only relates constant symbols to themselves *) +Definition empty_cong s := closure (graph (@const s)). +Arguments empty_cong : clear implicits. diff --git a/examples/congprog.v b/examples/congprog.v index 78dbcc0..1aa902d 100644 --- a/examples/congprog.v +++ b/examples/congprog.v @@ -1,3 +1,16 @@ +(* +Copyright 2009 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 HB Require Import structures. From mathcomp Require Import ssreflect ssrbool ssrnat eqtype ssrfun. From mathcomp Require Import div finset seq fintype finfun choice. @@ -23,137 +36,6 @@ From htt Require Import kvmaps hashtab congmath. Notation finE := finset.inE. Prenex Implicits sepit. -Arguments nil {K V}. - -From mathcomp Require Import bigop. -From Coq Require Import Setoid. -Prenex Implicits star. - -Add Parametric Morphism U : (@star U) with signature - @Eq_Pred _ _ ==> @Eq_Pred _ _ ==> @Eq_Pred _ _ as star_morph. -Proof. -move=>x y E x1 y1 E1 m /=. -split; case=>h1 [h2][-> H1 H2]; exists h1, h2; split=>//. - by apply/E. by apply/E1. - by apply/E. by apply/E1. -Qed. - -Lemma sepit0' (U : pcm) A (f : A -> Pred U) : - IterStar.sepit [::] f <~> emp. -Proof. exact: IterStar.sepit0. Qed. - -Lemma sepit_cons' (U : pcm) A x (s : seq A) (f : A -> Pred U) : - IterStar.sepit (x :: s) f <~> - f x # IterStar.sepit s f. -Proof. exact: IterStar.sepit_cons. Qed. - -Lemma sepitX (U : pcm) A (s : seq A) (f : A -> Pred U) : - IterStar.sepit s f <~> - \big[star/emp]_(i <- s) f i. -Proof. -elim: s=>[|x s IH]. -- rewrite sepit0'. rewrite big_nil. by []. -rewrite sepit_cons'. -rewrite big_cons. -move=>m. -split; case=>h1 [h2][Em H1 H2]; exists h1, h2; split=>//. -- by rewrite -IH. -by rewrite IH. -Qed. - -Lemma Iter_sepit0 (U : pcm) A (f : A -> Pred U) : - IterStar.sepit [::] f <~> emp. -Proof. by rewrite sepitX big_nil. Qed. - -Lemma sepit_cons (U : pcm) A x (s : seq A) (f : A -> Pred U) : - IterStar.sepit (x::s) f <~> f x # IterStar.sepit s f. -Proof. by rewrite sepitX big_cons sepitX. Qed. - -(* U has the laws of commutative monoids from bigop *) -(* but doesn't really, as it's up to extensional equality *) -Lemma starA (U : pcm) (x y z : Pred U) : - x # y # z <~> (x # y) # z. -Proof. -move=>m /=; split; case=>h1 [h2][H1 H2]. -- case=>h3 [h4][H3 H4 H5]. - subst h2. - eexists (h1 \+ h3), h4. rewrite -joinA. split=>//. - by exists h1, h3. -case: H2=>h3 [h4][? H2 H3 H4]. subst h1. -exists h3, (h4 \+ h2). rewrite joinA. split=>//. -by exists h4, h2. -Qed. - -Lemma starC (U : pcm) (x y : Pred U) : - x # y <~> y # x. -Proof. -move=>m /=. -split. -- case=>h1 [h2][H H1 H2]. - exists h2, h1. rewrite joinC. by []. -case=>h1 [h2][H H1 H2]. -exists h2, h1. by rewrite joinC. -Qed. - -Lemma starL (U : pcm) (x : Pred U) : - emp # x <~> x. -Proof. -move=>m /=. -split. -- case=>h1 [h2][->->]. rewrite unitL. by []. -move=>H. -exists Unit, m. rewrite unitL. by []. -Qed. - -Lemma starR (U : pcm) (x : Pred U) : - x # emp <~> x. -Proof. by rewrite starC starL. Qed. - -(* -HB.instance Definition _ (U : pcm) := - Monoid.isComLaw.Build (Pred U) emp (@star U) (@starA U) (@starC U) (@starL U). -*) - -Lemma big_sepit_seq (U : pcm) A (s : seq A) (f : A -> U) m : - m = \big[join/Unit]_(i <- s) f i <-> - m \In IterStar.sepit s (fun i h => h = f i). -Proof. -rewrite sepitX. -elim: s m=>[|x xs IH] /= m. -- by rewrite !big_nil. -rewrite !big_cons. -split. -- move=>E. rewrite InE. exists (f x), (\big[join/Unit]_(j <- xs) f j). - split=>//. rewrite -IH. by []. -case=>h1 [h2][Em H1 H2]. rewrite -toPredE /= in H1. -rewrite Em H1. -congr (_ \+ _). -by rewrite IH. -Qed. - -Lemma Iter_sepit_emp (U : pcm) (A : eqType) (s : seq A) (f : A -> Pred U) : - (forall x, x \in s -> f x <~> emp (U:=U)) -> - IterStar.sepit s f <~> emp. -Proof. -move=>H. -rewrite sepitX. -elim: s H=>[|a xs IH] H. -- by rewrite big_nil. -rewrite big_cons. -rewrite H; last by rewrite inE eqxx. -rewrite starL IH //. -move=>x X. apply: H. -by rewrite inE X orbT. -Qed. - - -Lemma big_sepit (U : pcm) (I : finType) (s : {set I}) (f : I -> U) m : - m = \big[join/Unit]_(i in s) (f i) <-> - m \In sepit s (fun i h => h = f i). -Proof. by rewrite /sepit /= -big_sepit_seq -big_enum. Qed. - -(* empty congruence only relates constant symbols to themselves *) -Definition empty_cong s := closure (graph (@const s)). (*************) (* Signature *) @@ -295,8 +177,8 @@ Program Definition init : htb <-- kvm_new (htab V hash10); p <-- alloc null; ret (Ptrs rx cl ul htb p)). -Next Obligation. -case=>i [pf][/= f][hc][hct][->{i} Hc Hct]. +Next Obligation. +move=>_ cl loop k [i][pf][/= f][hc][hct][->{i} Hc Hct]. case: decP=>[{}pf|] /=; last first. - case: (ltngtP k n) pf=>// Ekn _ _; step=>_. exists f, hc, hct; split=>//. @@ -313,20 +195,20 @@ apply: tableP2 Hct=>// a. - by rewrite !finE ltnS indx_injE; case: ltngtP. by rewrite !finE !ffunE indx_injE; case: eqP=>// ->; rewrite ltnn. Qed. -Next Obligation. -case=>_ ->; apply: [stepE]=>//= rx hr Er; apply: [stepU]=>//= cl hc Ec. +Next Obligation. +case=>_ ->; apply: [stepE]=>//= rx hr Er; apply: [stepU]=>//= cl hc Ec. apply: [stepX]@hc=>//=. - split=>//; exists [ffun x => null], hc, Unit; rewrite unitR. by split=>//; rewrite (_ : [set c | indx c < 0] = set0) // sepit0. case=>_ [f][hc'][hrest][-> Hc' Hrest]. apply: [stepU]=>//= ul hu Ehu; apply: [stepU]=>//= htb ht Ht. -set d := Data [ffun x => x] [ffun c => [:: c]] [ffun c => [::]] (@nil K V) [::]. +set d := Data [ffun x => x] [ffun c => [:: c]] [ffun c => [::]] (@nil K V) [::]. step=>px; step; exists d; split; last by case: (initP s). split=>[a b|/=]; first by rewrite !ffunE !inE. exists f, [ffun s => null]. rewrite (_ : px :-> null \+ _ = hr \+ ((hc' \+ hrest) \+ (hu \+ Unit \+ (ht \+ (px :-> null \+ Unit))))); last by rewrite unitR; heap_congr. -hhauto; rewrite /sepit Iter_sepit_emp // => k. +hhauto; rewrite /sepit sepitseq_emp // => k. by rewrite /utab/table !ffunE; split=>//; case=>_ ->. Qed. @@ -374,7 +256,7 @@ Program Definition join_hclass (a' b' : symb) : Array.write r s b';; loop tt)) tt). Next Obligation. -case=>[[[/= d ct] ut]][i][H N] /=. +move=>a' b' loop [][[[/= d ct] ut]][i][H N] /=. case: H=>/= rh [_][->{i} Rh][_][h][->][th][ctx][->] Th Ctx H. rewrite (sepitT1 a') in Ctx; case: Ctx=>cta [w][->{ctx}Cta]. rewrite (sepitS b') !finE eq_sym {1}N /=. @@ -419,7 +301,7 @@ case: (x =P a')=>// _; case: (x =P b')=>// _. by rewrite Eca rev_cons cat_rcons. Qed. Next Obligation. -case=>d [i][[C]][/= ct][ut H] N. +move=>a' b' [d][i][[C]][/= ct][ut H] N. apply: [gE (d, ct, ut)]=>//= [[m]][ct1][ut1] W _. set d' := (Data _ _ _ _ _) in W; suff E : join_class d a' b' = d'. - by split; [apply: join_class_classP|rewrite E; eauto]. @@ -459,7 +341,7 @@ Program Definition join_huse (a' b' : symb) : Array.write ulist b' a;; loop tt)) tt). Next Obligation. -case=>[/= d][i][don][N][H Eu]. +move=>a' b' loop [][/= d][i][don][N][H Eu]. set d1 := join_use' d a' b' don in H Eu. case: H=>C [/= ct][ut][rh][_][->{i} Rh][cth][_][-> Htc][_][_][->] [ru][hu][-> Ut Hu][ht][_][-> Ht][p'][_][hp][->-> Hp]. @@ -521,7 +403,7 @@ apply: tableP R=>a /=; rewrite !finE andbT /ut3/ut2 !ffunE /= ?ffunE /=; by case/andP=>/negbTE -> /negbTE ->. Qed. Next Obligation. -by case=>d [i][N H]; apply: [gE d]=>[||??[]] //; exists [::]. +by move=>a' b' [d][i][N H]; apply: [gE d]=>[||??[]] //; exists [::]. Qed. Definition pT : Type := forall x : unit, @@ -552,7 +434,7 @@ Program Definition hpropagate := join_huse a' b';; loop tt)) tt. Next Obligation. -case=>/= d [i][C][/= ct][ut][hr][_][-> Hr][hc][_][-> Hc] +move=>loop [[d]][i][C][/= ct][ut][hr][_][-> Hr][hc][_][-> Hc] [hu][_][-> Hu][ht][_][-> Ht][q][_][hp][->] -> Hp. step; case: (q =P null) Hp=>[->{q}|/eqP N] Hp. - apply: vrfV=>V; case/(lseq_null (validX V)): Hp=>{V} Ep /= ->{hp}. @@ -611,7 +493,7 @@ Program Definition merge (e : Eq s) : Array.write ulist c2' x) end. Next Obligation. -case=>R [_][d][[C]][/= ct][ut][hr][_][-> Hr][hc][_][-> Hc][_][_] +move=>e a b [R][_][d][[C]][/= ct][ut][hr][_][-> Hr][hc][_][-> Hc][_][_] [->][hu][hu'][-> Hu Hu'][ht][_][-> Ht][q][_][hp][->] -> Hp [PI][Ep Erel]. step; apply: [stepX (pending d)]@hp=>//= x _ [r0][{Hp}hp][-> Hp]. set d1:=Data (rep d) (class d) (use d) (lookup d) (simp_pend a b :: pending d). @@ -627,7 +509,7 @@ case: (propagatePE L)=>L1 [L2 L3]; exists (propagate d1); do 3!split=>//. by rewrite -L3 -Erel clos_clos /CRel -!orrA orrAC. Qed. Next Obligation. -case=>R [_][d][[C]][/= ct][ut][hr][_][-> Hr][hc][_][-> H] +move=>e c c1 c2 [R][_][d][[C]][/= ct][ut][hr][_][-> Hr][hc][_][-> H] [_][_][->][hu][hu'][-> Hu Hu'][ht][_][-> Ht][q][_] [hp][->] -> Hp [PI][Ep Erel]; set cx := (c, c1, c2). apply: [stepX rep d, hr]@hr=>//= _ _ [->->]. @@ -716,7 +598,7 @@ Program Definition hnorm := else ret (app u1 u2) end)). Next Obligation. -case=>d [_][Ci][/= ct][ut][hr][hrest][-> Hr Hrest]. +move=>hnorm t [d][_][Ci][/= ct][ut][hr][hrest][-> Hr Hrest]. case: t=>[a|t1 t2]. - apply: [stepX rep d, hr]@hr=>//= _ _ [->->]. by step; do 2!split=>//; exists ct, ut; hhauto. @@ -744,7 +626,7 @@ Program Definition check t1 t2 : u2 <-- hnorm t2; ret (u1 == u2)). Next Obligation. -case=>R [h][d][H][[RI X]][Ep PI]. +move=>t1 t2 [R][h][d][H][[RI X]][Ep PI]. apply: [stepX d]@h=>//= _ {H}h [-> H]. apply: [stepX d]@h=>//= _ {H}h [-> H]. step; split; first by exists d. diff --git a/examples/hashtab.v b/examples/hashtab.v index e73e343..4c8fcca 100644 --- a/examples/hashtab.v +++ b/examples/hashtab.v @@ -29,26 +29,22 @@ From htt Require Import array kvmaps. Module Type Hashtab_sig. Parameter root : forall {K : ordType} {V : Type} (buckets : dkvm K V) {n : nat} (hash : K -> 'I_n), Set. -Parameter null_root : forall {K : ordType} {V : Type} {buckets : dkvm K V} - {n : nat} {hash : K -> 'I_n}, root buckets hash. -Parameter shape : forall {K : ordType} {V : Type} {buckets : dkvm K V} - {n : nat} {hash : K -> 'I_n}, root buckets hash -> {finMap K -> V} -> - Pred heap. -Parameter new : forall {K : ordType} {V : Type} {buckets : dkvm K V} - {n : nat} {hash : K -> 'I_n}, - STsep (emp, [vfun (x : root buckets hash) h => h \In shape x (nil K V)]). -Parameter free : forall {K : ordType} {V : Type} {buckets : dkvm K V} - {n : nat} {hash : K -> 'I_n} (x : root buckets hash), +Section HashTab. +Context {K : ordType} {V : Type} {buckets : dkvm K V}. +Context {n : nat} {hash : K -> 'I_n}. +Parameter null_root : root buckets hash. +Parameter shape : root buckets hash -> {finMap K -> V} -> Pred heap. +Parameter new : + STsep (emp, [vfun (x : root buckets hash) h => h \In shape x nil]). +Parameter free : forall x : root buckets hash, STsep {s} (shape x s, [vfun _ : unit => emp]). -Parameter insert : forall {K : ordType} {V : Type} {buckets : dkvm K V} - {n : nat} {hash : K -> 'I_n} (x : root buckets hash) k v, +Parameter insert : forall (x : root buckets hash) k v, STsep {s} (shape x s, [vfun (_ : unit) h => h \In shape x (ins k v s)]). -Parameter remove : forall {K : ordType} {V : Type} {buckets : dkvm K V} - {n : nat} {hash : K -> 'I_n} (x : root buckets hash) k, +Parameter remove : forall (x : root buckets hash) k, STsep {s} (shape x s, [vfun (_ : unit) h => h \In shape x (rem k s)]). -Parameter lookup : forall {K : ordType} {V : Type} {buckets : dkvm K V} - {n : nat} {hash : K -> 'I_n} (x : root buckets hash) k, +Parameter lookup : forall (x : root buckets hash) k, STsep {s} (shape x s, [vfun y h => h \In shape x s /\ y = fnd k s]). +End HashTab. End Hashtab_sig. Module HashTab : Hashtab_sig. @@ -61,7 +57,6 @@ Context (K : ordType) (V : Type) {buckets : dkvm K V} {n : nat} {hash : K -> 'I_n}. Notation KVshape := (@dkvm_shape _ _ buckets). Notation table := (table KVshape). -Notation nil := (nil K V). Notation root := (root buckets hash). (* hash table is specified by a single finMap *) @@ -80,7 +75,7 @@ Definition shape (x : root) (s : finMap K V) : Pred heap := Definition new_loopinv x := forall k, STsep (fun h => k <= n /\ exists tab, h \In Array.shape x tab # - sepit [set x:'I_n | x < k] (table tab (fun=> nil)), + sepit [set x:'I_n | x < k] (table tab (fun=>nil)), [vfun y => shape y nil]). Program Definition new : STsep (emp, [vfun y => shape y nil]) := @@ -306,7 +301,7 @@ HB.instance Definition _ K V (buckets : dkvm K V) n (hash : K -> 'I_n) := HashTab.new HashTab.free HashTab.insert HashTab.remove HashTab.lookup. (* htab is specific simple hash tab where buckets are association lists *) -Definition htab K V n hash := @hashtab K V (dalist K V) n hash. +Definition htab {K} V {n} hash := @hashtab K V (dalist K V) n hash. HB.instance Definition _ K V n hash := KVM.on (@htab K V n hash). diff --git a/examples/kvmaps.v b/examples/kvmaps.v index e7f8444..22a6516 100644 --- a/examples/kvmaps.v +++ b/examples/kvmaps.v @@ -34,7 +34,7 @@ From htt Require Import options model heapauto. HB.mixin Record isDKVM (K : ordType) (V : Type) (root : Set) := { dkvm_null : root; dkvm_shape : root -> {finMap K -> V} -> Pred heap; - dkvm_new : STsep (emp, [vfun x => dkvm_shape x (nil K V)]); + dkvm_new : STsep (emp, [vfun x => dkvm_shape x nil]); dkvm_free : forall x, STsep {s} (dkvm_shape x s, [vfun _ : unit => emp]); dkvm_insert : forall x k v, @@ -42,7 +42,8 @@ HB.mixin Record isDKVM (K : ordType) (V : Type) (root : Set) := { dkvm_remove : forall x k, STsep {s} (dkvm_shape x s,[vfun y => dkvm_shape y (rem k s)]); dkvm_lookup : forall x k, - STsep {s} (dkvm_shape x s, [vfun y m => m \In dkvm_shape x s /\ y = fnd k s])}. + STsep {s} (dkvm_shape x s, + [vfun y m => m \In dkvm_shape x s /\ y = fnd k s])}. #[short(type="dkvm")] HB.structure Definition DKVM K V := {root of @isDKVM K V root}. @@ -59,7 +60,7 @@ Arguments dkvm_new {K V}. HB.mixin Record isKVM (K : ordType) (V : Type) (root : Set) := { kvm_null : root; kvm_shape : root -> {finMap K -> V} -> Pred heap; - kvm_new : STsep (emp, [vfun x => kvm_shape x (nil K V)]); + kvm_new : STsep (emp, [vfun x => kvm_shape x nil]); kvm_free : forall x, STsep {s} (kvm_shape x s, [vfun _ : unit => emp]); kvm_insert : forall x k v, @@ -67,7 +68,8 @@ HB.mixin Record isKVM (K : ordType) (V : Type) (root : Set) := { kvm_remove : forall x k, STsep {s} (kvm_shape x s,[vfun _ : unit => kvm_shape x (rem k s)]); kvm_lookup : forall x k, - STsep {s} (kvm_shape x s, [vfun y m => m \In kvm_shape x s /\ y = fnd k s])}. + STsep {s} (kvm_shape x s, + [vfun y m => m \In kvm_shape x s /\ y = fnd k s])}. #[short(type="kvm")] HB.structure Definition KVM K V := {root of @isKVM K V root}. @@ -80,19 +82,20 @@ Arguments kvm_new {K V}. Module Type DAList_sig. Parameter root : ordType -> Type -> Set. -Parameter null_root : forall {K : ordType} {V : Type}, root K V. -Parameter shape : forall {K : ordType} {V : Type}, - root K V -> {finMap K -> V} -> Pred heap. -Parameter new : forall {K : ordType} {V : Type}, - STsep (emp, [vfun (x : root K V) h => h \In shape x (nil K V)]). -Parameter free : forall {K : ordType} {V : Type} (x : root K V), +Section DAlist. +Context {K : ordType} {V : Type}. +Parameter null_root : root K V. +Parameter shape : root K V -> {finMap K -> V} -> Pred heap. +Parameter new : STsep (emp, [vfun (x : root K V) h => h \In shape x nil]). +Parameter free : forall x : root K V, STsep {s} (shape x s, [vfun _ : unit => emp]). -Parameter insert : forall {K : ordType} {V : Type} (x : root K V) k v, +Parameter insert : forall (x : root K V) k v, STsep {s} (shape x s, [vfun (y : root K V) h => h \In shape y (ins k v s)]). -Parameter remove : forall {K : ordType} {V : Type} (x : root K V) k, +Parameter remove : forall (x : root K V) k, STsep {s} (shape x s, [vfun (y : root K V) h => h \In shape y (rem k s)]). -Parameter lookup : forall {K : ordType} {V : Type} (x : root K V) k, +Parameter lookup : forall (x : root K V) k, STsep {s} (shape x s, [vfun y m => m \In shape x s /\ y = fnd k s]). +End DAlist. End DAList_sig. Module DAList : DAList_sig. @@ -101,7 +104,6 @@ Definition null_root K V : @root K V := null. Section AssocList. Variables (K : ordType) (V : Type). Notation fmap := (finMap K V). -Notation nil := (nil K V). (* single entry of the map as a triple of heap cells *) Definition entry (p q : ptr) (k : K) (v : V) : heap := @@ -635,19 +637,20 @@ HB.instance Definition _ K V := Module Type AList_sig. Parameter root : ordType -> Type -> Set. -Parameter null_root : forall {K : ordType} {V : Type}, root K V. -Parameter shape : forall {K : ordType} {V : Type}, - root K V -> {finMap K -> V} -> Pred heap. -Parameter new : forall {K : ordType} {V : Type}, - STsep (emp, [vfun (x : root K V) h => h \In shape x (nil K V)]). -Parameter free : forall {K : ordType} {V : Type} (x : root K V), +Section AList. +Context {K : ordType} {V : Type}. +Parameter null_root : root K V. +Parameter shape : root K V -> {finMap K -> V} -> Pred heap. +Parameter new : STsep (emp, [vfun (x : root K V) h => h \In shape x nil]). +Parameter free : forall (x : root K V), STsep {s} (shape x s, [vfun _ : unit => emp]). -Parameter insert : forall {K : ordType} {V : Type} (x : root K V) k v, +Parameter insert : forall (x : root K V) k v, STsep {s} (shape x s, [vfun (_ : unit) h => h \In shape x (ins k v s)]). -Parameter remove : forall {K : ordType} {V : Type} (x : root K V) k, +Parameter remove : forall (x : root K V) k, STsep {s} (shape x s, [vfun (_ : unit) h => h \In shape x (rem k s)]). -Parameter lookup : forall {K : ordType} {V : Type} (x : root K V) k, +Parameter lookup : forall (x : root K V) k, STsep {s} (shape x s, [vfun y m => m \In shape x s /\ y = fnd k s]). +End AList. End AList_sig. Module AList : AList_sig. @@ -655,7 +658,6 @@ Section AssocList. Definition root : ordType -> Type -> Set := fun _ _ => ptr. Variables (K : ordType) (V : Type). Notation fmap := (finMap K V). -Notation nil := (nil K V). Definition null_root : root K V := null. Definition shape (x : ptr) (f : fmap) : Pred heap := @@ -742,4 +744,3 @@ Notation alist := AList.root. HB.instance Definition _ K V := isKVM.Build K V (alist K V) AList.null_root AList.new AList.free AList.insert AList.remove AList.lookup. - diff --git a/examples/tree.v b/examples/tree.v index 9908835..2b59094 100644 --- a/examples/tree.v +++ b/examples/tree.v @@ -18,7 +18,7 @@ From pcm Require Import pred prelude seqext. (* arbitrarily branching tree *) Inductive tree A := TNode of A & seq (tree A). -Arguments TNode [A]. +Arguments TNode {A}. Section Tree. Context {A : Type}. @@ -29,26 +29,18 @@ Definition ch (t : tree A) := let: TNode _ ts := t in ts. Lemma tree_ext (t : tree A) : TNode (rt t) (ch t) = t. Proof. by case: t. Qed. -(* a leaf is a node with an empty list *) +(* leaf is a node with an empty list *) Definition lf a : tree A := TNode a [::]. Lemma tree_ind' (P : tree A -> Prop) : - (forall a l, All P l -> P (TNode a l)) -> - forall t, P t. -Proof. -move=> indu; fix H 1. -elim => a l; apply indu. -by elim: l. -Qed. + (forall a l, All P l -> P (TNode a l)) -> + forall t, P t. +Proof. by move=>indu; fix H 1; elim => a l; apply indu; elim: l. Qed. Lemma tree_rec' (P : tree A -> Type) : - (forall a l, AllT P l -> P (TNode a l)) -> - forall t, P t. -Proof. -move=>indu; fix H 1. -elim => a l; apply: indu. -by elim: l. -Qed. + (forall a l, AllT P l -> P (TNode a l)) -> + forall t, P t. +Proof. by move=>indu; fix H 1; elim => a l; apply: indu; elim: l. Qed. (* custom induction principles *) @@ -65,7 +57,7 @@ Fixpoint preorder (t : tree A) : seq A := let: TNode a ts := t in foldl (fun s t => s ++ preorder t) [::a] ts. -Corollary foldl_cat {B C} z (fs : seq B) (a : B -> seq C): +Lemma foldl_cat {B C} z (fs : seq B) (a : B -> seq C): foldl (fun s t => s ++ a t) z fs = z ++ foldl (fun s t => s ++ a t) [::] fs. Proof. @@ -73,7 +65,9 @@ apply/esym/fusion_foldl; last by rewrite cats0. by move=>x y; rewrite catA. Qed. -Lemma preorderE t : preorder t = rt t :: \big[cat/[::]]_(c <- ch t) (preorder c). +Lemma preorderE t : + preorder t = + rt t :: \big[cat/[::]]_(c <- ch t) (preorder c). Proof. case: t=>a cs /=; rewrite foldl_cat /=; congr (_ :: _). elim: cs=>/= [| c cs IH]; first by rewrite big_nil. @@ -108,7 +102,6 @@ Qed. End EncodeDecodeTree. - HB.instance Definition _ (A : eqType) := Equality.copy (tree A) (pcan_type (pcancel_tree A)). @@ -125,7 +118,8 @@ Coercion pred_of_tree (t : tree_eqclass) : {pred A} := mem_tree t. Canonical tree_predType := ssrbool.PredType (pred_of_tree : tree A -> pred A). Lemma in_tnode a t ts : - (t \in TNode a ts) = (t == a) || has (fun q => t \in q) ts. + (t \in TNode a ts) = + (t == a) || has (fun q => t \in q) ts. Proof. by []. Qed. (* frequently used facts about membership in a tree *) @@ -160,7 +154,7 @@ Qed. End TreeEq. -Arguments in_tnode2 [A x y a ts]. +Arguments in_tnode2 {A x y a ts}. #[export] Hint Resolve in_tnode1 : core. (* a simplified induction principle for eq types *) diff --git a/examples/union_find.v b/examples/union_find.v index b86f834..27041ea 100644 --- a/examples/union_find.v +++ b/examples/union_find.v @@ -17,58 +17,12 @@ From pcm Require Import options axioms pred seqext. From pcm Require Import prelude pcm unionmap natmap heap autopcm automap. From htt Require Import options model heapauto tree. -(* TODO: remove these after upgrading FCSL-PCM to 1.9+ *) -Corollary foldr_join A (U : pcm) h (s : seq A) (a : A -> U): - foldr (fun t h => h \+ a t) h s = - h \+ foldr (fun t h => h \+ a t) Unit s. -Proof. -apply/esym/fusion_foldr; last by rewrite unitR. -by move=>x y; rewrite joinA. -Qed. - -Section BigOpsUM. -Variables (K : ordType) (C : pred K) (T : Type). -Variables (U : union_map C T). -Variables (I : Type) (f : I -> U). - -Lemma big_find_someXP (xs : seq I) P a v : - find a (\big[join/Unit]_(i <- xs | P i) f i) = Some v -> - exists i, [/\ i \In xs, P i & find a (f i) = Some v]. -Proof. -by rewrite -big_filter=>/big_find_someX [i] /Mem_filter [H1 H2 H3]; exists i. -Qed. - -Lemma bigInXP (xs : seq I) P a v : - (a, v) \In \big[join/Unit]_(i <- xs | P i) f i -> - exists i, [/\ i \In xs, P i & (a, v) \In f i]. -Proof. by case/In_find/big_find_someXP=>x [X1 X2 /In_find]; exists x. Qed. - -End BigOpsUM. - -Section OMapBig. -Variables (K : ordType) (C : pred K) (T T' : Type). -Variables (U : union_map C T) (U' : union_map C T'). -Variables (I : Type) (f : I -> U). - -Lemma omapVUnBig (a : K * T -> option T') ts : - omap a (\big[join/Unit]_(x <- ts) f x) = - if valid (\big[join/Unit]_(x <- ts) f x) - then \big[join/Unit]_(x <- ts) omap a (f x) - else undef : U'. -Proof. -elim: ts=>[|t ts IH]; first by rewrite !big_nil omap0 valid_unit. -by rewrite !big_cons omapVUn IH; case: ifP=>// /validR ->. -Qed. - -End OMapBig. - (**************) (**************) (* Union-find *) (**************) (**************) - (******************) (* inverted trees *) (******************) @@ -263,7 +217,8 @@ Lemma In_tsetP x r t: reflect ((x, r) \In tset t (rt t)) [&& valid (tset t (rt t)), x \in t & r == rt t]. Proof. -apply/(iffP idP); rewrite In_find find_tset; first by case/and3P=>->-> /eqP ->. +apply/(iffP idP); rewrite In_find find_tset. +- by case/and3P=>->-> /eqP ->. by case: ifP=>// /andP [->->][->] /=. Qed. @@ -364,11 +319,13 @@ Proof. by move=>V; rewrite range_fsetE V inE. Qed. Lemma valid_fset_tset (ts : seq (tree ptr)) : - valid (fset ts) -> {in ts, forall i, valid (tset i (rt i))}. + valid (fset ts) -> + {in ts, forall i, valid (tset i (rt i))}. Proof. by move=>+ i Hi; rewrite fsetE big_valid_seq=>/andP [/allP /(_ i Hi)]. Qed. Lemma valid_flay_tlay (ts : seq (tree ptr)) : - valid (flay ts) -> {in ts, forall i, valid (tlay i (rt i))}. + valid (flay ts) -> + {in ts, forall i, valid (tlay i (rt i))}. Proof. by move=>+ i Hi; rewrite flayE big_valid_seq=>/andP [/allP /(_ i Hi)]. Qed. Lemma dom_flay (ts : seq (tree ptr)) x : @@ -512,7 +469,6 @@ apply/In_fsetP; rewrite Vs; apply/hasP; exists j=>//. by rewrite P E (flay_mem_eq V Xi Ti Xj Tj) eqxx. Qed. - Lemma froot_loop x r ts : (x, r) \In fset ts -> find x (flay ts) = Some (idyn x) -> @@ -522,7 +478,7 @@ elim: ts r=>[|t ts IH] r //= /In_find H1 H2. move: (dom_valid (find_some H1)) (dom_valid (find_some H2))=>V1 V2. move: H1 H2; rewrite !findUnL ?(dom_tset,dom_tlay,validL V1,validL V2) //. case A: (x \in t). - rewrite find_tset ifT; first by case=>E; move/tlay_rt_loop=>K; rewrite K -E. +- rewrite find_tset ifT; first by case=>E; move/tlay_rt_loop=>K; rewrite K -E. apply/andP; split=>//; first by apply: validL V1. by rewrite -In_find; apply: IH. Qed. @@ -562,13 +518,11 @@ Lemma change_tset ta a b: mapv [fun v => v with a |-> b] (tset ta a) = tset ta b. Proof. elim/tree_ind2: ta a=>c ts IH a //; rewrite !tsetE /= => V. -rewrite omapVUn omapPt /= eq_refl omapVUnBig !(validX V). +rewrite omapVUn omapPt /= eq_refl big_omapVUn !(validX V). congr (_ \+ _); apply: eq_big_seq=>i /[dup] K /mem_seqP K'. by rewrite IH // (big_validV (validX V) K'). Qed. - - (*******************) (* Shape predicate *) (*******************) @@ -581,6 +535,7 @@ Proof. by case=>ts [->->]; rewrite valid_flay_fset=>/andP []. Qed. (*******) (* NEW *) (*******) + (* Creates a new equivalence class with a single element *) Program Definition newT : @@ -596,14 +551,13 @@ Qed. (********) (* FIND *) (********) + (* Returns the canonical representative of the equivalence class of an element*) Definition find_tp (x : ptr) := STsep {rs r} (fun h => shape rs h /\ (x, r) \In rs, [vfun res h => shape rs h /\ res = r]). - - Program Definition find1 (x : ptr) : find_tp x := Do (let root := ffix (fun (go : forall x, find_tp x) (x : ptr) => Do (p <-- !x; @@ -626,6 +580,7 @@ Qed. (*********) (* UNION *) (*********) + (* Joins the equivalence classes of the two arguments *) Definition union_tp (x y : ptr) := STsep {rx ry m} @@ -633,7 +588,6 @@ Definition union_tp (x y : ptr) := STsep {rx ry m} [vfun res h => shape (mapv [fun v => v with rx |-> ry] m) h /\ res = ry]). - Program Definition union (x y : ptr) : union_tp x y := Do (x_rt <-- find1 x; y_rt <-- find1 y; diff --git a/pcm/heap.v b/pcm/heap.v index 0364474..c58692c 100644 --- a/pcm/heap.v +++ b/pcm/heap.v @@ -114,7 +114,7 @@ End NullLemmas. Notation base := (@UM.base ptr (fun k => k != null) (dynamic id)). Definition def h := if h is Def _ _ then true else false. -Definition empty := @Def (finmap.nil _ _) is_true_true. +Definition empty := @Def finmap.nil is_true_true. Definition upd k v h := if h is Def hs ns then if decP (@idP (k != null)) is left pf then diff --git a/pcm/pcm.v b/pcm/pcm.v index adb5b7a..98e204b 100644 --- a/pcm/pcm.v +++ b/pcm/pcm.v @@ -18,7 +18,7 @@ limitations under the License. (******************************************************************************) From HB Require Import structures. -From Coq Require Import ssreflect ssrbool ssrfun. +From Coq Require Import ssreflect ssrbool ssrfun Setoid. From mathcomp Require Import ssrnat eqtype seq bigop fintype finset finfun. From pcm Require Import options axioms prelude seqperm pred seqext. @@ -1246,7 +1246,7 @@ Hint Resolve pleq_undef : core. (* folding functions that are morphisms in the first argument *) -Section PCMfold. +Section PCMfoldLeft. Variables (A : Type) (R : pcm) (a : R -> A -> R). Hypothesis H : (forall x y k, a (x \+ y) k = a x k \+ y). @@ -1254,14 +1254,15 @@ Hypothesis H : (forall x y k, a (x \+ y) k = a x k \+ y). 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 foldl_cat /= foldl_cat; congr foldl. +elim: {s2} s1 z0=>[|k ks IH] z0 //=; rewrite IH; congr foldl. rewrite -[a z0 k]unitL H -[z0]unitL !H. -by rewrite -{2}[a Unit x]unitL H joinCA joinA. +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. + 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. @@ -1273,7 +1274,54 @@ 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. +End PCMfoldLeft. + +Section PCMfoldRight. +Variables (A : Type) (R : pcm) (a : A -> R -> R). +Hypothesis H : (forall k x y, a k (x \+ y) = a k x \+ y). + +Lemma foldr_helper (s1 s2 : seq A) (z0 : R) x : + foldr a z0 (s1 ++ x :: s2) = foldr a z0 (s1 ++ rcons s2 x). +Proof. +rewrite -!foldl_rev rev_cat rev_cons cat_rcons. +rewrite foldl_helper; last by move=>*; rewrite H. +by rewrite -[x :: _ ++ _]revK rev_cons rev_cat !revK rcons_cat. +Qed. + +Lemma foldr_perm (s1 s2 : seq A) (z0 : R) : + perm s1 s2 -> + foldr a z0 s1 = foldr a z0 s2. +Proof. +move=>P; rewrite -!foldl_rev; apply: foldl_perm. +- by move=>*; rewrite H. +apply/(pperm_trans _ (pperm_rev _))/pperm_sym. +by apply/(pperm_trans _ (pperm_rev _))/pperm_sym. +Qed. + +Lemma foldr_init (s : seq A) (x y : R) : + foldr a (x \+ y) s = foldr a x s \+ y. +Proof. by rewrite -!foldl_rev foldl_init. Qed. + +End PCMfoldRight. + +(* fold-join lemmas *) + +Lemma foldl_join A (U : pcm) h (s : seq A) (a : A -> U): + foldl (fun h t => a t \+ h) h s = + foldl (fun h t => a t \+ h) Unit s \+ h. +Proof. +set g := fun x => x \+ h. +apply/esym/(fusion_foldl (g:=g)); last by rewrite /g unitL. +by move=>x y; rewrite joinA. +Qed. + +Lemma foldr_join A (U : pcm) h (s : seq A) (a : A -> U): + foldr (fun t h => h \+ a t) h s = + h \+ foldr (fun t h => h \+ a t) Unit s. +Proof. +apply/esym/fusion_foldr; last by rewrite unitR. +by move=>x y; rewrite joinA. +Qed. Section BigOps. Variables (U : pcm). @@ -1315,214 +1363,292 @@ End BigOps. (* separating conjunction aka star *) (***********************************) -Section Star. -Variable U : pcm. - -Definition star (p1 p2 : Pred U) : Pred U := +Definition star {U : pcm} (p1 p2 : Pred U) : Pred U := [Pred h | exists h1 h2, [ /\ h = h1 \+ h2, h1 \In p1 & h2 \In p2] ]. -Definition emp : Pred U := eq^~ Unit. -Definition top : Pred U := PredT. - -End Star. - -Arguments emp {U}. -Arguments top {U}. +Definition emp {U : pcm} : Pred U := eq^~ Unit. +Definition top {U : pcm} : Pred U := PredT. Notation "p1 '#' p2" := (star p1 p2) (at level 57, right associativity) : rel_scope. -(* iterated star *) +Add Parametric Morphism U : (@star U) with signature + @Eq_Pred _ _ ==> @Eq_Pred _ _ ==> @Eq_Pred _ _ as star_morph. +Proof. +move=>x y E x1 y1 E1 m /=; split; case=>h1 [h2][-> H1 H2]; +by exists h1, h2; split=>//; [apply/E|apply/E1]. +Qed. -Module IterStar. -Section IterStar. -Variables (U : pcm) (A : Type). +(* monoid laws for star hold under <~>, not = *) +(* thus, require fun/prop extensionality to make hold under = *) +(* working with <~> entails use of setoids *) +(* and prevents use of monoids from bigops *) +(* the latter isn't a big deal, however, as most lemmas *) +(* in bigops are about decidable/finite sets *) +(* and won't work with Prop anyway *) -Definition seqjoin (s : seq U) : U := - \big[join/Unit]_(i <- s) i. +Section StarMonoid. +Context {U : pcm}. +Implicit Type x y z : Pred U. -Definition sepit' (s : seq A) (f : A -> Pred U) : Pred U := - [Pred h | exists hs : seq U, - [ /\ size hs = size s, h = seqjoin hs & - All (fun '(a, h) => h \In f a) (zip s hs)]]. +Lemma starA x y z : x # y # z <~> (x # y) # z. +Proof. +move=>m /=; split; case=>h1 [h2][H1 H2]. +- case=>h3 [h4][H3 H4 H5]; subst h2. + eexists (h1 \+ h3), h4; rewrite -joinA; split=>//. + by exists h1, h3. +case: H2=>h3 [h4][? H2 H3 H4]; subst h1. +exists h3, (h4 \+ h2); rewrite joinA; split=>//. +by exists h4, h2. +Qed. + +Lemma starC x y : x # y <~> y # x. +Proof. +move=>m; split. +- by case=>h1 [h2][H H1 H2]; exists h2, h1; rewrite joinC. +by case=>h1 [h2][H H1 H2]; exists h2, h1; rewrite joinC. +Qed. + +Lemma starL x : emp # x <~> x. +Proof. +move=>m; split=>[|H]; first by case=>h1 [h2][->->]; rewrite unitL. +by exists Unit, m; rewrite unitL. +Qed. + +Lemma starR x : x # emp <~> x. +Proof. by rewrite starC starL. Qed. -Definition sepit s f := locked (sepit' s f). +Lemma starAC x y z : x # y # z <~> x # z # y. +Proof. by rewrite (starC y). Qed. -Lemma sepitE s f : sepit s f = sepit' s f. -Proof. by rewrite /sepit -lock. Qed. +Lemma starCA x y z : x # (y # z) <~> y # (x # z). +Proof. by rewrite starA (starC x) -starA. Qed. -Lemma sepit0 f : sepit [::] f =p emp. +(* +HB.instance Definition _ := + Monoid.isComLaw.Build (Pred U) emp star starA starC starL. +*) + +End StarMonoid. + + +(* star iterated over sequence *) + +Section IterStarSeq. +Context {U : pcm} {A : Type}. + +Notation seq_join hs := (\big[join/Unit]_(i <- hs) i). + +(* definition is locked to prevent automation from going inside *) +Definition sepit_seq (s : seq A) (f : A -> Pred U) : Pred U := + locked + [Pred h | exists hs : seq U, + [/\ size hs = size s, h = seq_join hs & + All (fun '(a, h) => h \In f a) (zip s hs)]]. + +Lemma sepit_seq_unlock s f : + sepit_seq s f = + [Pred h | exists hs : seq U, + [/\ size hs = size s, h = seq_join hs & + All (fun '(a, h) => h \In f a) (zip s hs)]]. +Proof. by rewrite /sepit_seq -lock. Qed. + +Lemma sepit_seq0 f : sepit_seq [::] f <~> emp. Proof. -move=>h; rewrite sepitE; split. -- by case=>/= hs [/size0nil -> -> _]; rewrite /seqjoin !big_nil. -by move=>->; exists [::]; rewrite /seqjoin !big_nil. +rewrite sepit_seq_unlock=>h; split. +- by case=>/= hs [/size0nil -> -> _]; rewrite !big_nil. +by move=>->; exists [::]; rewrite !big_nil. Qed. -Lemma sepit_cons x s f : sepit (x::s) f =p f x # sepit s f. +Lemma sepitseq_cons x s f : + sepit_seq (x::s) f <~> f x # sepit_seq s f. Proof. -move=>h; rewrite !sepitE; split. +rewrite !sepit_seq_unlock=>h; split. - case=>/=; case=>[|h0 hs]; case=>//= /eqP; rewrite eqSS =>/eqP Hs. - rewrite /seqjoin big_cons =>->[H0 H1]. - by exists h0, (seqjoin hs); do!split=>//; exists hs. + rewrite big_cons =>->[H0 H1]. + by exists h0, (seq_join hs); do!split=>//; exists hs. case=>h1[_][{h}-> H1][hs][E -> H]. -by exists (h1 :: hs); rewrite /= E /seqjoin !big_cons. +by exists (h1 :: hs); rewrite /= E !big_cons. +Qed. + +(* alternative def *) +(* appears simpler, but no much gain from \big ops *) +(* as star/emp are only monoids under <~>, not = *) +Lemma sepit_seqE s f : + sepit_seq s f <~> \big[star/emp]_(i <- s) f i. +Proof. +elim: s=>[|x s IH]; first by rewrite sepit_seq0 big_nil. +by rewrite sepitseq_cons big_cons IH. Qed. -Lemma sepit_cat s1 s2 f : sepit (s1 ++ s2) f =p sepit s1 f # sepit s2 f. +Lemma sepitseq_cat s1 s2 f : + sepit_seq (s1 ++ s2) f <~> + sepit_seq s1 f # sepit_seq s2 f. Proof. -rewrite !sepitE; elim: s1 s2=>[|x s1 IH] s2 h /=; split. -- case=>hs [E {h}-> H2]; rewrite -sepitE. - exists Unit, (seqjoin hs); rewrite unitL. - by split=>//; [rewrite sepit0 | exists hs]. -- by case=>h1[h2][{h}->]; rewrite -sepitE sepit0=>->; rewrite unitL. +rewrite sepit_seq_unlock. +elim: s1 s2=>[|x s1 IH] s2 h /=; split. +- case=>hs [E {h}-> H2]. + exists Unit, (seq_join hs); rewrite unitL. + split=>//; first by rewrite sepit_seq0. + by rewrite sepit_seq_unlock; exists hs. +- case=>h1[h2][{h}->]; rewrite sepit_seq0=>->. + by rewrite sepit_seq_unlock unitL. - case=>/=; case=>[|h0 hs]; case=>//= /eqP; rewrite eqSS=>/eqP E. - rewrite /seqjoin !big_cons /= =>->[H0 HS]. - case: (IH s2 (seqjoin hs)).1; first by exists hs. - move=>h1[h2][HJ H1 H2]; rewrite /seqjoin in HJ. + rewrite !big_cons /= =>->[H0 HS]. + case: (IH s2 (seq_join hs)).1; first by exists hs. + move=>h1[h2][HJ H1 H2]. exists (h0 \+ h1), h2; rewrite HJ joinA; split=>//. - by rewrite -sepitE sepit_cons sepitE; exists h0, h1. -case=>h1[h2][{h}->[]]; case=>[|h0 hs1]; case=>//= /eqP; rewrite eqSS=>/eqP E1. -rewrite /seqjoin !big_cons /= =>{h1}->[H0 H1]; case=>hs2[E2 {h2}-> H2]. -exists (h0 :: hs1 ++ hs2); rewrite /seqjoin big_cons big_cat joinA; split=>//=. + by rewrite sepitseq_cons; exists h0, h1. +case=>h1 [h2][{h} ->]; rewrite sepit_seq_unlock. +case; case=>[|h0 hs1]; case=>//= /eqP; rewrite eqSS=>/eqP E1. +rewrite !big_cons /= =>{h1}->[H0 H1]. +rewrite sepit_seq_unlock; case=>hs2[E2 {h2}-> H2]. +exists (h0 :: hs1 ++ hs2); rewrite big_cons big_cat joinA; split=>//=. - by rewrite !size_cat E1 E2. rewrite zip_cat //=; split=>//. by apply/All_cat. Qed. -Lemma sepit_perm s1 s2 (f : A -> Pred U) : - perm s1 s2 -> sepit s1 f =p sepit s2 f. +Lemma sepitseq_perm s1 s2 (f : A -> Pred U) : + perm s1 s2 -> + sepit_seq s1 f <~> sepit_seq s2 f. Proof. elim: s1 s2 =>[|x s1 IH] s2 /=; first by move/pperm_nil=>->. move=>H; have Hx: x \In s2 by apply/(pperm_in H)/In_cons; left. case: (In_split Hx)=>s21[s22] E; rewrite {s2 Hx}E in H *. move/pperm_cons_cat_cons: H=>/IH {}IH. -rewrite sepit_cons sepit_cat /= =>h; split. -- case=>h1[h2][{h}-> H1]; rewrite IH sepit_cat !sepitE. - case=>_[_][{h2}-> [hs3][E3 -> H3] [hs4][E4 -> H4]]; rewrite joinCA. - exists (seqjoin hs3), (h1 \+ seqjoin hs4); split=>//; first by exists hs3. - rewrite -sepitE sepit_cons sepitE; exists h1, (seqjoin hs4). +rewrite sepitseq_cons sepitseq_cat=>h; rewrite !toPredE; split. +- case=>h1 [h2][->{h} H1]; rewrite IH sepitseq_cat 2!sepit_seq_unlock. + case=>_ [_][->{h2}][hs3][E3 -> H3][hs4][E4 -> H4]; rewrite joinCA. + exists (seq_join hs3), (h1 \+ seq_join hs4); split=>//; first by exists hs3. + rewrite sepitseq_cons sepit_seq_unlock; exists h1, (seq_join hs4). by split=>//; exists hs4. -rewrite sepitE; case=>_[h2][{h}-> [hs1][Hs1 -> H1]]. -rewrite sepit_cons sepitE. -case=>h3[_][{h2}-> H3 [hs2][Hs2 -> H2]]; rewrite joinCA. -exists h3, (seqjoin hs1 \+ seqjoin hs2); split=>//. -rewrite IH sepitE; exists (hs1 ++ hs2); split. +rewrite sepitseq_cons sepit_seq_unlock. +case=>_ [h2][{h}-> [hs1][Hs1 -> H1]]. +rewrite sepit_seq_unlock. +case=>h3 [_][->{h2} H3 [hs2][Hs2 -> H2]]; rewrite joinCA. +exists h3, (seq_join hs1 \+ seq_join hs2); split=>//. +rewrite IH sepit_seq_unlock; exists (hs1 ++ hs2); split. - by rewrite !size_cat Hs1 Hs2. -- by rewrite /seqjoin big_cat. +- by rewrite big_cat. by rewrite zip_cat //; apply/All_cat. Qed. -Lemma sepitF s (f1 f2 : A -> Pred U) : - (forall x, x \In s -> f1 x =p f2 x) -> - sepit s f1 =p sepit s f2. +Lemma sepit_seqF s (f1 f2 : A -> Pred U) : + (forall x, x \In s -> f1 x <~> f2 x) -> + sepit_seq s f1 <~> sepit_seq s f2. Proof. -elim: s=>[|x s IH] H h; first by rewrite !sepit0. +elim: s=>[|x s IH] H h; rewrite !toPredE. +- by rewrite !sepit_seq0. have /IH {IH}H': forall x : A, x \In s -> f1 x =p f2 x. by move=>? H0; apply: H; apply/In_cons; right. have Hx : x \In x :: s by apply/In_cons; left. -by rewrite !sepit_cons; split; case=>h1[h2][{h}-> H1 H2]; exists h1, h2; +by rewrite !sepitseq_cons; split; case=>h1[h2][{h}-> H1 H2]; exists h1, h2; split=>//; [rewrite -H | rewrite -H' | rewrite H | rewrite H']. Qed. -End IterStar. +Lemma big_sepitseq (s : seq A) (f : A -> U) m : + m = \big[join/Unit]_(i <- s) f i <-> + m \In sepit_seq s (fun i h => h = f i). +Proof. +rewrite sepit_seqE; elim: s m=>[|x xs IH] /= m; first by rewrite !big_nil. +rewrite !big_cons; split=>[E|]; last first. +- by case=>h1 [h2][Em H1 H2]; rewrite Em H1; congr (_ \+ _); rewrite IH. +by rewrite InE; exists (f x), (\big[join/Unit]_(j <- xs) f j); rewrite -IH. +Qed. + +Lemma eq_sepitseqF s (f1 f2 : A -> Pred U) : + (forall x, x \In s -> f1 x <~> f2 x) -> + sepit_seq s f1 <~> sepit_seq s f2. +Proof. by move=>H; apply: sepit_seqF=>x Hx; apply: H. Qed. + +Lemma sepitseq_emp (s : seq A) (f : A -> Pred U) : + (forall x, x \In s -> f x <~> emp (U:=U)) -> + sepit_seq s f <~> emp. +Proof. +move=>H; rewrite sepit_seqE. +elim: s H=>[|a xs IH] H; first by rewrite big_nil. +rewrite big_cons H ?InE; last by left. +by rewrite starL IH // => x X; apply: H; rewrite InE; right. +Qed. + +End IterStarSeq. (* iterated star on eqType *) Section IterStarEq. -Variables (U : pcm) (A : eqType). +Context {U : pcm} {A : eqType}. -Lemma sepitP x s (f : A -> Pred U) : +Lemma sepit_seqP x s (f : A -> Pred U) : uniq s -> - sepit s f =p if x \in s - then f x # sepit (filter (predC1 x) s) f - else sepit s f. + sepit_seq s f <~> + if x \in s then f x # sepit_seq (filter (predC1 x) s) f + else sepit_seq s f. Proof. +(* can't use big library for this proof, because all *) +(* necessary lemmas require AC wrt =; thus, direct proof below *) case E: (x \in s)=>//. -elim: s E=>[|y s IH] //= /[swap]; case/andP=>Hy Hu; rewrite sepit_cons inE; case/orP. -- by move/eqP=>->; rewrite eq_refl filter_predC1. +elim: s E=>[|y s IH] //= /[swap]; case/andP=>Hy Hu. +- rewrite sepitseq_cons inE; case/orP. + by move/eqP=>->; rewrite eq_refl filter_predC1. move=>Hx h0. have ->: (y != x) by apply/eqP=>Hxy; rewrite Hxy Hx in Hy. -by split; case=>ha[h1][{h0}-> Ha]; [rewrite (IH Hx Hu) | rewrite sepit_cons]; +by split; case=>ha[h1][{h0}-> Ha]; [rewrite (IH Hx Hu) | rewrite sepitseq_cons]; case=>hb[h][{h1}-> Hb H]; rewrite joinCA; exists hb, (ha \+ h); split=>//; -[rewrite sepit_cons | rewrite (IH Hx Hu)]; exists ha, h. +[rewrite sepitseq_cons | rewrite (IH Hx Hu)]; exists ha, h. Qed. -Lemma eq_sepitF s (f1 f2 : A -> Pred U) : - (forall x, x \in s -> f1 x =p f2 x) -> - sepit s f1 =p sepit s f2. -Proof. by move=>H; apply: sepitF=>x Hx; apply/H/mem_seqP. Qed. - -Corollary perm_sepit s1 s2 (f : A -> Pred U) : - perm_eq s1 s2 -> sepit s1 f =p sepit s2 f. -Proof. by move/perm_eq_perm; exact: sepit_perm. Qed. +Lemma perm_sepitseq s1 s2 (f : A -> Pred U) : + perm_eq s1 s2 -> + sepit_seq s1 f <~> sepit_seq s2 f. +Proof. by move/perm_eq_perm; exact: sepitseq_perm. Qed. End IterStarEq. -End IterStar. - (* iterated star on finsets *) Section FinIterStar. -Variables (U : pcm) (I : finType). +Context {U : pcm} {I : finType}. -Definition sepit (s : {set I}) (Ps : I -> Pred U) := - IterStar.sepit (enum s) Ps. +Definition sepit (s : {set I}) (Ps : I -> Pred U) := + sepit_seq (enum s) Ps. -Lemma sepit0 f : sepit set0 f =p emp. -Proof. -rewrite /sepit (IterStar.perm_sepit (s2 := filter pred0 (enum I))). -- by rewrite filter_pred0 IterStar.sepit0. -apply: uniq_perm; first by exact: enum_uniq. -- by rewrite filter_uniq // enum_uniq. -by move=>x; rewrite !mem_filter /= in_set0. -Qed. +Lemma sepitE (s : {set I}) f : + sepit s f <~> \big[star/emp]_(i in s) f i. +Proof. by rewrite -big_filter /sepit sepit_seqE. Qed. + +Lemma sepit0 f : sepit set0 f <~> emp. +Proof. by rewrite sepitE big_pred0 // => x; rewrite in_set0. Qed. Lemma sepitF (s : {set I}) f1 f2 : - (forall x, x \in s -> f1 x =p f2 x) -> - sepit s f1 =p sepit s f2. -Proof. -move=>H; apply: IterStar.eq_sepitF=>x H1; apply: H. -by rewrite -mem_enum. + (forall x, x \in s -> f1 x <~> f2 x) -> + sepit s f1 <~> sepit s f2. +Proof. +move=>H; apply: eq_sepitseqF=>x /mem_seqP. +by rewrite mem_enum; apply: H. Qed. Lemma eq_sepit (s1 s2 : {set I}) f : s1 =i s2 -> - sepit s1 f =p sepit s2 f. -Proof. by move/eq_enum=>H; apply: IterStar.perm_sepit; rewrite H. Qed. + sepit s1 f <~> sepit s2 f. +Proof. by move/eq_enum=>H; apply: perm_sepitseq; rewrite H. Qed. Lemma sepitS x (s : {set I}) f : - sepit s f =p if x \in s then f x # sepit (s :\ x) f - else sepit s f. + sepit s f <~> if x \in s then f x # sepit (s :\ x) f + else sepit s f. Proof. -case E: (x \in s)=>//. -rewrite (IterStar.sepitP x (s:=enum s) f (enum_uniq s)) mem_enum E. -have Hp: perm_eq [seq y <- enum s | predC1 x y] (enum (s :\ x)). -- rewrite -filter_predI. - apply: uniq_perm=>[||y]; try by rewrite enum_uniq. - by rewrite !mem_filter /= in_setD1. -by move=>h0; split; case=>h1[h2][{h0}-> H1 H]; exists h1, h2; split=>//; -rewrite IterStar.perm_sepit; try by [exact: H]; [rewrite perm_sym |]. +rewrite /sepit {1}(sepit_seqP x f (enum_uniq s)) mem_enum. +case: (x \in s)=>//; rewrite !sepit_seqE !big_filter big_filter_cond. +rewrite (eq_bigl (mem (s :\ x))) //. +by move=>z; rewrite !inE andbC. Qed. -Lemma sepitT1 x f : sepit setT f =p f x # sepit (setT :\ x) f. +Lemma sepitT1 x f : sepit setT f <~> f x # sepit (setT :\ x) f. Proof. by rewrite (sepitS x) in_setT. Qed. Lemma big_sepit (s : {set I}) (f : I -> U) m : m = \big[join/Unit]_(i in s) (f i) <-> m \In sepit s (fun i h => h = f i). -Proof. -rewrite {1}/sepit IterStar.sepitE/IterStar.sepit'/IterStar.seqjoin. -split; last first. -- case=>hs [Es Em] H. - suff {Em}E : hs = map f (enum s) by rewrite Em E big_map big_enum. - elim: hs (enum s) Es H=>[|h hs IH] es; first by move/esym/size0nil=>->. - by case: es=>//= e es [E][H1 H2]; rewrite H1 -IH. -move=>Em; exists (map f (enum s)); split. -- by rewrite size_map. -- by rewrite big_map big_enum. -apply/AllP; case=>a b. -rewrite {1}(_ : enum s = map id (enum s)); last by rewrite map_id. -by rewrite zip_map; case/MapP=>x /= _ [->->]. -Qed. +Proof. by rewrite /sepit /= -big_sepitseq -big_enum. Qed. Lemma big_sepitT (f : I -> U) m : m = \big[join/Unit]_i (f i) <-> diff --git a/pcm/unionmap.v b/pcm/unionmap.v index 9e70998..a86963f 100644 --- a/pcm/unionmap.v +++ b/pcm/unionmap.v @@ -177,7 +177,7 @@ Qed. Definition valid f := if f is Def _ _ then true else false. -Definition empty := @Def (finmap.nil K V) is_true_true. +Definition empty := @Def (@finmap.nil K V) is_true_true. Definition upd k v f := if f is Def fs fpf then @@ -536,7 +536,7 @@ Abort. Lemma xx (x : umap nat nat) : x \+ x == x \+ x. Abort. -(* can i store maps into maps without universe inconsistencies? *) +(* can maps be stored into maps without universe inconsistencies? *) (* yes, the idea of the class works *) Lemma xx : 1 \\-> (1 \\-> 3) = 2 \\-> (7 \\-> 3). Abort. @@ -6773,6 +6773,13 @@ move/(big_find_some (dom_valid (find_some E1)) (or_introl erefl)). by rewrite E1; case=>->; exists x=>//; rewrite InE; left. Qed. +Lemma big_find_someXP (xs : seq I) P a v : + find a (\big[join/Unit]_(i <- xs | P i) f i) = Some v -> + exists i, [/\ i \In xs, P i & find a (f i) = Some v]. +Proof. +by rewrite -big_filter=>/big_find_someX [i] /Mem_filter [H1 H2 H3]; exists i. +Qed. + Lemma bigIn (xs : seq I) a i v : valid (\big[join/Unit]_(i <- xs) f i) -> i \In xs -> @@ -6792,10 +6799,15 @@ Lemma bigInX (xs : seq I) a v : exists2 i, i \In xs & (a, v) \In f i. Proof. by case/In_find/big_find_someX=>x X /In_find; exists x. Qed. +Lemma bigInXP (xs : seq I) P a v : + (a, v) \In \big[join/Unit]_(i <- xs | P i) f i -> + exists i, [/\ i \In xs, P i & (a, v) \In f i]. +Proof. by case/In_find/big_find_someXP=>x [X1 X2 /In_find]; exists x. Qed. + End BigOpsUM. Prenex Implicits big_find_some big_find_someD. -Prenex Implicits big_find_someX bigIn bigInD bigInX. +Prenex Implicits big_find_someX big_find_someXP bigIn bigInD bigInX bigInXP. (* if the index type is eqtype, we can have a somewhat better *) (* big validity lemma *) @@ -6882,6 +6894,24 @@ Proof. by move=>x; rewrite big_valid_dom_seq. Qed. End BigCatSeqUM. +Section OMapBig. +Variables (K : ordType) (C : pred K) (T T' : Type). +Variables (U : union_map C T) (U' : union_map C T'). +Variables (I : Type) (f : I -> U). + +Lemma big_omapVUn (a : K * T -> option T') ts : + omap a (\big[join/Unit]_(x <- ts) f x) = + if valid (\big[join/Unit]_(x <- ts) f x) + then \big[join/Unit]_(x <- ts) omap a (f x) + else undef : U'. +Proof. +elim: ts=>[|t ts IH]; first by rewrite !big_nil omap0 valid_unit. +by rewrite !big_cons omapVUn IH; case: ifP=>// /validR ->. +Qed. + +End OMapBig. + + (* DEVCOMMENT: *) (* remove "tests" for release *) Lemma xx (f : umap nat nat) : 3 \in dom (free f 2). From 77b76dc6828e03f22c5fb44bf9cc8585cd76d081 Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Fri, 16 Aug 2024 12:47:38 +0200 Subject: [PATCH 58/93] minor --- examples/congprog.v | 22 ++++++++++------------ 1 file changed, 10 insertions(+), 12 deletions(-) diff --git a/examples/congprog.v b/examples/congprog.v index 1aa902d..31f0328 100644 --- a/examples/congprog.v +++ b/examples/congprog.v @@ -35,8 +35,6 @@ From htt Require Import kvmaps hashtab congmath. Notation finE := finset.inE. -Prenex Implicits sepit. - (*************) (* Signature *) (*************) @@ -46,10 +44,10 @@ Module Type CongrSig. (* seq const is the list of constants over which *) (* structure computes *) (* this is a global argument, hence exposed by tp *) -Parameter tp : seq constant -> Set. +Parameter root : seq constant -> Set. (* abstract predicate tying roots, (congruence) relation, heap *) -Parameter shape : forall {s}, tp s -> rel_exp s -> Pred heap. +Parameter shape : forall {s}, root s -> rel_exp s -> Pred heap. (* initialize empty congruence structure; return roots *) Parameter init : forall {s}, @@ -72,7 +70,7 @@ End CongrSig. (* Implementation *) (******************) -(* faithful to Barcelogic's actual implementation *) +(* faithful to Barcelogic's implementation *) Module Congr : CongrSig. Section Congr. @@ -95,7 +93,7 @@ Definition hash10 k : 'I_10 := Ordinal (@hash_ord 10 k erefl). (* the tables relating arrays with the content of the lists *) (* ctab is for class lists, utab is for use lists *) -Definition llist (T : Set) : Set := ptr. +Definition llist_root (T : Set) : Set := ptr. Definition ctab := @table symb ptr (seq symb) (@lseq symb). Definition utab := @@ -112,10 +110,10 @@ Inductive ptrs : Set := {array symb -> symb} & (* class list; for each c *) (* singly-linked list storing whole class of c *) - {array symb -> llist symb} & + {array symb -> llist_root symb} & (* use list; internal structure *) (* see the paper for description *) - {array symb -> llist (symb*symb*symb)} & + {array symb -> llist_root (symb*symb*symb)} & (* hash table; for each pair of representatives *) (* stores equations; see paper for description *) htab V hash10 & @@ -123,7 +121,7 @@ Inductive ptrs : Set := ptr. (* renaming type to satisfy the signature check *) -Definition tp := ptrs. +Definition root := ptrs. Section ShapePredicates. Variable rt : ptrs. @@ -156,7 +154,7 @@ End ShapePredicates. (* Initialization procedure to generate *) (* the empty structure and its root pointers *) -Definition iT (clist : {array symb -> llist symb}): Type := forall k, +Definition iT (clist : {array symb -> llist_root symb}): Type := forall k, STsep (fun i => k <= n /\ exists f, i \In Array.shape clist f # sepit [set c | indx c < k] (ctab f [ffun c => [:: c]]), [vfun (_ : unit) m => exists f, m \In Array.shape clist f # @@ -183,7 +181,7 @@ case: decP=>[{}pf|] /=; last first. - case: (ltngtP k n) pf=>// Ekn _ _; step=>_. exists f, hc, hct; split=>//. apply: tableP2 Hct=>// x; rewrite !finE Ekn. - by rewrite /n cardE index_mem /index mem_enum. + by rewrite /n cardE index_mem /index mem_enum. step=>x; step; apply: [stepX f]@hc=>//= [[m]] Em. apply: [gE]=>//=; split=>//. eexists [ffun z => if z == ith k pf then x else f z]. @@ -202,7 +200,7 @@ apply: [stepX]@hc=>//=. by split=>//; rewrite (_ : [set c | indx c < 0] = set0) // sepit0. case=>_ [f][hc'][hrest][-> Hc' Hrest]. apply: [stepU]=>//= ul hu Ehu; apply: [stepU]=>//= htb ht Ht. -set d := Data [ffun x => x] [ffun c => [:: c]] [ffun c => [::]] (@nil K V) [::]. +set d := Data [ffun x => x] [ffun c => [:: c]] [ffun c => [::]] (@nil K V) [::]. step=>px; step; exists d; split; last by case: (initP s). split=>[a b|/=]; first by rewrite !ffunE !inE. exists f, [ffun s => null]. From 6b9d585ed50404eb8cabf970bb8a0d50d72c4b85 Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Fri, 16 Aug 2024 13:29:34 +0200 Subject: [PATCH 59/93] sectioning of congmath.v --- examples/congmath.v | 162 ++++++++++++++++++++++++++++---------------- 1 file changed, 105 insertions(+), 57 deletions(-) diff --git a/examples/congmath.v b/examples/congmath.v index a0c8f61..c267c07 100644 --- a/examples/congmath.v +++ b/examples/congmath.v @@ -145,6 +145,8 @@ Qed. HB.instance Definition _ s := hasDecEq.Build (@exp s) (@eqexpP s). +Notation rel_exp s := (Pred (exp s * exp s)). + (****************************************) (* Congruence relations over expression *) (****************************************) @@ -154,8 +156,7 @@ Section Congruence. Variable s : seq constant. Notation symb := (symb s). Notation exp := (exp s). - -Notation rel_exp := (Pred (exp*exp)). +Notation rel_exp := (rel_exp s). Definition Reflexive (r : rel_exp) := forall x, (x, x) \In r. @@ -167,7 +168,6 @@ Definition Equivalence (r : rel_exp) := Reflexive r /\ Symmetric r /\ Transitive r. Definition Antisymmetric (r : rel_exp) := forall x y, (x, y) \In r -> (y, x) \In r -> x = y. - Definition monotone (R : rel_exp) : Prop := forall f1 f2 e1 e2, (f1, f2) \In R -> (e1, e2) \In R -> (app f1 e1, app f2 e2) \In R. @@ -228,6 +228,7 @@ move=>{}R1 {}R2 H [x y] H1 C H2 H3; apply: H1 H3. by rewrite H. Qed. +(* to be repeated outside of the section *) Add Morphism closure with signature Morphisms.respectful (fun e1 e2 => e1 <~> e2) (fun e1 e2 => e1 <~> e2) as closure_morph. @@ -315,18 +316,42 @@ Qed. Definition graph (f : symb -> exp) : rel_exp := [Pred e | e \in image (fun s => (const s, f s)) predT]. +End Congruence. + +Hint Resolve reflC symC reflC symC : core. + +(* repeat the morphism declaration outside the section *) +Add Parametric Morphism s : (@closure s) with + signature Morphisms.respectful (fun e1 e2 => e1 <~> e2) + (fun e1 e2 => e1 <~> e2) as closure_morph'. +Proof. by apply: closE. Qed. + +(********************************************************************) +(* The definition of the data structures involved in the algorithm. *) +(********************************************************************) + +Section Structures. +Variable s : seq constant. +Notation symb := (symb s). +Notation exp := (exp s). +Notation rel_exp := (rel_exp s). + (* Equations in canonical form, as required by the congruence closure *) (* algorithm. An equation is in canonical form if it is an equation *) (* between two constants, or between a constant and an application of two *) (* constants. *) -Inductive Eq : Type := simp_eq of symb & symb | comp_eq of symb & symb & symb. +Inductive Eq : Type := + simp_eq of symb & symb | + comp_eq of symb & symb & symb. (* equations are an equality type *) Definition eqEq (eq1 eq2 : Eq) : bool := match eq1, eq2 with - simp_eq c1 c2, simp_eq d1 d2 => (c1 == d1) && (c2 == d2) - | comp_eq c c1 c2, comp_eq d d1 d2 => (c == d) && (c1 == d1) && (c2 == d2) + simp_eq c1 c2, simp_eq d1 d2 => + (c1 == d1) && (c2 == d2) + | comp_eq c c1 c2, comp_eq d d1 d2 => + (c == d) && (c1 == d1) && (c2 == d2) | _, _ => false end. @@ -410,10 +435,6 @@ Definition pend2eq (p : pend) := | comp_pend (c, c1, c2) (d, d1, d2) => simp_eq c d end. -(********************************************************************) -(* The definition of the data structures involved in the algorithm. *) -(********************************************************************) - Record data : Type := Data {(* An array, indexed by symbs, containing for each symb its *) (* current representative (for the relation obtained by the *) @@ -432,8 +453,8 @@ Record data : Type := (* The list of equations pending to be processed *) pending : seq pend}. -(* We can collect the representatives of all the symbols into a finite list. *) -(* We remove the duplicates. *) +(* Collect the representatives of all the symbols *) +(* into a finite list. Remove the duplicates. *) Definition reps D : seq symb := undup (map (rep D) (enum predT)). Lemma uniq_reps D : uniq (reps D). @@ -490,10 +511,22 @@ Hint Resolve cong_rel : core. Lemma clos_rel D R a : (const a, const (rep D a)) \In closure (CRel D \+p R). Proof. by rewrite /CRel clos_clos orrA; apply: clos_rep. Qed. +End Structures. + +Hint Resolve uniq_reps rep_in_reps clos_rep cong_rel : core. + +(************************************) +(* Mathematical (purely-functional) *) +(* model of the algorithm. *) +(* Also, proof of termination. *) +(************************************) -(*******************************************) -(* Congruence Closure Code and Termination *) -(*******************************************) +Section MathModel. +Variable s : seq constant. +Notation symb := (symb s). +Notation exp := (exp s). +Notation rel_exp := (rel_exp s). +Notation data := (data s). (* termination metric is the number of equations in the use and pending lists *) Definition metric (D : data) : nat := @@ -566,8 +599,8 @@ Lemma join_class_metric (D : data) (a' b' : symb) : Proof. move=>H1 H2 H3; rewrite /metric /= -addnA; congr plus. by rewrite (perm_big _ (join_class_perm H3 H2)) - -(perm_big _ (permEl (perm_filterC (pred1 a') (reps D)))) - big_cat filter_pred1_uniq // !big_filter big_cons /= big_nil addn0 addnC. + -(perm_big _ (permEl (perm_filterC (pred1 a') (reps D)))) + big_cat filter_pred1_uniq // !big_filter big_cons /= big_nil addn0 addnC. Qed. (* joining the use lists *) @@ -633,20 +666,19 @@ case S2: (x == a')=>//=. by rewrite ffunE S2. Qed. -Let pend0 (e : pend) := - match e with simp_pend a b => a | comp_pend (a,_,_) (b,_,_) => a end. -Let pend1 (e : pend) := - match e with simp_pend a b => b | comp_pend (a,_,_) (b,_,_) => b end. -Notation "e ..0" := (pend0 e) (at level 2). -Notation "e ..1" := (pend1 e) (at level 2). +(* left/right symbol in pending equation *) +Definition pendL (e : pend s) := + let: (simp_pend a _ | comp_pend (a,_,_) _) := e in a. +Definition pendR (e : pend s) := + let: (simp_pend _ b | comp_pend _ (b,_,_)) := e in b. (* loop through the equations in the pending list *) Function propagate (D : data) {measure metric D} : data := match (pending D) with | [::] => D | e :: p' => - let: a' := rep D (e..0) in - let: b' := rep D (e..1) in + let: a' := rep D (pendL e) in + let: b' := rep D (pendR e) in let: D' := Data (rep D) (class D) (use D) (lookup D) p' in if a' == b' then propagate D' else @@ -660,7 +692,7 @@ Proof. by rewrite /metric eq addSn; apply: ltP; apply: ltnSn. move=>D e p' eq H. rewrite join_use_metric ?H //; last first. -- by rewrite join_class_eq ?mem_filter ?rep_in_reps //= ?(eq_sym (rep D e..1)) H. +- by rewrite join_class_eq ?mem_filter ?rep_in_reps //= ?(eq_sym (rep D (pendR e))) H. - by rewrite join_class_eq ?mem_filter ?rep_in_reps ?H //= eq_refl. rewrite -join_class_metric ?H ?rep_in_reps //. by apply: ltP; rewrite /metric /reps /= eq /= addSn. @@ -680,9 +712,20 @@ Fixpoint norm (D : data) (t : exp) {struct t} : exp := end end. -(************************************) -(* Some invariants of the algorithm *) -(************************************) +End MathModel. + +(*******************************) +(* Invariants of the algorithm *) +(*******************************) + +Section Invariants. +Variable s : seq constant. +Notation symb := (symb s). +Notation exp := (exp s). +Notation rel_exp := (rel_exp s). +Notation data := (data s). +Notation pend2eq := (@pend2eq s). +Implicit Type D : data. (* the rep function is idempotent *) Definition rep_idemp D := forall a, rep D (rep D a) = rep D a. @@ -823,13 +866,28 @@ Definition propagate_inv D := rep_idemp D /\ use_inv D /\ lookup_inv D /\ use_lookup_inv D /\ lookup_use_inv D. -(****************) -(* Verification *) -(****************) +End Invariants. + +(**********************************) +(* Verification of the math model *) +(**********************************) + +(* proving properties of the mathematal model *) + +Section MathModelProofs. +Variable s : seq constant. +Notation symb := (symb s). +Notation exp := (exp s). +Notation rel_exp := (rel_exp s). +Notation data := (data s). +Notation pend2eq := (@pend2eq s). +Notation symb2eq := (@symb2eq s). +Notation closure := (@closure s). +Implicit Type D : data. (* first some basic rewrite rules *) -Lemma reps_rep (D : data) (a : symb) : +Lemma reps_rep D (a : symb) : rep_idemp D -> a \in reps D -> rep D a = a. Proof. by move=>H; rewrite mem_undup; case/mapP=>x _ ->; apply: H. Qed. @@ -906,7 +964,7 @@ move=>H2 H3 H1 [x y]; split=>/= T. case/imageP=>z _ /= [->->] {x y}; rewrite ffunE /=. case: eqP=>[<-|_]; last by apply: clos_rep. apply: (transC (y:=const (rep D z))); first by apply: clos_rep. - by apply: (@closI _ (const (rep D z), const b')); right. + by apply: closI; right. move: (clos_idemp (rep2rel (join_class D a' b')) (x,y))=>/=<-. apply: sub_closI T=>{x y} [[x y]]. case; last first. @@ -914,12 +972,12 @@ case; last first. by rewrite /= !ffunE /= reps_rep // eq_refl. case/imageP=>z _ /= [->->] {x y}. case E: (rep D z == a'); last first. -- apply: (@closI _ (const z, const (rep D z))); apply/imageP. +- apply: closI; apply/imageP. by exists z; rewrite // ffunE /= E. apply: (transC (y:=const b')). -- apply: (@closI _ (const z, const b')); apply/imageP. +- apply: closI; apply/imageP. by exists z; rewrite // ffunE /= E. -apply: symC; apply: (@closI _ (const (rep D z), const b')); apply/imageP. +apply: symC; apply: closI; apply/imageP. by exists (rep D z); rewrite // ffunE /= H1 // E. Qed. @@ -1072,9 +1130,7 @@ move=>L1 L2 L3 H1 H4 H5 [x y]; rewrite !toPredE; split. move: Q4; apply: sub_closI=>{x y} [[x y]] Q4; apply: sub_orl. by rewrite orrA. rewrite const_rel /CRel clos_clos !orrA symR. - apply: (@closI _ (app (const (rep D c1)) - (const (rep D c2)), const (rep D d))). - apply: sub_orr; apply: sub_orl. + apply: closI; apply: sub_orr; apply: sub_orl. rewrite invert_look. exists (rep D d1), (rep D d2), d, d1, d2. by rewrite -Q2 -Q3 !rep_in_reps. @@ -1581,8 +1637,8 @@ have L2: forall D pend_eq p' a b a' b' D' D'', by case=>[_][_][[d][d1][d2]][]; rewrite U. rewrite -H6 /= join_classE // T H6=>->. by apply: IH. -apply/(@propagate_ind (fun d d' => inv d -> inv d' /\ - pending d' = [::] /\ CRel d <~> CRel d')); first by []. +apply/(@propagate_ind s (fun d d' => inv d -> inv d' /\ + pending d' = [::] /\ CRel d <~> CRel d'))=>//. - by move=>D e p' H1 a' H2 b' H3 D'; apply: L1 H1 _ H2 H3; case: e=>// [[[c c1]] c2] [[d d1] d2]. move=>D e p' H1 a' H2 b' H3 D' E [] // H _ D''; @@ -1593,8 +1649,8 @@ Qed. (* Lemmas about interaction of propagate with pending and closure *) Lemma propagate_pendP d eq : propagate_inv d -> - propagate_inv (Data (rep d) (class d) (use d) - (lookup d) (eq :: pending d)). + @propagate_inv s (Data (rep d) (class d) (use d) + (lookup d) (eq :: pending d)). Proof. move=>H; set d' := Data _ _ _ _ _. have L: forall a b, similar d a b -> similar d' a b. @@ -1618,7 +1674,7 @@ Lemma propagate_clos_pendP d c c1 c2 e e1 e2 : Proof. move=>PI H. have [R1 R2]: rep d e1 = rep d c1 /\ rep d e2 = rep d c2. -- by move: PI=>[_][_][L1] _; apply: (L1 _ _ e). +- by move: PI=>[_][_][L1] _; apply: L1 H; apply: rep_in_reps. rewrite /CRel /= /eq2rel /=. rewrite clos_clos -!(rpull (eqs2rel _)) !orrA; apply: closKR. move=>[z y]; split=>O; rewrite toPredE -clos_idemp; move: O; apply: sub_closI; @@ -1661,13 +1717,13 @@ move: PI=>[R1][U1][L1][UL1] LU1; do 2!split=>//. - move=>a e e1 e2; rewrite /reps /= !ffunE /= => T1. case: ifP=>E1. - rewrite E (eqP E1) inE. - by case/orP=>[|O]; [case/eqP=>_->->; right | apply: U1 O]. + by case/orP=>[/eqP [_ ->->]|O]; [right|apply/U1/O/rep_in_reps]. case: ifP=>E2; last by apply: U1 T1. rewrite (eqP E2) inE. - by case/orP=>[|O]; [case/eqP=>_->->; left | apply: U1 O]. + by case/orP=>[/eqP [_ ->->]|O]; [left|apply/U1/O/rep_in_reps]. split=>[a b e e1 e2|]. - rewrite /reps /= fnd_ins. - by case: ifP=>E1; [case/eqP: E1=>->-> T1 T2 [_->->] | apply: L1]. + by case: ifP=>E1; [case/eqP: E1=>->-> T1 T2 [_->->]|apply: L1]. split=>[a e e1 e2|]. - rewrite /reps /= !ffunE /= => T1. case: ifP=>E1. @@ -1782,15 +1838,7 @@ move=>[a][b][c][c1][c2][/= -> Q1 Q2 Q3 ->] /=; do 2!rewrite reps_rep //. by rewrite Q3 H1. Qed. -End Congruence. - -Notation rel_exp s := (Pred (exp s * exp s)). - -(* repeat the morphism declaration outside the section *) -Add Parametric Morphism s : (@closure s) with - signature Morphisms.respectful (fun e1 e2 => e1 <~> e2) - (fun e1 e2 => e1 <~> e2) as closure_morph'. -Proof. by apply: closE. Qed. +End MathModelProofs. (* empty congruence only relates constant symbols to themselves *) Definition empty_cong s := closure (graph (@const s)). From 50909d9bec0610853c155455437556cf5a3b32de Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Fri, 16 Aug 2024 13:36:09 +0200 Subject: [PATCH 60/93] minor --- _CoqProject | 2 +- examples/congmath.v | 2 +- 2 files changed, 2 insertions(+), 2 deletions(-) diff --git a/_CoqProject b/_CoqProject index c488066..908b8d5 100644 --- a/_CoqProject +++ b/_CoqProject @@ -51,6 +51,6 @@ examples/bubblesort.v examples/quicksort.v examples/congmath.v examples/congprog.v -examples/tree.v +examples/tree.v examples/union_find.v examples/graph.v diff --git a/examples/congmath.v b/examples/congmath.v index c267c07..fa4c2d3 100644 --- a/examples/congmath.v +++ b/examples/congmath.v @@ -872,7 +872,7 @@ End Invariants. (* Verification of the math model *) (**********************************) -(* proving properties of the mathematal model *) +(* proving properties of the mathematical model *) Section MathModelProofs. Variable s : seq constant. From c7725aab110e46f295a471761166d74dc51700a0 Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Fri, 16 Aug 2024 13:51:54 +0200 Subject: [PATCH 61/93] modifying comments --- examples/kvmaps.v | 18 ++++++++++-------- 1 file changed, 10 insertions(+), 8 deletions(-) diff --git a/examples/kvmaps.v b/examples/kvmaps.v index 22a6516..59a3c38 100644 --- a/examples/kvmaps.v +++ b/examples/kvmaps.v @@ -25,10 +25,10 @@ From htt Require Import options model heapauto. (* Dynamic KV map is determined by its root pointer(s). *) (* Functions such insert and remove may modify *) (* the root, and will correspondingly return the new one. *) -(* Tp is abstracted to facilitate structures that may *) +(* root is abstracted to facilitate structures that may *) (* have more than one root pointers. Also, it enables *) (* passing K and V to methods thus reducing annotations *) -(* There's no deep reason to make tp : Set, except that *) +(* There's no deep reason to make root : Set, except that *) (* it should be thought of a collection of pointers, hence small. *) HB.mixin Record isDKVM (K : ordType) (V : Type) (root : Set) := { @@ -38,9 +38,9 @@ HB.mixin Record isDKVM (K : ordType) (V : Type) (root : Set) := { dkvm_free : forall x, STsep {s} (dkvm_shape x s, [vfun _ : unit => emp]); dkvm_insert : forall x k v, - STsep {s} (dkvm_shape x s,[vfun y => dkvm_shape y (ins k v s)]); + STsep {s} (dkvm_shape x s, [vfun y => dkvm_shape y (ins k v s)]); dkvm_remove : forall x k, - STsep {s} (dkvm_shape x s,[vfun y => dkvm_shape y (rem k s)]); + STsep {s} (dkvm_shape x s, [vfun y => dkvm_shape y (rem k s)]); dkvm_lookup : forall x k, STsep {s} (dkvm_shape x s, [vfun y m => m \In dkvm_shape x s /\ y = fnd k s])}. @@ -53,9 +53,9 @@ Arguments dkvm_new {K V}. (* Static KV map (or just KV map) are those that *) (* don't need to modify the root pointer. *) (* Typical example will be hash tables *) -(* Another example is a structure obtained by packaing *) +(* Another example is a structure obtained by packaging *) (* the root pointer along with the dynamic KV map that the *) -(* root points to. *) +(* root points to. *) HB.mixin Record isKVM (K : ordType) (V : Type) (root : Set) := { kvm_null : root; @@ -64,9 +64,9 @@ HB.mixin Record isKVM (K : ordType) (V : Type) (root : Set) := { kvm_free : forall x, STsep {s} (kvm_shape x s, [vfun _ : unit => emp]); kvm_insert : forall x k v, - STsep {s} (kvm_shape x s,[vfun _ : unit => kvm_shape x (ins k v s)]); + STsep {s} (kvm_shape x s, [vfun _ : unit => kvm_shape x (ins k v s)]); kvm_remove : forall x k, - STsep {s} (kvm_shape x s,[vfun _ : unit => kvm_shape x (rem k s)]); + STsep {s} (kvm_shape x s, [vfun _ : unit => kvm_shape x (rem k s)]); kvm_lookup : forall x k, STsep {s} (kvm_shape x s, [vfun y m => m \In kvm_shape x s /\ y = fnd k s])}. @@ -80,6 +80,7 @@ Arguments kvm_new {K V}. (* Association lists with DKVM interface *) (*****************************************) +(* module type for making the defs opaque *) Module Type DAList_sig. Parameter root : ordType -> Type -> Set. Section DAlist. @@ -635,6 +636,7 @@ HB.instance Definition _ K V := (* static variant packages the root pointer *) (* with the dynamic part of the structure *) +(* module type for making the defs opaque *) Module Type AList_sig. Parameter root : ordType -> Type -> Set. Section AList. From fad9a80259c18919934703287fe32ff2499a9d2d Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Tue, 27 Aug 2024 17:34:48 +0200 Subject: [PATCH 62/93] changed natmap to add view for last_val --- pcm/natmap.v | 16 ++++++++++++++++ 1 file changed, 16 insertions(+) diff --git a/pcm/natmap.v b/pcm/natmap.v index 41e582a..9fb2142 100644 --- a/pcm/natmap.v +++ b/pcm/natmap.v @@ -334,6 +334,22 @@ Lemma lastkey_memk A (U : natmap A) (h : U) k : last_key h <= k. Proof. by case: lastkeyP=>//= v _; case: ltngtP=>// _; apply. Qed. +(* view for last_val can simplify obligations from last_keyP *) +(* by joining the cases for unit and undef *) +Variant lastval_spec A (U : natmap A) (h : U) : option A -> Type := +| lastval_none of last_val h = None & (forall kv, kv \Notin h) : + lastval_spec h None +| lastval_some r kv of last_val h = Some r & kv \In h : + lastval_spec h (Some r). + +Lemma lastvalP A (U : natmap A) (h : U) : lastval_spec h (last_val h). +Proof. +case: lastkeyP=>[->|->|v]. +- by apply: lastval_none; [rewrite lastval_undef|move=>kv /In_undef]. +- by apply: lastval_none; [rewrite lastval0|move=>kv /In0]. +by move=>H; apply/lastval_some/H/In_findE. +Qed. + (************) (* last_key *) (************) From 4fd5dc973a0a14ece2c9a9c93d9a93aa9d1ad09c Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Tue, 27 Aug 2024 20:17:57 +0200 Subject: [PATCH 63/93] hm --- pcm/unionmap.v | 7 ++++++- 1 file changed, 6 insertions(+), 1 deletion(-) diff --git a/pcm/unionmap.v b/pcm/unionmap.v index a86963f..e285b57 100644 --- a/pcm/unionmap.v +++ b/pcm/unionmap.v @@ -2595,6 +2595,9 @@ case=>k'; rewrite InPtUnE //; case; first by case; left. by move=>H; right; apply/IH; exists k'. Qed. +Lemma In_rangeEX k v f : (k, v) \In f -> v \In range f. +Proof. by move=>H; apply/In_rangeX; 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. @@ -2703,6 +2706,9 @@ 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 mem_rangeEX k v f : (k, v) \In f -> v \in range f. +Proof. by move=>H; apply/mem_rangeX; exists k. Qed. + Lemma range_valid k f : k \in range f -> valid f. Proof. by move/mem_seqP/In_range_valid. Qed. @@ -6911,7 +6917,6 @@ Qed. End OMapBig. - (* DEVCOMMENT: *) (* remove "tests" for release *) Lemma xx (f : umap nat nat) : 3 \in dom (free f 2). From 4984d6d9eaadbfdbf700ebf1b862a651fefe0f93 Mon Sep 17 00:00:00 2001 From: Aleks Nanevski Date: Tue, 27 Aug 2024 21:46:17 +0200 Subject: [PATCH 64/93] removed extraneous lemmas from unionmap.v --- pcm/unionmap.v | 6 ------ 1 file changed, 6 deletions(-) diff --git a/pcm/unionmap.v b/pcm/unionmap.v index e285b57..06e2c64 100644 --- a/pcm/unionmap.v +++ b/pcm/unionmap.v @@ -2595,9 +2595,6 @@ case=>k'; rewrite InPtUnE //; case; first by case; left. by move=>H; right; apply/IH; exists k'. Qed. -Lemma In_rangeEX k v f : (k, v) \In f -> v \In range f. -Proof. by move=>H; apply/In_rangeX; 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. @@ -2706,9 +2703,6 @@ 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 mem_rangeEX k v f : (k, v) \In f -> v \in range f. -Proof. by move=>H; apply/mem_rangeX; exists k. Qed. - Lemma range_valid k f : k \in range f -> valid f. Proof. by move/mem_seqP/In_range_valid. Qed. From 74ad918d03fb28d8ac57d45eef1bb5658c0b0d31 Mon Sep 17 00:00:00 2001 From: Aleks Nanevski Date: Tue, 27 Aug 2024 21:53:34 +0200 Subject: [PATCH 65/93] minor --- pcm/unionmap.v | 1 + 1 file changed, 1 insertion(+) diff --git a/pcm/unionmap.v b/pcm/unionmap.v index 06e2c64..a86963f 100644 --- a/pcm/unionmap.v +++ b/pcm/unionmap.v @@ -6911,6 +6911,7 @@ Qed. End OMapBig. + (* DEVCOMMENT: *) (* remove "tests" for release *) Lemma xx (f : umap nat nat) : 3 \in dom (free f 2). From 7ac8e43d5686630ebaa488158a597fd68eb53f70 Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Wed, 28 Aug 2024 17:58:20 +0200 Subject: [PATCH 66/93] changed In_dom_umfilt lemma to use exists2 instead of exists. --- pcm/natmap.v | 119 ++++++++++++++++++++++++++++--------------------- pcm/unionmap.v | 49 ++++++++++---------- 2 files changed, 95 insertions(+), 73 deletions(-) diff --git a/pcm/natmap.v b/pcm/natmap.v index 9fb2142..d1b0072 100644 --- a/pcm/natmap.v +++ b/pcm/natmap.v @@ -2282,27 +2282,29 @@ 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 - (fun kv : (nat * V) => - let '(k, v) := kv in t1 < k <= t2). + t1 <= t2 -> + predD (le t2) (le t1) =1 + (fun kv : (nat * V) => + let '(k, v) := kv in 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 - (fun kv : nat * V => - let '(k, v) := kv in t1 < k < t2). + t1 < t2 -> + predD (lt t2) (le t1) =1 + (fun kv : nat * V => + let '(k, v) := kv in t1 < k < t2). Proof. by move=>T [k v] /=; rewrite -ltnNge. Qed. Lemma lastkey_umfilt_leq A (U : natmap A) (h : U) t : last_key (um_filterk (leq^~ t) h) <= t. Proof. -case V : (valid h); last first. -- by move/negbT/invalidE: V=>->; rewrite pfundef lastkey_undef. +case: (normalP h)=>[->|V]. +- by rewrite pfundef lastkey_undef. set j := um_filterk _ _. have V' : valid j by rewrite pfV. case E : (unitb j); first by move/unitbP: E=>->; rewrite lastkey0. have : last_key j \in dom j by case: lastkeyP V' E. -by case/In_dom_umfilt=>v []. +by case/In_dom_umfilt. Qed. Lemma lastkey_umfilt_dom A (U : natmap A) (h : U) t : @@ -2310,10 +2312,10 @@ Lemma lastkey_umfilt_dom A (U : natmap A) (h : U) t : Proof. by apply: lastkey_mono; apply: omf_subdom. Qed. Lemma umfilt_le_last A (U : natmap A) (h : U) t : - last_key h <= t -> um_filter (le t) h = h. + last_key h <= t -> + um_filter (le t) h = h. Proof. -case V : (valid h); last first. -- by move/invalidE: (negbT V)=>->; rewrite pfundef. +case: (normalP h)=>[->|V]; first by rewrite pfundef. move=>N; rewrite -{2}[h]umfilt_predT; apply/eq_in_umfilt. by case=>k v /In_dom/dom_lastkey/leq_trans; apply. Qed. @@ -2327,7 +2329,9 @@ Lemma umfilt_le_fresh A (U : natmap A) (h : U) : Proof. by apply: umfilt_le_last; apply: ltnW. Qed. Lemma umfilt_le0 A (U : natmap A) (h : U) t : - valid h -> {in dom h, forall k, t < k} -> um_filter (le t) h = Unit. + 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. @@ -2373,7 +2377,8 @@ by rewrite subX //; case: (ltngtP k1 k.-1). Qed. Lemma umfilt_leUn A (U : natmap A) (h1 h2 : U) t : - valid (h1 \+ h2) -> t <= last_key h1 -> + 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. @@ -2382,7 +2387,10 @@ by move=>k /D /(leq_ltn_trans K). Qed. Lemma umfilt_le_gapless A (U : natmap A) (h1 h2 : U) t : - gapless h1 -> valid h2 -> [pcm h1 <= h2] -> t <= last_key h1 -> + 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]. @@ -2390,7 +2398,8 @@ by subst h2; rewrite umfilt_leUn. Qed. Lemma dom_umfilt_le_eq A (U : natmap A) (h : U) t1 t2 : - t1 \in 0::dom h -> t2 \in 0::dom h -> + t1 \in 0::dom h -> + t2 \in 0::dom h -> um_filter (le t1) h = um_filter (le t2) h -> t1 = t2. Proof. @@ -2407,20 +2416,23 @@ by rewrite !leqnn; case: ltngtP. Qed. Lemma eval_leUn A (U : natmap A) R a (h1 h2 : U) t (z0 : R) : - valid (h1 \+ h2) -> t <= last_key h1 -> + 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 (U : natmap A) R a (h1 h2 : U) t (z0 : R) : - gapless h1 -> valid h2 -> [pcm h1 <= h2] -> t <= last_key h1 -> + 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 (U : natmap A) R a (h : U) (z0 : R) : eval a (le 0) h z0 = z0. Proof. -case W : (valid h); last first. -- by move/negbT/invalidE: W=>->; rewrite eval_undef. +case: (normalP h)=>[->|V]; first by rewrite eval_undef. rewrite eval_umfilt umfilt_mem0L ?eval0 //. by move=>k v /In_dom/dom_cond; rewrite /=; case: ltngtP. Qed. @@ -2430,10 +2442,9 @@ Lemma eval_le_split A (U : natmap A) R a (h : U) t1 t2 (z0 : R) : eval a (le t2) h z0 = eval a (fun '(k, _)=>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. +move=>T; case: (normalP h)=>[->|V]; first by rewrite !eval_undef. rewrite eval_umfilt (umfilt_predD h (pts_sub T)) evalUn; last first. -- move=>x y /In_dom_umfilt [vx][X _] /In_dom_umfilt [wy][/= /andP][]. +- move=>x y /In_dom_umfilt [vx X _] /In_dom_umfilt [wy /= /andP][]. by rewrite /= -ltnNge; move/(leq_ltn_trans X). - by rewrite -(umfilt_predD h (pts_sub T)) pfV. by rewrite -!eval_umfilt; apply: eq_in_eval; case=>k v _; apply: ptsD. @@ -2444,10 +2455,9 @@ Lemma eval_lt_split A (U : natmap A) R a (h : U) t1 t2 (z0 : R) : eval a (lt t2) h z0 = eval a (fun '(k, _)=>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. +move=>T; case: (normalP h)=>[->|V]; first by rewrite !eval_undef. rewrite eval_umfilt (umfilt_predD h (pts_sub_lt T)) evalUn; last first. -- move=>x y /In_dom_umfilt [vx][X _] /In_dom_umfilt [wy][/= /andP][]. +- move=>x y /In_dom_umfilt [vx X _] /In_dom_umfilt [wy /= /andP][]. by rewrite /= -ltnNge; move/(leq_ltn_trans X). - by rewrite -(umfilt_predD h (pts_sub_lt T)) pfV. by rewrite -!eval_umfilt; apply: eq_in_eval; case=>k v _; apply: ptsD_lt. @@ -2457,11 +2467,10 @@ Lemma eval_le_lt_split A (U : natmap A) R a (h : U) t (z0 : R) : eval a (le t) h z0 = eval a (fun '(k, _)=>k == t) h (eval a (lt t) h z0). Proof. -case V : (valid h); last first. -- by move/invalidE: (negbT V)=>->; rewrite !eval_undef. +case: (normalP h)=>[->|V]; first by 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 /In_dom_umfilt [vx][X _] /In_dom_umfilt [wy][/= /andP][]. +- move=>x y /In_dom_umfilt [vx X _] /In_dom_umfilt [wy /= /andP][]. by rewrite /= -ltnNge; move/(leq_ltn_trans X). - by rewrite -(umfilt_predD h D) pfV. rewrite -!eval_umfilt; apply: eq_in_eval; case=>k v _ /=. @@ -2469,7 +2478,8 @@ by case: ltngtP. Qed. Lemma eval_eq A (U : natmap A) R a (h : U) t v (z0 : R) : - (t, v) \In h -> eval a (fun '(k, _)=>k == t) h z0 = a z0 t v. + (t, v) \In h -> + eval a (fun '(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 (fun '(k, _)=>k == t) h = pts t v by rewrite evalPt /= N. @@ -2480,7 +2490,8 @@ by move=>[/eqP -> H1]; rewrite (In_fun H H1); apply: In_condPt. Qed. Lemma eval_le_last A (U : natmap A) R a (h : U) t (z0 : R) : - last_key h <= t -> eval a (le t) h z0 = eval a xpredT h z0. + 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/dom_lastkey/leq_trans; apply. Qed. @@ -2490,8 +2501,8 @@ Lemma eval_fresh_le A (U : natmap A) R a (h : U) t v (z0 : R) : 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: (normalP h)=>[->|V]. +- by rewrite join_undef !eval_undef; case: ifP. case: ifP=>H. - by rewrite eval_umfilt umfiltPtUn freshPtUnV // V ltnNge H -eval_umfilt. rewrite joinC evalUnPt; last first. @@ -2507,8 +2518,8 @@ Lemma eval_fresh_lt A (U : natmap A) R a (h : U) t v (z0 : R) : 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: (normalP h)=>[->|V]. +- by 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. @@ -2580,24 +2591,30 @@ by rewrite (umcnt_predD _ (pts_sub T)) (eq_in_umcnt2 _ (ptsD T)). Qed. Lemma umcnt_le0 A (U : natmap A) p (h : U) t : - valid h -> {in dom h, forall k, t < k} -> + 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 (U : natmap A) p (h1 h2 : U) t : - valid (h1 \+ h2) -> t <= last_key h1 -> + 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 (U : natmap A) p (h1 h2 : U) t : - gapless h1 -> valid h2 -> [pcm h1 <= h2] -> t <= last_key h1 -> + 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 (U : natmap A) p (h : U) t : - last_key h <= t -> um_count (predI p (le t)) h = um_count p h. + 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 (U : natmap A) p (h : U) t v : @@ -2605,9 +2622,8 @@ Lemma umcnt_fresh_le A (U : natmap A) p (h : U) t v : 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: (normalP h)=>[->|V]. +- by rewrite join_undef !umcnt_undef lastkey_undef andbF; case: ifP. case: ifP=>H. - by rewrite -!umcnt_umfilt umfiltPtUn valid_fresh // V ltnNge H. rewrite umcntPtUn ?valid_fresh //= ltnNge H /=. @@ -2719,7 +2735,7 @@ move=>G N; case W: (valid h); last first. 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_umfiltX [_ /G]. - by rewrite -(umfilt_le_split h N) pfV. -by move=>??/In_dom_umfilt[?][/leq_ltn_trans Y _]/In_dom_umfilt[?][/andP[/Y]]. +by move=>??/In_dom_umfilt[? /leq_ltn_trans Y _]/In_dom_umfilt[? /andP[/Y]]. Qed. (* "eliminating" growth *) @@ -2742,15 +2758,15 @@ have Eh: um_filter (le t2) h = h0 \+ (h1 \+ (pts k v \+ h2)). by rewrite (umfilt_le_split h K2) (umfilt_pt_split H) -!joinA. have W1 : valid (h0 \+ (h1 \+ (pts k v \+ h2))) by rewrite -Eh pfV. rewrite eval_umfilt (umfilt_le_split h K2) evalUn ?(validAL W1) //; last first. -- by move=>??/In_dom_umfilt[?][/leq_ltn_trans Y] _ /In_dom_umfilt[?][] /andP [/Y]. +- by move=>??/In_dom_umfilt[?/leq_ltn_trans Y] _ /In_dom_umfilt[?] /andP [/Y]. rewrite -(eval_umfilt (le t1)); apply: helper2 (validR W1) _ _ _ _ =>//. - by apply: growR W1 _; rewrite -Eh=>k1 v1 z1 /In_umfiltX [] _ /G. -- by move=>x /In_dom_umfilt; rewrite (subX t1 x) //; case=>v0 [] /andP []. -- by apply/allP=>x /In_dom_umfilt; case=>v0 [] /andP []. +- by move=>x /In_dom_umfilt; rewrite (subX t1 x) //; case=>v0 /andP []. +- by apply/allP=>x /In_dom_umfilt; case=>v0 /andP []. rewrite eval_umfilt Eh evalUn ?(validX W1) -?eval_umfilt // in E. -move=>x y /In_dom_umfilt; case=>v1 [/leq_ltn_trans Y _]. +move=>x y /In_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) =>/In_dom_umfilt; case=>v0 [/andP][/Y]. +by rewrite -(umfilt_lt_split h K4) =>/In_dom_umfilt; case=>v0 /andP [/Y]. Qed. End Growing. @@ -2762,14 +2778,17 @@ Variables (V R : Type) (U : natmap V) (X : ordType) (a : R -> V -> R) (f : R -> Implicit Types p : pred (nat*V). Lemma growIn (h : U) t1 t2 z0 : - growing a f h -> t1 <= t2 -> + 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 : U) t1 t2 z0 k v : - growing a f h -> (k, v) \In h -> t1 < k <= t2 -> + 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. diff --git a/pcm/unionmap.v b/pcm/unionmap.v index a86963f..54e0251 100644 --- a/pcm/unionmap.v +++ b/pcm/unionmap.v @@ -4167,12 +4167,12 @@ Lemma In_umfilt p x f : x \In f -> p x -> x \In um_filter p f. Proof. by move=>X1 X2; apply/In_umfiltX. Qed. Lemma In_dom_umfilt p f k : - reflect (exists v, [/\ p (k, v) & (k, v) \In f]) + reflect (exists2 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_umfiltX []; exists v. -by case=>v [Hp Hf]; elim: H; exists v; apply/In_umfilt. +apply: (iffP (In_domX _ _)). +- by case=>v /In_umfiltX []; exists v. +by case=>v Hp Hf; exists v; apply/In_umfilt. Qed. Lemma dom_omf_umfilt V' (U' : union_map C V') (f : omap_fun U U') x : @@ -4180,8 +4180,8 @@ Lemma dom_omf_umfilt V' (U' : union_map C V') (f : omap_fun U U') x : Proof. apply/domE=>k; apply/idP/idP. - case/In_dom_omfX=>//= v [H1 H2]. - by apply/In_dom_umfilt; exists v; rewrite /= H2. -case/In_dom_umfilt=>w [/=]. + by apply/In_dom_umfilt; exists v. +case/In_dom_umfilt=>w /=. case E: (omf f (k, w))=>[a|] // _ H. by move/In_dom: (In_omf _ H E). Qed. @@ -4198,7 +4198,7 @@ Proof. apply: ord_sorted_eq=>//=. - by apply: sorted_filter; [apply: trans | apply: sorted_dom]. move=>k; rewrite mem_filter; apply/idP/idP. -- by case/In_dom_umfilt=>w [H1 H2]; move/In_find: H2 (H2) H1=>-> /In_dom ->->. +- by case/In_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/In_dom_umfilt; exists v. Qed. @@ -4329,7 +4329,7 @@ move: f z0; elim/um_indf=>[||k v f IH W P] z0 /=. - by rewrite !pfunit !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 /In_dom_umfilt [w][_] /In_dom. +- apply/allP=>x /In_dom_umfilt [w _] /In_dom. by apply: path_mem (@trans K) P. have : valid (um_filter p (pts k v \+ f)) by rewrite pfVE W. by rewrite !umfiltPtUn W; case: ifP=>E V3; rewrite !umfoldlPtUn // E IH. @@ -4390,11 +4390,11 @@ Lemma dom_umfilt2 p1 p2 f x : (x \in dom (um_filter p1 f)) && (x \in dom (um_filter p2 f)). Proof. rewrite -umfilt_predI; apply/idP/idP. -- case/In_dom_umfilt=>v [/andP [X1 X2] H]. +- case/In_dom_umfilt=>v /andP [X1 X2] H. by apply/andP; split; apply/In_dom_umfilt; exists v. -case/andP=>/In_dom_umfilt [v1][X1 H1] /In_dom_umfilt [v2][X2 H2]. +case/andP=>/In_dom_umfilt [v1 X1 H1] /In_dom_umfilt [v2 X2 H2]. move: (In_fun H1 H2)=>E; rewrite -{v2}E in X2 H2 *. -by apply/In_dom_umfilt; exists v1; split=>//; apply/andP. +by apply/In_dom_umfilt; exists v1=>//; apply/andP. Qed. End FilterDefLemmas. @@ -4418,7 +4418,7 @@ Proof. apply: ord_sorted_eq=>//=. - by apply: sorted_filter; [apply: trans | apply: sorted_dom]. move=>k; rewrite mem_filter; apply/idP/idP. -- by case/In_dom_umfilt=>v [/= -> /In_dom]. +- by case/In_dom_umfilt=>v /= -> /In_dom. by case/andP=>H1 /In_domX [v H2]; apply/In_dom_umfilt; exists v. Qed. @@ -4428,8 +4428,8 @@ Lemma valid_umfiltkUn p1 p2 f : valid (um_filterk p1 f \+ um_filterk p2 f). Proof. move=>W H; rewrite validUnAE !pfVE W /=. -apply/allP=>x; case/In_dom_umfilt=>v1 /= [H2 F1]. -apply/negP; case/In_dom_umfilt=>v2 /= [H1 F2]. +apply/allP=>x; case/In_dom_umfilt=>v1 /= H2 F1. +apply/negP; case/In_dom_umfilt=>v2 /= H1 F2. by move: (H x (In_dom F1) H1 H2). Qed. @@ -4440,7 +4440,8 @@ Proof. by move=>k; rewrite dom_umfiltkE mem_filter. Qed. (* this also holds for invalid f1, as the corollary shows *) (* /DEVCOMMENT *) Lemma umfiltk_dom f1 f2 : - valid (f1 \+ f2) -> um_filterk (mem (dom f1)) (f1 \+ f2) = f1. + valid (f1 \+ f2) -> + um_filterk (mem (dom f1)) (f1 \+ f2) = f1. Proof. move=>W; apply/umem_eq; first by rewrite pfVE. - by rewrite (validL W). @@ -4456,14 +4457,16 @@ by rewrite -{2}[f]unitR umfiltk_dom // unitR. Qed. Lemma eq_in_umfiltk p1 p2 f : - {in dom f, p1 =1 p2} -> um_filterk p1 f = um_filterk 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 umfiltk_subdom p f : {subset dom (um_filterk p f) <= p}. Proof. by move=>a; rewrite dom_umfiltk; case/andP. Qed. -Lemma umfiltk_subdomE p f : {subset dom f <= p} <-> um_filterk p f = f. +Lemma umfiltk_subdomE 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. @@ -4478,7 +4481,8 @@ Lemma find_umfiltk k (p : pred K) f : Proof. by rewrite find_umfilt /=; case: (find _ _)=>[a|]; case: ifP. Qed. Lemma umfiltk_subdom0 p f : - valid f -> {subset dom f <= predC p} <-> um_filterk p f = Unit. + valid f -> + {subset dom f <= predC p} <-> um_filterk p f = Unit. Proof. move=>W; split=>[H|H k X]. - rewrite (eq_in_umfiltk (p2:=pred0)) ?umfilt_pred0 //. @@ -5064,8 +5068,7 @@ Lemma eval_oevUmfilt 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_umfiltX. -split; last by case. -split=>//; move: H; move/In_dom_umfilt => [v' [H H1]]. +split=>[|[]//]; split=>//; case/In_dom_umfilt: H=>v' H H1. by rewrite (In_fun H1 H0) in H. Qed. @@ -5421,7 +5424,7 @@ Lemma umcntF p k v f : Proof. move=>H; move/In_dom: (H)=>/= D; rewrite /um_count. case E: (k \in dom (um_filter p f)). -- case/In_dom_umfilt: E=>w [H1 H2]. +- case/In_dom_umfilt: E=>w H1 H2. rewrite -{w H2}(In_fun H H2) in H1 *. rewrite H1 omfF size_domF ?subn1 //=. by apply/In_dom_umfilt; exists v. @@ -5498,7 +5501,7 @@ Proof. case: (normalP f)=>[->|W]; first by rewrite !umcnt_undef. rewrite /um_count umfilt_predU size_domUn //. case: validUn=>//; rewrite ?(pfV,W) //. -move=>k /In_dom_umfilt [v1 [P1 H1]] /In_dom_umfilt [v2 [/= P2 H2]]. +move=>k /In_dom_umfilt [v1 P1 H1] /In_dom_umfilt [v2 /= P2 H2]. by rewrite -(In_fun H1 H2) P1 in P2. Qed. @@ -5878,7 +5881,7 @@ Proof. case=>Vh V D. rewrite validUnAE !pfV //=; apply/allP=>x /In_dom_omapX []. move=>t1 [/In_dom /= H _]; apply/In_dom_umfilt. -case; case=>tx vx [/eqP /= N]. +case; case=>tx vx /= /eqP /= N. by move/(In_side Ut)/In_dom/D/(_ H)/esym/N. Qed. From 75bfd6b9f3f46fb712fcb14f90942b2b9adb3ded Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Wed, 28 Aug 2024 18:29:14 +0200 Subject: [PATCH 67/93] found a way to view "weird" lemmas in natmap.v as non-weird. removed the corresponding comment, and repositioned the lemmas. --- pcm/natmap.v | 78 ++++++++++++++++++++++++++++++---------------------- 1 file changed, 45 insertions(+), 33 deletions(-) diff --git a/pcm/natmap.v b/pcm/natmap.v index d1b0072..4e483ed 100644 --- a/pcm/natmap.v +++ b/pcm/natmap.v @@ -323,17 +323,10 @@ by rewrite H1 lt0n (negbTE H2) H4; apply/lastkey_dom_valid/In_last. Qed. (* main membership invariant of last_key *) -Lemma lastkey_mem0 A (U : natmap A) (h : U) : last_key h \in 0 :: dom h. +Lemma lastkey_mem0 A (U : natmap A) (h : U) : + last_key h \in 0 :: dom h. Proof. by rewrite inE; case: lastkeyP. Qed. -(* DEVCOMMENT: *) -(* Weird variation. Worth preserving? *) -(* /DEVCOMMENT *) -Lemma lastkey_memk A (U : natmap A) (h : U) k : - (k < last_key h -> last_key h \notin dom h) -> - last_key h <= k. -Proof. by case: lastkeyP=>//= v _; case: ltngtP=>// _; apply. Qed. - (* view for last_val can simplify obligations from last_keyP *) (* by joining the cases for unit and undef *) Variant lastval_spec A (U : natmap A) (h : U) : option A -> Type := @@ -359,7 +352,9 @@ Variables (A : Type) (U : natmap A). Implicit Type h : U. (* last_key is upper bound *) -Lemma dom_lastkey h k : k \in dom h -> k <= last_key h. +Lemma dom_lastkey h k : + k \in dom h -> + k <= last_key h. Proof. rewrite /last_key !umEX /UM.dom; case: (UMC_from h)=>//; case=>s H _ /=. rewrite /supp/ord /= (leq_eqVlt k) orbC. @@ -374,22 +369,21 @@ Proof. by apply: contraL; rewrite -leqNgt; apply: dom_lastkey. Qed. Lemma lastkey_dom0 h k : last_key h < k -> k \notin 0 :: dom h. Proof. by apply: contraL; rewrite -leqNgt; apply: dom0_lastkey. Qed. -(* DEVCOMMENT: *) -(* Weird variation. Worth preserving? *) -(* /DEVCOMMENT *) -Lemma dom_lastkey_eq h k : - k \in dom h -> - (k < last_key h -> last_key h \notin dom h) -> - k = last_key h. -Proof. -move/dom_lastkey=>D /lastkey_memk H. -by rewrite (@anti_leq (last_key _) k) // H D. -Qed. - (* last_key is least upper bound *) -Lemma lastkey_lub h k : {in dom h, forall x, x <= k} -> last_key h <= k. +Lemma lastkey_lub h k : + {in dom h, forall x, x <= k} -> + last_key h <= k. Proof. by case lastkeyP=>// v /In_dom H; apply. Qed. +(* variant with contrapositive side-condition *) +Lemma lastkey_memk h k : + (forall x, k < x -> x \notin dom h) -> + last_key h <= k. +Proof. +move=>H; apply: lastkey_lub=>x; rewrite leqNgt. +by apply/contraL/H. +Qed. + (* useful variant with equality *) Lemma lub_lastkey h k : k \in dom h -> @@ -402,27 +396,49 @@ apply: lastkey_lub=>x; apply: contraTT. by rewrite -ltnNge; apply: X. Qed. +(* no need to quantify over x in lastkey_memk *) +(* it suffices to use x := last_key h *) +Lemma lastkey_memkX h k : + (k < last_key h -> last_key h \notin dom h) -> + last_key h <= k. +Proof. by case: lastkeyP=>//= v _; case: ltngtP=>// _; apply. Qed. + +(* no need to quantify over x in lub_lastkey *) +(* it suffices to use x := last_key h *) +Lemma lub_lastkeyX h k : + k \in dom h -> + (k < last_key h -> last_key h \notin dom h) -> + k = last_key h. +Proof. +move/dom_lastkey=>D /lastkey_memkX H. +by rewrite (@anti_leq (last_key _) k) // H D. +Qed. + (* equivalences *) -Lemma lastkey_leq h k : (last_key h <= k) = all (fun x => x <= k) (dom h). +Lemma lastkey_leq h k : + (last_key h <= k) = all (fun x => x <= k) (dom h). Proof. apply/idP/allP=>[H x /dom_lastkey D|]; by [apply: leq_trans D H | apply: lastkey_lub]. Qed. -Lemma leq_lastkey h k : (k <= last_key h) = has (fun x => k <= x) (0 :: dom h). +Lemma leq_lastkey h k : + (k <= last_key h) = has (fun x => k <= x) (0 :: dom h). Proof. apply/idP/hasP=>[|[x] D H]; last by apply: leq_trans H (dom0_lastkey D). by exists (last_key h)=>//; apply: lastkey_mem0. Qed. -Lemma ltn_lastkey h k : (k < last_key h) = has (fun x => k < x) (dom h). +Lemma ltn_lastkey h k : + (k < last_key h) = has (fun x => k < x) (dom h). Proof. apply/idP/hasP=>[|[x] D H]; last by apply: leq_trans H (dom_lastkey D). by exists (last_key h)=>//; case: lastkeyP H. Qed. -Lemma lastkey_ltn h k : (last_key h < k) = all (fun x => x < k) (0 :: dom h). +Lemma lastkey_ltn h k : + (last_key h < k) = all (fun x => x < k) (0 :: dom h). Proof. apply/idP/allP=>[H x D|]; last by apply; apply: lastkey_mem0. by apply: leq_ltn_trans (dom0_lastkey D) H. @@ -692,18 +708,14 @@ Proof. by move/lastkey_dom; apply/(subsetC omf_subdom). Qed. Lemma lastkey_odom0 f h k : last_key h < k -> k \notin 0 :: dom (f h). Proof. move/lastkey_dom0; apply/(subsetC omf_subdom0). Qed. -(* DEVCOMMENT: *) -(* Weird variation. Worth preserving? *) -(* composes dom_lastkey_eq with In_dom_omfX *) -(* /DEVCOMMENT *) -Lemma odom_lastkey_eq f h k v : +Lemma olub_lastkeyX f h k v : (k, v) \In h -> omf f (k, v) -> (forall w, k < last_key (f h) -> (last_key (f h), w) \In h -> omf f (last_key (f h), w) -> False) -> k = last_key (f h). Proof. -move=>D1 D2 H; apply: dom_lastkey_eq=>[|X]. +move=>D1 D2 H; apply: lub_lastkeyX=>[|X]. - by apply/In_dom_omfX; exists v. by apply/In_dom_omfX; case=>w []; apply: H X. Qed. From f01615cd02d622d2faecf2d3bd5fc9df19327ce7 Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Wed, 28 Aug 2024 20:57:57 +0200 Subject: [PATCH 68/93] propagating changed to natmap.v from mathador --- pcm/natmap.v | 57 +++++++++++++++++++++++++++++++++++++++++++++++++--- 1 file changed, 54 insertions(+), 3 deletions(-) diff --git a/pcm/natmap.v b/pcm/natmap.v index 4e483ed..d1d6d2c 100644 --- a/pcm/natmap.v +++ b/pcm/natmap.v @@ -508,7 +508,7 @@ Proof. by rewrite joinC; apply: lastkeyUnLT. Qed. (* disjoint maps with equal last keys are both empty *) Lemma lastkeyUn0T A1 A2 (U1 : natmap A1) (U2 : natmap A2) (h1 : U1) (h2 : U2) : valid h1 -> valid h2 -> - (forall x, x \in dom h1 -> x \in dom h2 -> false) -> + (forall x, x \in dom h1 -> x \in dom h2 -> False) -> last_key h1 = last_key h2 -> (h1 = Unit) * (h2 = Unit). Proof. @@ -527,7 +527,7 @@ Proof. by case: validUn=>// ?? D _; apply: lastkeyUn0T=>// x /D/negbTE ->. Qed. (* variation to avoid validity *) Lemma lastkey00 A1 A2 (U1 : natmap A1) (U2 : natmap A2) (h1 : U1) (h2 : U2) : - (forall x, x \in dom h1 -> x \in dom h2 -> false) -> + (forall x, x \in dom h1 -> x \in dom h2 -> False) -> last_key h1 = last_key h2 -> (last_key h1 = 0) * (last_key h2 = 0). Proof. @@ -539,13 +539,28 @@ Qed. (* variation with subdoms *) Lemma lastkey_subdom00 A1 A2 (U1 : natmap A1) (U2 : natmap A2) (h1 : U1) (h2 : U2) (s1 s2 : pred nat) : - (forall x, x \in s1 -> x \in s2 -> false) -> + (forall x, x \in s1 -> x \in s2 -> False) -> {subset dom h1 <= s1} -> {subset dom h2 <= s2} -> last_key h1 = last_key h2 -> (last_key h1 = 0) * (last_key h2 = 0). Proof. by move=>D S1 S2; apply: lastkey00=>x /S1 H /S2; apply: D H. Qed. +(* variation with subdoms and indexing *) +Lemma lastkey_subdom00E A1 A2 (U1 : natmap A1) (U2 : natmap A2) + (g : nat -> pred nat) (h1 : U1) (h2 : U2) (t1 t2 : nat) : + (forall x, x \in g t1 -> x \in g t2 -> t1 = t2) -> + {subset dom h1 <= g t1} -> + {subset dom h2 <= g t2} -> + last_key h1 != 0 -> + last_key h1 = last_key h2 -> + t1 = t2. +Proof. +move=>H S1 S2 X E; apply: contraNeq X=>/eqP X. +move/(lastkey_subdom00 _ S1 S2): E=>-> //. +by move=>x D1 /(H x D1)/X. +Qed. + (* interaction with constructors *) Section LastkeyConstructors. @@ -800,6 +815,42 @@ Qed. End LastKeyOmapFun. +(* lastkey_subdom00 for application of omap functions *) +Lemma lastkey00E A A1 A2 (U : natmap A) (U1 : natmap A1) (U2 : natmap A2) + (f1 : omap_fun U U1) (f2 : omap_fun U U2) (h1 : U) (h2 : U) : + (forall x, x \in dom h1 -> x \in dom h2 -> False) -> + last_key (f1 h1) = last_key (f2 h2) -> + (last_key (f1 h1) = 0) * (last_key (f2 h2) = 0). +Proof. by move=>H; apply: lastkey_subdom00 H omf_subdom omf_subdom. Qed. + +(* formulation with application and indexing *) +(* NOTE: not a consequence of lastkey_subdom00E *) +Lemma lastkeyFE T A A1 A2 (U : natmap A) (g : nat -> T -> U) + (U1 : natmap A1) (f1 : omap_fun U U1) + (U2 : natmap A2) (f2 : omap_fun U U2) + (h1 h2 : T) t1 t2 : + (forall x, x \in dom (g t1 h1) -> x \in dom (g t2 h2) -> t1 = t2) -> + last_key (f1 (g t1 h1)) != 0 -> + last_key (f1 (g t1 h1)) = last_key (f2 (g t2 h2)) -> + t1 = t2. +Proof. +move=>X N1 E; apply: contraNeq N1=>/eqP N2. +by move/lastkey00E: E=>-> // x D1 /(X x D1)/N2. +Qed. + +(* formulation with application of composition and with indexing *) +Lemma lastkeyFFE T A A1 A2 A1' A2' (U : natmap A) (g : nat -> T -> U) + (U1 : natmap A1) (f1 : omap_fun U U1) + (U2 : natmap A2) (f2 : omap_fun U U2) + (U1' : natmap A1') (f1' : omap_fun U1 U1') + (U2' : natmap A2') (f2' : omap_fun U2 U2') + (h1 h2 : T) t1 t2 : + (forall x, x \in dom (g t1 h1) -> x \in dom (g t2 h2) -> t1 = t2) -> + last_key (f1' (f1 (g t1 h1))) != 0 -> + last_key (f1' (f1 (g t1 h1))) = last_key (f2' (f2 (g t2 h2))) -> + t1 = t2. +Proof. by rewrite (compE f1') (compE f2'); apply: lastkeyFE. Qed. + (* last_key and filter *) Lemma lastkey_umfilt_le A (U : natmap A) (h : U) k : From b426ffe046ce00606016f7c9d8316f7310ffafa0 Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Thu, 29 Aug 2024 15:28:16 +0200 Subject: [PATCH 69/93] added alternative lemma for seq_lt irreflexivity, one that isn't given as an equation, but as implication into False. --- core/useqord.v | 4 ++++ 1 file changed, 4 insertions(+) diff --git a/core/useqord.v b/core/useqord.v index ec0b84e..35f00e0 100644 --- a/core/useqord.v +++ b/core/useqord.v @@ -102,6 +102,10 @@ Proof. by move=>x; rewrite seqle_unlock. Qed. Lemma slt_irr x ks : x <[ks] x = false. Proof. by rewrite seqlt_unlock; apply: ltnn. Qed. +(* non-equational variant *) +Lemma slt_irrN x ks : ~ x <[ks] x. +Proof. by rewrite slt_irr. Qed. + (****************** antisymmetry ****************) Lemma sle_antisym ks : {in ks, antisymmetric (seq_le ks)}. From af748a8dbdc95cf6cd9b06e7719d102917b54ecc Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Thu, 29 Aug 2024 15:49:08 +0200 Subject: [PATCH 70/93] renaming slt_irrN into sltnn --- core/useqord.v | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/core/useqord.v b/core/useqord.v index 35f00e0..3b97c06 100644 --- a/core/useqord.v +++ b/core/useqord.v @@ -103,7 +103,7 @@ Lemma slt_irr x ks : x <[ks] x = false. Proof. by rewrite seqlt_unlock; apply: ltnn. Qed. (* non-equational variant *) -Lemma slt_irrN x ks : ~ x <[ks] x. +Lemma sltnn x ks : ~ x <[ks] x. Proof. by rewrite slt_irr. Qed. (****************** antisymmetry ****************) From ef24ddceedd0f4021f7f95a9d1a1d60a5ffe9710 Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Thu, 29 Aug 2024 18:06:16 +0200 Subject: [PATCH 71/93] added validPt2 and domPt2 --- pcm/unionmap.v | 16 ++++++++++++++++ 1 file changed, 16 insertions(+) diff --git a/pcm/unionmap.v b/pcm/unionmap.v index 54e0251..667873c 100644 --- a/pcm/unionmap.v +++ b/pcm/unionmap.v @@ -1745,6 +1745,14 @@ 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. +Lemma validPt2 k1 k2 v1 v2 : + valid (pts k1 v1 \+ pts k2 v2 : U) = + [&& C k1, C k2 & k1 != k2]. +Proof. +rewrite validPtUn validPt domPt inE negb_and. +by case: (C k2)=>//=; rewrite eq_sym. +Qed. + (* the projections from validPtUn are often useful *) Lemma validPtUnI v1 k f : @@ -1794,6 +1802,14 @@ 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 domPt2 k1 k2 v1 v2 x : + x \in dom (pts k1 v1 \+ pts k2 v2 : U) = + [&& C k1, C k2, k1 != k2 & x \in pred2 k1 k2]. +Proof. +rewrite domPtUn !inE validPt2 domPt inE !(eq_sym x). +by case: (C k1)=>//=; case: (C k2). +Qed. + Lemma validxx f : valid (f \+ f) -> dom f =i pred0. Proof. by case: validUn=>// _ _ L _ z; case: (_ \in _) (L z)=>//; elim. Qed. From ebdc17cd72f32b47fd062c2ccbaa749e24c4d557 Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Wed, 4 Sep 2024 19:15:03 +0200 Subject: [PATCH 72/93] propagating changes from mathador --- examples/graph.v | 192 +++++++++++++++++++++++++++++------------------ pcm/unionmap.v | 67 ++++++++++------- 2 files changed, 158 insertions(+), 101 deletions(-) diff --git a/examples/graph.v b/examples/graph.v index 84f728c..5780ec8 100644 --- a/examples/graph.v +++ b/examples/graph.v @@ -39,25 +39,28 @@ From pcm Require Import pcm unionmap natmap autopcm automap. (* not from a fixed finite set. *) Notation node := nat. -Record pregraph := Pregraph {pregraph_base : @UM.base node nat_pred (seq node)}. +(* A is the contents/labeling of the nodes *) +Record pregraph (A : Type) := + Pregraph {pregraph_base : @UM.base node nat_pred (A * seq node)}. Section PregraphUMC. -Implicit Type f : pregraph. +Variable A : Type. +Implicit Type f : pregraph A. Local Coercion pregraph_base : pregraph >-> UM.base. -Let pg_valid f := @UM.valid nat nat_pred (seq nat) f. -Let pg_empty := Pregraph (@UM.empty nat nat_pred (seq nat)). -Let pg_undef := Pregraph (@UM.Undef nat nat_pred (seq nat)). -Let pg_upd k v f := Pregraph (@UM.upd nat nat_pred (seq nat) k v f). -Let pg_dom f := @UM.dom nat nat_pred (seq nat) f. -Let pg_assocs f := @UM.assocs nat nat_pred (seq nat) f. -Let pg_free f k := Pregraph (@UM.free nat nat_pred (seq nat) f k). -Let pg_find k f := @UM.find nat nat_pred (seq nat) k f. -Let pg_union f1 f2 := Pregraph (@UM.union nat nat_pred (seq nat) f1 f2). -Let pg_empb f := @UM.empb nat nat_pred (seq nat) f. -Let pg_undefb f := @UM.undefb nat nat_pred (seq nat) f. -Let pg_from (f : pregraph) : UM.base _ _ := f. -Let pg_to (b : @UM.base nat nat_pred (seq nat)) : pregraph := Pregraph b. -Let pg_pts k v := Pregraph (@UM.pts nat nat_pred (seq nat) k v). +Let pg_valid f := @UM.valid nat nat_pred (A * seq nat) f. +Let pg_empty := Pregraph (@UM.empty nat nat_pred (A * seq nat)). +Let pg_undef := Pregraph (@UM.Undef nat nat_pred (A * seq nat)). +Let pg_upd k v f := Pregraph (@UM.upd nat nat_pred (A * seq nat) k v f). +Let pg_dom f := @UM.dom nat nat_pred (A * seq nat) f. +Let pg_assocs f := @UM.assocs nat nat_pred (A * seq nat) f. +Let pg_free f k := Pregraph (@UM.free nat nat_pred (A * seq nat) f k). +Let pg_find k f := @UM.find nat nat_pred (A * seq nat) k f. +Let pg_union f1 f2 := Pregraph (@UM.union nat nat_pred (A * seq nat) f1 f2). +Let pg_empb f := @UM.empb nat nat_pred (A * seq nat) f. +Let pg_undefb f := @UM.undefb nat nat_pred (A * seq nat) f. +Let pg_from (f : pregraph A) : UM.base _ _ := f. +Let pg_to (b : @UM.base nat nat_pred (A * seq nat)) : pregraph A := Pregraph b. +Let pg_pts k v := Pregraph (@UM.pts nat nat_pred (A * seq nat) k v). Lemma pregraph_is_umc : union_map_axiom pg_valid pg_empty pg_undef pg_upd pg_dom @@ -65,19 +68,21 @@ Lemma pregraph_is_umc : pg_undefb pg_pts pg_from pg_to. Proof. by split=>//; split=>[|[]]. Qed. -HB.instance Definition _ := - isUnion_map.Build node nat_pred (seq node) pregraph pregraph_is_umc. +HB.instance Definition blah := + isUnion_map.Build node nat_pred (A * seq node)%type + (pregraph A) pregraph_is_umc. End PregraphUMC. -HB.instance Definition _ := isNatMap.Build (seq node) pregraph. -HB.instance Definition _ := - hasDecEq.Build pregraph (@union_map_eqP node _ (seq node) pregraph). -Canonical pregraph_PredType : PredType (node * (seq node)) := - um_PredType pregraph. -Coercion Pred_of_history (x : pregraph) : {Pred _} := +HB.instance Definition _ A := isNatMap.Build (A * seq node)%type (pregraph A). +HB.instance Definition _ (A : eqType) := + hasDecEq.Build (pregraph A) + (@union_map_eqP node _ (A * seq node)%type (pregraph A)). +Canonical pregraph_PredType A : PredType (node * (A * seq node)) := + um_PredType (pregraph A). +Coercion Pred_of_history A (x : pregraph A) : {Pred _} := [eta Mem_UmMap x]. -Notation "x &-> v" := (ptsT pregraph x v) (at level 30). +Notation "x &-> v" := (ptsT (@pregraph _) x v) (at level 30). (**************************) (* Links, children, edges *) @@ -86,23 +91,27 @@ Notation "x &-> v" := (ptsT pregraph x v) (at level 30). (* Links of x includes all edges outgoing from x *) (* and may explicitly include nodes that aren't in dom g *) (* (i.e., are dangling, null or non-null) *) -Definition links (g : pregraph) x := oapp id [::] (find x g). +Definition links A (g : pregraph A) x := oapp snd [::] (find x g). (* children x removes dangling edges (null or non-null) from links *) -Definition children (g : pregraph) x : seq node := +Definition children A (g : pregraph A) x : seq node := filter (mem (dom g)) (links g x). (* edge is applicative variant of children *) (* thus, dangling edges (null or non-null) are *not* edges. *) -Definition edge g : rel node := mem \o children g. -Arguments edge g x y : simpl never. +Definition edge A (g : pregraph A) : rel node := mem \o children g. +Arguments edge {A} g x y : simpl never. + +Section PregraphLemmas. +Context {A : Type}. +Implicit Type g : pregraph A. (* links lemmas *) -Lemma links_undef x : links undef x = [::]. +Lemma links_undef x : links (undef : pregraph A) x = [::]. Proof. by []. Qed. -Lemma links_unit x : links Unit x = [::]. +Lemma links_unit x : links (Unit : pregraph A) x = [::]. Proof. by []. Qed. Lemma linksND g x : @@ -127,27 +136,25 @@ Lemma size_links g x : x \in dom g. Proof. by rewrite /links/oapp; case: dom_find. Qed. -Lemma linksD (g : pregraph) x y : +Lemma linksD g x y : y \in links g x -> x \in dom g. Proof. by move=>X; apply: size_links; case: (links g x) X. Qed. -Lemma In_graph (g : pregraph) x xs : - (x, xs) \In g -> +Lemma In_graph g x v xs : + (x, (v, xs)) \In g -> xs = links g x. Proof. by rewrite /links/oapp=>/In_find ->. Qed. -Lemma In_graphX (g : pregraph) x : +Lemma In_graphX g x : x \in dom g -> - (x, links g x) \In g. -Proof. -by move=>Dx; apply/In_find; rewrite /links/oapp; case: dom_find Dx. -Qed. + exists v, (x, (v, links g x)) \In g. +Proof. by case/In_domX=>-[v xs] /[dup] /In_graph ->; eauto. Qed. -Lemma range_links (g : pregraph) x : +Lemma range_links g x : x \in dom g -> - links g x \in range g. -Proof. by move/In_graphX/mem_range. Qed. + links g x \in map snd (range g). +Proof. by case/In_graphX=>v /In_range/(Mem_map snd)/mem_seqP. Qed. Lemma links_umfiltk g p x : links (um_filterk p g) x =i @@ -156,10 +163,10 @@ Proof. by move=>i; rewrite /links find_umfiltk; case: (p x). Qed. (* children lemmas *) -Lemma children_undef x : children undef x = [::]. +Lemma children_undef x : children (undef : pregraph A) x = [::]. Proof. by []. Qed. -Lemma children_unit x : children Unit x = [::]. +Lemma children_unit x : children (Unit : pregraph A) x = [::]. Proof. by []. Qed. Lemma childrenND g x : @@ -184,15 +191,15 @@ Lemma childrenUnR g1 g2 x : {subset children g2 x <= children (g1 \+ g2) x}. Proof. by rewrite joinC; apply: childrenUnL. Qed. -Lemma children_links (g : pregraph) x : +Lemma children_links g x : {subset children g x <= links g x}. Proof. by move=>z; rewrite /children mem_filter=>/andP []. Qed. (* if x is node in g then g x contains all children of x *) (* and maybe some more nodes that aren't in g *) -Lemma range_children (g : pregraph) x : +Lemma range_children g x : x \in dom g -> - exists2 xs, xs \in range g & + exists2 xs, xs \in map snd (range g) & {subset children g x <= xs}. Proof. move=>Dx; exists (links g x); first by apply: range_links. @@ -209,10 +216,10 @@ Qed. (* edge lemmas *) -Lemma edge_undef x y : edge undef x y = false. +Lemma edge_undef x y : edge (undef : pregraph A) x y = false. Proof. by rewrite /edge/= children_undef. Qed. -Lemma edge_unit x y : edge Unit x y = false. +Lemma edge_unit x y : edge (Unit : pregraph A) x y = false. Proof. by rewrite /edge/= children_unit. Qed. Lemma edge_children g x y : @@ -263,6 +270,8 @@ Lemma edge_umfiltkE g p : {in predC p &, edge (um_filterk (predC p) g) =2 edge g}. Proof. by move=>x y; rewrite !inE => X Y; rewrite edge_umfiltk /= X Y. Qed. +End PregraphLemmas. + (***********************) (* Depth-first search *) (***********************) @@ -277,10 +286,14 @@ Proof. by move=>x y; rewrite !inE => X Y; rewrite edge_umfiltk /= X Y. Qed. (* If the latter is desired, it can be separately defined *) (* as a conjunct of dfs and links properties. *) -Fixpoint dfs (g : pregraph) (n : nat) (v : seq node) x := +Fixpoint dfs A (g : pregraph A) (n : nat) (v : seq node) x := if (x \notin dom g) || (x \in v) then v else if n is n'.+1 then foldl (dfs g n') (x :: v) (children g x) else v. +Section DFSLemmas. +Context {A : Type}. +Implicit Type g : pregraph A. + Lemma dfs_notin g n v x : x \notin dom g -> dfs g n v x = v. @@ -413,6 +426,8 @@ apply/IHn=>//; exists (p ++ pr). by rewrite -cat_cons disjoint_catR D (disjoint_consRE D2). Qed. +End DFSLemmas. + (******************) (* Connectivity *) (* (reachability) *) @@ -421,17 +436,19 @@ Qed. (* start dfs with p as avoidance set, then filter out p. *) (* this computes only the nodes visited during dfs. *) (* not for client use, hence primed name *) -Definition component' (p : pred node) g x : seq node := +Definition component' A (p : pred node) (g : pregraph A) x : seq node := filter (predC p) (dfs g (size (dom g)) (filter p (dom g)) x). (* y is connected from x if y is visited during dfs *) (* avoiding nodes from p; assuming x in dom g *) -Definition connect p (g : pregraph) x : pred node := +Definition connect A p (g : pregraph A) x : pred node := fun y => (x \in dom g) && (y \in component' p g x). -(* connect lemmas *) +Section ConnectLemmas. +Context {A : Type}. +Implicit Type g : pregraph A. -Lemma connectP (p : pred node) (g : pregraph) x y : +Lemma connectP (p : pred node) g x y : reflect [/\ x \in dom g, ~~ p x & exists xs, [/\ path (edge g) x xs, y = last x xs & {in xs, forall z, ~~ p z}]] @@ -460,7 +477,7 @@ Qed. (* there's path from x to y iff *) (* there's path that doesn't loop through x *) -Lemma connectX (p : pred node) (g : pregraph) x y : +Lemma connectX (p : pred node) g x y : reflect [/\ x \in dom g, ~~ p x & exists xs, [/\ path (edge g) x xs, y = last x xs, x \notin xs & @@ -482,10 +499,10 @@ case/and3P: Px=>Px Ex Pp2; exists p2; split=>//. by move=>z Z; rewrite Hxs // mem_cat Z orbT. Qed. -Lemma connect_undef p x : connect p undef x =i pred0. +Lemma connect_undef p x : connect p (undef : pregraph A) x =i pred0. Proof. by move=>y; apply/connectP; case; rewrite dom_undef. Qed. -Lemma connect_unit p x : connect p Unit x =i pred0. +Lemma connect_unit p x : connect p (Unit : pregraph A) x =i pred0. Proof. by move=>y; apply/connectP; case; rewrite dom0. Qed. Lemma connectDx p g x y : @@ -534,13 +551,13 @@ Lemma connect0 p g x : x \in connect p g x = (x \in dom g) && ~~ p x. Proof. by apply/connectP/andP; case=>// H1 H2; split=>//; exists [::]. Qed. -Lemma connect0I (p : pred node) (g : pregraph) x : +Lemma connect0I (p : pred node) g x : x \in dom g -> ~~ p x -> x \in connect p g x. Proof. by rewrite connect0=>->->. Qed. -Lemma connect0N (p : pred node) (g : pregraph) x y : +Lemma connect0N (p : pred node) g x y : x \in dom g -> ~~ p x -> y \notin connect p g x -> @@ -601,10 +618,10 @@ Proof. case: (normalP g)=>[->|V y]; first by rewrite pfundef. apply/connectP/connectP; case; rewrite dom_umfiltk inE /=. - case/andP=>Hx Dx _ [xs][Px Ey Hxs]; split=>//. - exists xs; rewrite -(eq_in_path (@edge_umfiltkE g p)) //. + exists xs; rewrite -(eq_in_path (@edge_umfiltkE A g p)) //. by apply/allP/in_consE. move=>Dx Hx [xs][Px Ey Hxs]; rewrite Dx Hx; split=>//. -exists xs; rewrite (eq_in_path (@edge_umfiltkE g p)) //. +exists xs; rewrite (eq_in_path (@edge_umfiltkE A g p)) //. by apply/allP/in_consE. Qed. @@ -702,15 +719,17 @@ exists (rcons xs z); split=>/=. by move=>w; rewrite mem_rcons inE; case/orP=>[/eqP ->|/Hxs]. Qed. +End ConnectLemmas. + (***********************) (* Connected component *) (***********************) (* part of g reachable from x (avoiding nothing) *) -Definition subconnect g x : pregraph := +Definition subconnect A (g : pregraph A) x : pregraph A := um_filterk (mem (connect pred0 g x)) g. -Lemma edge_subconnect g x y z : +Lemma edge_subconnect A (g : pregraph A) x y z : edge (subconnect g x) y z = [&& y \in connect pred0 g x, z \in connect pred0 g x & @@ -723,7 +742,7 @@ Qed. (* connectivity relation out of x in a connected subcomponent *) (* is that same connectivity relation out of x in the graph *) -Lemma connect_subconnect g x : +Lemma connect_subconnect A (g : pregraph A) x : connect pred0 (subconnect g x) x =i connect pred0 g x. Proof. @@ -748,6 +767,10 @@ Qed. (* Avoidance sets (marking) *) (****************************) +Section MarkingLemmas. +Context {A : Type}. +Implicit Type g : pregraph A. + (* deconstructing connecting path from left *) (* strenghtening avoidance *) Lemma edge_connectX p g (x y : node) : @@ -769,7 +792,7 @@ Qed. (* extending connecting path from left *) (* weakening avoidance *) -Lemma edge_connectXI (p : pred node) (g : pregraph) x y z : +Lemma edge_connectXI (p : pred node) g x y z : x \in dom g -> ~~ p x -> edge g x y -> @@ -838,6 +861,8 @@ rewrite (connect_avoid _ C); apply/connect_in_sub=>z Dz. by rewrite !inE=>/orP [/eqP <-|->//]; rewrite connect0 Dz orbN. Qed. +End MarkingLemmas. + (******************) (* Biconnectivity *) (******************) @@ -846,9 +871,13 @@ Qed. (* y is biconnected from x iff *) (* x and y are mutually connected *) -Definition biconnect p g x : pred node := +Definition biconnect A p (g : pregraph A) x : pred node := [pred y | (x \in connect p g y) && (y \in connect p g x)]. +Section BiconnectLemmas. +Context {A : Type}. +Implicit Type g : pregraph A. + Lemma biconnect0 p g x : x \in dom g -> ~~ p x -> @@ -879,10 +908,16 @@ Lemma biconnect_eq g (p1 p2 : pred node) (x : node) : biconnect p1 g x =i biconnect p2 g x. Proof. by move=>S y; rewrite !inE !(connect_eq g _ S). Qed. +End BiconnectLemmas. + (**********) (* Cycles *) (**********) +Section CycleLemmas. +Context {A : Type}. +Implicit Type g : pregraph A. + (* elements in a cycle are interconnected *) (* avoiding everything outside the cycle *) @@ -954,18 +989,24 @@ move=>Nxy; apply/(iffP (biconnect_cycle2P pred0 g Nxy)). by case=>xs H1 H2; exists xs. Qed. +End CycleLemmas. + (**************) (* Acyclicity *) (**************) (* graph is preacyclic if only self-loops are biconnected *) -Definition preacyclic g := +Definition preacyclic A (g : pregraph A) := all2rel (fun x y => (y \in biconnect pred0 g x) ==> (x == y)) (dom g). (* graph is acyclic if it doesn't even have self-loops *) -Definition acyclic g := +Definition acyclic A (g : pregraph A) := preacyclic g && all (fun x => ~~ edge g x x) (dom g). +Section AcyclicityLemmas. +Context {A : Type}. +Implicit Type g : pregraph A. + Lemma preacyclicP g : reflect {in dom g &, forall x y, y \in biconnect pred0 g x -> x = y} (preacyclic g). @@ -1041,12 +1082,18 @@ have [Dx|Nx] := boolP (x \in dom g); last by rewrite linksND. by apply: contra (H _ Dx)=>Lx; rewrite edge_links Dx Lx. Qed. +End AcyclicityLemmas. + (***************) (* N-pregraphs *) (***************) (* getting the n-th link *) -Definition get_nth g x := nth null (links g x). +Definition get_nth A (g : pregraph A) x := nth null (links g x). + +Section NpregraphLemmas. +Context {A : Type}. +Implicit Type g : pregraph A. Lemma getnth_mem0 g x n : (get_nth g x n == null) || @@ -1065,15 +1112,16 @@ Proof. by move=>H; move: (getnth_mem0 g x n); rewrite (negbTE H). Qed. (* n-pregraph is pregraph with global bound n *) (* for the number of links of a node *) -Definition n_pregraph_axiom (n : nat) (g : pregraph) := +Definition n_pregraph_axiom (n : nat) g := {in dom g, forall x, size (links g x) = n}. Lemma npregraphP n g : n_pregraph_axiom n g <-> - {in range g, forall xs, size xs = n}. + {in map snd (range g), forall xs, size xs = n}. Proof. split=>H x; last by move=>Dx; apply: H (range_links Dx). -by case/mem_rangeX=>k X; rewrite (In_graph X); apply: H (In_dom X). +case/mem_seqP/Mem_map_inv=>-[v xs][->] /In_rangeX [k X]. +by rewrite (In_graph X); apply: H (In_dom X). Qed. Lemma npregraphUnL n g1 g2 : diff --git a/pcm/unionmap.v b/pcm/unionmap.v index 667873c..6a96302 100644 --- a/pcm/unionmap.v +++ b/pcm/unionmap.v @@ -1338,17 +1338,17 @@ Lemma free_undef k : free undef k = undef :> U. Proof. by rewrite !umEX. Qed. Lemma freeU k1 k2 v f : - free (upd k2 v f) k1 = + free (upd k1 v f) k2 = if k1 == k2 then - if C k2 then free f k1 else undef - else upd k2 v (free f k1). + if C k1 then free f k2 else undef + else upd k1 v (free f k2). Proof. rewrite !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. +- by rewrite H2 !umEX rem_ins eq_sym H1. - by case: ifP H2. -by rewrite !umEX rem_ins H1. +by rewrite !umEX rem_ins eq_sym H1. Qed. Lemma freeF k1 k2 f : @@ -1660,21 +1660,25 @@ by move=>->; rewrite valid_unit andbT; case: ifP=>// _ [->]. Qed. Lemma freePt2 k1 k2 v : - C k2 -> - free (pts k2 v) k1 = if k1 == k2 then Unit else pts k2 v :> U. + C k1 -> + free (pts k1 v) k2 = if k1 == k2 then Unit else pts k1 v :> U. Proof. by move=>N; rewrite ptsU freeU free0 N. Qed. Lemma freePt k v : - C k -> free (pts k v : U) k = Unit. + C k -> + free (pts k v : U) k = Unit. Proof. by move=>N; rewrite freePt2 // eq_refl. Qed. Lemma cancelPt k v1 v2 : - valid (pts k v1 : U) -> pts k v1 = pts k v2 :> U -> v1 = v2. + valid (pts k v1 : U) -> + pts k v1 = pts k v2 :> U -> + v1 = v2. Proof. by rewrite validPt !ptsU; apply: upd_inj. Qed. Lemma cancelPt2 k1 k2 v1 v2 : valid (pts k1 v1 : U) -> - pts k1 v1 = pts k2 v2 :> U -> (k1, v1) = (k2, v2). + pts k1 v1 = pts k2 v2 :> U -> + (k1, v1) = (k2, v2). Proof. move=>H E; have : dom (pts k1 v1 : U) = dom (pts k2 v2 : U) by rewrite E. rewrite !domPtK -(validPt _ v1) -(validPt _ v2) -E H. @@ -1698,7 +1702,8 @@ by case E: decP=>[a|a] /=; [rewrite a | case: ifP]. Qed. Lemma assocsUnPt f k v : - valid (f \+ pts k v) -> all (ord^~ k) (dom f) -> + valid (f \+ pts k v) -> + all (ord^~ k) (dom f) -> assocs (f \+ pts k v) = rcons (assocs f) (k, v). Proof. rewrite /ptsx !umEX /UM.assocs/UM.union/UM.pts/UM.dom/supp /=. @@ -1707,7 +1712,8 @@ by case: allP=>//; case: g H2=>// s1 H2 H3 H4 _; apply: first_ins' H4. Qed. Lemma assocsPtUn f k v : - valid (pts k v \+ f) -> all (ord k) (dom f) -> + valid (pts k v \+ f) -> + all (ord k) (dom f) -> assocs (pts k v \+ f) = (k, v) :: (assocs f). Proof. rewrite /ptsx !umEX /UM.assocs/UM.union/UM.pts/UM.dom/supp /=. @@ -1870,36 +1876,37 @@ Proof. by rewrite joinC; apply: findPtUn2. Qed. (* free *) Lemma freePtUn2 k1 k2 v f : - valid (pts k2 v \+ f) -> - free (pts k2 v \+ f) k1 = - if k1 == k2 then f else pts k2 v \+ free f k1. + valid (pts k1 v \+ f) -> + free (pts k1 v \+ f) k2 = + if k1 == k2 then f else pts k1 v \+ free f k2. Proof. -move=>V'; rewrite freeUn domPtUn inE V' /= eq_sym. +move=>V'; rewrite freeUn domPtUn inE V' /=. rewrite ptsU freeU (validPtUn_cond V') free0. -case: eqP=>/= E; first by rewrite E unitL (dom_free (validPtUnD V')). +case: (k1 =P k2)=>/= E. +- by rewrite -E unitL (dom_free (validPtUnD V')). by case: ifP=>// N; rewrite dom_free // N. Qed. -Lemma freeUnPt2 k1 k2 v f : - valid (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 (pts k v \+ f) k = 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 (f \+ pts k v) k = f. + valid (f \+ pts k v) -> + free (f \+ pts k v) k = f. Proof. by rewrite joinC; apply: freePtUn. Qed. -Lemma freeX k v f1 f2 : valid f1 -> f1 = pts k v \+ f2 -> f2 = free f1 k. +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. +Lemma updPtUn k v1 v2 f : + upd k v1 (pts k v2 \+ f) = pts k v1 \+ f. Proof. case V1 : (valid (pts k v2 \+ f)). - by rewrite updUnL domPt inE eq_refl updPt (validPtUn_cond V1). @@ -1923,7 +1930,9 @@ apply: find_eta=>//; first by rewrite validU H W. by move=>k'; rewrite findU H W findPtUn2 // findF; case: eqP. Qed. -Lemma upd_eta2 k v f : k \notin dom f -> upd k v f = pts k v \+ f. +Lemma upd_eta2 k v f : + k \notin dom f -> + upd k v f = pts k v \+ f. Proof. by move=>D; rewrite upd_eta freeND. Qed. (* um_unitb *) @@ -1986,7 +1995,7 @@ move: (V1); rewrite E => V2. 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. + by rewrite ptsU freeU X free0 -ptsU freePtUn. suff E'' : free f1 k2 = free f2 k1. - by rewrite -E''; rewrite E' joinCA in E; case/(cancel V1): E. move: (erefl (free f2 k1)). From bea457052987b30d8979819cefccb36304742eec Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Wed, 4 Sep 2024 19:18:56 +0200 Subject: [PATCH 73/93] blah --- examples/graph.v | 13 +++++++++++-- 1 file changed, 11 insertions(+), 2 deletions(-) diff --git a/examples/graph.v b/examples/graph.v index 5780ec8..88bd857 100644 --- a/examples/graph.v +++ b/examples/graph.v @@ -1129,7 +1129,8 @@ Lemma npregraphUnL n g1 g2 : n_pregraph_axiom n (g1 \+ g2) -> n_pregraph_axiom n g1. Proof. -by rewrite !npregraphP=>V H x X; apply: H; rewrite rangeUn inE V X. +rewrite !npregraphP=>V H _ /mem_seqP/Mem_map_inv [[v xs]][-> X]. +by apply: H; apply/mem_seqP/Mem_map/In_rangeL/X. Qed. Lemma npregraphUnR n g1 g2 : @@ -1141,7 +1142,15 @@ Proof. by rewrite joinC; apply: npregraphUnL. Qed. Lemma npregraphF n g x : n_pregraph_axiom n g -> n_pregraph_axiom n (free g x). -Proof. by rewrite !npregraphP=>H z /rangeF; apply: H. Qed. +Proof. +rewrite !npregraphP=>H z /mem_seqP/Mem_map_inv [[v xs]][-> X]. +apply/H/mem_seqP/Mem_map. + + + + + +by rewrite !npregraphP=>H z /rangeF; apply: H. Qed. Lemma links_nth n g x : n_pregraph_axiom n g -> From 93cd997783ff2e2813686705869fa639bccb38bf Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Wed, 4 Sep 2024 22:01:23 +0200 Subject: [PATCH 74/93] ibalh --- examples/graph.v | 1 + 1 file changed, 1 insertion(+) diff --git a/examples/graph.v b/examples/graph.v index 88bd857..0ac5373 100644 --- a/examples/graph.v +++ b/examples/graph.v @@ -1146,6 +1146,7 @@ Proof. rewrite !npregraphP=>H z /mem_seqP/Mem_map_inv [[v xs]][-> X]. apply/H/mem_seqP/Mem_map. +UP TO HERE From a45938bc0efc853f3bfbce22d6d5154b830641f3 Mon Sep 17 00:00:00 2001 From: Aleks Nanevski Date: Thu, 5 Sep 2024 05:15:50 +0200 Subject: [PATCH 75/93] blah --- examples/graph.v | 99 ++++++++--- examples/schorr.v | 420 +++++++++++++++++++++++++++++++++++++++------- pcm/unionmap.v | 75 ++++++++- 3 files changed, 511 insertions(+), 83 deletions(-) diff --git a/examples/graph.v b/examples/graph.v index 0ac5373..c16352f 100644 --- a/examples/graph.v +++ b/examples/graph.v @@ -68,7 +68,7 @@ Lemma pregraph_is_umc : pg_undefb pg_pts pg_from pg_to. Proof. by split=>//; split=>[|[]]. Qed. -HB.instance Definition blah := +HB.instance Definition _ := isUnion_map.Build node nat_pred (A * seq node)%type (pregraph A) pregraph_is_umc. End PregraphUMC. @@ -272,6 +272,8 @@ Proof. by move=>x y; rewrite !inE => X Y; rewrite edge_umfiltk /= X Y. Qed. End PregraphLemmas. +Prenex Implicits In_graph. + (***********************) (* Depth-first search *) (***********************) @@ -1111,11 +1113,19 @@ Proof. by move=>H; move: (getnth_mem0 g x n); rewrite (negbTE H). Qed. (* n-pregraph is pregraph with global bound n *) (* for the number of links of a node *) - +Definition n_pregraphb (n : nat) g := + all (fun x => size (links g x) == n) (dom g). Definition n_pregraph_axiom (n : nat) g := {in dom g, forall x, size (links g x) = n}. -Lemma npregraphP n g : +Lemma npregraphP (n : nat) g : + reflect (n_pregraph_axiom n g) (n_pregraphb n g). +Proof. +apply: (iffP allP)=>H; first by move=>x /H /eqP. +by move=>x D; apply/eqP/H. +Qed. + +Lemma npregraphE n g : n_pregraph_axiom n g <-> {in map snd (range g), forall xs, size xs = n}. Proof. @@ -1124,12 +1134,23 @@ case/mem_seqP/Mem_map_inv=>-[v xs][->] /In_rangeX [k X]. by rewrite (In_graph X); apply: H (In_dom X). Qed. +Lemma npregraphUn n g1 g2 : + n_pregraph_axiom n g1 -> + n_pregraph_axiom n g2 -> + n_pregraph_axiom n (g1 \+ g2). +Proof. +move=>H1 H2 z; rewrite domUn inE; case/andP=>V. +case/orP=>D. +- by rewrite /links findUnL // D H1. +by rewrite /links findUnR // D H2. +Qed. + Lemma npregraphUnL n g1 g2 : valid (g1 \+ g2) -> n_pregraph_axiom n (g1 \+ g2) -> n_pregraph_axiom n g1. Proof. -rewrite !npregraphP=>V H _ /mem_seqP/Mem_map_inv [[v xs]][-> X]. +rewrite !npregraphE=>V H _ /mem_seqP/Mem_map_inv [[v xs]][-> X]. by apply: H; apply/mem_seqP/Mem_map/In_rangeL/X. Qed. @@ -1143,15 +1164,9 @@ Lemma npregraphF n g x : n_pregraph_axiom n g -> n_pregraph_axiom n (free g x). Proof. -rewrite !npregraphP=>H z /mem_seqP/Mem_map_inv [[v xs]][-> X]. -apply/H/mem_seqP/Mem_map. - -UP TO HERE - - - - -by rewrite !npregraphP=>H z /rangeF; apply: H. Qed. +rewrite !npregraphE=>H z /mem_seqP/Mem_map_inv [[v xs]][->]. +by case/In_rangeF=>k' _ /In_range X; apply/H/mem_seqP/Mem_map. +Qed. Lemma links_nth n g x : n_pregraph_axiom n g -> @@ -1163,6 +1178,30 @@ move=>H Dx; apply/(eq_from_nth (x0:=null))=>[|i Hi]. by rewrite map_nth_iota0 ?H // -(H _ Dx) take_size. Qed. +(* n_pregraph gives rise to seprel *) + +Definition n_pregraph n g1 g2 := + n_pregraphb n g1 && n_pregraphb n g2. + +Lemma npregraph_is_sep n : seprel_axiom (@n_pregraph n). +Proof. +split=>[|x y V|x y V|x y z V]; rewrite /n_pregraph. +- by rewrite /n_pregraphb dom0. +- by rewrite andbC. +- by rewrite /n_pregraphb dom0; case/andP=>->. +case/andP=>Hx Hy /andP [_ Hz]; rewrite Hx Hy Hz /=. +apply/allP=>w; rewrite domUn inE (validX V) /=. +case/orP=>Dw. +- by rewrite /links findUnL ?(validX V) // Dw (allP Hy). +by rewrite /links findUnR ?(validX V) // Dw (allP Hz). +Qed. + +HB.instance Definition _ n := + isSeprel.Build (pregraph A) (n_pregraph n) (npregraph_is_sep n). + +End NpregraphLemmas. + + (**********) (**********) (* Graphs *) @@ -1170,22 +1209,26 @@ Qed. (**********) (* x is in_node if it's in the graph or is null *) -Definition in_node (g : pregraph) (x : node) := +Definition in_node A (g : pregraph A) (x : node) := (x == null) || (x \in dom g). (* x is out-node if no edge goes into it *) -Definition out_node (g : pregraph) (x : node) := - {in range g, forall xs, x \notin xs}. +Definition out_node A (g : pregraph A) (x : node) := + {in map snd (range g), forall xs, x \notin xs}. (* pregraph is graph if *) (* all nodes in all adjacency lists are in-nodes *) -Definition graph_axiom (g : pregraph) := - forall xs x, xs \in range g -> x \in xs -> in_node g x. +Definition graph_axiom A (g : pregraph A) := + forall xs x, xs \in map snd (range g) -> x \in xs -> in_node g x. -HB.mixin Record isGraph (g : pregraph) := { +HB.mixin Record isGraph A (g : pregraph A) := { graph_subproof : graph_axiom g}. #[short(type=graph)] -HB.structure Definition Graph := {g of isGraph g }. +HB.structure Definition Graph A := {g of isGraph A g}. + +Section GraphLemmas. +Context {A : Type}. +Implicit Type g : pregraph A. (* removing out-node causes no dangling edges; *) (* hence preserves graph axiom *) @@ -1194,8 +1237,9 @@ Lemma graphF g x : graph_axiom g -> graph_axiom (free g x). Proof. -move=>Hna Ha xs q /rangeF Hs Hq. -move: (Ha xs q Hs Hq) (Hna _ Hs)=>{}Hs. +move=>Hna Ha xs q /mem_seqP/MapP [[v vs]] /In_rangeF [k'] N +/In_range/(Mem_map snd)/mem_seqP Hs ->{xs} Hq. +move: (Ha vs q Hs Hq) (Hna _ Hs)=>{}Hs. rewrite /in_node domF inE in Hs *. by case: (x =P q) Hq=>//= ->->. Qed. @@ -1209,9 +1253,12 @@ Proof. have E : g = subconnect g x \+ um_filterk (negb \o connect pred0 g x) g by apply: umfilt_predC. move=>V Ha xs /= n Hxs Hn; have {}Dn : in_node g n. -- by apply: Ha Hn; rewrite E rangeUn inE -E V Hxs. -case/mem_rangeX: Hxs=>k /In_umfiltX [/= Ck] /In_graph Hf. -rewrite /in_node in Dn *; case/boolP: (n == null) Dn=>//= Hnn Dn. +- case/mem_seqP/MapP: Hxs=>-[a b Hxs /= H]. + apply: Ha Hn; apply/mem_seqP/MapP; exists (a, b)=>//. + by rewrite E; apply/In_rangeL/Hxs; rewrite -E. +case/mem_seqP/MapP: Hxs=>-[v vs] /In_rangeX [k] /In_umfiltX [/= Ck]. +move/In_graph=>->{vs} Hf; rewrite /in_node in Dn *. +case/boolP: (n == null) Dn=>//= Hnn Dn. case: (connectD Ck)=>_ Dk. rewrite /subconnect dom_umfiltk inE /= Dn andbT. apply: connect_trans Ck _; apply/connectP; split=>//. @@ -1243,3 +1290,5 @@ case: (ltnP n (size (links g x)))=>Hm; last first. by move/(graph_links (x:=x)); apply; rewrite mem_nth. Qed. +End GraphLemmas. + diff --git a/examples/schorr.v b/examples/schorr.v index ff49d98..ea66f26 100644 --- a/examples/schorr.v +++ b/examples/schorr.v @@ -22,7 +22,7 @@ From htt Require Import options model heapauto graph. (* helper lemmas *) (*****************) -Lemma eq_connect_aux (g1 g2 : pregraph) p: +Lemma eq_connect_aux A (g1 g2 : pregraph A) p: valid (g1 \+ g2) -> {subset dom g2 <= predC p} -> um_filterk p (g1 \+ g2) = um_filterk p g1. @@ -31,7 +31,7 @@ move=>V /(umfiltk_subdom0 p (validR V)) S. by rewrite umfiltUn // S unitR. Qed. -Lemma connectUn_eq (g g1 g2 : pregraph) (p : pred node) x : +Lemma connectUn_eq A (g g1 g2 : pregraph A) (p : pred node) x : valid (g \+ g1) -> valid (g \+ g2) -> {subset dom g1 <= p} -> @@ -45,105 +45,411 @@ rewrite -[RHS]connect_umfiltk eq_connect_aux //. by move=>y /S2; rewrite !inE negbK. Qed. +(*****************) +(* Binary graphs *) +(*****************) + +(* short notation for left/right node of x *) +Notation lft g x := (get_nth g x 0). +Notation rgh g x := (get_nth g x 1). + +Definition contents {A} (g : pregraph A) : nmap A := mapv fst g. + +HB.instance Definition _ A := + OmapFun.copy (@contents A) (@contents A). + +Lemma In_contentsX A (g : pregraph A) x v : + (x, v) \In contents g <-> + exists lks, (x, (v, lks)) \In g. +Proof. +rewrite In_omfX; split; last by case=>lks H; exists (v, lks). +by case; case=>w lks /= H [<-{v}]; exists lks. +Qed. + +Lemma In_contents A (g : pregraph A) x v lks : + (x, (v, lks)) \In g -> + (x, v) \In contents g. +Proof. by rewrite In_contentsX; exists lks. Qed. + +Definition links {A} (g : pregraph A) : nmap (seq node) := mapv snd g. + +HB.instance Definition _ A := + OmapFun.copy (@links A) (@links A). + +Lemma In_linksX A (g : pregraph A) x lks : + (x, lks) \In links g <-> + exists v, (x, (v, lks)) \In g. +Proof. +rewrite In_omfX; split; last by case=>v H; exists (v, lks). +by case; case=>v lks' /= H [<-{lks}]; exists v. +Qed. + +Lemma In_links A (g : pregraph A) x v lks : + (x, (v, lks)) \In g -> + (x, lks) \In links g. +Proof. by rewrite In_linksX; exists v. Qed. + +(* laying binary pregraph onto heap *) +Definition lay2_k {A} (v : A * seq node) : dynamic id := + idyn (v.1, (nth null v.2 0, nth null v.2 1)). +Definition lay2 {A} (g : pregraph A) : heap := mapv lay2_k g. + +HB.instance Definition _ A := + OmapFun.copy (@lay2 A) (@lay2 A). + + +Canonical heap_PredType : PredType (nat * dynamic id) := + um_PredType heap. +Coercion Pred_of_nmap (x : heap) : {Pred _} := + [eta Mem_UmMap x]. + +Lemma In_layX A (g : pregraph A) x (v : A) lft rgh : + (x, idyn (v, (lft, rgh))) \In lay2 g <-> + exists lks, + [/\ lft = nth null lks 0, + rgh = nth null lks 1 & + (x, (v, lks)) \In g]. +Proof. +rewrite In_omfX; split; last first. +- by case=>lks [->-> H]; exists (v, lks). +case=>-[w lks] H /Some_inj/inj_pair2 /= [<-{v} <-<-]. +by exists lks. +Qed. + +Lemma In_lay A (g : pregraph A) x (v : A) lks : + (x, (v, lks)) \In g -> + (x, idyn (v, (nth null lks 0, nth null lks 1))) \In lay2 g. +Proof. by rewrite In_layX=>H; exists lks. Qed. + +Lemma In_lay2X A (g : pregraph A) x (v : A) lft rgh : + n_pregraph_axiom 2 g -> + (x, idyn (v, (lft, rgh))) \In lay2 g <-> + (x, (v, [:: lft; rgh])) \In g. +Proof. +move=>H; split; last first. +- by move=>X; apply/In_layX; exists [:: lft; rgh]. +case/In_layX=>lks [L R X]. +rewrite -(_ : lks = [:: lft; rgh]) //. +rewrite (In_graph X) (links_nth H (In_dom X)) /=. +by rewrite /get_nth -(In_graph X) -L -R. +Qed. + +Lemma In_lay2 A (g : pregraph A) x (v : A) lft rgh : + n_pregraph_axiom 2 g -> + (x, (v, [:: lft; rgh])) \In g -> + (x, idyn (v, (lft, rgh))) \In lay2 g. +Proof. by move=>H /In_lay2X-/(_ H). Qed. + + + +(* updating binary pregraph at node x *) + +(* pass v = None to just update links *) +Definition bupd A (g : pregraph A) x (v : option A) (lft rgh : node) := + if find x (contents g) is Some v' then + if v is Some w then upd x (w, [:: lft; rgh]) g + else upd x (v', [:: lft; rgh]) g + else undef. + +Lemma bupd_is_binary A (g : pregraph A) x v lft rgh : + n_pregraph_axiom 2 g -> + n_pregraph_axiom 2 (@bupd A g x v lft rgh). +Proof. +move=>H z; rewrite /bupd find_omf /omfx /=. +case: (find x g)=>[[m lnk]|//]. +case: v=>[w|]; rewrite domU inE /graph.links findU; +case: (x =P 0)=>//= /eqP Nx; case: (z =P x)=>[_ V|_ Dz]. +- by rewrite V. +- by rewrite H. +- by rewrite V. +by rewrite H. +Qed. + +(* updating just contents *) +Notation updC g x v := (bupd g x (Some v) (lft g x) (rgh g x)). +(* updating just left link *) +Notation updL g x l := (bupd g x (contents g x) l (rgh g x)). +(* updating just right link *) +Notation updR g x r := (bupd g x (contents g x) (lft g x) r). +(* updating just contents and left link *) +Notation updCL g x v l := (bupd g x (Some v) l (rgh g x)). +(* updating just contents and right link *) +Notation updCR g x v r := (bupd g x (Some v) (lft g x) r). + +Lemma glE A (g : pregraph A) (x : node) : + n_pregraph_axiom 2 g -> + x \in dom g -> + graph.links g x = [:: lft g x; rgh g x]. +Proof. by move=>H D; rewrite (links_nth H D). Qed. + +Lemma glD A (g : pregraph A) (x : node) : + n_pregraph_axiom 2 g -> + x \in dom g -> + (x, [:: lft g x; rgh g x]) \In links g. +Proof. +move=>H /[dup]/In_graphX R /(glE H) <-. +by apply/In_linksX. +Qed. + +Lemma getU A (g : pregraph A) (x : node) v lf rg i : + get_nth (bupd g x v lf rg) x i = + if x \notin dom g then null + else nth null [:: lf; rg] i. +Proof. +rewrite /get_nth/graph.links/bupd find_omf/omfx /=. +case: (dom_find x)=>[|w /In_find H _]; first by rewrite nth_nil. +by case: v=>[v|] /=; rewrite findU eqxx /= (In_cond H) (In_valid H). +Qed. + (****************) (* Schorr-Waite *) (****************) -(* short notation for left/right child of x *) -Notation gl g x := (get_nth g x 0). -Notation gr g x := (get_nth g x 1). - (* type of markings *) -Inductive mark := L | R | LR. +Inductive mark := U | L | R | LR. (* decidable equality on marks *) Definition eq_mark l1 l2 : bool := - if (l1, l2) is ((L, L)|(R, R)|(LR, LR)) then true else false. + if (l1, l2) is ((U,U)|(L, L)|(R, R)|(LR, LR)) then true else false. Lemma eq_markP : Equality.axiom eq_mark. Proof. by case; case=>//=; constructor. Qed. HB.instance Definition _ := hasDecEq.Build mark eq_markP. +Notation pregraph := (pregraph mark). + +Definition marked (g : pregraph) : pregraph := + um_filterv (fun v => v.1 != U) g. +HB.instance Definition _ := OmapFun.copy marked marked. + + (* marking of children *) (* given marking map m, x is m-child of y iff: *) (* - m y = L and x is left child of y *) (* - m y = R and x is right child of y *) -Definition ch (m : nmap mark) (g : pregraph) (x y : nat) := - [|| (find y m == Some L) && (gl g y == x) | - (find y m == Some R) && (gr g y == x)]. - -Lemma chP (m : nmap mark) g x y : - reflect [\/ (y, L) \In m /\ gl g y = x | - (y, R) \In m /\ gr g y = x] - (ch m g x y). +Definition ch (g : pregraph) (x y : nat) := + [|| (find y (contents g) == Some L) && (lft g y == x) | + (find y (contents g) == Some R) && (rgh g y == x)]. + +Lemma chP (g : pregraph) x y : + reflect [\/ (y, L) \In contents g /\ lft g y = x | + (y, R) \In contents g /\ rgh g y = x] + (ch g x y). Proof. rewrite /ch; case: orP=>H; constructor. - by case: H=>H; [left|right]; case/andP: H=>/eqP/In_find H /eqP. by case=>[][/In_find/eqP M /eqP G]; apply: H; [left|right]; apply/andP. Qed. -Lemma chN0 m g x y : - ch m g x y -> +Lemma chN0 g x y : + ch g x y -> y != null. Proof. by case/chP=>[][/In_dom/dom_cond]. Qed. -Lemma ch_fun m g a b x : - ch m g a x -> - ch m g b x -> +Lemma ch_fun g a b x : + ch g a x -> + ch g b x -> a = b. Proof. by case/chP=>[][H <-] /chP [][/(In_fun H) {}H <-]. Qed. -Lemma ch_path m g s x : +Lemma ch_path g s x : x \in s -> - path (ch m g) null s -> - exists y, y \in belast null s /\ ch m g y x. + path (ch g) null s -> + exists y, y \in belast null s /\ ch g y x. Proof. exact: path_prev. Qed. -Lemma ch_path_uniq m g s : - path (ch m g) null s -> +Lemma ch_path_uniq g s : + path (ch g) null s -> uniq (null :: s). Proof. by apply: path_uniq; [apply: chN0|apply: ch_fun]. Qed. -Definition upd_edge (m : nmap mark) g x y : seq node := - if find x m is Some L then [:: y; gr g x] else [:: gl g x; y]. -Fixpoint rev_graph' m g ps t : pregraph := +Definition upd_edge g x y : mark * seq node := + match find x (contents g) with + | Some L => (L, [:: y; rgh g x]) + | _ => (R, [:: lft g x; y]) + end. + +Fixpoint rev_graph' (g : pregraph) (ps : seq node) t : pregraph := if ps is p :: ps then - rev_graph' m (free g p) ps p \+ pts p (upd_edge m g p t) + rev_graph' (free g p) ps p \+ pts p (upd_edge g p t) else g. -Definition rev_graph m g s t := rev_graph' m g (rev s) t. +Definition rev_graph g s t := rev_graph' g (rev s) t. (* layout of graph+marking as heap *) -Definition lay (m : nmap mark) (g : pregraph) : heap := + + +(* +Definition lay (g : pregraph) : heap := \big[join/Unit]_(x <- dom g) (x :-> (gl g x, gr g x, odflt L (find x m))). -(* reach m g s t holds iff *) +Lemma bigF (A : eqType) (xs : seq A) (f : A -> heap) k : + valid (\big[join/Unit]_(i <- xs) f i) -> + free (\big[join/Unit]_(i <- xs) f i) k = + \big[join/Unit]_(i <- xs) free (f i) k. +Proof. +elim: xs=>[|x xs IH]. +- by rewrite !big_nil free0. +rewrite !big_cons => V. +rewrite freeUn. +case: ifP=>D. +- rewrite IH //. rewrite (validR V). by []. +rewrite -IH; last by rewrite (validX V). +rewrite !freeND //. +- apply: contraFN D=>D. + rewrite domUn inE D orbT andbT. + by []. +apply: contraFN D=>D. +by rewrite domUn inE (validX V) D. +Qed. +*) + + +Lemma In_layE1 (g : pregraph) x pl pr c : + n_pregraph_axiom 2 g -> + find x (lay2 g) = Some (idyn (c, (pl, pr))) -> + (x, (c, [:: pl; pr])) \In g. +Proof. by move=>H /In_find /In_lay2X-/(_ H). Qed. + +Lemma In_layE2 (g : pregraph) x pl pr c : + n_pregraph_axiom 2 g -> + (x, (c, [:: pl; pr])) \In g -> + find x (lay2 g) = Some (idyn (c, (pl, pr))). +Proof. by move=>H /In_lay2X-/(_ H)/In_find. Qed. + +Lemma In_layX2 (g : pregraph) x c pl pr : + n_pregraph_axiom 2 g -> + (x, (c, [:: pl; pr])) \In g -> + exists h, lay2 g = x :-> (c, (pl, pr)) \+ h. +Proof. by move=>H Gx; eexists _; apply/heap_eta2/In_layE2. Qed. + +Lemma In_layX3 (g : pregraph) x l r c : + lay2 (bupd g x (Some c) l r) = + if x \in dom g then + upd x (idyn (c, (l, r))) (lay2 g) + else undef. +Proof. +rewrite /lay2/bupd find_omf/omfx /=. +case: (dom_find x)=>[|v]; first by rewrite omap_undef. +move/In_find=>E _; rewrite (upd_eta x). +rewrite omapPtUn -(upd_eta x) validU. +rewrite (In_cond E) (In_valid E) /=. +rewrite (upd_eta x) !omap_free !omap_omap. +congr (_ \+ _); apply/eq_in_omap. +by case=>k w /= H; case: (k =P x). +Qed. + +Lemma In_layX4 (g : pregraph) x l r c : + lay2 (x &-> (c, [:: l; r]) \+ g) = + x :-> (c, (l, r)) \+ lay2 g. +Proof. +rewrite omfPtUn /=; case: ifP=>// V; set j := (_ \+ _). +case: (normalP j)=>[->//|]. +rewrite !validPtUn valid_omap dom_omf_some // in V *. +by rewrite V. +Qed. + +(* reach g s t holds iff *) (* each unmarked node x in g is reachable through unmarked nodes *) (* - either starting from t, or *) (* - starting from a right child of some node in s *) -Definition reach (m : nmap mark) (g : pregraph) (s : seq node) t := + +Definition reach (g : pregraph) (s : seq node) (t : node) := forall x, - (* x is node in g *) - x \in dom g -> - (* x is unmarked *) - x \notin dom m -> - (* x is reachable from t avoiding marked nodes *) - (* Aleks note: if unmarked nodes are represented by a marking U *) - (* then this conjunct will have to change, as it currently assumes *) - (* that unmarked nodes are just those that aren't in m *) - x \in connect (mem (dom m)) g t \/ - (* or x is reachable from some right child of a node y in s *) + (* if x is unmarked node in g *) + (x, U) \In contents g -> + (* then x is reachable from t avoiding marked nodes *) + x \in connect (mem (dom (marked g))) g t \/ + (* or x is reachable from some right child of a node in s *) (* avoiding marked nodes *) - exists y, y \in s /\ x \in connect (mem (dom m)) g (gr g y). - -Definition shape g0 r p t := - fun h:heap => (exists s g m, - dom g = dom h /\ - h = lay m g /\ - path (ch m g) null s /\ - rev_graph m g s t = g0 /\ - reach m g s t /\ - {subset dom (um_filterv (predU (pred1 L) (pred1 LR)) m) <= s} /\ - p = last null s /\ - r = head t s /\ - graph_axiom g). + exists2 y, y \in s & + x \in connect (mem (dom (marked g))) g (rgh g y). + +UP TO HERE + +Definition shape (g0 g : pregraph) (m : nmap mark) (r p t : node) := + fun h => exists stack, + [/\ h = lay m g, + {subset dom m <= dom g}, + path (ch m g) null stack, + rev_graph m g stack t = g0, + reach m g stack t, + {subset dom (um_filterv (predU (pred1 L) (pred1 LR)) m) <= stack}, + p = last null stack, + r = head t stack, + n_pregraph_axiom 2 g & + graph_axiom g]. + +Program Definition pop (p t : ptr): + STsep {g0 g m r} (fun h => + [/\ shape g0 g m r p t h, + (p, R) \In m & + t \in dom m \/ t = null], + [vfun res h => + let: g' := upd p [:: gl g p; t] g in + let: m' := upd p LR m in + shape g0 g' m' r res.1 res.2 h /\ + res = (gr g p, p)]) := + Do (q <-- read (A:=node*node*mark) p; + p ::= (q.1.1, t, LR);; + ret (q.1.2, p)). +Next Obligation. +move=>p t [g0][g][m][r][/= h][H Pm D]. +case: H=>stack [-> H1 H2 H3 H4 H5 H6 H7 H8 H9]. +move: (H1 p (In_dom Pm))=>Pg. +case/(glD H8)/(In_layX2 H8 H1): (Pg)=>h' E. +rewrite E. +do 3!step; split=>//; exists stack; split. +- rewrite In_layX3. + - by rewrite upd_eta E freePtUn // (validX H). + - by rewrite (dom_valid Pg). + - by rewrite (In_valid Pm). + - by rewrite (In_cond Pm). + +Search upd. + + +admit. +Admitted. + +Program Definition swing (p t : ptr): + STsep {g0 g m r} (fun h => + [/\ shape g0 g m r p t h, + (p, L) \In m & + t \in dom m \/ t = null], + [vfun res h => + let: g' := upd p [:: t; gl g p] g in + let: m' := upd p R m in + shape g0 g' m' r res.1 res.2 h /\ + res = (t, gl g t) ]) := + Do (q <-- read (A:=node*node*mark) p; + p ::= (t, q.1.1, R);; + ret (p, q.1.2)). +Next Obligation. +move=>p t [g0][g][m][r][/= h][H P D]. +case: H=>stack [-> H1 H2 H3 H4 H5 H6 H7 H8 H9]. +move: (H1 p (In_dom P))=>/(glD H8) /(In_layX2 H8 H1) [h' ->]. +do 3!step; split=>//. +exists stack; split =>//. +admit. +Admitted. + +Program Definition push (p t : ptr): + STsep {g0 g m r} (fun h => + [/\ shape g0 g m r p t h, + t \in dom g & + t \notin dom m], + [vfun res h => + let: g' := upd t [:: p; gr g t] g in + let: m' := upd t L m in + shape g0 g' m' r res.1 res.2 h /\ + res = (t, gl g t) ]) := + Do (q <-- read (A:=node*node*mark) t; + t ::= (p, q.1.2, L);; + ret (t, q.1.1)). +Next Obligation. +move=>p t [g0][g][m][r][/= h][H P D]. +case: H=>stack [-> H1 H2 H3 H4 H5 H6 H7 H8 H9]. +admit. +Admitted. + diff --git a/pcm/unionmap.v b/pcm/unionmap.v index 6a96302..5d166db 100644 --- a/pcm/unionmap.v +++ b/pcm/unionmap.v @@ -2673,9 +2673,18 @@ Lemma rangeUnPtK k v f : range (f \+ pts k v) = rcons (range f) v. Proof. by move=>W H; rewrite /range assocsUnPt // map_rcons. Qed. +Lemma In_rangeF k v f : + v \In range (free f k) <-> + exists2 k', k' != k & (k', v) \In f. +Proof. +rewrite In_rangeX; split; first by case=>x /InF []; exists x. +case=>k' Nk H; exists k'; apply/InF; split=>//. +by rewrite validF (In_valid H). +Qed. + End Range. -Prenex Implicits In_range_valid In_range In_rangeUn. +Prenex Implicits In_range_valid In_range In_rangeUn In_rangeF. (* decidable versions, when V is eqtype *) @@ -6832,6 +6841,60 @@ Lemma bigInXP (xs : seq I) P a v : exists i, [/\ i \In xs, P i & (a, v) \In f i]. Proof. by case/In_find/big_find_someXP=>x [X1 X2 /In_find]; exists x. Qed. +Lemma bigF (xs : seq I) k : + valid (\big[join/Unit]_(i <- xs) f i) -> + free (\big[join/Unit]_(i <- xs) f i) k = + \big[join/Unit]_(i <- xs) free (f i) k. +Proof. +elim: xs=>[|x xs IH]; first by rewrite !big_nil free0. +rewrite !big_cons => V; rewrite freeUn. +rewrite -IH ?(validR V) //; case: ifP=>// D. +rewrite !freeND //; apply: contraFN D=>D; +by rewrite domUn inE D V //= orbT. +Qed. + +(* if xs is domain of some map *) + +Lemma bigD1FE (k : K) (F : K -> U) (g : U) : + \big[join/Unit]_(i <- dom g) F i = + if k \in dom g then + F k \+ \big[join/Unit]_(i <- dom (free g k)) F i + else \big[join/Unit]_(i <- dom (free g k)) F i. +Proof. +case: ifP=>D; last by rewrite freeND // (negbT D). +rewrite (bigD1_seq k) //= -big_filter (_ : filter _ _ = dom (free g k)) //. +apply/ord_sorted_eq; first by rewrite sorted_filter. +- by rewrite sorted_dom. +by move=>z; rewrite mem_filter domF inE eq_sym; case: (k =P z). +Qed. + +Lemma bigPtUnE (x : K) (a : T) (g : U) (F : K -> U) : + \big[join/Unit]_(i <- dom (pts x a \+ g)) F i = + if valid (pts x a \+ g) then + F x \+ \big[join/Unit]_(i <- dom g) F i + else Unit. +Proof. +rewrite (bigD1FE x) domPtUnE. +case: (normalP (pts x a \+ g))=>[->|V]. +- by rewrite free_undef dom_undef big_nil. +by rewrite freePtUn. +Qed. + +Lemma bigPtUn (x : K) (a : T) (g : U) (F : K -> U) : + valid (pts x a \+ g) -> + \big[join/Unit]_(i <- dom (pts x a \+ g)) F i = + F x \+ \big[join/Unit]_(i <- dom g) F i. +Proof. by move=>V; rewrite bigPtUnE V. Qed. + +Lemma bigDUE (x : K) (a : T) (g : U) (F : K -> U) : + \big[join/Unit]_(i <- dom (upd x a g)) F i = + if C x && valid g then + F x \+ \big[join/Unit]_(i <- dom (free g x)) F i + else Unit. +Proof. +by rewrite (upd_eta x) bigPtUnE validPtUn validF domF inE eqxx andbT. +Qed. + End BigOpsUM. Prenex Implicits big_find_some big_find_someD. @@ -6922,6 +6985,16 @@ Proof. by move=>x; rewrite big_valid_dom_seq. Qed. End BigCatSeqUM. +Lemma bigPt_valid (K : ordType) (C : pred K) T (U : union_map C T) + (g : U) (F : K -> T) : + valid (\big[join/Unit]_(i <- dom g) pts i (F i) : U). +Proof. +rewrite big_valid_seq; apply/andP; split. +- by apply/allP=>z Dz; rewrite validPt (dom_cond Dz). +apply/uniq_big_catEX=>//; split=>[i Di|i j k Di Dj] //. +by rewrite !domPt !inE (dom_cond Di) (dom_cond Dj)=>/eqP -> /eqP. +Qed. + Section OMapBig. Variables (K : ordType) (C : pred K) (T T' : Type). Variables (U : union_map C T) (U' : union_map C T'). From 1c9b614b0b6c8c58c0c3b8e14f4c0d5ffcb1893d Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Thu, 5 Sep 2024 18:47:35 +0200 Subject: [PATCH 76/93] blah --- core/prelude.v | 8 +- examples/graph.v | 177 +++++++++++++++++++- examples/schorr.v | 401 ++++++++++++++-------------------------------- htt/heapauto.v | 2 +- htt/model.v | 6 +- pcm/heap.v | 5 +- pcm/unionmap.v | 27 ++-- 7 files changed, 319 insertions(+), 307 deletions(-) diff --git a/core/prelude.v b/core/prelude.v index 1302659..cc5458a 100644 --- a/core/prelude.v +++ b/core/prelude.v @@ -879,16 +879,18 @@ Definition inhab {I : ifinType} : I := inhab0 card_inhab. Inductive Side := LL | RR. Definition Side_eq x y := match x, y with LL, LL => true | RR, RR => true | _, _ => false end. -Definition nat2Side x := if odd x then LL else RR. -Definition Side2nat x := if x is RR then 0 else 1. -Lemma ssrcanc : ssrfun.cancel Side2nat nat2Side. Proof. by case. Qed. +Coercion nat_of_side x := if x is LL then 0 else 1. +Definition side_of_nat x := if odd x then RR else LL. +Lemma ssrcanc : ssrfun.cancel nat_of_side side_of_nat. Proof. by case. Qed. HB.instance Definition _ : isCountable Side := CanIsCountable ssrcanc. Lemma Side_enumP : Finite.axiom [:: LL; RR]. Proof. by case. Qed. HB.instance Definition _ : isFinite Side := isFinite.Build Side Side_enumP. HB.instance Definition _ := isInhabited.Build Side (inhabits LL). +Definition flip x := if x is LL then RR else LL. + (*************) (* Sequences *) (*************) diff --git a/examples/graph.v b/examples/graph.v index c16352f..7315997 100644 --- a/examples/graph.v +++ b/examples/graph.v @@ -15,7 +15,7 @@ From HB Require Import structures. From Coq Require Import ssreflect ssrfun. From mathcomp Require Import ssrbool eqtype ssrnat seq path. From pcm Require Import options axioms pred prelude seqext. -From pcm Require Import pcm unionmap natmap autopcm automap. +From pcm Require Import heap pcm unionmap natmap autopcm automap. (*************) (*************) @@ -38,7 +38,7 @@ From pcm Require Import pcm unionmap natmap autopcm automap. (* in that the set of nodes is drawn from an infinite set *) (* not from a fixed finite set. *) -Notation node := nat. +Definition node := nat. (* A is the contents/labeling of the nodes *) Record pregraph (A : Type) := Pregraph {pregraph_base : @UM.base node nat_pred (A * seq node)}. @@ -84,9 +84,14 @@ Coercion Pred_of_history A (x : pregraph A) : {Pred _} := Notation "x &-> v" := (ptsT (@pregraph _) x v) (at level 30). -(**************************) -(* Links, children, edges *) -(**************************) +(**********************************) +(* labels, links, children, edges *) +(**********************************) + +(* maps each node to its contents (ie. label) *) +Definition labels {A} (g : pregraph A) : nmap A := mapv fst g. +HB.instance Definition _ A := OmapFun.copy (@labels A) (@labels A). +Definition olabel {A} (g : pregraph A) x := find x (labels g). (* Links of x includes all edges outgoing from x *) (* and may explicitly include nodes that aren't in dom g *) @@ -106,6 +111,21 @@ Section PregraphLemmas. Context {A : Type}. Implicit Type g : pregraph A. +(* marks lemmas *) + +Lemma In_labelsX g x v : + (x, v) \In labels g <-> + exists lks, (x, (v, lks)) \In g. +Proof. +rewrite In_omfX; split; last by case=>lks H; exists (v, lks). +by case; case=>w lks /= H [<-{v}]; exists lks. +Qed. + +Lemma In_labels g x v lks : + (x, (v, lks)) \In g -> + (x, v) \In labels g. +Proof. by rewrite In_labelsX; exists lks. Qed. + (* links lemmas *) Lemma links_undef x : links (undef : pregraph A) x = [::]. @@ -1201,6 +1221,153 @@ HB.instance Definition _ n := End NpregraphLemmas. +(********************) +(********************) +(* Binary pregraphs *) +(********************) +(********************) + +(* notation for left/right node of x *) +Notation sel2 m g x := (get_nth g x m). +Notation lft g x := (sel2 LL g x). +Notation rgh g x := (sel2 RR g x). + +(* updating binary pregraph at node x *) + +(* if v = None, updates links, keeping the label *) +Definition upd2 A (g : pregraph A) x (v : option A) (lft rgh : node) := + if find x (labels g) is Some v' then + if v is Some w then upd x (w, [:: lft; rgh]) g + else upd x (v', [:: lft; rgh]) g + else undef. + +Lemma upd2_is_binary A (g : pregraph A) x v lft rgh : + n_pregraph_axiom 2 g -> + n_pregraph_axiom 2 (@upd2 A g x v lft rgh). +Proof. +move=>H z; rewrite /upd2 find_omf /omfx /=; case: (find x g)=>[[m lnk]|//]. +case: v=>[w|]; rewrite domU inE /graph.links findU; +case: (x =P 0)=>//= /eqP Nx; case: (z =P x)=>[_ V|_ Dz]. +- by rewrite V. +- by rewrite H. +- by rewrite V. +by rewrite H. +Qed. + +(* updating just contents *) +Notation updC g x v := (upd2 g x (Some v) (lft g x) (rgh g x)). +(* updating just left link *) +Notation updL g x l := (upd2 g x None l (rgh g x)). +(* updating just right link *) +Notation updR g x r := (upd2 g x None (lft g x) r). +(* updating just contents and left link *) +Notation updCL g x v l := (upd2 g x (Some v) l (rgh g x)). +(* updating just contents and right link *) +Notation updCR g x v r := (upd2 g x (Some v) (lft g x) r). + +Lemma In_links2 A (g : pregraph A) (x : node) v lks : + n_pregraph_axiom 2 g -> + (x, (v, lks)) \In g -> + lks = [:: lft g x; rgh g x]. +Proof. +move=>H X; rewrite (In_graph X). +by rewrite (links_nth H (In_dom X)). +Qed. + +(* laying binary pregraph onto heap *) + +Definition lay2_k {A} (v : A * seq node) : dynamic id := + idyn (v.1, (nth null v.2 0, nth null v.2 1)). +Definition lay2 {A} (g : pregraph A) : heap := mapv lay2_k g. +HB.instance Definition _ A := OmapFun.copy (@lay2 A) (@lay2 A). + +Lemma In_layX A (g : pregraph A) x (c : A) lft rgh : + (x, idyn (c, (lft, rgh))) \In lay2 g <-> + exists lks, + [/\ lft = nth null lks 0, + rgh = nth null lks 1 & + (x, (c, lks)) \In g]. +Proof. +rewrite In_omfX; split; last first. +- by case=>lks [->-> H]; exists (c, lks). +case=>-[w lks] H /Some_inj/inj_pair2 /= [<-<-<-]. +by exists lks. +Qed. + +Lemma In_lay A (g : pregraph A) x (c : A) lks : + (x, (c, lks)) \In g -> + (x, idyn (c, (nth null lks 0, nth null lks 1))) \In lay2 g. +Proof. by rewrite In_layX=>H; exists lks. Qed. + +Lemma In_lay2X A (g : pregraph A) x (v : A) lft rgh : + n_pregraph_axiom 2 g -> + (x, idyn (v, (lft, rgh))) \In lay2 g <-> + (x, (v, [:: lft; rgh])) \In g. +Proof. +move=>H; split; last first. +- by move=>X; apply/In_layX; exists [:: lft; rgh]. +case/In_layX=>lks [L R X]. +rewrite -(_ : lks = [:: lft; rgh]) //. +rewrite (In_graph X) (links_nth H (In_dom X)) /=. +by rewrite /get_nth -(In_graph X) -L -R. +Qed. + +Lemma In_lay2 A (g : pregraph A) x (c : A) lft rgh : + n_pregraph_axiom 2 g -> + (x, (c, [:: lft; rgh])) \In g -> + (x, idyn (c, (lft, rgh))) \In lay2 g. +Proof. by move=>H /In_lay2X-/(_ H). Qed. + +Lemma lay2_eta A (g : pregraph A) x c pl pr : + n_pregraph_axiom 2 g -> + (x, (c, [:: pl; pr])) \In g -> + exists h, lay2 g = x :-> (c, (pl, pr)) \+ h. +Proof. by move=>H Gx; eexists _; apply/heap_eta2/In_find/In_lay2. Qed. + +(* lay with update that changes label *) +Lemma lay2CU A (g : pregraph A) x l r c : + lay2 (upd2 g x (Some c) l r) = + if x \in dom g then + upd x (idyn (c, (l, r))) (lay2 g) + else undef. +Proof. +rewrite /lay2/upd2 find_omf/omfx /=. +case: (dom_find x)=>[|v]; first by rewrite omap_undef. +move/In_find=>E _; rewrite (upd_eta x). +rewrite omapPtUn -(upd_eta x) validU. +rewrite (In_cond E) (In_valid E) /=. +rewrite (upd_eta x) !omap_free !omap_omap. +congr (_ \+ _); apply/eq_in_omap. +by case=>k w /= H; case: (k =P x). +Qed. + +(* lay with update that keeps label fixed *) +Lemma lay2U A (g : pregraph A) x l r : + lay2 (upd2 g x None l r) = + if find x (labels g) is Some c then + upd x (idyn (c, (l, r))) (lay2 g) + else undef. +Proof. +rewrite /olabel/lay2/upd2 find_omf /=. +case: (dom_find x g)=>[/In_findN D|v /In_find E _]. +- by rewrite omap_undef. +rewrite /omfx /= (upd_eta x). +rewrite omapPtUn -(upd_eta x) validU. +rewrite (In_cond E) (In_valid E) /=. +rewrite (upd_eta x) !omap_free !omap_omap. +congr (_ \+ _); apply/eq_in_omap. +by case=>k w /= H; case: (k =P x). +Qed. + +Lemma lay2PtUn A (g : pregraph A) x l r c : + lay2 (x &-> (c, [:: l; r]) \+ g) = + x :-> (c, (l, r)) \+ lay2 g. +Proof. +rewrite omfPtUn /=; case: ifP=>// V; set j := (_ \+ _). +case: (normalP j)=>[->//|]. +rewrite !validPtUn valid_omap dom_omf_some // in V *. +by rewrite V. +Qed. (**********) (**********) diff --git a/examples/schorr.v b/examples/schorr.v index ea66f26..5f3b1af 100644 --- a/examples/schorr.v +++ b/examples/schorr.v @@ -49,154 +49,12 @@ Qed. (* Binary graphs *) (*****************) -(* short notation for left/right node of x *) -Notation lft g x := (get_nth g x 0). -Notation rgh g x := (get_nth g x 1). - -Definition contents {A} (g : pregraph A) : nmap A := mapv fst g. - -HB.instance Definition _ A := - OmapFun.copy (@contents A) (@contents A). - -Lemma In_contentsX A (g : pregraph A) x v : - (x, v) \In contents g <-> - exists lks, (x, (v, lks)) \In g. -Proof. -rewrite In_omfX; split; last by case=>lks H; exists (v, lks). -by case; case=>w lks /= H [<-{v}]; exists lks. -Qed. - -Lemma In_contents A (g : pregraph A) x v lks : - (x, (v, lks)) \In g -> - (x, v) \In contents g. -Proof. by rewrite In_contentsX; exists lks. Qed. - -Definition links {A} (g : pregraph A) : nmap (seq node) := mapv snd g. - -HB.instance Definition _ A := - OmapFun.copy (@links A) (@links A). - -Lemma In_linksX A (g : pregraph A) x lks : - (x, lks) \In links g <-> - exists v, (x, (v, lks)) \In g. -Proof. -rewrite In_omfX; split; last by case=>v H; exists (v, lks). -by case; case=>v lks' /= H [<-{lks}]; exists v. -Qed. - -Lemma In_links A (g : pregraph A) x v lks : - (x, (v, lks)) \In g -> - (x, lks) \In links g. -Proof. by rewrite In_linksX; exists v. Qed. - -(* laying binary pregraph onto heap *) -Definition lay2_k {A} (v : A * seq node) : dynamic id := - idyn (v.1, (nth null v.2 0, nth null v.2 1)). -Definition lay2 {A} (g : pregraph A) : heap := mapv lay2_k g. - -HB.instance Definition _ A := - OmapFun.copy (@lay2 A) (@lay2 A). - - -Canonical heap_PredType : PredType (nat * dynamic id) := - um_PredType heap. -Coercion Pred_of_nmap (x : heap) : {Pred _} := - [eta Mem_UmMap x]. - -Lemma In_layX A (g : pregraph A) x (v : A) lft rgh : - (x, idyn (v, (lft, rgh))) \In lay2 g <-> - exists lks, - [/\ lft = nth null lks 0, - rgh = nth null lks 1 & - (x, (v, lks)) \In g]. -Proof. -rewrite In_omfX; split; last first. -- by case=>lks [->-> H]; exists (v, lks). -case=>-[w lks] H /Some_inj/inj_pair2 /= [<-{v} <-<-]. -by exists lks. -Qed. - -Lemma In_lay A (g : pregraph A) x (v : A) lks : - (x, (v, lks)) \In g -> - (x, idyn (v, (nth null lks 0, nth null lks 1))) \In lay2 g. -Proof. by rewrite In_layX=>H; exists lks. Qed. - -Lemma In_lay2X A (g : pregraph A) x (v : A) lft rgh : - n_pregraph_axiom 2 g -> - (x, idyn (v, (lft, rgh))) \In lay2 g <-> - (x, (v, [:: lft; rgh])) \In g. -Proof. -move=>H; split; last first. -- by move=>X; apply/In_layX; exists [:: lft; rgh]. -case/In_layX=>lks [L R X]. -rewrite -(_ : lks = [:: lft; rgh]) //. -rewrite (In_graph X) (links_nth H (In_dom X)) /=. -by rewrite /get_nth -(In_graph X) -L -R. -Qed. - -Lemma In_lay2 A (g : pregraph A) x (v : A) lft rgh : - n_pregraph_axiom 2 g -> - (x, (v, [:: lft; rgh])) \In g -> - (x, idyn (v, (lft, rgh))) \In lay2 g. -Proof. by move=>H /In_lay2X-/(_ H). Qed. - - - -(* updating binary pregraph at node x *) - -(* pass v = None to just update links *) -Definition bupd A (g : pregraph A) x (v : option A) (lft rgh : node) := - if find x (contents g) is Some v' then - if v is Some w then upd x (w, [:: lft; rgh]) g - else upd x (v', [:: lft; rgh]) g - else undef. - -Lemma bupd_is_binary A (g : pregraph A) x v lft rgh : - n_pregraph_axiom 2 g -> - n_pregraph_axiom 2 (@bupd A g x v lft rgh). -Proof. -move=>H z; rewrite /bupd find_omf /omfx /=. -case: (find x g)=>[[m lnk]|//]. -case: v=>[w|]; rewrite domU inE /graph.links findU; -case: (x =P 0)=>//= /eqP Nx; case: (z =P x)=>[_ V|_ Dz]. -- by rewrite V. -- by rewrite H. -- by rewrite V. -by rewrite H. -Qed. - -(* updating just contents *) -Notation updC g x v := (bupd g x (Some v) (lft g x) (rgh g x)). -(* updating just left link *) -Notation updL g x l := (bupd g x (contents g x) l (rgh g x)). -(* updating just right link *) -Notation updR g x r := (bupd g x (contents g x) (lft g x) r). -(* updating just contents and left link *) -Notation updCL g x v l := (bupd g x (Some v) l (rgh g x)). -(* updating just contents and right link *) -Notation updCR g x v r := (bupd g x (Some v) (lft g x) r). - -Lemma glE A (g : pregraph A) (x : node) : - n_pregraph_axiom 2 g -> - x \in dom g -> - graph.links g x = [:: lft g x; rgh g x]. -Proof. by move=>H D; rewrite (links_nth H D). Qed. - -Lemma glD A (g : pregraph A) (x : node) : - n_pregraph_axiom 2 g -> - x \in dom g -> - (x, [:: lft g x; rgh g x]) \In links g. -Proof. -move=>H /[dup]/In_graphX R /(glE H) <-. -by apply/In_linksX. -Qed. - Lemma getU A (g : pregraph A) (x : node) v lf rg i : - get_nth (bupd g x v lf rg) x i = + get_nth (upd2 g x v lf rg) x i = if x \notin dom g then null else nth null [:: lf; rg] i. Proof. -rewrite /get_nth/graph.links/bupd find_omf/omfx /=. +rewrite /get_nth/graph.links/upd2 find_omf/omfx /=. case: (dom_find x)=>[|w /In_find H _]; first by rewrite nth_nil. by case: v=>[v|] /=; rewrite findU eqxx /= (In_cond H) (In_valid H). Qed. @@ -206,12 +64,19 @@ Qed. (****************) (* type of markings *) -Inductive mark := U | L | R | LR. +(* U = unmarked *) +(* M m = marked, but still exploring m (LL or RR) subgraph *) +(* MM = fully marked and explored both subgraphs *) +Inductive mark := U | M of Side | MM. (* decidable equality on marks *) Definition eq_mark l1 l2 : bool := - if (l1, l2) is ((U,U)|(L, L)|(R, R)|(LR, LR)) then true else false. + if (l1, l2) is ((U,U)|(M LL, M LL)|(M RR, M RR)|(MM, MM)) + then true else false. Lemma eq_markP : Equality.axiom eq_mark. -Proof. by case; case=>//=; constructor. Qed. +Proof. +case; case=>//=; try by [constructor]; +by case=>[|[]|]; constructor. +Qed. HB.instance Definition _ := hasDecEq.Build mark eq_markP. Notation pregraph := (pregraph mark). @@ -220,19 +85,22 @@ Definition marked (g : pregraph) : pregraph := um_filterv (fun v => v.1 != U) g. HB.instance Definition _ := OmapFun.copy marked marked. +Definition omark (g : pregraph) x : option Side := + if olabel g x is Some (M m) then Some m else None. (* marking of children *) +(* (* given marking map m, x is m-child of y iff: *) (* - m y = L and x is left child of y *) (* - m y = R and x is right child of y *) Definition ch (g : pregraph) (x y : nat) := - [|| (find y (contents g) == Some L) && (lft g y == x) | - (find y (contents g) == Some R) && (rgh g y == x)]. + [|| (olabel g y == Some LL) && (lft g y == x) | + (olabel g y == Some RR) && (rgh g y == x)]. Lemma chP (g : pregraph) x y : - reflect [\/ (y, L) \In contents g /\ lft g y = x | - (y, R) \In contents g /\ rgh g y = x] + reflect [\/ (y, L) \In labels g /\ lft g y = x | + (y, R) \In labels g /\ rgh g y = x] (ch g x y). Proof. rewrite /ch; case: orP=>H; constructor. @@ -261,152 +129,115 @@ Lemma ch_path_uniq g s : path (ch g) null s -> uniq (null :: s). Proof. by apply: path_uniq; [apply: chN0|apply: ch_fun]. Qed. - - -Definition upd_edge g x y : mark * seq node := - match find x (contents g) with - | Some L => (L, [:: y; rgh g x]) - | _ => (R, [:: lft g x; y]) - end. - -Fixpoint rev_graph' (g : pregraph) (ps : seq node) t : pregraph := - if ps is p :: ps then - rev_graph' (free g p) ps p \+ pts p (upd_edge g p t) - else g. - -Definition rev_graph g s t := rev_graph' g (rev s) t. - -(* layout of graph+marking as heap *) - - -(* -Definition lay (g : pregraph) : heap := - \big[join/Unit]_(x <- dom g) (x :-> (gl g x, gr g x, odflt L (find x m))). - -Lemma bigF (A : eqType) (xs : seq A) (f : A -> heap) k : - valid (\big[join/Unit]_(i <- xs) f i) -> - free (\big[join/Unit]_(i <- xs) f i) k = - \big[join/Unit]_(i <- xs) free (f i) k. -Proof. -elim: xs=>[|x xs IH]. -- by rewrite !big_nil free0. -rewrite !big_cons => V. -rewrite freeUn. -case: ifP=>D. -- rewrite IH //. rewrite (validR V). by []. -rewrite -IH; last by rewrite (validX V). -rewrite !freeND //. -- apply: contraFN D=>D. - rewrite domUn inE D orbT andbT. - by []. -apply: contraFN D=>D. -by rewrite domUn inE (validX V) D. -Qed. *) - -Lemma In_layE1 (g : pregraph) x pl pr c : - n_pregraph_axiom 2 g -> - find x (lay2 g) = Some (idyn (c, (pl, pr))) -> - (x, (c, [:: pl; pr])) \In g. -Proof. by move=>H /In_find /In_lay2X-/(_ H). Qed. - -Lemma In_layE2 (g : pregraph) x pl pr c : - n_pregraph_axiom 2 g -> - (x, (c, [:: pl; pr])) \In g -> - find x (lay2 g) = Some (idyn (c, (pl, pr))). -Proof. by move=>H /In_lay2X-/(_ H)/In_find. Qed. - -Lemma In_layX2 (g : pregraph) x c pl pr : - n_pregraph_axiom 2 g -> - (x, (c, [:: pl; pr])) \In g -> - exists h, lay2 g = x :-> (c, (pl, pr)) \+ h. -Proof. by move=>H Gx; eexists _; apply/heap_eta2/In_layE2. Qed. - -Lemma In_layX3 (g : pregraph) x l r c : - lay2 (bupd g x (Some c) l r) = - if x \in dom g then - upd x (idyn (c, (l, r))) (lay2 g) - else undef. -Proof. -rewrite /lay2/bupd find_omf/omfx /=. -case: (dom_find x)=>[|v]; first by rewrite omap_undef. -move/In_find=>E _; rewrite (upd_eta x). -rewrite omapPtUn -(upd_eta x) validU. -rewrite (In_cond E) (In_valid E) /=. -rewrite (upd_eta x) !omap_free !omap_omap. -congr (_ \+ _); apply/eq_in_omap. -by case=>k w /= H; case: (k =P x). -Qed. - -Lemma In_layX4 (g : pregraph) x l r c : - lay2 (x &-> (c, [:: l; r]) \+ g) = - x :-> (c, (l, r)) \+ lay2 g. -Proof. -rewrite omfPtUn /=; case: ifP=>// V; set j := (_ \+ _). -case: (normalP j)=>[->//|]. -rewrite !validPtUn valid_omap dom_omf_some // in V *. -by rewrite V. -Qed. - -(* reach g s t holds iff *) -(* each unmarked node x in g is reachable through unmarked nodes *) -(* - either starting from t, or *) -(* - starting from a right child of some node in s *) - +(* node x is in the stack iff its marked L or R *) +Definition stack_marked (g : pregraph) (s : seq node) := + forall x, x \in s = isSome (omark g x). + +(* consecutive stack nodes respect marking *) +Definition stack_consec (g : pregraph) (s : seq node) := + forall x y m, + (* if x is just under y in the stack *) + consec (null :: s) x y -> + (* and y is marked by m *) + omark g y = Some m -> + (* then x is y's child, left or right determined by m *) + x = sel2 m g y. + +(* graph g differs from g0 only in reversing edges on the stack *) +Definition graph_diff (g0 g : pregraph) (s : seq node) t := + [/\ (* graphs have equal nodes, and *) + dom g0 = dom g & forall x, + (* if x is marked by m, then *) + if omark g x is Some m then + (* sel2 m g x is predecessor of x (or null, if stack empty) *) + [/\ consec (null :: s) (sel2 m g x) x, + (* sel2 m g0 x is successor of x (or t, if x last) *) + consec (rcons s t) x (sel2 m g0 x) & + (* graphs agree on children of flipped marking *) + sel2 (flip m) g x = sel2 (flip m) g0 x] + (* if x is unmarked or fully marked then graphs agree *) + else forall m, sel2 m g x = sel2 m g0 x]. + +(* each unmarked node is reachable either from t *) +(* or from stack's right spine, *) +(* in both cases by avoiding marked nodes *) Definition reach (g : pregraph) (s : seq node) (t : node) := forall x, (* if x is unmarked node in g *) - (x, U) \In contents g -> + (x, U) \In labels g -> (* then x is reachable from t avoiding marked nodes *) - x \in connect (mem (dom (marked g))) g t \/ - (* or x is reachable from some right child of a node in s *) + x \in connect [dom (marked g)] g t \/ + (* or x is reachable from some right child of a node in stack s *) (* avoiding marked nodes *) exists2 y, y \in s & - x \in connect (mem (dom (marked g))) g (rgh g y). - -UP TO HERE - -Definition shape (g0 g : pregraph) (m : nmap mark) (r p t : node) := - fun h => exists stack, - [/\ h = lay m g, - {subset dom m <= dom g}, - path (ch m g) null stack, - rev_graph m g stack t = g0, - reach m g stack t, - {subset dom (um_filterv (predU (pred1 L) (pred1 LR)) m) <= stack}, - p = last null stack, - r = head t stack, - n_pregraph_axiom 2 g & - graph_axiom g]. + x \in connect [dom (marked g)] g (rgh g y). + +Definition shape (g0 g : pregraph) (r p t : node) := + fun h => exists s, + [/\ h = lay2 g, p = last null s, r = head t s, + stack_marked g s, stack_consec g s, + graph_diff g0 g s t, reach g s t, + n_pregraph_axiom 2 g & graph_axiom g]. Program Definition pop (p t : ptr): - STsep {g0 g m r} (fun h => - [/\ shape g0 g m r p t h, - (p, R) \In m & - t \in dom m \/ t = null], + STsep {g0 g r} (fun h => + [/\ shape g0 g r p t h, + (p, M RR) \In labels g & + t \in dom (marked g) \/ t = null], [vfun res h => - let: g' := upd p [:: gl g p; t] g in - let: m' := upd p LR m in - shape g0 g' m' r res.1 res.2 h /\ - res = (gr g p, p)]) := - Do (q <-- read (A:=node*node*mark) p; - p ::= (q.1.1, t, LR);; - ret (q.1.2, p)). + shape g0 (updCR g p MM t) r res.1 res.2 h /\ + res = (rgh g p, p)]) := + Do ('(_, (l, r)) <-- read (mark*(node*node)) p; + p ::= (MM, (l, t));; + ret (r, p)). Next Obligation. -move=>p t [g0][g][m][r][/= h][H Pm D]. -case: H=>stack [-> H1 H2 H3 H4 H5 H6 H7 H8 H9]. -move: (H1 p (In_dom Pm))=>Pg. -case/(glD H8)/(In_layX2 H8 H1): (Pg)=>h' E. -rewrite E. -do 3!step; split=>//; exists stack; split. -- rewrite In_layX3. - - by rewrite upd_eta E freePtUn // (validX H). - - by rewrite (dom_valid Pg). - - by rewrite (In_valid Pm). - - by rewrite (In_cond Pm). - -Search upd. +move=>p t [g0][g][r][_][[s [->]] S Sel Dg C Rc Ep Er H G Pm D] /=. +case/In_marksX: Pm=>lks /[dup] /(In_blinks H) -> Pm. +have Uq : uniq (null :: s). +- admit. +have Ps : p \in s. +- by rewrite Ep last_change // -Ep (In_cond Pm). +rewrite Ep in Ps. +case/rcons_lastP: Ps=>s'. rewrite -Ep. +move=>E. subst s. +clear Ep. +have Nxs' : p \notin s'. +- simpl in Uq. rewrite rcons_uniq in Uq. + by case/and3P: Uq. + +case: (In_layX2 H Pm)=>h /[dup] E ->; do 3!step; move=>V. +split=>//; exists s'; split. +- by rewrite In_layX3 /= (In_dom Pm) upd_eta E freePtUn ?(validX V). +- move=>x. + rewrite /bupd (In_findE (In_marks Pm)) omfU ?(In_cond Pm) //=. + split. + - move=>B. + have Nxp : x != p. + - by case: eqP B Nxs'=>// ->->. + have : x \in rcons s' p. rewrite mem_rcons inE B orbT. by []. + move/S. case=>X; [left|right]. + - apply/InU=>/=. rewrite validU (In_cond Pm) (negbTE Nxp). rewrite pfVE (In_valid Pm). by []. + - apply/InU=>/=. rewrite validU (In_cond Pm) (negbTE Nxp). rewrite pfVE (In_valid Pm). by []. + rewrite !InU /=. + case; case; case: eqP=>// /eqP Nxp _ X; + suff : x \in rcons s' p by [rewrite mem_rcons inE (negbTE Nxp)]. + - by apply/S; left. + by apply/S; right. +k + + + + have : x != p /\ x \in s. + - + + + - move/mem_belast. + + - rewrite /bupd (In_findE (In_marks Pm)). + + admit. diff --git a/htt/heapauto.v b/htt/heapauto.v index 8ae51d9..bcac768 100644 --- a/htt/heapauto.v +++ b/htt/heapauto.v @@ -150,7 +150,7 @@ Variables (A B : Type). Lemma val_readR v x i (f : form (x :-> v) i) (r : post A) : (valid (untag f) -> r (Val v) f) -> - vrf f (read x) r. + vrf f (read A x) r. Proof. by rewrite formE; apply: vrf_read. Qed. Lemma try_readR e1 e2 v x i (f : form (x :-> v) i) (r : post B) : diff --git a/htt/model.v b/htt/model.v index 7f96384..a2c66e4 100644 --- a/htt/model.v +++ b/htt/model.v @@ -73,7 +73,7 @@ Parameter allocb : forall A, A -> nat -> ST ptr. Parameter dealloc : ptr -> ST unit. Arguments throw {A}. -Arguments read {A}. +Arguments read : clear implicits. (* given program e : ST A, input heap i, postcondition Q *) (* vrf' judges if running e in i produces output heap and result *) @@ -116,7 +116,7 @@ Parameter vrf_try : forall A B (e : ST A) (e1 : A -> ST B) vrf i (try e e1 e2) Q. Parameter vrf_read : forall A x j (v : A) (Q : post A), (valid (x :-> v \+ j) -> Q (Val v) (x :-> v \+ j)) -> - vrf (x :-> v \+ j) (read x) Q. + vrf (x :-> v \+ j) (read A x) Q. Parameter vrf_write : forall A x (v : A) B (u : B) j (Q : post unit), (valid (x :-> v \+ j) -> Q (Val tt) (x :-> v \+ j)) -> vrf (x :-> u \+ j) (write x v) Q. @@ -149,7 +149,7 @@ Notation "' x '<--' c1 ';' c2" := (bnd c1 (fun x => c2)) "'[v' ' x '<--' c1 ';' '//' c2 ']'"). Notation "c1 ';;' c2" := (bnd c1 (fun _ => c2)) (at level 81, right associativity). -Notation "'!' x" := (read x) (at level 50). +Notation "'!' x" := (read _ x) (at level 50). Notation "x '::=' e" := (write x e) (at level 60). (* using locked for Do, to make the Do definitions opaque *) (* otherwise, clients will inline the code of Do, preventing *) diff --git a/pcm/heap.v b/pcm/heap.v index c58692c..ea4c3c1 100644 --- a/pcm/heap.v +++ b/pcm/heap.v @@ -20,7 +20,7 @@ limitations under the License. From HB Require Import structures. From Coq Require Import ssreflect ssrbool ssrfun Eqdep. From mathcomp Require Import ssrnat eqtype fintype tuple finfun seq path bigop. -From pcm Require Import options axioms prelude finmap. +From pcm Require Import options axioms prelude pred finmap. From pcm Require Import pcm unionmap natmap. (************) @@ -171,6 +171,9 @@ Export Heap.Exports. Notation "x :-> v" := (ptsT heap x (idyn v)) (at level 30). +Canonical heap_PredType : PredType (nat * dynamic id) := um_PredType heap. +Coercion Pred_of_nmap (x : heap) : {Pred _} := [eta Mem_UmMap x]. + (********************) (* points-to lemmas *) (********************) diff --git a/pcm/unionmap.v b/pcm/unionmap.v index 5d166db..5d76e97 100644 --- a/pcm/unionmap.v +++ b/pcm/unionmap.v @@ -597,6 +597,8 @@ Abort. (* dom *) (*******) +Notation "[ 'dom' f ]" := [mem (dom f)] : fun_scope. + Section DomLemmas. Variables (K : ordType) (C : pred K) (V : Type) (U : union_map C V). Implicit Types (k : K) (v : V) (f g : U). @@ -607,7 +609,10 @@ Proof. by rewrite !umEX. Qed. Lemma dom0 : dom (Unit : U) = [::]. Proof. by rewrite !umEX. Qed. -Lemma dom0E f : valid f -> dom f =i pred0 -> f = Unit. +Lemma dom0E f : + valid f -> + dom f =i pred0 -> + f = Unit. Proof. rewrite !umEX /UM.valid /UM.dom /UM.empty -{3}[f]tfE !eqE UM.umapE. case: (UMC_from f)=>[|f' S] //= _; rewrite fmapE /= {S}. @@ -844,7 +849,7 @@ Proof. by move=>E /domE D1 /domE D2; apply/domE/domKUni. Qed. Lemma subdom_filter (f1 : U1) (f2 : U2) : {subset dom f1 <= dom f2} -> - dom f1 = filter (mem (dom f1)) (dom f2). + dom f1 = filter [dom f1] (dom f2). Proof. have Tx : transitive (@ord K) by apply: trans. move=>S; apply: (sorted_eq (leT := @ord K))=>//. @@ -2267,7 +2272,7 @@ Qed. Lemma umpreim_pred0 f : um_preim pred0 f =1 pred0. Proof. by move=>x; rewrite /um_preim; by case: (find x f). Qed. -Lemma umpreim_dom f : um_preim predT f =1 mem (dom f). +Lemma umpreim_dom f : um_preim predT f =1 [dom f]. Proof. move=>x; rewrite /um_preim /=. case X: (find x f)=>[a|] /=; first by rewrite (find_some X). @@ -3544,6 +3549,9 @@ Proof. rewrite omfE; apply: omap_subdom. Qed. Arguments omf_subdom {f x}. +Lemma In_odom f x k : k \In f x -> k.1 \in dom x. +Proof. by move/In_dom/omf_subdom. Qed. + Lemma omf_cond f x k : k \in dom (f x) -> C k. Proof. by move/omf_subdom/dom_cond. Qed. @@ -3780,6 +3788,7 @@ End OmapFunLemmas. Arguments omf_subdom {K C V V' U U' f x}. Arguments In_omf {K C V V' U U'} _ {k v w x}. +Arguments In_odom {K C V V' U U' f x k} . (* omap_fun's with different ranges *) @@ -4475,7 +4484,7 @@ Proof. by move=>k; rewrite dom_umfiltkE mem_filter. Qed. (* /DEVCOMMENT *) Lemma umfiltk_dom f1 f2 : valid (f1 \+ f2) -> - um_filterk (mem (dom f1)) (f1 \+ f2) = f1. + um_filterk [dom f1] (f1 \+ f2) = f1. Proof. move=>W; apply/umem_eq; first by rewrite pfVE. - by rewrite (validL W). @@ -4484,7 +4493,7 @@ case=>k v; rewrite In_umfiltX; split=>[|H]. by split; [apply: In_dom H | apply: InL]. Qed. -Corollary umfiltk_dom' f : um_filterk (mem (dom f)) f = f. +Corollary umfiltk_dom' f : um_filterk [dom f] f = f. Proof. case: (normalP f)=>[->|H]; first by rewrite pfundef. by rewrite -{2}[f]unitR umfiltk_dom // unitR. @@ -4507,7 +4516,7 @@ 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. +Lemma umfiltk_memdomE f : um_filterk [dom f] f = f. Proof. by apply/umfiltk_subdomE. Qed. Lemma find_umfiltk k (p : pred K) f : @@ -4564,7 +4573,7 @@ by rewrite /obind/oapp /=; case=>k v; case: ifP; case: (omf f _). Qed. Lemma umfiltk_dom_omf V' (U' : union_map C V') (f : omap_fun U U') x : - um_filterk (mem (dom x)) (f x) = f x. + um_filterk [dom x] (f x) = f x. Proof. by rewrite -umfiltk_omf umfiltk_dom'. Qed. Lemma umfiltkU p k v f : @@ -4669,7 +4678,7 @@ 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. + um_filterk [dom a] x = a. Proof. by move=>W; split=>[|<-] // H; case: H W=>b ->; apply: umfiltk_dom. Qed. (* join is the least upper bound *) @@ -4754,7 +4763,7 @@ 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)). + reflect (x1 = um_filterk [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. From a2283a8f848cde99fc4687c64f9b600fae78231c Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Thu, 5 Sep 2024 19:04:16 +0200 Subject: [PATCH 77/93] blah --- examples/graph.v | 10 ++++++++++ examples/schorr.v | 29 +++++++++-------------------- 2 files changed, 19 insertions(+), 20 deletions(-) diff --git a/examples/graph.v b/examples/graph.v index 7315997..de401b8 100644 --- a/examples/graph.v +++ b/examples/graph.v @@ -1369,6 +1369,16 @@ rewrite !validPtUn valid_omap dom_omf_some // in V *. by rewrite V. Qed. +Lemma sel2U A (g : pregraph A) (x : node) v lf rg i : + sel2 i (upd2 g x v lf rg) x = + if x \notin dom g then null + else nth null [:: lf; rg] i. +Proof. +rewrite /get_nth/graph.links/upd2 find_omf/omfx /=. +case: (dom_find x)=>[|w /In_find H _]; first by rewrite nth_nil. +by case: v=>[v|] /=; rewrite findU eqxx /= (In_cond H) (In_valid H). +Qed. + (**********) (**********) (* Graphs *) diff --git a/examples/schorr.v b/examples/schorr.v index 5f3b1af..921a0cd 100644 --- a/examples/schorr.v +++ b/examples/schorr.v @@ -45,20 +45,6 @@ rewrite -[RHS]connect_umfiltk eq_connect_aux //. by move=>y /S2; rewrite !inE negbK. Qed. -(*****************) -(* Binary graphs *) -(*****************) - -Lemma getU A (g : pregraph A) (x : node) v lf rg i : - get_nth (upd2 g x v lf rg) x i = - if x \notin dom g then null - else nth null [:: lf; rg] i. -Proof. -rewrite /get_nth/graph.links/upd2 find_omf/omfx /=. -case: (dom_find x)=>[|w /In_find H _]; first by rewrite nth_nil. -by case: v=>[v|] /=; rewrite findU eqxx /= (In_cond H) (In_valid H). -Qed. - (****************) (* Schorr-Waite *) (****************) @@ -131,9 +117,11 @@ Lemma ch_path_uniq g s : Proof. by apply: path_uniq; [apply: chN0|apply: ch_fun]. Qed. *) -(* node x is in the stack iff its marked L or R *) +(* node in the stack are marked L or R *) +(* other direction also holds, but isn't stated explicitly *) +(* as it's implied by graph_diff (defined below) *) Definition stack_marked (g : pregraph) (s : seq node) := - forall x, x \in s = isSome (omark g x). + forall x, x \in s -> isSome (omark g x). (* consecutive stack nodes respect marking *) Definition stack_consec (g : pregraph) (s : seq node) := @@ -149,15 +137,16 @@ Definition stack_consec (g : pregraph) (s : seq node) := Definition graph_diff (g0 g : pregraph) (s : seq node) t := [/\ (* graphs have equal nodes, and *) dom g0 = dom g & forall x, - (* if x is marked by m, then *) + (* if x is marked m, then *) if omark g x is Some m then - (* sel2 m g x is predecessor of x (or null, if stack empty) *) + (* m-child of x in g is predecessor of x (or null, if stack empty) *) [/\ consec (null :: s) (sel2 m g x) x, - (* sel2 m g0 x is successor of x (or t, if x last) *) + (* m-child of x in g0 is successor of x (or t, if x last) *) consec (rcons s t) x (sel2 m g0 x) & (* graphs agree on children of flipped marking *) sel2 (flip m) g x = sel2 (flip m) g0 x] - (* if x is unmarked or fully marked then graphs agree *) + (* otherwise, if x is unmarked or fully marked, then *) + (* graphs agree on children of x *) else forall m, sel2 m g x = sel2 m g0 x]. (* each unmarked node is reachable either from t *) From 48f54a24eb60f5d42330598120ce98130dbfcc80 Mon Sep 17 00:00:00 2001 From: Aleks Nanevski Date: Fri, 6 Sep 2024 06:23:11 +0200 Subject: [PATCH 78/93] many changes introduced to deal with graphs --- core/seqext.v | 10 +++ core/slice.v | 3 +- core/uconsec.v | 174 +++++++++++++++++++++++++++++++++++++++--- core/useqord.v | 3 +- core/uslice.v | 4 +- examples/bubblesort.v | 2 +- examples/graph.v | 70 +++++++++++++---- examples/quicksort.v | 2 +- examples/schorr.v | 141 ++++++++++++++++------------------ pcm/natmap.v | 1 + 10 files changed, 298 insertions(+), 112 deletions(-) diff --git a/core/seqext.v b/core/seqext.v index ee757ee..7033112 100644 --- a/core/seqext.v +++ b/core/seqext.v @@ -21,6 +21,10 @@ From pcm Require Import options prelude pred. (* TODO upstream to mathcomp *) +Lemma head_rcons {A} (s : seq A) (x y : A) : + head x (rcons s y) = head y s. +Proof. by rewrite headI. Qed. + Lemma rcons_nseq {A} n (x : A) : rcons (nseq n x) x = nseq n.+1 x. Proof. by elim: n=>//=n ->. Qed. @@ -1012,6 +1016,12 @@ case: eqP=>[->|N]; first by apply: rcons_lastP. by constructor; case=>s' E; elim: N; rewrite E last_rcons. Qed. +Lemma rcons_lastN (s : seq A) a p : + p != a -> + p = last a s -> + exists ss, s = rcons ss p. +Proof. by move/[swap]=>-> N; apply/rcons_lastX/last_change/N. Qed. + Lemma index_last_size_uniq z (s : seq A) : uniq s -> index (last z s) s = (size s).-1. diff --git a/core/slice.v b/core/slice.v index 9f053ce..dca83f4 100644 --- a/core/slice.v +++ b/core/slice.v @@ -15,8 +15,7 @@ From Coq Require Import ssreflect ssrbool ssrfun. From mathcomp Require Import ssrnat eqtype seq path interval order. From mathcomp Require Import fintype finfun tuple. From pcm Require Import options prelude seqext. - -Open Scope order_scope. +Local Open Scope order_scope. Import Order.Theory. (* DEVCOMMENT *) diff --git a/core/uconsec.v b/core/uconsec.v index 9da15eb..fdfc444 100644 --- a/core/uconsec.v +++ b/core/uconsec.v @@ -14,8 +14,7 @@ limitations under the License. From Coq Require Import ssreflect ssrbool ssrfun. From mathcomp Require Import ssrnat eqtype seq path interval order. From pcm Require Import options prelude ordtype seqext slice useqord uslice. - -Open Scope order_scope. +Local Open Scope order_scope. Import Order.Theory. (* We assume the sequences are unique and most lemmas do require this *) @@ -499,7 +498,8 @@ Qed. Lemma consec_hdswap k1 k2 ks x : uniq ks -> - k1 \notin ks -> k2 \notin ks -> + k1 \notin ks -> + k2 \notin ks -> x != k2 -> consec (k1::ks) k1 x -> consec (k2::ks) k2 x. Proof. @@ -516,16 +516,13 @@ by apply: contra K1 => /eqP <-. Qed. Lemma consec_hd2 k1 k2 ks : - uniq ks -> - k1 \notin ks -> k2 \notin ks -> - k1 != k2 -> consec [:: k1, k2 & ks] k1 k2. + k1 != k2 -> + consec [:: k1, k2 & ks] k1 k2. Proof. -move=>U K1 K2 N; rewrite /consec !sltL eq_sym N /= nilp_hasPn. +move=>N; rewrite /consec !sltL eq_sym N /= nilp_hasPn. apply/hasPn=>z; apply: contraL=>/= _. -rewrite eqslice_mem_uniq; last first. -- by rewrite /= inE negb_or N K1 K2. -rewrite negb_and; apply/orP; right. -by rewrite in_itv /= (negbTE N) !eqxx /= negb_and ltEnat /= -!leqNgt leqn0 lt0n orbN. +rewrite eqslice_mem /= eqxx (negbTE N) /= eqxx. +by apply/hasP; case=>x _; rewrite in_itv; case: x. Qed. (* a useful lemma that collects the case analysis *) @@ -739,6 +736,36 @@ case: ifP=>H1 //; case: ifP=>H2 //. by split; [case=>->->; rewrite !eq_refl|case/andP=>/eqP -> /eqP ->]. Qed. +Lemma consec_rcons (s : seq A) a x y : + uniq (a :: rcons s y) -> + consec (a :: rcons s y) x y <-> + x = last a s. +Proof. +rewrite /= mem_rcons inE negb_or rcons_uniq -andbA. +case/and4P=>U1 U2 U3 U4. +rewrite -rcons_cons consec_rconsE; last first. +- by rewrite inE negb_or eq_sym U1 U3. +- by rewrite /= U2 U4. +case: (x =P y)=>[->{x}|/eqP N] /=. +- rewrite inE negb_or eq_sym U1 mem_rcons inE eqxx /=. + split=>// E; move: (mem_last a s). + by rewrite -E inE eq_sym (negbTE U1) (negbTE U3). +rewrite inE negb_or eq_sym U1 U3 /= eqxx /=. +by split=>[/eqP|->]. +Qed. + +Lemma consec_ext (s : seq A) x y p : + uniq (rcons s p) -> + consec s x y -> + y \in s -> + consec (rcons s p) x y. +Proof. +rewrite rcons_uniq=>/andP [U1 U2] C S. +rewrite consec_rconsE //=. +case: eqP C U1=>[-> /consec_mem ->//|/eqP N]. +by rewrite S. +Qed. + Lemma consec_behead k ks x y : uniq ks -> k \notin ks -> x != k -> @@ -951,3 +978,128 @@ by case=>-> {t T} [ks' -> _]; left; exists ks'. Qed. End ConsecNat. + + +(**********************) +(* consec, index, nth *) +(**********************) + +Lemma consec_index (A : eqType) (ks : seq A) (x y : A) : + uniq ks -> + consec ks x y <-> + (index y ks) = (index x ks).+1. +Proof. +move=>U; split. +- elim: ks x y U=>[|k ks IH] /= x y; first by rewrite consec_nil. + case/andP=>U1 U2; rewrite consec_consE // !(eq_sym k). + case: (x =P k); case: (y =P k)=>Ex Ey //. + - by subst x y; rewrite (negbTE U1). + - subst x; case: ifP=>Ky /eqP -> //. + by case: ks {Ky Ex U1 U2 IH}=>//= ??; rewrite eqxx. + by move/IH=>->. +elim: ks x y U=>[|k ks IH] //= x y /andP [U1 U2]. +case: (k =P y)=>N; first by subst k; case: eqP. +move/eqP: N=>N. +case: (k =P x); last first. +- move/eqP=>Nx [I]; rewrite consec_consE // eq_sym (negbTE Nx) eq_sym N /=. + by apply: IH I. +move=>?; subst x; case. +case: ks {IH} U1 U2 N=>[|a ks] /= U1 U2 N. +- by rewrite consec_consE //= eqxx eq_sym N eqxx. +rewrite inE negb_or in U1; case/andP: U1=>U1 U1'. +case/andP: U2=>U2 U2'; case: eqP=>// -> _. +by rewrite consec_hd2. +Qed. + +(* push nat_scope to the top of scope stacks explicitly *) +(* because of clash with order_scope *) +Local Open Scope nat_scope. + +(* consec/nth in hypotheses, i.e., elimination *) + +Lemma consec_nthE (A : eqType) (ks : seq A) a (x : A) i : + uniq (a :: ks) -> + i < size ks -> + consec ks x (nth a ks i.+1) -> + x = nth a ks i. +Proof. +rewrite /= =>/andP [U1 U2] N. +case N1 : (i.+1 < size ks). +- move/[dup]/consec_mem=>X /consec_index-/(_ U2). + by rewrite index_uniq //; case=>->; rewrite nth_index. +have -> : nth a ks i.+1 = a by rewrite nth_default // leqNgt N1. +move/negbT: N1; rewrite -leqNgt leq_eqVlt ltnS leqNgt N orbF. +move/eqP=>S; rewrite (consecP_lastE a) // -nth_last S. +by rewrite -subn1 subSS subn0; case/andP=>_ /eqP. +Qed. + +(* simmilar, but uses consed list *) + +Lemma consec_nthE' (A : eqType) (ks : seq A) a (x : A) i : + uniq (a :: ks) -> + i < size ks -> + consec (a :: ks) x (nth a ks i) -> + x = if i == 0 then a else nth a ks i.-1. +Proof. +move/[dup]=>U /= /andP [U1 U2] N. +move/[dup]/consec_mem=>D /consec_index-/(_ U) /=. +have X : nth a ks i \in ks by rewrite mem_nth. +case: (a =P nth _ _ _) U1=>[->|]; first by rewrite X. +move/eqP=>Y U1; rewrite index_uniq //; case. +case: (a =P x)=>[->->//|/eqP Na -> /=]. +by rewrite nth_index //; move: D; rewrite inE eq_sym (negbTE Na). +Qed. + +(* consec/nth in the conclusion; introduction lemma *) + +Lemma consec_nthI (A : eqType) (ks : seq A) a i : + uniq (a :: ks) -> + i.+1 < size ks -> + consec (a :: ks) (nth a ks i) (nth a ks i.+1). +Proof. +elim: ks a {3 4}a i=>[|k ks IH] a b i //=. +rewrite inE negb_or -andbA; case/and4P=>U1 U2 U3 U4. +rewrite ltnS=>S; rewrite consec_consE //=; last first. +- by rewrite inE negb_or U1 U2. +- by rewrite U3 U4. +rewrite nth_cons; case: i S=>[|i] S /=. +- rewrite (eq_sym k) (negbTE U1) /= nth0. + case: ks {U1 IH} U2 U3 U4 S=>//= c ks U2 U3 U4 S. + - rewrite inE negb_or in U2; case/andP: U2=>U2 U2'. + by rewrite eq_sym U2 consec_consE // eqxx inE eqxx /=. +case: ifP=>E. +- have N : nth b ks i \notin ks by rewrite (eqP E). + by rewrite mem_nth // in N; apply: leq_trans S. +move/negbT: E=>N. +have M : nth b ks i.+1 \in ks by rewrite mem_nth. +case: eqP=>X /=; first by rewrite X (negbTE U2) in M. +by apply: IH=>//=; rewrite U3 U4. +Qed. + +(*******************) +(* consec and path *) +(*******************) + +Lemma consec_path (A : eqType) a ks (R : rel A) : + uniq (a :: ks) -> + (forall x y, x \in a::ks -> y \in ks -> + consec (a :: ks) x y -> R x y) -> + path R a ks. +Proof. +move=>U H; apply/(pathP a)=>i S; rewrite nth_cons. +case: i S=>[|i] S /=; last first. +- apply: H. + - by rewrite inE mem_nth ?orbT //; apply: leq_trans S. + - by rewrite mem_nth. + by apply: consec_nthI U S. +case: ks U H S=>[|k ks] //= U H S; apply: H. +- by rewrite inE eqxx. +- by rewrite inE eqxx. +rewrite inE negb_or -andbA in U. +case/and4P: U=>U1 U2 U3 U4. +rewrite consec_consE /=; last first. +- by rewrite inE negb_or U1. +- by rewrite U3 U4. +by rewrite eqxx inE eqxx. +Qed. + diff --git a/core/useqord.v b/core/useqord.v index 3b97c06..d34c170 100644 --- a/core/useqord.v +++ b/core/useqord.v @@ -14,8 +14,7 @@ limitations under the License. From Coq Require Import ssreflect ssrbool ssrfun. From mathcomp Require Import ssrnat eqtype seq path interval order. From pcm Require Import options prelude ordtype seqext. - -Open Scope order_scope. +Local Open Scope order_scope. Import Order.Theory. (* We assume the sequences are unique and use the first index, however most *) diff --git a/core/uslice.v b/core/uslice.v index 3624796..b17c674 100644 --- a/core/uslice.v +++ b/core/uslice.v @@ -14,12 +14,10 @@ limitations under the License. From Coq Require Import ssreflect ssrbool ssrfun. From mathcomp Require Import ssrnat eqtype seq path interval order. From pcm Require Import options prelude ordtype seqext slice useqord. - (* We assume the sequences are unique and use the first index, *) (* however most lemmas don't require this condition explicitly. *) (* The ones that do are grouped in separate sections. *) - -Open Scope order_scope. +Local Open Scope order_scope. Import Order.Theory. (* slicing by element index *) diff --git a/examples/bubblesort.v b/examples/bubblesort.v index ab192b1..8fb6810 100644 --- a/examples/bubblesort.v +++ b/examples/bubblesort.v @@ -18,9 +18,9 @@ From pcm Require Import options axioms prelude pred ordtype slice. From pcm Require Import seqext pcm unionmap heap. From htt Require Import options model heapauto. From htt Require Import array. - From mathcomp Require order. Import order.Order.NatOrder order.Order.TTheory. +Local Open Scope order_scope. (* Brief mathematics of (bubble) array sorting: *) (* Theory of permutations built out of (adjacent-element) swaps acting on *) diff --git a/examples/graph.v b/examples/graph.v index de401b8..8820bdc 100644 --- a/examples/graph.v +++ b/examples/graph.v @@ -111,7 +111,7 @@ Section PregraphLemmas. Context {A : Type}. Implicit Type g : pregraph A. -(* marks lemmas *) +(* labels lemmas *) Lemma In_labelsX g x v : (x, v) \In labels g <-> @@ -121,10 +121,15 @@ rewrite In_omfX; split; last by case=>lks H; exists (v, lks). by case; case=>w lks /= H [<-{v}]; exists lks. Qed. -Lemma In_labels g x v lks : - (x, (v, lks)) \In g -> - (x, v) \In labels g. -Proof. by rewrite In_labelsX; exists lks. Qed. +Lemma In_labels g x xs : + (x, xs) \In g -> + (x, xs.1) \In labels g. +Proof. by case: xs=>v lks; rewrite In_labelsX; exists lks. Qed. + +Lemma In_olabel g x xs : + (x, xs) \In g -> + olabel g x = Some xs.1. +Proof. by rewrite /olabel=>/In_labels/In_find ->. Qed. (* links lemmas *) @@ -1228,7 +1233,7 @@ End NpregraphLemmas. (********************) (* notation for left/right node of x *) -Notation sel2 m g x := (get_nth g x m). +Definition sel2 A (m : Side) (g : pregraph A) x := get_nth g x m. Notation lft g x := (sel2 LL g x). Notation rgh g x := (sel2 RR g x). @@ -1265,6 +1270,49 @@ Notation updCL g x v l := (upd2 g x (Some v) l (rgh g x)). (* updating just contents and right link *) Notation updCR g x v r := (upd2 g x (Some v) (lft g x) r). +Lemma find_upd2 A (g : pregraph A) x p v lf rg : + find x (upd2 g p (Some v) lf rg) = + if p \in dom g then + if x == p then Some (v, [:: lf; rg]) + else find x g + else None. +Proof. +rewrite /upd2/labels find_omf /omfx/=. +case: (dom_find p)=>[|[k w]]; first by rewrite find_undef. +by move/In_find=>H _; rewrite findU (In_cond H) (In_valid H). +Qed. + +Lemma find_upd2N A (g : pregraph A) x p lf rg : + find x (upd2 g p None lf rg) = + if find p (labels g) is Some v then + if x == p then Some (v, [:: lf; rg]) + else find x g + else None. +Proof. +rewrite /upd2/labels find_omf /omfx /=. +case: (dom_find p)=>[|[k w]]; first by rewrite find_undef. +by move/In_find=>H _; rewrite findU (In_cond H) (In_valid H). +Qed. + +Lemma sel2U A (g : pregraph A) (x p : node) v lf rg i : + sel2 i (upd2 g p v lf rg) x = + if x == p then + if x \in dom g then nth null [:: lf; rg] i + else null + else + if (p \in dom g) && (x \in dom g) then sel2 i g x + else null. +Proof. +rewrite /sel2/get_nth/graph.links/upd2 find_omf/omfx /=. +case: (x =P p)=>[->|/eqP N]. +- case: (dom_find p)=>[|w /In_find H _]; first by rewrite nth_nil. + by case: v=>[v|] /=; rewrite findU eqxx /= (In_cond H) (In_valid H). +case: (dom_find p g)=>/=; first by rewrite nth_nil. +case=>k w /In_find H _; case: v=>[v|]; +rewrite /oapp findU (negbTE N) (In_cond H); +by case: (dom_find x g)=>//; rewrite nth_nil. +Qed. + Lemma In_links2 A (g : pregraph A) (x : node) v lks : n_pregraph_axiom 2 g -> (x, (v, lks)) \In g -> @@ -1369,16 +1417,6 @@ rewrite !validPtUn valid_omap dom_omf_some // in V *. by rewrite V. Qed. -Lemma sel2U A (g : pregraph A) (x : node) v lf rg i : - sel2 i (upd2 g x v lf rg) x = - if x \notin dom g then null - else nth null [:: lf; rg] i. -Proof. -rewrite /get_nth/graph.links/upd2 find_omf/omfx /=. -case: (dom_find x)=>[|w /In_find H _]; first by rewrite nth_nil. -by case: v=>[v|] /=; rewrite findU eqxx /= (In_cond H) (In_valid H). -Qed. - (**********) (**********) (* Graphs *) diff --git a/examples/quicksort.v b/examples/quicksort.v index 5b004d2..8f88cd7 100644 --- a/examples/quicksort.v +++ b/examples/quicksort.v @@ -19,9 +19,9 @@ From pcm Require Import options axioms prelude seqext pred ordtype slice. From pcm Require Import pcm unionmap heap. From htt Require Import options model heapauto. From htt Require Import array. - From mathcomp Require order. Import order.Order.NatOrder order.Order.TTheory. +Local Open Scope order_scope. (* Brief mathematics of quickorting *) (* There is some overlap with mathematics developed for bubblesort *) diff --git a/examples/schorr.v b/examples/schorr.v index 921a0cd..7fc5e32 100644 --- a/examples/schorr.v +++ b/examples/schorr.v @@ -74,36 +74,31 @@ HB.instance Definition _ := OmapFun.copy marked marked. Definition omark (g : pregraph) x : option Side := if olabel g x is Some (M m) then Some m else None. -(* marking of children *) +Lemma In_omark (g : pregraph) x c lks : + (x, (c, lks)) \In g -> + omark g x = if c is M m then Some m else None. +Proof. by rewrite /omark=>/In_olabel=>->. Qed. + + +(* old helpers kept for now, but of unclear utility *) (* -(* given marking map m, x is m-child of y iff: *) -(* - m y = L and x is left child of y *) -(* - m y = R and x is right child of y *) -Definition ch (g : pregraph) (x y : nat) := - [|| (olabel g y == Some LL) && (lft g y == x) | - (olabel g y == Some RR) && (rgh g y == x)]. - -Lemma chP (g : pregraph) x y : - reflect [\/ (y, L) \In labels g /\ lft g y = x | - (y, R) \In labels g /\ rgh g y = x] - (ch g x y). -Proof. -rewrite /ch; case: orP=>H; constructor. -- by case: H=>H; [left|right]; case/andP: H=>/eqP/In_find H /eqP. -by case=>[][/In_find/eqP M /eqP G]; apply: H; [left|right]; apply/andP. -Qed. +Definition ch (g : pregraph) x y : bool := + if omark g y is Some m then sel2 m g y == x else false. Lemma chN0 g x y : ch g x y -> y != null. -Proof. by case/chP=>[][/In_dom/dom_cond]. Qed. +Proof. +rewrite /ch/omark/olabel find_omf /omfx /=. +by case: (dom_find y)=>[//|v] /In_find/In_cond. +Qed. Lemma ch_fun g a b x : ch g a x -> ch g b x -> a = b. -Proof. by case/chP=>[][H <-] /chP [][/(In_fun H) {}H <-]. Qed. +Proof. by rewrite /ch; case: (omark g x)=>// v /eqP -> /eqP. Qed. Lemma ch_path g s x : x \in s -> @@ -117,11 +112,10 @@ Lemma ch_path_uniq g s : Proof. by apply: path_uniq; [apply: chN0|apply: ch_fun]. Qed. *) -(* node in the stack are marked L or R *) -(* other direction also holds, but isn't stated explicitly *) -(* as it's implied by graph_diff (defined below) *) + +(* stack contains all and only nodes marked L or R *) Definition stack_marked (g : pregraph) (s : seq node) := - forall x, x \in s -> isSome (omark g x). + forall x, x \in s = isSome (omark g x). (* consecutive stack nodes respect marking *) Definition stack_consec (g : pregraph) (s : seq node) := @@ -139,10 +133,8 @@ Definition graph_diff (g0 g : pregraph) (s : seq node) t := dom g0 = dom g & forall x, (* if x is marked m, then *) if omark g x is Some m then - (* m-child of x in g is predecessor of x (or null, if stack empty) *) - [/\ consec (null :: s) (sel2 m g x) x, (* m-child of x in g0 is successor of x (or t, if x last) *) - consec (rcons s t) x (sel2 m g0 x) & + [/\ consec (rcons s t) x (sel2 m g0 x) & (* graphs agree on children of flipped marking *) sel2 (flip m) g x = sel2 (flip m) g0 x] (* otherwise, if x is unmarked or fully marked, then *) @@ -165,11 +157,27 @@ Definition reach (g : pregraph) (s : seq node) (t : node) := Definition shape (g0 g : pregraph) (r p t : node) := fun h => exists s, - [/\ h = lay2 g, p = last null s, r = head t s, + [/\ h = lay2 g, p = last null s, r = head t s, uniq (null :: s), stack_marked g s, stack_consec g s, graph_diff g0 g s t, reach g s t, n_pregraph_axiom 2 g & graph_axiom g]. +(* helper lemma of unclear utility *) + +(* +Lemma stack_path g s : + stack_marked g s -> + stack_consec g s -> + uniq (null :: s) -> + path (ch g) null s. +Proof. +move=>H1 H2 U; apply: consec_path=>//= x y Dx Dy C. +move: Dy; rewrite H1 /ch; case D : (omark g y)=>[a|//]. +by rewrite -(H2 x y). +Qed. +*) + + Program Definition pop (p t : ptr): STsep {g0 g r} (fun h => [/\ shape g0 g r p t h, @@ -182,55 +190,36 @@ Program Definition pop (p t : ptr): p ::= (MM, (l, t));; ret (r, p)). Next Obligation. -move=>p t [g0][g][r][_][[s [->]] S Sel Dg C Rc Ep Er H G Pm D] /=. -case/In_marksX: Pm=>lks /[dup] /(In_blinks H) -> Pm. -have Uq : uniq (null :: s). -- admit. -have Ps : p \in s. -- by rewrite Ep last_change // -Ep (In_cond Pm). -rewrite Ep in Ps. -case/rcons_lastP: Ps=>s'. rewrite -Ep. -move=>E. subst s. -clear Ep. -have Nxs' : p \notin s'. -- simpl in Uq. rewrite rcons_uniq in Uq. - by case/and3P: Uq. - -case: (In_layX2 H Pm)=>h /[dup] E ->; do 3!step; move=>V. -split=>//; exists s'; split. -- by rewrite In_layX3 /= (In_dom Pm) upd_eta E freePtUn ?(validX V). -- move=>x. - rewrite /bupd (In_findE (In_marks Pm)) omfU ?(In_cond Pm) //=. - split. - - move=>B. - have Nxp : x != p. - - by case: eqP B Nxs'=>// ->->. - have : x \in rcons s' p. rewrite mem_rcons inE B orbT. by []. - move/S. case=>X; [left|right]. - - apply/InU=>/=. rewrite validU (In_cond Pm) (negbTE Nxp). rewrite pfVE (In_valid Pm). by []. - - apply/InU=>/=. rewrite validU (In_cond Pm) (negbTE Nxp). rewrite pfVE (In_valid Pm). by []. - rewrite !InU /=. - case; case; case: eqP=>// /eqP Nxp _ X; - suff : x \in rcons s' p by [rewrite mem_rcons inE (negbTE Nxp)]. - - by apply/S; left. - by apply/S; right. -k - - - - have : x != p /\ x \in s. - - - - - - move/mem_belast. - - - rewrite /bupd (In_findE (In_marks Pm)). - - - - -admit. -Admitted. +move=>p t [g0][g][r][_][[s [->]]] Ep Er /= /andP [U1 U2] +Sm Sc Gd Rc G2 G Pm D /=; case/In_labelsX: Pm=>lks /[dup] /(In_links2 G2) -> Pm. +(* prepare for popping p from the end of s to obtain ss *) +move: (In_cond Pm)=>/= Np; case/(rcons_lastN Np): Ep=>ss ?; subst s. +rewrite headI /= in Er; rewrite mem_rcons inE negb_or eq_sym Np /= in U1. +rewrite rcons_uniq in U2; case/andP: U2=>U2 U3. +case: (lay2_eta G2 Pm)=>h /[dup] E ->; do !step; move=>V. +(* the new stack is ss *) +split=>//; exists ss; split=>//=. +- by rewrite lay2CU (In_dom Pm) upd_eta E freePtUn ?(validX V). +- rewrite (Sc (last null ss) p RR) ?(In_omark Pm) ?consec_rcons //=. + by rewrite mem_rcons inE negb_or eq_sym Np U1 rcons_uniq U2. +- by rewrite U1 U3. +- move=>x; move: (Sm x); rewrite mem_rcons inE. + rewrite /omark/olabel !find_omf find_upd2 (In_dom Pm) /omfx/=. + by case: (x =P p) U2=>// -> /negbTE ->. +- move=>x y m C; move: (Sm y); rewrite mem_rcons inE. + rewrite sel2U /omark/olabel !find_omf find_upd2 (In_dom Pm) /=. + case: (y =P p)=>// /eqP Ny; case: (dom_find y)=>//= [][][] //=. + move=>k v /In_find H _ S [<-]; apply: Sc (In_omark H). + rewrite -rcons_cons consec_ext ?inE ?S ?orbT //. + by rewrite rcons_uniq /= inE negb_or Np U1 U2 U3. + + +UP TO HERE + + + + + Program Definition swing (p t : ptr): STsep {g0 g m r} (fun h => diff --git a/pcm/natmap.v b/pcm/natmap.v index d1d6d2c..a8bcb40 100644 --- a/pcm/natmap.v +++ b/pcm/natmap.v @@ -24,6 +24,7 @@ From mathcomp Require Import ssrnat eqtype seq path interval. From pcm Require Import axioms options prelude pred finmap rtc seqext. From pcm Require Export useqord uslice uconsec pcm morphism unionmap. Import order.Order.NatOrder. (* listed last to avoid notation clash *) +Local Open Scope order_scope. (************************) (* Maps over non-0 nats *) From 73ab03426c9a508abc2e21804484191f90c4dd85 Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Tue, 17 Sep 2024 15:02:06 +0200 Subject: [PATCH 79/93] preparing for release --- _CoqProject | 5 +++ core/finmap.v | 1 + core/pred.v | 2 +- core/prelude.v | 2 +- core/slice.v | 5 ++- meta.yml | 24 +++++------ pcm/morphism.v | 105 ++++++++++++++++++++++++++++--------------------- pcm/pcm.v | 1 - pcm/unionmap.v | 12 +++--- 9 files changed, 87 insertions(+), 70 deletions(-) diff --git a/_CoqProject b/_CoqProject index 908b8d5..b0ff611 100644 --- a/_CoqProject +++ b/_CoqProject @@ -7,6 +7,11 @@ -arg -w -arg -notation-overridden -arg -w -arg -redundant-canonical-projection +# release-specific arguments +-arg -w -arg -notation-incompatible-prefix # specific to coq8.20.0 +-arg -w -arg -deprecated-from-Coq # specific to coq8.21 +-arg -w -arg -deprecated-dirpath-Coq # specific to coq8.21 + core/options.v core/axioms.v core/prelude.v diff --git a/core/finmap.v b/core/finmap.v index 346e367..9ff1422 100644 --- a/core/finmap.v +++ b/core/finmap.v @@ -1297,6 +1297,7 @@ case: eqP=>// <-{k2}; case E1: (zip_f v1 v2)=>[w|//]. case E2: (zip' s1 s2)=>[t|//][<-{s}] /=. case: eqP=>[_ [<-]|_]. (* DEVCOMMENT: exists v1, v2 errors! *) +(* /DEVCOMMENT *) - by exists v1; exists v2. by case: (filter (predk V x) t) (IH _ _ E2). Qed. diff --git a/core/pred.v b/core/pred.v index e1c6a3e..fcf6ca1 100644 --- a/core/pred.v +++ b/core/pred.v @@ -1566,7 +1566,7 @@ 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. +Proof. by split. Qed. Lemma id_rel_trans : Transitive id_rel. Proof. by move=> x y z ->->. Qed. Lemma id_rel_func : functional id_rel. diff --git a/core/prelude.v b/core/prelude.v index cc5458a..380c43a 100644 --- a/core/prelude.v +++ b/core/prelude.v @@ -1280,7 +1280,7 @@ Lemma onth_tnth {n} (s : n.-tuple A) (i : 'I_n) : Proof. elim: n s i =>[|n IH] s i; first by case: i. case/tupleP: s=>/=x s; case: (unliftP ord0 i)=>[j|]-> /=. -- by rewrite tnthS. +- by rewrite tnthS ?add0n. by rewrite tnth0. Qed. diff --git a/core/slice.v b/core/slice.v index dca83f4..47958bb 100644 --- a/core/slice.v +++ b/core/slice.v @@ -15,7 +15,8 @@ From Coq Require Import ssreflect ssrbool ssrfun. From mathcomp Require Import ssrnat eqtype seq path interval order. From mathcomp Require Import fintype finfun tuple. From pcm Require Import options prelude seqext. -Local Open Scope order_scope. + +Open Scope order_scope. Import Order.Theory. (* DEVCOMMENT *) @@ -23,7 +24,7 @@ Import Order.Theory. (* but that changed with mathcomp 2.0 *) (* /DEVCOMMENT *) Section BSimp_Extension. -Variables (disp : unit) (T : porderType disp). +Context disp (T : porderType disp). Implicit Types (x y : T) (b c : bool). Lemma binf_inf b c : (BInfty T b == BInfty T c) = (b == c). diff --git a/meta.yml b/meta.yml index 3944bf1..f36d9a1 100644 --- a/meta.yml +++ b/meta.yml @@ -78,36 +78,32 @@ license: file: LICENSE supported_coq_versions: - text: Coq 8.15 to 8.17 - opam: '{ (>= "8.15" & < "8.19~") | (= "dev") }' + text: Coq 8.19 to 8.20 + opam: '{ (>= "8.19" & < "8.21~") | (= "dev") }' tested_coq_opam_versions: -- version: '1.17.0-coq-8.15' +- version: '2.2.0-coq-8.19' repo: 'mathcomp/mathcomp' -- version: '1.17.0-coq-8.18' +- version: '2.2.0-coq-8.20' + repo: 'mathcomp/mathcomp' +- version: 'latest-coq-dev' repo: 'mathcomp/mathcomp' -# - version: 'coq-dev' -# repo: 'mathcomp/mathcomp-dev' dependencies: - opam: name: coq-mathcomp-ssreflect - version: '{ (>= "1.17.0" & < "1.18~") | (= "dev") }' + version: '{ (>= "2.2.0" & < "2.3~") | (= "dev") }' description: |- - [MathComp ssreflect 1.17](https://math-comp.github.io) + [MathComp ssreflect 2.2](https://math-comp.github.io) - opam: name: coq-mathcomp-algebra description: |- [MathComp algebra](https://math-comp.github.io) -- opam: - name: coq-mathcomp-fingroup - description: |- - [MathComp fingroup](https://math-comp.github.io) - opam: name: coq-fcsl-pcm - version: '{ (>= "1.8.0" & < "1.9~") | (= "dev") }' + version: '{ (>= "2.0.0" & < "2.1~") | (= "dev") }' description: |- - [FCSL-PCM 1.8](https://github.com/imdea-software/fcsl-pcm) + [FCSL-PCM 2.0](https://github.com/imdea-software/fcsl-pcm) namespace: htt diff --git a/pcm/morphism.v b/pcm/morphism.v index b6a4a8b..431fa91 100644 --- a/pcm/morphism.v +++ b/pcm/morphism.v @@ -1571,25 +1571,37 @@ HB.instance Definition _ := isFull_PCM_morphism.Build U W (g \o f) comp_is_full_morphism. End FullCompIsFull. -(* instances for combinations must be declared explicitly *) -(* DEVCOMMENT *) -(* this seems unintended, should check with mathcomp people *) -(* /DEVCOMMENT *) - +(* instances for combinations must declare PCM_morphism.on *) +(* before declaring fullness structure. *) +(* DEVCOMMENT: *) +(* Alternative is to write the definition explicitly as below, *) +(* but that's too low level. *) +(* HB.instance Definition _ (f : full_norm_pcm_morph U V) (g : full_norm_pcm_morph V W) := Full_Norm_PCM_morphism.copy (g \o f) (Full_Norm_PCM_morphism.pack_ (Norm_PCM_morphism.class (g \o f)) - (Full_PCM_morphism.class (g \o f))). + (Full_PCM_morphism.class (g \o f))). *) +(* /DEVCOMMENT *) -HB.instance Definition _ (f : full_binorm_pcm_morph U V) - (g : full_binorm_pcm_morph V W) := - Full_Binorm_PCM_morphism.copy (g \o f) - (Full_Binorm_PCM_morphism.pack_ - (Binorm_PCM_morphism.class (g \o f)) - (Norm_PCM_morphism.class (g \o f)) - (Full_PCM_morphism.class (g \o f))). +HB.instance Definition _ + (f : full_norm_pcm_morph U V) + (g : full_norm_pcm_morph V W) := + PCM_morphism.on (g \o f). +HB.instance Definition _ + (f : full_norm_pcm_morph U V) + (g : full_norm_pcm_morph V W) := + Full_Norm_PCM_morphism.on (g \o f). + +HB.instance Definition _ + (f : full_binorm_pcm_morph U V) + (g : full_binorm_pcm_morph V W) := + PCM_morphism.on (g \o f). +HB.instance Definition _ + (f : full_binorm_pcm_morph U V) + (g : full_binorm_pcm_morph V W) := + Full_Binorm_PCM_morphism.on (g \o f). End CompMorphism. @@ -1812,46 +1824,51 @@ HB.instance Definition _ := End CompTPCM. (* combinations declared explicitly *) -(* DEVCOMMENT: boilerplate -- could it be automated? *) HB.instance Definition _ - (f : norm_tpcm_morph U V) (g : norm_tpcm_morph V W) := - Norm_TPCM_morphism.copy (g \o f) - (Norm_TPCM_morphism.pack_ - (TPCM_morphism.class (g \o f)) - (Norm_PCM_morphism.class (g \o f))). - + (f : norm_tpcm_morph U V) + (g : norm_tpcm_morph V W) := + PCM_morphism.on (g \o f). HB.instance Definition _ - (f : binorm_tpcm_morph U V) (g : binorm_tpcm_morph V W) := - Binorm_TPCM_morphism.copy (g \o f) - (Binorm_TPCM_morphism.pack_ - (TPCM_morphism.class (g \o f)) - (Binorm_PCM_morphism.class (g \o f)) - (Norm_PCM_morphism.class (g \o f))). + (f : norm_tpcm_morph U V) + (g : norm_tpcm_morph V W) := + Norm_TPCM_morphism.on (g \o f). HB.instance Definition _ - (f : full_tpcm_morph U V) (g : full_tpcm_morph V W) := - Full_TPCM_morphism.copy (g \o f) - (Full_TPCM_morphism.pack_ - (TPCM_morphism.class (g \o f)) - (Full_PCM_morphism.class (g \o f))). + (f : binorm_tpcm_morph U V) + (g : binorm_tpcm_morph V W) := + PCM_morphism.on (g \o f). +HB.instance Definition _ + (f : binorm_tpcm_morph U V) + (g : binorm_tpcm_morph V W) := + Binorm_TPCM_morphism.on (g \o f). HB.instance Definition _ - (f : full_norm_tpcm_morph U V) (g : full_norm_tpcm_morph V W) := - Full_Norm_TPCM_morphism.copy (g \o f) - (Full_Norm_TPCM_morphism.pack_ - (TPCM_morphism.class (g \o f)) - (Full_PCM_morphism.class (g \o f)) - (Norm_PCM_morphism.class (g \o f))). + (f : full_tpcm_morph U V) + (g : full_tpcm_morph V W) := + PCM_morphism.on (g \o f). +HB.instance Definition _ + (f : full_tpcm_morph U V) + (g : full_tpcm_morph V W) := + Full_TPCM_morphism.on (g \o f). HB.instance Definition _ - (f : full_binorm_tpcm_morph U V) (g : full_binorm_tpcm_morph V W) := - Full_Binorm_TPCM_morphism.copy (g \o f) - (Full_Binorm_TPCM_morphism.pack_ - (TPCM_morphism.class (g \o f)) - (Full_PCM_morphism.class (g \o f)) - (Binorm_PCM_morphism.class (g \o f)) - (Norm_PCM_morphism.class (g \o f))). + (f : full_norm_tpcm_morph U V) + (g : full_norm_tpcm_morph V W) := + PCM_morphism.on (g \o f). +HB.instance Definition _ + (f : full_norm_tpcm_morph U V) + (g : full_norm_tpcm_morph V W) := + Full_Norm_TPCM_morphism.on (g \o f). + +HB.instance Definition _ + (f : full_binorm_tpcm_morph U V) + (g : full_binorm_tpcm_morph V W) := + PCM_morphism.on (g \o f). +HB.instance Definition _ + (f : full_binorm_tpcm_morph U V) + (g : full_binorm_tpcm_morph V W) := + Full_Binorm_TPCM_morphism.on (g \o f). End Comp. diff --git a/pcm/pcm.v b/pcm/pcm.v index 98e204b..ba2189d 100644 --- a/pcm/pcm.v +++ b/pcm/pcm.v @@ -253,7 +253,6 @@ Proof. by case: (@tpcm_subproof U). Qed. Lemma undef_join (U : tpcm) (x : U) : undef \+ x = undef. Proof. by case: (@tpcm_subproof U). Qed. - Section TPCMLemmas. Variable U : tpcm. diff --git a/pcm/unionmap.v b/pcm/unionmap.v index 5d76e97..ae16212 100644 --- a/pcm/unionmap.v +++ b/pcm/unionmap.v @@ -543,7 +543,6 @@ Abort. (* /DEVCOMMENT *) - (***************) (* Finite sets *) (***************) @@ -589,7 +588,7 @@ Notation "# x" := (ptsT (fset _) x tt) (at level 30, format "# x"). (* DEVCOMMENT *) (* test *) -(* /DECOMMENT *) +(* /DEVCOMMENT *) Lemma xx : #1 \+ #2 = Unit. Abort. @@ -937,10 +936,10 @@ Proof. by rewrite validUnHE has_sym. Qed. Lemma validFUn 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 ->. -by rewrite validF V1. -by rewrite V2. +case: validUn=>// V1 V2 H _; case: validUn=>//. +- by rewrite validF V1. +- by rewrite V2. +by move=>k'; rewrite domF inE; case: eqP=>// _ /H/negbTE ->. Qed. Lemma validUnF k f1 f2 : @@ -4634,7 +4633,6 @@ Proof. exact: mapv_omapvE. Qed. End OmapMembershipLemmas. - (************************) (* PCM-induced ordering *) (************************) From cbd2c96d188fdd0ec295c031c4d2ed1603b4d1d7 Mon Sep 17 00:00:00 2001 From: Alekandar Nanevski Date: Mon, 23 Sep 2024 02:43:56 +0200 Subject: [PATCH 80/93] removed files inheritted from fcsl-pcm --- _CoqProject | 26 - core/auto.v | 102 - core/axioms.v | 135 - core/finmap.v | 1356 --------- core/options.v | 3 - core/ordtype.v | 497 ---- core/pred.v | 1833 ------------ core/prelude.v | 1391 --------- core/rtc.v | 273 -- core/seqext.v | 2138 -------------- core/seqperm.v | 253 -- core/slice.v | 727 ----- core/uconsec.v | 1105 ------- core/useqord.v | 578 ---- core/uslice.v | 1005 ------- examples/graph.v | 1509 ---------- examples/schorr.v | 264 -- pcm/automap.v | 764 ----- pcm/autopcm.v | 356 --- pcm/heap.v | 552 ---- pcm/invertible.v | 124 - pcm/morphism.v | 3341 --------------------- pcm/mutex.v | 583 ---- pcm/natmap.v | 2860 ------------------ pcm/pcm.v | 1660 ----------- pcm/unionmap.v | 7028 --------------------------------------------- 26 files changed, 30463 deletions(-) delete mode 100644 core/auto.v delete mode 100644 core/axioms.v delete mode 100644 core/finmap.v delete mode 100644 core/options.v delete mode 100644 core/ordtype.v delete mode 100644 core/pred.v delete mode 100644 core/prelude.v delete mode 100644 core/rtc.v delete mode 100644 core/seqext.v delete mode 100644 core/seqperm.v delete mode 100644 core/slice.v delete mode 100644 core/uconsec.v delete mode 100644 core/useqord.v delete mode 100644 core/uslice.v delete mode 100644 examples/graph.v delete mode 100644 examples/schorr.v delete mode 100644 pcm/automap.v delete mode 100644 pcm/autopcm.v delete mode 100644 pcm/heap.v delete mode 100644 pcm/invertible.v delete mode 100644 pcm/morphism.v delete mode 100644 pcm/mutex.v delete mode 100644 pcm/natmap.v delete mode 100644 pcm/pcm.v delete mode 100644 pcm/unionmap.v diff --git a/_CoqProject b/_CoqProject index b0ff611..f4eab2f 100644 --- a/_CoqProject +++ b/_CoqProject @@ -1,5 +1,3 @@ --Q core pcm --Q pcm pcm -Q examples htt -Q htt htt -docroot docs # where the documentation should go @@ -12,29 +10,6 @@ -arg -w -arg -deprecated-from-Coq # specific to coq8.21 -arg -w -arg -deprecated-dirpath-Coq # specific to coq8.21 -core/options.v -core/axioms.v -core/prelude.v -core/pred.v -core/auto.v -core/seqext.v -core/slice.v -core/uslice.v -core/useqord.v -core/uconsec.v -core/ordtype.v -core/seqperm.v -core/finmap.v -core/rtc.v -pcm/pcm.v -pcm/autopcm.v -pcm/morphism.v -pcm/invertible.v -pcm/unionmap.v -pcm/natmap.v -pcm/automap.v -pcm/heap.v -pcm/mutex.v htt/options.v htt/domain.v htt/model.v @@ -58,4 +33,3 @@ examples/congmath.v examples/congprog.v examples/tree.v examples/union_find.v -examples/graph.v diff --git a/core/auto.v b/core/auto.v deleted file mode 100644 index dd77e11..0000000 --- a/core/auto.v +++ /dev/null @@ -1,102 +0,0 @@ -(* -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. -*) - -From Coq Require Import ssreflect ssrbool ssrfun. -From mathcomp Require Import ssrnat seq eqtype. -From pcm Require Import options prelude. - -(***************************************************************************) -(* Common infrastructure for syntactifying expressions in automated lemmas *) -(***************************************************************************) - -(* In cancellation algorithms for automated lemmas, we iterate over the *) -(* first expression e1 and remove each of its components from the second *) -(* expression e2. But, typically, we want to remove only one occurrence *) -(* of that component. *) -(* *) -(* First, almost always, only one occurrence will exist, lest e2 be *) -(* undefined. Thus, it's a waste of effort to traverse e2 in full once *) -(* we've found an occurrence. *) -(* *) -(* Second, in some symmetric cancellation problems, e.g., dom_eq e1 e2, *) -(* we *want* to remove only one occurrence from e2 for each component in *) -(* e1. Otherwise, we will not produce a sound reduction. E.g., *) -(* dom (x \+ x) (x \+ x) is valid, since both expressions are undef. *) -(* However, after removing x from the left side, and both x's from the *) -(* right side, we get dom x Unit, which is not valid. *) -(* *) -(* Thus, as a matter of principle, we want a filter that removes just one *) -(* occurrence of a list element. *) -(* *) -(* We write it with p pulled out in a section in order to get it to *) -(* simplify automatically. *) - -Section OneShotFilter. -Variables (A : Type) (p : pred A). - -(* a variant of filter that removes only the first occurence *) - -Fixpoint rfilter {A} (p : pred A) (ts : seq A) : seq A := - if ts is t :: ts' then if p t then ts' else t :: rfilter p ts' else [::]. - -End OneShotFilter. - -(* rfilter can also be thought of as a generalization of rem *) -Lemma rfilter_rem {T : eqType} (ts : seq T) x : - rfilter (pred1 x) ts = rem x ts. -Proof. by elim: ts=> [|t ts IH] //=; case: eqP=>//= _; rewrite IH. Qed. - -(* A canonical structure program for searching and inserting in a list *) - -Section XFind. -Variable A : Type. - -Structure tagged_elem := XTag {xuntag :> A}. - -(* DEVCOMMENT *) -(* remove? *) -(* Local Coercion untag : tagged_elem >-> A. *) -(* /DEVCOMMENT *) - -Definition extend_tag := XTag. -Definition recurse_tag := extend_tag. -Canonical found_tag x := recurse_tag x. - -(* Main structure: *) -(* - xs1 : input sequence *) -(* - xs2 : output sequence; if pivot is found, then xs2 = xs1, else *) -(* xs2 = pivot :: xs1 *) -(* - i : output index of pivot in xs2 *) - -Definition axiom xs1 xs2 i (pivot : tagged_elem) := - onth xs2 i = Some (xuntag pivot) /\ Prefix xs1 xs2. -Structure xfind (xs1 xs2 : seq A) (i : nat) := - Form {pivot :> tagged_elem; _ : axiom xs1 xs2 i pivot}. - -(* found the elements *) -Lemma found_pf x t : axiom (x :: t) (x :: t) 0 (found_tag x). -Proof. by []. Qed. -Canonical found_form x t := Form (@found_pf x t). - -(* recurse *) -Lemma recurse_pf i x xs1 xs2 (f : xfind xs1 xs2 i) : - axiom (x :: xs1) (x :: xs2) i.+1 (recurse_tag (xuntag f)). -Proof. by case: f=>pv [H1 H2]; split=>//; apply/Prefix_cons. Qed. -Canonical recurse_form i x xs1 xs2 f := Form (@recurse_pf i x xs1 xs2 f). - -(* failed to find; attach the element to output *) -Lemma extend_pf x : axiom [::] [:: x] 0 (extend_tag x). -Proof. by []. Qed. -Canonical extend_form x := Form (@extend_pf x). - -End XFind. diff --git a/core/axioms.v b/core/axioms.v deleted file mode 100644 index 1154652..0000000 --- a/core/axioms.v +++ /dev/null @@ -1,135 +0,0 @@ -(* -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 axioms that are used in some parts of the library. *) -(* The selected set of axioms is known to be consistent with Coq's logic. *) -(* These axioms are: *) -(* - propositional extensionality (pext); *) -(* - functional extensionality (fext). *) -(* This file also defines the dynamic type as an alias for sigT and *) -(* Jonh Major equality via equality cast. *) -(******************************************************************************) - -From Coq Require Import ssreflect ssrfun Eqdep ClassicalFacts. -From mathcomp Require Import eqtype. -From pcm Require Import options. - -(*****************************) -(* Axioms and extensionality *) -(*****************************) - -(* We're additionally using the eq_rect_eq axiom (equivalent to UIP) from - Coq.Logic.Eqdep for its two consequences: inj_pair2 and StreicherK *) - -(* extensionality is needed for domains *) -Axiom pext : forall p1 p2 : Prop, (p1 <-> p2) -> p1 = p2. -Axiom fext : forall A (B : A -> Type) (f1 f2 : forall x, B x), - (forall x, f1 x = f2 x) -> f1 = f2. - -Lemma pf_irr (P : Prop) (p1 p2 : P) : p1 = p2. -Proof. by apply/ext_prop_dep_proof_irrel_cic/@pext. Qed. - -Lemma sval_inj A P : injective (@sval A P). -Proof. -move=>[x Hx][y Hy] /= H; move: Hx Hy; rewrite H=>*. -congr exist; apply: pf_irr. -Qed. - -Lemma svalE A (P : A -> Prop) x H : sval (exist P x H) = x. -Proof. by []. Qed. - -Lemma compf1 A B (f : A -> B) : f = f \o id. -Proof. by apply: fext. Qed. - -Lemma comp1f A B (f : A -> B) : f = id \o f. -Proof. by apply: fext. Qed. - -(*****************************************) -(* Cast and John Major Equality via cast *) -(*****************************************) - -(* depends on StreicherK axiom *) - -Section Cast. -Variable (T : Type) (interp : T -> Type). - -Definition cast A B (pf : A = B) (v : interp B) : interp A := - ecast _ _ (esym pf) v. - -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) := exists pf, v = cast pf w. - -Lemma jm_refl A (v : interp A) : jmeq v v. -Proof. by exists (erefl _); 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=>[[?]|] ->; [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 subst B; rewrite !eqc. Qed. - -End Cast. - -Arguments cast {T} interp [A][B] pf v. -Arguments jmeq {T} interp [A][B] v w. - -#[export] 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). - -(* in case of eqTypes StreicherK not needed *) -Section EqTypeCast. -Variable (T : eqType) (interp : T -> Type). -Lemma eqd a (pf : a = a) (v : interp a) : cast interp pf v = v. -Proof. by rewrite eq_axiomK. Qed. -End EqTypeCast. - - -(* type dynamic is sigT *) - -Section Dynamic. -Variables (A : Type) (P : A -> Type). - -(** eta expand definitions to prevent universe inconsistencies when using - the injectivity of constructors of datatypes depending on [[dynamic]] *) - -Definition dynamic := sigT P. -Definition dyn := existT P. -Definition dyn_tp := @projT1 _ P. -Definition dyn_val := @projT2 _ P. -Definition dyn_eta := @sigT_eta _ P. -Definition dyn_injT := @eq_sigT_fst _ P. -Definition dyn_inj := @inj_pair2 _ P. - -End Dynamic. - -Prenex Implicits dyn_tp dyn_val dyn_injT dyn_inj. -Arguments dyn {T} interp {A} _ : rename. -Notation idyn v := (@dyn _ id _ v). - -Lemma dynE (A B : Type) interp (v : interp A) (w : interp B) : - jmeq interp v w <-> dyn interp v = dyn interp w. -Proof. by split=>[[pf ->]|[pf]]; subst B; [rewrite !eqc | move/dyn_inj=>->]. Qed. diff --git a/core/finmap.v b/core/finmap.v deleted file mode 100644 index 9ff1422..0000000 --- a/core/finmap.v +++ /dev/null @@ -1,1356 +0,0 @@ -(* -Copyright 2009 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 finitely supported maps with keys drawn from *) -(* an ordered type and values from an arbitrary type. *) -(******************************************************************) - -From HB Require Import structures. -From Coq Require Import ssreflect ssrbool ssrfun. -From Coq Require Setoid. -From mathcomp Require Import ssrnat eqtype seq path. -From pcm Require Export ordtype seqperm. -From pcm Require Import options. - -Section Def. -Variables (K : ordType) (V : Type). - -Definition key (x : K * V) := x.1. -Definition value (x : K * V) := x.2. -Definition predk k := preim key (pred1 k). -Definition predCk k := preim key (predC1 k). - -Record finMap := FinMap { - seq_of : seq (K * V); - _ : sorted ord (map key seq_of)}. - -Definition finMap_for of phant (K -> V) := finMap. - -Identity Coercion finMap_for_finMap : finMap_for >-> finMap. -End Def. - -Notation "{ 'finMap' fT }" := (finMap_for (Phant fT)) - (at level 0, format "{ 'finMap' '[hv' fT ']' }") : type_scope. - -Prenex Implicits key value predk predCk seq_of. - -Section Ops. -Variables (K : ordType) (V : Type). -Notation fmap := (finMap K V). -Notation key := (@key K V). -Notation value := (@value K V). -Notation predk := (@predk K V). -Notation predCk := (@predCk K V). - -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 (eq_irrelevance H1 H2). -Qed. - -Lemma sorted_nil : sorted ord (map key [::]). Proof. by []. Qed. -Definition nil := FinMap sorted_nil. - -Definition fnd k (s : fmap) := - if (filter (predk k) (seq_of s)) is (_, v):: _ - then Some v else None. - -Fixpoint ins' (k : K) (v : V) (s : seq (K * V)) {struct s} : seq (K * V) := - if s is (k1, v1)::s1 then - if ord k k1 then (k, v)::s else - if k == k1 then (k, v)::s1 else (k1, v1)::(ins' k v s1) - else [:: (k, v)]. - -Lemma path_ins' s k1 k2 v : - ord k1 k2 -> - path ord k1 (map key s) -> - path ord k1 (map key (ins' k2 v s)). -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: (ord_total k2 k'); rewrite H4 H5. -Qed. - -Lemma sorted_ins' s k v : - sorted ord (map key s) -> - sorted ord (map key (ins' k v s)). -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: (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'). - -Lemma sorted_filter' k s : - sorted ord (map key s) -> - sorted ord (map key (filter (predCk k) s)). -Proof. by move=>H; rewrite -filter_map sorted_filter //; apply: trans. Qed. - -Definition rem k s := let: FinMap s' p' := s in FinMap (sorted_filter' k p'). - -Lemma sorted_behd s : - sorted ord (map key s) -> - sorted ord (map key (behead s)). -Proof. by case: s=>//= [[??]] ?; apply: path_sorted. Qed. - -Definition behd s := let: FinMap s' p' := s in FinMap (sorted_behd p'). -Definition supp s := map key (seq_of s). -Definition cosupp s := map value (seq_of s). - -End Ops. - -Arguments nil {K V}. -Prenex Implicits fnd ins rem supp nil. - -Section Laws. -Variables (K : ordType) (V : Type). -Notation fmap := (finMap K V). -Notation nil := (@nil K V). - -(* `path_le` specialized to `transitive ord` *) -Lemma ord_path (x y : K) s : - ord x y -> - path ord y s -> - path ord x s. -Proof. by apply: path_le. Qed. - -Lemma last_ins' (x : K) (v : V) s : - path ord x (map key s) -> - ins' x v s = (x, v) :: s. -Proof. by elim: s=>[|[k1 v1] s IH] //=; case: ifP. Qed. - -Lemma first_ins' (k : K) (v : V) s : - (forall x, x \in map key s -> ord x k) -> - ins' k v s = rcons s (k, v). -Proof. -elim: s=>[|[k1 v1] s IH] H //=. -move: (H k1); rewrite inE eq_refl; move/(_ (erefl _)). -case: ordP=>// O _; rewrite IH //. -by move=>x H'; apply: H; rewrite inE /= H' orbT. -Qed. - -Lemma notin_path (x : K) s : - path ord x s -> - x \notin s. -Proof. -elim: s=>[|k s IH] //=. -rewrite inE negb_or; case/andP=>T1 T2; case: eqP=>H /=. -- by rewrite H irr in T1. -by apply: IH; apply: ord_path T2. -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: 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. - -Lemma all_path_supp (s : fmap) k : - all (ord k) (supp s) -> - path ord k (supp s). -Proof. by rewrite path_sortedE // =>->/=; case: s. Qed. - -Lemma notin_filter (x : K) s : - x \notin (map key s) -> filter (predk V x) s = [::]. -Proof. -elim: s=>[|[k v] s IH] //=. -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 /=. -elim: s1 P1 s2 P2 H=>[|[k v] s1 IH] /= P1. -- by case=>[|[k2 v2] s2 P2] //=; move/(_ k2); rewrite eq_refl. -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 H. -move: (notin_path P1) (notin_path P2)=>N1 N2. -case E: (k == k2). -- rewrite -{k2 E}(eqP E) in P2 H N2 *. - move: (H k); rewrite eq_refl=>[[E2]]; rewrite -E2 {v2 E2} in H *. - rewrite IH // => k'. - 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: (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. -Qed. - -Lemma predkN (k1 k2 : K) : predI (predk V k1) (predCk V k2) =1 - if k1 == k2 then pred0 else predk V k1. -Proof. -by move=>x; case: ifP=>H /=; [|case: eqP=>//->]; rewrite ?(eqP H) ?andbN ?H. -Qed. - -Variant supp_spec x (s : fmap) : bool -> Type := -| supp_spec_some v of fnd x s = Some v : supp_spec x s true -| supp_spec_none of fnd x s = None : supp_spec x s false. - -Lemma suppP x (s : fmap) : supp_spec x s (x \in supp s). -Proof. -move E: (x \in supp s)=>b; case: b E; move/idP; last first. -- move=>H; apply: supp_spec_none. - case E: (fnd _ _)=>[v|] //; case: H. - rewrite /supp /fnd in E *; case: s E=>/=. - elim=>[|[y w] s IH] H1 //=. - case: ifP=>H; first by rewrite (eqP H) inE eq_refl. - rewrite -topredE /= eq_sym H; apply: IH. - by apply: path_sorted H1. -case: s; elim=>[|[y w] s IH] /= H1 //; rewrite /supp /= inE in IH *. -case: eqP=>[-> _|H] //=. -- by apply: (@supp_spec_some _ _ w); rewrite /fnd /= eq_refl. -move: (path_sorted H1)=>H1'; move/(IH H1'); case=>[v H2|H2]; -[apply: (@supp_spec_some _ _ v) | apply: supp_spec_none]; -by rewrite /fnd /=; case: eqP H=>// ->. -Qed. - -Lemma suppE (s1 s2 : fmap) : supp s1 =i supp s2 <-> supp s1 = supp s2. -Proof. -split; last by move=>->. -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: (ordP k1 k2)=>/= O H1 H2. -- move/(_ k1); rewrite !inE eq_refl /= eq_sym. - 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. - - by apply: path_sorted H1. - - by apply: path_sorted 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: ordP O=>//= O _. -by move/(ord_path O)/notin_path/negbTE: H1=>->. -Qed. - -Lemma supp_nil : supp nil = [::]. Proof. by []. Qed. - -Lemma supp_nilE (s : fmap) : (supp s = [::]) <-> (s = nil). -Proof. by split=>[|-> //]; case: s; case=>// H; rewrite fmapE. Qed. - -Lemma supp_rem k (s : fmap) : - supp (rem k s) =i predI (predC1 k) (mem (supp s)). -Proof. -case: s => s H1 x; rewrite /supp inE /=. -by case H2: (x == k)=>/=; rewrite -filter_map mem_filter /= H2. -Qed. - -Lemma supp_ins k v (s : fmap) : - supp (ins k v s) =i [predU pred1 k & supp s]. -Proof. -case: s => s H x; rewrite /supp inE /=. -elim: s x k v H=>[|[k1 v1] s IH] //= x k v H. -case: ifP=>H1 /=; first by rewrite inE. -case: ifP=>H2 /=; first by rewrite !inE (eqP H2) orbA orbb. -by rewrite !inE (IH _ _ _ (path_sorted H)) orbCA. -Qed. - -Lemma fnd_empty k : fnd k nil = None. Proof. by []. Qed. - -Lemma fnd_rem k1 k2 (s : fmap) : - fnd k1 (rem k2 s) = if k1 == k2 then None else fnd k1 s. -Proof. -case: s => s H; rewrite /fnd -filter_predI (eq_filter (predkN k1 k2)). -by case: eqP=>//; rewrite filter_pred0. -Qed. - -Lemma fnd_ins k1 k2 v (s : fmap) : - fnd k1 (ins k2 v s) = if k1 == k2 then Some v else fnd k1 s. -Proof. -case: s => s H; rewrite /fnd /=. -elim: s k1 k2 v H=>[|[k' v'] s IH] //= k1 k2 v H. -- by case: ifP=>H1; [rewrite (eqP H1) eq_refl | rewrite eq_sym H1]. -case: ifP=>H1 /=. -- by case: ifP=>H2; [rewrite (eqP H2) eq_refl | rewrite (eq_sym k1) H2]. -case: ifP=>H2 /=. -- rewrite (eqP H2). - by case: ifP=>H3; [rewrite (eqP H3) eq_refl | rewrite eq_sym H3]. -case: ifP=>H3; first by rewrite -(eqP H3) eq_sym H2. -by apply: IH; apply: path_sorted H. -Qed. - -Lemma ins_rem k1 k2 v (s : fmap) : - ins k1 v (rem k2 s) = - if k1 == k2 then ins k1 v s else rem k2 (ins k1 v s). -Proof. -move: k1 k2 v s. -have L3: forall (x : K) s, - path ord x (map key s) -> filter (predCk V x) s = s. -- move=>x t; move/notin_path; elim: t=>[|[k3 v3] t IH] //=. - rewrite inE negb_or; case/andP=>T1 T2. - by rewrite eq_sym T1 IH. -have L5: forall (x : K) (v : V) s, - sorted ord (map key s) -> ins' x v (filter (predCk V x) s) = ins' x v s. -- move=>x v s; elim: s x v=>[|[k' v'] s IH] x v //= H. - case H1: (ord x k'). - - case H2: (k' == x)=>/=; first by rewrite (eqP H2) irr in H1. - by rewrite H1 L3 //; apply: ord_path H1 H. - case H2: (k' == x)=>/=. - - rewrite (eqP H2) eq_refl in H *. - by rewrite L3 //; apply: last_ins' H. - rewrite eq_sym H2 H1 IH //. - by apply: path_sorted H. -move=>k1 k2 v [s H]. -case: ifP=>H1; rewrite /ins /rem fmapE /=. -- rewrite {k1 H1}(eqP H1). - elim: s k2 v H=>[|[k' v'] s IH] //= k2 v H. - case H1: (k' == k2)=>/=. - - rewrite eq_sym H1 (eqP H1) irr in H *. - by rewrite L3 // last_ins'. - rewrite eq_sym H1; case: ifP=>H3. - - by rewrite L3 //; apply: ord_path H3 H. - by rewrite L5 //; apply: path_sorted H. -elim: s k1 k2 H1 H=>[|[k' v'] s IH] //= k1 k2 H1 H; first by rewrite H1. -case H2: (k' == k2)=>/=. -- rewrite (eqP H2) in H *; rewrite H1. - case H3: (ord k1 k2)=>/=. - - by rewrite H1 eq_refl /= last_ins' // L3 //; apply: ord_path H. - by rewrite eq_refl /= IH //; apply: path_sorted H. -case H3: (ord k1 k')=>/=; first by rewrite H1 H2. -case H4: (k1 == k')=>/=; first by rewrite H1. -by rewrite H2 IH //; apply: path_sorted H. -Qed. - -Lemma ins_ins k1 k2 v1 v2 (s : fmap) : - ins k1 v1 (ins k2 v2 s) = if k1 == k2 then ins k1 v1 s - else ins k2 v2 (ins k1 v1 s). -Proof. -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: (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: (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: (ordP k2 k3) H2 H3. - by case: (ordP k1 k3) (trans H3 H2) H4. -- rewrite -(eqP H2) {H2} (H1). - case: (ordP k1 k2) (H1)=>//= H2 _; rewrite ?irr ?eq_refl //. - rewrite eq_sym H1. - by case: (ordP k2 k1) H1 H2. -case: (ordP k1 k3)=>H3 /=. -- rewrite eq_sym H1. - case: (ordP k2 k1) H1 (trans H3 H2)=>//. - by case: (ordP k2 k3) H2=>//=. -- rewrite (eqP H3). - 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. -Qed. - -Lemma rem_empty k : rem k nil = nil. -Proof. by rewrite fmapE. Qed. - -Lemma rem_rem k1 k2 (s : fmap) : - rem k1 (rem k2 s) = if k1 == k2 then rem k1 s else rem k2 (rem k1 s). -Proof. -rewrite /rem; case: s => s H /=. -case H1: (k1 == k2); rewrite fmapE /= -!filter_predI; apply: eq_filter=>x /=. -- by rewrite (eqP H1) andbb. -by rewrite andbC. -Qed. - -Lemma rem_ins k1 k2 v (s : fmap) : - rem k1 (ins k2 v s) = - if k1 == k2 then rem k1 s else ins k2 v (rem k1 s). -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: (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: (ordP k2 k3)=>H2 /=. -- rewrite eq_sym H1 /=. - 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: (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: (ordP k2 k3) H2=>//= H2 _. -by rewrite IH //; apply: path_sorted H. -Qed. - -Lemma rem_supp k (s : fmap) : - k \notin supp s -> rem k s = s. -Proof. -case: s => s H1; rewrite /supp !fmapE /= => H2. -elim: s H1 H2=>[|[k1 v1] s1 IH] //=; move/path_sorted=>H1. -rewrite inE negb_or; case/andP=>H2; move/(IH H1)=>H3. -by rewrite eq_sym H2 H3. -Qed. - -Lemma fnd_supp k (s : fmap) : - k \notin supp s -> fnd k s = None. -Proof. by case: suppP. Qed. - -Lemma fnd_supp_in k (s : fmap) : - k \in supp s -> exists v, fnd k s = Some v. -Proof. by case: suppP=>[v|]; [exists v|]. Qed. - -Lemma cancel_ins k v (s1 s2 : fmap) : - k \notin (supp s1) -> k \notin (supp s2) -> - ins k v s1 = ins k v s2 -> s1 = s2. -Proof. -move: s1 s2=>[s1 p1][s2 p2]; rewrite !fmapE /supp /= {p1 p2}. -elim: s1 k v s2=>[k v s2| [k1 v1] s1 IH1 k v s2] /=. -- case: s2=>[| [k2 v2] s2] //= _. - rewrite inE negb_or; case/andP=>H1 _; case: ifP=>// _. - by rewrite (negbTE H1); case=>E; rewrite E eq_refl in H1. -rewrite inE negb_or; case/andP=>H1 H2 H3. -case: ifP=>H4; case: s2 H3=>[| [k2 v2] s2] //=. -- rewrite inE negb_or; case/andP=>H5 H6. - case: ifP=>H7; first by case=>->->->. - by rewrite (negbTE H5); case=>E; rewrite E eq_refl in H5. -- by rewrite (negbTE H1)=>_; case=>E; rewrite E eq_refl in H1. -rewrite inE negb_or (negbTE H1); case/andP=>H5 H6. -rewrite (negbTE H5); case: ifP=>H7 /=. -- by case=>E; rewrite E eq_refl in H1. -by case=>->-> H; congr (_ :: _); apply: IH1 H. -Qed. - -End Laws. - -Section Append. -Variable (K : ordType) (V : Type). -Notation fmap := (finMap K V). -Notation nil := (@nil K V). - -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. -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: (ordP k k2)=>H /=. -- by rewrite H2 H H4. -- by rewrite H2 (eqP H) H4. -rewrite H3 /=. -have H6: sorted ord (map key s) by apply: path_sorted H5. -by move: (IH H6 k2 k v H H4); case: s {IH p H4 H5} H6. -Qed. - -Lemma path_supp_ins_inv k1 k v (s : fmap) : - path ord k (supp s) -> path ord k1 (supp (ins k v s)) -> - ord k1 k && path ord k1 (supp s). -Proof. -case: s=>s p; rewrite /supp /= => H1; rewrite last_ins' //=. -by case/andP=>H2 H3; rewrite H2; apply: ord_path H3. -Qed. - -(* forward induction principle *) -Lemma fmap_ind' (P : fmap -> Prop) : - P nil -> - (forall k v s, path ord k (supp s) -> P s -> P (ins k v s)) -> - forall s, P s. -Proof. -move=>H1 H2; case; elim=>[|[k v] s IH] /= H. -- by rewrite (_ : FinMap _ = nil); last by rewrite fmapE. -have S: sorted ord (map key s) by apply: path_sorted H. -rewrite (_ : FinMap _ = ins k v (FinMap S)); last by rewrite fmapE /= last_ins'. -by apply: H2. -Qed. - -(* backward induction principle *) -Lemma fmap_ind'' (P : fmap -> Prop) : - P nil -> (forall k v s, (forall x, x \in supp s -> ord x k) -> - P s -> P (ins k v s)) -> - forall s, P s. -Proof. -move=>H1 H2; case; elim/last_ind=>[|s [k v] IH] /= H. -- by rewrite (_ : FinMap _ = nil); last by rewrite fmapE. -have Sb: subseq (map key s) (map key (rcons s (k, v))). -- by elim: s {IH H}=>[|x s IH] //=; rewrite eq_refl. -have S : sorted ord (map key s). -- by apply: subseq_sorted Sb H; apply: ordtype.trans. -have T : forall x : K, x \in map key s -> ord x k. -- elim: {IH Sb S} s H=>[|[k1 v1] s IH] //= L x. - rewrite inE; case/orP; last by apply: IH; apply: path_sorted L. - move/eqP=>->; elim: s {IH} L=>[|[x1 w1] s IH] /=; first by rewrite andbT. - by case/andP=>O /(ord_path O) /IH. -rewrite (_ : FinMap _ = ins k v (FinMap S)). -- by apply: H2 (IH _)=>x /T. -by rewrite fmapE /= first_ins'. -Qed. - -Fixpoint fcat' (s1 : fmap) (s2 : seq (K * V)) {struct s2} : fmap := - if s2 is (k, v)::t then fcat' (ins k v s1) t else s1. - -Definition fcat s1 s2 := fcat' s1 (seq_of s2). - -Lemma fcat_ins' k v s1 s2 : - k \notin (map key s2) -> fcat' (ins k v s1) s2 = ins k v (fcat' s1 s2). -Proof. -move=>H; elim: s2 k v s1 H=>[|[k2 v2] s2 IH] k1 v1 s1 //=. -rewrite inE negb_or; case/andP=>H1 H2. -by rewrite -IH // ins_ins eq_sym (negbTE H1). -Qed. - -Lemma fcat_nil' s : fcat' nil (seq_of s) = s. -Proof. -elim/fmap_ind': s=>[|k v s L IH] //=. -by rewrite seqof_ins //= (_ : FinMap _ = ins k v nil) // - fcat_ins' ?notin_path // IH. -Qed. - -Lemma fcat0s s : fcat nil s = s. Proof. by apply: fcat_nil'. Qed. -Lemma fcats0 s : fcat s nil = s. Proof. by []. Qed. - -Lemma fcat_inss k v s1 s2 : - k \notin supp s2 -> fcat (ins k v s1) s2 = ins k v (fcat s1 s2). -Proof. by case: s2=>s2 p2 H /=; apply: fcat_ins'. Qed. - -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: (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: (ordP k2 k1) H2 => // H2 _. -rewrite {1}/fcat seqof_ins //= fcat_ins' ?notin_path // IH ?notin_path //. -rewrite ins_ins; case: (ordP k2 k1) H2 => // H2 _; congr (ins _ _ _). -by rewrite -/(fcat s2 (ins k2 v2 s1)) IH. -Qed. - -Lemma fcat_rems k s1 s2 : - k \notin supp s2 -> fcat (rem k s1) s2 = rem k (fcat s1 s2). -Proof. -elim/fmap_ind': s2 k s1=>[|k2 v2 s2 H IH] k1 v1. -- by rewrite !fcats0. -rewrite supp_ins inE /= negb_or; case/andP=>H1 H2. -by rewrite fcat_sins IH // ins_rem eq_sym (negbTE H1) -fcat_sins. -Qed. - -Lemma fcat_srem k s1 s2 : - k \notin supp s1 -> fcat s1 (rem k s2) = rem k (fcat s1 s2). -Proof. -elim/fmap_ind': s2 k s1=>[|k2 v2 s2 H IH] k1 s1. -- rewrite rem_empty fcats0. - elim/fmap_ind': s1=>[|k3 v3 s3 H1 IH]; first by rewrite rem_empty. - rewrite supp_ins inE /= negb_or. - case/andP=>H2; move/IH=>E; rewrite {1}E . - by rewrite ins_rem eq_sym (negbTE H2). -move=>H1; rewrite fcat_sins rem_ins; case: ifP=>E. -- by rewrite rem_ins E IH. -by rewrite rem_ins E -IH // -fcat_sins. -Qed. - -Lemma fnd_fcat k s1 s2 : - fnd k (fcat s1 s2) = - if k \in supp s2 then fnd k s2 else fnd k s1. -Proof. -elim/fmap_ind': s2 k s1=>[|k2 v2 s2 H IH] k1 s1. -- by rewrite fcats0. -rewrite supp_ins inE /=; case: ifP; last first. -- move/negbT; rewrite negb_or; case/andP=>H1 H2. - by rewrite fcat_sins fnd_ins (negbTE H1) IH (negbTE H2). -case/orP; first by move/eqP=><-; rewrite fcat_sins !fnd_ins eq_refl. -move=>H1; rewrite fcat_sins !fnd_ins. -by case: ifP=>//; rewrite IH H1. -Qed. - -Lemma supp_fcat s1 s2 : supp (fcat s1 s2) =i [predU supp s1 & supp s2]. -Proof. -elim/fmap_ind': s2 s1=>[|k v s L IH] s1. -- by rewrite supp_nil fcats0 => x; rewrite inE /= orbF. -rewrite fcat_sins ?notin_path // => x. -rewrite supp_ins !inE /=. -case E: (x == k)=>/=. -- by rewrite supp_ins inE /= E orbT. -by rewrite IH supp_ins inE /= inE /= E. -Qed. - -End Append. - -(* an induction principle for pairs of finite maps with equal support *) - -Section FMapInd. -Variables (K : ordType) (V : Type). -Notation fmap := (finMap K V). -Notation nil := (@nil K V). - -Lemma supp_eq_ins (s1 s2 : fmap) k1 k2 v1 v2 : - path ord k1 (supp s1) -> path ord k2 (supp s2) -> - supp (ins k1 v1 s1) =i supp (ins k2 v2 s2) -> - k1 = k2 /\ supp s1 =i supp s2. -Proof. -move=>H1 H2 H; move: (H k1) (H k2). -rewrite !supp_ins !inE /= !eq_refl (eq_sym k2). -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). -rewrite !supp_ins !inE /=; case: eqP=>//= -> _. -by rewrite (negbTE (notin_path H1)) (negbTE (notin_path H2)). -Qed. - -Lemma fmap_ind2 (P : fmap -> fmap -> Prop) : - P nil nil -> - (forall k v1 v2 s1 s2, - path ord k (supp s1) -> path ord k (supp s2) -> - P s1 s2 -> P (ins k v1 s1) (ins k v2 s2)) -> - forall s1 s2, supp s1 =i supp s2 -> P s1 s2. -Proof. -move=>H1 H2; elim/fmap_ind'=>[|k1 v1 s1 T1 IH1]; -elim/fmap_ind'=>[|k2 v2 s2 T2 _] //. -- by move/(_ k2); rewrite supp_ins inE /= eq_refl supp_nil. -- by move/(_ k1); rewrite supp_ins inE /= eq_refl supp_nil. -by case/supp_eq_ins=>// E; rewrite -{k2}E in T2 *; move/IH1; apply: H2. -Qed. - -End FMapInd. - - -Section Filtering. -Variables (K : ordType) (V : Type). -Notation fmap := (finMap K V). -Notation nil := (@nil K V). - -Definition kfilter' (p : pred K) (s : fmap) := - filter (fun kv => p kv.1) (seq_of s). - -Lemma sorted_kfilter (p : pred K) s : sorted ord (map key (kfilter' p s)). -Proof. -by case: s=>s H; rewrite -filter_map path.sorted_filter //; apply: trans. -Qed. - -Definition kfilter (p : pred K) (s : fmap) := FinMap (sorted_kfilter p s). - -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 H. -by case E: (p k)=>//=; rewrite H. -Qed. - -Lemma kfilt_nil (p : pred K) : kfilter p nil = nil. -Proof. by apply/fmapE. Qed. - -Lemma fnd_kfilt (p : pred K) k (s : fmap) : - fnd k (kfilter p s) = - if p k then fnd k s else None. -Proof. -case: s; rewrite /fnd /kfilter /kfilter' /=. -elim=>[|[k1 v] s IH] /=; first by case: ifP. -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. -Qed. - -Lemma kfilt_ins (p : pred K) k v (s : fmap) : - kfilter p (ins k v s) = - if p k then ins k v (kfilter p s) else kfilter p s. -Proof. -apply/fmapP=>k2; case: ifP=>E1. -- by rewrite fnd_ins !fnd_kfilt fnd_ins; case: eqP=>// ->; rewrite E1. -by rewrite !fnd_kfilt fnd_ins; case: eqP=>// ->; rewrite E1. -Qed. - -Lemma rem_kfilt (p : pred K) k (s : fmap) : - rem k (kfilter p s) = - if p k then kfilter p (rem k s) else kfilter p s. -Proof. -apply/fmapP=>k2; case: ifP=>E1. -- by rewrite fnd_rem !fnd_kfilt fnd_rem; case: eqP=>// ->; rewrite E1. -by rewrite fnd_rem fnd_kfilt; case: eqP=>// ->; rewrite E1. -Qed. - -Lemma kfilt_rem (p : pred K) k (s : fmap) : - kfilter p (rem k s) = - if p k then rem k (kfilter p s) else kfilter p s. -Proof. -apply/fmapP=>k2; case: ifP=>E1. -- by rewrite fnd_kfilt !fnd_rem fnd_kfilt; case: eqP=>// ->; rewrite E1. -by rewrite !fnd_kfilt fnd_rem; case: eqP=>// ->; rewrite E1. -Qed. - -(* filter and fcat *) - -Lemma kfilt_fcat (p : pred K) (s1 s2 : fmap) : - kfilter p (fcat s1 s2) = fcat (kfilter p s1) (kfilter p s2). -Proof. -apply/fmapP=>k; rewrite fnd_kfilt !fnd_fcat !fnd_kfilt supp_kfilt mem_filter. -by case: ifP. -Qed. - -Lemma kfilter_pred0 (s : fmap) : kfilter pred0 s = nil. -Proof. by apply/fmapE; rewrite /= /kfilter' filter_pred0. Qed. - -Lemma kfilter_predT (s : fmap) : kfilter predT s = s. -Proof. by apply/fmapE; rewrite /= /kfilter' filter_predT. Qed. - -Lemma kfilter_predI p1 p2 (s : fmap) : - kfilter (predI p1 p2) s = kfilter p1 (kfilter p2 s). -Proof. by apply/fmapE; rewrite /= /kfilter' filter_predI. Qed. - -Lemma kfilter_predU p1 p2 (s : fmap) : - kfilter (predU p1 p2) s = fcat (kfilter p1 s) (kfilter p2 s). -Proof. -apply/fmapP=>k; rewrite fnd_kfilt fnd_fcat !fnd_kfilt supp_kfilt mem_filter. -rewrite inE /=; case: (ifP (p1 k)); case: (ifP (p2 k))=>//=; -by [case: ifP | case: suppP]. -Qed. - -Lemma eq_in_kfilter p1 p2 s : - {in supp s, p1 =1 p2} -> kfilter p1 s = kfilter p2 s. -Proof. -move=>H; apply/fmapE; rewrite /= /kfilter'. -case: s H; rewrite /supp /=; elim=>[|[k v] s IH] //=. -move/path_sorted/IH=>{IH} H H1. -have ->: p1 k = p2 k by apply: H1; rewrite inE /= eq_refl. -by rewrite H // => x E; apply: H1; rewrite inE /= E orbT. -Qed. - -End Filtering. - -Section DisjointUnion. -Variable (K : ordType) (V : Type). -Notation fmap := (finMap K V). -Notation nil := (@nil K V). - -Definition disj (s1 s2 : fmap) := - all (predC (fun x => x \in supp s2)) (supp s1). - -Variant disj_spec (s1 s2 : fmap) : bool -> Type := -| disj_true of (forall x, x \in supp s1 -> x \notin supp s2) : - disj_spec s1 s2 true -| disj_false x of x \in supp s1 & x \in supp s2 : - disj_spec s1 s2 false. - -Lemma disjP s1 s2 : disj_spec s1 s2 (disj s1 s2). -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. -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. -Proof. -case: disjP; case: disjP=>//. -- by move=>x H1 H2; move/(_ x H2); rewrite H1. -by move=>H1 x H2; move/H1; rewrite H2. -Qed. - -Lemma disj_nil (s : fmap) : disj s nil. -Proof. by case: disjP. Qed. - -Lemma disj_ins k v (s1 s2 : fmap) : - disj s1 (ins k v s2) = (k \notin supp s1) && (disj s1 s2). -Proof. -case: disjP=>[H|x H1]. -- case E: (k \in supp s1)=>/=. - - by move: (H _ E); rewrite supp_ins inE /= eq_refl. - case: disjP=>// x H1 H2. - by move: (H _ H1); rewrite supp_ins inE /= H2 orbT. -rewrite supp_ins inE /=; case/orP=>[|H2]. -- by move/eqP=><-; rewrite H1. -rewrite andbC; case: disjP=>[H|y H3 H4] //=. -by move: (H _ H1); rewrite H2. -Qed. - -Lemma disj_rem k (s1 s2 : fmap) : - disj s1 s2 -> disj s1 (rem k s2). -Proof. -case: disjP=>// H _; case: disjP=>// x; move/H. -by rewrite supp_rem inE /= andbC; move/negbTE=>->. -Qed. - -Lemma disj_remE k (s1 s2 : fmap) : - k \notin supp s1 -> disj s1 (rem k s2) = disj s1 s2. -Proof. -move=>H; case: disjP; case: disjP=>//; last first. -- move=>H1 x; move/H1; rewrite supp_rem inE /= => E. - by rewrite (negbTE E) andbF. -move=>x H1 H2 H3; move: (H3 x H1) H. -rewrite supp_rem inE /= negb_and H2 orbF negbK. -by move/eqP=><-; rewrite H1. -Qed. - -Lemma disj_fcat (s s1 s2 : fmap) : - disj s (fcat s1 s2) = disj s s1 && disj s s2. -Proof. -elim/fmap_ind': s s1 s2=>[|k v s L IH] s1 s2. -- by rewrite !(disjC nil) !disj_nil. -rewrite !(disjC (ins _ _ _)) !disj_ins supp_fcat inE /= negb_or. -case: (k \in supp s1)=>//=. -case: (k \in supp s2)=>//=; first by rewrite andbF. -by rewrite -!(disjC s) IH. -Qed. - -Lemma fcatC (s1 s2 : fmap) : - disj s1 s2 -> - fcat s1 s2 = fcat s2 s1. -Proof. -rewrite /fcat. -elim/fmap_ind': s2 s1=>[|k v s2 L IH] s1 /=; first by rewrite fcat_nil'. -rewrite disj_ins; case/andP=>D1 D2. -by rewrite fcat_ins' // -IH // seqof_ins //= -fcat_ins' ?notin_path. -Qed. - -Lemma fcatA (s1 s2 s3 : fmap) : - disj s2 s3 -> - fcat (fcat s1 s2) s3 = fcat s1 (fcat s2 s3). -Proof. -move=>H. -elim/fmap_ind': s3 s1 s2 H=>[|k v s3 L IH] s1 s2 /=; first by rewrite !fcats0. -rewrite disj_ins; case/andP=>H1 H2. -by rewrite fcat_sins ?notin_path // IH // fcat_sins ?notin_path // fcat_sins. -Qed. - -Lemma fcatAC (s1 s2 s3 : fmap) : - [&& disj s1 s2, disj s2 s3 & disj s1 s3] -> - fcat s1 (fcat s2 s3) = fcat s2 (fcat s1 s3). -Proof. by case/and3P=>H1 H2 H3; rewrite -!fcatA // (@fcatC s1 s2). Qed. - -Lemma fcatCA (s1 s2 s3 : fmap) : - [&& disj s1 s2, disj s2 s3 & disj s1 s3] -> - fcat (fcat s1 s2) s3 = fcat (fcat s1 s3) s2. -Proof. -by case/and3P=>H1 H2 H3; rewrite !fcatA // ?(@fcatC s2 s3) ?(disjC s3). -Qed. - -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. -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. -Qed. - -Lemma fcatKs (s s1 s2 : fmap) : - disj s s1 && disj s s2 -> - fcat s s1 = fcat s s2 -> s1 = s2. -Proof. -case/andP=>H1 H2. -rewrite (fcatC H1) (fcatC H2); apply: fcatsK. -by rewrite -!(disjC s) H1 H2. -Qed. - -(* DEVCOMMENT *) -(* heh, a theory of submaps would be good here *) -(* but i don't have time for a decent development *) -(* so let's do a quick lemma that's needed for feaps *) -(* /DEVCOMMENT *) - -Lemma disj_kfilt p1 p2 s1 s2 : - disj s1 s2 -> disj (kfilter p1 s1) (kfilter p2 s2). -Proof. -elim/fmap_ind': s2 s1=>[|k v s _ IH] s1 /=. -- by rewrite kfilt_nil => _; case: disjP. -rewrite disj_ins; case/andP=>H1 H2; rewrite kfilt_ins. -case: ifP=>E; last by apply: IH. -rewrite disj_ins supp_kfilt mem_filter negb_and H1 orbT /=. -by apply: IH. -Qed. - -Lemma in_disj_kfilt p1 p2 s : - {in supp s, forall x, ~~ p1 x || ~~ p2 x} -> - disj (kfilter p1 s) (kfilter p2 s). -Proof. -elim/fmap_ind': s=>[|k v s _ IH] //= H. -rewrite !kfilt_ins; case: ifP=>E1; case: ifP=>E2. -- move: (H k); rewrite E1 E2 supp_ins inE /= eq_refl /=. - by move/(_ (erefl _)). -- rewrite disjC disj_ins disjC supp_kfilt mem_filter negb_and E2 /=. - by apply: IH=>x S; apply: H; rewrite supp_ins inE /= S orbT. -- rewrite disj_ins supp_kfilt mem_filter negb_and E1 /=. - by apply: IH=>x S; apply: H; rewrite supp_ins inE /= S orbT. -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. -Variables (K : ordType) (V : eqType). - -Definition feq (s1 s2 : finMap K V) := seq_of s1 == seq_of s2. - -Lemma feqP : Equality.axiom feq. -Proof. -move=>s1 s2; rewrite /feq. -case: eqP; first by move/fmapE=>->; apply: ReflectT. -by move=>H; apply: ReflectF; move/fmapE; move/H. -Qed. - -HB.instance Definition _ := hasDecEq.Build (finMap K V) feqP. -End EqType. - -(* mapping a function over a contents of a finite map *) - -Section Map. -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 (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. -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 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: ordP=>O //=. -by rewrite IH // (path_sorted H). -Qed. - -Lemma mapf_fcat s1 s2 : mapf (fcat s1 s2) = fcat (mapf s1) (mapf s2). -Proof. -elim/fmap_ind': s2 s1=>[|k v s2 H IH] s1 /=. -- rewrite fcats0; set j := FinMap _. - by rewrite (_ : j = @nil K V) ?fcat0s //; apply/fmapE. -by rewrite fcat_sins mapf_ins IH -fcat_sins mapf_ins. -Qed. - -Lemma mapf_disjL s1 s2 s : mapf s1 = mapf s2 -> disj s1 s = disj s2 s. -Proof. -case: s1 s2 s=>s1 S1 [s2 S2][s S] /fmapE /=. -elim: s1 S1 s2 S2 s S=>[|[k v] s1 IH] /= S1; case=>//= [[k2 v2]] s2 S2 s S. -case=>E _; rewrite -{k2}E in S2 *. -move/(IH (path_sorted S1) _ (path_sorted S2) _ S). -by rewrite /disj /supp /= => ->. -Qed. - -Lemma mapf_disj s1 s2 s1' s2' : - mapf s1 = mapf s2 -> mapf s1' = mapf s2' -> - disj s1 s1' = disj s2 s2'. -Proof. -move=>X1 X2; apply: eq_trans (mapf_disjL _ X1) _. -by rewrite !(disjC s2); apply: mapf_disjL X2. -Qed. - -End Map. - -Section FoldFMap. -Variables (A B: ordType) (V C: Type). - -Definition foldfmap g (e: C) (s: finMap A V) := - foldr g e (seq_of s). - -Lemma foldf_nil g e : foldfmap g e (@nil A V) = e. -Proof. by rewrite /foldfmap //=. Qed. - -Lemma foldf_ins g e k v f: - path ord k (supp f) -> - foldfmap g e (ins k v f) = g (k, v) (foldfmap g e f). -Proof. by move=> H; rewrite /foldfmap //= seqof_ins //. Qed. -End FoldFMap. - -Section KeyMap. - -Section MapDef. -Variables (A B : ordType) (V : Type). - -Variable (f: A -> B). -Hypothesis Hf : forall x y, strictly_increasing f x y. - -Definition mapk (m : finMap A V) : finMap B V := - foldfmap (fun p s => ins (f (key p)) (value p) s) nil m. - -(* mapK preserves sorted *) - -Lemma sorted_mapk m : sorted ord (supp (mapk m)). -Proof. case: (mapk m)=>[s]I //=. Qed. - -Lemma path_mapk m k: - path ord k (supp m) -> - path ord (f k) (supp (mapk m)). -Proof. -elim/fmap_ind': m k =>// k1 v1 s P IH k. -rewrite {1}/supp //= {1}seqof_ins // /= => /andP [H]; move/IH=>H1. -by rewrite /mapk foldf_ins // /supp /= seqof_ins //= H1 andbT (Hf H). -Qed. - -Lemma mapk_nil : mapk (@nil A V) = @nil B V. -Proof. by rewrite /mapk //=. Qed. - -Lemma mapk_ins k v s : - path ord k (supp s) -> - mapk (ins k v s) = ins (f k) v (mapk s). -Proof. by move=> H; rewrite /mapk foldf_ins =>//. Qed. -End MapDef. - -Arguments mapk {A B V} f m. - -Variables (A B C : ordType) (V : Type) (f : A -> B) (g : B -> C). -Hypothesis Hf : forall x y, strictly_increasing f x y. - -Lemma mapk_id m : @mapk A A V id m = m. -Proof. -by elim/fmap_ind': m=>// k v s L IH; rewrite -{2}IH /mapk foldf_ins //. -Qed. - -Lemma mapk_comp m: - mapk g (@mapk A B V f m) = mapk (comp g f) m. -Proof. -elim/fmap_ind': m =>//= k v s P IH. -rewrite [mapk (g \o f) _]mapk_ins //. -rewrite mapk_ins // mapk_ins //; first by rewrite IH. -exact: (path_mapk Hf P). -Qed. -End KeyMap. -Arguments mapk {A B V} f m. - -(* zipping two finite maps with equal domains and ranges *) -(* zip_p predicate is a test for when contents of the two maps *) -(* at a given key are "combinable". typically zip_p will test *) -(* that the types of contents are equal, in the case the map is *) -(* storing dynamics *) - -Section Zip. -Variables (K : ordType) (V : Type) (zip_f : V -> V -> option 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 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 - [::], [::] => Some [::] - | (k1, v1)::s1', (k2, v2)::s2' => - if k1 == k2 then - if zip_f v1 v2 is Some v then - if zip' s1' s2' is Some s' then Some ((k1, v) :: s') - else None - else None - else None - | _, _ => None - end. - -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] //=; rewrite eq_sym; case: eqP=>// ->{k2}. -by rewrite comm IH. -Qed. - -Lemma zipA' s1 s2 s3 : - obind (zip' s1) (zip' s2 s3) = obind (zip'^~ s3) (zip' s1 s2). -Proof. -elim: s1 s2 s3=>[|[k1 v1] s1 IH]. -- case=>[|[k2 v2] s2]; case=>[|[k3 v3] s3] //=; case: eqP=>// ->{k2}. - by case: (zip_f v2 v3)=>// v; case: (zip' s2 s3). -case=>[|[k2 v2] s2]; case=>[|[k3 v3] s3] //=. -- by case: eqP=>// ->{k1}; case: (zip_f v1 v2)=>// v; case: (zip' s1 s2). -case: (k2 =P k3)=>[->{k2}|E1] /=; last first. -- case: (k1 =P k2)=>E2 //=. - case: (zip_f v1 v2)=>// v. - case: (zip' s1 s2)=>//= s. - by rewrite E2; case: eqP E1. -case: (k1 =P k3)=>[->{k1}|E1] /=; last first. -- by case: (zip_f v2 v3)=>// v; case: (zip' s2 s3)=>//= s; case: eqP E1. -case E1: (zip_f v2 v3)=>[w1|]; last first. -- case E3: (zip_f v1 v2)=>[w3|] //. - case S3: (zip' s1 s2)=>[t3|] //=; rewrite eq_refl. - by move: (assoc v1 v2 v3); rewrite /obind/oapp E1 E3=><-. -case S1: (zip' s2 s3)=>[t1|] /=; last first. -- case E3: (zip_f v1 v2)=>[w3|//]. - case S3: (zip' s1 s2)=>[t3|] //=; rewrite eq_refl. - move: (IH s2 s3); rewrite /obind/oapp S1 S3=><-. - by case: (zip_f w3 v3). -rewrite eq_refl. -case E3: (zip_f v1 v2)=>[w3|]; -move: (assoc v1 v2 v3); rewrite /obind/oapp E1 E3=>-> //=. -case S3: (zip' s1 s2)=>[t3|]; -move: (IH s2 s3); rewrite /obind/oapp S3 S1=>-> /=. -- by rewrite eq_refl. -by case: (zip_f w3 v3). -Qed. - -Lemma zip_unitL' s : zip' (zip_unit' s) s = Some s. -Proof. by elim: s=>[|[k v] s IH] //=; rewrite eq_refl unitL IH. Qed. - -Lemma map_key_zip' s1 s2 s : - zip' s1 s2 = Some s -> map key s = map key s1. -Proof. -elim: s1 s2 s=>[|[k1 v1] s1 IH]; case=>[|[k2 v2] s2] //= s; first by case=><-. -case: eqP=>// ->{k1}; case: (zip_f v1 v2)=>// w; case Z: (zip' s1 s2)=>[t|//]. -by case=><-; rewrite -(IH _ _ Z). -Qed. - -Lemma zip_unitE' s1 s2 : - (exists s, zip' s1 s2 = Some s) <-> zip_unit' s1 = zip_unit' s2. -Proof. -split. -- case; elim: s1 s2 =>[|[k1 v1] s1 IH]; case=>// [[k2 v2] s2] s //=. - case: eqP=>// <-{k2}. - case E1: (zip_f v1 v2)=>[w|//]; case E2: (zip' s1 s2)=>[t|//] _. - by move/IH: E2=>->; congr ((_, _)::_); apply/unitE; exists w. -elim: s1 s2=>[|[k1 v1] s1 IH]; case=>//; first by exists [::]. -case=>k2 v2 s2 //= [<-{k2}]; case/unitE=>s ->; case/IH=>t ->. -by exists ((k1, s)::t); rewrite eq_refl. -Qed. - -Lemma zip_sorted' s1 s2 : - sorted ord (map key s1) -> - forall s, zip' s1 s2 = Some s -> sorted ord (map key s). -Proof. by move=>H s; move/map_key_zip'=>->. Qed. - -Definition zip f1 f2 : option (finMap K V) := - match f1, f2 with - FinMap s1 pf1, FinMap s2 pf2 => - match zip' s1 s2 as z return zip' s1 s2 = z -> _ with - Some s => fun pf => Some (FinMap (zip_sorted' pf1 pf)) - | None => fun pf => None - end (erefl _) - end. - -Lemma zip_unit_sorted' s : - sorted ord (map key s) -> - sorted ord (map key (zip_unit' s)). -Proof. -rewrite (_ : map key s = map key (zip_unit' s)) //. -by apply: (map_key_zip' (s2:=s)); apply: zip_unitL'. -Qed. - -Definition zip_unit f := - let: FinMap s pf := f in FinMap (zip_unit_sorted' pf). - -Lemma zip_unitE f1 f2 : - (exists f, zip f1 f2 = Some f) <-> zip_unit f1 = zip_unit f2. -Proof. -case: f1 f2=>s1 H1 [s2 H2] /=; split. -- case=>s; move: (zip_sorted' _). - case E: (zip' s1 s2)=>[t|//] _ _; apply/fmapE=>/=. - by apply/zip_unitE'; exists t. -case; case/zip_unitE'=>s E; move/(zip_sorted' H1): (E)=>T. -exists (FinMap T); move: (zip_sorted' _); rewrite E=>pf. -by congr Some; apply/fmapE. -Qed. - -(* Now lemmas for interaction of zip with the other primitives *) - -Lemma zip_supp' s1 s2 s : zip' s1 s2 = Some s -> map key s = map key s1. -Proof. -elim: s1 s2 s=>[|[k1 v1] s1 IH] /=; first by case=>// s [<-]. -case=>[|[k2 v2] s2] // s; case: eqP=>// ->{k1}. -case E: (zip_f v1 v2)=>[w|//]; case F: (zip' s1 s2)=>[t|//]. -by move/IH: F=><- [<-]. -Qed. - -Lemma zip_supp f1 f2 f : - zip f1 f2 = Some f -> supp f =i supp f1. -Proof. -case: f1 f2 f=>s1 H1 [s2 H2] [s3 H3] /=; move: (zip_sorted' _). -case E: (zip' s1 s2)=>[t|//]; move=>pf [F x]; rewrite -{s3}F in H3 *. -by rewrite /supp (zip_supp' E). -Qed. - -Lemma zip_filter' s1 s2 s x : - zip' s1 s2 = Some s -> - zip' (filter (predCk V x) s1) (filter (predCk V x) s2) = - Some (filter (predCk V x) s). -Proof. -elim: s1 s2 s=>[|[k1 v1] s1 IH]; case=>[|[k2 v2] s2] //= s; first by case=><-. -case: eqP=>// <-{k2}; case E1: (zip_f v1 v2)=>[a|//]. -case E2: (zip' s1 s2)=>[t|//]; move/IH: E2=>E2 [<-{s}]. -case: eqP=>[->{k1}|] /=; first by rewrite E2 eq_refl. -by rewrite eq_refl E1 E2; case: eqP. -Qed. - -Lemma zip_rem f1 f2 f x : - zip f1 f2 = Some f -> - zip (rem x f1) (rem x f2) = Some (rem x f). -Proof. -case: f1 f2 f=>s1 H1 [s2 H2] [s3 H3] /=; do 2![move: (zip_sorted' _)]. -case E1: (zip' s1 s2)=>[t|//]; case E2 : (zip' _ _)=>[q|]; -rewrite (@zip_filter' _ _ _ _ E1) // in E2. -by case: E2=><-{q} pf1 pf2 [E]; congr Some; apply/fmapE; rewrite /= E. -Qed. - -Lemma zip_fnd f1 f2 f x (v : V) : - zip f1 f2 = Some f -> - fnd x f = Some v -> - exists v1 v2, - [/\ zip_f v1 v2 = Some v, fnd x f1 = Some v1 & fnd x f2 = Some v2]. -Proof. -case: f1 f2 f=>s1 H1 [s2 H2] [s3 H3] /=; move: (zip_sorted' _). -case E1: (zip' s1 s2)=>[s|//] pf [E]; rewrite /fnd /=. -clear H1 H2 H3 pf; rewrite -{s3}E. -elim: s1 s2 s E1=>[|[k1 v1] s1 IH]; case=>[|[k2 v2] s2] //= s. -- by case=><-. -case: eqP=>// <-{k2}; case E1: (zip_f v1 v2)=>[w|//]. -case E2: (zip' s1 s2)=>[t|//][<-{s}] /=. -case: eqP=>[_ [<-]|_]. -(* DEVCOMMENT: exists v1, v2 errors! *) -(* /DEVCOMMENT *) -- by exists v1; exists v2. -by case: (filter (predk V x) t) (IH _ _ E2). -Qed. - -(* The other direction of the zip_fnd lemma *) - -Lemma fnd_zip f1 f2 f x (v1 v2 : V) : - fnd x f1 = Some v1 -> fnd x f2 = Some v2 -> - zip f1 f2 = Some f -> fnd x f = zip_f v1 v2. -Proof. -case: f1 f2=>s1 H1 [s2 H2] /=; move: (zip_sorted' _). -case E: (zip' s1 s2)=>[s|//] pf; rewrite /fnd /= => F1 F2. -case=><-{f} /= {pf H1 H2}. -elim: s1 s2 s E F1 F2=>[|[k1 w1] s1 IH]; case=>[|[k2 w2] s2] //= s. -case: eqP=>// <-{k2}; case E1: (zip_f w1 w2)=>[w|//]. -case E2: (zip' s1 s2)=>[t|//] [<-{s}]. -case: eqP=>/=; last by case: eqP=>// _ _; apply: IH. -by move=>->{k1}; rewrite eq_refl; case=><- [<-]. -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 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: ordP=>//= O. -by rewrite IH // (path_sorted H). -Qed. - -Lemma zunit_fcat f1 f2 : - zip_unit (fcat f1 f2) = fcat (zip_unit f1) (zip_unit f2). -Proof. -elim/fmap_ind': f2 f1=>[|k v f2 H IH] f1 /=. -- rewrite fcats0; set j := FinMap _. - by rewrite (_ : j = @nil K V) ?fcat0s //; apply/fmapE. -by rewrite fcat_sins zunit_ins IH -fcat_sins zunit_ins. -Qed. - -Lemma zunit_supp f : supp (zip_unit f) = supp f. -Proof. -case: f=>s H; rewrite /supp /= {H}. -by elim: s=>[|[k v] s IH] //=; rewrite IH. -Qed. - -Lemma zunit_disj f1 f2 : - disj f1 f2 = disj (zip_unit f1) (zip_unit f2). -Proof. -case: disjP; case: disjP=>//; rewrite !zunit_supp. -- by move=>x H1 H2; move/(_ _ H1); rewrite H2. -by move=>H x; move/H/negbTE=>->. -Qed. - -End Zip. diff --git a/core/options.v b/core/options.v deleted file mode 100644 index 6269953..0000000 --- a/core/options.v +++ /dev/null @@ -1,3 +0,0 @@ -#[export] Set Implicit Arguments. -#[export] Unset Strict Implicit. -#[export] Unset Printing Implicit Defensive. diff --git a/core/ordtype.v b/core/ordtype.v deleted file mode 100644 index 1a83f23..0000000 --- a/core/ordtype.v +++ /dev/null @@ -1,497 +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 HB Require Import structures. -From Coq Require Import ssreflect ssrbool ssrfun. -From mathcomp Require Import ssrnat eqtype seq path fintype. -From pcm Require Import options seqext. - -Definition connected (T : eqType) (ord : rel T) := - forall x y, x != y -> ord x y || ord y x. - -Definition ordtype_axiom (T : eqType) (ord : rel T) := - [/\ irreflexive ord, transitive ord & connected ord]. - -HB.mixin Record isOrdered T of Equality T := { - ord : rel T; - ordtype_subproof : ordtype_axiom ord}. - -(* Ideally, we would just generate the structure thus. *) -(* But, we need to insert some Universe Polymorphism declaration *) -(* so we switch to plan B: print the generated code using #[log] *) -(* command, then cut-and-paste it and decorate by hand. *) -(* Unfortunately, this doesn't quite work, HB manuals notwithstanding, *) -(* because custom declarations of log code impoverish the Elpi state *) -(* making subsequent invocations of HB work only partially. *) -(* In this file in particular, the canonicals for each instance *) -(* must be declared by hand as HB.instance does that *) -(* only if Ordered is declared by HB.structure *) -(* -#[log(raw)] -#[short(type="ordType")] -HB.structure Definition Ordered := - {T of Equality T & isOrdered T}. -*) - -(* cut-and-pasted code starts *) -Module Ordered. -Section axioms_. -Local Unset Implicit Arguments. - -Record axioms_ (T : Type) : Type := Class { - eqtype_hasDecEq_mixin : Equality.mixin_of T; - ordtype_isOrdered_mixin : isOrdered.axioms_ T eqtype_hasDecEq_mixin}. -End axioms_. - -Global Arguments axioms_ : clear implicits. -Global Arguments Class [_] [_] _. -Global Arguments eqtype_hasDecEq_mixin [_] _. -Global Arguments ordtype_isOrdered_mixin [_] _. - -Section type. -Local Unset Implicit Arguments. -Polymorphic Cumulative Record - type : Type := Pack { sort : Type; _ : axioms_ sort; }. -End type. - -Definition class (cT : type) := - let: Pack _ c as cT' := cT return axioms_ (sort cT') in c. - -Global Arguments type : clear implicits. -Global Arguments Pack [_] _. -Global Arguments sort : clear implicits. -Global Arguments class : clear implicits. - -(* DEVCOMMENT *) -(* The polymorphism annotations here and below are needed for storing *) -(* ordType instances in finMaps which have an ordType constraint of *) -(* their own. An example of this is KVMap from HTT. *) -(* Ultimately there should be a better solution if we want to switch *) -(* to Mathcomp's orders. *) -(* /DEVCOMMENT *) - -(* Polymorphic Universe annotations added *) -Section PolymorphicClonePack. -Polymorphic Universe ou. -Polymorphic Variables (T : Type@{ou}) (cT : type@{ou}). - -Polymorphic Definition phant_clone : forall (c : axioms_ T), - unify Type Type T (sort cT) nomsg -> - unify type type cT (Pack (sort:=T) c) nomsg -> type := - fun (c : axioms_ T) => - fun=> (fun=> (Pack (sort:=T) c)). - -Polymorphic Definition pack_ := fun (m : Equality.mixin_of T) - (m0 : isOrdered.axioms_ T m) => - Pack (sort:=T) - {|eqtype_hasDecEq_mixin := m; ordtype_isOrdered_mixin := m0 |}. - -End PolymorphicClonePack. - -Local Arguments phant_clone : clear implicits. -Notation clone X2 X1 := (phant_clone X2 X1 _ id_phant id_phant). -Local Arguments pack_ : clear implicits. - -Module Exports. -Notation ordType := Ordered.type. -#[reversible] Coercion sort : Ordered.type >-> Sortclass. - -(* Polymorphic annotation added *) -Polymorphic Definition ordtype_Ordered_class__to__eqtype_Equality_class : - forall T : Type, axioms_ T -> Equality.axioms_ T := - fun (T : Type) (c : axioms_ T) => - {| Equality.eqtype_hasDecEq_mixin := eqtype_hasDecEq_mixin c |}. - -Local Arguments ordtype_Ordered_class__to__eqtype_Equality_class : - clear implicits. - -#[reversible] Coercion ordtype_Ordered_class__to__eqtype_Equality_class : - ordtype.Ordered.axioms_ >-> eqtype.Equality.axioms_. - -Polymorphic Definition ordtype_Ordered__to__eqtype_Equality : - ordType -> eqType := - fun s : ordType => {| Equality.sort := s; Equality.class := class s |}. - -Local Arguments ordtype_Ordered__to__eqtype_Equality : - clear implicits. - -#[reversible] Coercion ordtype_Ordered__to__eqtype_Equality : - ordtype.Ordered.type >-> eqtype.Equality.type. - -Global Canonical ordtype_Ordered__to__eqtype_Equality. - -#[reversible] Coercion eqtype_hasDecEq_mixin : - ordtype.Ordered.axioms_ >-> eqtype.hasDecEq.axioms_. - -#[reversible] Coercion ordtype_isOrdered_mixin : - ordtype.Ordered.axioms_ >-> ordtype.isOrdered.axioms_. - -End Exports. -Import Exports. - -Definition phant_on_ : forall T : ordType, phant T -> axioms_ T := - fun T : ordType => fun=> class T. -Local Arguments phant_on_ : clear implicits. - -Notation on_ X1 := ( phant_on_ _ (Phant X1)). -Notation copy X2 X1 := ( phant_on_ _ (Phant X1) : axioms_ X2). -Notation on X1 := ( phant_on_ _ (Phant _) : axioms_ X1). - -Module EtaAndMixinExports. -Section hb_instance_91. -Variable T : ordType. -Local Arguments T : clear implicits. - -Definition HB_unnamed_factory_92 : axioms_ (eta T) := class T. -Local Arguments HB_unnamed_factory_92 : clear implicits. -Definition ordtype_Ordered__to__eqtype_hasDecEq : - Equality.mixin_of (eta T) := Equality.class T. -Local Arguments ordtype_Ordered__to__eqtype_hasDecEq : - clear implicits. - -Definition ordtype_Ordered__to__ordtype_isOrdered : - isOrdered.axioms_ (eta T) HB_unnamed_factory_92 := - ordtype_isOrdered_mixin HB_unnamed_factory_92. -Local Arguments ordtype_Ordered__to__ordtype_isOrdered : - clear implicits. - -Definition HB_unnamed_mixin_95 := - ordtype_isOrdered_mixin HB_unnamed_factory_92. -Local Arguments HB_unnamed_mixin_95 : clear implicits. - -Definition structures_eta__canonical__ordtype_Ordered : ordType := - Pack (sort := eta T) - {| - eqtype_hasDecEq_mixin := Equality.class T; - ordtype_isOrdered_mixin := HB_unnamed_mixin_95 - |}. - -Local Arguments structures_eta__canonical__ordtype_Ordered : - clear implicits. -Global Canonical structures_eta__canonical__ordtype_Ordered. -End hb_instance_91. -End EtaAndMixinExports. -End Ordered. - -HB.export Ordered.Exports. -HB.export Ordered.EtaAndMixinExports. - -Definition ord : forall s : ordType, rel s := - fun s : ordType => isOrdered.ord (Ordered.class s). -Local Arguments ord : clear implicits. -Global Arguments ord {_}. - -Definition ordtype_subproof : forall s : ordType, ordtype_axiom ord := - fun s : ordType => isOrdered.ordtype_subproof (Ordered.class s). -Local Arguments ordtype_subproof : clear implicits. -Global Arguments ordtype_subproof {_}. - -Notation Ordered X1 := (Ordered.axioms_ X1). -(* end of generated and changed code *) - - -(* Repacking *) - -Lemma irr {T : ordType} : irreflexive (@ord T). -Proof. by case: (@ordtype_subproof T). Qed. - -Lemma trans {T : ordType} : transitive (@ord T). -Proof. by case: (@ordtype_subproof T). Qed. - -Lemma connex {T : ordType} : connected (@ord T). -Proof. by case: (@ordtype_subproof T). Qed. - -Definition oleq (T : ordType) (t1 t2 : T) := ord t1 t2 || (t1 == t2). - -Prenex Implicits ord oleq. - -Section Lemmas. -Context {T : ordType}. -Implicit Types x y : T. - -Lemma ord_total x y : [|| ord x y, x == y | ord y x]. -Proof. -case E: (x == y)=>/=; first by rewrite orbT. -by apply: connex; rewrite negbT. -Qed. - -Lemma ord_neq (x y : T) : - ord x y -> - x != y. -Proof. -move=>H; apply/negP=>/eqP E. -by rewrite E irr in H. -Qed. - -Lemma nsym x y : - ord x y -> - ~ ord y x. -Proof. by move=>E1 E2; move: (trans E1 E2); rewrite irr. Qed. - -Lemma antisym : antisymmetric (@ord T). -Proof. by move=>x y /andP [H] /(nsym H). 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. 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_path tr (filter_subseq p s)). Qed. - -Lemma ord_sorted_eq (s1 s2 : seq T) : - sorted ord s1 -> - sorted ord s2 -> - s1 =i s2 -> - s1 = s2. -Proof. exact/irr_sorted_eq/irr/(@trans _). Qed. - -Lemma oleq_ord_trans (n m p : T) : - oleq m n -> - ord n p -> - ord m p. -Proof. by case/orP; [apply: trans | move/eqP=>->]. Qed. - -Lemma ord_oleq_trans (n m p : T) : - ord m n -> - oleq n p -> - ord m p. -Proof. by move=>H /orP [|/eqP <- //]; apply: trans. Qed. - -End Lemmas. - -#[export] Hint Resolve orefl irr trans connex -otrans oantisym oleq_ord_trans : 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 ordType *) -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. - -Lemma oleq_eqVord x y : oleq x y = (x == y) || ord x y. -Proof. by rewrite /oleq orbC. 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. - -(**********************************************) -(* Building hierarchies for some basic orders *) -(**********************************************) - -(* trivial total order for unit type *) -Definition unit_ord (x y : unit) := false. -Lemma unit_is_ordtype : ordtype_axiom unit_ord. Proof. by []. Qed. -HB.instance Definition unit_ord_mix := isOrdered.Build unit unit_is_ordtype. -(* manual canonical declaration, as HB fails to declare it *) -Canonical unit_ordType : ordType := - Ordered.Pack (sort:=unit) (Ordered.Class unit_ord_mix). - -(* trivial total order for nat *) -Lemma nat_is_ordtype : ordtype_axiom ltn. -Proof. by split=>[x||x y] /=; [rewrite ltnn|apply: ltn_trans|case: ltngtP]. Qed. -HB.instance Definition nat_ord_mix := isOrdered.Build nat nat_is_ordtype. -Canonical nat_ordType : ordType := - Ordered.Pack (sort:=nat) (Ordered.Class nat_ord_mix). - -Lemma nat_ord : ord =2 ltn. Proof. by []. Qed. -Lemma nat_oleq : oleq =2 leq. -Proof. by move=>x y; rewrite /oleq nat_ord /=; case: ltngtP. Qed. - -(* product order *) -Section ProdOrdType. -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 prod_is_ordtype : ordtype_axiom lex. -Proof. -split=>[x|[x1 x2][y1 y2][z1 z2]|[x1 x2][y1 y2]]; rewrite /lex /=. -- by rewrite eq_refl irr. -- 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). -by rewrite -pair_eqE /= -(eq_sym x1 y1); case: ifPn=>[_|] /connex. -Qed. - -HB.instance Definition prod_ord_mix := - isOrdered.Build (K * T)%type prod_is_ordtype. -(* manual canonical declaration, as HB fails to declare it *) -Canonical prod_ordType : ordType := - Ordered.Pack (sort:=prod K T) (Ordered.Class prod_ord_mix). -End ProdOrdType. - - -(* Every fintype is ordered *) -Section FinTypeOrd. -Variable T : finType. - -Definition ordf : rel T := - fun x y => index x (enum T) < index y (enum T). - -Lemma fintype_is_ordtype : ordtype_axiom ordf. -Proof. -split=>[x|x y z|x y]; rewrite /ordf /=. -- by rewrite ltnn. -- by apply: ltn_trans. -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. - -(* There isn't canonical projection to latch onto *) -(* so we can't have generic inheritance. *) -(* -HB.instance Definition _ := - isOrdered.Build T%type fintype_is_ordtype. -*) - -End FinTypeOrd. - -(* However, we can get ordtype for any individual fintype *) -(* e.g. ordinals 'I_n *) -HB.instance Definition ordinal_ord_mix n := - isOrdered.Build 'I_n (fintype_is_ordtype 'I_n). -(* manual canonical declaration, as HB fails to declare it *) -Canonical ordinal_ordType n : ordType := - Ordered.Pack (sort:='I_n) (Ordered.Class (ordinal_ord_mix n)). - -Section SeqOrdType. -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 seq_is_ordtype : ordtype_axiom ords. -Proof. -split. -- by elim=>[|x xs IH] //=; rewrite eq_refl. -- elim=>[|y ys /= IH][|x xs][|z zs] //=. - case: (eqVneq x y)=>[->{x}|Hxy]; first by case: ifP=>// _; apply: IH. - case: (eqVneq y z)=>[<-{z}|Hyz]; first by rewrite (negbTE Hxy). - by case: (x =P z)=>[<-{Hyz z} H /(nsym H)//|_]; apply: trans. -elim=>[|x xs IH][|y ys] //=; rewrite (eq_sym y). -case: (x =P y)=>[-> H|/eqP/connex //]. -by apply: IH; apply: contra H=>/eqP ->. -Qed. - -HB.instance Definition seq_ord_mix := isOrdered.Build (seq T) seq_is_ordtype. -(* manual canonical declaration, as HB fails to declare it *) -Canonical seq_ordType : ordType := - Ordered.Pack (sort:=seq T) (Ordered.Class seq_ord_mix). -End SeqOrdType. - -Lemma sorted_cat_cons_cat (A : ordType) (l r : seq A) x : - sorted ord (l ++ x :: r) = - sorted ord (l ++ [::x]) && sorted ord (x::r). -Proof. -rewrite !(sorted_pairwise (@trans A)) cats1 pairwise_cat. -rewrite pairwise_rcons allrel_consr !pairwise_cons. -case/boolP: (all (ord^~ x) l)=>//= Hl. -case/boolP: (all (ord x) r)=>/= [Hr|_]; last by rewrite !andbF. -by rewrite (allrel_trans (@trans A) Hl Hr). -Qed. - diff --git a/core/pred.v b/core/pred.v deleted file mode 100644 index fcf6ca1..0000000 --- a/core/pred.v +++ /dev/null @@ -1,1833 +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 re-implements some of ssrbool's entities in Prop universe *) -(******************************************************************************) - -From Coq Require Import ssreflect ssrbool ssrfun Setoid Basics. -From mathcomp Require Import ssrnat seq eqtype. -From pcm Require Import options. - -(* First some basic propositional equalities *) - -(* DEVCOMMENT *) -(* Basically, we need to repeat most of ssrbool.v here but we'll do it as we go. *) -(* /DEVCOMMENT *) - -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. -Open Scope fun_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. - -Definition SubPred T (p1 p2 : Pred T) := forall x : T, p1 x -> p2 x. - -Notation xPred0 := (fun _ => False). -Notation xPred1 := (fun x => eq^~ x). -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)). - -(* The packed class interface for pred-like types. *) - -Structure PredType T := - PropPredType {Pred_Sort :> Type; toPred : Pred_Sort -> Pred T}. - -Definition Clone_Pred T U := - fun pT & @Pred_Sort T pT -> U => - fun toP (pT' := @PropPredType T U toP) & phant_id pT' pT => pT'. -Notation "[ 'PredType' 'of' T ]" := (@Clone_Pred _ T _ id _ id) : form_scope. - -Canonical PredPredType T := PropPredType (@id (Pred T)). -Canonical PropFunPredType T := PropPredType (@id (T -> Prop)). - -Notation "{ 'Pred' T }" := (Pred_Sort (PredPredType T)) - (format "{ 'Pred' T }") : type_scope. - -Definition Eq_Pred {T} (pT : PredType T) (mp1 mp2 : pT) := - forall x : T, toPred mp1 x <-> toPred mp2 x. -Definition Sub_Pred {T} (pT : PredType T) (mp1 mp2 : pT) := - forall x : T, toPred mp1 x -> toPred mp2 x. -Definition Eq_PredFun {T1 T2} (pT2 : PredType T2) p1 p2 := - forall x : T1, @Eq_Pred T2 pT2 (p1 x) (p2 x). -Definition Sub_PredFun T1 T2 (pT2 : PredType T2) p1 p2 := - forall x : T1, @Sub_Pred T2 pT2 (p1 x) (p2 x). - -Notation "A <~> B" := (@Eq_Pred _ _ A B) - (at level 70, no associativity) : rel_scope. -Notation "A ~> B" := (@Sub_Pred _ _ A B) - (at level 70, no associativity) : rel_scope. -Notation "A <~1> B" := (@Eq_PredFun _ _ _ A B) - (at level 70, no associativity) : rel_scope. -Notation "A ~1> B" := (@Sub_PredFun _ _ _ A B) - (at level 70, no associativity) : rel_scope. - -(* The type of self-simplifying collective predicates. *) - -Definition Simpl_Pred T := simpl_fun T Prop. -Definition PropSimplPred {T} (p : Pred T) : Simpl_Pred T := SimplFun p. - -(* Some Simple_Pred constructors *) - -Definition Pred0 {T} := @PropSimplPred T xPred0. -Definition PredT {T} := @PropSimplPred T xPredT. -Definition Pred1 {T} (x : T) := PropSimplPred (xPred1 x). -Definition PredI {T} (p1 p2 : Pred T) := PropSimplPred (xPredI p1 p2). -Definition PredU {T} (p1 p2 : Pred T) := PropSimplPred (xPredU p1 p2). -Definition PredC {T} (p : Pred T) := PropSimplPred (xPredC p). -Definition PredD {T} (p1 p2 : Pred T) := PropSimplPred (xPredD p1 p2). -Definition Preim {aT rT} (f : aT -> rT) (d : Pred rT) := - PropSimplPred (xPreim f d). - -Notation "[ 'Pred' : T | E ]" := (PropSimplPred (fun _ : T => E)) - (format "[ 'Pred' : T | E ]") : fun_scope. -Notation "[ 'Pred' x | E ]" := (PropSimplPred (fun x => E)) - (x name, format "[ 'Pred' x | E ]") : fun_scope. -Notation "[ 'Pred' x : T | E ]" := (PropSimplPred (fun x : T => E)) - (x name, only parsing) : fun_scope. -Notation "[ 'Pred' x | E1 & E2 ]" := [Pred x | E1 /\ E2 ] - (x name) : fun_scope. -Notation "[ 'Pred' x : T | E1 & E2 ]" := [Pred x : T | E1 /\ E2 ] - (x name, only parsing) : fun_scope. - -(* coercion for Simpl_Pred *) - -Module PredOfSimpl. -Definition Coerce T (sp : Simpl_Pred T) : Pred T := fun_of_simpl sp. -End PredOfSimpl. -Notation Pred_Of_Simpl := PredOfSimpl.Coerce. -Coercion Pred_Of_Simpl : Simpl_Pred >-> Pred. -Canonical SimplPredType T := PropPredType (@Pred_Of_Simpl T). - -Module Type PropPredSortOfSimplSignature. -Parameter Coerce : forall T, Simpl_Pred T -> {Pred T}. -End PropPredSortOfSimplSignature. -Module PropDeclarePredSortOfSimpl - (PropPredSortOfSimpl : PropPredSortOfSimplSignature). -Coercion PropPredSortOfSimpl.Coerce : Simpl_Pred >-> Pred_Sort. -End PropDeclarePredSortOfSimpl. -Module Export PropPredSortOfSimplCoercion := - PropDeclarePredSortOfSimpl PredOfSimpl. - -Definition PredArgType := Type. -Bind Scope type_scope with PredArgType. -Identity Coercion Sort_Of_PredArgType : PredArgType >-> Sortclass. -Coercion Pred_Of_ArgType (T : PredArgType) : Simpl_Pred T := PredT. -Notation "{ :: T }" := (T%type : PredArgType) : type_scope. - -(* prop relations *) - -Definition Rel T := T -> Pred T. -Identity Coercion Fun_Of_Rel : Rel >-> Funclass. - -Definition eqRel {T} (R1 R2 : Rel T) := forall x y, R1 x y <-> R2 x y. -Definition subRel {T} (R1 R2 : Rel T) := forall x y : T, R1 x y -> R2 x y. - -Notation "A <~2> B" := (eqRel A B) - (at level 70, no associativity) : rel_scope. -Notation "A ~2> B" := (subRel A B) - (at level 70, no associativity) : rel_scope. - -(* composing relation and function *) - -Definition Rel_app A B (D : Rel A) (f : B -> A) : Rel B := - fun x y => D (f x) (f y). - -Arguments Rel_app {A B} D f _ _ /. -Notation "D \\o f" := (@Rel_app _ _ D f) - (at level 42, left associativity). - -Lemma rel_app1 T (g : T -> T -> Prop) : g \\o id <~2> g. -Proof. by []. Qed. - -(* simplifying relations *) - -Definition Simpl_Rel T := T -> Simpl_Pred T. - -Coercion Rel_Of_Simpl T (sr : Simpl_Rel T) : Rel T := fun x : T => sr x. -Arguments Rel_Of_Simpl {T} sr x /. - -Notation xRelU := (fun (r1 r2 : Rel _) x y => r1 x y \/ r2 x y). -Notation xRelPre := (fun f (r : Rel _) x y => r (f x) (f y)). - -Definition PropSimplRel {T} (r : Rel T) : Simpl_Rel T := - fun x => PropSimplPred (r x). -Definition RelU {T} (r1 r2 : Rel T) := PropSimplRel (xRelU r1 r2). -Definition RelPre {aT rT} (f : aT -> rT) (r : Rel rT) := PropSimplRel (xRelPre f r). - -Notation "[ 'Rel' x y | E ]" := (PropSimplRel (fun x y => E)) - (x name, y name, only parsing) : fun_scope. -Notation "[ 'Rel' x y : T | E ]" := (PropSimplRel (fun x y : T => E)) - (x name, y name, only parsing) : fun_scope. - -Lemma subRelUl T (r1 r2 : Rel T) : subRel r1 (RelU r1 r2). -Proof. by left. Qed. - -Lemma subRelUr T (r1 r2 : Rel T) : subRel r2 (RelU r1 r2). -Proof. by right. Qed. - -(* Variant of Simpl_Pred specialised to the membership operator. *) - -Variant Mem_Pred T := PropMem of Pred T. - -Coercion Pred_Of_Mem {T} mp : {Pred T} := let: PropMem p := mp in [eta p]. -Canonical MemPredType T := PropPredType (@Pred_Of_Mem T). - -Definition In_Mem {T} (x : T) mp := Pred_Of_Mem mp x. -Definition Eq_Mem {T} mp1 mp2 := forall x : T, In_Mem x mp1 <-> In_Mem x mp2. -Definition Sub_Mem {T} mp1 mp2 := forall x : T, In_Mem x mp1 -> In_Mem x mp2. - -Arguments In_Mem {T} x mp : simpl never. -Global Typeclasses Opaque Eq_Mem Sub_Mem. - -(* -The Simpl_Of_Mem; Pred_Of_Simpl path provides a new Mem_Pred >-> Pred - coercion, but does not override the Pred_Of_Mem : Mem_Pred >-> Pred_sort - explicit coercion declaration above. -*) - -Coercion Simpl_Of_Mem {T} mp := PropSimplPred (fun x : T => In_Mem x mp). - -Lemma Sub_Refl T (mp : Mem_Pred T) : Sub_Mem mp mp. -Proof. by []. Qed. -Arguments Sub_Refl {T mp} [x] mp_x. - -Definition Mem T (pT : PredType T) : pT -> Mem_Pred T := - let: PropPredType _ toP := pT in fun A => PropMem [eta toP A]. -Arguments Mem {T pT} A : rename, simpl never. - -Notation "x \In A" := (In_Mem 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" := (Eq_Mem (Mem A) (Mem B)) - (at level 70, no associativity) : type_scope. -Notation "A <=p B" := (Sub_Mem (Mem A) (Mem B)) - (at level 70, no associativity) : type_scope. - -(* conversions *) -Lemma eq_conv T (A B : Pred T) : (A <~> B) = (A =p B). -Proof. by []. Qed. -Lemma sub_conv T (A B : Pred T) : (A ~> B) = (A <=p B). -Proof. by []. Qed. - -Notation "[ 'Mem' A ]" := - (Pred_Of_Simpl (Simpl_Of_Mem (Mem A))) (only parsing) : fun_scope. - -Notation "[ 'PredI' A & B ]" := (PredI [Mem A] [Mem B]) - (format "[ 'PredI' A & B ]") : fun_scope. -Notation "[ 'PredU' A & B ]" := (PredU [Mem A] [Mem B]) - (format "[ 'PredU' A & B ]") : fun_scope. -Notation "[ 'PredD' A & B ]" := (PredD [Mem A] [Mem B]) - (format "[ 'PredD' A & B ]") : fun_scope. -Notation "[ 'PredC' A ]" := (PredC [Mem A]) - (format "[ 'PredC' A ]") : fun_scope. -Notation "[ 'Preim' f 'of' A ]" := (Preim f [Mem A]) - (format "[ 'Preim' f 'of' A ]") : fun_scope. - -Notation "[ 'Pred' x 'In' A ]" := [Pred x | x \In A] - (x name) : fun_scope. -Notation "[ 'Pred' x 'In' A | E ]" := [Pred x | (x \In A) /\ E] - (x name) : fun_scope. -Notation "[ 'Pred' x 'In' A | E1 & E2 ]" := - [Pred x | x \In A & E1 /\ E2 ] - (x name) : fun_scope. - -(* infix notation for PredU and PredI *) -(* is sometimes more readable *) - -Notation "A \+p B" := (PredOfSimpl.Coerce (PredU A B)) - (at level 55, right associativity) : rel_scope. -Notation "A \*p B" := (PredOfSimpl.Coerce (PredI A B)) - (at level 45, right associativity) : rel_scope. - -Notation "[ 'Rel' x y 'In' A & B | E ]" := - [Rel x y | (x \In A) /\ (y \In B) /\ E] - (at level 0, x name, y name, - format "[ 'Rel' x y 'In' A & B | E ]") : fun_scope. -Notation "[ 'Rel' x y 'In' A & B ]" := [Rel x y | (x \In A) /\ (y \In B)] - (at level 0, x name, y name, - format "[ 'Rel' x y 'In' A & B ]") : fun_scope. -Notation "[ 'Rel' x y 'In' A | E ]" := [Rel x y In A & A | E] - (at level 0, x name, y name, - format "[ 'Rel' x y 'In' A | E ]") : fun_scope. -Notation "[ 'Rel' x y 'In' A ]" := [Rel x y In A & A] - (at level 0, x name, y name, - format "[ 'Rel' x y 'In' A ]") : fun_scope. - -Definition Applicative_Pred T := Pred T. -Definition Collective_Pred T := Pred T. -Coercion Applicative_Pred_Of_Simpl T (sp : Simpl_Pred T) : Applicative_Pred T := - fun_of_simpl sp. -Coercion Collective_Pred_Of_Simpl T (sp : Simpl_Pred T) : Collective_Pred T := - let: SimplFun p := sp in p. - -(* Explicit simplification rules for predicate application and membership. *) -Section PredicateSimplification. -Variables T : Type. -Implicit Types (p : Pred T) (pT : PredType T) (sp : Simpl_Pred T). -Implicit Types (mp : Mem_Pred T). - -Structure Registered_Applicative_Pred p := PropRegisteredApplicativePred { - Applicative_Pred_Value :> Pred T; - _ : Applicative_Pred_Value = p -}. -Definition PropApplicativePred p := PropRegisteredApplicativePred (erefl p). -Canonical Applicative_Pred_Applicative sp := - PropApplicativePred (Applicative_Pred_Of_Simpl sp). - -Structure Manifest_Simpl_Pred p := PropManifestSimplPred { - Simpl_Pred_Value :> Simpl_Pred T; - _ : Simpl_Pred_Value = PropSimplPred p -}. -Canonical Expose_Simpl_Pred p := PropManifestSimplPred (erefl (PropSimplPred p)). - -Structure Manifest_Mem_Pred p := PropManifestMemPred { - Mem_Pred_Value :> Mem_Pred T; - _ : Mem_Pred_Value = PropMem [eta p] -}. -Canonical Expose_Mem_Pred p := PropManifestMemPred (erefl (PropMem [eta p])). - -Structure Applicative_Mem_Pred p := - PropApplicativeMemPred {Applicative_Mem_Pred_Value :> Manifest_Mem_Pred p}. -Canonical Check_Applicative_Mem_Pred p (ap : Registered_Applicative_Pred p) := - [eta @PropApplicativeMemPred ap]. - -Lemma Mem_toPred pT (pp : pT) : Mem (toPred pp) = Mem pp. -Proof. by case: pT pp. Qed. - -Lemma toPredE pT x (pp : pT) : toPred pp x = (x \In pp). -Proof. by case: pT pp. Qed. - -Lemma App_PredE x p (ap : Registered_Applicative_Pred p) : ap x = (x \In p). -Proof. by case: ap=>ap /= ->. Qed. - -Lemma In_Applicative x p (amp : Applicative_Mem_Pred p) : In_Mem x amp = p x. -Proof. by case: amp; case=>amp /= ->. Qed. - -Lemma In_Collective x p (msp : Manifest_Simpl_Pred p) : - (x \In Collective_Pred_Of_Simpl msp) = p x. -Proof. by case: msp=>msp pf; rewrite -toPredE /= pf. Qed. - -Lemma In_Simpl x p (msp : Manifest_Simpl_Pred p) : - In_Mem x (PropMem [eta Pred_Of_Simpl msp]) = p x. -Proof. by case: msp=>msp /= ->. Qed. - -Lemma Unfold_In x p : (x \In ([eta p] : Pred T)) = p x. -Proof. by []. Qed. - -Lemma Simpl_PredE p : PropSimplPred p =1 p. -Proof. by []. Qed. - -Lemma Mem_Simpl sp : Mem sp = sp :> Pred T. -Proof. by []. Qed. - -Definition MemE := Mem_Simpl. -Lemma Mem_Mem mp : - (Mem mp = mp) * (Mem (mp : Simpl_Pred T) = mp) * (Mem (mp : Pred T) = mp). -Proof. by case: mp. Qed. - -End PredicateSimplification. - -(* Qualifiers and keyed predicates. *) - -Variant Qualifier (q : nat) T := PropQualifier of {Pred T}. - -Coercion Has_Quality n T (q : Qualifier n T) : {Pred T} := - fun x => let: PropQualifier p := q in p x. -Arguments Has_Quality n {T}. - -Lemma QualifE n T p x : (x \In @PropQualifier n T p) = p x. -Proof. by []. Qed. - -Notation "x \Is A" := (x \In Has_Quality 0 A) - (at level 0, only parsing) : fun_scope. -Notation "x \Is A" := (x \In Has_Quality 0 A) - (at level 0, only printing) : fun_scope. -Notation "x \Is 'a' A" := (x \In Has_Quality 1 A) - (at level 0, only parsing) : fun_scope. -Notation "x \Is 'a' A" := (x \In Has_Quality 1 A) - (at level 0, only printing) : fun_scope. -Notation "x \Is 'an' A" := (x \In Has_Quality 2 A) - (at level 0, only parsing) : fun_scope. -Notation "x \Is 'an' A" := (x \In Has_Quality 2 A) - (at level 0, only printing) : fun_scope. -Notation "x \Isn't A" := (x \Notin Has_Quality 0 A) - (at level 0) : fun_scope. -Notation "x \Isn't 'a' A" := (x \Notin Has_Quality 1 A) - (at level 0) : fun_scope. -Notation "x \Isn't 'an' A" := (x \Notin Has_Quality 2 A) - (at level 0) : fun_scope. -Notation "[ 'Qualify' x | P ]" := (PropQualifier 0 (fun x => P)) : form_scope. -Notation "[ 'Qualify' x : T | P ]" := - (PropQualifier 0 (fun x : T => P)) (only parsing) : form_scope. -Notation "[ 'Qualify' 'a' x | P ]" := - (PropQualifier 1 (fun x => P)) : form_scope. -Notation "[ 'Qualify' 'a' x : T | P ]" := - (PropQualifier 1 (fun x : T => P)) (only parsing) : form_scope. -Notation "[ 'Qualify' 'an' x | P ]" := - (PropQualifier 2 (fun x => P)) : form_scope. -Notation "[ 'Qualify' 'an' x : T | P ]" := - (PropQualifier 2 (fun x : T => P)) (only parsing) : form_scope. - -(* Keyed predicates: support for property-bearing predicate interfaces. *) - -Section KeyPred. -Variable T : Type. -Variant Pred_Key (p : {Pred T}) : Prop := PropDefaultPredKey. - -Variable p : {Pred T}. -Structure Keyed_Pred (k : Pred_Key p) := - PropPackKeyedPred {Unkey_Pred :> {Pred T}; _ : Unkey_Pred =p p}. - -Variable k : Pred_Key p. -Definition PropKeyedPred := @PropPackKeyedPred k p (fun x => iff_refl _). - -Variable k_p : Keyed_Pred k. -Lemma Keyed_PredE : k_p =p p. -Proof. by case: k_p. Qed. - -Canonical Keyed_Mem := - @PropPackKeyedPred k (Pred_Of_Mem (Mem k_p)) Keyed_PredE. -Canonical Leyed_Mem_Simpl := - @PropPackKeyedPred k (Pred_Of_Simpl (Mem k_p)) Keyed_PredE. - -End KeyPred. - -Local Notation In_Unkey x S := (x \In @Unkey_Pred _ S _ _) (only parsing). -Notation "x \In S" := (In_Unkey x S) (only printing) : fun_scope. - -Section KeyedQualifier. -Variables (T : Type) (n : nat) (q : Qualifier n T). - -Structure Keyed_Qualifier (k : Pred_Key q) := - PropPackKeyedQualifier {Unkey_Qualifier; _ : Unkey_Qualifier = q}. -Definition PropKeyedQualifier k := PropPackKeyedQualifier k (erefl q). -Variables (k : Pred_Key q) (k_q : Keyed_Qualifier k). -Fact Keyed_Qualifier_Subproof : Unkey_Qualifier k_q =p q. -Proof. by case: k_q=>kq pf x /=; rewrite pf. Qed. -Canonical Keyed_Qualifier_Keyed := PropPackKeyedPred k Keyed_Qualifier_Subproof. - -End KeyedQualifier. - -Notation "x \Is A" := - (In_Unkey x (Has_Quality 0 A)) (only printing) : fun_scope. -Notation "x \Is 'a' A" := - (In_Unkey x (Has_Quality 1 A)) (only printing) : fun_scope. -Notation "x \Is 'an' A" := - (In_Unkey x (Has_Quality 2 A)) (only printing) : fun_scope. - -Module DefaultKeying. -Canonical Default_Keyed_Pred T p := PropKeyedPred (@PropDefaultPredKey T p). -Canonical Default_Keyed_Qualifier T n (q : Qualifier n T) := - PropKeyedQualifier (PropDefaultPredKey q). -End DefaultKeying. - - -(**************************************) -(* Definitions and lemmas for setoids *) -(**************************************) - -(* Declaration of relations *) - -Section EqPred. -Context {T : Type} {pT : PredType T}. -Lemma EqPred_refl (r : pT) : Eq_Pred r r. Proof. by []. Qed. -Lemma EqPred_sym (r1 r2 : pT) : Eq_Pred r1 r2 -> Eq_Pred r2 r1. -Proof. by move=>H1 x; split; move/H1. Qed. -Lemma EqPred_trans' (r1 r2 r3 : pT) : - Eq_Pred r1 r2 -> Eq_Pred r2 r3 -> Eq_Pred r1 r3. -Proof. by move=>H1 H2 x; split; [move/H1; move/H2 | move/H2; move/H1]. Qed. -Definition EqPred_trans r2 r1 r3 := @EqPred_trans' r1 r2 r3. -End EqPred. - -#[export] Hint Resolve EqPred_refl : core. - -(* declare relations for all Eq_Pred and Eq_Mem *) - -Add Parametric Relation T (pT : PredType T) : pT (@Eq_Pred _ pT) - reflexivity proved by (@EqPred_refl _ _) - symmetry proved by (@EqPred_sym _ _) - transitivity proved by (@EqPred_trans' _ _) as EqPred_rel. - -(* do we need to add relations for mem_pred as well? *) -Add Parametric Relation T : (Mem_Pred T) (@Eq_Mem T) - reflexivity proved by (@EqPred_refl _ _) - symmetry proved by (@EqPred_sym _ _) - transitivity proved by (@EqPred_trans' _ _) as EqMem_rel. - -Section SubPred. -Context {T : Type} {pT : PredType T}. -Lemma SubPred_refl (r : pT) : Sub_Pred r r. Proof. by []. Qed. -Lemma SubPred_trans' (r1 r2 r3 : pT) : - Sub_Pred r1 r2 -> Sub_Pred r2 r3 -> Sub_Pred r1 r3. -Proof. by move=>H1 H2 x; move/H1; move/H2. Qed. -Definition SubPred_trans r2 r1 r3 := @SubPred_trans' r1 r2 r3. -End SubPred. - -#[export] Hint Resolve SubPred_refl : core. - -Add Parametric Relation T (pT : PredType T) : pT (@Sub_Pred _ pT) - reflexivity proved by (@SubPred_refl _ _) - transitivity proved by (@SubPred_trans' _ _) as SubPred_rel. - -Add Parametric Relation T : (Mem_Pred T) (@Sub_Mem _) - reflexivity proved by (@SubPred_refl _ _) - transitivity proved by (@SubPred_trans' _ _) as SubMem_rel. - -Section EqRel. -Context {T : Type}. -Lemma eqRel_refl (R : Rel T) : eqRel R R. Proof. by []. Qed. -Lemma eqRel_sym {R1 R2 : Rel T} : eqRel R1 R2 -> eqRel R2 R1. -Proof. by move=>H x y; rewrite H. Qed. -Lemma eqRel_trans {R1 R2 R3 : Rel T} : - eqRel R1 R2 -> eqRel R2 R3 -> eqRel R1 R3. -Proof. by move=>H1 H2 x y; rewrite H1 H2. Qed. -End EqRel. - -#[export] Hint Resolve eqRel_refl : core. - -Add Parametric Relation (T : Type) : (Rel T) (@eqRel T) - reflexivity proved by (@eqRel_refl T) - symmetry proved by (@eqRel_sym T) - transitivity proved by (@eqRel_trans T) as eqRel_rel. - -Section SubRel. -Context {T : Type}. -Lemma subRel_refl (R : Rel T) : subRel R R. Proof. by []. Qed. -Lemma subRel_trans (R1 R2 R3 : Rel T) : - subRel R1 R2 -> subRel R2 R3 -> subRel R1 R3. -Proof. by move=>H1 H2 x y /H1/H2. Qed. -End SubRel. - -#[export] Hint Resolve subRel_refl : core. - -Add Parametric Relation (T : Type) : (Rel T) (@subRel T) - reflexivity proved by (@subRel_refl T) - transitivity proved by (@subRel_trans T) as subRel_rel. - -(* Declaring morphisms. *) - -(* DEVCOMMENT *) -(* Annoyingly, even the coercions must be declared *) -(* /DEVCOMMENT *) - -(* -Add Parametric Morphism T : (@Pred_of_Simpl T) with signature - @EqSimplPred _ ==> @EqPredType T (PredPredType T) as Pred_of_Simpl_morph. -Proof. by []. Qed. -*) - -(* DEVCOMMENT *) -(* Do we need other coercions? We'll discover as we go *) - -(* Now the other morphisms. Again, not clear which ones are needed. *) -(* However, for all this to work, it seems that morphisms must be *) -(* declared with most specific signatures, or else the system *) -(* complains. For example, we use EqPred _ instead of EqPredType _ _, *) -(* even though the former is an instance of the later. *) - -(* -Add Parametric Morphism T : (@EqPred T) with signature - @EqPred _ ==> @EqPred _ ==> iff as EqPred_morph. -Proof. by move=>r1 s1 H1 r2 s2 H2; rewrite H1 H2. Qed. -*) -(* /DEVCOMMENT *) - -Add Parametric Morphism T (pT : PredType T) : (@Eq_Pred T pT) - with signature - @Eq_Pred _ _ ==> @Eq_Pred _ _ ==> iff as EqPred_morph. -Proof. by move=>r1 s1 H1 r2 s2 H2; rewrite H1 H2. Qed. - -Add Parametric Morphism T (pT : PredType T) : (@Sub_Pred T pT) - with signature - @Eq_Pred _ _ ==> @Eq_Pred _ _ ==> 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 : (@Eq_Mem T) - with signature @Eq_Mem _ ==> @Eq_Mem _ ==> iff as EqMem_morph. -Proof. by move=>r1 s1 H1 r2 s2 H2; rewrite H1 H2. Qed. - -Add Parametric Morphism T : (@Sub_Mem T) - with signature @Eq_Mem _ ==> @Eq_Mem _ ==> iff as SubMem_morph. -Proof. by move=>r1 s1 H1 r2 s2 H2; split=>H x; move/H1; move/H; move/H2. Qed. - -Add Parametric Morphism T : (@Eq_Pred T _) - with signature @Eq_Mem _ ==> @Eq_Mem _ ==> iff as EqPred_Mem_morph. -Proof. -move=>x y H x0 y0 H0; split=>X p /=; move: (H p) (H0 p)=>{}H {}H0; -rewrite /In_Mem /= in H H0. -- by rewrite -H -H0; apply/X. -by rewrite H H0; apply/X. -Qed. - -Add Parametric Morphism T : (@In_Mem T) with signature - @eq _ ==> @Eq_Mem _ ==> iff as InMem_morph. -Proof. by move=>x r s H; split; move/H. Qed. - -(* to rewrite under [Mem], must enable rewriting under *) -(* Mem, Simpl_Of_Mem and Pred_Of_Simpl *) - -(* rewriting by Eq_Pred under Mem *) -Add Parametric Morphism T (pT : PredType T) : (@Mem T pT) - with signature @Eq_Pred _ _ ==> @Eq_Pred _ _ as Mem_morph. -Proof. by case: pT=>pT F /= x y; apply. Qed. - -(* and under Eq_Mem \o Mem *) -Add Parametric Morphism T (pT : PredType T) : (@Mem T pT) - with signature @Eq_Pred _ _ ==> @Eq_Mem _ as Mem_morph'. -Proof. by move=>x y H p; rewrite /Eq_Pred -!toPredE in H *; apply: H. Qed. - -(* rewriting under Simpl_Of_Mem *) -Add Parametric Morphism T : (@Simpl_Of_Mem T) - with signature @Eq_Pred _ _ ==> @Eq_Pred _ _ as SimplOfMem_morph. -Proof. by []. Qed. - -(* rewriting under Pred_Of_Simpl *) -Add Parametric Morphism T : (@Pred_Of_Simpl T) - with signature @Eq_Pred _ _ ==> @Eq_Pred _ _ as Pred_Of_Simpl_morph. -Proof. by []. Qed. - -(* rewriting by Eq_Mem under Eq_Pred *) -Add Parametric Morphism T : (@Mem T _) - with signature @Eq_Mem _ ==> @Eq_Pred _ _ as Mem_Morph. -Proof. by []. Qed. - -Add Parametric Morphism T : (@Simpl_Of_Mem T) - with signature @Eq_Mem _ ==> @Eq_Pred _ _ as SimplOfMem_morph'. -Proof. by []. Qed. - -(* rewriting under Pred_Of_Simpl *) -Add Parametric Morphism T : (@Pred_Of_Simpl T) - with signature @Eq_Mem _ ==> @Eq_Pred _ _ as Pred_Of_Simpl_morph'. -Proof. by []. Qed. - -Add Parametric Morphism T : (@PredU T) with signature - @Eq_Pred T _ ==> @Eq_Pred T _ ==> @Eq_Pred T _ 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 - @Eq_Pred _ _ ==> @Eq_Pred _ _ ==> @Eq_Pred _ _ 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 - @Eq_Pred _ _ ==> @Eq_Pred _ _ as predC_morph. -Proof. by move=>r s H x; split=>H1; apply/H. Qed. - -Add Parametric Morphism A : (@subRel A) with signature - eqRel ==> eqRel ==> iff as subrel_eq. -Proof. by move=>x y H x1 y1 H1; split=>X a b /H/X/H1. Qed. - -Add Parametric Morphism A B : (@Rel_app A B) with signature - eqRel ==> eq ==> eqRel as Relapp_eq. -Proof. by move=>x y H f s s'; split=>/H. Qed. - -Add Parametric Morphism A B : (@Rel_app A B) with signature - subRel ==> eq ==> subRel as Relapp_sub. -Proof. by move=>x y H f s s' /H. Qed. - -(* Can we declare relation application as morphism? *) -(* We can, but it isn't picked up in practice. *) -(* This is so because application is a special term *) -(* that apprently can't be matched on. *) -(* If we want to rewrite under application *) -(* the choices are: *) -(* a) don't use setoids, but prove equalites *) -(* b) move to collective predicates where *) -(* function application is user-level term *) -(* -Add Parametric Morphism A : (fun (R : Rel A) (x y : A) => R x y) - with signature eqrel ==> eq ==> eq ==> iff as app_eq. -Proof. by move=>r1 r2 E x y; split=>/E. 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. by split; (case=>/=; [right|left]). 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. simpl. -Proof. by move=>x /=; intuition. Qed. - -Lemma orrAC r1 r2 r3 : - (r1 \+p r2) \+p r3 <~> (r1 \+p r3) \+p r2. -Proof. by move=>x /=; intuition. Qed. - -Lemma orrA r1 r2 r3 : - (r1 \+p r2) \+p r3 <~> r1 \+p r2 \+p r3. -Proof. by rewrite (orrC r2) orrCA (orrC r3). Qed. - -(* absorption *) -(* DEVCOMMENT: when lemmas use sub_rel, it's usually better *) -(* to use the collective form sub_mem *) -(* Typically, with sub_rel, there will be lemma application *) -(* that will turn expression of the from x \In A into A x *) -(* and the latter immediately prevents setoid rewriting *) -(* /DEVCOMMENT *) -Lemma orrAb r1 r2 : - r1 <~> r1 \+p r2 <-> - r2 <=p r1. -Proof. -split; first by move=>-> x; rewrite -!toPredE /=; auto. -move=>H x /=; split; first by auto. -by case=>//; move/H. -Qed. - -Lemma sub_orl r1 r2 : r1 <=p r1 \+p r2. Proof. by left. Qed. -Lemma sub_orr r1 r2 : r2 <=p 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 <~> 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. - -#[export] 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_Seq (s : EqSeq_Class) : {Pred T} := [eta Mem_Seq s]. -Canonical Structure seq_PredType := - PropPredType (Pred_Of_Seq : seq T -> Pred T). -(* The line below makes mem_seq a canonical instance of topred. *) -Canonical Structure Mem_Seq_PredType := PropPredType 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_Applicative, In_Simpl, Simpl_PredE)). - -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. - -Prenex Implicits In_split. - -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. - -Lemma MapP T1 T2 (f : T1 -> T2) (s : seq T1) (y : T2) : - y \In map f s <-> exists2 x, x \In s & y = f x. -Proof. -elim: s => [|x s IHs] /=; first by split=>//; case. -rewrite In_cons; split. -- case=>[->|]; first by exists x=>//; apply/In_cons; left. - by case/IHs=>k H ->; exists k=>//; apply/In_cons; right. -case=>k /In_cons [->|H E]; first by left. -by right; apply/IHs; exists k. -Qed. - -Lemma Mem_filter (T : Type) (a : pred T) (x : T) (s : seq T) : - x \In filter a s <-> a x /\ x \In s. -Proof. -elim: s; first by split; case. -move=>y s IHs /=; rewrite 2!fun_if /=. -case E: (a y). -- rewrite InE IHs; split; last first. - - by case=>H1 /In_cons [->|]; [left | right]. - case=>[->|]; first by split=>//; apply/In_cons; left. - by case=>H1 H2; split=>//; apply/In_cons; right. -rewrite IHs; split. -- by case=>H1 H2; split=>//; apply/In_cons; right. -by case=>H1 /In_cons [] // ?; subst y; rewrite E in H1. -Qed. - -Lemma eq_In_filter (T : Type) a1 a2 (s : seq T) : - (forall x, x \In s -> a1 x = a2 x) -> - filter a1 s = filter a2 s. -Proof. -elim: s => //= x s IHs eq_a. -rewrite eq_a; last by rewrite InE; left. -rewrite IHs // => y s_y; apply: eq_a. -by rewrite InE; right. -Qed. - -(* 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. - -(* DEVCOMMENT *) -(* this interferes with the usage of inE, see - https://gitlab.software.imdea.org/mathador/fcsl/-/issues/97 *) -(* -(* 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. -*) -(* /DEVCOMMENT *) - -(* Big \In equivalences for all and has *) - -Section allhasP. -Context {T : Type}. -Variables (p : pred T). - -Lemma allPIn (s : seq T) : - reflect (forall x, x \In s -> p x) (all p s). -Proof. -elim: s=>[|x s IHs] /=; first by constructor. -case: andP=>[[H1 H2]|H]; constructor. -- by move=>z /In_cons [->|/(IHs H2)]. -move=>H1; elim: H; split; first by apply/H1/In_cons; left. -by apply/IHs=>z H; apply/H1/In_cons; right. -Qed. - -Lemma allPnIn (s : seq T) : - reflect (exists2 x, x \In s & ~~ p x) (~~ all p s). -Proof. -elim: s => [|x s IHs] /=; first by constructor=> [[x Hx _]]. -rewrite /= andbC negb_and; case: IHs => IHs /=. -- by constructor; case: IHs => y Hy Hay; exists y=>//; apply/In_cons; right. -apply: (iffP idP) => [|[y]]; first by exists x=>//; apply/In_cons; left. -by case/In_cons=>[->//|H1 H2]; elim: IHs; exists y. -Qed. - -Lemma hasPIn (s : seq T) : - reflect (exists2 x, x \In s & p x) (has p s). -Proof. -elim: s => [|y s IHs] /=; first by right; case. -case Py: (p y); first by left; exists y=>//; apply/In_cons; left. -apply: (iffP IHs)=>[] [x ysx Px]; exists x => //; first by apply/In_cons; right. -by case/In_cons: ysx Px Py=>[->->|]. -Qed. - -Lemma hasPnIn (s : seq T) : - reflect (forall x, x \In s -> ~~ p x) (~~ has p s). -Proof. -apply: (iffP (negPP (hasPIn s)))=>H. -- by move=>x Hx; apply: contra_notN H=>Px; exists x. -by case=>x Hx; apply/negP/H. -Qed. - -(* implication form of hasPIn is frequently used *) -(* and if you don't have it, causes fluff in proofs *) - -Lemma hasPInX x xs : x \In xs -> p x -> has p xs. -Proof. by move=>H1 H2; apply/hasPIn; exists x. Qed. - -End allhasP. - -Arguments hasPInX {T p x xs}. - -Section AllHasP. -Context {T : Type}. -Variables (P : Pred T). - -Fixpoint All xs := if xs is x :: xs then P x /\ All xs else True. - -Lemma AllP xs : All xs <-> forall x, x \In xs -> P x. -Proof. -elim: xs=>[|x xs IH] //=. -split; first by case=>H1 /IH H2 z; rewrite InE; case=>[->//|]; apply: H2. -move=>H; split; first by apply: H; rewrite InE; left. -by apply/IH=>z Z; apply: H; rewrite InE; right. -Qed. - -Lemma All_cat (s1 s2 : seq T) : - All (s1 ++ s2) <-> All s1 /\ All s2. -Proof. -split. -- by move/AllP=>H; split; apply/AllP=>x Hx; apply/H/Mem_cat; [left|right]. -by case=>/AllP H1 /AllP H2; apply/AllP=>x /Mem_cat; case=>Hx; [apply: H1| apply: H2]. -Qed. - -Fixpoint Has xs := if xs is x :: xs then P x \/ Has xs else False. - -Lemma HasP xs : Has xs <-> exists2 x, x \In xs & P x. -Proof. -elim: xs=>[|x xs IH] /=; first by split=>//; case. -split. -- case=>[H|/IH]; first by exists x=>//; rewrite InE; left. - by case=>y H1 H2; exists y=>//; rewrite InE; right. -case=>y; rewrite InE; case=>[->|H1 H2]; first by left. -by right; apply/IH; exists y. -Qed. - -Lemma Has_cat (s1 s2 : seq T) : - Has (s1 ++ s2) <-> Has s1 \/ Has s2. -Proof. -split. -- by move/HasP=>[x] /Mem_cat; case=>Hx Px; [left|right]; apply/HasP; exists x. -by case=>/HasP [x Hx Px]; apply/HasP; exists x=>//; apply/Mem_cat; [left|right]. -Qed. - -End AllHasP. - -Section AllHasT. -Context {T : Type}. -Variables (P : T -> Type). - -Fixpoint AllT xs : Type := if xs is x :: xs then P x * AllT xs else unit. - -Fixpoint HasT xs : Type := if xs is x :: xs then P x + HasT xs else Empty_set. - -End AllHasT. - -(* prop uniqueness *) -Fixpoint Uniq T (xs : seq T) := - if xs is x :: xs then x \Notin xs /\ Uniq xs else True. - -Lemma UniqP (T : eqType) (xs : seq T) : - reflect (Uniq xs) (uniq xs). -Proof. -elim: xs=>[|x xs IH] //=; first by apply: ReflectT. -case: andP=>H; constructor. -- by case: H=>/mem_seqP H /IH. -by case=>/mem_seqP H1 /IH H2; elim: H. -Qed. - -(***********************************) -(* 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 name, - 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 name, - 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 name, 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 name, 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 <~> Image f X2 -> - X1 <~> 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) <~> 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) <~> 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 <~> Image f2 X. -Proof. by move=>H x; split; case=>a ->; exists a. Qed. - -(********************************) -(* Relation preserving function *) -(********************************) - -Definition rel_on_func A B (f : B -> A) (p : Pred B) (g : Rel A) := - forall x y, p x -> g (f x) y -> - exists y', [/\ y = f y', p y' & (g \\o f) x y']. - -Notation "'on-' f" := (rel_on_func f) - (at level 1, format "'on-' f"). - -(* last conjunct can be dropped *) -Lemma onE A B (f : B -> A) p g : - on-f p g <-> - forall x y, p x -> g (f x) y -> exists2 y', y = f y' & p y'. -Proof. -split=>H x y C /[dup] X /(H _ _ C). -- by case=>y' []; exists y'. -by case=>y' E; exists y'; rewrite /Rel_app -E. -Qed. - -Add Parametric Morphism A B : (@rel_on_func A B) with signature - eq ==> @Eq_Pred _ _ ==> eqRel ==> iff as relon_eq. -Proof. -move=>f p1 p2 H r1 r2 E; split=>X; -by apply/onE=>x y /H /= P /E /(X _ _ P) [y'][-> /H]; eauto. -Qed. - -Add Parametric Morphism A B : (@rel_on_func A B) with signature - eq ==> @Eq_Pred _ _ ==> subRel ==> flip impl as relon_sub. -Proof. -move=>f p1 p2 H r1 r2 E X. -by apply/onE=>x y /H /= P /E /(X _ _ P) [y'][-> /H]; eauto. -Qed. - -(**********************************) -(**********************************) -(* Theory 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. - -(* DEVCOMMENT *) -(* Using sym_left_Transitive as a view doesn't work. *) -(* see https://github.com/coq/coq/issues/8352 *) -(* /DEVCOMMENT *) - -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). - -(* DEVCOMMENT *) -(* TODO: see if these can be simplified *) -(* /DEVCOMMENT *) - -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. - -#[export] 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. - -Section IteratedRels. -Variable T : Type. -Implicit Type g : T -> T -> Prop. - -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. - -Lemma iterS g n s1 s2 : - iter' g n.+1 s1 s2 <-> exists s, iter' g n s1 s /\ g s s2. -Proof. -rewrite -addn1; split; first by case/iter'_split=>s [H1] [s'][H2 <-]; exists s. -by case=>s [H1 H2]; apply: iter'_add H1 _; exists s2. -Qed. - -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 iter_sub g g' : g ~2> g' -> iter g ~2> iter g'. -Proof. -move=>H s s' [n X]; exists n; elim: n s s' X=>[|n IH] s s' //=. -by case=>x [/H G] /IH; exists x. -Qed. - -Lemma iter_eq g g' : g <~2> g' -> iter g <~2> iter g'. -Proof. by move=>H s s'; split; apply/iter_sub; rewrite H. Qed. - -Lemma iter1 g : g ~2> iter g. -Proof. by move=>s s'; exists 1, s'. Qed. - -Lemma iter_app T' g (f : T' -> T) : iter (g \\o f) ~2> iter g \\o f. -Proof. -move=>s s' [n] H; exists n; elim: n s s' H=>[|n IH] s s' /=; first by move=>->. -by case=>x [H] /IH; exists (f x). -Qed. - -Lemma iter_on T' (f : T' -> T) p g : on-f p g -> on-f p (iter g). -Proof. -move=>O; apply/onE=>x y C [n]. -elim: n x y C=>[|n IH] x y /= C; first by move=><-; exists x. -case=>_ [/(O _ _ C)] [x'][->] X' H /(IH _ _ X') [y' ->] Y'. -by exists y'. -Qed. - -Lemma iter_subf T' g (f : T' -> T) : iter (g \\o f) ~2> iter g \\o f. -Proof. -move=>s s' [n] H; exists n. -elim: n s s' H=>[|n IH] s s'; first by move=>->. -by case=>x [H1] {}/IH; exists (f x). -Qed. - -Lemma iter_fsub T' g (f : T' -> T) : - injective f -> - on-f xpredT g -> - iter g \\o f ~2> iter (g \\o f). -Proof. -move=>H1 H2 s s' [n] H; exists n. -elim: n s s' H =>[|n IH] s s' /=; first by move/H1. -by case=>_ [] /(H2 _ _ erefl) [x][-> _ X] {}/IH; exists x. -Qed. - -Lemma iter_eqf T' g (f : T' -> T) : - injective f -> - on-f xpredT g -> - iter g \\o f <~2> iter (g \\o f). -Proof. by move=>??; split; [apply: iter_fsub|apply: iter_subf]. Qed. - -(* localized variant of iter_fsub *) -Lemma iter_fsubl T' g (f : T' -> T) p s s' : - injective f -> - on-f p g -> - p s -> - (iter g \\o f) s s' -> iter (g \\o f) s s'. -Proof. -move=>H1 H2 P [n] H; exists n. -elim: n s s' P H=>[|n IH] s s' /= Ps; first by move/H1=>->. -by case=>_ [] /(H2 _ _ Ps) [x][-> Px H] /(IH _ _ Px); exists x. -Qed. - -End IteratedRels. - -Lemma iter_pair {T} {g : unit*T -> unit*T -> Prop} : - iter g \\o pair tt <~2> iter (g \\o pair tt). -Proof. by apply: iter_eqf=>[x y []|x [[y]]] //; exists y. Qed. - -Add Parametric Morphism A : (@iter A) with signature - eqRel ==> eqRel as iter_eq'. -Proof. by move=>x y H x1 y1; rewrite (iter_eq H _). Qed. - -Add Parametric Morphism A : (@iter A) with signature - subRel ==> subRel as iter_sub'. -Proof. by move=>x y H x1 y1; apply: iter_sub. 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. - -Arguments iter1 {T}%_type_scope {g _ _}. -Arguments iter_app {T T'}%_type_scope {g f}. -Arguments iter_refl {T g s}. - -#[export] 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 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 split. 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 _ _ /. - -Lemma union_eq A (x1 y1 x2 y2 : Rel A) : - eqRel x1 y1 -> eqRel x2 y2 -> - eqRel (Union_rel x1 x2) (Union_rel y1 y2). -Proof. by rewrite /eqRel/Union_rel=>E1 E2 x y; rewrite E1 E2. Qed. - -Arguments union_eq {A x1 y1 x2 y2}. -Prenex Implicits union_eq. - -Add Parametric Morphism A : (@Union_rel A) with signature - eqRel ==> eqRel ==> eqRel as union_eq'. -Proof. by move=>*; apply: union_eq. Qed. - -Lemma union_sub A (x1 y1 x2 y2 : Rel A) : - subRel x1 y1 -> subRel x2 y2 -> - subRel (Union_rel x1 x2) (Union_rel y1 y2). -Proof. by move=>E1 E2 x y [/E1|/E2]; [left | right]. Qed. - -Arguments union_sub {A x1 y1 x2 y2}. -Prenex Implicits union_sub. - -Add Parametric Morphism A : (@Union_rel A) with signature - subRel ==> subRel ==> subRel as union_sub1. -Proof. by move=>*; apply: union_sub. Qed. - -Lemma subrelL A (x y : Rel A) : x ~2> Union_rel x y. -Proof. by left. Qed. - -Lemma subrelR A (x y : Rel A) : y ~2> Union_rel x y. -Proof. by right. Qed. - -#[export] -Instance eqrel_subrel A : subrelation (@eqRel A) (@subRel A). -Proof. by move=>x y H ?? /H. Qed. - -Lemma union_app A B D1 D2 (f : B -> A) : - Union_rel D1 D2 \\o f <~2> - Union_rel (D1 \\o f) (D2 \\o f). -Proof. by apply: union_eq. Qed. - -Arguments union_app {A B}%_type_scope {D1 D2 f}%_function_scope. -Prenex Implicits union_app. - -Lemma on_union A B D1 D2 (f : B -> A) p : - on-f p D1 -> - on-f p D2 -> - on-f p (Union_rel D1 D2). -Proof. -move=>H1 H2; apply/onE=>x y P; case. -- by case/(H1 _ _ P)=>y' [->] P'; exists y'. -by case/(H2 _ _ P)=>y' [->] P'; exists y'. -Qed. - -(*************************) -(* 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). - -(* DEVCOMMENT *) -(* (d2 d2' : mem_pred T2) (d3 d3' : mem_pred T3). *) -(* /DEVCOMMENT *) - -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. - -(* we can now state localized version of iter_fsubl *) -Lemma iter_fsubl' T T' (g : Rel T) (f : T' -> T) p : - injective f -> - on-f p g -> - {In p & xPredT, iter g \\o f ~2> iter (g \\o f)}. -Proof. by move=>H1 H2 x y H _; apply: iter_fsubl H1 H2 H. Qed. - - diff --git a/core/prelude.v b/core/prelude.v deleted file mode 100644 index 380c43a..0000000 --- a/core/prelude.v +++ /dev/null @@ -1,1391 +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 HB Require Import structures. -From Coq Require Import ssreflect ssrfun Eqdep. -From mathcomp Require Import ssrbool ssrnat seq eqtype choice path. -From mathcomp Require Import fintype finset finfun tuple perm fingroup. -From pcm Require Import options axioms. - -(***********) -(* Prelude *) -(***********) - -(* 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}. - -(* 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}. - -(* 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. - -Lemma trans_eq A (x y z : A) : - x = y -> - x = z -> - y = z. -Proof. by move/esym; apply: eq_trans. 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. - -(************) -(* Products *) -(************) - -Inductive Prod3 U1 U2 U3 := - mk3 {proj31 : U1; proj32 : U2; proj33 : U3}. -Inductive Prod4 U1 U2 U3 U4 := - mk4 {proj41 : U1; proj42 : U2; proj43 : U3; proj44 : U4}. -Inductive Prod5 U1 U2 U3 U4 U5 := - mk5 {proj51 : U1; proj52 : U2; proj53 : U3; proj54 : U4; proj55 : U5}. -Inductive Prod6 U1 U2 U3 U4 U5 U6 := - mk6 {proj61 : U1; proj62 : U2; proj63 : U3; - proj64 : U4; proj65 : U5; proj66 : U6}. -Inductive Prod7 U1 U2 U3 U4 U5 U6 U7 := - mk7 {proj71 : U1; proj72 : U2; proj73 : U3; - proj74 : U4; proj75 : U5; proj76 : U6; proj77 : U7}. -Prenex Implicits proj31 proj32 proj33. -Prenex Implicits proj41 proj42 proj43 proj44. -Prenex Implicits proj51 proj52 proj53 proj54 proj55. -Prenex Implicits proj61 proj62 proj63 proj64 proj65 proj66. -Prenex Implicits proj71 proj72 proj73 proj74 proj75 proj76 proj77. - -Definition eq3 (U1 U2 U3 : eqType) (x y : Prod3 U1 U2 U3) := - [&& proj31 x == proj31 y, proj32 x == proj32 y & proj33 x == proj33 y]. -Definition eq4 (U1 U2 U3 U4 : eqType) (x y : Prod4 U1 U2 U3 U4) := - [&& proj41 x == proj41 y, proj42 x == proj42 y, proj43 x == proj43 y & - proj44 x == proj44 y]. -Definition eq5 (U1 U2 U3 U4 U5 : eqType) (x y : Prod5 U1 U2 U3 U4 U5) := - [&& proj51 x == proj51 y, proj52 x == proj52 y, proj53 x == proj53 y, - proj54 x == proj54 y & proj55 x == proj55 y]. -Definition eq6 (U1 U2 U3 U4 U5 U6 : eqType) (x y : Prod6 U1 U2 U3 U4 U5 U6) := - [&& proj61 x == proj61 y, proj62 x == proj62 y, proj63 x == proj63 y, - proj64 x == proj64 y, proj65 x == proj65 y & proj66 x == proj66 y]. -Definition eq7 (U1 U2 U3 U4 U5 U6 U7 : eqType) (x y : Prod7 U1 U2 U3 U4 U5 U6 U7) := - [&& proj71 x == proj71 y, proj72 x == proj72 y, proj73 x == proj73 y, - proj74 x == proj74 y, proj75 x == proj75 y, proj76 x == proj76 y & - proj77 x == proj77 y]. - -Lemma eq3P U1 U2 U3 : - Equality.axiom (@eq3 U1 U2 U3). -Proof. -rewrite /eq3; case=>x1 x2 x3 [y1 y2 y3] /=. -by do ![case: eqP=>[->|]]; constructor=>//; case. -Qed. - -Lemma eq4P U1 U2 U3 U4 : - Equality.axiom (@eq4 U1 U2 U3 U4). -Proof. -rewrite /eq4; case=>x1 x2 x3 x4 [y1 y2 y3 y4] /=. -by do ![case: eqP=>[->|]]; constructor=>//; case. -Qed. - -Lemma eq5P U1 U2 U3 U4 U5 : - Equality.axiom (@eq5 U1 U2 U3 U4 U5). -Proof. -rewrite /eq5; case=>x1 x2 x3 x4 x5 [y1 y2 y3 y4 y5] /=. -by do ![case: eqP=>[->|]]; constructor=>//; case. -Qed. - -Lemma eq6P U1 U2 U3 U4 U5 U6 : - Equality.axiom (@eq6 U1 U2 U3 U4 U5 U6). -Proof. -rewrite /eq6; case=>x1 x2 x3 x4 x5 x6 [y1 y2 y3 y4 y5 y6] /=. -by do ![case: eqP=>[->|]]; constructor=>//; case. -Qed. - -Lemma eq7P U1 U2 U3 U4 U5 U6 U7 : - Equality.axiom (@eq7 U1 U2 U3 U4 U5 U6 U7). -Proof. -rewrite /eq7; case=>x1 x2 x3 x4 x5 x6 x7 [y1 y2 y3 y4 y5 y6 y7] /=. -by do ![case: eqP=>[->|]]; constructor=>//; case. -Qed. - -HB.instance Definition _ (U1 U2 U3 : eqType) := - hasDecEq.Build (Prod3 U1 U2 U3) (@eq3P U1 U2 U3). -HB.instance Definition _ (U1 U2 U3 U4 : eqType) := - hasDecEq.Build (Prod4 U1 U2 U3 U4) (@eq4P U1 U2 U3 U4). -HB.instance Definition _ (U1 U2 U3 U4 U5 : eqType) := - hasDecEq.Build (Prod5 U1 U2 U3 U4 U5) (@eq5P U1 U2 U3 U4 U5). -HB.instance Definition _ (U1 U2 U3 U4 U5 U6 : eqType) := - hasDecEq.Build (Prod6 U1 U2 U3 U4 U5 U6) (@eq6P U1 U2 U3 U4 U5 U6). -HB.instance Definition _ (U1 U2 U3 U4 U5 U6 U7 : eqType) := - hasDecEq.Build (Prod7 U1 U2 U3 U4 U5 U6 U7) (@eq7P U1 U2 U3 U4 U5 U6 U7). - -(***************************) -(* 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. - -Lemma compE A B C (g : B -> C) (f : A -> B) x : - g (f x) = (g \o f) x. -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, aka. fork/fanout/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. - -(* composing relation and function *) -Definition rel_app A B (S : rel A) f : rel B := - fun x y => S (f x) (f y). -Arguments rel_app {A B} S f _ _ /. - -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. - -(************************) -(* 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 , P6 , P7 , P8 , P9 , P10 & P11 ]" (at level 0, format - "'[hv' [ /\ '[' P1 , '/' P2 , '/' P3 , '/' P4 , '/' P5 , '/' P6 , '/' P7 , '/' P8 , '/' P9 , '/' P10 ']' '/ ' & P11 ] ']'"). -Reserved Notation "[ /\ P1 , P2 , P3 , P4 , P5 , P6 , P7 , P8 , P9 , P10 , P11 & P12 ]" (at level 0, format - "'[hv' [ /\ '[' P1 , '/' P2 , '/' P3 , '/' P4 , '/' P5 , '/' P6 , '/' P7 , '/' P8 , '/' P9 , '/' P10 , '/' P11 ']' '/ ' & P12 ] ']'"). -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 and11 (P1 P2 P3 P4 P5 P6 P7 P8 P9 P10 P11 : Prop) : Prop := - And11 of P1 & P2 & P3 & P4 & P5 & P6 & P7 & P8 & P9 & P10 & P11. -Inductive and12 (P1 P2 P3 P4 P5 P6 P7 P8 P9 P10 P11 P12 : Prop) : Prop := - And12 of P1 & P2 & P3 & P4 & P5 & P6 & P7 & P8 & P9 & P10 & P11 & P12. - -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 , P6 , P7 , P8 , P9 , P10 & P11 ]" := - (and11 P1 P2 P3 P4 P5 P6 P7 P8 P9 P10 P11) : type_scope. -Notation "[ /\ P1 , P2 , P3 , P4 , P5 , P6 , P7 , P8 , P9 , P10 , P11 & P12 ]" := - (and12 P1 P2 P3 P4 P5 P6 P7 P8 P9 P10 P11 P12) : 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 *) - -(* DEVCOMMENT *) -(* TODO: we should move some of the following to [ssrbool] in Coq *) -(* /DEVCOMMENT *) - -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 *) -#[export] Program Instance and3_impl_morphism : Proper (iter_respectful impl impl 3) and3 | 1. -#[export] Program Instance and3_iff_morphism : Proper (iter_respectful iff iff 3) and3 | 1. - -#[export] Program Instance and4_impl_morphism : Proper (iter_respectful impl impl 4) and4 | 1. -#[export] Program Instance and4_iff_morphism : Proper (iter_respectful iff iff 4) and4 | 1. - -#[export] Program Instance and5_impl_morphism : Proper (iter_respectful impl impl 5) and5 | 1. -#[export] Program Instance and5_iff_morphism : Proper (iter_respectful iff iff 5) and5 | 1. - -#[export] Program Instance and6_impl_morphism : Proper (iter_respectful impl impl 6) and6 | 1. -#[export] Program Instance and6_iff_morphism : Proper (iter_respectful iff iff 6) and6 | 1. - -#[export] Program Instance and7_impl_morphism : Proper (iter_respectful impl impl 7) and7 | 1. -#[export] Program Instance and7_iff_morphism : Proper (iter_respectful iff iff 7) and7 | 1. - -#[export] Program Instance and8_impl_morphism : Proper (iter_respectful impl impl 8) and8 | 1. -#[export] Program Instance and8_iff_morphism : Proper (iter_respectful iff iff 8) and8 | 1. - -#[export] Program Instance and9_impl_morphism : Proper (iter_respectful impl impl 9) and9 | 1. -#[export] Program Instance and9_iff_morphism : Proper (iter_respectful iff iff 9) and9 | 1. - -#[export] Program Instance and10_impl_morphism : Proper (iter_respectful impl impl 10) and10 | 1. -#[export] Program Instance and10_iff_morphism : Proper (iter_respectful iff iff 10) and10 | 1. - -#[export] Program Instance and11_impl_morphism : Proper (iter_respectful impl impl 11) and11 | 1. -#[export] Program Instance and11_iff_morphism : Proper (iter_respectful iff iff 11) and11 | 1. - -#[export] Program Instance and12_impl_morphism : Proper (iter_respectful impl impl 12) and12 | 1. -#[export] Program Instance and12_iff_morphism : Proper (iter_respectful iff iff 12) and12 | 1. - -(** Logical disjunction *) -#[export] Program Instance or3_impl_morphism : Proper (iter_respectful impl impl 3) or3 | 1. -#[export] Program Instance or3_iff_morphism : Proper (iter_respectful iff iff 3) or3 | 1. - -#[export] Program Instance or4_impl_morphism : Proper (iter_respectful impl impl 4) or4 | 1. -#[export] Program Instance or4_iff_morphism : Proper (iter_respectful iff iff 4) or4 | 1. - -#[export] Program Instance or5_impl_morphism : Proper (iter_respectful impl impl 5) or5 | 1. -#[export] Program Instance or5_iff_morphism : Proper (iter_respectful iff iff 5) or5 | 1. - -#[export] Program Instance or6_impl_morphism : Proper (iter_respectful impl impl 6) or6 | 1. -#[export] Program Instance or6_iff_morphism : Proper (iter_respectful iff iff 6) or6 | 1. - -#[export] Program Instance or7_impl_morphism : Proper (iter_respectful impl impl 7) or7 | 1. -#[export] 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 b11 b12 : bool. - -Lemma and6P : reflect [/\ b1, b2, b3, b4, b5 & b6] [&& b1, b2, b3, b4, b5 & b6]. -Proof. -case: b1=>/=; last by constructor; case. -case: b2=>/=; last by constructor; case. -case: b3=>/=; last by constructor; case. -case: b4=>/=; last by constructor; case. -case: b5=>/=; last by constructor; case. -case: b6=>/=; last by constructor; case. -by constructor. -Qed. - -Lemma and7P : reflect [/\ b1, b2, b3, b4, b5, b6 & b7] [&& b1, b2, b3, b4, b5, b6 & b7]. -Proof. -case: b1=>/=; last by constructor; case. -case: b2=>/=; last by constructor; case. -case: b3=>/=; last by constructor; case. -case: b4=>/=; last by constructor; case. -case: b5=>/=; last by constructor; case. -case: b6=>/=; last by constructor; case. -case: b7=>/=; last by constructor; case. -by constructor. -Qed. - -Lemma and8P : reflect [/\ b1, b2, b3, b4, b5, b6, b7 & b8] [&& b1, b2, b3, b4, b5, b6, b7 & b8]. -Proof. -case: b1=>/=; last by constructor; case. -case: b2=>/=; last by constructor; case. -case: b3=>/=; last by constructor; case. -case: b4=>/=; last by constructor; case. -case: b5=>/=; last by constructor; case. -case: b6=>/=; last by constructor; case. -case: b7=>/=; last by constructor; case. -case: b8=>/=; last by constructor; case. -by constructor. -Qed. - -Lemma and9P : reflect [/\ b1, b2, b3, b4, b5, b6, b7, b8 & b9] [&& b1, b2, b3, b4, b5, b6, b7, b8 & b9]. -Proof. -case: b1=>/=; last by constructor; case. -case: b2=>/=; last by constructor; case. -case: b3=>/=; last by constructor; case. -case: b4=>/=; last by constructor; case. -case: b5=>/=; last by constructor; case. -case: b6=>/=; last by constructor; case. -case: b7=>/=; last by constructor; case. -case: b8=>/=; last by constructor; case. -case: b9=>/=; last by constructor; case. -by constructor. -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. -case: b1=>/=; last by constructor; case. -case: b2=>/=; last by constructor; case. -case: b3=>/=; last by constructor; case. -case: b4=>/=; last by constructor; case. -case: b5=>/=; last by constructor; case. -case: b6=>/=; last by constructor; case. -case: b7=>/=; last by constructor; case. -case: b8=>/=; last by constructor; case. -case: b9=>/=; last by constructor; case. -case: b10=>/=; last by constructor; case. -by constructor. -Qed. - -Lemma and11P : reflect [/\ b1, b2, b3, b4, b5, b6, b7, b8, b9, b10 & b11] - [&& b1, b2, b3, b4, b5, b6, b7, b8, b9, b10 & b11]. -Proof. -case: b1=>/=; last by constructor; case. -case: b2=>/=; last by constructor; case. -case: b3=>/=; last by constructor; case. -case: b4=>/=; last by constructor; case. -case: b5=>/=; last by constructor; case. -case: b6=>/=; last by constructor; case. -case: b7=>/=; last by constructor; case. -case: b8=>/=; last by constructor; case. -case: b9=>/=; last by constructor; case. -case: b10=>/=; last by constructor; case. -case: b11=>/=; last by constructor; case. -by constructor. -Qed. - -Lemma and12P : reflect [/\ b1, b2, b3, b4, b5, b6, b7, b8, b9, b10, b11 & b12] - [&& b1, b2, b3, b4, b5, b6, b7, b8, b9, b10, b11 & b12]. -Proof. -case: b1=>/=; last by constructor; case. -case: b2=>/=; last by constructor; case. -case: b3=>/=; last by constructor; case. -case: b4=>/=; last by constructor; case. -case: b5=>/=; last by constructor; case. -case: b6=>/=; last by constructor; case. -case: b7=>/=; last by constructor; case. -case: b8=>/=; last by constructor; case. -case: b9=>/=; last by constructor; case. -case: b10=>/=; last by constructor; case. -case: b11=>/=; last by constructor; case. -case: b12=>/=; last by constructor; case. -by constructor. -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 and11P {b1 b2 b3 b4 b5 b6 b7 b8 b9 b10 b11}. -Arguments and12P {b1 b2 b3 b4 b5 b6 b7 b8 b9 b10 b11 b12}. - -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}. - -Lemma iffPb (b1 b2 : bool) : reflect (b1 <-> b2) (b1 == b2). -Proof. -case: eqP=>[->|N]; constructor=>//. -case: b1 b2 N; case=>//= _. -- by case=>/(_ erefl). -by case=>_ /(_ erefl). -Qed. - -Lemma implyb_trans a b c : - a ==> b -> - b ==> c -> - a ==> c. -Proof. by case: a=>//=->. Qed. - -Lemma iffE (b1 b2 : bool) : b1 = b2 <-> (b1 <-> b2). -Proof. by split=>[->|] //; move/iffPb/eqP. Qed. - -(* subsets *) - -Lemma subsetC T (p q : mem_pred T) : - {subset p <= q} -> - {subset predC q <= predC p}. -Proof. by move=>H x; apply: contra (H x). Qed. - -Lemma subsetR T (p q : mem_pred T) : - {subset p <= predC q} -> - {subset q <= predC p}. -Proof. by move=>H x; apply: contraL (H x). Qed. - -Lemma subsetL T (p q : mem_pred T) : - {subset predC p <= q} -> - {subset predC q <= p}. -Proof. by move=>H x; apply: contraR (H x). Qed. - -Lemma subsetLR T (p q : mem_pred T) : - {subset predC p <= predC q} -> - {subset q <= p}. -Proof. by move=>H x; apply: contraLR (H x). Qed. - -Lemma subset_disj T (p q r : mem_pred T) : - {subset p <= q} -> - (forall x, x \in q -> x \in r -> false) -> - (forall x, x \in p -> x \in r -> false). -Proof. by move=>H D x /H/D. Qed. - -Lemma subset_disj2 T (p p1 q q1 : mem_pred T) : - {subset p1 <= p} -> - {subset q1 <= q} -> - (forall x, x \in p -> x \in q -> false) -> - (forall x, x \in p1 -> x \in q1 -> false). -Proof. -move=>H1 H2 D1; apply: subset_disj H1 _ => x H /H2. -by apply: D1 H. -Qed. - -(**************) -(* empty type *) -(**************) - -Lemma False_eqP : Equality.axiom (fun _ _ : False => true). -Proof. by case. Qed. - -HB.instance Definition _ := hasDecEq.Build False False_eqP. - -(*************) -(* 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. - -#[export] Hint Extern 0 (ocancel _ _) => - (apply: lft_inl_ocanc || apply: rgt_inr_ocanc) : core. - -(****************) -(* option types *) -(****************) - -(* Alternative mechanism for manipulating *) -(* properties of the form isSome X. *) -(* To use a lemma of the form isSome X *) -(* one tyoically has to explicitly put the conclusion *) -(* of that lemma onto the context, and then case on X. *) -(* If such lemma is restated using is_some X *) -(* then it suffices to case on the lemma's name. *) -(* This saves typing X explicitly, which *) -(* may be significant if X is large. *) - -Inductive is_some_spec A x : option A -> Prop := -| is_some_case v of x = Some v : is_some_spec x (Some v). - -Hint Resolve is_some_case : core. - -Notation is_some x := (is_some_spec x x). - -Lemma is_someP A (x : option A) : reflect (is_some x) (isSome x). -Proof. by case: x=>[a|]; constructor=>//; case. Qed. - -(* some simplifications *) -Lemma oapp_some A (x : option A) : oapp [eta Some] None x = x. -Proof. by case: x. Qed. - -Lemma obind_some A (x : option A) : obind [eta Some] x = x. -Proof. exact: oapp_some. Qed. - -(********) -(* nats *) -(********) - -Lemma gt0 m n : m < n -> 0 < n. -Proof. by case: n. Qed. - -Lemma neq0 m n : m < n -> n != 0. -Proof. by move/gt0; rewrite lt0n. Qed. - -Lemma neqSn n : n.+1 != n. -Proof. by elim: n. Qed. - -Lemma neqnS n : n != n.+1. -Proof. by elim: n. Qed. - -Lemma subn_eq0P m n : reflect (m - n = 0) (m <= n). -Proof. by rewrite -subn_eq0; apply/eqP. Qed. - -(**************************************) -(* Inhabited (non-empty) finite types *) -(**************************************) - -(* if a type is non-empty, one can pick an element *) - -Lemma inhabF {T : finType} : - 0 < #|T| -> ~ @predT T =1 xpred0. -Proof. by case/card_gt0P=>x _ /(_ x). Qed. - -Definition inhab0 {T : finType} (pf : 0 < #|T|) : T := - match pickP predT with - | Pick x _ => x - | Nopick qf => False_rect T (inhabF pf qf) - end. - -(* variant of nth that needs no seed value *) -Definition ith {T : finType} i (pf : i < #|T|) : T := - nth (inhab0 (leq_ltn_trans (leq0n i) pf)) (enum T) i. - -Arguments ith {T}. - -(* dually, variant of index that doesn't overflow *) -Definition indx {T : finType} (x : T) := index x (enum T). - -(* lemmas associated with ith/indx *) - -Lemma indx_card {T : finType} (i : T) : indx i < #|T|. -Proof. by rewrite cardT index_mem mem_enum. Qed. - -Lemma indx_ith {T : finType} i (pf : i < #|T|) : - indx (ith i pf) = i. -Proof. by rewrite /indx/ith index_uniq ?enum_uniq -?cardE. Qed. - -Lemma ith_indx {T : finType} (i : T) (pf : indx i < #|T|) : - ith (indx i) pf = i. -Proof. by rewrite /ith/indx nth_index // mem_enum. Qed. - -Lemma indx_inj {T} : injective (@indx T). -Proof. -rewrite /indx=>x1 x2. -have [] : x1 \in enum T /\ x2 \in enum T by rewrite !mem_enum. -elim: (enum T)=>[|x xs IH] //=; rewrite !inE !(eq_sym x). -case: (x1 =P x)=>[<-|] _ /=; first by case: (x2 =P x1). -by case: (x2 =P x)=>//= _ X1 X2 []; apply: IH X1 X2. -Qed. - -Lemma ith_inj {T : finType} i1 i2 (pf1 : i1 < #|T|) (pf2 : i2 < #|T|) : - ith i1 pf1 = ith i2 pf2 -> - i1 = i2. -Proof. -rewrite /ith; move: (inhab0 _) (inhab0 _)=>o1 o2 H. -rewrite cardE in pf1 pf2. -elim: (enum T) i1 i2 pf1 pf2 o1 o2 H (enum_uniq T)=>[|x xs IH] //=. -case=>[|i1][|i2] //= pf1 pf2 o1 o2. -- by move=>->; rewrite mem_nth. -- by move=><-; rewrite mem_nth. -by move=>H /andP [_ U]; rewrite (IH _ _ pf1 pf2 o1 o2). -Qed. - -Lemma indx_injE {T : finType} s i (pf : i < #|T|) : - (s == ith i pf) = (indx s == i). -Proof. -apply/eqP/eqP=>[->|E]; first by rewrite indx_ith. -by rewrite -E in pf *; rewrite ith_indx. -Qed. - -Lemma map_indx {T : finType} : map indx (enum T) = iota 0 #|T|. -Proof. -rewrite cardE /indx. -elim: (enum T) (enum_uniq T)=>[|x xs IH] //. -set f := index^~(x :: xs)=>/= /andP [H1 H2]. -rewrite {1}/f /= eqxx; congr (0 :: _). -case: (eq_in_map f (fun x=>(index x xs).+1) xs)=>E _. -rewrite E; last first. -- by move=>z R; rewrite /f /=; case: (x =P z) R H1=>//= ->->. -by rewrite -add1n iotaDl -IH // -map_comp. -Qed. - -Lemma take_enum {T : finType} x i : - x \in take i (enum T) = (indx x < i). -Proof. -pose f x := indx x. -rewrite -(mem_map indx_inj) map_take map_indx take_iota. -case: (leqP i #|T|)=>H; rewrite mem_iota /=; first by rewrite add0n. -rewrite add0n cardE index_mem mem_enum inE /=; apply: sym_eq. -apply: (@ltn_trans #|T|) H. -by rewrite cardE index_mem mem_enum. -Qed. - -Lemma drop_enum {T : finType} x i : - x \in drop i (enum T) = (i <= indx x). -Proof. -pose f x := index x (enum T). -rewrite -(mem_map indx_inj) map_drop map_indx drop_iota mem_iota add0n. -case H : (i <= indx x)=>//=; rewrite subnKC ?indx_card //. -by apply/(leq_trans H)/ltnW/indx_card. -Qed. - -Lemma take_enum_filter {T : finType} k : - filter (preim indx [pred x | x < k]) (enum T) = - take k (enum T). -Proof. -apply: (inj_map indx_inj). -rewrite map_take map_indx -filter_map map_indx. -apply: (sorted_eq leq_trans anti_leq). -- by apply/(sorted_filter leq_trans)/iota_sorted. -- by apply/take_sorted/iota_sorted. -apply: uniq_perm. -- by rewrite filter_uniq // iota_uniq. -- by rewrite take_uniq // iota_uniq. -move=>x; rewrite mem_filter take_iota cardE. -case: (leqP k (size _))=>H1; rewrite !mem_iota !add0n /=. -- by case: (ltnP x k)=>H2 //=; apply: leq_trans H1. -case: (ltnP x (size _))=>H2 //=; last by rewrite andbF. -by rewrite andbT; apply: ltn_trans H1. -Qed. - -Lemma drop_enum_filter {T : finType} k : - filter (preim indx [pred x | x >= k]) (enum T) = - drop k (enum T). -Proof. -apply: (inj_map indx_inj). -rewrite map_drop map_indx -filter_map map_indx. -apply: (sorted_eq leq_trans anti_leq). -- by apply/(sorted_filter leq_trans)/iota_sorted. -- by apply/drop_sorted/iota_sorted. -apply: uniq_perm. -- by rewrite filter_uniq // iota_uniq. -- by rewrite drop_uniq // iota_uniq. -move=>x; rewrite mem_filter drop_iota cardE !add0n inE. -case: (leqP k (size (enum T)))=>H1; rewrite !mem_iota add0n. -- by rewrite subnKC. -case: subn_eq0P (ltnW H1)=>// -> _; rewrite addn0 leq0n /=. -by case: (leqP k x)=>//= K1; case: ltngtP (leq_trans H1 K1). -Qed. - -Lemma enum_split {T : finType} k : - enum T = take (indx k) (enum T) ++ k :: drop (indx k).+1 (enum T). -Proof. -rewrite -{2}(@nth_index T k k (enum T)) ?mem_enum //. -by rewrite -drop_nth ?index_mem ?mem_enum // cat_take_drop. -Qed. - -Lemma takeord {I : finType} T k x (f : {ffun I -> T}) : - take (indx k) (fgraph [ffun y => [eta f with k |-> x] y]) = - take (indx k) (fgraph f). -Proof. -set f' := (finfun _). -suff E: {in take (indx k) (enum I), f =1 f'}. -- by rewrite !fgraph_codom /= !codomE -2!map_take; move/eq_in_map: E. -move: (enum_uniq I); rewrite {1}(enum_split k) cat_uniq /= =>H4. -move=>y H5; rewrite /f' /= !ffunE /=; case: eqP H5 H4=>// -> ->. -by rewrite andbF. -Qed. - -Lemma dropord {I : finType} T k x (f : {ffun I -> T}) : - drop (indx k).+1 (fgraph [ffun y => [eta f with k |->x] y]) = - drop (indx k).+1 (fgraph f). -Proof. -set f' := (finfun _). -suff E: {in drop (indx k).+1 (enum I), f =1 f'}. -- by rewrite !fgraph_codom /= !codomE -2!map_drop; move/eq_in_map: E. -move: (enum_uniq I); rewrite {1}(enum_split k) cat_uniq /= => H4. -move=>y H5; rewrite /f' /= !ffunE /=; case: eqP H5 H4=>// -> ->. -by rewrite !andbF. -Qed. - -Lemma size_fgraph {I : finType} T1 T2 - (r1 : {ffun I -> T1}) (r2 : {ffun I -> T2}) : - size (fgraph r1) = size (fgraph r2). -Proof. by rewrite !fgraph_codom /= !codomE !size_map. Qed. - -Lemma fgraphE {I : finType} T (r1 r2 : {ffun I -> T}) : - fgraph r1 = fgraph r2 -> - r1 = r2. -Proof. -move=> eq_r12; apply/ffunP=> x. -by rewrite -[x]enum_rankK -!tnth_fgraph eq_r12. -Qed. - -(* building inhabited finite types *) - -Lemma inhabits (T : finType) (t : T) : 0 < #|T|. -Proof. by apply/card_gt0P; exists t. Qed. - -Lemma inhabits_irr (T : finType) (t1 t2 : T) : - inhabits t1 = inhabits t2. -Proof. by apply: bool_irrelevance. Qed. - -Definition inhabited_axiom (T : finType) := 0 < #|T|. - -HB.mixin Record isInhabited T of Finite T := { - card_inhab : inhabited_axiom T}. - -#[short(type="ifinType")] -HB.structure Definition Inhabited := {T of isInhabited T & Finite T}. - -HB.instance Definition _ := isInhabited.Build unit (inhabits tt). -HB.instance Definition _ := isInhabited.Build bool (inhabits false). -HB.instance Definition _ n := isInhabited.Build 'I_n.+1 (inhabits ord0). -HB.instance Definition _ (T : finType) := - isInhabited.Build (option T) (inhabits None). - -Definition inhab {I : ifinType} : I := inhab0 card_inhab. - -(*************************************) -(* A copy of booleans with mnemonics *) -(* LL and RR for working with sides *) -(*************************************) - -Inductive Side := LL | RR. -Definition Side_eq x y := - match x, y with LL, LL => true | RR, RR => true | _, _ => false end. - -Coercion nat_of_side x := if x is LL then 0 else 1. -Definition side_of_nat x := if odd x then RR else LL. -Lemma ssrcanc : ssrfun.cancel nat_of_side side_of_nat. Proof. by case. Qed. -HB.instance Definition _ : isCountable Side := CanIsCountable ssrcanc. - -Lemma Side_enumP : Finite.axiom [:: LL; RR]. Proof. by case. Qed. -HB.instance Definition _ : isFinite Side := isFinite.Build Side Side_enumP. -HB.instance Definition _ := isInhabited.Build Side (inhabits LL). - -Definition flip x := if x is LL then RR else LL. - -(*************) -(* Sequences *) -(*************) - -(* folds *) -(* TODO upstream to mathcomp *) - -Lemma map_foldr {T1 T2} (f : T1 -> T2) xs : - map f xs = foldr (fun x ys => f x :: ys) [::] xs. -Proof. by []. Qed. - -Lemma fusion_foldr {T R Q} (g : R -> Q) f0 f1 z0 z1 (xs : seq T) : - (forall x y, g (f0 x y) = f1 x (g y)) -> - g z0 = z1 -> - g (foldr f0 z0 xs) = foldr f1 z1 xs. -Proof. by move=>Hf Hz; elim: xs=>//= x xs <-. Qed. - -Lemma fusion_foldl {T R Q} (g : R -> Q) f0 f1 z0 z1 (xs : seq T) : - (forall x y, g (f0 x y) = f1 (g x) y) -> - g z0 = z1 -> - g (foldl f0 z0 xs) = foldl f1 z1 xs. -Proof. -move=>Hf Hz; elim: xs z0 z1 Hz =>//= x xs IH z0 z1 Hz. -by apply: IH; rewrite Hf Hz. -Qed. - -Lemma foldl_foldr {T R} (f : R -> T -> R) z xs : - foldl f z xs = foldr (fun b g x => g (f x b)) id xs z. -Proof. by elim: xs z=>/=. Qed. - -Lemma foldr_foldl {T R} (f : T -> R -> R) z xs : - foldr f z xs = foldl (fun g b x => g (f b x)) id xs z. -Proof. -elim/last_ind: xs z=>//= xs x IH z. -by rewrite foldl_rcons -IH foldr_rcons. -Qed. - -(* pmap *) -(* TODO upstream to mathcomp *) -Lemma pmap_pcomp {S T U} (f : T -> option U) (g : S -> option T) s : - pmap (pcomp f g) s = pmap f (pmap g s). -Proof. by elim: s=>//= x s ->; rewrite /pcomp; case: (g x). Qed. - -(* sequence prefixes *) - -(* Two helper concepts for searching in sequences: *) -(* - onth: like nth, but returns None when the element is not found *) -(* - Prefix: a prefix relation on sequences, used for growing *) -(* interpretation contexts *) - -Fixpoint onth A (s : seq A) n : option A := - if s is x::sx then if n is nx.+1 then onth sx nx else Some x else None. - -Definition Prefix A (s1 s2 : seq A) : Prop := - forall n x, onth s1 n = some x -> onth s2 n = some x. - -(* Lemmas *) - -Section SeqPrefix. -Variable A : Type. -Implicit Type s : seq A. - -Variant onth_spec s n : bool -> Type := -| onth_none : onth s n = None -> onth_spec s n false -| onth_some v : onth s n = Some v -> onth_spec s n true. - -Lemma onth_sizeP s n : onth_spec s n (n < size s). -Proof. -elim: s n=>/= [|a s IH] n; first by rewrite ltn0; constructor. -case: n=>[|n] /=; first by apply: (@onth_some _ _ a). -rewrite ltnS; case: (IH n)=>[|v] H. -- by constructor. -by apply: (@onth_some _ _ v). -Qed. - -Lemma size_onth s n : n < size s -> exists x, onth s n = Some x. -Proof. -by case: onth_sizeP=>// v H _; exists v. -Qed. - -Lemma onth_size s n x : onth s n = Some x -> n < size s. -Proof. -by case: onth_sizeP=>//->. -Qed. - -Lemma size_onthPn s n : reflect (onth s n = None) (size s <= n). -Proof. -by rewrite leqNgt; apply: (iffP idP); case: onth_sizeP=>//= v ->. -Qed. - -Lemma onth_sizeN s : onth s (size s) = None. -Proof. by apply/size_onthPn. Qed. - -Lemma nth_onth x0 n s : nth x0 s n = odflt x0 (onth s n). -Proof. -elim: s n=>/= [|a s IH] n /=; first by apply: nth_nil. -by case: n. -Qed. - -Lemma onth_nth x0 n s : - n < size s -> - onth s n = Some (nth x0 s n). -Proof. -elim: s n=>//= a s IH n. -by rewrite ltnS; case: n. -Qed. - -Lemma Prefix_refl s : Prefix s s. -Proof. by move=>n x <-. Qed. - -Lemma Prefix_trans s2 s1 s3 : - Prefix s1 s2 -> Prefix s2 s3 -> Prefix s1 s3. -Proof. by move=>H1 H2 n x E; apply: H2; apply: H1. Qed. - -Lemma Prefix_cons x s1 s2 : Prefix (x :: s1) (x :: s2) <-> Prefix s1 s2. -Proof. by split=>E n; [apply: (E n.+1) | case: n]. Qed. - -Lemma Prefix_cons' x y s1 s2 : - Prefix (x :: s1) (y :: s2) -> x = y /\ Prefix s1 s2. -Proof. by move=>H; case: (H 0 x (erefl _)) (H)=>-> /Prefix_cons. Qed. - -Lemma Prefix_rcons x s : Prefix s (rcons s x). -Proof. by elim: s=>//= y ys IH; apply/Prefix_cons; apply: IH. Qed. - -Lemma Prefix_cat s1 s2 : Prefix s1 (s1 ++ s2). -Proof. -elim: s2 s1=>[|x xs IH] s1; first by rewrite cats0. -rewrite -cat_rcons; apply: Prefix_trans (IH _). -by apply: Prefix_rcons. -Qed. - -Lemma Prefix_size s1 s2 : Prefix s1 s2 -> size s1 <= size s2. -Proof. -elim: s1 s2=>[//|a s1 IH] [|b s2] H; first by move: (H 0 a (erefl _)). -by rewrite ltnS; apply: (IH _ (proj2 (Prefix_cons' H))). -Qed. - -Lemma Prefix_onth s t x : - x < size s -> - Prefix s t -> onth s x = onth t x. -Proof. -elim:s t x =>[//|a s IH] [|b t] x H1 H2; first by move: (H2 0 a (erefl _)). -by case/Prefix_cons': H2=><- H2; case: x H1=>[|n] //= H1; apply: IH. -Qed. - -Lemma PrefixE s1 s2 : Prefix s1 s2 <-> exists s3, s2 = s1 ++ s3. -Proof. -split; last by case=>s3 ->; apply: Prefix_cat. -elim: s1 s2=>[|x xs IH] s2; first by exists s2. -case: s2=>[/(_ 0 x erefl)//|y ys /Prefix_cons' [?]]. -by subst y=>/IH [s3 ->]; exists s3. -Qed. - -End SeqPrefix. - -#[export] Hint Resolve Prefix_refl : core. - -(* when A : eqType *) - -Section SeqPrefixEq. -Variable A : eqType. -Implicit Type s : seq A. - -Lemma onth_mem s n x : - onth s n = Some x -> - x \in s. -Proof. -by elim: s n=>//= a s IH [[->]|n /IH]; rewrite inE ?eq_refl // orbC =>->. -Qed. - -Lemma onth_index (s : seq A) x : - onth s (index x s) = - if x \in s then Some x else None. -Proof. -by elim: s=>//=h s IH; rewrite inE eq_sym; case: eqP=>//= ->. -Qed. - -Lemma PrefixP (s1 s2 : seq A) : - reflect (Prefix s1 s2) (prefix s1 s2). -Proof. -apply/(equivP (prefixP (s1:=s1) (s2:=s2))). -by apply: iff_sym; exact: PrefixE. -Qed. - -End SeqPrefixEq. - -(******************************) -(* Some commuting conversions *) -(******************************) - -Lemma fun_op A B C (b : option A) (vS : A -> B) (vN : B) (f : B -> C) : - f (if b is Some v then vS v else vN) = - if b is Some v then f (vS v) else f vN. -Proof. by case: b=>//. Qed. - -Lemma op_if A B (b : bool) (vS vN : option A) (vS1 : A -> B) (vN1 : B) : - (if (if b then vS else vN) is Some v then vS1 v else vN1) = - if b then if vS is Some v then vS1 v else vN1 - else if vN is Some v then vS1 v else vN1. -Proof. by case: b. Qed. - -Lemma if_triv b : (if b then true else false) = b. -Proof. by case: b. Qed. - -(*************************************************************) -(* quick ways to extract projections and transitivity proofs *) -(* out of iterated inequalities *) -(*************************************************************) - -Lemma ltn13 a b c : a < b < c -> a < c. -Proof. by case/andP; apply: ltn_trans. Qed. - -Lemma ltn12 a b c : a < b < c -> a < b. -Proof. by case/andP. Qed. - -Lemma ltn23 a b c : a < b < c -> b < c. -Proof. by case/andP. Qed. - -Lemma leq13 a b c : a <= b <= c -> a <= c. -Proof. by case/andP; apply: leq_trans. Qed. - -Lemma leq12 a b c : a <= b <= c -> a <= b. -Proof. by case/andP. Qed. - -Lemma leq23 a b c : a <= b <= c -> b <= c. -Proof. by case/andP. Qed. - -Lemma lqt13 a b c : a <= b < c -> a < c. -Proof. by case/andP; apply: leq_ltn_trans. Qed. - -Lemma lqt12 a b c : a <= b < c -> a <= b. -Proof. by case/andP. Qed. - -Lemma lqt23 a b c : a <= b < c -> b < c. -Proof. by case/andP. Qed. - -(********************) -(* Finite functions *) -(********************) - -Section FinFun. -Variables (T : finType) (Us : T -> Type). -Implicit Type f : {dffun forall t, Us t}. - -(* Explicit name for finite function application. *) -(* Will be used to hang canonical projections onto. *) -(* The function/argument order is reversed to facilitate rewriting. *) -Definition sel tg f : Us tg := f tg. - -Lemma ffinP f1 f2 : (forall t, sel t f1 = sel t f2) <-> f1 = f2. -Proof. by rewrite ffunP. Qed. - -(* beta equality *) -Lemma sel_fin t (f : forall t, Us t) : sel t (finfun f) = f t. -Proof. by rewrite /sel ffunE. Qed. - -(* eta equality *) -Lemma fin_eta f : f = finfun (sel^~ f). -Proof. by apply/ffinP=>t; rewrite sel_fin. Qed. - -(* function *) -Definition splice tg f (v : Us tg) : {dffun _} := - finfun (fun x => - if decP (x =P tg) is left pf then cast Us pf v - else sel x f). - -Lemma sel_splice t f x (v : Us x) : - sel t (splice f v) = - if decP (t =P x) is left pf then cast Us pf v - else sel t f. -Proof. by rewrite sel_fin. Qed. - -Lemma sel_spliceE t f v : sel t (splice f v) = v. -Proof. by rewrite sel_fin; case: eqP=>//= pf; rewrite eqd. Qed. - -Lemma sel_spliceN t x f (w : Us x) : - t <> x -> sel t (splice f w) = sel t f. -Proof. by move=>N; rewrite sel_fin; case: eqP. Qed. - -Lemma splice_eta t f : splice f (sel t f) = f. -Proof. -apply/ffinP=>x; rewrite sel_fin; case: eqP=>//=. -by move/[dup]=>-> pf; rewrite eqd. -Qed. - -End FinFun. - -Arguments sel {T Us} tg f. -Arguments splice {T Us tg} f v. - -(* notation for building finfuns *) -(* DEVCOMMENT *) -(* copied from finfun to fix some bad spacing in formatting *) -(* /DEVCOMMENT *) - -Notation "[ 'ffun' x : aT => E ]" := (finfun (fun x : aT => E)) - (at level 0, x name, format "[ 'ffun' x : aT => E ]") : fun_scope. - -Notation "[ 'ffun' x => E ]" := (@finfun _ (fun=> _) (fun x => E)) - (at level 0, x name, format "[ 'ffun' x => E ]") : fun_scope. - -Notation "[ 'ffun' => E ]" := [ffun _ => E] - (at level 0, format "[ 'ffun' => E ]") : fun_scope. - -Section IteratedNotation. -Variables (T : finType) (Us : T -> Type). - -Variant dfun_delta : Type := DFunDelta t of Us t. - -(* for iteration that starts with function ends with function *) -Definition dapp_fdelta df (f : forall t, Us t) z := - let: DFunDelta t v := df in - if decP (z =P t) is left pf then cast Us pf v - else f z. - -(* for iteration that starts with finfun ends with function *) -Definition splice' df (f : {ffun forall t, Us t}) z := - dapp_fdelta df f z. - -End IteratedNotation. - -Delimit Scope fun_delta_scope with FUN_DELTA. -Arguments dapp_fdelta {T Us} df%_FUN_DELTA f. - -Notation "y \\ x" := (@DFunDelta _ _ x y) (at level 1). - -(* notation for simultaneous update of f with d1,..,dn *) -(* rewrite by sel_fin peels all layers *) -Notation "[ 'ext' f 'with' d1 , .. , dn ]" := - (finfun ( - dapp_fdelta d1%_FUN_DELTA .. (dapp_fdelta dn%_FUN_DELTA f) ..)) - (at level 0, format - "'[hv' [ '[' 'ext' '/ ' f ']' '/' 'with' '[' d1 , '/' .. , '/' dn ']' ] ']'" - ) : fun_scope. - -(* notation for iterated update of f with d1, then d2, ... *) -(* rewrite by sel_fin peels top layer only *) -Notation "[ 'splice' F 'with' d1 , .. , dn ]" := - (finfun (splice' - d1%_FUN_DELTA .. (finfun (splice' dn%_FUN_DELTA (finfun F))) ..)) - (at level 0, format - "'[hv' [ '[' 'splice' '/ ' F ']' '/' 'with' '[' d1 , '/' .. , '/' dn ']' ] ']'" - ) : fun_scope. - -Section TestingNotation. -Variables (T : finType) (Us : T -> Type). -Variables (f : {dffun forall t, Us t}) (t1 t2 : T) (v1 : Us t1) (v2 : Us t2). - -(* we have three different options in displaying splices *) -(* splice, [splice], and [ext] *) -Lemma test : - [/\ sel t2 [splice f with v1 \\ t1, v2 \\ t2] = v2, - sel t2 (splice (splice f v1) v2) = v2, - sel t2 [ext f with v1 \\ t1, v2 \\ t2] = v2 & -(* and we can use underscores to elide some info *) - sel t2 [ext f with v1 \\ _, v2 \\ _] = v2]. -Abort. -End TestingNotation. - -(* mapping over simply-typed finite functions *) - -Section FinFunMap. -Variables (T : finType) (A B : Type). -Implicit Types (f : A -> B) (x : {ffun T -> A}). - -Definition fmap f x := [ffun tg => f (sel tg x)]. - -Lemma sel_fmap f x tg : sel tg (fmap f x) = f (sel tg x). -Proof. exact: sel_fin. Qed. - -Lemma fmap_splice f x tg (v : A) : - fmap f (splice (tg:=tg) x v) = splice (tg:=tg) (fmap f x) (f v). -Proof. -apply/ffinP=>t; rewrite !sel_fin; case: eqP=>//= ?; subst t. -by rewrite !eqc. -Qed. - -End FinFunMap. - -(* surgery on tuples and finfuns *) - -Section OnthCodom. -Variable A : Type. - -Lemma onth_tnth {n} (s : n.-tuple A) (i : 'I_n) : - onth s i = Some (tnth s i). -Proof. -elim: n s i =>[|n IH] s i; first by case: i. -case/tupleP: s=>/=x s; case: (unliftP ord0 i)=>[j|]-> /=. -- by rewrite tnthS ?add0n. -by rewrite tnth0. -Qed. - -Lemma onth_codom {n} (i : 'I_n) (f: {ffun 'I_n -> A}) : - onth (fgraph f) i = Some (f i). -Proof. -pose i' := cast_ord (esym (card_ord n)) i. -move: (@tnth_fgraph _ _ f i'); rewrite (enum_val_ord) {2}/i' cast_ordKV=><-. -by rewrite (onth_tnth (fgraph f) i'). -Qed. - -End OnthCodom. - -(* ffun and permutation *) -Section PermFfun. -Variables (I : finType) (A : Type). - -Definition pffun (p : {perm I}) (f : {ffun I -> A}) := - [ffun i => f (p i)]. - -Lemma pffunE1 (f : {ffun I -> A}) : pffun 1%g f = f. -Proof. by apply/ffunP=>i; rewrite !ffunE permE. Qed. - -Lemma pffunEM (p p' : {perm I}) (f : {ffun I -> A}) : - pffun (p * p') f = pffun p (pffun p' f). -Proof. by apply/ffunP => i; rewrite !ffunE permM. Qed. - -End PermFfun. - -(* Finite sets *) - -Lemma enum_T (I : finType) : enum setT = enum I. -Proof. by rewrite enum_setT enumT. Qed. - -Lemma enum_0 (I : finType) (s : {set I}) : - s =i set0 -> - enum s = [::]. -Proof. by move/eq_enum=>->; rewrite enum_set0. Qed. - -Lemma setTE (I : finType) (p : pred I) : - p =1 xpredT -> - [set x | p x] = setT :> {set I}. -Proof. by move=>H; apply/setP=>x; rewrite !inE H. Qed. - -Lemma set0E (I : finType) (p : pred I) : - p =1 xpred0 -> - [set x | p x] = set0 :> {set I}. -Proof. by move=>H; apply/setP=>x; rewrite !inE H. Qed. - -(* streamlining for ordinals *) - -Lemma set_ord0N m (p : pred nat) : - (forall x, x < m -> ~~ p x) -> - [set x | p (\val x)] = set0 :> {set 'I_m}. -Proof. by move=>H; apply: set0E; case=>z pf; rewrite (negbTE (H z pf)). Qed. - -Lemma set_ord0 m (p : pred nat) : - (forall x, p x -> m <= x) -> - [set x | p (\val x)] = set0 :> {set 'I_m}. -Proof. -move=>H; apply: set_ord0N=>x; rewrite ltnNge. -by apply/contra/H. -Qed. - -Lemma set_ordT m (p : pred nat) : - (forall x, x < m -> p x) -> - [set x | p (\val x)] = setT :> {set 'I_m}. -Proof. by move=>H; apply: setTE; case=>z pf; rewrite H. Qed. - -Lemma set_ordTN m (p : pred nat) : - (forall x, ~~ p x -> m <= x) -> - [set x | p (\val x)] = setT :> {set 'I_m}. -Proof. -move=>H; apply/set_ordT=>x; rewrite ltnNge. -by apply/contraR/H. -Qed. - -(* Tagging *) - -Notation Tag := (@existT _ _). - -Lemma Tag_inj T Us (t1 t2 : T) i1 i2 : - Tag t1 i1 = Tag t2 i2 -> - t1 = t2 /\ jmeq Us i1 i2. -Proof. by case=>?; subst t2=>/inj_pair2 ->. Qed. -Arguments Tag_inj {T Us t1 t2 i1 i2}. - -(* tagged union of equality types is equality type *) - -Section TaggedEq. -Variables (T : eqType) (Us : T -> eqType). - -Definition tag_eq : sigT Us -> sigT Us -> bool := - fun '(Tag tx opx) '(Tag ty opy) => - if decP (tx =P ty) is left pf then opx == cast Us pf opy - else false. - -Lemma tag_eqP : Equality.axiom tag_eq. -Proof. -case=>tx opx [ty opy] /=; case: (tx =P ty)=>pf; last first. -- by constructor; case=>/pf. -subst ty; rewrite /= eqc; case: eqP=>pf; constructor; -by [rewrite pf|case=>/inj_pair2/pf]. -Qed. - -HB.instance Definition _ := hasDecEq.Build (sigT Us) tag_eqP. -End TaggedEq. - diff --git a/core/rtc.v b/core/rtc.v deleted file mode 100644 index 72e7c72..0000000 --- a/core/rtc.v +++ /dev/null @@ -1,273 +0,0 @@ -(* -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 pcm Require Import options axioms pred prelude finmap seqext. - -(* 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 := Finite.clone (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 H; apply: Rsub. -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/seqext.v b/core/seqext.v deleted file mode 100644 index 7033112..0000000 --- a/core/seqext.v +++ /dev/null @@ -1,2138 +0,0 @@ -(* -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. -*) - -From Coq Require Import ssreflect ssrbool ssrfun. -From mathcomp Require Import ssrnat seq eqtype path choice fintype bigop perm. -From pcm Require Import options prelude pred. - -(*********************) -(* Extensions to seq *) -(*********************) - -(* TODO upstream to mathcomp *) - -Lemma head_rcons {A} (s : seq A) (x y : A) : - head x (rcons s y) = head y s. -Proof. by rewrite headI. Qed. - -Lemma rcons_nseq {A} n (x : A) : - rcons (nseq n x) x = nseq n.+1 x. -Proof. by elim: n=>//=n ->. Qed. - -Lemma behead_rcons {A} (xs : seq A) x : - 0 < size xs -> - behead (rcons xs x) = rcons (behead xs) x. -Proof. by case: xs. Qed. - -Lemma nilp_hasPn {A} (s : seq A) : nilp s = ~~ has predT s. -Proof. by case: s. Qed. - -Lemma filter_predIC {A} (s : seq A) p1 p2 : - filter (predI p1 p2) s = filter (predI p2 p1) s. -Proof. by apply: eq_filter => z /=; rewrite andbC. Qed. - -Lemma filter_swap {A} (s : seq A) p1 p2 : - filter p1 (filter p2 s) = filter p2 (filter p1 s). -Proof. by rewrite -!filter_predI filter_predIC. Qed. - -(* TODO contribute to mathcomp? *) -Lemma map_nilp {A B} (f : A -> B) (s : seq A) : - nilp (map f s) = nilp s. -Proof. by rewrite /nilp; case: s. Qed. - -Lemma filter_nilp {A} (p : pred A) (s : seq A) : - nilp (filter p s) = ~~ has p s. -Proof. by rewrite /nilp size_filter -leqn0 leqNgt has_count. Qed. - -Lemma head_map {P Q} (f : P -> Q) z (s : seq P) : - f (head z s) = head (f z) (map f s). -Proof. by case: s. Qed. - -Lemma zip_map2 {P Q R S} (f : P -> R) (g : Q -> S) (s1 : seq P) (s2 : seq Q) : - zip (map f s1) (map g s2) = - map (fun '(x1,x2) => (f x1,g x2)) (zip s1 s2). -Proof. -elim: s1 s2=>/= [|x1 s1 IH] [|x2 s2] //=. -by congr cons. -Qed. - -Lemma zip_mapl {P Q R} (f : P -> R) (s1 : seq P) (s2 : seq Q) : - zip (map f s1) s2 = - map (fun '(x1,x2) => (f x1,x2)) (zip s1 s2). -Proof. by rewrite -{1}(map_id s2) zip_map2. Qed. - -Lemma zip_mapr {P Q S} (g : Q -> S) (s1 : seq P) (s2 : seq Q) : - zip s1 (map g s2) = - map (fun '(x1,x2) => (x1,g x2)) (zip s1 s2). -Proof. by rewrite -{1}(map_id s1) zip_map2. Qed. - -Lemma drop_take_id {A} x (s : seq A) : drop x (take x s) = [::]. -Proof. by rewrite -{2}(add0n x) -take_drop take0. Qed. - -Lemma drop_take_mask {A} (s : seq A) x y : - drop x (take y s) = mask (nseq x false ++ nseq (y-x) true) s. -Proof. -case: (ltnP x (size s))=>Hx; last first. -- rewrite drop_oversize; last by rewrite size_take_min geq_min Hx orbT. - rewrite -{1}(subnKC Hx) nseqD -catA -{3}(cats0 s) mask_cat; last by rewrite size_nseq. - by rewrite mask0 mask_false. -have Hx': size (nseq x false) = size (take x s). -- by rewrite size_nseq size_take_min; symmetry; apply/minn_idPl/ltnW. -rewrite -{2}(cat_take_drop x s) mask_cat // mask_false /= -takeEmask take_drop. -case: (leqP x y)=>[Hxy|/ltnW Hxy]; first by rewrite subnK. -move: (Hxy); rewrite -subn_eq0=>/eqP->; rewrite add0n drop_take_id. -by rewrite drop_oversize // size_take_min geq_min Hxy. -Qed. - -Section LemmasEq. -Variables A : eqType. -Implicit Type xs : seq A. - -(* With A : Type, we have the In_split lemma. *) -(* With A : eqType, the lemma can be strenghtened to *) -(* not only return the split of xs, but the split of xs *) -(* that uses the first occurrence of x is xs *) -Lemma in_split xs x : - x \in xs -> - exists xs1 xs2, - xs = xs1 ++ x :: xs2 /\ x \notin xs1. -Proof. -rewrite -has_pred1; case/split_find=>_ s1 s2 /eqP ->. -by rewrite has_pred1=>H; exists s1, s2; rewrite -cats1 -catA. -Qed. - -(* a weaker form of in_split *) -Lemma mem_split (x : A) s : - x \in s -> - exists s1 s2, s = s1 ++ [:: x] ++ s2. -Proof. by case/in_split=>s1 [s2][H _]; exists s1, s2. Qed. - -Lemma mem_split_uniq (x : A) s : - x \in s -> uniq s -> - exists s1 s2, [/\ s = s1 ++ [:: x] ++ s2, - uniq (s1 ++ s2) & - x \notin s1 ++ s2]. -Proof. -move=>/[swap] Hu /mem_split [s1 [s2 H]]. -exists s1, s2; move: Hu. -by rewrite H uniq_catCA cons_uniq; case/andP. -Qed. - -Lemma perm_cons2 (x y : A) s : - perm_eq [:: x, y & s] [:: y, x & s]. -Proof. -by rewrite (_ : [::x,y & s] = [::x] ++ [::y] ++ s) // - (_ : [::y,x & s] = [::y] ++ [::x] ++ s) // perm_catCA. -Qed. - -Lemma all_notin (p : pred A) xs y : - all p xs -> - ~~ p y -> - y \notin xs. -Proof. by move/allP=>Ha; apply/contra/Ha. Qed. - -Lemma in_consE P x xs : - {in x :: xs, forall z, P z} <-> - P x /\ {in xs, forall z, P z}. -Proof. -split=>[H|[H1 H2]]; last by move=>z /orP [/eqP ->|/H2]. -split=>[|z Z]; first by apply: H; rewrite inE eqxx. -by apply: H; rewrite inE Z orbT. -Qed. - -Lemma subset_all a (s1 s2 : seq A) : - {subset s1 <= s2} -> - all a s2 -> - all a s1. -Proof. by move=>Hs /allP Ha1; apply/allP=>s /Hs /Ha1. Qed. - -Lemma subset_nilR xs : - {subset xs <= [::]} -> - xs = [::]. -Proof. by case: xs=>// x xs /(_ x); rewrite inE eqxx=>/(_ erefl). Qed. - -Lemma subset_nil xs ys : - {subset xs <= ys} -> - ys = [::] -> - xs = [::]. -Proof. by move=>X E; move: E X=>->; apply: subset_nilR. Qed. - -Lemma subset_consL x (s1 s2 : seq A) : - {subset x :: s1 <= s2} <-> - x \in s2 /\ {subset s1 <= s2}. -Proof. -split=>[S|[X S]]. -- by split=>[|z Z]; apply: S; rewrite inE ?eqxx ?Z ?orbT. -by move=>z; rewrite inE; case/orP=>[/eqP ->|/S//]. -Qed. - -Lemma subset_consLI x (s1 s2 : seq A) : - x \in s2 -> - {subset s1 <= s2} -> - {subset x :: s1 <= s2}. -Proof. by move=>H1 H2; rewrite subset_consL. Qed. - -Lemma subset_consR x (s : seq A) : - {subset s <= x :: s}. -Proof. by move=>z E; rewrite inE E orbT. Qed. - -Lemma subset_consLR x (s1 s2 : seq A) : - {subset s1 <= s2} -> - {subset x :: s1 <= x :: s2}. -Proof. -move=>X z; rewrite !inE; case/orP=>[|/X] -> //. -by rewrite orbT. -Qed. - -Lemma all_mem xs ys : - reflect {subset ys <= xs} (all [mem xs] ys). -Proof. -by case: allP=>H; constructor; [move=>x /H | move=>X; apply: H=>x /X]. -Qed. - -Lemma all_predC_sym xs ys : - all [predC xs] ys = all [predC ys] xs. -Proof. by rewrite all_predC has_sym -all_predC. Qed. - -Lemma nilp_filter (p : pred A) s : - reflect {in s, forall z, ~~ p z} (nilp (filter p s)). -Proof. -case E : (nilp _); constructor. -- move: E; rewrite nilp_hasPn=>/hasPn H x Kx; apply/negP=>Px. - by move: (H x); rewrite mem_filter Px=>/(_ Kx). -move=>X; move/negP: E; elim. -rewrite nilp_hasPn; apply/negP=>/hasP [x]. -rewrite mem_filter=>/andP [Px Kx] _. -by move: (X x Kx); rewrite Px. -Qed. - -Lemma index_rcons a x 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 x xs : - x \notin xs <-> index x xs = size xs. -Proof. -split; first by exact: memNindex. -by move=>E; rewrite -index_mem E ltnn. -Qed. - -Lemma index_sizeE x xs : - reflect (index x xs = size xs) (x \notin xs). -Proof. by apply/(equivP idP)/index_memN. Qed. - -Lemma size0nilP xs : - reflect (xs = [::]) (size xs == 0). -Proof. -case: eqP=>X; constructor; first by move/size0nil: X. -by move=>N; rewrite N in X. -Qed. - -Lemma has_nilP xs : - reflect (has predT xs) (xs != [::]). -Proof. by case: xs=>[|x xs]; constructor. Qed. - -Lemma map_nilP {B : eqType} (f : B -> A) s : - reflect (exists k, k \in map f s) (map f s != [::]). -Proof. -case: has_nilP=>X; constructor. -- by case/hasP: X=>x; exists x. -by case=>k K; elim: X; apply/hasP; exists k. -Qed. - -Lemma filter_nilP (p : pred A) xs : - reflect (forall x, p x -> x \in xs -> false) - ([seq x <- xs | p x] == [::]). -Proof. -case: eqP=>E; constructor. -- move=>x H1 H2; suff : x \in [seq x <- xs | 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 filter_pred1 x xs : - x \notin xs -> - filter (pred1 x) xs = [::]. -Proof. -move=>H; apply/eqP; apply/filter_nilP=>z /eqP ->. -by rewrite (negbTE H). -Qed. - -Lemma filter_predC1 x xs : - x \notin xs -> - filter (predC1 x) xs = xs. -Proof. -by move=>H; apply/all_filterP/allP=>y /=; case: eqP=>// ->; apply/contraL. -Qed. - -Lemma filter_mem_sym (s1 s2 : seq A) : - filter (mem s1) s2 =i filter (mem s2) s1. -Proof. by move=>x; rewrite !mem_filter andbC. Qed. - -Lemma index_inj xs x y : - x \in xs -> - index x xs = index y xs -> - x = y. -Proof. -elim: xs=>[|k xs IH] //=; rewrite inE eq_sym. -case: eqP=>[->{k} _|_ /= S]; case: eqP=>// _ []; apply: IH S. -Qed. - -Lemma cat_cancel (xs1 xs2 ys1 ys2 : seq A) (k : A) : - k \notin xs1 -> k \notin xs2 -> - xs1 ++ k :: ys1 = xs2 ++ k :: ys2 -> - (xs1 = xs2) * (ys1 = ys2). -Proof. -move=>Nk1 Nk2 E. -have Es : size xs1 = size xs2. -- have : index k (xs1++k::ys1) = index k (xs2++k::ys2) by rewrite E. - by rewrite !index_cat /= (negbTE Nk1) (negbTE Nk2) eqxx !addn0. -have Ex : xs1 = take (size xs1) (xs1 ++ k :: ys1). -- by rewrite take_cat ltnn subnn /= cats0. -rewrite E Es take_cat ltnn subnn /= cats0 in Ex. -rewrite {xs1 Nk1 Es}Ex in E *. -have : ys1 = drop (size (xs2++[::k])) (xs2++k::ys1). -- by rewrite drop_cat size_cat /= addn1 ltnNge ltnW //= subSn // subnn /= drop0. -by rewrite E drop_cat size_cat /= addn1 ltnNge ltnW //= subSn // subnn /= drop0. -Qed. - -(* if the list is not empty, the default value in head doesn't matter *) -Lemma head_dflt (x1 x2 x : A) xs : - x \in xs -> - head x1 xs = head x2 xs. -Proof. by case: xs. Qed. - -Lemma mem_head (x : A) xs : head x xs \in x :: xs. -Proof. by case: xs=>[|y ys]; rewrite !inE //= eqxx orbT. Qed. - -(* a common pattern of using mem_head that avoids forward reasoning *) -Lemma mem_headI (x : A) xs a : - a = head x xs -> - a \in x :: xs. -Proof. by move=>->; apply: mem_head. Qed. - -Lemma head_nilp (x : A) xs : - x \notin xs -> - head x xs = x -> - nilp xs. -Proof. -elim: xs=>[|y ys IH] //= H1 H2. -by rewrite H2 inE eqxx /= in H1. -Qed. - -Lemma head_notin (x y : A) xs : - y \in xs -> - x \notin xs -> - head x xs != x. -Proof. -move=>Y X; apply/negP=>/eqP; move/(head_nilp X)/nilP=>E. -by rewrite E in Y. -Qed. - -(* weaker form of in_mask *) -(* TODO upstream to mathcomp *) -Lemma in_mask_count x m xs : - count_mem x xs <= 1 -> - x \in mask m xs = (x \in xs) && nth false m (index x xs). -Proof. -elim: xs m => [|y xs IHs] m /=; first by rewrite mask0 in_nil. -case: m=>/=[|b m]; first by rewrite in_nil nth_nil andbF. -case: b; rewrite !inE eq_sym; case: eqP=>//= _. -- by rewrite add0n; apply: IHs. -- rewrite -{2}(addn0 1%N) leq_add2l leqn0 => /eqP Hc. - rewrite IHs; last by rewrite Hc. - by move/count_memPn/negbTE: Hc=>->. -by rewrite add0n; apply: IHs. -Qed. - -Lemma mem_take_index x xs : - x \notin take (index x xs) xs. -Proof. -elim: xs=>//=h xs; case: eqP=>//= /eqP H IH. -by rewrite inE negb_or eq_sym H. -Qed. - -Lemma prefix_drop_sub (s1 s2 : seq A) : - seq.prefix s1 s2 -> - forall n, {subset (drop n s1) <= drop n s2}. -Proof. -case/seq.prefixP=>s0 {s2}-> n x H. -rewrite drop_cat; case: ltnP=>Hn. -- by rewrite mem_cat H. -by move: H; rewrite drop_oversize. -Qed. - -End LemmasEq. - -(* decidable sequence disjointness *) - -Definition disjoint {A : eqType} (s1 s2 : seq A) := - all (fun x => x \notin s1) s2. - -Arguments disjoint {A} : simpl never. - -Lemma disjointC (A : eqType) (s1 s2 : seq A) : - disjoint s1 s2 = disjoint s2 s1. -Proof. -apply/idP/idP=>/allP S; apply/allP=>x X; -by apply/negP=>/S; rewrite X. -Qed. - -Lemma disjoint_catR (A : eqType) (s s1 s2 : seq A) : - disjoint s (s1 ++ s2) = - disjoint s s1 && disjoint s s2. -Proof. by rewrite /disjoint all_cat. Qed. - -Lemma disjoint_catL (A : eqType) (s s1 s2 : seq A) : - disjoint (s1 ++ s2) s = - disjoint s1 s && disjoint s2 s. -Proof. by rewrite -!(disjointC s) disjoint_catR. Qed. - -Lemma disjoint1L (A : eqType) x (s : seq A) : - disjoint [:: x] s = (x \notin s). -Proof. -apply/idP/idP. -- by apply: contraL=>X; apply/allPn; exists x=>//; rewrite negbK inE. -by apply: contraR=>/allPn [y H]; rewrite inE negbK =>/eqP <-. -Qed. - -Lemma disjoint1R (A : eqType) x (s : seq A) : - disjoint s [:: x] = (x \notin s). -Proof. by rewrite disjointC disjoint1L. Qed. - -Lemma disjoint_consL (A : eqType) x (s1 s2 : seq A) : - disjoint (x :: s1) s2 = - (x \notin s2) && disjoint s1 s2. -Proof. by rewrite -cat1s disjoint_catL disjoint1L. Qed. - -Lemma disjoint_consR (A : eqType) x (s1 s2 : seq A) : - disjoint s1 (x :: s2) = - (x \notin s1) && disjoint s1 s2. -Proof. by rewrite -cat1s disjoint_catR disjoint1R. Qed. - -Lemma disjoint_consLI (A : eqType) x (s1 s2 : seq A) : - x \notin s2 -> - disjoint s1 s2 -> - disjoint (x :: s1) s2. -Proof. by rewrite disjoint_consL=>->->. Qed. - -Lemma disjoint_consRI (A : eqType) x (s1 s2 : seq A) : - x \notin s1 -> - disjoint s1 s2 -> - disjoint s1 (x :: s2). -Proof. by rewrite disjoint_consR=>->->. Qed. - -Lemma disjoint_consLE (A : eqType) x (s1 s2 : seq A) : - disjoint (x :: s1) s2 -> - (x \notin s2) * (disjoint s1 s2). -Proof. by rewrite disjoint_consL=>/andX. Qed. - -Lemma disjoint_consRE (A : eqType) x (s1 s2 : seq A) : - disjoint s1 (x :: s2) -> - (x \notin s1) * (disjoint s1 s2). -Proof. by rewrite disjoint_consR=>/andX. Qed. - -Lemma disjoint_subL (A : eqType) (s s1 s2 : seq A) : - {subset s2 <= s1} -> - disjoint s s1 -> - disjoint s s2. -Proof. by move=>X /allP H; apply/allP=>z /X /H. Qed. - -Lemma disjoint_subR (A : eqType) (s s1 s2 : seq A) : - {subset s2 <= s1} -> - disjoint s1 s -> - disjoint s2 s. -Proof. -move=>X; rewrite disjointC=>/(disjoint_subL X). -by rewrite disjointC. -Qed. - -Lemma disjoint_eqL {A : eqType} {s s1 s2 : seq A} : - s1 =i s2 -> - disjoint s1 s = disjoint s2 s. -Proof. by move=>X; apply/idP/idP; apply: disjoint_subR=>z; rewrite X. Qed. - -Lemma disjoint_eqR {A : eqType} {s s1 s2 : seq A} : - s1 =i s2 -> - disjoint s s1 = disjoint s s2. -Proof. by move=>X; apply/idP/idP; apply: disjoint_subL=>z; rewrite X. Qed. - -Lemma disjointN (A : eqType) (s1 s2 : seq A) : - ~~ disjoint s1 s2 -> - exists2 x, x \in s1 & x \in s2. -Proof. by case/allPn=>x; rewrite negbK; exists x. Qed. - -(* finding last occurrence of element in a sequence *) - -Section FindLast. -Variables (T : Type). -Implicit Types (x : T) (p : pred T) (s : seq T). - -(* helper function for finding the last occurrence in a single pass *) -(* calculates index and size *) -Definition findlast_aux oi0 p s : option nat * nat := - foldl (fun '(o,i) x => (if p x then Some i else o, i.+1)) oi0 s. - -Lemma findlast_auxE oi0 p s : - findlast_aux oi0 p s = - let k := seq.find p (rev s) in - (if k == size s then oi0.1 - else Some (oi0.2 + (size s - k).-1), oi0.2 + size s). -Proof. -rewrite /findlast_aux; elim: s oi0=>/= [|x s IH] [o0 i0] /=. -- by rewrite addn0. -rewrite IH /= rev_cons -cats1 find_cat /= has_find. -move: (find_size p (rev s)); rewrite size_rev; case: ltngtP=>// H _. -- case: eqP=>[E|_]; first by rewrite E ltnNge leqnSn in H. - apply: injective_projections=>/=; [congr Some | rewrite addSnnS]=>//. - by rewrite !predn_sub /= -predn_sub addSnnS prednK // subn_gt0. -case: ifP=>_; rewrite addSnnS; last by rewrite addn1 eqxx. -by rewrite addn0 eqn_leq leqnSn /= ltnn subSnn addn0. -Qed. - -Definition findlast p s : nat := - let '(o, i) := findlast_aux (None, 0) p s in odflt i o. - -(* finding last is finding first in reversed list and flipping indices *) -Lemma findlastE p s : - findlast p s = - if has p s then (size s - seq.find p (rev s)).-1 else size s. -Proof. -rewrite /findlast findlast_auxE /= !add0n -has_rev; case/boolP: (has p (rev s)). -- by rewrite has_find size_rev; case: ltngtP. -by move/hasNfind=>->; rewrite size_rev eqxx. -Qed. - -Lemma findlast_size p s : - findlast p s <= size s. -Proof. -rewrite findlastE; case: ifP=>// _. -by rewrite -subnS; apply: leq_subr. -Qed. - -Lemma has_findlast p s : - has p s = (findlast p s < size s). -Proof. -symmetry; rewrite findlastE; case: ifP=>H /=; last by rewrite ltnn. -by rewrite -subnS /= ltn_subrL /=; case: s H. -Qed. - -Lemma hasNfindlast p s : - ~~ has p s -> - findlast p s = size s. -Proof. by rewrite has_findlast; case: ltngtP (findlast_size p s). Qed. - -Lemma findlast0 p x : findlast p [::] = 0. -Proof. by []. Qed. - -Lemma findlast1 p x : findlast p [::x] = ~~ p x. -Proof. by rewrite findlastE /= orbF; case: ifP=>// ->. Qed. - -Lemma findlast_cat p s1 s2 : - findlast p (s1 ++ s2) = - if has p s1 && ~~ has p s2 - then findlast p s1 - else size s1 + findlast p s2. -Proof. -rewrite !findlastE has_cat rev_cat find_cat has_rev size_cat size_rev. -case/boolP: (has p s2)=>H2; last first. -- rewrite orbF; case/boolP: (has p s1)=>//= H1. - by rewrite addnC subnDl. -have H2' : find p (rev s2) < size s2. -- by rewrite -size_rev -has_find has_rev. -rewrite /= orbT andbF -addnBA; last by apply: ltnW. -rewrite -!subn1 -subnDA -addnBA; last by rewrite subn_gt0. -by rewrite subnDA. -Qed. - -Lemma findlast_cons p x s : - findlast p (x::s) = - if p x && ~~ has p s then 0 else (findlast p s).+1. -Proof. -rewrite -cat1s findlast_cat /= !add1n orbF findlast1. -by case: ifP=>// /andP [->]. -Qed. - -Lemma findlast_rcons p x s : - findlast p (rcons s x) = - if p x then size s - else if has p s then findlast p s - else (size s).+1. -Proof. -rewrite -cats1 findlast_cat /= orbF findlast1. -case: (p x)=>/=; first by rewrite andbF addn0. -by rewrite andbT addn1. -Qed. - -Lemma nth_findlast x0 p s : - has p s -> - p (nth x0 s (findlast p s)). -Proof. -rewrite findlastE=>/[dup] E ->; rewrite -has_rev in E. -rewrite -subnS -nth_rev; last by rewrite -size_rev -has_find. -by apply: nth_find. -Qed. - -Lemma has_drop p s i : - has p s -> - has p (drop i s) = (i <= findlast p s). -Proof. -rewrite findlastE=>/[dup] E ->; rewrite -has_rev in E. -have Hh: 0 < size s - find p (rev s). -- by rewrite -size_rev subn_gt0 -has_find. -rewrite -size_rev; move/(has_take (size s - i)): (E). -rewrite take_rev -subnS size_rev. -case/boolP: (i < size s)=>[Hi|]. -- rewrite subnA //; last by apply: ltnW. - rewrite subnn add0n has_rev=>->. - rewrite ltn_subRL addnC -ltn_subRL subnS. - by case: (size s - find p (rev s)) Hh. -rewrite -ltnNge ltnS => Hi _. -rewrite drop_oversize //=; symmetry; apply/negbTE. -rewrite -ltnNge subnS prednK //. -by apply/leq_trans/Hi; exact: leq_subr. -Qed. - -Lemma find_geq p s i : - has p (drop i s) -> i <= findlast p s. -Proof. -case/boolP: (has p s)=>Hp; first by rewrite (has_drop _ Hp). -suff: ~~ has p (drop i s) by move/negbTE=>->. -move: Hp; apply: contra; rewrite -{2}(cat_take_drop i s) has_cat=>->. -by rewrite orbT. -Qed. - -Lemma find_leq_last p s : - find p s <= findlast p s. -Proof. -rewrite findlastE. -case/boolP: (has p s)=>[|_]; last by apply: find_size. -elim: s=>//= h s IH. -rewrite rev_cons -cats1 find_cat has_rev size_rev /=. -case/orP; first by move=>->. -move=>/[dup] H ->; case: ifP=>_ //. -rewrite subSn /=; last first. -- by rewrite -size_rev; apply: find_size. -apply: (leq_ltn_trans (IH H)); rewrite ltn_predL subn_gt0. -by rewrite -size_rev -has_find has_rev. -Qed. - -Variant split_findlast_nth_spec p : seq T -> seq T -> seq T -> T -> Type := - FindLastNth x s1 s2 of p x & ~~ has p s2 : - split_findlast_nth_spec p (rcons s1 x ++ s2) s1 s2 x. - -Lemma split_findlast_nth x0 p s (i := findlast p s) : - has p s -> - split_findlast_nth_spec p s (take i s) (drop i.+1 s) (nth x0 s i). -Proof. -move=> p_s; rewrite -[X in split_findlast_nth_spec _ X](cat_take_drop i s). -rewrite (drop_nth x0 _); last by rewrite -has_findlast. -rewrite -cat_rcons; constructor; first by apply: nth_findlast. -by rewrite has_drop // ltnn. -Qed. - -Variant split_findlast_spec p : seq T -> seq T -> seq T -> Type := - FindLastSplit x s1 s2 of p x & ~~ has p s2 : - split_findlast_spec p (rcons s1 x ++ s2) s1 s2. - -Lemma split_findlast p s (i := findlast p s) : - has p s -> - split_findlast_spec p s (take i s) (drop i.+1 s). -Proof. -by case: s => // x ? in i * =>?; case: split_findlast_nth=>//; constructor. -Qed. - -End FindLast. - -Section FindLastEq. -Variables T : eqType. -Implicit Type s : seq T. - -Definition indexlast (x : T) : seq T -> nat := findlast (pred1 x). - -Lemma indexlast_size x s : indexlast x s <= size s. -Proof. by rewrite /indexlast; apply: findlast_size. Qed. - -Lemma indexlast_mem x s : (indexlast x s < size s) = (x \in s). -Proof. by rewrite /indexlast -has_findlast has_pred1. Qed. - -Lemma memNindexlast x s : x \notin s -> indexlast x s = size s. -Proof. by rewrite -has_pred1=>/hasNfindlast. Qed. - -Lemma indexlast0 x : indexlast x [::] = 0. -Proof. by []. Qed. - -Lemma indexlast1 x y : indexlast x [::y] = (x != y). -Proof. by rewrite /indexlast findlast1 /= eq_sym. Qed. - -Lemma indexlast_cat x s1 s2 : - indexlast x (s1 ++ s2) = - if (x \in s1) && (x \notin s2) - then indexlast x s1 - else size s1 + indexlast x s2. -Proof. by rewrite /indexlast findlast_cat !has_pred1. Qed. - -Lemma indexlast_cons x y s : - indexlast x (y::s) = - if (y == x) && (x \notin s) then 0 else (indexlast x s).+1. -Proof. by rewrite /indexlast findlast_cons has_pred1. Qed. - -Lemma indexlast_rcons x y s : - indexlast x (rcons s y) = - if y == x then size s - else if x \in s then indexlast x s - else (size s).+1. -Proof. by rewrite /indexlast findlast_rcons has_pred1. Qed. - -Lemma index_geq x s i : - x \in drop i s -> - i <= indexlast x s. -Proof. by rewrite -has_pred1; apply: find_geq. Qed. - -Lemma index_leq_last x s : index x s <= indexlast x s. -Proof. by apply: find_leq_last. Qed. - -Lemma indexlast_count x s : - count_mem x s <= 1 <-> index x s = indexlast x s. -Proof. -elim: s=>//= h t IH; rewrite indexlast_cons. -case: eqP=>/= ?; last first. -- by rewrite add0n IH; split=>[->|[]]. -rewrite add1n ltnS leqn0; split. -- by move/eqP/count_memPn=>->. -by case: ifP=>//= /count_memPn->. -Qed. - -Lemma index_lt_last x s : - 1 < count_mem x s -> - index x s < indexlast x s. -Proof. -move=>H; move: (index_leq_last x s); rewrite leq_eqVlt. -by case: eqP=>//= /indexlast_count; case: leqP H. -Qed. - -Lemma indexlast_uniq x s : - uniq s -> - index x s = indexlast x s. -Proof. -move=>H; apply/indexlast_count. -by rewrite count_uniq_mem //; apply: leq_b1. -Qed. - -Lemma indexlast_memN x xs : - x \notin xs <-> indexlast x xs = size xs. -Proof. -split; first by exact: memNindexlast. -by move=>E; rewrite -indexlast_mem E ltnn. -Qed. - -Lemma index_last_inj x y s : - (x \in s) || (y \in s) -> - index x s = indexlast y s -> x = y. -Proof. -elim: s=>[|k s IH] //=; rewrite !inE indexlast_cons !(eq_sym k). -case: eqP=>[{k}<- _|_ /= S]; first by case: eqP=>//=. -move: S; case/boolP: (y \in s)=>/=. -- by rewrite andbF=>H _ []; apply: IH; rewrite H orbT. -move=>Ny; rewrite orbF andbT. -by case: eqP=>//; rewrite orbF=>_ H []; apply: IH; rewrite H. -Qed. - -Lemma indexlast_inj x y s : - x \in s -> - indexlast x s = indexlast y s -> - x = y. -Proof. -elim: s=>[|k s IH] //=; rewrite inE eq_sym !indexlast_cons. -case: eqP=>[->{k} _|_ /= S] /=. -- case: eqP=>//= _. - by case: ifP=>// /negbT; rewrite negbK=>H; case; apply: IH. -by case: ifP=>// _ []; apply: IH. -Qed. - -Lemma mem_drop_indexlast x s : - x \notin drop (indexlast x s).+1 s. -Proof. -elim: s=>//=h s; rewrite indexlast_cons. -case: eqP=>//= _ H. -by case: ifP=>//=; rewrite drop0. -Qed. - -Variant splitLast x : seq T -> seq T -> seq T -> Type := - SplitLast p1 p2 of x \notin p2 : splitLast x (rcons p1 x ++ p2) p1 p2. - -Lemma splitLastP s x (i := indexlast x s) : - x \in s -> - splitLast x s (take i s) (drop i.+1 s). -Proof. -case: s => // y s in i * => H. -case: split_findlast_nth=>//; first by rewrite has_pred1. -move=>_ s1 s2 /= /eqP->; rewrite has_pred1 => H2. -by constructor. -Qed. - -End FindLastEq. - -(* finding all occurrences *) - -Section FindAll. -Variables T : Type. -Implicit Types (x : T) (p : pred T) (s : seq T). - -(* helper function for finding all occurrences in a single pass with difference lists *) -Definition findall_aux oi0 p s : (seq nat -> seq nat) * nat := - foldl (fun '(s,i) x => (if p x then s \o cons i else s, i.+1)) oi0 s. - -Lemma findall_auxE oi0 p s : - (forall s1 s2 : seq nat, oi0.1 (s1 ++ s2) = oi0.1 s1 ++ s2) -> - let: (rs, ix) := findall_aux oi0 p s in - (forall s' : seq nat, - rs s' = oi0.1 (unzip1 (filter (p \o snd) - (zip (iota oi0.2 (size s)) s))) ++ s') - /\ ix = oi0.2 + size s. -Proof. -move; rewrite /findall_aux; elim: s oi0=>[|x s IH] [o0 i0] /= H0. -- split; last by rewrite addn0. - by move=>s'; rewrite -{1}(cat0s s') H0. -case/boolP: (p x)=>/= Hpx. -- move: (IH (o0 \o cons i0, i0.+1))=>/=. - rewrite addSn addnS; apply=>s1 s2. - by rewrite -cat_cons H0. -by move: (IH (o0, i0.+1))=>/=; rewrite addSn addnS; apply. -Qed. - -Definition findall p s : seq nat := (findall_aux (id, 0) p s).1 [::]. - -Lemma findallE p s : - findall p s = unzip1 (filter (p \o snd) (zip (iota 0 (size s)) s)). -Proof. -rewrite /findall. -move: (@findall_auxE (id, 0) p s)=>/= /(_ (fun s1 s2 : seq nat => erefl)). -case E: (findall_aux (id, 0) p s)=>[rs ix] /=; case=>/(_ [::]) -> _. -by rewrite cats0. -Qed. - -Lemma findall_cat p s1 s2 : - findall p (s1 ++ s2) = - findall p s1 ++ map (fun n => size s1 + n) (findall p s2). -Proof. -rewrite !findallE size_cat iotaD add0n zip_cat; last by rewrite size_iota. -rewrite filter_cat {1}/unzip1 map_cat; congr (_ ++ _). -set n := size s1. -rewrite -{1}(addn0 n) iotaDl zip_mapl filter_map -!map_comp. -rewrite (eq_filter (a2:=(p \o snd))); last by case. -by apply: eq_map; case. -Qed. - -Lemma findall_cons p x s : - findall p (x::s) = - if p x then 0 :: map S (findall p s) else map S (findall p s). -Proof. by rewrite -cat1s findall_cat /= findallE /=; case: (p x). Qed. - -Lemma findall_rcons p x s : - findall p (rcons s x) = - if p x then rcons (findall p s) (size s) else findall p s. -Proof. -rewrite -cats1 findall_cat /= !findallE /=; case: (p x)=>/=. -- by rewrite addn0 cats1. -by rewrite cats0. -Qed. - -Lemma findall_size p s : - size (findall p s) = count p s. -Proof. -by elim: s=>//=x s IH; rewrite findall_cons; case: (p x)=>/=; -rewrite size_map IH. -Qed. - -Lemma findall_nilp p s : - nilp (findall p s) = ~~ has p s. -Proof. by rewrite /nilp findall_size has_count -leqNgt leqn0. Qed. - -Lemma findall_head p s : - find p s = head (size s) (findall p s). -Proof. -elim: s=>//=x s IH; rewrite findall_cons; case: (p x)=>//=. -by rewrite -head_map IH. -Qed. - -Lemma findall_last p s : - findlast p s = last (size s) (findall p s). -Proof. -elim/last_ind: s=>//=s x IH; rewrite findall_rcons findlast_rcons size_rcons. -case: (p x)=>/=; first by rewrite last_rcons. -by rewrite IH -(negbK (has _ _)) -findall_nilp; case: (findall p s). -Qed. - -Lemma findall_count1 p s : - count p s <= 1 -> - findall p s = if has p s then [::find p s] else [::]. -Proof. -rewrite -(findall_size p s); case E: (findall p s)=>[|h t] /=. -- by move/nilP: E; rewrite findall_nilp=>/negbTE ->. -have ->: has p s by rewrite -(negbK (has p s)) -findall_nilp E. -rewrite ltnS leqn0 => /eqP/size0nil Et; move: E; rewrite {t}Et. -by rewrite findall_head=>->. -Qed. - -End FindAll. - -Section FindAllEq. -Variables T : eqType. -Implicit Type s : seq T. - -Definition indexall (x : T) : seq T -> seq nat := findall (pred1 x). - -Lemma indexall_size x s : - size (indexall x s) = count_mem x s. -Proof. by rewrite /indexall findall_size. Qed. - -Lemma indexall_mem x s : nilp (indexall x s) = (x \notin s). -Proof. by rewrite /indexall findall_nilp has_pred1. Qed. - -Lemma indexall_head x s : - index x s = head (size s) (indexall x s). -Proof. by rewrite /index /indexall findall_head. Qed. - -Lemma indexall_last x s : - indexlast x s = last (size s) (indexall x s). -Proof. by rewrite /indexlast /indexall findall_last. Qed. - -Lemma indexall_count1 x s : - count_mem x s <= 1 -> - indexall x s = if x \in s then [:: index x s] else [::]. -Proof. by rewrite /indexall /index -has_pred1; apply: findall_count1. Qed. - -Corollary indexall_uniq x s : - uniq s -> - indexall x s = if x \in s then [:: index x s] else [::]. -Proof. by move=>U; apply: indexall_count1; rewrite (count_uniq_mem _ U) leq_b1. Qed. - -End FindAllEq. - -(* Interaction of filter/last/index *) - -Section FilterLastIndex. -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. - -Lemma last_notin_nilp k (s : seq A) : - ~~ nilp s -> - k \notin s -> - last k s != k. -Proof. -move=>N; apply: (last_notin (x := head k s)). -by case: s N=>//= x s _; rewrite inE eqxx. -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 eqxx. 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 and rcons *) - -Lemma rcons_lastX x y (s : seq A) : - x \in s -> - exists s', s = rcons s' (last y s). -Proof. -elim/last_ind: s=>[|ks k IH] //=. -by rewrite last_rcons; exists ks. -Qed. - -Lemma rcons_lastP x (s : seq A) : - reflect (exists s', s = rcons s' (last x s)) - (last x s \in s). -Proof. -case X : (last x s \in s); constructor; first by apply: rcons_lastX X. -case=>s' E; move/negP: X; elim. -by rewrite E last_rcons mem_rcons inE eqxx. -Qed. - -Lemma rcons_lastXP x y (s : seq A) : - reflect (exists s', s = rcons s' x) - ((x == last y s) && (x \in s)). -Proof. -case: eqP=>[->|N]; first by apply: rcons_lastP. -by constructor; case=>s' E; elim: N; rewrite E last_rcons. -Qed. - -Lemma rcons_lastN (s : seq A) a p : - p != a -> - p = last a s -> - exists ss, s = rcons ss p. -Proof. by move/[swap]=>-> N; apply/rcons_lastX/last_change/N. Qed. - -Lemma index_last_size_uniq z (s : seq A) : - uniq s -> - index (last z s) s = (size s).-1. -Proof. -elim: s z=>//= x s IH z. -case/andP=>Nx U; rewrite eq_sym; case: eqVneq=>H. -- by rewrite (last_nochange_nil H Nx). -rewrite {}IH //; apply: prednK. -by case: {Nx U}s H=>//=; rewrite eqxx. -Qed. - -(* last has bigger index than anything in x *) -Lemma index_last_mono x k (s : seq A) : - uniq s -> - x \in s -> - index x s <= index (last k s) s. -Proof. -elim: s k=>[|k s IH] //= k'; rewrite inE !(eq_sym k). -case/andP=>K U; case: (x =P k)=>//= /eqP N X. -case: (last k s =P k)=>[/last_nochange|/eqP L]. -- by case: eqP X=>[->//|]; rewrite (negbTE K). -by apply: leq_trans (IH k U X) _. -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 : nilp s by move/nilP=>->. -rewrite /nilp 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 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. - -Lemma index_filter_ltL (p : pred A) (ks : seq A) (t1 t2 : A) : - (t1 \notin ks) || p t1 -> - (index t1 ks < index t2 ks) -> - (index t1 (filter p ks) < index t2 (filter p ks)). -Proof. -elim: ks t1 t2=>[|k ks IH] t1 t2 //=; rewrite inE negb_or (eq_sym t1). -case: eqP=>[->{k} /= Pt1|/eqP Nkt1 /= H]. -- by rewrite Pt1 /= eqxx; case: eqP. -case: eqP=>// /eqP Nkt2; case: ifP=>H1 /=. -- by rewrite (negbTE Nkt1) (negbTE Nkt2) !ltnS; apply: IH H. -by rewrite ltnS; apply: IH H. -Qed. - -Lemma index_filter_leL (p : pred A) (ks : seq A) (t1 t2 : A) : - (t1 \notin ks) || p t1 -> - (index t1 ks <= index t2 ks) -> - (index t1 (filter p ks) <= index t2 (filter p ks)). -Proof. -elim: ks t1 t2=>[|k ks IH] t1 t2 //=; rewrite inE negb_or (eq_sym t1). -case: eqP=>[->{k} /= Pt1|/eqP Nkt1 /= H]. -- by rewrite Pt1 /= eqxx; case: eqP. -case: eqP=>// /eqP Nkt2; case: ifP=>H1 /=. -- by rewrite (negbTE Nkt1) (negbTE Nkt2) !ltnS; apply: IH H. -by rewrite ltnS; apply: IH H. -Qed. - -Lemma index_filter_ltR (p : pred A) (ks : seq A) (t1 t2 : A) : - (t2 \notin 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 negb_or /=. -rewrite (eq_sym t2). -case: eqP=>[->{k} /= Pt2|/eqP Nkt2 /=]. -- by rewrite Pt2 /= eqxx; case: eqP. -case: eqP=>[->{t1}//|/eqP Nt1k]. -case: ifP=>H1 H2 /=. -- by rewrite (negbTE Nt1k) (negbTE Nkt2) !ltnS; apply: IH H2. -by rewrite ltnS; apply: IH H2. -Qed. - -Lemma index_filter_leR (p : pred A) (ks : seq A) (t1 t2 : A) : - (t2 \notin 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 negb_or /=. -rewrite (eq_sym t2). -case: eqP=>[->{k} /= Pt2|/eqP Nkt2 /=]. -- by rewrite Pt2 /= eqxx; case: eqP. -case: eqP=>[->{t1}//|/eqP Nt1k]. -case: ifP=>H1 H2 /=. -- by rewrite (negbTE Nt1k) (negbTE Nkt2) !ltnS; apply: IH H2. -by rewrite ltnS; apply: IH H2. -Qed. - -(* we can put the left and right lemmas together *) -Lemma index_filter_lt (p : pred A) (ks : seq A) (t1 t2 : A) : - (t1 \notin ks) || p t1 -> - (t2 \notin ks) || p t2 -> - (index t1 (filter p ks) < index t2 (filter p ks)) = - (index t1 ks < index t2 ks). -Proof. -move=>H1 H2; apply/idP/idP. -- by apply: index_filter_ltR. -by apply: index_filter_ltL. -Qed. - -Lemma index_filter_le (p : pred A) (ks : seq A) (t1 t2 : A) : - (t1 \notin ks) || p t1 -> - (t2 \notin ks) || p t2 -> - (index t1 (filter p ks) <= index t2 (filter p ks)) = - (index t1 ks <= index t2 ks). -Proof. -move=>H1 H2; apply/idP/idP. -- by apply: index_filter_leR. -by apply: index_filter_leL. -Qed. - -(* index and masking *) - -Lemma index_mask (s : seq A) m a b : - uniq s -> - a \in mask m s -> - b \in mask m s -> - index a (mask m s) < index b (mask m s) <-> - index a s < index b s. -Proof. -elim: m s=>[|x m IH][|k s] //= /andP [K U]; case: x=>[|Ma Mb] /=. -- rewrite !inE; case/orP=>[/eqP <-|Ma]. - - by case/orP=>[/eqP ->|]; rewrite eqxx //; case: eqP. - case/orP=>[/eqP ->|Mb]; first by rewrite eqxx. - by case: eqP; case: eqP=>//; rewrite ltnS IH. -case: eqP Ma K=>[-> /mem_mask -> //|Ka]. -case: eqP Mb=>[-> /mem_mask -> //|Kb Mb Ma]. -by rewrite ltnS IH. -Qed. - -Lemma indexlast_mask (s : seq A) m a b : - uniq s -> - a \in mask m s -> - b \in mask m s -> - indexlast a (mask m s) < indexlast b (mask m s) <-> - indexlast a s < indexlast b s. -Proof. -elim: m s=>[|x m IH][|k s] //= /andP [K U]; case: x=>[|Ma Mb] /=; -rewrite !indexlast_cons. -- rewrite !inE; case/orP=>[/eqP Ea|Ma]. - - rewrite -{k}Ea in K *. - have Km: (a \notin mask m s) by apply: contra K; apply: mem_mask. - case/orP=>[/eqP ->|]; rewrite eqxx /=; first by rewrite K ltnn. - case: eqP=>[->|Nab] H /=; first by rewrite !ltnn. - by rewrite K Km. - case/orP=>[/eqP Eb|Mb]. - - rewrite -{k}Eb in K *. - have ->: (b \notin mask m s) by apply: contra K; apply: mem_mask. - by rewrite eqxx /= K. - rewrite Ma Mb (mem_mask Ma) (mem_mask Mb) /= !andbF. - by apply: IH. -case: eqP Ma K=>[-> /mem_mask -> //|Ka] /=. -case: eqP Mb=>[-> /mem_mask -> //|Kb Mb Ma _] /=. -by rewrite ltnS IH. -Qed. - -Lemma index_subseq (s1 s2 : seq A) a b : - subseq s1 s2 -> - uniq s2 -> - a \in s1 -> - b \in s1 -> - index a s1 < index b s1 <-> index a s2 < index b s2. -Proof. by case/subseqP=>m _ ->; apply: index_mask. Qed. - -Lemma indexlast_subseq (s1 s2 : seq A) a b : - subseq s1 s2 -> - uniq s2 -> - a \in s1 -> - b \in s1 -> - indexlast a s1 < indexlast b s1 <-> indexlast a s2 < indexlast b s2. -Proof. by case/subseqP=>m _ ->; apply: indexlast_mask. Qed. - -End FilterLastIndex. - -(* index and mapping *) - -Section IndexPmap. -Variables A B : eqType. - -Lemma index_pmap_inj (s : seq A) (f : A -> option B) a1 a2 b1 b2 : - injective f -> - f a1 = Some b1 -> - f a2 = Some b2 -> - index b1 (pmap f s) < index b2 (pmap f s) <-> index a1 s < index a2 s. -Proof. -move=>Inj E1 E2; elim: s=>[|k s IH] //=; rewrite /oapp. -case: eqP=>[->{k}|]. -- rewrite E1 /= eqxx. - case: (a1 =P a2) E1 E2=>[-> -> [/eqP ->] //|]. - by case: (b1 =P b2)=>[-> Na <- /Inj /esym/Na|]. -case: eqP=>[->{k} Na|N2 N1]; first by rewrite E2 /= eqxx !ltn0. -case E : (f k)=>[b|] //=. -case: eqP E1 E=>[-><- /Inj/N1 //|_ _]. -by case: eqP E2=>[-><- /Inj/N2 //|_ _ _]; rewrite IH. -Qed. - -Lemma index_pmap_inj_mem (s : seq A) (f : A -> option B) a1 a2 b1 b2 : - {in s &, injective f} -> - a1 \in s -> - a2 \in s -> - f a1 = Some b1 -> - f a2 = Some b2 -> - index b1 (pmap f s) < index b2 (pmap f s) <-> index a1 s < index a2 s. -Proof. -move=>Inj A1 A2 E1 E2. -elim: s Inj A1 A2=>[|k s IH] //= Inj; rewrite /oapp !inE !(eq_sym k). -case: eqP Inj=>[<-{k} /= Inj _|]. -- rewrite E1 /= !eqxx eq_sym. - case: eqP E1 E2=>[->-> [->]|]; first by rewrite eqxx. - case: eqP=>[-> Na <- E /= A2|//]. - by move/Inj: E Na=>-> //; rewrite inE ?(eqxx,A2,orbT). -case eqP=>[<-{k} Na Inj /= A1 _|]; first by rewrite E2 /= eqxx !ltn0. -move=>N2 N1 Inj /= A1 A2. -have Inj1 : {in s &, injective f}. -- by move=>x y X Y; apply: Inj; rewrite inE ?X ?Y ?orbT. -case E : (f k)=>[b|] /=; last by rewrite IH. -case: eqP E1 E=>[-> <- E|_ _]. -- by move/Inj: E N1=>-> //; rewrite inE ?(eqxx,A1,orbT). -case: eqP E2=>[-><- E|_ _ _]; last by rewrite IH. -by move/Inj: E N2=>-> //; rewrite inE ?(eqxx,A2,orbT). -Qed. - -(* we can relax the previous lemma a bit *) -(* the relaxation will be more commonly used than the previous lemma *) -(* because the option type gives us the implication that the second *) -(* element is in the map *) -Lemma index_pmap_inj_in (s : seq A) (f : A -> option B) a1 a2 b1 b2 : - {in s & predT, injective f} -> - f a1 = Some b1 -> - f a2 = Some b2 -> - index b1 (pmap f s) < index b2 (pmap f s) <-> index a1 s < index a2 s. -Proof. -move=>Inj E1 E2. -case A1 : (a1 \in s); last first. -- move/negbT/index_sizeE: (A1)=>->. - suff /index_sizeE -> : b1 \notin pmap f s by rewrite !ltnNge !index_size. - rewrite mem_pmap; apply/mapP; case=>x X /esym; rewrite -E1=>E. - by move/(Inj _ _ X): E A1=><- //; rewrite X. -case A2 : (a2 \in s). -- by apply: index_pmap_inj_mem=>// x y X _; apply: Inj. -move/negbT/index_sizeE: (A2)=>->. -suff /index_sizeE -> : b2 \notin pmap f s. -- by rewrite !index_mem /= A1 mem_pmap; split=>// _; apply/mapP; exists a1. -rewrite mem_pmap; apply/mapP; case=>x X /esym; rewrite -E2=>E. -by move/(Inj _ _ X): E A2=><- //; rewrite X. -Qed. - -End IndexPmap. - -Section Allrel. -Variables S T : Type. - -Lemma allrel_rconsl (r : T -> S -> bool) x xs ys : - allrel r (rcons xs x) ys = allrel r xs ys && all (r x) ys. -Proof. by rewrite -cats1 allrel_catl allrel1l. Qed. - -Lemma allrel_rconsr (r : T -> S -> bool) y xs ys : - allrel r xs (rcons ys y) = allrel r xs ys && all (r^~ y) xs. -Proof. by rewrite -cats1 allrel_catr allrel1r. Qed. - -End Allrel. - - -Section AllrelEq. -Variables S T : eqType. - -Lemma allrel_in_l (r : T -> S -> bool) (xs xs' : seq T) (ys : seq S) : - xs =i xs' -> - allrel r xs ys = allrel r xs' ys. -Proof. -by move=>H; rewrite !allrel_allpairsE; apply/eq_all_r/mem_allpairs_dep. -Qed. - -Lemma allrel_in_r (r : T -> S -> bool) (xs : seq T) (ys ys' : seq S) : - ys =i ys' -> - allrel r xs ys = allrel r xs ys'. -Proof. -by move=>H; rewrite !allrel_allpairsE; apply/eq_all_r/mem_allpairs_dep. -Qed. - -Lemma allrel_sub_l (r : T -> S -> bool) (xs xs' : seq T) (ys : seq S) : - {subset xs' <= xs} -> - allrel r xs ys -> allrel r xs' ys. -Proof. -move=>H Ha; apply/allrelP=>x y Hx Hy. -by move/allrelP: Ha; apply=>//; apply: H. -Qed. - -Lemma allrel_sub_r (r : T -> S -> bool) (xs : seq T) (ys ys' : seq S) : - {subset ys' <= ys} -> - allrel r xs ys -> allrel r xs ys'. -Proof. -move=>H Ha; apply/allrelP=>x y Hx Hy. -by move/allrelP: Ha; apply=>//; apply: H. -Qed. - -Lemma allrel_trans (xs ys : seq S) z r : - transitive r -> - all (r^~ z) xs -> all (r z) ys -> allrel r xs ys. -Proof. -move=>Ht /allP Ha /allP Hp; apply/allrelP=>x y + Hy. -by move/Ha/Ht; apply; apply: Hp. -Qed. - -End AllrelEq. - -Section SeqRel. -Variable A : eqType. -Implicit Type ltT leT : rel A. - -(* ordering with path, seq and last *) - -Lemma eq_last (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 (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 (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 eqxx. -by rewrite (H (last a s)) ?orbT // mem_last. -Qed. - -Lemma path_lastR (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. - -Lemma path_prev leT s a x : - x \in s -> - path leT a s -> - exists y, y \in belast a s /\ leT y x. -Proof. -case/splitPr=>p1 p2; rewrite cat_path /= =>/and3P []. -by exists (last a p1); rewrite belast_cat /= mem_cat inE eqxx orbT. -Qed. - -Lemma path_next leT s a x b : - x \in a :: s -> - path leT a (rcons s b) -> - exists y, y \in rcons s b /\ leT x y. -Proof. -case/splitPl=>p1 p2 Ex; rewrite rcons_cat cat_path rcons_path Ex. -case/and3P=>H1 H2 H3; case: p2 H2 H3=>[|y p2] /= H2 H3. -- by exists b; rewrite mem_cat inE eqxx orbT. -case/andP: H2=>H2 _; exists y; split=>//. -by rewrite mem_cat inE eqxx orbT. -Qed. - -Lemma path_uniq leT a s : - (forall x y, leT x y -> y != a) -> - (forall a b x, leT a x -> leT b x -> a = b) -> - path leT a s -> - uniq (a :: s). -Proof. -move=>Na Fp; elim/last_ind: s=>[|s x IH] //=. -rewrite rcons_path mem_rcons=>/andP [Px Cx]. -move: {IH}(IH Px) (IH Px); rewrite {1}lastI /=. -rewrite !rcons_uniq inE negb_or=>/andP [N _]/andP [->->]. -rewrite eq_sym (Na _ _ Cx) !andbT /=; apply/negP. -move/path_prev=>/(_ _ _ Px) [/= y][/[swap]] /(Fp _ _ _ Cx) <-. -by rewrite (negbTE N). -Qed. - -(* in a sorted list, the last element is maximal *) -(* and the maximal element is last *) - -Lemma sorted_last_key_max (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 (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 (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 (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 eqxx. -by move=>z D; apply: H2; rewrite inE D orbT. -Qed. - -Lemma seq_last_mono (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 (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 (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 (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 path_mem_irr (s : seq A) ltT x : - irreflexive ltT -> - transitive ltT -> - path ltT x s -> - x \notin s. -Proof. -move=>I T P; apply: contraFT (I x). -by rewrite negbK; apply: path_mem T P. -Qed. - -Lemma sorted_rcons (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_rconsE (leT : rel A) xs x : - transitive leT -> - sorted leT (rcons xs x) = - all (leT^~ x) xs && sorted leT xs. -Proof. -move/rev_trans=>Ht; rewrite -(revK (rcons _ _)) rev_rcons rev_sorted /=. -by rewrite path_sortedE // all_rev rev_sorted. -Qed. - -Lemma sorted1 (r : rel A) xs : - size xs == 1 -> - sorted r xs. -Proof. by case: xs=>// x; case. Qed. - -Lemma sorted_subset_subseq (s1 s2 : seq A) ltT : - irreflexive ltT -> - transitive ltT -> - sorted ltT s1 -> - sorted ltT 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: irr_sorted_eq S1 _ _=>//; first by rewrite sorted_filter. -by move=>k; rewrite mem_filter; case S : (_ \in _)=>//; rewrite (H _ S). -Qed. - -Lemma sorted_ord_index (s : seq A) ltT x y : - irreflexive ltT -> - transitive ltT -> - sorted ltT s -> - x \in s -> - ltT 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 path_ord_index_leq (s : seq A) leT x y : - transitive leT -> - antisymmetric leT -> - leT x y -> - path leT y s -> - x \in s -> - x = y. -Proof. -move=>T; elim: s x y=>[|a l IH] //= x y As Lxy. -case/andP=>Lya Pal; rewrite inE. -case: eqP Lya Pal As=>[<-{a} Lyx _ As _|Nxa Lya Pal /= As' X]. -- by apply: As=>//; rewrite Lxy Lyx. -by move/Nxa: (IH x a As' (T _ _ _ Lxy Lya) Pal X). -Qed. - -Lemma sorted_ord_index_leq (s : seq A) leT x y : - transitive leT -> - antisymmetric leT -> - sorted leT s -> - x \in s -> - leT x y -> - x != y -> - index x s < index y s. -Proof. -move=>T As S P H N; elim: s S As P=>[|z s IH] //= P As; rewrite inE !(eq_sym z). -case: eqP H P As=>[<-{z} H P As _|Nxz H P As /= X]; first by rewrite eq_sym (negbTE N). -case: eqP Nxz P=>[<-{z} Nxy P|Nyz Nxz P]. -- by move/Nxy: (path_ord_index_leq T As H P X). -by apply: IH X=>//; apply: path_sorted P. -Qed. - -Lemma sorted_index_ord (s : seq A) leT x y : - transitive leT -> - sorted leT s -> - y \in s -> - index x s < index y s -> - leT x y. -Proof. -move=>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. - -(* sorted, uniq, filter *) - -Lemma lt_sorted_uniq_le (s : seq A) ltT : - irreflexive ltT -> - antisymmetric ltT -> - transitive ltT -> - sorted ltT s = uniq s && - (sorted (fun k t => (k == t) || ltT k t) s). -Proof. -move=>I As T; case: s=>// n s; elim: s n=>//= m s IHs n. -rewrite inE negb_or IHs -!andbA /=. -case: (n =P m)=>[->|/eqP Nm /=]; first by rewrite I. -case lTnm : (ltT n m)=>/=; last by rewrite !andbF. -case Ns: (n \in s)=>//=; do !bool_congr. -have T' : transitive (fun k t => (k == t) || ltT k t). -- move=>x y z /orP [/eqP -> //|H]. - case/orP=>[/eqP <-|]; first by rewrite H orbT. - by move/(T _ _ _ H)=>->; rewrite orbT. -apply/negP=>/(order_path_min T')/allP/(_ n Ns). -rewrite eq_sym (negbTE Nm) /= =>lTmn. -by rewrite (As m n) ?eqxx // lTnm lTmn in Nm. -Qed. - -Lemma sort_sorted_in_lt (s : seq A) ltT : - irreflexive ltT -> - antisymmetric ltT -> - transitive ltT -> - uniq s -> - {in s &, total (fun k t => (k == t) || ltT k t)} -> - sorted ltT (sort (fun k t => (k == t) || ltT k t) s). -Proof. -move=>I S T U Tot; rewrite lt_sorted_uniq_le //. -by rewrite sort_uniq U (sort_sorted_in Tot _). -Qed. - -(* filtering and consecutive elements in an order *) -Lemma filterCN (ltT : rel A) f t1 t2 : - t1 \notin f -> - {in f, forall z, ltT z t2 = (z == t1) || ltT z t1} -> - filter (ltT^~ t2) f = filter (ltT^~ t1) f. -Proof. -move=>N C; apply: eq_in_filter=>x T; rewrite C ?inE ?orbT //. -by case: eqP N T=>// -> /negbTE ->. -Qed. - -Lemma filterCE (ltT : rel A) f t1 t2 : - irreflexive ltT -> - transitive ltT -> - sorted ltT f -> - {in f, forall z, ltT z t2 = (z == t1) || ltT z t1} -> - t1 \in f -> - filter (ltT^~ t2) f = filter (ltT^~ t1) f ++ [:: t1]. -Proof. -move=>I T S Z F; have U : uniq f by apply: sorted_uniq T I _ S. -rewrite -(filter_pred1_uniq U F); apply: irr_sorted_eq (T) I _ _ _ _ _. -- by apply: sorted_filter T _ _ S. -- rewrite -[filter (ltT^~ t1) _]revK -[filter (pred1 t1) _]revK -rev_cat. - rewrite rev_sorted -filter_rev filter_pred1_uniq ?(mem_rev,rev_uniq) //. - rewrite /= path_min_sorted ?(rev_sorted, sorted_filter T _ S) //. - by apply/allP=>x; rewrite mem_rev mem_filter=>/andP []. -move=>x; rewrite mem_cat !mem_filter /=. -case X: (x \in f); last by rewrite !andbF. -by rewrite Z // orbC !andbT. -Qed. - -(* frequently we have nested filtering and sorting *) -(* for which the following forms of the lemmas is more effective *) - -Lemma filter2CN (ltT : rel A) p f t1 t2 : - t1 \notin p -> - {in p, forall z, ltT z t2 = (z == t1) || ltT z t1} -> - filter (ltT^~ t2) (filter p f) = filter (ltT^~ t1) (filter p f). -Proof. -move=>N C; apply: filterCN; first by rewrite mem_filter negb_and N. -by move=>z; rewrite mem_filter=>/andP [D _]; apply: C. -Qed. - -Lemma filter2CE (ltT : rel A) (p : pred A) f t1 t2 : - irreflexive ltT -> - antisymmetric ltT -> - transitive ltT -> - {in f &, total (fun k t => (k == t) || ltT k t)} -> - {in p, forall z, ltT z t2 = (z == t1) || ltT z t1} -> - uniq f -> - p t1 -> t1 \in f -> - filter (ltT^~ t2) - (filter p (sort (fun k t => (k == t) || ltT k t) f)) = - filter (ltT^~ t1) - (filter p (sort (fun k t => (k == t) || ltT k t) f)) ++ [:: t1]. -Proof. -move=>I Asym T Tot Z U P F; apply: filterCE (I) (T) _ _ _. -- by rewrite (sorted_filter T _ _) //; apply: sort_sorted_in_lt. -- by move=>z; rewrite mem_filter=>/andP [Pz _]; apply: Z. -by rewrite mem_filter mem_sort P F. -Qed. - -(* nth *) - -Lemma nth_cons (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 (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 (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 eqxx. -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 (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 (s : seq A) x0 k i : nth x0 s i = nth x0 (k::s) i.+1. -Proof. by []. Qed. - -Lemma nth_leT (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 (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 (s : seq A) ltT x0 k i j : - irreflexive ltT -> - transitive ltT -> - i <= size s -> - j <= size s -> - path ltT k s -> - ltT (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 (s : seq A) ltT x0 k z i : - irreflexive ltT -> - transitive ltT -> - path ltT k s -> - ltT (nth x0 (k::s) i) z -> - ltT 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 (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,eqxx,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. - -End SeqRel. - -(* there always exists a nat not in a given list *) -Lemma not_memX (ks : seq nat) : exists k, k \notin ks. -Proof. -have L a xs : foldl addn a xs = a + foldl addn 0 xs. -- elim: xs a=>[|z xs IH] //= a; first by rewrite addn0. - by rewrite add0n // [in LHS]IH [in RHS]IH addnA. -set k := foldl addn 1 ks. -suff K a : a \in ks -> a < k by exists k; apply/negP=>/K; rewrite ltnn. -rewrite {}/k; elim: ks=>[|k ks IH] //=; rewrite inE. -case/orP=>[/eqP ->|/IH]; first by rewrite L add1n addSn ltnS leq_addr. -rewrite L=>N; rewrite L; apply: leq_trans N _. -by rewrite addnAC leq_addr. -Qed. - -Section BigCat. -Context {A B : Type}. -Implicit Types (xs : seq A) (f : A -> seq B). - -Lemma flatten_map_big xs f : - flatten (map f xs) = \big[cat/[::]]_(x <- xs) f x. -Proof. -elim: xs=>/= [|x xs IH]; first by rewrite big_nil. -by rewrite big_cons IH. -Qed. - -Lemma size_big_cat xs f : - size (\big[cat/[::]]_(x <- xs) f x) = - \sum_(x <- xs) (size (f x)). -Proof. -elim: xs=>[|x xs IH] /=; first by rewrite !big_nil. -by rewrite !big_cons size_cat IH. -Qed. - -Lemma has_big_cat (p : pred B) xs f : - has p (\big[cat/[::]]_(x <- xs) f x) = - has (fun x => has p (f x)) xs. -Proof. -elim: xs=>[|x xs IH]; first by rewrite big_nil. -by rewrite big_cons has_cat /= IH. -Qed. - -End BigCat. - -Lemma big_cat_mem_has A (B : eqType) xs (f : A -> seq B) b : - b \in \big[cat/[::]]_(x <- xs) f x = - has (fun x => b \in f x) xs. -Proof. -rewrite -has_pred1 has_big_cat; apply: eq_has=>x. -by rewrite has_pred1. -Qed. - -(* big_cat_mem for A : Type *) -Lemma big_cat_memT A (B : eqType) x xs (f : A -> seq B) : - reflect (exists2 i, i \In xs & x \in f i) - (x \in \big[cat/[::]]_(i <- xs) f i). -Proof. by rewrite big_cat_mem_has; apply/hasPIn. Qed. - -(* big_cat_mem for A : eqType *) -Lemma big_cat_memE (A B : eqType) x xs (f : A -> seq B) : - reflect (exists2 i, i \in xs & x \in f i) - (x \in \big[cat/[::]]_(i <- xs) f i). -Proof. by rewrite big_cat_mem_has; apply: hasP. Qed. - -(* big_cat_mem for A : finType *) -Lemma big_cat_mem (A : finType) (B : eqType) (f : A -> seq B) x : - reflect (exists i, x \in f i) - (x \in \big[cat/[::]]_i f i). -Proof. -case: big_cat_memE=>H; constructor. -- by case: H=>i H1 H2; exists i. -by case=>i X; elim: H; exists i. -Qed. - -(* uniqueness for A : Type *) -Lemma uniq_big_catT A (B : eqType) xs (f : A -> seq B) : - Uniq xs -> - (forall x, x \In xs -> uniq (f x)) -> - (forall x1 x2, x1 \In xs -> x2 \In xs -> - has (mem (f x1)) (f x2) -> x1 = x2) -> - uniq (\big[cat/[::]]_(x <- xs) f x). -Proof. -elim: xs=>[|x xs IH] /=. -- by rewrite big_nil /=; constructor. -case=>Nx U H1 H2; rewrite big_cons cat_uniq. -apply/and3P; split. -- by apply: H1; left. -- rewrite has_big_cat -all_predC; apply/allPIn=>x3 H3 /=. - by apply: contra_notN Nx=>H; rewrite (H2 _ _ _ _ H) //; [left | right]. -apply: IH=>//. -- by move=>z Hz; apply: H1; right. -by move=>z1 z2 Hz1 Hz2 N; apply: H2=>//; right. -Qed. - -Lemma big_cat_uniq_pairwise A (B : eqType) xs (f : A -> seq B) x1 x2 : - uniq (\big[cat/[::]]_(x <- xs) f x) -> - x1 \In xs -> - x2 \In xs -> - has (mem (f x1)) (f x2) -> - x1 = x2. -Proof. -elim: xs=>[|x xs IH] //. -rewrite big_cons cat_uniq; case/and3P=>U N Us. -case/In_cons=>[->|H1]; case/In_cons=>[->|H2] //; last by apply: IH. -- case/hasP=>/= b Hb1 Hb2. - exfalso; move/negP: N; apply; apply/hasP. - by exists b=>//; apply/big_cat_memT; exists x2. -case/hasP=>/= b Hb1 Hb2. -exfalso; move/negP: N; apply. -apply/hasP; exists b=>//; apply/big_cat_memT. -by exists x1. -Qed. - -(* uniqueness for A : eqType *) -Lemma uniq_big_catE (A B : eqType) (f : A -> seq B) (xs : seq A) : - reflect - [/\ forall i, i \in xs -> uniq (f i), - forall i k, i \in xs -> k \in f i -> count_mem i xs = 1 & - forall i j k, i \in xs -> j \in xs -> - k \in f i -> k \in f j -> i = j] - (uniq (\big[cat/[::]]_(i <- xs) f i)). -Proof. -elim: xs=>[|x xs IH] /=; first by rewrite big_nil; constructor. -rewrite big_cons cat_uniq. -case H1 : (uniq (f x))=>/=; last first. -- by constructor; case=>/(_ x); rewrite inE eqxx H1=>/(_ erefl). -case: hasPn=>/= V; last first. -- constructor; case=>H2 H3 H4; elim: V=>z /big_cat_memE [i X Zi]. - apply/negP=>Zx; move/(H3 i z): (Zi); rewrite inE X orbT=>/(_ erefl). - move: (H4 x i z); rewrite !inE eqxx X orbT=>/(_ erefl erefl Zx Zi)=>E. - by rewrite -{i Zi}E eqxx add1n in X *; case=>/count_memPn; rewrite X. -case: IH=>H; constructor; last first. -- case=>H2 H3 H4; apply: H; split; last 1 first. - - by move=>i j k Xi Xj; apply: H4; rewrite inE ?Xi ?Xj orbT. - - by move=>i X; apply: H2; rewrite inE X orbT. - move=>i k X D; move/(H3 i k): (D); rewrite inE X orbT=>/(_ erefl). - case: (x =P i) X D=>[<-{i}|N] X D; last by rewrite add0n. - by rewrite add1n=>[[]] /count_memPn; rewrite X. -case: H=>H2 H3 H4; split; first by move=>i; rewrite inE=>/orP [/eqP ->|/H2]. -- move=>i k; rewrite inE eq_sym; case: (x =P i)=>[<- _|N /= Xi] K; last first. - - by rewrite add0n; apply: H3 Xi K. - rewrite add1n; congr S; apply/count_memPn; apply: contraL (K)=>X. - by apply/V/big_cat_memE; exists x. -move=>i j k; rewrite !inE=>/orP [/eqP ->{i}|Xi] /orP [/eqP ->{j}|Xj] Ki Kj //. -- by suff : k \notin f x; [rewrite Ki | apply/V/big_cat_memE; exists j]. -- by suff : k \notin f x; [rewrite Kj | apply/V/big_cat_memE; exists i]. -by apply: H4 Kj. -Qed. - -(* alternative formulation *) -Lemma uniq_big_catEX (A B : eqType) (f : A -> seq B) (xs : seq A) : - uniq xs -> - reflect - [/\ forall i, i \in xs -> uniq (f i) & - forall i j k, i \in xs -> j \in xs -> - k \in f i -> k \in f j -> i = j] - (uniq (\big[cat/[::]]_(i <- xs) f i)). -Proof. -move=>U; case: uniq_big_catE=>H; constructor; first by case: H. -by case=>H1 H2; elim: H; split=>// i k I K; rewrite count_uniq_mem // I. -Qed. - -(* uniqueness for A : finType *) -Lemma uniq_big_cat (A : finType) (B : eqType) (f : A -> seq B) : - reflect ([/\ forall t, uniq (f t) & - forall t1 t2 k, k \in f t1 -> k \in f t2 -> t1 = t2]) - (uniq (\big[cat/[::]]_tag f tag)). -Proof. -case: uniq_big_catEX=>[|[R1 R2]|R]. -- by rewrite /index_enum -enumT enum_uniq. -- by constructor; split=>[t|t1 t2 k K1 K2]; [apply: R1|apply: R2 K1 K2]. -constructor; case=>R1 R2; elim: R; split=>[t ?|t1 t2 k ??]; -by [apply: R1 | apply: R2]. -Qed. - -Lemma uniq_big_cat_uniqT A (B : eqType) xs (f : A -> seq B) y : - uniq (\big[cat/[::]]_(x <- xs) f x) -> - y \In xs -> - uniq (f y). -Proof. -elim: xs=>[|x xs IH] in y * => //. -rewrite big_cons InE cat_uniq. -case/and3P=>U _ Us; case=>[->|Hy] //. -by apply: IH. -Qed. - -Prenex Implicits uniq_big_cat_uniqT. - -Lemma uniq_big_cat_uniq (A : finType) (f : A -> seq nat) t : - uniq (\big[cat/[::]]_t f t) -> - uniq (f t). -Proof. by move/uniq_big_cat_uniqT; apply; apply/mem_seqP. Qed. - -Lemma uniq_big_cat_uniq0 (A : finType) (f : A -> seq nat) t : - uniq (0 :: \big[cat/[::]]_t f t) -> uniq (0 :: f t). -Proof. -case/andP=>U1 /uniq_big_cat_uniq /= U2. -rewrite U2 andbT (contra _ U1) // => U3. -by apply/big_cat_mem; exists t. -Qed. - -Lemma uniq_big_cat_disj (A : finType) (B : eqType) (f : A -> seq B) t1 t2 x : - uniq (\big[cat/[::]]_t f t) -> - x \in f t1 -> - x \in f t2 -> - t1 = t2. -Proof. by case/uniq_big_cat=>_; apply. Qed. - - diff --git a/core/seqperm.v b/core/seqperm.v deleted file mode 100644 index 962d1a7..0000000 --- a/core/seqperm.v +++ /dev/null @@ -1,253 +0,0 @@ -(* -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 pcm Require Import options pred. - -(****************************************************) -(* 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. - -Lemma pperm_rcons s1 s2 x : - perm s1 s2 <-> - perm (rcons s1 x) (rcons s2 x). -Proof. -rewrite -!cats1; split=>[|H]; first by apply: pperm_cat2rL. -by rewrite -[s1]cats0 -[s2]cats0 -(pperm_cat_cons x). -Qed. - -Lemma pperm_rev s : perm s (rev s). -Proof. -elim: s=>[|x xs IH] //=. -rewrite rev_cons -cats1. -apply: (@pperm_trans ([:: x] ++ rev xs)); last by apply: pperm_catC. -by rewrite pperm_cons. -Qed. - -End Permutations. - -#[export] 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/core/slice.v b/core/slice.v deleted file mode 100644 index 47958bb..0000000 --- a/core/slice.v +++ /dev/null @@ -1,727 +0,0 @@ -(* -Copyright 2022 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 path interval order. -From mathcomp Require Import fintype finfun tuple. -From pcm Require Import options prelude seqext. - -Open Scope order_scope. -Import Order.Theory. - -(* DEVCOMMENT *) -(* intervals used to do these simplifications automatically *) -(* but that changed with mathcomp 2.0 *) -(* /DEVCOMMENT *) -Section BSimp_Extension. -Context disp (T : porderType disp). -Implicit Types (x y : T) (b c : bool). - -Lemma binf_inf b c : (BInfty T b == BInfty T c) = (b == c). -Proof. by rewrite eqE /= eqE /= andbT. Qed. - -Lemma bside_inf x b c : - ((BSide b x == BInfty T c) = false) * - ((BInfty T c == BSide b x) = false). -Proof. by rewrite !eqE /= !eqE /= !andbF. Qed. - -Lemma bside_side x y b c : - (BSide b x == BSide c y) = (b == c) && (x == y). -Proof. by rewrite !eqE. Qed. - -End BSimp_Extension. - -Definition bnd_simp := ((@bside_inf, @binf_inf, @bside_side), bnd_simp). - -(* sequence slicing by nat indices *) -(* reuses mathcomp interval infrastructure *) - -(* convert bound to nat *) -(* maps -oo -> 0, +oo -> m *) -Definition bnd (i : itv_bound nat) (m : nat) : nat := - match i with - | BSide b n => n + ~~ b - | BInfty b => if b then 0 else m - end. - -(* slicing is applying an interval to a seq *) -Definition slice {A : Type} (s : seq A) (i : interval nat) := - let sz := size s in - let: Interval l u := i in - drop (bnd l sz) (take (bnd u sz) s). - -Arguments slice {A} s i : simpl never. - -Notation "&: s i" := (slice s i) - (at level 1, i at next level, s at next level, - format "'&:' s i") : fun_scope. - -(* subtract n from bound, convert values past zero to -oo *) -Definition shl_bnd (i : itv_bound nat) (n : nat) := - match i with - | BSide b m => if (m < n)%N then -oo else BSide b (m - n) - | BInfty b => BInfty _ b - end. - -Lemma shl_bnd0 i : shl_bnd i 0%N = i. -Proof. by case: i=>[[] i|[]] //=; rewrite subn0. Qed. - -Lemma shl_bnd_add i n m : shl_bnd (shl_bnd i n) m = shl_bnd i (n+m). -Proof. -case: i=>[[] i|[]] //=. -- case: ltnP=>Hin /=; first by rewrite ltn_addr. - by rewrite ltn_subLR // subnDA. -case: ltnP=>Hin /=; first by rewrite ltn_addr. -by rewrite ltn_subLR // subnDA. -Qed. - -Definition shl_itv (i : interval nat) (n : nat) := - let: Interval l u := i in - Interval (shl_bnd l n) (shl_bnd u n). - -Lemma shl_itv0 i : shl_itv i 0%N = i. -Proof. by case: i=>i j /=; rewrite !shl_bnd0. Qed. - -Lemma shl_itv_add i n m : shl_itv (shl_itv i n) m = shl_itv i (n+m). -Proof. by case: i=>i j /=; rewrite !shl_bnd_add. Qed. - -Lemma mem_shl n m i : - (n \in shl_itv i m) = (m + n \in i). -Proof. -rewrite !in_itv; case: i=>i j /=. -case: i=>[l i|[]]; case: j=>[r j|[]] //=; -rewrite ?andbF ?andbT //. -- case: (ltnP j m)=>/= Hj. - - rewrite andbF; symmetry; apply/negbTE; rewrite negb_and lteifNE negbK -lteifNE. - by apply/orP; right; apply/lteifS/ltn_addr. - have ->: (n < j - m ?<= if ~~ r) = (m + n < j ?<= if ~~ r). - - case: r=>/=; first by rewrite ltEnat /= ltn_subRL. - by rewrite leEnat leq_subRL. - case: (m + n < j ?<= if ~~ r); rewrite ?andbF // !andbT. - case: (ltnP i m)=>/= Hi. - - by symmetry; apply/lteifS/ltn_addr. - case: l=>/=; last by rewrite ltEnat /= ltn_subLR. - by rewrite leEnat leq_subLR. -- case: (ltnP i m)=>/= Hi. - - by symmetry; apply/lteifS/ltn_addr. - case: l=>/=; last by rewrite ltEnat /= ltn_subLR. - by rewrite leEnat leq_subLR. -case: (ltnP j m)=>/= Hj. -- symmetry; apply/negbTE; rewrite lteifNE negbK. - by apply/lteifS/ltn_addr. -case: r=>/=; first by rewrite ltEnat /= ltn_subRL. -by rewrite leEnat leq_subRL. -Qed. - -(* generalize to orderType? *) -Lemma mem_itv_inf (n : nat) : n \in `]-oo,+oo[. -Proof. by rewrite in_itv. Qed. - -Section Lemmas. -Variable (A : Type). -Implicit Type (s : seq A). - -Lemma slice_decompose s l r x y : - &:s (Interval (BSide l x) (BSide r y)) = - &:(&:s (Interval -oo (BSide r y))) (Interval (BSide l x) +oo). -Proof. by rewrite /slice /= drop0 take_size. Qed. - -(* "test" lemmas *) -Lemma slice_ouou s a b : - &:s `]a,b[ = &:(&:s `]-oo,b[) `]a,+oo[. -Proof. by rewrite slice_decompose. Qed. - -(* TODO generalize? *) -Lemma slice_uouo s a b : - &:s `]a,b+a] = &:(&:s `]a,+oo[) `]-oo,b[. -Proof. -by rewrite /slice /= !addn1 !addn0 drop0 take_size take_drop addnS. -Qed. - -(* ... *) - -(* some simplifying equalities on slices *) - -Lemma slice_0L s y : &:s (Interval -oo y) = &:s (Interval (BLeft 0%N) y). -Proof. by rewrite /slice /= addn0. Qed. - -Lemma slice_FR s x : &:s (Interval x +oo) = &:s (Interval x (BLeft (size s))). -Proof. by rewrite /slice /= addn0. Qed. - -Lemma slice_uu s : &:s `]-oo, +oo[ = s. -Proof. by rewrite /slice /= drop0 take_size. Qed. - -Lemma itv_underL s (i j : itv_bound nat) : - bnd i (size s) = 0%N -> - &:s (Interval i j) = &:s (Interval -oo j). -Proof. by move=>Hi; rewrite /slice /= Hi drop0. Qed. - -Lemma itv_underR s (i j : itv_bound nat) : - bnd j (size s) = 0%N -> - &:s (Interval i j) = [::]. -Proof. by move=>Hj; rewrite /slice /= Hj take0. Qed. - -Corollary itv_under0R s (i : itv_bound nat) : - &:s (Interval i (BLeft 0%N)) = [::]. -Proof. by apply: itv_underR. Qed. - -Lemma itv_overL s (i j : itv_bound nat) : - size s <= bnd i (size s) -> - &:s (Interval i j) = [::]. -Proof. -move=>Hi /=; rewrite /slice /=; apply: drop_oversize. -rewrite size_take; case: ifP=>// H. -by apply/ltnW/(leq_trans H). -Qed. - -Lemma itv_overR s (i j : itv_bound nat) : - size s <= bnd j (size s) -> - &:s (Interval i j) = &:s (Interval i +oo). -Proof. by move=>Hj; rewrite /slice /= take_size take_oversize. Qed. - -Lemma itv_minfR s (i : itv_bound nat) : - &:s (Interval i -oo) = [::]. -Proof. by rewrite /slice /= take0. Qed. - -Lemma itv_pinfL s (j : itv_bound nat) : - &:s (Interval +oo j) = [::]. -Proof. -rewrite /slice /=; apply: drop_oversize. -by rewrite size_take_min; apply: geq_minr. -Qed. - -(* TODO unify the next two? *) -Lemma itv_swapped_bnd s (i j : itv_bound nat) : - j <= i -> - &:s (Interval i j) = [::]. -Proof. -move=>H; rewrite /slice; apply: drop_oversize. -move: H; case: i=>[ib ix|[]]; case: j=>[jb jx|[]] //=; -rewrite ?bnd_simp ?take_size ?take0 ?drop_size //= size_take_min. -- rewrite leBSide; case: ib=>/=; case: jb=>/=; rewrite ?addn0 ?addn1=>Hji. - - apply/leq_trans/Hji; exact: geq_minl. - - apply/leq_trans/Hji; exact: geq_minl. - - apply: leq_trans; first by exact: geq_minl. - by apply: ltnW. - by apply: leq_trans; first by exact: geq_minl. -by move=>_; exact: geq_minr. -Qed. - -Lemma itv_swapped s (i j : itv_bound nat) : - bnd j (size s) <= bnd i (size s) -> - &:s (Interval i j) = [::]. -Proof. -move=>H; rewrite /slice; apply: drop_oversize. -rewrite size_take_min; apply/leq_trans/H. -by exact: geq_minl. -Qed. - -Corollary slice_usize s : s = &:s `]-oo, size s[. -Proof. by rewrite itv_overR /=; [rewrite slice_uu | rewrite addn0]. Qed. - -(* slice size *) - -Lemma slice_size s (i : interval nat) : - size (&:s i) = minn (bnd i.2 (size s)) (size s) - bnd i.1 (size s). -Proof. -rewrite /slice; case: i=>[[l i|[]][r j|[]]] //=; rewrite ?take0 ?drop0 ?take_size ?drop_size ?minnn ?min0n ?subn0 //=. -- by rewrite size_drop size_take_min. -- by rewrite size_drop. -- by rewrite size_take_min. -- by rewrite size_drop size_take_min. -by rewrite subnn. -Qed. - -(* splitting slice *) - -Lemma slice_split_bnd s (i : interval nat) (x : itv_bound nat) : - i.1 <= x <= i.2 -> - &:s i = &:s (Interval i.1 x) ++ &:s (Interval x i.2). -Proof. -case: i=>i1 i2 /=. -case: x=>[b x|[]]; rewrite ?bnd_simp /=; first last. -- by move/eqP=>->; rewrite itv_pinfL cats0. -- by rewrite andbT =>/eqP->; rewrite itv_minfR. -case/boolP: (size s <= bnd (BSide b x) (size s)). -- move=>H Hx; rewrite (itv_overL _ H) cats0 (itv_overR _ H). - apply: itv_overR; rewrite /= in H. - case/andP: Hx=>_. - case: i2=>[[] i2|[]] //= Hx; apply: (leq_trans H); case: {H}b Hx; rewrite bnd_simp ?addn0 ?addn1 //. - by move=>Hx; apply/ltnW. -rewrite -ltNge ltEnat /= =>/ltnW/minn_idPl Hx. -case: i1=>[l i|[]]; case: i2=>[r j|[]] //=; -rewrite ?bnd_simp ?andbF ?andbT // ?lteBSide /slice /= ?drop0 ?take_size. -- case/andP=>Hi Hj. - have Hxj : (x + ~~ b <= j + ~~r)%N. - - case: r Hj=>/=; rewrite ?leEnat ?ltEnat; case: {Hx Hi}b=>/=; - rewrite ?addn0 ?addn1 // => Hj; - by apply: ltnW. - have Hxi : (i + ~~ l <= x + ~~ b)%N. - - case: l Hi=>/=; rewrite ?implybT ?implybF /= ?addn0 ?addn1 /= => Hi. - - by apply/(leq_trans Hi)/leq_addr. - by case: {Hx Hj Hxj}b Hi=>/=; rewrite ?leEnat ?ltEnat /= ?addn0 ?addn1. - rewrite -{1}(subnKC Hxj) takeD drop_cat size_take_min Hx take_drop subnK //. - rewrite ltn_neqAle Hxi andbT; case: eqP=>//= ->. - by rewrite subnn drop0 drop_take_id. -- move=>Hi; rewrite -{1}(cat_take_drop (x + ~~ b) s) drop_cat size_take_min Hx. - have Hxi : (i + ~~ l <= x + ~~ b)%N. - - case: l Hi=>/=; rewrite ?implybT ?implybF /= ?addn0 ?addn1 /= => Hi. - - by apply/(leq_trans Hi)/leq_addr. - by case: {Hx}b Hi=>/=; rewrite ?leEnat ?ltEnat /= ?addn0 ?addn1. - rewrite ltn_neqAle Hxi andbT; case: eqP=>//= ->. - by rewrite subnn drop0 drop_take_id. -- move=>Hj. - have Hbj : (x + ~~ b <= j + ~~ r)%N. - - case: r Hj=>/=; rewrite ?leEnat ?ltEnat; case: {Hx}b=>/=; - rewrite ?addn0 ?addn1 // => Hx; - by apply: ltnW. - by rewrite -{1}(subnKC Hbj) takeD take_drop subnK. -by move=>_; rewrite cat_take_drop. -Qed. - -Corollary slice_split s (i : interval nat) b (x : nat) : - x \in i -> - &:s i = &:s (Interval i.1 (BSide b x)) ++ &:s (Interval (BSide b x) i.2). -Proof. -move=>Hx; apply: slice_split_bnd. -case: i Hx=>[[l i|[]][r j|[]]] //=; -rewrite in_itv /= ?andbT ?andbF // ?leBSide; case: b=>/=; -rewrite ?implybF ?implybT //=. -- by case/andP=>->/= /lteifW. -- by case/andP=>/lteifW->. -- by move/lteifW. -by move/lteifW. -Qed. - -(* splitting whole list *) - -Corollary slice_split_full s b (x : nat) : - s = &:s (Interval -oo (BSide b x)) ++ &:s (Interval (BSide b x) +oo). -Proof. by rewrite -[LHS](slice_uu s); apply: slice_split. Qed. - -(* slice extrusion *) - -Lemma slice_extrude s (i : interval nat) : - i.1 < i.2 -> - s = &:s (Interval -oo i.1) ++ &:s i ++ &:s (Interval i.2 +oo). -Proof. -case: i=>[[[] i|[]][[] j|[]]] //=; rewrite bnd_simp=>H; - rewrite ?itv_minfR ?itv_pinfL /= ?cats0. -- by rewrite -{1}(slice_uu s) (slice_split _ true (x:=i)) //= - (slice_split _ true (x:=j) (i:=`[i, +oo[)) //= in_itv /= andbT; apply/ltnW. -- by rewrite -{1}(slice_uu s) (slice_split _ true (x:=i)) //= - (slice_split _ false (x:=j) (i:=`[i, +oo[)) //= in_itv /= andbT. -- by rewrite -{1}(slice_uu s) (slice_split _ true (x:=i)). -- by rewrite -{1}(slice_uu s) (slice_split _ false (x:=i)) //= - (slice_split _ true (x:=j) (i:=`]i, +oo[)) //= in_itv /= andbT. -- by rewrite -{1}(slice_uu s) (slice_split _ false (x:=i)) //= - (slice_split _ false (x:=j) (i:=`]i, +oo[)) //= in_itv /= andbT. -- by rewrite -{1}(slice_uu s) (slice_split _ false (x:=i)). -- by rewrite -{1}(slice_uu s) (slice_split _ true (x:=j)). -- by rewrite -{1}(slice_uu s) (slice_split _ false (x:=j)). -by rewrite slice_uu. -Qed. - -(* "test" lemmas *) -Corollary slice_uxou s i : s = &:s `]-oo, i] ++ &:s `]i, +oo[. -Proof. by rewrite -(slice_split _ _ (i:=`]-oo, +oo[)) // slice_uu. Qed. - -Corollary slice_uoxu s i : s = &:s `]-oo, i[ ++ &:s `[i, +oo[. -Proof. by rewrite -(slice_split _ _ (i:=`]-oo, +oo[)) // slice_uu. Qed. - -Corollary slice_uxxo s a b : - a <= b -> - &:s `]-oo, b] = &:s `]-oo, a] ++ &:s `]a, b]. -Proof. by move=>Hab; rewrite (slice_split _ false (x:=a)). Qed. - -Corollary slice_uxox s a b : - a <= b -> - &:s `]-oo, b] = &:s `]-oo, a[ ++ &:s `[a, b]. -Proof. by move=>Hab; rewrite (slice_split _ true (x:=a)). Qed. - -Corollary slice_uoox (s : seq A) a b : - a < b -> - &:s `]-oo, b[ = &:s `]-oo, a[ ++ &:s `[a, b[. -Proof. by move=>Hab; rewrite (slice_split _ true (x:=a)). Qed. - -(* ... *) - -(* singleton slice *) - -Lemma slice_kk s k l r : - &:s (Interval (BSide l k) (BSide r k)) = - if l && ~~ r then oapp (cons^~ [::]) [::] (onth s k) else [::]. -Proof. -rewrite /slice /=; case: (onth_sizeP s k)=>[|v] H; rewrite H /=. -- move/size_onthPn: H=>H; rewrite if_same. - apply: drop_oversize; rewrite size_take ltnNge. - have ->/= : (size s <= k + ~~ r)%N by apply/leq_trans/leq_addr. - by apply/leq_trans/leq_addr. -move: (onth_size H)=>Hk; case: l; case: r=>/=; -rewrite ?addn0 ?addn1. -- by apply: drop_oversize; rewrite size_take Hk. -- rewrite -addn1 addnC -take_drop. - rewrite (take_nth v); last by rewrite size_drop subn_gt0. - by rewrite take0 /= nth_drop addn0 nth_onth H. -- by apply: drop_oversize; rewrite size_take Hk. -apply: drop_oversize; rewrite size_take; case: ifP=>// /negbT. -by rewrite -leqNgt. -Qed. - -(* "test" lemmas *) -Corollary slice_kkxx s k : - &:s `[k, k] = oapp (cons^~ [::]) [::] (onth s k). -Proof. by rewrite slice_kk. Qed. - -Corollary slice_kkxo s k : &:s `[k,k[ = [::]. -Proof. by rewrite slice_kk. Qed. - -Corollary slice_kkox s k : &:s `]k,k] = [::]. -Proof. by rewrite slice_kk. Qed. - -Corollary slice_kkoo s k : &:s `]k,k[ = [::]. -Proof. by rewrite slice_kk. Qed. - -(* endpoint +1 *) - -Lemma slice_oSL x a s : - &:s (Interval (BLeft x.+1) a) = &:s (Interval (BRight x) a). -Proof. -case/boolP: (x.+1 \in Interval (BRight x) a)=>[H|]. -- rewrite [RHS](slice_split _ true (x:=x.+1)) //=. - suff ->: &:s `]x, x.+1[ = [::] by rewrite cat0s. - by rewrite /slice /= addn0 addn1 drop_take_id. -rewrite in_itv /= negb_and ltEnat /= ltnS leqnn /=. -case: a=>[[] a|[]] //=. -- by rewrite -leNgt leEnat => H; rewrite !itv_swapped //= !addn0 ?addn1 leEnat. -- by rewrite -ltNge ltEnat /= => H; rewrite !itv_swapped //= !addn1 ?addn0 leEnat ltnS. -by move=>_; rewrite !itv_minfR. -Qed. - -Lemma slice_oSR a x s : - &:s (Interval a (BLeft x.+1)) = &:s (Interval a (BRight x)). -Proof. -case/boolP: (x \in Interval a (BLeft x.+1))=>[H|]. -- rewrite (slice_split _ false (x:=x)) //=. - suff ->: &:s `]x, x.+1[ = [::] by rewrite cats0. - by rewrite /slice /= addn0 addn1 drop_take_id. -rewrite in_itv /= negb_and ltEnat /= ltnS leqnn /= orbF. -case: a=>[[] a|[]] //=. -- by rewrite -ltNge ltEnat /= => H; rewrite !itv_swapped //= !addn0 ?addn1 leEnat. -- by rewrite -leNgt leEnat => H; rewrite !itv_swapped //= !addn1 ?addn0 leEnat ltnS. -by move=>_; rewrite !itv_pinfL. -Qed. - -(* endpoint -1 *) - -Corollary slice_oPL a x s : - &:s (Interval (BRight x.-1) a) = if 0 < x - then &:s (Interval (BLeft x) a) - else &:s (Interval (BRight x) a). -Proof. by case: x=>//= n; rewrite slice_oSL. Qed. - -Corollary slice_oPR a x s : - &:s (Interval a (BRight x.-1)) = if 0 < x - then &:s (Interval a (BLeft x)) - else &:s (Interval a (BRight x)). -Proof. by case: x=>//= n; rewrite slice_oSR. Qed. - -(* endpoint split *) - -Lemma slice_xR a x s : - a <= BLeft x -> - &:s (Interval a (BRight x)) = - oapp (rcons (&:s (Interval a (BLeft x)))) - (&:s (Interval a +oo)) - (onth s x). -Proof. -move=>Hax; rewrite (slice_split _ true (x:=x)) /=; last first. -- rewrite in_itv /= lexx andbT. - by case: a Hax=>/=[ax av|ax]; case: ax. -rewrite slice_kk /=; case: (onth_sizeP s x)=>[|v] H; - rewrite H /=; last by rewrite cats1. -rewrite cats0 itv_overR //= addn0. -by apply/size_onthPn. -Qed. - -Lemma slice_xL b x s : - BRight x <= b -> - &:s (Interval (BLeft x) b) = - oapp (cons^~ (&:s (Interval (BRight x) b))) - [::] - (onth s x). -Proof. -move=>Hxb; rewrite (slice_split _ false (x:=x)) /=; last first. -- rewrite in_itv /= lexx /=. - by case: b Hxb=>/=[bx bv|bx]; case: bx. -rewrite slice_kk /=; case: (onth_sizeP s x)=>[|v] H; rewrite H //=. -rewrite itv_overL //= addn1; apply: ltW. -by apply/size_onthPn. -Qed. - -(* slice of empty list *) - -Lemma slice0 i : - &:([::] : seq A) i = [::]. -Proof. -by rewrite /slice /=; case: i=>[[[] av|[]]]; case=>[[] bv|[]]. -Qed. - -(* slice of one-element list *) - -Lemma slice1 (k : A) i : - &:[::k] i = if 0%N \in i then [::k] else [::]. -Proof. -rewrite /slice in_itv /=. -case: i=>[[[] i|[]][[] j|[]]] //=; -rewrite ?drop0 ?addn0 ?addn1 ?andbF ?andbT //=. -- case: j=>[|j]/=; first by rewrite ltxx andbF. - by rewrite andbT; case: i. -- by case: i. -- by case: i. -- by apply: drop_oversize; rewrite size_take /=; case: ifP=>//; case: j. -- by case: j. -by case: j. -Qed. - -(* slice and constructors *) - -Lemma slice_cat s1 s2 i : - &:(s1++s2) i = &:s1 i ++ &:s2 (shl_itv i (size s1)). -Proof. -rewrite /slice; case: i=>[[l i|[]][r j|[]]] /=; -rewrite ?take0 ?drop0 ?take_size ?drop_size // ?size_cat. -- rewrite take_cat; case: ltnP=>Hj. - - have ->/=: (j < size s1)%N by apply/leq_ltn_trans/Hj/leq_addr. - by rewrite take0 /= cats0. - rewrite (take_oversize (s:=s1)) //. - move: Hj; rewrite leq_eqVlt; case/orP=>[/eqP Ej|Hj]. - - rewrite -Ej subnn take0 cats0 /=. - case: r Ej=>/=. - - by rewrite addn0=>->; rewrite ltnn subnn /= take0 /= cats0. - by rewrite addn1=>->; rewrite !ltnS leqnn /= take0 /= cats0. - rewrite (ltnNge j); have H1: (size s1 <= j)%N. - - by case: r Hj=>/=; rewrite ?addn1 // addn0 => /ltnW. - rewrite H1 drop_cat; case: ltnP=>Hi /=. - - have ->/=: (i < size s1)%N by apply/leq_ltn_trans/Hi/leq_addr. - by rewrite drop0 addnBAC. - rewrite (drop_oversize (s:=s1)) //=. - move: Hi; rewrite leq_eqVlt; case/orP=>[/eqP Ei|Hi]. - - rewrite -Ei subnn drop0 /=. - case: l Ei=>/=. - - by rewrite addn0=><-; rewrite ltnn subnn /= drop0 addnBAC. - by rewrite addn1=><-; rewrite leqnn /= drop0 addnBAC. - rewrite (ltnNge i); have H2: (size s1 <= i)%N. - - by case: l Hi=>/=; rewrite ?addn1 // addn0 => /ltnW. - by rewrite H2 /= !addnBAC. -- rewrite drop_cat; case: ltnP=>Hi /=. - - have ->/=: (i < size s1)%N by apply/leq_ltn_trans/Hi/leq_addr. - by rewrite drop0. - rewrite (drop_oversize (s:=s1)) //=. - move: Hi; rewrite leq_eqVlt; case/orP=>[/eqP Ei|Hi]. - - rewrite -Ei subnn drop0 /=. - case: l Ei=>/=. - - by rewrite addn0=><-; rewrite ltnn subnn /= drop0. - by rewrite addn1=><-; rewrite leqnn /= drop0. - rewrite (ltnNge i); have H2: (size s1 <= i)%N. - - by case: l Hi=>/=; rewrite ?addn1 // addn0 => /ltnW. - by rewrite H2 /= addnBAC. -- rewrite take_cat; case: ltnP=>Hj. - - have ->/=: (j < size s1)%N by apply/leq_ltn_trans/Hj/leq_addr. - by rewrite take0 /= cats0. - rewrite (take_oversize (s:=s1)) //. - move: Hj; rewrite leq_eqVlt; case/orP=>[/eqP Ej|Hj]. - - rewrite -Ej subnn take0 cats0 /=. - case: r Ej=>/=. - - by rewrite addn0=>->; rewrite ltnn subnn /= take0 /= cats0. - by rewrite addn1=>->; rewrite !ltnS leqnn /= take0 /= cats0. - rewrite (ltnNge j); have H1: (size s1 <= j)%N. - - by case: r Hj=>/=; rewrite ?addn1 // addn0 => /ltnW. - by rewrite H1 /= addnBAC. -by rewrite !drop_oversize //= size_take_min ?size_cat; apply: geq_minr. -Qed. - -Lemma slice_cat_piecewise s1 s2 i : - &:(s1++s2) i = if size s1 \in i - then &:s1 (Interval i.1 +oo) ++ &:s2 (Interval -oo (shl_bnd i.2 (size s1))) - else if bnd i.2 (size s1 + size s2) <= size s1 - then &:s1 i - else &:s2 (shl_itv i (size s1)). -Proof. -rewrite slice_cat; case: i=>i j /=; case: ifP. -- rewrite in_itv; case/andP=>Hi Hj. - rewrite (itv_overR _ (j:=j)); last first. - - case: j Hj=>[[] j|[]] //=. - - by rewrite addn0; move/ltnW. - by rewrite addn1 leEnat; move/(ltnW (n:=j.+1)). - rewrite (itv_underL (i:=shl_bnd i _)) //. - case: i Hi=>[[] i|[]] //=; last by rewrite ltEnat /=; move=>->. - rewrite leEnat leq_eqVlt; case/orP=>[/eqP->|->] //=. - by rewrite ltnn /= subnn. -rewrite in_itv=>/negbT; rewrite negb_and=>H. -case: ifP=>Hj. -- rewrite (itv_underR (s:=s2)); first by rewrite cats0. - case: {H}j Hj=>[[] j|[]] //=. - - rewrite addn0 leEnat leq_eqVlt; case/orP=>[/eqP->|->] //=. - by rewrite ltnn /= subnn. - - by rewrite addn1 leEnat =>->. - by rewrite leEnat -{2}(addn0 (size s1)) leq_add2l leqn0 => /eqP. -case/orP: H; last first. -- case: j Hj=>[[] j|[]] //=. - - by rewrite addn0 -leNgt=>->. - by rewrite leEnat addn1 -ltnNge =>->. -case: i=>[[] i|[]] //=; last by rewrite !itv_pinfL. -- rewrite leEnat -ltnNge => Hi. - rewrite (ltnNge i) (ltnW Hi) /= itv_overL //= addn0 leEnat. - by apply: ltnW. -rewrite ltEnat /= -leqNgt => Hi. -rewrite ltnNge Hi /= itv_overL //= addn1 leEnat. -by apply: ltnW. -Qed. - -Lemma slice_cons x s i : - &:(x::s) i = if 0%N \in i then x::&:s (shl_itv i 1) else &:s (shl_itv i 1). -Proof. by rewrite -cat1s slice_cat /= slice1; case: ifP. Qed. - -Lemma slice_rcons x s i : - &:(rcons s x) i = if size s \in i then rcons (&:s i) x else &:s i. -Proof. -rewrite -cats1 slice_cat slice1 mem_shl addn0. -by case: ifP=>_; [rewrite cats1 | rewrite cats0]. -Qed. - -(* mask *) - -Lemma slice_mask s i : - &:s i = let x := bnd i.1 (size s) in - let y := bnd i.2 (size s) in - mask (nseq x false ++ nseq (y-x) true) s. -Proof. by rewrite /slice /=; case: i=>i j /=; rewrite drop_take_mask. Qed. - -(* count *) - -Lemma slice_count p s i : - count p (&:s i) = - count (fun j => j \in i) (findall p s). -Proof. -elim: s i=>/= [|x s IH] i; first by rewrite slice0. -rewrite findall_cons slice_cons. -case/boolP: (p x)=>/= Hpx; case/boolP: (0 \in i)=>I0 /=. -- rewrite Hpx !add1n; congr S. - rewrite IH count_map; apply: eq_count=>z /=. - by rewrite mem_shl add1n. -- rewrite IH /= count_map; apply: eq_count=>z /=. - by rewrite mem_shl add1n. -- rewrite (negbTE Hpx) /= IH /= count_map; apply: eq_count=>z /=. - by rewrite mem_shl add1n. -rewrite IH /= count_map; apply: eq_count=>z /=. -by rewrite mem_shl add1n. -Qed. - -(* has *) - -Corollary slice_has p s i : - has p (&:s i) = - has (fun j => j \in i) (findall p s). -Proof. by rewrite !has_count slice_count. Qed. - -End Lemmas. - -(* map *) - -Lemma slice_map {A B} (f : A -> B) i s : - [seq f x | x <- &:s i] = &: [seq f x | x <- s] i. -Proof. by rewrite !slice_mask /= map_mask /= size_map. Qed. - -Section LemmasEq. -Variable (A : eqType). -Implicit Type (s : seq A). - -(* membership *) - -Corollary slice_memE x s i : - x \in &:s i = - has (fun j => j \in i) (indexall x s). -Proof. by rewrite /indexall -has_pred1; apply: slice_has. Qed. - -Corollary slice_memE1 x s i : - (count_mem x s <= 1)%N -> - x \in &:s i = - (x \in s) && (bnd i.1 (size s) <= index x s < bnd i.2 (size s)). -Proof. -move=>H; rewrite slice_memE indexall_count1 //; case: ifP=>//= Hx; rewrite orbF. -by case: i=>/= [[[] i|[]][[] j|[]]]; rewrite in_itv /= ?ltEnat ?leEnat /= - ?ltn0 ?addn0 ?addn1 ?andbT ?andbF // ?(leqNgt (size _)) index_mem Hx //= andbT. -Qed. - -Corollary slice_uniq_memE x s i : - uniq s -> - x \in &:s i = - (x \in s) && (bnd i.1 (size s) <= index x s < bnd i.2 (size s)). -Proof. by move=>U; apply: slice_memE1; rewrite (count_uniq_mem _ U) leq_b1. Qed. - -(* subset *) - -Lemma slice_subset s i1 i2 : - i1 <= i2 -> - {subset (&:s i1) <= &:s i2}. -Proof. -case: i1=>i1 j1; case: i2=>i2 j2. -rewrite subitvE; case/andP=>Hi Hj. -move=>x Hx. -have Hij : i1 <= j1. -- apply: contraLR Hx; rewrite -ltNge=>/ltW Hji. - by rewrite itv_swapped_bnd. -rewrite (@slice_split_bnd _ _ _ i1) /=; last first. -- by rewrite Hi /=; apply/le_trans/Hj. -rewrite mem_cat; apply/orP; right. -rewrite (@slice_split_bnd _ _ _ j1) /=; last first. -- by rewrite Hj andbT. -by rewrite mem_cat Hx. -Qed. - -Corollary slice_subset_full s i : - {subset &:s i <= s}. -Proof. -by rewrite -{2}(slice_uu s); apply/slice_subset/itv_lex1. -Qed. - -(* slicing preserves uniqueness *) - -Lemma slice_uniq s i : - uniq s -> uniq (&:s i). -Proof. -move=>U; rewrite /slice /=; case: i=>l u. -by apply/drop_uniq/take_uniq. -Qed. - -(* slicing preserves sortedness *) - -Lemma slice_sorted r s i : - sorted r s -> sorted r (&:s i). -Proof. -move=>S; rewrite /slice /=; case: i=>l u. -by apply/drop_sorted/take_sorted. -Qed. - -End LemmasEq. - -(* slicing and rcons *) -Lemma codom_ux_rcons A n (f : {ffun 'I_n -> A}) (i : 'I_n) : - &:(fgraph f) `]-oo, i : nat] = - rcons (&:(fgraph f) `]-oo, i : nat[) (f i). -Proof. by rewrite slice_xR // slice_uu onth_codom. Qed. - - - - diff --git a/core/uconsec.v b/core/uconsec.v deleted file mode 100644 index fdfc444..0000000 --- a/core/uconsec.v +++ /dev/null @@ -1,1105 +0,0 @@ -(* -Copyright 2022 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 path interval order. -From pcm Require Import options prelude ordtype seqext slice useqord uslice. -Local Open Scope order_scope. -Import Order.Theory. - -(* We assume the sequences are unique and most lemmas do require this *) -(* condition explicitly. Should it be added to `consec` itself? *) - -(**************************************) -(**************************************) -(* Consecutive elements in a sequence *) -(**************************************) -(**************************************) - -Section ConsecEq. -Variable (A : eqType). -Implicit Type (ks : seq A). - -(* The interval lemmas let us equate evaluations of interval endpoints *) -(* But, when a property we want to prove involves other components *) -(* we need to properly induct over ks. *) -(* We thus first define what it means for elements in ks to be consecutive, *) -(* so that the properties can be used in the proofs of the inductive cases. *) - -(* t2 is consecutive to t1 in ks if it appears after t1 and there are *) -(* no other elements in ks between t1 and t2 *) - -Definition consec ks t1 t2 := - [&& t1 <[ks] t2 & nilp (&=ks `]t1, t2[)]. - -(* several alternative formulations *) -Lemma consecP ks t1 t2 : - reflect (t1 <[ks] t2 /\ ~~ has predT (&=ks `]t1, t2[)) - (consec ks t1 t2). -Proof. -apply: (iffP andP); case=>-> H; split=>//. -- by rewrite -nilp_hasPn. -by rewrite nilp_hasPn. -Qed. - -Lemma consecP_inlt (ks : seq A) t1 t2 : - uniq ks -> - reflect ([/\ t1 \in ks & {in ks, forall z, z <[ks] t2 = z <=[ks] t1}]) - (consec ks t1 t2). -Proof. -move=>U; apply: (iffP (consecP ks t1 t2)); case=>H1 H2; split. -- by apply: slt_memE H1. -- by move=>z Z; apply: (hasNL_oo U H1 H2 Z). -- by rewrite H2 // sle_refl. -by apply: (hasNR_oo U)=>z /H2. -Qed. - -Lemma consecP_ingt (ks : seq A) t1 t2 : - uniq ks -> - reflect ([/\ t1 \in ks & {in ks, forall z, t2 <=[ks] z = t1 <[ks] z}]) - (consec ks t1 t2). -Proof. -move=>U; apply (iffP (consecP_inlt t1 t2 U)); case=>H1 H2; - split=>// z /H2 E. -- by rewrite sleNgt E sltNge. -by rewrite sltNge E sleNgt. -Qed. - -Lemma consec_nil t1 t2 : consec [::] t1 t2 = false. -Proof. by case: consecP_ingt=>//=; case. Qed. - -(* frequent projections *) - -Lemma consec_slt ks t1 t2 : consec ks t1 t2 -> t1 <[ks] t2. -Proof. by case/andP. Qed. - -Lemma consec_sltW ks t1 t2 : consec ks t1 t2 -> t1 <=[ks] t2. -Proof. by move/consec_slt/sltW. Qed. - -Lemma consec_mem ks t1 t2 : consec ks t1 t2 -> t1 \in ks. -Proof. by case/andP=>/slt_memE. Qed. - -Lemma consec_oo ks t1 t2 : consec ks t1 t2 -> &=ks `]t1, t2[ = [::]. -Proof. by case/consecP=>_; rewrite -nilp_hasPn=>/nilP. Qed. - -Lemma consec_in (ks : seq A) t1 t2 : - uniq ks -> - consec ks t1 t2 -> {in ks, forall z, z <[ks] t2 = z <=[ks] t1}. -Proof. by move=>U; case/(consecP_inlt _ _ U). Qed. - -(* and some streamlining *) - -Lemma consec_prev (ks : seq A) x y z : - uniq ks -> - consec ks x y -> z <[ks] y -> z <=[ks] x. -Proof. by move=>U; case/(consecP_inlt _ _ U)=>X E N; rewrite -E // (slt_memE N). Qed. - -Lemma consec_prevN (ks : seq A) x y z : - uniq ks -> - z != x -> consec ks x y -> z <[ks] y -> z <[ks] x. -Proof. -move=>U N C /(consec_prev U C). -by rewrite sle_eqVlt; [rewrite (negbTE N) | rewrite (consec_mem C) orbT]. -Qed. - -Lemma consec_next (ks : seq A) x y z : - uniq ks -> - consec ks x y -> x <[ks] z -> y <=[ks] z. -Proof. -move=>U; case/(consecP_ingt _ _ U)=>X E N; case Dz : (z \in ks); first by rewrite E. -by apply: sle_memI; rewrite Dz. -Qed. - -Lemma consec_nextN (ks : seq A) x y z : - uniq ks -> - y \in ks \/ z \in ks -> - y != z -> consec ks x y -> x <[ks] z -> y <[ks] z. -Proof. -move=>U /orP D N C /(consec_next U C). -by rewrite (sle_eqVlt D) (negbTE N). -Qed. - -(* main splitting properties of consec *) - -(* if t2 \in ks then ks splits *) -Lemma consecP_split (ks : seq A) t1 t2 : - uniq ks -> - t2 \in ks -> - reflect (exists xs1 xs2, ks = xs1++t1::t2::xs2) - (consec ks t1 t2). -Proof. -move=>U T2; apply: (iffP idP). -- case/andP=>+ H; case/slt_splitL=>xs1 [xs2][E] N T1 T2'. - rewrite {ks}E mem_cat /= inE (negbTE T2') eq_sym (negbTE N) /= in U T2 H *. - case/in_split: T2=>ks3 [ks4][E T2'']; rewrite {xs2}E in U H *. - by rewrite eqsl_oo_split // in H; move/nilP: H=>->; exists xs1, ks4. -case=>xs1 [xs2] E; move: U; rewrite {ks}E in T2 *. -rewrite cat_uniq /= inE !negb_or -!andbA (eq_sym t1 t2). -case/and8P=>U1 U2 U3 U4 U5 U6 U7 U8. -by rewrite /consec slt_splitR // eqsl_oo_split_consec=>//; rewrite eq_sym. -Qed. - -(* if t2 \notin ks, then t1 is last *) -Lemma consec_last (ks : seq A) t1 t2 : - uniq ks -> - consec ks t1 t2 -> - t2 \notin ks <-> exists ks', ks = rcons ks' t1. -Proof. -move=>U /andP [T]; rewrite nilp_hasPn => /hasPn H. -case: (lastP ks) U H T=>[|xs x] /= {ks} + H. -- by rewrite slt_nil. -rewrite rcons_uniq slt_rcons mem_rcons inE negb_or !(eq_sym x). -case/andP=>Nx Ux; case: ifP=>X; rewrite ?andbF ?andbT. -- move=>Nt; split=>//; case=>ks' /rcons_inj [??]; subst ks' x. - by rewrite (slt_memE Nt) in Nx. -move/contra: (H x)=>/(_ erefl). -rewrite eqslice_mem_uniq /=; last by rewrite rcons_uniq Nx. -rewrite mem_rcons inE eqxx /= in_itv /= negb_and ltEnat /=. -rewrite -!seqlt_unlock -!sleNgt !sle_rcons (negbTE Nx) X /=. -rewrite eqxx /= orbF andbT (eq_sym x). -case/orP=>[/negbTE->|/eqP->] /=. -- by case/andP=>H1 /eqP->; split=>// _; exists xs. -rewrite eqxx /= orbF => H1; split=>//; case=>ks' /rcons_inj [_ Ex]. -by rewrite Ex H1 in Nx. -Qed. - - -(* restatement using last *) -Lemma consec_lastE ks t1 t2 t3 : - uniq ks -> - consec ks t1 t2 -> - t2 \notin ks <-> t1 = last t3 ks. -Proof. -move=>U C; rewrite (consec_last U C). -split=>[[ks' ->]|E]; first by rewrite last_rcons. -case: (lastP ks) E C=>[-> /andP [] /=|s x]. -- by rewrite slt_nil. -by rewrite last_rcons => ->; exists s. -Qed. - -(* not quite the same lemmas as consec_last, but a useful split *) -Lemma consecP_last ks t1 t2 : - uniq ks -> t2 \notin ks -> - reflect (exists xs, ks = rcons xs t1) - (consec ks t1 t2). -Proof. -move=>U T2; apply: (iffP idP). -- by move/consec_last: T2; apply. -case=>xs E; rewrite E /consec rcons_uniq mem_rcons inE negb_or eq_sym in U T2 *. -case/andP: U T2=>T1 U /andP [N T2]. -rewrite slt_rcons (negbTE T2) (negbTE T1) N eq_refl /= nilp_hasPn. -rewrite -all_predC; apply/allP=>x /=; apply: contraTeq=>_. -rewrite eqslice_mem_uniq; last by rewrite rcons_uniq T1. -rewrite mem_rcons inE in_itv /= ltEnat /= !negb_and negb_or. -rewrite -!seqlt_unlock -!sleNgt !sle_rcons (eq_sym x) eqxx. -rewrite orbF T1 (negbTE T2) (negbTE N) /= andbC orbCA orbb. -case: ifP=>/= _; last by rewrite orbN. -by rewrite orbF sleNgt; apply: contra T1; exact: slt_memE. -Qed. - -Lemma consecP_lastE y ks t1 t2 : - uniq ks -> t2 \notin ks -> - consec ks t1 t2 = (t1 \in ks) && (t1 == last y ks). -Proof. -move=>U T2; case: consecP_last=>//. -- by case=>xs ->; rewrite mem_rcons last_rcons inE eq_refl. -move=>H; apply/esym/negP=>/andP [H1 H2]; elim: H. -by apply/(rcons_lastXP _ y); rewrite H1 H2. -Qed. - -Lemma consecP_nilp_filter ks (p : pred _) t1 t2 : - uniq ks -> - consec (filter p ks) t1 t2 <-> - if p t2 then [/\ t1 <[ks] t2, p t1 & nilp (filter p (&=ks `]t1, t2[))] - else [/\ t1 \in ks, p t1 & nilp (filter p (&=ks `]t1, +oo[))]. -Proof. -move=>U; case: ifP=>P2; split. -- case/consecP=>Cx Nx. - move: (slt_memE Cx); rewrite mem_filter=>/andP [P1 K1]. - rewrite slt_filter ?(P1,P2,orbT) // in Cx. - split=>//; rewrite nilp_hasPn; apply: contra Nx. - case/hasP=>x; rewrite mem_filter; case/andP=>Px. - move/(mem_oo _ _ _ U); case=>N1 N2 Kx _; apply/hasP; exists x=>//. - apply/mem_oo; first by apply: filter_uniq. - by rewrite mem_filter !slt_filter ?(P1,P2,Px,orbT). -- case=>N P1 Nx; apply/consecP; split. - - by rewrite slt_filter ?(P1,P2,orbT). - rewrite nilp_hasPn in Nx; apply: contra Nx. - case/hasP=>x /(mem_oo _ _ _ (filter_uniq _ U)); rewrite mem_filter. - case=>/andP [Px Kx] X1 X2 _. - rewrite !slt_filter ?(P1,P2,Px,orbT) // in X1 X2. - by apply/hasP; exists x=>//; rewrite mem_filter Px /=; apply/mem_oo. -- case/consecP=>Cx Nx. - move: (slt_memE Cx); rewrite mem_filter=>/andP [P1 K1]. - split=>//; rewrite nilp_hasPn; apply: contra Nx. - case/hasP=>x; rewrite mem_filter=>/andP [Px] /(mem_ou _ _ U) [Nx Kx] _. - apply/hasP; exists x=>//; apply/mem_oo; first by apply: filter_uniq. - rewrite mem_filter Px Nx; split=>//. - - by rewrite slt_filter ?(P1,Px,orbT). - by apply: slt_memI; rewrite mem_filter ?Px // P2. -case=>K1 P1 Nx; apply/consecP; split. -- apply: slt_memI; first by rewrite mem_filter P1 K1. - by rewrite mem_filter P2. -rewrite nilp_hasPn in Nx; apply: contra Nx. -case/hasP=>x /(mem_oo _ _ _ (filter_uniq _ U)); rewrite mem_filter. -case=>/andP [Px Kx N1 N2] _. -apply/hasP; exists x=>//; rewrite mem_filter Px /=. -apply/mem_ou=>//; split=>//. -by rewrite slt_filter ?(P1,Px,orbT) in N1. -Qed. - -Lemma consecP_filter ks (p : pred _) t1 t2 : - uniq ks -> - consec (filter p ks) t1 t2 <-> - if p t2 - then [/\ t1 <[ks] t2, p t1 & {in &=ks `]t1, t2[ , forall z, ~~ p z}] - else [/\ t1 \in ks , p t1 & {in &=ks `]t1, +oo[, forall z, ~~ p z}]. -Proof. -move=>U; split=>[|H]. -- by move/(consecP_nilp_filter _ _ _ U); case: ifP=>P [?? /nilp_filter]. -by apply/consecP_nilp_filter; case: ifP H=>P [?? /nilp_filter]. -Qed. - -Lemma olt_consec_prev ks t1 t2 : - uniq ks -> - t1 <[ks] t2 <-> exists t, t1 <=[ks] t /\ consec ks t t2. -Proof. -move=>U; split=>[H|]; last first. -- by case=>t [H1] /consecP [H2 _]; apply: (sle_slt_trans H1 H2). -case/slt_splitL: H U=>ks1 [ks2][-> Nt1t2 N1 N2] /=. -rewrite cat_uniq /= negb_or /= -!andbA. -case/and5P=>Uks1 Nt1ks1 /hasPn Hks2 Nt1ks2 Uks2. -case X : (t2 \in ks2); last first. -- have L : last t1 ks2 \notin ks1. - - move: (mem_last t1 ks2); rewrite inE. - by case/orP=>[/eqP ->//|H]; apply: Hks2. - exists (last t1 ks2); split. - - by rewrite sle_cat (negbTE N1) sleL andbT. - apply/andP; split. - - rewrite slt_cat (negbTE N2) (negbTE L) /= slt_cons (eq_sym t2) Nt1t2 /=. - move: (mem_last t1 ks2); rewrite inE=>/orP [->//|H]. - by rewrite slt_memI ?X ?orbT. - rewrite nilp_hasPn; apply: contra L; case/hasP=>x + _; case/mem_oo. - - rewrite cat_uniq /= negb_or Uks1 Uks2 Nt1ks1 Nt1ks2 /= andbT. - by apply/hasPn. - move=>_; rewrite !slt_cat (negbTE N2) !slt_cons (eq_sym t2) Nt1t2 /=. - case Xks1 : (x \in ks1)=>/=; first by move/slt_memE. - case/orP=>// /andP [Nxt1]; rewrite (negbTE Nxt1) /=. - case/orP=>[/eqP/last_nochange|/[swap] Xp1]. - - by rewrite (negbTE Nt1ks2) /==>/eqP ->; rewrite slt_nil. - move: (@sle_last _ x t1 ks2)=>/(_ Uks2 (slt_memE Xp1)) Z. - by move/(sle_slt_trans Z); rewrite slt_irr. -case/splitP: {ks2} X Hks2 Nt1ks2 Uks2=>p1 p2 H2. -rewrite mem_cat cat_uniq /= negb_or rcons_uniq mem_rcons inE. -rewrite (negbTE Nt1t2) /= -!andbA. -case/andP=>Nt1p1 Nt1p2 /and4P [Nt2p1 Up1 /hasPn Hp2 Up2]. -have L : last t1 p1 \notin ks1. -- move: (mem_last t1 p1); rewrite inE. - case/orP=>[/eqP ->//|H]; apply: H2. - by rewrite mem_cat mem_rcons inE H orbT. -exists (last t1 p1); split. -- by rewrite sle_cat (negbTE Nt1ks1) sleL andbT. -apply/andP; split. -- rewrite slt_cat (negbTE N2) (negbTE L) /= slt_cons (eq_sym t2) Nt1t2 /=. - rewrite slt_cat mem_rcons inE eq_refl /= slt_rcons (negbTE Nt2p1) eq_refl /=. - by move: (mem_last t1 p1); rewrite inE=>/orP [->|->] //=; rewrite orbT. -rewrite nilp_hasPn; apply: contra L; case/hasP=>x. -case/mem_oo. -- rewrite cat_uniq /= cat_uniq mem_cat !negb_or mem_rcons rcons_uniq - negb_or Uks1 N1 Nt1t2 Nt1p1 Nt1p2 Nt2p1 Up1 Up2 /= andbT. - by apply/andP; split; apply/hasPn. -move=>_; rewrite !slt_cat !slt_cons !slt_cat !mem_rcons !inE (negbTE N2) - (eq_sym t2) Nt1t2 /= eqxx /=. -case Xks1 : (x \in ks1)=>/=; first by move/slt_memE. -case/orP=>// /andP [Nxt1]; rewrite (negbTE Nxt1) /=. -case/orP=>[/eqP/last_nochange|/[swap]]. -- rewrite (negbTE Nt1p1)=>/eqP ->/=. - by rewrite slt_cons eqxx. -rewrite slt_rcons (negbTE Nt2p1) eqxx /= orbF => Xp1. -rewrite Xp1 orbT slt_rcons Xp1 sltNge. -by move: (@sle_last _ x t1 p1)=>/(_ Up1 Xp1) ->. -Qed. - -Lemma olt_consec_next ks t1 t2 : - uniq ks -> - t1 <[ks] t2 <-> exists t, consec ks t1 t /\ t <=[ks] t2. -Proof. -move=>U; split=>[H|]; last first. -- by case=>t [/consecP [X _] /(slt_sle_trans X)]. -case/slt_splitL: H U=>ks1 [ks2][-> Nt1t2 N1 N2] /=. -rewrite cat_uniq /= negb_or -!andbA. -case/and5P=>Uks1 _ /hasPn Nks2 Nt1ks2 Uks2. -have H : head t2 ks2 \notin ks1. -- move: (mem_head t2 ks2); rewrite inE. - by case/orP=>[/eqP ->//|]; apply: Nks2. -exists (head t2 ks2); split; last first. -- rewrite sle_cat (negbTE H) N2 /= sle_cons (eq_sym t2) Nt1t2 /=. - by rewrite sle_head orbT. -apply/andP; split. -- rewrite slt_cat (negbTE H) (negbTE N1) /= sltL. - case: eqP Nt1ks2 (mem_head t2 ks2)=>// -> X. - by rewrite inE (negbTE Nt1t2) (negbTE X). -rewrite nilp_hasPn; apply: contra H. -case/hasP=>x; case/mem_oo. -- rewrite cat_uniq /= negb_or Uks1 N1 Nt1ks2 Uks2 /= andbT. - by apply/hasPn. -move=>_; rewrite !slt_cat /= !slt_cons eqxx /= andbT (negbTE N1) /=. -case Xks1 : (x \in ks1)=>/=. -- by move/slt_memE=>E; rewrite E in N1. -move=>Nxt1; rewrite (negbTE Nxt1) /=; case: ifP=>// _. -case/andP=>_ /(sle_slt_trans (@sle_head _ t2 ks2 x))=>+ _. -by rewrite slt_irr. -Qed. - -(* previous element is uniquely determined *) -Lemma consec_prev_inj ks t t1 t2 : - uniq ks -> - consec ks t1 t -> - consec ks t2 t -> - t1 = t2. -Proof. -move=>U /andP [T1 +] /andP [T2]; rewrite !nilp_hasPn => /hasPn H1 /hasPn H2. -move: (@slt_total _ ks t1 t2 (slt_memE T1)). -case/or3P=>[/eqP->|L1|L2] //. -- move: (H1 t2)=>/contraL/(_ erefl); apply: contraNeq=>_. - by apply/mem_oo=>//; split=>//; apply: slt_memE T2. -move: (H2 t1)=>/contraL/(_ erefl); apply: contraNeq=>_. -by apply/mem_oo=>//; split=>//; apply: slt_memE T1. -Qed. - -(* next of a non-last element is uniquely determined *) -Lemma consec_next_inj_nonlast ks t t1 t2 t3 : - uniq ks -> - t != last t3 ks -> - consec ks t t1 -> - consec ks t t2 -> t1 = t2. -Proof. -move=>U N C1 C2. -have K1 : t1 \in ks by apply: contraR N=>/(consec_lastE t3 U C1) ->. -have K2 : t2 \in ks by apply: contraR N=>/(consec_lastE t3 U C2) ->. -case/andP: C1 C2=>T1 + /andP [T2]. -rewrite !nilp_hasPn => /hasPn H1 /hasPn H2. -move: (@slt_total _ ks t1 t2 K1); case/or3P=>[/eqP->|L1|L2] //. -- move: (H2 t1)=>/contraL/(_ erefl); apply: contraNeq=>_. - by apply/mem_oo. -move: (H1 t2)=>/contraL/(_ erefl); apply: contraNeq=>_. -by apply/mem_oo. -Qed. - -(* a restatement in a more useful form *) -Lemma consec_next_inj ks t t1 t2 : - uniq ks -> - t1 \in ks -> - consec ks t t1 -> - consec ks t t2 -> t1 = t2. -Proof. -move=>U T C1 C2; suff N : t != last t1 ks. -- by apply: consec_next_inj_nonlast U N C1 C2. -by apply: contraL T=>/eqP /(consec_lastE t1 U C1). -Qed. - -(* consecutiveness and sortedness under general relation *) - -Lemma consec_sorted_lt ltT ks t1 t2 : - uniq ks -> - irreflexive ltT -> - antisymmetric ltT -> - transitive ltT -> - sorted ltT ks -> - t2 \in ks -> - consec ks t1 t2 -> - {in ks, forall z, ltT z t2 = (z == t1) || ltT z t1}. -Proof. -move=>U I Asym T S T2 C; move: (consec_mem C)=>T1. -have {}Asym : antisymmetric (fun x y => (x == y) || ltT x y). -- move=>x y; rewrite (eq_sym y); case: eqP=>//= _. - by apply: (Asym x y). -have {}T : transitive (fun x y => (x == y) || ltT x y). -- move=>x y z; case: eqP=>[->|/eqP _] //=. - case: eqP=>[->->|/eqP _ /=]; first by rewrite orbT. - by case: eqP=>//= _; apply: T. -have {}S : sorted (fun x y => (x == y) || ltT x y) ks. -- by apply: sub_sorted S=>x y ->; rewrite orbT. -move=>z Z; move/(consec_in U)/(_ z Z): C. -rewrite (slt_sorted_leE Asym T S) //. -rewrite (sle_sorted_leE Asym T S) //. -by rewrite !orbA orbb; case: eqP=>//= ->; rewrite I. -Qed. - -Lemma consec_sorted_le (leT : rel A) ks t1 t2 : - uniq ks -> - {in ks, reflexive leT} -> - antisymmetric leT -> - transitive leT -> - sorted leT ks -> - t2 \in ks -> - consec ks t1 t2 -> - {in ks, forall z, leT z t1 = (z != t2) && leT z t2}. -Proof. -move=>U R Asym T S T2 C; move: (consec_mem C)=>T1. -move=>z Z; move/(consec_in U)/(_ z Z): C. -rewrite (slt_sorted_leE Asym T S) //. -rewrite (sle_sorted_leE Asym T S) //. -by move=>->; case: eqP=>// ->; rewrite R. -Qed. - -(* alternative formulation where we sort ks in consec *) -(* this form is required in some proofs for linearizability *) -Lemma consec_sort_lt ltT ks t1 t2 : - irreflexive ltT -> - antisymmetric ltT -> - transitive ltT -> - {in ks &, total (fun x y => (x == y) || ltT x y)} -> - uniq ks -> - t2 \in ks -> - consec (sort (fun x y => (x == y) || ltT x y) ks) t1 t2 -> - {in ks, forall z, ltT z t2 = (z == t1) || ltT z t1}. -Proof. -set ks' := sort _ _=>I asym T Tot U T2 C z Z. -apply: (@consec_sorted_lt ltT ks')=>//. -- by rewrite sort_uniq. -- by apply: sort_sorted_in_lt. -- by rewrite mem_sort. -by rewrite mem_sort. -Qed. - -Lemma consec_sort_le leT ks t1 t2 : - uniq ks -> - {in ks, reflexive leT} -> - antisymmetric leT -> - transitive leT -> - {in ks &, total leT} -> - t2 \in ks -> - consec (sort leT ks) t1 t2 -> - {in ks, forall z, leT z t1 = (z != t2) && leT z t2}. -Proof. -set ks' := sort _ _=>U R Asym T Tot T2 C z Z. -apply: (@consec_sorted_le leT ks')=>//. -- by rewrite sort_uniq. -- by move=>x; rewrite mem_sort; apply: R. -- by apply: sort_sorted_in Tot _ _. -- by rewrite mem_sort. -by rewrite mem_sort. -Qed. - -(*******************************) -(* consec and cons constructor *) -(*******************************) - -Lemma consec_hdswap k1 k2 ks x : - uniq ks -> - k1 \notin ks -> - k2 \notin ks -> - x != k2 -> - consec (k1::ks) k1 x -> consec (k2::ks) k2 x. -Proof. -move=>U K1 K2 N2 C. -have N1 : x != k1 by move/consec_slt: C; rewrite sltL. -move: C; rewrite /consec !sltL N1 N2 /= !nilp_hasPn. -apply: contra=>/hasP [z Z] _; apply/hasP; exists z=>//. -case/mem_oo: Z; first by rewrite /= K2. -rewrite inE sltL slt_cons=>Z1 Nz /andP [_ Z2]. -rewrite (negbTE Nz) /= in Z1 Z2. -apply/mem_oo; first by rewrite /= K1. -rewrite inE sltL slt_cons Z1 N1 Z2 /= orbT; split=>//. -by apply: contra K1 => /eqP <-. -Qed. - -Lemma consec_hd2 k1 k2 ks : - k1 != k2 -> - consec [:: k1, k2 & ks] k1 k2. -Proof. -move=>N; rewrite /consec !sltL eq_sym N /= nilp_hasPn. -apply/hasPn=>z; apply: contraL=>/= _. -rewrite eqslice_mem /= eqxx (negbTE N) /= eqxx. -by apply/hasP; case=>x _; rewrite in_itv; case: x. -Qed. - -(* a useful lemma that collects the case analysis *) -(* for consec (k::ks) *) -Lemma consec_consEP' k k1 k2 ks : - uniq ks -> - k \notin ks -> - consec (k::ks) k1 k2 <-> - if k1 == k then - if k2 \in ks then ks = k2 :: behead ks - else k2 != k /\ ks = [::] - else k2 != k /\ consec ks k1 k2. -Proof. -move=>U N; split; last first. -- case: (k1 =P k)=>[->{k1}|/eqP Nk1k [Nk2]]; last first. - - case/consecP=>H1; move=> H2; apply/consecP. - rewrite slt_cons (negbTE Nk1k) /= H1 andbT; split=>//. - apply: contra H2=>/hasP [x H2 _]; apply/hasP; exists x=>//. - move: H2; case/mem_oo; first by rewrite /= N. - rewrite !slt_cons inE (negbTE Nk1k) Nk2 /=; case: eqVneq=>//= _ Hx Hx1 Hx2. - by apply/mem_oo. - case K2 : (k2 \in ks). - - move=>E; rewrite {K2}E /= in N U *. - rewrite inE negb_or in N. - case/andP: U=>U1 U2; case/andP: N=>N1 N2. - by apply: consec_hd2. - case=>Nk2k E; rewrite {ks N U K2}E. - apply/consecP; rewrite sltL; split=>//. - apply/hasPn=>x; apply: contraL=>_; apply: contra Nk2k. - by case/mem_oo=>//; rewrite inE sltL => /eqP->; rewrite eqxx. -case: (k1 =P k)=>[->|/eqP Nk1k]; last first. -- move/consecP; rewrite slt_cons (negbTE Nk1k) /=. - case=>/andP [Nk2k H1 H2]; split=>//; apply/consecP; split=>//. - apply: contra H2=>/hasP [x] H _. - apply/hasP; exists x=>//. - case/mem_oo: H=>// Hx Hx1 Hx2. - apply/mem_oo=>/=; first by rewrite N. - rewrite inE !slt_cons (negbTE Nk1k) (negbTE Nk2k) /=. - by case: eqVneq=>[E|_] //=; rewrite -E Hx in N. -case K2: (k2 \in ks)=>/consecP []; rewrite sltL => Nk2k; last first. -- move=>H; split=>//; apply/nilP; rewrite nilp_hasPn. - apply: contra H=>/hasP [x X _]; apply/hasP; exists x=>//. - apply/mem_oo; first by rewrite /= N. - rewrite inE sltL slt_cons X Nk2k /= orbT. - case: (x =P k)=>//= [E|_]; first by rewrite -E X in N. - by split=>//; apply: slt_memI=>//; rewrite K2. -case: ks U N K2=>//= x ks; rewrite !inE negb_or. -case/andP=>Nkxs U /andP [Nkx Nks] K2 H; congr (_ :: _). -case/orP: K2=>[/eqP->|K2] //. -apply: contraNeq H=>?; apply/hasP; exists x=>//. -apply/mem_oo; first by rewrite /= inE negb_or Nkx Nks Nkxs. -rewrite sltL slt_cons sltL !inE eqxx /= orbT Nk2k /= (eq_sym x) (negbTE Nkx) /=. -by split=>//; rewrite eq_sym. -Qed. - -Lemma consec_consEP k k1 k2 ks : - uniq ks -> - k \notin ks -> - consec (k::ks) k1 k2 <-> - if k1 == k then - if k2 \in ks then k2 = head k ks - else k2 != k /\ ks = [::] - else k2 != k /\ consec ks k1 k2. -Proof. -move=>U /(consec_consEP' _ _ U)=>->. -case: ifP=>// E1; case: ifP=>//. -by case: {U}ks=>//= x ks E2; split=>[[]|->]. -Qed. - -(* let's make them boolean *) -Lemma consec_consE' k k1 k2 ks : - uniq ks -> - k \notin ks -> - consec (k::ks) k1 k2 = - if k1 == k then - if k2 \in ks then ks == k2 :: behead ks - else (k2 != k) && (ks == [::]) - else (k2 != k) && consec ks k1 k2. -Proof. -move=>U K; rewrite iffE consec_consEP' //. -case: ifP=>H1; last by case: andP. -case: ifP=>H2; first by case: eqP. -by split; [case=>->->|case/andP=>-> /eqP]. -Qed. - -Lemma consec_consE k k1 k2 ks : - uniq ks -> - k \notin ks -> - consec (k::ks) k1 k2 = - if k1 == k then - if k2 \in ks then k2 == head k ks - else (k2 != k) && (ks == [::]) - else (k2 != k) && consec ks k1 k2. -Proof. -move=>U K; rewrite iffE consec_consEP //. -case: ifP=>H1; last by case: andP. -case: ifP=>H2; first by case: eqP. -by split; [case=>->->|case/andP=>-> /eqP]. -Qed. - -(* for rcons, we need a uniqueness condition *) -Lemma consec_rconsEP' k k1 k2 ks : - uniq ks -> k \notin ks -> - consec (rcons ks k) k1 k2 <-> - if k1 != k then - if k2 \notin ks then k2 = k /\ exists ks', ks = rcons ks' k1 - else consec ks k1 k2 - else k2 \notin rcons ks k. -Proof. -move=>U K; split; last first. -- case: eqP=>[-> /= K2|/eqP N /=]. - - by apply/consecP_last=>//; [rewrite rcons_uniq K U | exists ks]. - case: ifP=>K2; last first. - - move/negbT: K2; rewrite negbK=>K2 /(consecP_inlt _ _ U) [K1 E]. - apply/consecP_inlt; first by rewrite rcons_uniq K. - split; first by rewrite mem_rcons inE K1 orbT. - move=>z; rewrite mem_rcons inE; case/orP=>[/eqP ->{z}|Z]. - - rewrite slt_rcons K2 sle_rcons (negbTE K) eq_refl andbT K1 /=. - by apply/negP=>/slt_memE; rewrite (negbTE K). - by rewrite slt_rcons K2 sle_rcons Z; apply: E. - case=>->{k2 K2} [ks' E]; rewrite {ks}E in U K *. - apply/consecP_inlt; first by rewrite rcons_uniq K. - split. - - by rewrite mem_rcons inE mem_rcons inE eq_refl orbT. - rewrite rcons_uniq in U; case/andP: U=>N1 U. - rewrite mem_rcons inE negb_or eq_sym N /= in K. - move=>z; rewrite mem_rcons inE mem_rcons inE => Z. - rewrite slt_rcons mem_rcons inE eq_sym (negbTE N) (negbTE K) /=. - rewrite eq_refl orbF !sle_rcons N1 /= eq_refl orbF !mem_rcons !inE eq_refl /=. - case: eqP Z=>[->{z} /= _|/eqP Nz /=]. - - by rewrite eq_sym (negbTE N) /=; case: ifP=>// K'; rewrite K' sle_memI. - rewrite (eq_sym z); case: eqP=>[/= ->|/=]; first by rewrite sle_refl if_same. - by move=>_ Z; rewrite Z sle_memI. -case/consecP_inlt; first by rewrite rcons_uniq K. -move=>+ E; rewrite mem_rcons inE=>K1. -case/orP: K1 E=>[/eqP ->{k1}|K1] E. -- rewrite eq_refl /=. - move: (E k); rewrite mem_rcons inE eq_refl=>/(_ erefl). - rewrite sle_refl; apply: contraL; rewrite mem_rcons inE. - case/orP=>[/eqP ->|K2]; first by rewrite slt_irr. - rewrite slt_rcons K2; apply/negP. - by move/(slt_trans (slt_memI K2 K)); rewrite slt_irr. -case: eqP K1 K=>[->->//|/eqP N K1 K /=]; case: ifP=>K2; last first. -- move/negbT: K2; rewrite negbK=>K2; apply/consecP_inlt=>//. - split=>// z Z; move: (E z); rewrite mem_rcons inE Z orbT=>/(_ erefl). - by rewrite slt_rcons K2 sle_rcons Z. -move: (E k); rewrite mem_rcons inE eq_refl=>/(_ erefl). -rewrite slt_rcons (negbTE K2) (negbTE K) /= eq_refl andbT. -rewrite sle_rcons (negbTE K) K1 /=; case: eqP=>// -> _; split=>//. -suff -> : k1 = last k1 ks. -- move: ks {E N U K K2} k1 K1; apply: last_ind=>[//|ks x IH] k1. - rewrite mem_rcons inE; case/orP=>[/eqP ->|]. - - by rewrite last_rcons; exists ks. - by case/IH=>ks' E; rewrite last_rcons; exists ks. -apply/eqP/contraT; rewrite eq_sym=>M; exfalso. -move: (last_change M)=>L. -move: (E (last k1 ks)); rewrite mem_rcons inE L orbT=>/(_ erefl). -rewrite slt_rcons sle_rcons (negbTE K2) L /=. -move/esym; rewrite sle_eqVlt; last by rewrite L. -rewrite (negbTE M) /=. -by move/(sle_slt_trans (sle_last k1 U K1)); rewrite slt_irr. -Qed. - -Lemma consec_rconsEP k k1 k2 ks : - uniq ks -> k \notin ks -> - consec (rcons ks k) k1 k2 <-> - if k1 != k then - if k2 \notin ks then k2 = k /\ k1 = last k ks - else consec ks k1 k2 - else k2 \notin rcons ks k. -Proof. -move=>U K; rewrite consec_rconsEP' //. -case: eqP=>//= /eqP N1; case: ifP=>// N2. -split; case=>->; first by case=>ks' ->; rewrite last_rcons. -move/esym=>E; split=>//; rewrite -E. -have : last k ks != k by rewrite E. -move/last_change=>{N1 N2}E. -move: ks k U K E; apply: last_ind=>// xs x IH k U K E. -by rewrite last_rcons; exists xs. -Qed. - -(* boolean version *) -Lemma consec_rconsE' y k k1 k2 ks : - uniq ks -> k \notin ks -> - consec (rcons ks k) k1 k2 = - if k1 != k then - if k2 \notin ks then [&& k2 == k, k1 == last y ks & k1 \in ks] - else consec ks k1 k2 - else k2 \notin rcons ks k. -Proof. -move=>U K; rewrite iffE consec_rconsEP' //. -case: ifP=>_ //; case: ifP=>_ //; split. -- by case=>X1 /(rcons_lastXP _ y); rewrite X1 eq_refl. -case/and3P=>/eqP X1 X2 X3; split=>//. -by apply/(rcons_lastXP _ y); rewrite X2 X3. -Qed. - -(* This one is the same, except it drops k1 \in ks condition *) -(* and replaces y by k. The simplifications may be desirable *) -(* depending on the direction of rewriting *) -Lemma consec_rconsE k k1 k2 ks : - uniq ks -> k \notin ks -> - consec (rcons ks k) k1 k2 = - if k1 != k then - if k2 \notin ks then (k2 == k) && (k1 == last k ks) - else consec ks k1 k2 - else k2 \notin rcons ks k. -Proof. -move=>U K; rewrite iffE consec_rconsEP //. -case: ifP=>H1 //; case: ifP=>H2 //. -by split; [case=>->->; rewrite !eq_refl|case/andP=>/eqP -> /eqP ->]. -Qed. - -Lemma consec_rcons (s : seq A) a x y : - uniq (a :: rcons s y) -> - consec (a :: rcons s y) x y <-> - x = last a s. -Proof. -rewrite /= mem_rcons inE negb_or rcons_uniq -andbA. -case/and4P=>U1 U2 U3 U4. -rewrite -rcons_cons consec_rconsE; last first. -- by rewrite inE negb_or eq_sym U1 U3. -- by rewrite /= U2 U4. -case: (x =P y)=>[->{x}|/eqP N] /=. -- rewrite inE negb_or eq_sym U1 mem_rcons inE eqxx /=. - split=>// E; move: (mem_last a s). - by rewrite -E inE eq_sym (negbTE U1) (negbTE U3). -rewrite inE negb_or eq_sym U1 U3 /= eqxx /=. -by split=>[/eqP|->]. -Qed. - -Lemma consec_ext (s : seq A) x y p : - uniq (rcons s p) -> - consec s x y -> - y \in s -> - consec (rcons s p) x y. -Proof. -rewrite rcons_uniq=>/andP [U1 U2] C S. -rewrite consec_rconsE //=. -case: eqP C U1=>[-> /consec_mem ->//|/eqP N]. -by rewrite S. -Qed. - -Lemma consec_behead k ks x y : - uniq ks -> - k \notin ks -> x != k -> - consec (k::ks) x y -> y != k /\ consec ks x y. -Proof. by move=>U K Nx /(consec_consEP _ _ U K); rewrite (negbTE Nx). Qed. - -(* the following isn't a consequence of consec_consE *) -(* as it's independent of k \notin ks *) -(* TODO not anymore *) -Lemma consec_cons k ks x y : - uniq ks -> k \notin ks -> - x != k -> y != k -> consec ks x y -> consec (k::ks) x y. -Proof. -move=>U N Nx Ny; rewrite /consec slt_cons Ny (negbTE Nx) /=. -case/andP=>->/=; rewrite !nilp_hasPn; apply: contra. -case/hasP=>z Z _; apply/hasP; exists z=>//. -case/mem_oo: Z; first by rewrite /= N. -rewrite inE !slt_cons (negbTE Nx) Ny !(eq_sym z) /=. -case: eqVneq=>//= _ Hz Hxz Hzy. -by apply/mem_oo. -Qed. - -(* Pairs of consecutive elements are totally ordered. *) -(* That is: either the two pairs are the same pair, *) -(* or one of the second projections is ordered before the *) -(* first projection of the other pair. *) -Lemma consec2_total ks x1 x2 y1 y2 : - uniq ks -> - y1 \in ks \/ y2 \in ks -> - consec ks x1 y1 -> consec ks x2 y2 -> - [|| (x1 == x2) && (y1 == y2), y2 <=[ks] x1 | y1 <=[ks] x2]. -Proof. -move=>U. -suff L a1 a2 b1 b2 : b1 \in ks -> - consec ks a1 b1 -> consec ks a2 b2 -> - [|| (a1 == a2) && (b1 == b2), b2 <=[ks] a1 | b1 <=[ks] a2]. -- case=>Y Cx1 Cx2; first by apply: L. - case/or3P: (L _ _ _ _ Y Cx2 Cx1)=>[|->|->]; try by rewrite !orbT. - by case/andP=>/eqP/esym -> /eqP/esym ->; rewrite !eq_refl. -move=>Y1 Cx1 Cx2; move: (consec_mem Cx1) (consec_mem Cx2)=>X1 X2. -case/or3P: (slt_total a2 X1) Cx2 X2=>[/eqP <-{a2}|H|H] Cx2 X2. -- by rewrite (consec_next_inj U Y1 Cx1 Cx2) !eq_refl. -- by rewrite (consec_next U Cx1 H) !orbT. -by rewrite (consec_next U Cx2 H) !orbT. -Qed. - -(***************************************) -(* Consecutiveness induction principle *) -(***************************************) - -Lemma consec_ind_raw k ks (P : A -> Prop) : - k \notin ks -> uniq ks -> - P k -> - (forall t1 t2, t2 \in ks -> consec (k::ks) t1 t2 -> P t1 -> P t2) -> - forall t, t \in k::ks -> P t. -Proof. -elim: ks=>[|x ks IH] //= K U H0 Hstep t; first by rewrite inE=>/eqP ->. -rewrite inE negb_or in K U; case/andP: K U=>K1 K2 /andP [U1 U2]. -rewrite !inE; case/or3P; first by move/eqP=>->. -- move/eqP=>->{t}; apply: Hstep H0; first by rewrite inE eq_refl. - by rewrite /consec sltL eq_sym K1 (eqsl_oo_split_consec (s1:=[::])). -have U : uniq (k::x::ks) by rewrite /= inE negb_or K1 K2 U1 U2. -move=>T; apply: IH=>[|||t1 t2 T2 C|] //; last by rewrite inE T orbT. -(* useful sides *) -have Nx2 : t2 != x by case: eqP T2 U1=>// ->->. -have Nk2 : t2 != k by case: eqP T2 K2=>// ->->. -(* main proof *) -case: (t1 =P k) C=>[->{t1} C _|/eqP Nt1k C]. -- (* in this case, we need two induction steps *) - suff [Ckx Cxt2] : consec [:: k, x & ks] k x /\ consec [:: k, x & ks] x t2. - - have : P x by apply: Hstep Ckx H0; rewrite inE eq_refl. - by apply: Hstep Cxt2; rewrite inE T2 orbT. - split; first by apply: consec_hd2. - apply: consec_cons=>//=. - - by rewrite U1. - - by rewrite inE negb_or K1. - - by rewrite eq_sym. - by apply: consec_hdswap C. -(* another useful side which holds only if k != t1 *) -have Nx1 : t1 != x. -- case: eqP U1=>// <-; move: (consec_mem C). - by rewrite inE (negbTE Nt1k) /= =>->. -(* then the proof is straightforward *) -apply: Hstep; first by rewrite inE T2 orbT. -apply: consec_cons=>//=. -- by rewrite U1. -- by rewrite inE negb_or K1. -by apply: consec_cons=>//; case/consec_behead: C. -Qed. - -(* somewhat modified variant of consec_ind_raw *) -(* where we supply the starting k by hand *) -Lemma consec_ind k ks (P : A -> Prop) : - uniq (k :: ks) -> - P k -> - (forall t1 t2, t2 \in ks -> consec (k::ks) t1 t2 -> P t1 -> P t2) -> - forall t, t \in ks -> P t. -Proof. -move=>/= /andP [U1 U2] P0 IH t D; apply: consec_ind_raw IH _ _=>//. -by rewrite inE D orbT. -Qed. - -(* a version that deconstructs ks to find the starting point *) -(* and gives us useful (though derivable) membership predicates t1 \in ks *) -Lemma consec_indX (ks : seq A) (P : A -> Prop) : - uniq ks -> - (forall t1, t1 \in ks -> ks = t1 :: behead ks -> P t1) -> - (forall t1 t2, t1 \in ks -> t2 \in ks -> consec ks t1 t2 -> - P t1 -> P t2) -> - forall t, t \in ks -> P t. -Proof. -case: ks=>//= k ks /andP [K U] H1 H2; apply: consec_ind_raw=>//. -- by apply: (H1 k)=>//; rewrite inE eq_refl. -move=>t1 t2 N Cx; apply: H2=>//. -- by rewrite (consec_mem Cx). -by rewrite inE N orbT. -Qed. - -(* special variants when we induct over an interval *) -(* that is open/closed/unbounded on the right *) -Lemma consec_indo ks k1 k2 (P : A -> Prop) : - uniq ks -> k1 <[ks] k2 -> - P k1 -> - (forall t1 t2, - t2 \in &=ks `]k1, k2[ -> - consec (&=ks `[k1, k2[) t1 t2 -> P t1 -> P t2) -> - forall t, t \in &=ks `]k1,k2[ -> P t. -Proof. -move=>U N H0 IH t; apply: (consec_ind (k:=k1))=>//=. -- by rewrite eqslice_uniq // andbT eqslice_mem_uniq // negb_and in_itv /= ltxx /= orbT. -by move=>t1 t2 H C; apply: IH=>//; rewrite eqsl_xoL N. -Qed. - -Lemma consec_indx ks k1 k2 (P : A -> Prop) : - uniq ks -> k1 <=[ks] k2 -> k1 \in ks -> - P k1 -> - (forall t1 t2, - t2 \in &=ks `]k1, k2] -> - consec (&=ks `[k1, k2]) t1 t2 -> P t1 -> P t2) -> - forall t, t \in &=ks `]k1, k2] -> P t. -Proof. -move=>U N K H0 IH t; apply: (consec_ind (k:=k1))=>//=. - - by rewrite eqslice_uniq // andbT eqslice_mem_uniq //= negb_and in_itv /= ltxx /= orbT. -by move=>t1 t2 H C; apply: IH=>//; rewrite eqsl_xxL N /= K. -Qed. - -Lemma consec_indu ks k (P : A -> Prop) : - uniq ks -> k \in ks -> - P k -> - (forall t1 t2, - t2 \in &=ks `]k, +oo[ -> - consec (&=ks `[k, +oo[) t1 t2 -> P t1 -> P t2) -> - forall t, t \in &=ks `]k, +oo[ -> P t. -Proof. -move=>U K H0 IH t; apply: (consec_ind (k:=k))=>//=. -- by rewrite eqslice_uniq // andbT eqslice_mem_uniq //= negb_and in_itv /= ltxx /= orbT. -by move=>t1 t2 H C; apply: IH=>//; rewrite eqsl_xuL K. -Qed. - -End ConsecEq. - -Section ConsecOrd. -Variable (A : ordType). -Implicit Type (ks : seq A). - -(* consecutiveness and sortedness under A *) - -Lemma consec_sorted ks t1 t2 : - uniq ks -> - sorted ord ks -> t2 \in ks -> consec ks t1 t2 -> - {in ks, forall z, ord z t2 = oleq z t1}. -Proof. -move=>U S T2 /(consecP_inlt _ _ U) [T1 H] z Z. -rewrite -(slt_sortedE S Z T2) -(sle_sortedE S Z T1). -by apply: H Z. -Qed. - -End ConsecOrd. - -Section ConsecNat. - -(* an element is either first, or has a predecessor *) -Lemma consec_prevX (ks : seq nat) x : - uniq ks -> - x \in ks -> - ks = x :: behead ks \/ exists y, consec ks y x. -Proof. -case: (not_memX ks)=>k N U X. -have U' : uniq (k :: ks) by rewrite /= N U. -have : k <[k::ks] x by rewrite sltL; case: eqP X N=>// ->->. -case/(olt_consec_prev _ _ U')=>t [_]; rewrite consec_consEP' //. -by case: eqP X=>[_ ->|_ _ []]; [left|right; exists t]. -Qed. - -(* an element is either last, or has a successor in ks *) -Lemma consec_nextX (ks : seq nat) x : - uniq ks -> - x \in ks -> - (exists ks', ks = rcons ks' x) \/ - exists y, y \in ks /\ consec ks x y. -Proof. -case: (not_memX ks)=>k N U X. -have Ur : uniq (rcons ks k) by rewrite rcons_uniq N U. -have : x <[rcons ks k] k by rewrite slt_rcons (negbTE N) eq_refl orbF. -case/(olt_consec_next _ _ Ur)=>t []. -rewrite consec_rconsEP' //. -case: eqP X N=>[->->//|/eqP Nkx X N /=]. -case T: (t \in ks)=>/=; first by right; exists t. -by case=>-> {t T} [ks' -> _]; left; exists ks'. -Qed. - -End ConsecNat. - - -(**********************) -(* consec, index, nth *) -(**********************) - -Lemma consec_index (A : eqType) (ks : seq A) (x y : A) : - uniq ks -> - consec ks x y <-> - (index y ks) = (index x ks).+1. -Proof. -move=>U; split. -- elim: ks x y U=>[|k ks IH] /= x y; first by rewrite consec_nil. - case/andP=>U1 U2; rewrite consec_consE // !(eq_sym k). - case: (x =P k); case: (y =P k)=>Ex Ey //. - - by subst x y; rewrite (negbTE U1). - - subst x; case: ifP=>Ky /eqP -> //. - by case: ks {Ky Ex U1 U2 IH}=>//= ??; rewrite eqxx. - by move/IH=>->. -elim: ks x y U=>[|k ks IH] //= x y /andP [U1 U2]. -case: (k =P y)=>N; first by subst k; case: eqP. -move/eqP: N=>N. -case: (k =P x); last first. -- move/eqP=>Nx [I]; rewrite consec_consE // eq_sym (negbTE Nx) eq_sym N /=. - by apply: IH I. -move=>?; subst x; case. -case: ks {IH} U1 U2 N=>[|a ks] /= U1 U2 N. -- by rewrite consec_consE //= eqxx eq_sym N eqxx. -rewrite inE negb_or in U1; case/andP: U1=>U1 U1'. -case/andP: U2=>U2 U2'; case: eqP=>// -> _. -by rewrite consec_hd2. -Qed. - -(* push nat_scope to the top of scope stacks explicitly *) -(* because of clash with order_scope *) -Local Open Scope nat_scope. - -(* consec/nth in hypotheses, i.e., elimination *) - -Lemma consec_nthE (A : eqType) (ks : seq A) a (x : A) i : - uniq (a :: ks) -> - i < size ks -> - consec ks x (nth a ks i.+1) -> - x = nth a ks i. -Proof. -rewrite /= =>/andP [U1 U2] N. -case N1 : (i.+1 < size ks). -- move/[dup]/consec_mem=>X /consec_index-/(_ U2). - by rewrite index_uniq //; case=>->; rewrite nth_index. -have -> : nth a ks i.+1 = a by rewrite nth_default // leqNgt N1. -move/negbT: N1; rewrite -leqNgt leq_eqVlt ltnS leqNgt N orbF. -move/eqP=>S; rewrite (consecP_lastE a) // -nth_last S. -by rewrite -subn1 subSS subn0; case/andP=>_ /eqP. -Qed. - -(* simmilar, but uses consed list *) - -Lemma consec_nthE' (A : eqType) (ks : seq A) a (x : A) i : - uniq (a :: ks) -> - i < size ks -> - consec (a :: ks) x (nth a ks i) -> - x = if i == 0 then a else nth a ks i.-1. -Proof. -move/[dup]=>U /= /andP [U1 U2] N. -move/[dup]/consec_mem=>D /consec_index-/(_ U) /=. -have X : nth a ks i \in ks by rewrite mem_nth. -case: (a =P nth _ _ _) U1=>[->|]; first by rewrite X. -move/eqP=>Y U1; rewrite index_uniq //; case. -case: (a =P x)=>[->->//|/eqP Na -> /=]. -by rewrite nth_index //; move: D; rewrite inE eq_sym (negbTE Na). -Qed. - -(* consec/nth in the conclusion; introduction lemma *) - -Lemma consec_nthI (A : eqType) (ks : seq A) a i : - uniq (a :: ks) -> - i.+1 < size ks -> - consec (a :: ks) (nth a ks i) (nth a ks i.+1). -Proof. -elim: ks a {3 4}a i=>[|k ks IH] a b i //=. -rewrite inE negb_or -andbA; case/and4P=>U1 U2 U3 U4. -rewrite ltnS=>S; rewrite consec_consE //=; last first. -- by rewrite inE negb_or U1 U2. -- by rewrite U3 U4. -rewrite nth_cons; case: i S=>[|i] S /=. -- rewrite (eq_sym k) (negbTE U1) /= nth0. - case: ks {U1 IH} U2 U3 U4 S=>//= c ks U2 U3 U4 S. - - rewrite inE negb_or in U2; case/andP: U2=>U2 U2'. - by rewrite eq_sym U2 consec_consE // eqxx inE eqxx /=. -case: ifP=>E. -- have N : nth b ks i \notin ks by rewrite (eqP E). - by rewrite mem_nth // in N; apply: leq_trans S. -move/negbT: E=>N. -have M : nth b ks i.+1 \in ks by rewrite mem_nth. -case: eqP=>X /=; first by rewrite X (negbTE U2) in M. -by apply: IH=>//=; rewrite U3 U4. -Qed. - -(*******************) -(* consec and path *) -(*******************) - -Lemma consec_path (A : eqType) a ks (R : rel A) : - uniq (a :: ks) -> - (forall x y, x \in a::ks -> y \in ks -> - consec (a :: ks) x y -> R x y) -> - path R a ks. -Proof. -move=>U H; apply/(pathP a)=>i S; rewrite nth_cons. -case: i S=>[|i] S /=; last first. -- apply: H. - - by rewrite inE mem_nth ?orbT //; apply: leq_trans S. - - by rewrite mem_nth. - by apply: consec_nthI U S. -case: ks U H S=>[|k ks] //= U H S; apply: H. -- by rewrite inE eqxx. -- by rewrite inE eqxx. -rewrite inE negb_or -andbA in U. -case/and4P: U=>U1 U2 U3 U4. -rewrite consec_consE /=; last first. -- by rewrite inE negb_or U1. -- by rewrite U3 U4. -by rewrite eqxx inE eqxx. -Qed. - diff --git a/core/useqord.v b/core/useqord.v deleted file mode 100644 index d34c170..0000000 --- a/core/useqord.v +++ /dev/null @@ -1,578 +0,0 @@ -(* -Copyright 2022 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 path interval order. -From pcm Require Import options prelude ordtype seqext. -Local Open Scope order_scope. -Import Order.Theory. - -(* We assume the sequences are unique and use the first index, however most *) -(* lemmas don't require this condition explicitly. The ones that do are *) -(* grouped in a separate section. *) - -(***********************************) -(***********************************) -(* Sequence-induced ordering *) -(* definition and basic properties *) -(***********************************) -(***********************************) - -(* x <[ks] y if first x appears to the left of last y in the sequence ks *) - -(* It turns out it's useful to have 0 <[ks] x, for every x. *) -(* Basically, we use these orderings for reasoning about *) -(* timestamps in histories, and we always keep the null timestamp *) -(* to stand for the initialization step *) -(* That said, the null timestamp is never in any history as *) -(* the initialization step is implicit *) - -Module Type SeqOrdTp. -Parameter seq_le : forall (A : eqType) (ks : seq A), A -> A -> bool. -Parameter seq_lt : forall (A : eqType) (ks : seq A), A -> A -> bool. -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"). -Parameter seqle_unlock : forall (A : eqType) ks (t1 t2 : A), - t1 <=[ks] t2 = (index t1 ks <= index t2 ks)%N. -Parameter seqlt_unlock : forall (A : eqType) ks (t1 t2 : A), - t1 <[ks] t2 = (index t1 ks < index t2 ks)%N. -End SeqOrdTp. - -Module SeqOrd : SeqOrdTp. -Section SeqOrd. -Variables (A : eqType) (ks : seq A) (t1 t2 : A). -Definition seq_le := (index t1 ks <= index t2 ks)%N. -Definition seq_lt := (index t1 ks < index t2 ks)%N. -Definition seqle_unlock := erefl seq_le. -Definition seqlt_unlock := erefl seq_lt. -End SeqOrd. -End SeqOrd. -Export SeqOrd. - -(* alternative rewrites *) -Lemma seqle_unlockE (A : eqType) ks (t1 t2 : A) : - t1 <=[ks] t2 = (index t1 ks <= index t2 ks). -Proof. exact: seqle_unlock. Qed. - -Lemma seqlt_unlockE (A : eqType) ks (t1 t2 : A) : - t1 <[ks] t2 = (index t1 ks < index t2 ks). -Proof. exact: seqlt_unlock. Qed. - -Section SeqLeBase. -Variable (A : eqType). -Implicit Type (ks : seq A). - -(****************** transitivity ****************) - -Lemma sle_trans ks : transitive (seq_le ks). -Proof. by move=>y x z; rewrite !seqle_unlock; apply: leq_trans. Qed. - -Lemma slt_trans ks : transitive (seq_lt ks). -Proof. by move=>y x z; rewrite !seqlt_unlock; apply: ltn_trans. Qed. - -Lemma sle_slt_trans ks t1 t2 t3 : - t1 <=[ks] t2 -> t2 <[ks] t3 -> t1 <[ks] t3. -Proof. by rewrite !seqlt_unlock !seqle_unlock; apply: leq_ltn_trans. Qed. - -Lemma slt_sle_trans ks t1 t2 t3 : - t1 <[ks] t2 -> t2 <=[ks] t3 -> t1 <[ks] t3. -Proof. by rewrite !seqlt_unlock !seqle_unlock; apply: leq_trans. Qed. - - -(****************** reflexivity ****************) - -Lemma sle_refl ks : reflexive (seq_le ks). -Proof. by move=>x; rewrite seqle_unlock. Qed. - -(****************** irreflexivity ***************) - -Lemma slt_irr x ks : x <[ks] x = false. -Proof. by rewrite seqlt_unlock; apply: ltnn. Qed. - -(* non-equational variant *) -Lemma sltnn x ks : ~ x <[ks] x. -Proof. by rewrite slt_irr. Qed. - -(****************** antisymmetry ****************) - -Lemma sle_antisym ks : {in ks, antisymmetric (seq_le ks)}. -Proof. -move=>x Hx y; rewrite !seqle_unlock. -by rewrite -eqn_leq =>/eqP /index_inj; apply. -Qed. - -(****************** asymmetry ***************) - -Lemma slt_asym x y ks : x <[ks] y -> ~~ y <[ks] x. -Proof. by rewrite !seqlt_unlock; case: ltngtP. Qed. - -(***************** totality ********************) - -Lemma sle_total ks x y : x <=[ks] y || y <=[ks] x. -Proof. by rewrite !seqle_unlock; case: ltngtP. Qed. - -Lemma slt_total ks x y : - x \in ks -> - [|| x == y, x <[ks] y | y <[ks] x]. -Proof. -rewrite !seqlt_unlock=>H; case: ltngtP; rewrite ?orbT ?orbF //. -by move/index_inj=>->. -Qed. - -(* transfer properties of sequence ordering *) - -(****************** sle_eqVlt ***************) - -Lemma sle_eqVlt ks t1 t2 : - (t1 \in ks) || (t2 \in ks) -> - t1 <=[ks] t2 = (t1 == t2) || (t1 <[ks] t2). -Proof. -move=>H; rewrite seqlt_unlock seqle_unlock leq_eqVlt /=. -case: (t1 =P t2)=>[->|N] /=; first by rewrite eq_refl. -case: eqP=>//=; case/orP: H=>H; first by move/(index_inj H)/N. -by move/esym/(index_inj H)/esym/N. -Qed. - -(****************** slt_neqAle ***************) - -Lemma slt_neqAle ks t1 t2 : - (t1 \in ks) || (t2 \in ks) -> - t1 <[ks] t2 = (t1 != t2) && (t1 <=[ks] t2). -Proof. -move=>H. -rewrite seqlt_unlock seqle_unlock ltn_neqAle. -case: (t1 =P t2)=>[->|N] /=; first by rewrite eq_refl. -case: eqP=>//=; case/orP: H=>H; first by move/(index_inj H)/N. -by move/esym/(index_inj H)/esym/N. -Qed. - -(****************** sltNge ***************) - -Lemma sltNge ks t1 t2 : t1 <[ks] t2 = ~~ t2 <=[ks] t1. -Proof. by rewrite seqlt_unlock seqle_unlock ltnNge. Qed. - -(****************** sleNgt ***************) - -Lemma sleNgt ks t1 t2 : t1 <=[ks] t2 = ~~ t2 <[ks] t1. -Proof. by rewrite sltNge negbK. Qed. - -(* order properties of the sequence orderings *) - -(****************** slt_neq ***************) - -Corollary slt_neq x y ks : x <[ks] y -> x != y. -Proof. by apply/contraL=>/eqP->; rewrite slt_irr. Qed. - -End SeqLeBase. - -#[export] Hint Resolve sle_refl : core. - -Section SeqLeProp. -Variable (A : eqType). -Implicit Type (ks : seq A). - -Lemma sltW ks t1 t2 : t1 <[ks] t2 -> t1 <=[ks] t2. -Proof. by rewrite seqlt_unlock seqle_unlock; apply: ltnW. Qed. - - -(* membership properties of the sequence orderings *) - -Lemma slt_memI x y ks : x \in ks -> y \notin ks -> x <[ks] y. -Proof. by move=>X /index_memN E; rewrite seqlt_unlock E index_mem. Qed. - -Lemma sle_memI x y ks : y \notin ks -> x <=[ks] y. -Proof. by move/index_memN=>E; rewrite seqle_unlock E index_size. Qed. - -Lemma slt_memE x y ks : x <[ks] y -> x \in ks. -Proof. -rewrite seqlt_unlock -index_mem=>/leq_trans. -by apply; rewrite index_size. -Qed. - -Lemma sle_memE x y ks : x <=[ks] y -> y \in ks -> x \in ks. -Proof. by rewrite seqle_unlock -!index_mem; apply: leq_ltn_trans. Qed. - -(* sequence orderings and constructors *) - -Lemma slt_nil x y : x <[Nil A] y = false. -Proof. by rewrite seqlt_unlock. Qed. -Lemma sle_nil x y : x <=[Nil A] y. -Proof. by rewrite seqle_unlock. Qed. - -(* cons *) - -Lemma slt_cons x y k ks : - x <[k :: ks] y = (y != k) && ((x == k) || (x <[ks] y)). -Proof. by rewrite !seqlt_unlock /= !(eq_sym k); case: eqP; case: eqP. Qed. - -Lemma sle_cons x y k ks : - x <=[k :: ks] y = (x == k) || (y != k) && x <=[ks] y. -Proof. by rewrite sleNgt slt_cons negb_and negbK negb_or -sleNgt. Qed. - -Lemma sltL x y ks : x <[x :: ks] y = (y != x). -Proof. by rewrite slt_cons eq_refl andbT. Qed. - -Lemma sleL x y ks : x <=[x :: ks] y. -Proof. by rewrite sle_cons eq_refl. Qed. - -Lemma sltR x y ks : x <[y :: ks] y = false. -Proof. by rewrite sltNge sleL. Qed. - -Lemma sleR x y ks : x <=[y :: ks] y = (y == x). -Proof. by rewrite sleNgt sltL negbK. Qed. - -(* sequence ordering and head *) - -Lemma sle_head x ks y : (head x ks) <=[ks] y. -Proof. -case: ks=>[|k ks] /=; first by rewrite sle_nil. -by rewrite sleL. -Qed. - -(* sequence orderings and rcons *) - -Lemma slt_rcons x y k ks : - x <[rcons ks k] y = if y \in ks then x <[ks] y - else (x \in ks) || (k != y) && (k == x). -Proof. -rewrite !seqlt_unlock !index_rcons. -case X: (x \in ks); case Y: (y \in ks)=>//=. -- by case: eqP; [rewrite index_mem | rewrite ltnS index_size]. -- move/negbT/index_memN: X=>X; rewrite [LHS]ltnNge [RHS]ltnNge. - rewrite X index_size /=. - case: eqP=>//; first by rewrite index_size. - by rewrite ltnW // ltnS index_size. -rewrite !(eq_sym k). -case: eqP=>_; case: eqP=>_ /=. -- by rewrite ltnn. -- by rewrite ltnS leqnn. -- by rewrite ltnNge leqnSn. -by rewrite ltnn. -Qed. - -Lemma sle_rcons x y k ks : - x <=[rcons ks k] y = if x \in ks then x <=[ks] y - else (y \notin ks) && ((k == x) || (k != y)). -Proof. -by rewrite !sleNgt slt_rcons; case: ifP=>//; rewrite negb_or negb_and negbK. -Qed. - -(* some shortcuts for slt/sle_rcons *) - -Lemma slt_rconsI x y k ks : x <[ks] y -> x <[rcons ks k] y. -Proof. by move=>H; rewrite slt_rcons H (slt_memE H) if_same. Qed. - -Lemma sle_rconsI x y k ks : k != y -> x <=[ks] y -> x <=[rcons ks k] y. -Proof. -move=>N H; rewrite sle_rcons H N orbT andbT. -case: ifP=>// K; apply: contraFT K. -by rewrite negbK; apply: sle_memE H. -Qed. - -Lemma slt_rcons_in x y k ks : - x \in ks -> x <[rcons ks k] y = x <[ks] y. -Proof. -move=>H; rewrite slt_rcons H /=; case: ifP=>// K. -by apply/esym; apply: slt_memI=>//; rewrite K. -Qed. - -Lemma sle_rcons_in x y k ks : - x \in ks -> x <=[rcons ks k] y = x <=[ks] y. -Proof. by move=>X; rewrite sle_rcons X. Qed. - -Lemma slt_rcons_inE ks x y k1 k2 : - (x \in ks) || (y \in ks) -> - x <[rcons ks k1] y = x <[rcons ks k2] y. -Proof. by rewrite !slt_rcons=>/orP [] ->. Qed. - -Lemma sle_rcons_inE ks x y k1 k2 : - (x \in ks) || (y \in ks) -> - x <=[rcons ks k1] y = x <=[rcons ks k2] y. -Proof. by rewrite !sle_rcons=>/orP [] ->. Qed. - -Lemma slt_rconsR ks x k : x <[rcons ks k] k -> x \in ks. -Proof. by rewrite slt_rcons eq_refl orbF; case: ifP=>[_ /slt_memE|]. Qed. - -Lemma sle_rconsR ks x k : x <=[rcons ks k] k -> x \in rcons ks k. -Proof. -rewrite sle_rcons eq_refl orbF mem_rcons inE. -case X: (x \in ks); first by rewrite orbT. -by rewrite orbF eq_sym; case/andP. -Qed. - -(* sequence orderings and concatenation *) -Lemma sle_cat ks1 ks2 x y : - x <=[ks1++ks2] y = if x \in ks1 then x <=[ks1] y - else (y \notin ks1) && x <=[ks2] y. -Proof. -rewrite !seqle_unlock !index_cat. -case X: (x \in ks1); case Y: (y \in ks1)=>//=. -- move/negbT/index_memN: Y=>Y. - by rewrite Y index_size ltnW // ltn_addr // index_mem. -- rewrite -index_mem in Y. - apply/negP=>H; move/(leq_ltn_trans H): Y. - by rewrite ltnNge leq_addr. -by rewrite leq_add2l. -Qed. - -Lemma slt_cat ks1 ks2 x y : - x <[ks1++ks2] y = if y \in ks1 then x <[ks1] y - else (x \in ks1) || x <[ks2] y. -Proof. by rewrite !sltNge sle_cat; case: ifP=>//; rewrite negb_and negbK. Qed. - -(* shortcuts *) - -Lemma slt_catL ks1 ks2 x y : x <[ks1] y -> x <[ks1++ks2] y. -Proof. by move=>H; rewrite slt_cat H (slt_memE H) if_same. Qed. - -Lemma slt_splitR x y ks1 ks2 : y != x -> y \notin ks1 -> x <[ks1++x::ks2] y. -Proof. -by move=>N Y; rewrite slt_cat slt_cons eq_refl andbT (negbTE Y) N orbT. -Qed. - -Lemma sle_splitR x y ks1 ks2 : y \notin ks1 -> x <=[ks1++x::ks2] y. -Proof. -move=>Y; rewrite sle_eqVlt; last first. -- by apply/orP; left; rewrite mem_cat inE eq_refl orbT. -by case: eqP=>[|/eqP N] //=; rewrite (slt_splitR _ _ Y) // eq_sym. -Qed. - -(* the other direction of slt_splitR, further strenghtened *) -(* with an additional fact that x \notin ks1 *) -(* by picking a split with the first occurrence of x *) -(* in fact, we can have both directions here, so we prove a reflect lemma *) -(* but it should really only be used in the direction x <[ks] y -> .. *) -(* because in the other direction slt_splitR is already stronger. *) -Lemma slt_splitL ks x y : - reflect (exists ks1 ks2, [/\ ks = ks1++x::ks2, x != y, - x \notin ks1 & y \notin ks1]) - (x <[ks] y). -Proof. -case H : (x <[ks] y); constructor; last first. -- apply/contraFnot: H; case=>ks1 [ks2][-> N _]. - by apply: slt_splitR; rewrite eq_sym. -rewrite seqlt_unlock in H. -have : x \in ks by rewrite -index_mem (leq_trans H _) // index_size. -case/in_split=>ks1 [ks2][E X]; exists ks1, ks2. -rewrite /seq_lt {ks}E !index_cat /= eq_refl in H *. -case: eqP H=>[->|/eqP N] /=; first by rewrite ltnn. -rewrite (negbTE X) addn0; case: ifP=>//= _. -by rewrite ltnNge index_size. -Qed. - -(* ditto for ole_split *) -Lemma sle_splitL ks x y : - x \in ks -> - reflect (exists ks1 ks2, [/\ ks = ks1++x::ks2, - x \notin ks1 & y \notin ks1]) - (x <=[ks] y). -Proof. -move=>X; case H : (x <=[ks] y); constructor; last first. -- apply/contraFnot: H; case=>ks1 [ks2][-> _ N]. - by apply: sle_splitR. -case/in_split: X=>ks1 [ks2][E X]; exists ks1, ks2; split=>//. -rewrite seqle_unlock {ks}E !index_cat /= eq_refl (negbTE X) addn0 in H. -by case: ifP H=>//; rewrite -index_mem; case: ltngtP. -Qed. - -(* sequence orderings and filter *) - -Lemma slt_filterL (p : pred A) ks x y : - (x \notin ks) || p x -> - x <[ks] y -> x <[filter p ks] y. -Proof. by rewrite !seqlt_unlock; apply: index_filter_ltL. Qed. - -Lemma sle_filterL (p : pred A) ks x y : - (x \notin ks) || p x -> - x <=[ks] y -> x <=[filter p ks] y. -Proof. by rewrite !seqle_unlock; apply: index_filter_leL. Qed. - -Lemma slt_filterR (p : pred A) ks x y : - (y \notin ks) || p y -> - x <[filter p ks] y -> x <[ks] y. -Proof. by rewrite !seqlt_unlock; apply: index_filter_ltR. Qed. - -Lemma sle_filterR (p : pred A) ks x y : - (y \notin ks) || p y -> - x <=[filter p ks] y -> x <=[ks] y. -Proof. by rewrite !seqle_unlock; apply: index_filter_leR. Qed. - -Lemma slt_filter (p : pred A) ks x y : - (x \notin ks) || p x -> (y \notin ks) || p y -> - x <[filter p ks] y = x <[ks] y. -Proof. -by move=>H1 H2; apply/idP/idP; [apply: slt_filterR | apply: slt_filterL]. -Qed. - -Lemma sle_filter (p : pred A) ks x y : - (x \notin ks) || p x -> (y \notin ks) || p y -> - x <=[filter p ks] y = x <=[ks] y. -Proof. -by move=>H1 H2; apply/idP/idP; [apply: sle_filterR | apply: sle_filterL]. -Qed. - -(* sequence orderings and sortedness *) - -(* slt/sle under general sorted relations *) -Lemma slt_sorted_lt ltT ks x y : - transitive ltT -> - sorted ltT ks -> - y \in ks -> x <[ks] y -> ltT x y. -Proof. by rewrite seqlt_unlock; apply: sorted_index_ord. Qed. - -Lemma sle_sorted_lt ltT ks x y : - transitive ltT -> - sorted ltT ks -> - y \in ks -> x <=[ks] y -> (x == y) || ltT x y. -Proof. -move=>T S Y; rewrite sle_eqVlt; last by rewrite Y orbT. -by case/orP=>[->//|/(slt_sorted_lt T S Y) ->]; rewrite orbT. -Qed. - -(* we can get the other direction as well *) -(* if we add antisymmetry *) -(* and the condition that x \in ks *) -Lemma slt_sorted_leE leT ks x y : - antisymmetric leT -> - transitive leT -> - sorted leT ks -> - x \in ks -> y \in ks -> - x <[ks] y = (x != y) && leT x y. -Proof. -move=>As T S X Y; apply/idP/idP. -- by case: eqP=>[->|/eqP N] /=; [apply: contraLR; rewrite slt_irr | apply: slt_sorted_lt]. -by rewrite seqlt_unlock; case/andP=>H K; apply: sorted_ord_index_leq K H. -Qed. - -(* if we add antisymmetry and t1 \in ks *) -Lemma sle_sorted_leE leT ks x y : - antisymmetric leT -> - transitive leT -> - sorted leT ks -> - x \in ks -> y \in ks -> - x <=[ks] y = (x == y) || leT x y. -Proof. -move=>As T S X Y; rewrite sle_eqVlt; last by rewrite X. -by rewrite (slt_sorted_leE As T S X Y); case: eqP. -Qed. - -End SeqLeProp. - -#[export] Hint Resolve slt_nil sle_nil : core. - - -Section SeqLeUniq. -Variable (A : eqType). -Implicit Type (ks : seq A). - -Lemma slt_subseq ks1 ks2 k t : - subseq ks1 ks2 -> uniq ks2 -> k \in ks1 -> t \in ks1 -> - k <[ks1] t = k <[ks2] t. -Proof. -rewrite !seqlt_unlock=>S U T K. -by apply/idP/idP=>/(index_subseq S U T K). -Qed. - -Lemma sle_subseq ks1 ks2 k t : - subseq ks1 ks2 -> uniq ks2 -> k \in ks1 -> t \in ks1 -> - k <=[ks1] t = k <=[ks2] t. -Proof. by move=>S U T K; rewrite !sleNgt (slt_subseq S U K T). Qed. - -(* sequence orderings and last *) - -Lemma sle_last x k ks : - uniq ks -> x \in ks -> x <=[ks] (last k ks). -Proof. by rewrite seqle_unlock; apply: index_last_mono. Qed. - -Lemma sle_last_cons x k ks : - uniq (k :: ks) -> x \in k::ks -> x <=[k::ks] (last k ks). -Proof. -move=>/= /andP [U1 U2]. -rewrite inE sle_cons; case: eqP=>//= /eqP Nxk K. -by rewrite (last_notin K) //=; apply: sle_last. -Qed. - -Lemma slt_last x k ks : - uniq ks -> x \in ks -> last k ks != x -> x <[ks] (last k ks). -Proof. -move=>U X N; move: (sle_last k U X); rewrite sle_eqVlt; last by rewrite X. -by rewrite eq_sym (negbTE N). -Qed. - -Lemma slt_last_cons x k ks : - uniq (k :: ks) -> x \in k::ks -> - last k ks != x -> x <[k::ks] (last k ks). -Proof. -move=>U X N; rewrite slt_neqAle; last by rewrite X. -by rewrite eq_sym N sle_last_cons. -Qed. - -(* every list is sorted by its slt relation, assuming uniqueness *) -Lemma sorted_slt ks : uniq ks -> sorted (seq_lt ks) ks. -Proof. -case: ks=>//= k ks; elim: ks k=>[|k1 ks IH] k2 //=. -rewrite inE negb_or -andbA=>/and4P [N1 N2 N3 N4]. -rewrite sltL eq_sym N1 /=. -have : path (seq_lt [:: k1 & ks]) k1 ks by apply: IH; rewrite N3 N4. -apply: (@sub_in_path _ (mem (k1::ks))); last by apply/allP. -move=>x y /=; rewrite !inE !slt_cons. -case/orP=>[/eqP ->{x}|X]. -- rewrite (eq_sym k1 k2) (negbTE N1) /= eq_refl andbT. - case/orP=>[/eqP ->|Y ->]; first by rewrite eq_refl. - by case: eqP Y N2=>// ->->. -case/orP=>[/eqP ->|Y]; first by rewrite eq_refl. -case: eqP Y N3=>[->|/eqP N Y N3] //=. -case: eqP X N3=>[->->|/eqP M X K1] //= H. -by rewrite H orbT andbT; case: eqP Y N2=>// ->->. -Qed. - -Lemma sorted_sle ks : uniq ks -> sorted (seq_le ks) ks. -Proof. -move=>U; apply: sub_sorted (sorted_slt U). -by move=>x y /sltW. -Qed. - -End SeqLeUniq. - -Section SeqLeOrd. -Variable (A : ordType). -Implicit Type (ks : seq A). - -(* olt/ole and sortedness under ordering on A *) - -Lemma slt_sorted ks x y : - sorted ord ks -> y \in ks -> x <[ks] y -> ord x y. -Proof. by apply/slt_sorted_lt/trans. Qed. - -Lemma sle_sorted ks x y : - sorted ord ks -> y \in ks -> x <=[ks] y -> oleq x y. -Proof. by rewrite oleq_eqVord; apply/sle_sorted_lt/trans. Qed. - -Lemma slt_sortedE ks x y : - sorted ord ks -> - x \in ks -> y \in ks -> - x <[ks] y = ord x y. -Proof. -move=>S X Y; apply/idP/idP; first by apply: slt_sorted S Y. -by rewrite seqlt_unlock; apply: (sorted_ord_index (@irr _) (@trans _)) S X. -Qed. - -Lemma sle_sortedE ks x y : - sorted ord ks -> - x \in ks -> y \in ks -> - x <=[ks] y = oleq x y. -Proof. by move=>S X Y; rewrite oleqNord sleNgt (slt_sortedE S Y X). Qed. - -End SeqLeOrd. diff --git a/core/uslice.v b/core/uslice.v deleted file mode 100644 index b17c674..0000000 --- a/core/uslice.v +++ /dev/null @@ -1,1005 +0,0 @@ -(* -Copyright 2022 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 path interval order. -From pcm Require Import options prelude ordtype seqext slice useqord. -(* We assume the sequences are unique and use the first index, *) -(* however most lemmas don't require this condition explicitly. *) -(* The ones that do are grouped in separate sections. *) -Local Open Scope order_scope. -Import Order.Theory. - -(* slicing by element index *) - -Definition ix_bnd {A : eqType} (s : seq A) (i : itv_bound A) : itv_bound nat := - match i with - | BSide b x => BSide b (index x s) - | BInfty b => BInfty _ b - end. - -Definition ix_itv {A : eqType} (s : seq A) (i : interval A) : interval nat := - let: Interval l u := i in - Interval (ix_bnd s l) (ix_bnd s u). - -Definition eq_slice {A : eqType} (s : seq A) (i : interval A) := - slice s (ix_itv s i). - -Notation "&= s i" := (eq_slice s i) - (at level 1, i at next level, s at next level, - format "&= s i") : fun_scope. - -Section Lemmas. -Variable (A : eqType). -Implicit Type (s : seq A). - -Lemma eqsl_uu s : - &=s `]-oo,+oo[ = s. -Proof. by apply: slice_uu. Qed. - -Lemma eqsl_underL s (i j : itv_bound A) : - bnd (ix_bnd s i) (size s) = 0%N -> - &=s (Interval i j) = &=s (Interval -oo j). -Proof. by move=>H; apply: itv_underL. Qed. - -Lemma eqsl_underR s (i j : itv_bound A) : - bnd (ix_bnd s j) (size s) = 0%N -> - &=s (Interval i j) = [::]. -Proof. by move=>H; apply: itv_underR. Qed. - -Lemma eqsl_overL s (i j : itv_bound A) : - size s <= bnd (ix_bnd s i) (size s) -> - &=s (Interval i j) = [::]. -Proof. by move=>H; apply: itv_overL. Qed. - -Lemma eqsl_overR s (i j : itv_bound A) : - size s <= bnd (ix_bnd s j) (size s) -> - &=s (Interval i j) = &=s (Interval i +oo). -Proof. by move=>Hj; apply: itv_overR. Qed. - -Lemma eqsl_swapped s (i j : itv_bound A) : - bnd (ix_bnd s j) (size s) <= bnd (ix_bnd s i) (size s) -> - &=s (Interval i j) = [::]. -Proof. -by move=>H; rewrite /eq_slice /=; apply: itv_swapped. -Qed. - -Corollary eqsl_notinL y b t s : - t \notin s -> - &=s (Interval (BSide b t) y) = [::]. -Proof. -move=>T; apply: eqsl_overL=>/=. -by rewrite (memNindex T); apply: leq_addr. -Qed. - -Corollary eqsl_notinR x b t s : - t \notin s -> - &=s (Interval x (BSide b t)) = &=s (Interval x +oo). -Proof. -move=>T; apply: eqsl_overR=>/=. -by rewrite (memNindex T); apply: leq_addr. -Qed. - -Lemma eqsl_minfR s (i : itv_bound A) : - &=s (Interval i -oo) = [::]. -Proof. by rewrite /eq_slice /=; apply: itv_minfR. Qed. - -Lemma eqsl_pinfL s (j : itv_bound A) : - &=s (Interval +oo j) = [::]. -Proof. by rewrite /eq_slice /=; apply: itv_pinfL. Qed. - -Lemma eqsliceRO_notin s x a : - x \notin &=s (Interval a (BLeft x)). -Proof. -rewrite /eq_slice /slice /= addn0. -apply/negP=>/mem_drop; apply/negP. -by apply: mem_take_index. -Qed. - -Corollary eqslice_subset_full s i : - {subset &=s i <= s}. -Proof. by exact: slice_subset_full. Qed. - -(* splitting *) - -Lemma eqslice_split (s : seq A) (i : interval A) b (x : A) : - ix_bnd s i.1 <= BSide b (index x s) <= ix_bnd s i.2 -> - &=s i = &=s (Interval i.1 (BSide b x)) ++ &=s (Interval (BSide b x) i.2). -Proof. by case: i=>i j H; rewrite /eq_slice; apply: slice_split_bnd. Qed. - -Corollary eqslice_split_full s b (x : A) : - s = &=s (Interval -oo (BSide b x)) ++ &=s (Interval (BSide b x) +oo). -Proof. by rewrite -[LHS](eqsl_uu s); apply: eqslice_split. Qed. - -(* test lemmas / instantiations *) - -Corollary eqsl_uoxx (ks : seq A) t1 t2 : - t1 <=[ks] t2 -> - &=ks `]-oo, t2] = &=ks `]-oo, t1[ ++ &=ks `[t1, t2]. -Proof. by rewrite seqle_unlock=>H; apply: eqslice_split. Qed. - -Corollary eqsl_uxoo (ks : seq A) t1 t2 : - t1 <[ks] t2 -> - &=ks `]-oo, t2[ = &=ks `]-oo, t1] ++ &=ks `]t1, t2[. -Proof. by rewrite seqlt_unlock=>H; apply: eqslice_split. Qed. - -Corollary eqsl_uoxo (ks : seq A) t1 t2 : - t1 <=[ks] t2 -> - &=ks `]-oo, t2[ = &=ks `]-oo,t1[ ++ &=ks `[t1, t2[. -Proof. by rewrite seqle_unlock=>H; apply: eqslice_split. Qed. - -Corollary eqsl_uxox (ks : seq A) t1 t2 : - t1 <=[ks] t2 -> - &=ks `]-oo, t2] = &=ks `]-oo, t1] ++ &=ks `]t1, t2]. -Proof. by rewrite seqle_unlock=>H; apply: eqslice_split. Qed. - -Corollary eqsl_xxou (ks : seq A) t1 t2 : - t1 <=[ks] t2 -> - &=ks `[t1, +oo[ = &=ks `[t1, t2] ++ &=ks `]t2, +oo[. -Proof. -rewrite seqle_unlock=>H. -by apply: eqslice_split=>/=; rewrite andbT. -Qed. - -Corollary eqsl_uxou (ks : seq A) t : ks = &=ks `]-oo, t] ++ &=ks `]t, +oo[. -Proof. exact: eqslice_split_full. Qed. - -Corollary eqsl_uoxu (ks : seq A) t : ks = &=ks `]-oo, t[ ++ &=ks `[t, +oo[. -Proof. exact: eqslice_split_full. Qed. - -(* -Lemma eqsl_kk_out s l r k : - ~~ l || r -> - &=s (Interval (BSide l k) (BSide r k)) = [::]. -Proof. -move=>H; apply: eqsl_swapped=>/=. -case/orP: H. -- by move/negbTE=>-> /=; case: r=>//=; rewrite addn1 addn0. -move=>-> /=; rewrite addn0; case: l=>/=; first by rewrite addn0. -by rewrite addn1. -Qed. -*) - -Lemma eqsl_kk1 (s : seq A) l r k : - &=s (Interval (BSide l k) (BSide r k)) = - if [&& l, ~~ r & k \in s] then [:: k] else [::]. -Proof. -rewrite /eq_slice /= slice_kk onth_index. -apply/esym; case: ifP; first by case/and3P=>->->->. -move/negbT; rewrite !negb_and negbK. -case/or3P=>[/negbTE->|->|/negbTE->] //=. -- by rewrite andbF. -by rewrite if_same. -Qed. - -(* endpoint split of single-bounded interval *) - -Lemma eqsl_uxR t s : - &=s `]-oo, t] = if t \in s - then rcons (&=s `]-oo, t[) t - else &=s `]-oo, t[. -Proof. -rewrite /eq_slice /= (@slice_split _ _ _ true (index t s)) /=; last first. -- by rewrite in_itv /=. -rewrite slice_kk /= onth_index; case: ifP=>/= H. -- by rewrite cats1. -by rewrite cats0. -Qed. - -Lemma eqsl_xuL t s : - &=s `[t, +oo[ = if t \in s - then t :: &=s `]t, +oo[ - else &=s `]t, +oo[. -Proof. -rewrite /eq_slice /= (@slice_split _ _ _ false (index t s)) //=; last first. -- by rewrite in_itv /= andbT. -by rewrite slice_kk /= onth_index; case: ifP. -Qed. - -Lemma eqsl_xxL t1 t2 s : - &=s `[t1, t2] = if (t1 <=[s] t2) && (t1 \in s) - then t1 :: &=s `]t1, t2] - else [::]. -Proof. -rewrite /eq_slice seqle_unlock /=. -case: leqP=>I /=; last by rewrite itv_swapped_bnd. -rewrite (@slice_split _ _ _ false (index t1 s)) /=; last first. -- by rewrite in_itv /= lexx. -rewrite slice_kk /= onth_index; case: ifP=>//= /negbT N1. -by rewrite (memNindex N1) itv_overL //= addn1. -Qed. - -Lemma eqsl_xxR t1 t2 s : - &=s `[t1, t2] = if t1 <=[s] t2 then - if t2 \in s - then rcons (&=s `[t1, t2[) t2 - else &=s `[t1, +oo[ - else [::]. -Proof. -rewrite /eq_slice seqle_unlock /=. -case: leqP=>I /=; last by rewrite itv_swapped_bnd //. -rewrite (@slice_split _ _ _ true (index t2 s)) /=; last first. -- by rewrite in_itv /= lexx andbT. -rewrite slice_kk /= onth_index /=; case: ifP=>/=; first by rewrite cats1. -rewrite cats0 => /negbT/memNindex->. -by apply: itv_overR=>/=; rewrite addn0. -Qed. - -Lemma eqsl_xoL t1 t2 s : - &=s `[t1, t2[ = if t1 <[s] t2 - then t1 :: &=s `]t1, t2[ - else [::]. -Proof. -rewrite /eq_slice seqlt_unlock /=. -case: ltnP=>I; last by rewrite itv_swapped_bnd. -rewrite (@slice_split _ _ _ false (index t1 s)) /=; last first. -- by rewrite in_itv /= lexx. -rewrite slice_kk /= onth_index; case: ifP=>//= /negbT/memNindex E. -by move: I; rewrite E ltnNge index_size. -Qed. - -Lemma eqsl_oxR t1 t2 s : - &=s `]t1, t2] = if t1 <[s] t2 then - if t2 \in s - then rcons (&=s `]t1, t2[) t2 - else &=s `]t1, +oo[ - else [::]. -Proof. -rewrite /eq_slice seqlt_unlock /=. -case: ltnP=>I; last by rewrite itv_swapped_bnd. -rewrite (@slice_split _ _ _ true (index t2 s)) /=; last first. -- by rewrite in_itv /= lexx andbT. -rewrite slice_kk /= onth_index /=; case: ifP=>/=; first by rewrite cats1. -rewrite cats0 =>/negbT/memNindex->. -by apply: itv_overR=>/=; rewrite addn0. -Qed. - -(* some simplifying equalities on intervals *) - -Lemma eqsl_uL_notinE s b t : - t \notin s -> - &=s `(Interval -oo (BSide b t)) = s. -Proof. -move=>N; rewrite /eq_slice /= itv_overR /=; first by exact: slice_uu. -by rewrite (memNindex N); exact: leq_addr. -Qed. - -Lemma eqsl_uR_notinE s b t : - t \notin s -> - &=s `(Interval (BSide b t) +oo) = [::]. -Proof. -move=>N; rewrite /eq_slice /= itv_overL //=. -by rewrite (memNindex N); exact: leq_addr. -Qed. - -(* eqslice of nil *) -Lemma eqslice0 i : - &=([::] : seq A) i = [::]. -Proof. by rewrite /eq_slice slice0. Qed. - -(* concat *) -(* TODO unify *) - -Lemma eqsl_uL_catE s1 s2 b t : - &= (s1 ++ s2) (Interval -oo (BSide b t)) = - if t \in s1 then &= s1 (Interval -oo (BSide b t)) - else s1 ++ &= s2 (Interval -oo (BSide b t)). -Proof. -rewrite /eq_slice slice_cat /= index_cat; case: ifP=>H1. -- by rewrite index_mem H1 itv_minfR cats0. -rewrite ltnNge leq_addr /= addnC addnK itv_overR /=; first by rewrite slice_uu. -by rewrite -addnA addnCA; exact: leq_addr. -Qed. - -Lemma eqsl_uR_catE s1 s2 b t : - &= (s1 ++ s2) (Interval (BSide b t) +oo) = - if t \in s1 then &= s1 (Interval (BSide b t) +oo) ++ s2 - else &= s2 (Interval (BSide b t) +oo). -Proof. -rewrite /eq_slice slice_cat /= index_cat; case: ifP=>H1. -- by rewrite index_mem H1 slice_uu. -rewrite ltnNge leq_addr /= addnC addnK itv_overL //=. -by rewrite -addnA addnCA; exact: leq_addr. -Qed. - -Lemma eqsl_catE s1 s2 b1 t1 b2 t2 : - &= (s1 ++ s2) (Interval (BSide b1 t1) (BSide b2 t2)) = - if t1 \in s1 - then if t2 \in s1 - then &= s1 (Interval (BSide b1 t1) (BSide b2 t2)) - else &= s1 (Interval (BSide b1 t1) +oo) ++ &= s2 (Interval -oo (BSide b2 t2)) - else if t2 \in s1 - then [::] - else &= s2 (Interval (BSide b1 t1) (BSide b2 t2)). -Proof. -rewrite /eq_slice slice_cat /= !index_cat. -case/boolP: (t1 \in s1)=>H1; case/boolP: (t2 \in s1)=>H2. -- by rewrite !index_mem H1 H2 itv_minfR cats0. -- rewrite index_mem H1 ltnNge leq_addr /= itv_overR /=; last by rewrite -addnA; exact: leq_addr. - by congr (_ ++ _); rewrite addnC addnK. -- by rewrite ltnNge leq_addr /= index_mem H2 itv_minfR cats0 itv_overL //= -addnA; exact: leq_addr. -rewrite !ltnNge !leq_addr /= itv_overL /=; last by rewrite -addnA; exact: leq_addr. -by do 2!rewrite addnC addnK. -Qed. - -(* cons lemmas *) -(* TODO unify *) - -Lemma eqsl_uL_consE s b k t : - &=(k :: s) (Interval -oo (BSide b t)) = - if t == k then - if b - then [::] - else [::k] - else k :: &=s (Interval -oo (BSide b t)). -Proof. -rewrite -cat1s eqsl_uL_catE /= inE; case: eqVneq=>// {t}->. -by rewrite /eq_slice /= eqxx; case: b. -Qed. - -Corollary eqsl_uL_consL s b t : - &=(t :: s) (Interval -oo (BSide b t)) = - if b then [::] else [::t]. -Proof. by rewrite eqsl_uL_consE eqxx. Qed. - -Lemma eqsl1_uL (k : A) b t : - &=[::k] (Interval -oo (BSide b t)) = - if ~~ b || (t!=k) then [::k] else [::]. -Proof. -rewrite eqsl_uL_consE; case: eqVneq=>_. -- by rewrite orbF; case: b. -by rewrite orbT eqslice0. -Qed. - -Lemma eqsl_uR_consE s b k t : - &=(k :: s) (Interval (BSide b t) +oo) = - if t == k then - if b - then k::s - else s - else &=s (Interval (BSide b t) +oo). -Proof. -rewrite -cat1s eqsl_uR_catE /= inE; case: eqVneq=>// {t}->. -by rewrite /eq_slice /= eqxx; case: b. -Qed. - -Corollary eqsl_uR_consL s b t : - &=(t :: s) (Interval (BSide b t) +oo) = if b then t::s else s. -Proof. by rewrite eqsl_uR_consE eqxx. Qed. - -Lemma eqsl1_uR (k : A) b t : - &=[::k] (Interval (BSide b t) +oo) = - if b && (t==k) then [::k] else [::]. -Proof. -rewrite eqsl_uR_consE; case: eqVneq=>_; first by rewrite andbT. -by rewrite eqslice0 andbF. -Qed. - -Lemma eqsl_consE s k b1 t1 b2 t2 : - &=(k :: s) (Interval (BSide b1 t1) (BSide b2 t2)) = - if t1 == k then - if t2 == k - then if b1 && ~~ b2 then [::k] else [::] - else if b1 then k :: &=s (Interval -oo (BSide b2 t2)) - else &=s (Interval -oo (BSide b2 t2)) - else if t2 == k - then [::] - else &=s (Interval (BSide b1 t1) (BSide b2 t2)). -Proof. -rewrite -cat1s eqsl_catE /= !inE. -case: eqVneq=>[{t1}->|N1]; case: eqVneq=>[E2|N2] //=. -- by rewrite /eq_slice /= E2 eqxx; case: b1; case: b2. -by rewrite /eq_slice /= eqxx; case: b1. -Qed. - -(* rcons lemmas *) - -Lemma eqsl_uL_rconsE s b k t : - &=(rcons s k) (Interval -oo (BSide b t)) = - if t \in s - then &=s (Interval -oo (BSide b t)) - else if ~~ b || (t!=k) then rcons s k else s. -Proof. -rewrite -cats1 eqsl_uL_catE /=; case: ifP=>//= _. -by rewrite eqsl1_uL; case: ifP=>_ //; rewrite cats0. -Qed. - -Lemma eqsl_uR_rconsE s b k t : - &=(rcons s k) (Interval (BSide b t) +oo) = - if t \in s - then rcons (&=s (Interval (BSide b t) +oo)) k - else if b && (t == k) then [:: k] else [::]. -Proof. -rewrite -cats1 eqsl_uR_catE /=; case: ifP=>//= _. -- by rewrite cats1. -by rewrite eqsl1_uR. -Qed. - -Lemma eqsl_rconsE s k b1 t1 b2 t2 : - &=(rcons s k) (Interval (BSide b1 t1) (BSide b2 t2)) = - if t1 \in s then - if t2 \in s - then &=s (Interval (BSide b1 t1) (BSide b2 t2)) - else if b2 && (k==t2) - then &=s (Interval (BSide b1 t1) +oo) - else rcons (&=s (Interval (BSide b1 t1) +oo)) k - else if t2 \in s - then [::] - else if [&& b1, k==t1 & (k==t2) ==> ~~b2] - then [::k] else [::]. -Proof. -rewrite -cats1 eqsl_catE. -case/boolP: (t1 \in s)=>H1; case/boolP: (t2 \in s)=>H2 //. -- rewrite /eq_slice /=; case: eqVneq; case B2: b2=>/= _; try by rewrite cats1. - by rewrite cats0. -by rewrite /eq_slice /=; case: eqVneq=>_; case: eqVneq=>_; case: b1; case: b2. -Qed. - -(* test surgery lemmas *) - -Lemma eqsl_uo_split t (s1 s2 : seq A) : - t \notin s1 -> - &=(s1++t::s2) `]-oo, t[ = s1. -Proof. -by move=>X1; rewrite eqsl_uL_catE (negbTE X1) eqsl_uL_consE eqxx cats0. -Qed. - -Lemma eqsl_ou_split t (s1 s2 : seq A) : - t \notin s1 -> - &=(s1++t::s2) `]t, +oo[ = s2. -Proof. -by move=>X1; rewrite eqsl_uR_catE (negbTE X1) eqsl_uR_consE eqxx. -Qed. - -Lemma eqsl_oo_split t1 t2 (s1 s2 s : seq A) : - t1 != t2 -> - t1 \notin s1 -> - t2 \notin s1 -> - t2 \notin s -> - &=(s1++t1::s++t2::s2) `]t1, t2[ = s. -Proof. -move=>N X1 X21 X2. -rewrite eqsl_catE -cat_cons eqsl_uL_catE eqsl_catE /= - eqsl_consE !eqsl_uL_consE eqsl_uR_consE /= !inE !eqxx /= !cats0. -by rewrite (negbTE X1) (negbTE X21) eq_sym (negbTE N) /= (negbTE X2). -Qed. - -Corollary eqsl_oo_nil_split t1 t2 (s s2 : seq A) : - t1 != t2 -> - t2 \notin s -> - &=(t1::s++t2::s2) `]t1, t2[ = s. -Proof. -move=>N X2. -by rewrite -(cat0s (_ :: _ ++ _)); apply: eqsl_oo_split. -Qed. - -Corollary eqsl_oo_split_consec t1 t2 (s1 s2 : seq A) : - t1 != t2 -> - t1 \notin s1 -> t2 \notin s1 -> - &=(s1++t1::t2::s2) `]t1, t2[ = [::]. -Proof. -move=>N X1 X2. -by rewrite -(cat1s t1); apply: (@eqsl_oo_split _ _ _ _ [::]). -Qed. - -Corollary eqsl_oo_nil_split_consec t1 t2 s : - t1 != t2 -> - &=(t1::t2::s) `]t1, t2[ = [::]. -Proof. -move=>N. -by rewrite -(cat0s (_ :: _ :: _)); apply: eqsl_oo_split_consec. -Qed. - -(* intervals and filter *) - -(* TODO unify / better theory *) -Lemma eqsl_filterL (p : pred A) b (y : A) s : - (y \notin s) || p y -> - &= (filter p s) (Interval -oo (BSide b y)) = filter p (&= s (Interval -oo (BSide b y))). -Proof. -case/orP=>Hy. -- rewrite !eqsl_notinR //=; first by rewrite !eqsl_uu. - by apply: contra Hy; rewrite mem_filter; case/andP. -elim: s=>//= h s IH. -case/boolP: (p h)=>/= Hp; last first. -- rewrite {}IH eqsl_uL_consE; case: eqVneq=>[E|_] /=. - - by rewrite -E Hy in Hp. - by rewrite (negbTE Hp). -rewrite !eqsl_uL_consE; case: eqVneq=>_ /=. -- by case: {IH}b=>//=; rewrite Hp. -by rewrite Hp IH. -Qed. - -Lemma eqsl_filterR (p : pred A) b (x : A) s : - (x \notin s) || p x -> - &= (filter p s) (Interval (BSide b x) +oo) = filter p (&= s (Interval (BSide b x) +oo)). -Proof. -case/orP=>Hx. -- by rewrite !eqsl_notinL //= mem_filter negb_and Hx orbT. -elim: s=>//=h s IH. -case/boolP: (p h)=>/= Hp; last first. -- rewrite {}IH eqsl_uR_consE; case: eqVneq=>[E|_] //=. - by rewrite -E Hx in Hp. -rewrite !eqsl_uR_consE; case: eqVneq=>//= _. -by case: {IH}b=>//=; rewrite Hp. -Qed. - -Lemma eqsl_filter (p : pred A) b1 t1 b2 t2 s : - (t1 \notin s) || (p t1 && ((t2 \notin s) || p t2)) -> - &= (filter p s) (Interval (BSide b1 t1) (BSide b2 t2)) = - filter p (&= s (Interval (BSide b1 t1) (BSide b2 t2))). -Proof. -case/orP=>[N1|/andP [H1]]. -- by rewrite !eqsl_notinL //= mem_filter negb_and N1 orbT. -case/orP=>H2. -- rewrite !eqsl_notinR //=. - - by rewrite eqsl_filterR // H1 orbT. - by rewrite mem_filter negb_and H2 orbT. -elim: s=>//= h s IH. -case/boolP: (p h)=>/= Hp; last first. -- rewrite {}IH eqsl_consE; case: eqVneq=>[E1|_]. - - by rewrite -E1 H1 in Hp. - case: eqVneq=>[E2|_] //. - by rewrite -E2 H2 in Hp. -rewrite !eqsl_consE; case: eqVneq=>/=_; case: eqVneq=>//=_. -- by case: ifP=>//= _; rewrite Hp. -rewrite eqsl_filterL; last by rewrite H2 orbT. -by case: ifP=>//= _; rewrite Hp. -Qed. - -Lemma eqslice_mem (i : interval A) (ks : seq A) (k : A) : - k \in &=ks i = - has (fun j => j \in ix_itv ks i) (indexall k ks). -Proof. by rewrite /eq_slice slice_memE. Qed. - -Lemma eqslice_sorted r s i : - sorted r s -> sorted r (&=s i). -Proof. by exact: slice_sorted. Qed. - -End Lemmas. - -Section LemmasUniq. -Variable (A : eqType). -Implicit Type (s : seq A). - -Lemma eqsliceLO_notin (s : seq A) x b : - uniq s -> - x \notin &=s (Interval (BRight x) b). -Proof. -rewrite /eq_slice /slice /= addn1 => U. -move: (mem_drop_indexlast x s); rewrite indexlast_uniq //. -by apply/contra/prefix_drop_sub/prefix_take. -Qed. - -Lemma eqsl_lastR_uniq s x : - uniq s -> - s = &=s `]-oo, (last x s)]. -Proof. -move=>U; rewrite /eq_slice /= [LHS]slice_usize index_last_size_uniq // slice_oPR. -by case: ifP=>// /negbT; rewrite -ltnNge /= ltnS leqn0 => /eqP/size0nil->. -Qed. - -Lemma eqslice_uniq s i : - uniq s -> uniq (&=s i). -Proof. -rewrite /eq_slice; apply: slice_uniq. -Qed. - -Lemma eqslice_mem_uniq (i : interval A) s (x : A) : - uniq s -> - x \in &=s i = - (x \in s) && (index x s \in ix_itv s i). -Proof. -move=>U; rewrite eqslice_mem indexall_uniq //. -by case: ifP=>//= _; rewrite orbF. -Qed. - -End LemmasUniq. - -(* interaction of slicing and sequence ordering under uniqueness assumptions *) - -Section SliceSeqOrd. -Variable (A : eqType). -Implicit Type (s : seq A). - -(* TODO generalize / refactor to useq? *) -Lemma uniq_ux_filter s i : - uniq s -> - &=s `]-oo, i] = [seq x <- s | x <=[s] i]. -Proof. -move=>U; rewrite /eq_slice /= /slice /= drop0 addn1. -elim: s U=>//=h s IH. -case/andP=>Nh U; rewrite sle_cons eqxx /=. -congr (cons _); case: (eqVneq i h)=>[E|N]. -- rewrite -{h}E in Nh *; rewrite take0 -(filter_pred0 s). - apply: eq_in_filter=>z Hz /=; rewrite sle_cons eqxx /= orbF. - by apply/esym/negbTE; apply: contraNneq Nh=><-. -rewrite IH //; apply: eq_in_filter=>z Hz; rewrite sle_cons N /=. -suff: (z != h) by move/negbTE=>->. -by apply: contraNneq Nh=><-. -Qed. - -Lemma uniq_uo_filter s i : - uniq s -> - &=s `]-oo, i[ = [seq x <- s | x <[s] i]. -Proof. -move=>U; rewrite /eq_slice /= /slice /= drop0 addn0. -elim: s U=>//=h s IH. -case/andP=>Nh U; rewrite slt_cons eqxx /= andbT. -case: (eqVneq i h)=>[E|N] /=. -- rewrite -{h}E in Nh *; rewrite -(filter_pred0 s). - by apply: eq_in_filter=>z Hz /=; rewrite slt_cons eqxx. -congr (cons _); rewrite IH //; apply: eq_in_filter=>z Hz. -rewrite slt_cons N /=. -suff: (z != h) by move/negbTE=>->. -by apply: contraNneq Nh=><-. -Qed. - -(* sequence ordering, intervals, and last *) - -Lemma olt_ole_last k s x t : - uniq (k::s) -> t != k -> - x <[k::s] t = x <=[k::s] (last k (&=s `]-oo, t[)). -Proof. -elim: s k=>/= [|y s IH] k //=. -- rewrite slt_cons slt_nil sle_cons sle_nil andbT eqxx orbF. - by move=>_ ->. -rewrite inE negb_or -andbA; case/and4P=>Nky K Y U Nkt. -rewrite slt_cons sle_cons Nkt /=; case: (eqVneq x k)=>//= Nkx. -rewrite eqsl_uL_consE; case: (eqVneq t y)=>/= [E|Nty]. -- by rewrite eqxx /= E sltR. -suff -> /= : last y (&=s `]-oo, t[) != k by rewrite IH //= Y. -apply: contraTneq (mem_last y (&=s `]-oo, t[))=> E. -rewrite inE E negb_or Nky /=; apply: contra K. -by exact: slice_subset_full. -Qed. - -(* a slight variation to add the cons to last *) -Corollary olt_ole_last' k s x t : - uniq (k::s) -> t != k -> - x <[k::s] t = x <=[k::s] (last k (&=(k::s) `]-oo, t[)). -Proof. by move=>U K; rewrite olt_ole_last // eqsl_uL_consE (negbTE K). Qed. - -Corollary eqsl_uox_last k s t : - uniq (k::s) -> t != k -> - &=(k :: s) `]-oo, last k (&=s `]-oo, t[) ] = &=(k :: s) `]-oo,t[. -Proof. -move=>U K. -rewrite (uniq_ux_filter _ U) (uniq_uo_filter _ U). -by apply: eq_in_filter=>x _; rewrite -olt_ole_last. -Qed. - -(* membership *) - -Lemma mem_oo t1 t2 (ks : seq A) (k : A) : - uniq ks -> - reflect ([/\ k \in ks, t1 <[ks] k & k <[ks] t2]) - (k \in &=ks `]t1, t2[). -Proof. -rewrite !seqlt_unlock=>U. -by rewrite eqslice_mem_uniq // in_itv; apply: and3P. -Qed. - -Lemma mem_xo t1 t2 (ks : seq A) k : - uniq ks -> - reflect ([/\ k \in ks, t1 <=[ks] k & k <[ks] t2]) - (k \in &=ks `[t1, t2[). -Proof. -rewrite seqlt_unlock seqle_unlock=>U. -by rewrite eqslice_mem_uniq // in_itv; apply: and3P. -Qed. - -Lemma mem_ox t1 t2 (ks : seq A) k : - uniq ks -> - reflect ([/\ k \in ks, t1 <[ks] k & k <=[ks] t2]) - (k \in &=ks `]t1, t2]). -Proof. -rewrite seqlt_unlock seqle_unlock=>U. -by rewrite eqslice_mem_uniq // in_itv; apply: and3P. -Qed. - -Lemma mem_xx t1 t2 (ks : seq A) k : - uniq ks -> - reflect ([/\ k \in ks, t1 <=[ks] k & k <=[ks] t2]) - (k \in &=ks `[t1, t2]). -Proof. -rewrite !seqle_unlock=>U. -by rewrite eqslice_mem_uniq // in_itv; apply: and3P. -Qed. - -Lemma mem_uo t (ks : seq A) k : - uniq ks -> - reflect ([/\ k \in ks & k <[ks] t])(k \in &=ks `]-oo, t[). -Proof. -rewrite seqlt_unlock=>U. -by rewrite eqslice_mem_uniq // in_itv; apply: andP. -Qed. - -Lemma mem_ux t (ks : seq A) k : - uniq ks -> - reflect ([/\ k \in ks & k <=[ks] t])(k \in &=ks `]-oo, t]). -Proof. -rewrite seqle_unlock=>U. -by rewrite eqslice_mem_uniq // in_itv; apply: andP. -Qed. - -Lemma mem_ou t (ks : seq A) k : - uniq ks -> - reflect ([/\ k \in ks & t <[ks] k])(k \in &=ks `]t, +oo[). -Proof. -rewrite seqlt_unlock=>U. -by rewrite eqslice_mem_uniq // in_itv /= andbT; apply: andP. -Qed. - -Lemma mem_xu t (ks : seq A) k : - uniq ks -> - reflect ([/\ k \in ks & t <=[ks] k])(k \in &=ks `[t, +oo[). -Proof. -rewrite seqle_unlock=>U. -by rewrite eqslice_mem_uniq // in_itv /= andbT; apply: andP. -Qed. - -(* has predT lemmas *) - -Lemma has_predT_uslice (ks : seq A) i : - uniq ks -> - has predT (&=ks i) = has (fun z => index z ks \in ix_itv ks i) ks. -Proof. -move=>U; apply/hasP/hasP; case=>x. -- by rewrite eqslice_mem_uniq // =>/andP [Hx Hi] _; exists x. -by move=>Hx Hi; exists x=>//; rewrite eqslice_mem_uniq // Hx. -Qed. - -(* corollaries *) -Lemma has_oo t1 t2 (ks : seq A) : - uniq ks -> - has predT (&=ks `]t1, t2[) = has (fun z => t1 <[ks] z && z <[ks] t2) ks. -Proof. -move/has_predT_uslice=>->; apply: eq_has=>z. -by rewrite !seqlt_unlock in_itv. -Qed. - -Lemma has_ox t1 t2 (ks : seq A) : - uniq ks -> - has predT (&=ks `]t1, t2]) = has (fun z => t1 <[ks] z && z <=[ks] t2) ks. -Proof. -move/has_predT_uslice=>->; apply: eq_has=>z. -by rewrite seqlt_unlock seqle_unlock in_itv. -Qed. - -Lemma has_xo t1 t2 (ks : seq A) : - uniq ks -> - has predT (&=ks `[t1, t2[) = has (fun z => t1 <=[ks] z && z <[ks] t2) ks. -Proof. -move/has_predT_uslice=>->; apply: eq_has=>z. -by rewrite seqle_unlock seqlt_unlock in_itv. -Qed. - -Lemma has_xx t1 t2 (ks : seq A) : - uniq ks -> - has predT (&=ks `[t1, t2]) = has (fun z => t1 <=[ks] z && z <=[ks] t2) ks. -Proof. -move/has_predT_uslice=>->; apply: eq_has=>z. -by rewrite !seqle_unlock in_itv. -Qed. - -Lemma has_ou t (ks : seq A) : - uniq ks -> - has predT (&=ks `]t, +oo[) = has (fun z => t <[ks] z) ks. -Proof. -move/has_predT_uslice=>->; apply: eq_has=>z. -by rewrite seqlt_unlock in_itv /= andbT. -Qed. - -Lemma has_xu t (ks : seq A) : - uniq ks -> - has predT (&=ks `[t, +oo[) = has (fun z => t <=[ks] z) ks. -Proof. -move/has_predT_uslice=>->; apply: eq_has=>z. -by rewrite seqle_unlock in_itv /= andbT. -Qed. - -Lemma has_uo t (ks : seq A) : - uniq ks -> - has predT (&=ks `]-oo, t[) = has (fun z => z <[ks] t) ks. -Proof. -move/has_predT_uslice=>->; apply: eq_has=>z. -by rewrite seqlt_unlock in_itv. -Qed. - -Lemma has_ux t (ks : seq A) : - uniq ks -> - has predT (&=ks `]-oo, t]) = has (fun z => z <=[ks] t) ks. -Proof. -move/has_predT_uslice=>->; apply: eq_has=>z. -by rewrite seqle_unlock in_itv. -Qed. - -(* negation of has on the left side *) - -Lemma hasNL_oo (ks : seq A) t1 t2 (p : pred A) : - uniq ks -> t1 <[ks] t2 -> - ~~ has p (&=ks `]t1, t2[) -> - {in ks, forall z, p z -> z <[ks] t2 = z <=[ks] t1}. -Proof. -move=>U T /hasPn H z K P. -move: (H z); rewrite eqslice_mem_uniq // in_itv K /= =>/contraL/(_ P). -rewrite negb_and -!seqlt_unlockE -!sleNgt; case/orP=>Hz. -- by rewrite Hz; apply: (sle_slt_trans Hz). -rewrite sltNge Hz sleNgt; congr negb; apply/esym. -by apply: (slt_sle_trans T). -Qed. - -Lemma hasNL_ox (ks : seq A) t1 t2 (p : pred A) : - uniq ks -> t1 <=[ks] t2 -> - ~~ has p (&=ks `]t1, t2]) -> - {in ks, forall z, p z -> z <=[ks] t2 = z <=[ks] t1}. -Proof. -move=>U T /hasPn H z K P. -move: (H z); rewrite eqslice_mem_uniq // in_itv K /= =>/contraL/(_ P). -rewrite negb_and -seqlt_unlockE -seqle_unlockE -sleNgt -sltNge. -case/orP=>Hz. -- by rewrite Hz; apply: (sle_trans Hz). -rewrite !sleNgt Hz; congr negb; apply/esym. -by apply: (sle_slt_trans T). -Qed. - -Lemma hasNL_xo (ks : seq A) t1 t2 (p : pred A) : - uniq ks -> t1 <=[ks] t2 -> - ~~ has p (&=ks `[t1, t2[) -> - {in ks, forall z, p z -> z <[ks] t2 = z <[ks] t1}. -Proof. -move=>U T /hasPn H z K P. -move: (H z); rewrite eqslice_mem_uniq // in_itv K /= =>/contraL/(_ P). -rewrite negb_and -seqle_unlockE -seqlt_unlockE -sltNge -sleNgt. -case/orP=>Hz. -- by rewrite Hz; apply: (slt_sle_trans Hz). -rewrite !sltNge Hz; congr negb; apply/esym. -by apply: (sle_trans T). -Qed. - -Lemma hasNL_xx (ks : seq A) t1 t2 (p : pred A) : - uniq ks -> t1 <=[ks] t2 -> - ~~ has p (&=ks `[t1, t2]) -> - {in ks, forall z, p z -> z <=[ks] t2 = z <[ks] t1}. -Proof. -move=>U T /hasPn H z K P. -move: (H z); rewrite eqslice_mem_uniq // in_itv K /= =>/contraL/(_ P). -rewrite negb_and -!seqle_unlockE -!sltNge; case/orP=>Hz. -- by rewrite Hz; apply/sltW/(slt_sle_trans Hz). -rewrite sltNge sleNgt Hz; congr negb; apply/esym. -by apply/sltW/(sle_slt_trans T). -Qed. - -Lemma hasNL_ou (ks : seq A) t (p : pred A) : - uniq ks -> - ~~ has p (&=ks `]t, +oo[) -> {in ks, forall z, p z -> z <=[ks] t}. -Proof. -move=>U /hasPn H z K P. -move: (H z); rewrite eqslice_mem_uniq // in_itv K /= andbT =>/contraL/(_ P). -by rewrite -seqlt_unlockE -sleNgt. -Qed. - -Lemma hasNL_xu (ks : seq A) t (p : pred A) : - uniq ks -> - ~~ has p (&=ks `[t, +oo[) -> {in ks, forall z, p z -> z <[ks] t}. -Proof. -move=>U /hasPn H z K P. -move: (H z); rewrite eqslice_mem_uniq // in_itv K /= andbT =>/contraL/(_ P). -by rewrite -seqle_unlockE -sltNge. -Qed. - -Lemma hasNL_uo (ks : seq A) t (p : pred A) : - uniq ks -> - ~~ has p (&=ks `]-oo, t[) -> {in ks, forall z, p z -> t <=[ks] z}. -Proof. -move=>U /hasPn H z K P. -move: (H z); rewrite eqslice_mem_uniq // in_itv K /= =>/contraL/(_ P). -by rewrite -seqlt_unlockE -sleNgt. -Qed. - -Lemma hasNL_ux (ks : seq A) t (p : pred A) : - uniq ks -> - ~~ has p (&=ks `]-oo, t]) -> {in ks, forall z, p z -> t <[ks] z}. -Proof. -move=>U /hasPn H z K P. -move: (H z); rewrite eqslice_mem_uniq // in_itv K /= =>/contraL/(_ P). -by rewrite -seqle_unlockE -sltNge. -Qed. - -(* negation of has on the right side *) - -Lemma hasNR_oo (ks : seq A) t1 t2 (p : pred A) : - uniq ks -> - {in ks, forall z, p z -> z <[ks] t2 = z <=[ks] t1} -> - ~~ has p (&=ks `]t1, t2[). -Proof. -move=>U T; apply/hasPn=>z; rewrite eqslice_mem_uniq // in_itv /=. -rewrite -!seqlt_unlockE; case/and3P=>H1 H2 H3; apply: contraL H2=>P. -by rewrite -sleNgt -(T _ H1 P). -Qed. - -Lemma hasNR_ox (ks : seq A) t1 t2 (p : pred A) : - uniq ks -> - {in ks, forall z, p z -> z <=[ks] t2 = z <=[ks] t1} -> - ~~ has p (&=ks `]t1, t2]). -Proof. -move=>U T; apply/hasPn=>z; rewrite eqslice_mem_uniq // in_itv /=. -rewrite -seqlt_unlockE -seqle_unlockE. -case/and3P=>H1 H2 H3; apply: contraL H2=>P. -by rewrite -sleNgt -(T _ H1 P). -Qed. - -Lemma hasNR_xo (ks : seq A) t1 t2 (p : pred A) : - uniq ks -> - {in ks, forall z, p z -> z <[ks] t2 = z <[ks] t1} -> - ~~ has p (&=ks `[t1, t2[). -Proof. -move=>U T; apply/hasPn=>z; rewrite eqslice_mem_uniq // in_itv /=. -rewrite -seqlt_unlockE -seqle_unlockE. -case/and3P=>H1 H2 H3; apply: contraL H2=>P. -by rewrite -sltNge -(T _ H1 P). -Qed. - -Lemma hasNR_xx (ks : seq A) t1 t2 (p : pred A) : - uniq ks -> - {in ks, forall z, p z -> z <=[ks] t2 = z <[ks] t1} -> - ~~ has p (&=ks `[t1, t2]). -Proof. -move=>U T; apply/hasPn=>z; rewrite eqslice_mem_uniq // in_itv /=. -rewrite -!seqle_unlockE. -case/and3P=>H1 H2 H3; apply: contraL H2=>P. -by rewrite -sltNge -(T _ H1 P). -Qed. - -Lemma hasNR_ou (ks : seq A) t (p : pred A) : - uniq ks -> - {in ks, forall z, p z -> z <=[ks] t} -> - ~~ has p (&=ks `]t, +oo[). -Proof. -move=>U T; apply/hasPn=>z; rewrite eqslice_mem_uniq // in_itv /= andbT. -rewrite -seqlt_unlockE; case/andP=>H1 H2; apply: contraL H2=>P. -by rewrite -sleNgt; apply: T. -Qed. - -Lemma hasNR_xu (ks : seq A) t (p : pred A) : - uniq ks -> - {in ks, forall z, p z -> z <[ks] t} -> - ~~ has p (&=ks `[t, +oo[). -Proof. -move=>U T; apply/hasPn=>z; rewrite eqslice_mem_uniq // in_itv /= andbT. -rewrite -seqle_unlockE; case/andP=>H1 H2; apply: contraL H2=>P. -by rewrite -sltNge; apply: T. -Qed. - -Lemma hasNR_uo (ks : seq A) t (p : pred A) : - uniq ks -> - {in ks, forall z, p z -> t <=[ks] z} -> - ~~ has p (&=ks `]-oo, t[). -Proof. -move=>U T; apply/hasPn=>z; rewrite eqslice_mem_uniq // in_itv /=. -rewrite -seqlt_unlockE; case/andP=>H1 H2; apply: contraL H2=>P. -by rewrite -sleNgt; apply: T. -Qed. - -Lemma hasNR_ux (ks : seq A) t (p : pred A) : - uniq ks -> - {in ks, forall z, p z -> t <[ks] z} -> - ~~ has p (&=ks `]-oo, t]). -Proof. -move=>U T; apply/hasPn=>z; rewrite eqslice_mem_uniq // in_itv /=. -rewrite -seqle_unlockE; case/andP=>H1 H2; apply: contraL H2=>P. -by rewrite -sltNge; apply: T. -Qed. - -End SliceSeqOrd. diff --git a/examples/graph.v b/examples/graph.v deleted file mode 100644 index 8820bdc..0000000 --- a/examples/graph.v +++ /dev/null @@ -1,1509 +0,0 @@ -(* -Copyright 2022 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 HB Require Import structures. -From Coq Require Import ssreflect ssrfun. -From mathcomp Require Import ssrbool eqtype ssrnat seq path. -From pcm Require Import options axioms pred prelude seqext. -From pcm Require Import heap pcm unionmap natmap autopcm automap. - -(*************) -(*************) -(* Pregraphs *) -(*************) -(*************) - -(* Pregraphs are natmaps, mapping each graph node x into a node sequence *) -(* that enumerates the children of x (x's adjacency list). *) -(* Pregraph differs from graph, in that edges can *dangle*; that is *) -(* terminate with a node that isn't in the graph's domain. *) -(* Dangling edges are allowed because they enable encoding positional *) -(* information about nodes, as usual in pointer structures. *) -(* For example, if there are 3 links for x, and null is the 2nd link, *) -(* that encodes that the second child of x doesn't exist. *) -(* Non-null dangling links are technically possible, *) -(* but are treated same as null. *) - -(* Pregraph differs from fingraph (of mathcomp) *) -(* in that the set of nodes is drawn from an infinite set *) -(* not from a fixed finite set. *) - -Definition node := nat. -(* A is the contents/labeling of the nodes *) -Record pregraph (A : Type) := - Pregraph {pregraph_base : @UM.base node nat_pred (A * seq node)}. - -Section PregraphUMC. -Variable A : Type. -Implicit Type f : pregraph A. -Local Coercion pregraph_base : pregraph >-> UM.base. -Let pg_valid f := @UM.valid nat nat_pred (A * seq nat) f. -Let pg_empty := Pregraph (@UM.empty nat nat_pred (A * seq nat)). -Let pg_undef := Pregraph (@UM.Undef nat nat_pred (A * seq nat)). -Let pg_upd k v f := Pregraph (@UM.upd nat nat_pred (A * seq nat) k v f). -Let pg_dom f := @UM.dom nat nat_pred (A * seq nat) f. -Let pg_assocs f := @UM.assocs nat nat_pred (A * seq nat) f. -Let pg_free f k := Pregraph (@UM.free nat nat_pred (A * seq nat) f k). -Let pg_find k f := @UM.find nat nat_pred (A * seq nat) k f. -Let pg_union f1 f2 := Pregraph (@UM.union nat nat_pred (A * seq nat) f1 f2). -Let pg_empb f := @UM.empb nat nat_pred (A * seq nat) f. -Let pg_undefb f := @UM.undefb nat nat_pred (A * seq nat) f. -Let pg_from (f : pregraph A) : UM.base _ _ := f. -Let pg_to (b : @UM.base nat nat_pred (A * seq nat)) : pregraph A := Pregraph b. -Let pg_pts k v := Pregraph (@UM.pts nat nat_pred (A * seq nat) k v). - -Lemma pregraph_is_umc : - union_map_axiom pg_valid pg_empty pg_undef pg_upd pg_dom - pg_assocs pg_free pg_find pg_union pg_empb - pg_undefb pg_pts pg_from pg_to. -Proof. by split=>//; split=>[|[]]. Qed. - -HB.instance Definition _ := - isUnion_map.Build node nat_pred (A * seq node)%type - (pregraph A) pregraph_is_umc. -End PregraphUMC. - -HB.instance Definition _ A := isNatMap.Build (A * seq node)%type (pregraph A). -HB.instance Definition _ (A : eqType) := - hasDecEq.Build (pregraph A) - (@union_map_eqP node _ (A * seq node)%type (pregraph A)). -Canonical pregraph_PredType A : PredType (node * (A * seq node)) := - um_PredType (pregraph A). -Coercion Pred_of_history A (x : pregraph A) : {Pred _} := - [eta Mem_UmMap x]. - -Notation "x &-> v" := (ptsT (@pregraph _) x v) (at level 30). - -(**********************************) -(* labels, links, children, edges *) -(**********************************) - -(* maps each node to its contents (ie. label) *) -Definition labels {A} (g : pregraph A) : nmap A := mapv fst g. -HB.instance Definition _ A := OmapFun.copy (@labels A) (@labels A). -Definition olabel {A} (g : pregraph A) x := find x (labels g). - -(* Links of x includes all edges outgoing from x *) -(* and may explicitly include nodes that aren't in dom g *) -(* (i.e., are dangling, null or non-null) *) -Definition links A (g : pregraph A) x := oapp snd [::] (find x g). - -(* children x removes dangling edges (null or non-null) from links *) -Definition children A (g : pregraph A) x : seq node := - filter (mem (dom g)) (links g x). - -(* edge is applicative variant of children *) -(* thus, dangling edges (null or non-null) are *not* edges. *) -Definition edge A (g : pregraph A) : rel node := mem \o children g. -Arguments edge {A} g x y : simpl never. - -Section PregraphLemmas. -Context {A : Type}. -Implicit Type g : pregraph A. - -(* labels lemmas *) - -Lemma In_labelsX g x v : - (x, v) \In labels g <-> - exists lks, (x, (v, lks)) \In g. -Proof. -rewrite In_omfX; split; last by case=>lks H; exists (v, lks). -by case; case=>w lks /= H [<-{v}]; exists lks. -Qed. - -Lemma In_labels g x xs : - (x, xs) \In g -> - (x, xs.1) \In labels g. -Proof. by case: xs=>v lks; rewrite In_labelsX; exists lks. Qed. - -Lemma In_olabel g x xs : - (x, xs) \In g -> - olabel g x = Some xs.1. -Proof. by rewrite /olabel=>/In_labels/In_find ->. Qed. - -(* links lemmas *) - -Lemma links_undef x : links (undef : pregraph A) x = [::]. -Proof. by []. Qed. - -Lemma links_unit x : links (Unit : pregraph A) x = [::]. -Proof. by []. Qed. - -Lemma linksND g x : - x \notin dom g -> - links g x = [::]. -Proof. by rewrite /links/oapp; case: dom_find. Qed. - -Lemma linksUnL g1 g2 x : - valid (g1 \+ g2) -> - links (g1 \+ g2) x = - if x \in dom g1 then links g1 x else links g2 x. -Proof. by move=>V; rewrite /links/oapp findUnL //; case: dom_find. Qed. - -Lemma linksUnR g1 g2 x : - valid (g1 \+ g2) -> - links (g1 \+ g2) x = - if x \in dom g2 then links g2 x else links g1 x. -Proof. by rewrite joinC=>/linksUnL; apply. Qed. - -Lemma size_links g x : - size (links g x) > 0 -> - x \in dom g. -Proof. by rewrite /links/oapp; case: dom_find. Qed. - -Lemma linksD g x y : - y \in links g x -> - x \in dom g. -Proof. by move=>X; apply: size_links; case: (links g x) X. Qed. - -Lemma In_graph g x v xs : - (x, (v, xs)) \In g -> - xs = links g x. -Proof. by rewrite /links/oapp=>/In_find ->. Qed. - -Lemma In_graphX g x : - x \in dom g -> - exists v, (x, (v, links g x)) \In g. -Proof. by case/In_domX=>-[v xs] /[dup] /In_graph ->; eauto. Qed. - -Lemma range_links g x : - x \in dom g -> - links g x \in map snd (range g). -Proof. by case/In_graphX=>v /In_range/(Mem_map snd)/mem_seqP. Qed. - -Lemma links_umfiltk g p x : - links (um_filterk p g) x =i - if p x then links g x else [::]. -Proof. by move=>i; rewrite /links find_umfiltk; case: (p x). Qed. - -(* children lemmas *) - -Lemma children_undef x : children (undef : pregraph A) x = [::]. -Proof. by []. Qed. - -Lemma children_unit x : children (Unit : pregraph A) x = [::]. -Proof. by []. Qed. - -Lemma childrenND g x : - x \notin dom g -> - children g x = [::]. -Proof. by rewrite /children=>/linksND ->. Qed. - -Lemma childrenD g x : - {subset children g x <= dom g}. -Proof. by move=>y; rewrite /children mem_filter; case/andP. Qed. - -Lemma childrenUnL g1 g2 x : - valid (g1 \+ g2) -> - {subset children g1 x <= children (g1 \+ g2) x}. -Proof. -move=>V y; rewrite /children !mem_filter /= =>/andP [Dy Ly]. -by rewrite domUn inE V Dy linksUnL //= (linksD Ly). -Qed. - -Lemma childrenUnR g1 g2 x : - valid (g1 \+ g2) -> - {subset children g2 x <= children (g1 \+ g2) x}. -Proof. by rewrite joinC; apply: childrenUnL. Qed. - -Lemma children_links g x : - {subset children g x <= links g x}. -Proof. by move=>z; rewrite /children mem_filter=>/andP []. Qed. - -(* if x is node in g then g x contains all children of x *) -(* and maybe some more nodes that aren't in g *) -Lemma range_children g x : - x \in dom g -> - exists2 xs, xs \in map snd (range g) & - {subset children g x <= xs}. -Proof. -move=>Dx; exists (links g x); first by apply: range_links. -by apply: children_links. -Qed. - -Lemma children_umfiltk g p x y : - y \in children (um_filterk p g) x = - [&& p x, p y & y \in children g x]. -Proof. -rewrite /children !mem_filter /= links_umfiltk dom_umfiltk inE -andbA. -by case: (p x)=>//=; rewrite !andbF. -Qed. - -(* edge lemmas *) - -Lemma edge_undef x y : edge (undef : pregraph A) x y = false. -Proof. by rewrite /edge/= children_undef. Qed. - -Lemma edge_unit x y : edge (Unit : pregraph A) x y = false. -Proof. by rewrite /edge/= children_unit. Qed. - -Lemma edge_children g x y : - edge g x y = (y \in children g x). -Proof. by []. Qed. - -Lemma edgeUnL g1 g2 x y : - valid (g1 \+ g2) -> - edge g1 x y -> - edge (g1 \+ g2) x y. -Proof. by move=>V; apply: childrenUnL. Qed. - -Lemma edgeUnR g1 g2 x y : - valid (g1 \+ g2) -> - edge g2 x y -> - edge (g1 \+ g2) x y. -Proof. by move=>V; apply: childrenUnR. Qed. - -Lemma edgeD g x y : - edge g x y -> - (x \in dom g) * (y \in dom g). -Proof. -rewrite /edge/= => H; split; last by apply: childrenD H. -by apply: contraLR H=>/childrenND ->. -Qed. - -Lemma edge_links g x y : - edge g x y = (y \in dom g) && (y \in links g x). -Proof. -rewrite /edge/children/links/oapp/= mem_filter /=. -by case: dom_find. -Qed. - -Lemma path_dom g x xs : - path (edge g) x xs -> - {subset xs <= dom g}. -Proof. -elim: xs x=>[|a xs IH] x //= /andP [/edgeD [_ He] /IH]. -by apply: subset_consLI He. -Qed. - -Lemma edge_umfiltk g p x y : - edge (um_filterk p g) x y = - [&& p x, p y & edge g x y]. -Proof. by rewrite /edge /= children_umfiltk. Qed. - -Lemma edge_umfiltkE g p : - {in predC p &, edge (um_filterk (predC p) g) =2 edge g}. -Proof. by move=>x y; rewrite !inE => X Y; rewrite edge_umfiltk /= X Y. Qed. - -End PregraphLemmas. - -Prenex Implicits In_graph. - -(***********************) -(* Depth-first search *) -(***********************) - -(* lifts dfs from mathcomp fingraph to pregraphs *) - -(* list of nodes traversed by depth-first search of g *) -(* at depth n, starting from x, and avoiding v. *) -(* Definition uses children, not links; *) -(* thus, it doesn't follow dangling edges *) -(* and dfs can't express reachability to an outside node. *) -(* If the latter is desired, it can be separately defined *) -(* as a conjunct of dfs and links properties. *) - -Fixpoint dfs A (g : pregraph A) (n : nat) (v : seq node) x := - if (x \notin dom g) || (x \in v) then v else - if n is n'.+1 then foldl (dfs g n') (x :: v) (children g x) else v. - -Section DFSLemmas. -Context {A : Type}. -Implicit Type g : pregraph A. - -Lemma dfs_notin g n v x : - x \notin dom g -> - dfs g n v x = v. -Proof. by elim: n=>[|n IH] /= ->. Qed. - -Lemma subset_dfs g n v x : - {subset v <= foldl (dfs g n) v x}. -Proof. -elim: n x v => [|n IHn] /=; elim=>[|x xs IHx] v //=. -- by case: ifP. -move=>y Hy; apply: IHx; case: ifP=>//= _. -by apply: IHn; rewrite inE Hy orbT. -Qed. - -(* avoidance set is bound by g *) -Lemma subset_foldl_dfs_dom g n v x : - {subset v <= dom g} -> - {subset foldl (dfs g n) v x <= dom g}. -Proof. -elim: n x v=>[|n IHn]; elim=>[|x xs IHx] v //=. -- by case: ifP=>_; apply: IHx. -case: ifP; first by case: (x \in dom g)=>//= H; apply: IHx. -case Dx: (x \in dom g)=>//= H Gx; apply/IHx/IHn. -by move=>z; rewrite inE; case/orP=>[/eqP ->|/Gx]. -Qed. - -Lemma subset_dfs_dom g n v x : - {subset v <= dom g} -> - {subset dfs g n v x <= dom g}. -Proof. -case: n=>[|n] H /=; case: ifP=>//=. -case Dx : (x \in dom g)=>//= _; apply: subset_foldl_dfs_dom. -by move=>z; rewrite inE; case/orP=>[/eqP ->|/H]. -Qed. - -Lemma uniq_dfs_foldl g n v x : - uniq v -> - uniq (foldl (dfs g n) v x). -Proof. -elim: n x v=>[|n IHn]; elim=>[|x xs IHx] v U //=; apply: IHx. -- by rewrite if_same. -case: (x \in dom g)=>//=; case: ifP=>// Xv. -by rewrite IHn //= Xv. -Qed. - -Lemma uniq_dfs g n v x : - uniq v -> - uniq (dfs g n v x). -Proof. -case: n=>[|n] U /=; first by rewrite if_same. -case: (x \in dom g)=>//=; case: ifP=>// Xv. -by rewrite uniq_dfs_foldl //= Xv. -Qed. - -(* there's a path in g from x to y avoiding v *) -Inductive dfs_path g (v : seq node) x y : Prop := - DfsPath xs of - path (edge g) x xs & - y = last x xs & - disjoint v (x :: xs). - -Lemma dfs_path_id g v x : - x \notin v -> - dfs_path g v x x. -Proof. -move=>Vx; apply: (DfsPath (xs:=[::]))=>//=. -by rewrite disjoint1R. -Qed. - -Lemma dfs_pathP g n x y v : - size (dom g) <= size v + n -> - uniq v -> - {subset v <= dom g} -> - y \notin v -> - x \in dom g -> - reflect (dfs_path g v x y) (y \in dfs g n v x). -Proof. -elim: n=>[|n IHn] /= in x y v * => Hv Uv Sv Ny Dx. -- rewrite addn0 in Hv; rewrite Dx if_same (negbTE Ny). - apply: ReflectF; case=>xs E _; rewrite disjoint_consR. - by rewrite (uniq_min_size Uv Sv Hv) Dx. -rewrite Dx /=; have [Vx|Vx] := ifPn. -- by rewrite (negbTE Ny); apply: ReflectF=>[[xs]]; rewrite disjoint_consR Vx. -set v1 := x :: v; set a := children g x; have [->|/eqP Nyx] := eqVneq y x. -- by rewrite subset_dfs ?inE ?eqxx //; apply/ReflectT/dfs_path_id. -apply: (@equivP (exists2 x1, x1 \in a & dfs_path g v1 x1 y))=>/=; last first. -- split=>{IHn} [[x1 Hx1 [p1 P1 E1 D1]]|[p /shortenP []]]. - - apply: (DfsPath (xs:=x1::p1))=>//=; first by rewrite edge_children -/a Hx1. - by rewrite disjoint_consR Vx (disjoint_consLE D1). - case=>[_ _ _ /Nyx|] //= x1 xs /andP [Hx1 Hp1] /and3P [N1 _ _] S1 E1 D1. - exists x1=>//; apply: (DfsPath (xs:=xs))=>//; rewrite disjoint_consL N1. - by rewrite (disjoint_consRE (disjoint_subL (subset_consLR S1) D1)). -move: (Dx). -have {Nyx Ny} : y \notin v1 by apply/norP; move/eqP: Nyx. -have {Sv Dx} : {subset v1 <= dom g} by apply: subset_consLI. -have {Vx Uv} : uniq v1 by rewrite /= Vx. -have {Hv} : size (dom g) <= size v1 + n by rewrite addSnnS. -have : {subset a <= dom g} by apply: childrenD. -elim: {x v}a (x) v1=>[|x xs IHa] x' v /= Dxs Hv U Sv Nv Dx'. -- by rewrite (negbTE Nv); apply: ReflectF; case. -have Dx : x \in dom g by apply: Dxs; rewrite inE eqxx. -have Da : {subset xs <= dom g} by move=>z Z; apply/Dxs/subset_consR. -set v2 := dfs g n v x. -have Sv2 : {subset v <= v2} := @subset_dfs g n v [:: x]. -have [Hy2|Ny2] := boolP (y \in v2). -- rewrite subset_dfs //; apply: ReflectT. - by exists x; [rewrite inE eq_refl|apply/IHn]. -apply: {IHa} (equivP (IHa _ _ _ _ _ _ Ny2 Dx))=>//. -- by rewrite (leq_trans Hv) ?leq_add2r ?uniq_leq_size. -- by rewrite uniq_dfs. -- by apply: subset_dfs_dom. -split=>[][x1 Hx1 [p1 P1 Ey D1]]. -- exists x1; first by rewrite inE Hx1 orbT. - by apply: DfsPath (disjoint_subR Sv2 D1). -have Nx1 : x1 \notin v by rewrite (disjoint_consRE D1). -suff D2 : disjoint v2 (x1 :: p1). -- move: Hx1; rewrite inE; case/orP=>[/eqP ?|Hx1]; last first. - - by exists x1=>//; apply: DfsPath D2. - subst x1; have : x \notin v2 by rewrite (disjoint_consRE D2). - by move/negP; elim; apply/IHn=>//; apply: dfs_path_id. -apply: contraR Ny2=>/disjointN [/= x2 Hx2v Hx2]. -case/splitPl: Hx2 Ey P1 D1=>/= pl pr Ex2. -rewrite last_cat cat_path -cat_cons lastI cat_rcons {}Ex2. -move=>Ey /andP [_ P1]; rewrite disjoint_catR=>/andP [D1 D2]. -have Nx2 : x2 \notin v by rewrite (disjoint_consRE D2). -have [p P E D] := IHn _ _ v Hv U Sv Nx2 Dx Hx2v. -apply/IHn=>//; exists (p ++ pr). -- by rewrite cat_path P -E P1. -- by rewrite last_cat -E. -by rewrite -cat_cons disjoint_catR D (disjoint_consRE D2). -Qed. - -End DFSLemmas. - -(******************) -(* Connectivity *) -(* (reachability) *) -(******************) - -(* start dfs with p as avoidance set, then filter out p. *) -(* this computes only the nodes visited during dfs. *) -(* not for client use, hence primed name *) -Definition component' A (p : pred node) (g : pregraph A) x : seq node := - filter (predC p) (dfs g (size (dom g)) (filter p (dom g)) x). - -(* y is connected from x if y is visited during dfs *) -(* avoiding nodes from p; assuming x in dom g *) -Definition connect A p (g : pregraph A) x : pred node := - fun y => (x \in dom g) && (y \in component' p g x). - -Section ConnectLemmas. -Context {A : Type}. -Implicit Type g : pregraph A. - -Lemma connectP (p : pred node) g x y : - reflect [/\ x \in dom g, ~~ p x & exists xs, - [/\ path (edge g) x xs, y = last x xs & - {in xs, forall z, ~~ p z}]] - (y \in connect p g x). -Proof. -rewrite /connect/component'/= {2}/in_mem /= mem_filter /= andbA. -case: (boolP (x \in dom g))=>Dx; last by constructor; case. -case: (boolP (p y))=>Py. -- constructor; case=>_ Px [xs][P E Pxs]. - suff : ~~ p y by rewrite Py. - have : y \in x :: xs by rewrite E mem_last. - by rewrite inE=>/orP [/eqP ->|/Pxs]. -apply: (iffP (dfs_pathP _ _ _ _ _))=>//. -- by rewrite leq_addl. -- by rewrite filter_uniq. -- by move=>z; rewrite mem_filter=>/andP []. -- by rewrite mem_filter negb_and Py. -- case=>xs P E /disjoint_consRE []. - rewrite mem_filter Dx andbT => Px D. - split=>//; exists xs; split=>// z Xz; move: (allP D z Xz). - by rewrite mem_filter (path_dom P Xz) andbT. -case=>_ Px [xs][P E Pxs]; exists xs=>//. -rewrite disjoint_consR mem_filter negb_and Px /=. -by apply/allP=>z /Pxs; rewrite mem_filter negb_and=>->. -Qed. - -(* there's path from x to y iff *) -(* there's path that doesn't loop through x *) -Lemma connectX (p : pred node) g x y : - reflect [/\ x \in dom g, ~~ p x & exists xs, - [/\ path (edge g) x xs, y = last x xs, - x \notin xs & - {in xs, forall z, ~~ p z}]] - (y \in connect p g x). -Proof. -case: connectP=>H; constructor; last first. -- case=>Dx Hx [xs][Px Ey Nx Hxs]. - by apply: H; split=>//; exists xs. -case: H=>Dx Hx [xs][Px Ey Hxs]; split=>//. -(* if x doesn't appear in the path xs, then done *) -have [Mx|Nx] := boolP (x \in xs); last by exists xs. -(* path xs loops to x; find the last occurrence of x *) -(* and use the part of the path from that occurrence *) -case: {-1} _ _ _ / (splitLastP Mx) (erefl xs)=>/= {Mx} p1 p2 Nxp2 E. -rewrite {xs}E /= in Px Ey Hxs. -rewrite last_cat cat_path rcons_path last_rcons -andbA in Px Ey Hxs. -case/and3P: Px=>Px Ex Pp2; exists p2; split=>//. -by move=>z Z; rewrite Hxs // mem_cat Z orbT. -Qed. - -Lemma connect_undef p x : connect p (undef : pregraph A) x =i pred0. -Proof. by move=>y; apply/connectP; case; rewrite dom_undef. Qed. - -Lemma connect_unit p x : connect p (Unit : pregraph A) x =i pred0. -Proof. by move=>y; apply/connectP; case; rewrite dom0. Qed. - -Lemma connectDx p g x y : - y \in connect p g x -> - (x \in dom g) * ~~ p x. -Proof. by case/connectP. Qed. - -Lemma connectDy p g x y : - y \in connect p g x -> - (y \in dom g) * ~~ p y. -Proof. -case/connectP=>Dx Px [xs][P E Pxs]. -have : y \in x :: xs by rewrite E mem_last. -rewrite inE; case/orP=>[/eqP ->|Dy]. -- by rewrite Dx Px. -by rewrite (path_dom P) // Pxs. -Qed. - -Lemma connectD p g x y : - y \in connect p g x -> - (x \in dom g) * (y \in dom g). -Proof. by move=>C; rewrite (connectDx C) (connectDy C). Qed. - -Lemma connectDp p g x y : - y \in connect p g x -> - (~~ p x) * (~~ p y). -Proof. by move=>C; rewrite (connectDx C) (connectDy C). Qed. - -Lemma connectDN p g x : - x \notin dom g -> - connect p g x =i pred0. -Proof. -move=>Dx y; apply/negP=>/connectD []. -by rewrite (negbTE Dx). -Qed. - -Lemma connectDNp (p : pred node) g x : - p x -> - connect p g x =i pred0. -Proof. -move=>Hx y; apply/negP=>/connectX []. -by rewrite Hx. -Qed. - -Lemma connect0 p g x : - x \in connect p g x = (x \in dom g) && ~~ p x. -Proof. by apply/connectP/andP; case=>// H1 H2; split=>//; exists [::]. Qed. - -Lemma connect0I (p : pred node) g x : - x \in dom g -> - ~~ p x -> - x \in connect p g x. -Proof. by rewrite connect0=>->->. Qed. - -Lemma connect0N (p : pred node) g x y : - x \in dom g -> - ~~ p x -> - y \notin connect p g x -> - x != y. -Proof. by move=>Gx Px; apply: contra=>/eqP <-; rewrite connect0 Gx. Qed. - -Lemma connect_trans p g : transitive (connect p g). -Proof. -move=>x y z /connectP [Dy Hy][ys][Py Ey Hys]. -case/connectP=>Dx Hx [xs][Px Ex Hxs]. -apply/connectP; split=>//. -exists (ys ++ xs); split=>[||w]. -- by rewrite cat_path -Ey Py Px. -- by rewrite last_cat -Ey. -by rewrite mem_cat; case/orP; [apply: Hys|apply: Hxs]. -Qed. - -Lemma connectUnL p g1 g2 x : - valid (g1 \+ g2) -> - {subset connect p g1 x <= - connect p (g1 \+ g2) x}. -Proof. -move=>V y /connectP [Dx Hx][xs][Px Ey Hxs]. -apply/connectP; split=>//; first by rewrite domUn inE V Dx. -exists xs; split=>//. -by apply: sub_path Px=>z w; apply: edgeUnL V. -Qed. - -Lemma connectUnR p g1 g2 x : - valid (g1 \+ g2) -> - {subset connect p g2 x <= - connect p (g1 \+ g2) x}. -Proof. by rewrite joinC; apply: connectUnL. Qed. - -Lemma connectUn p g1 g2 x : - [pcm g1 <= g2] -> - valid g2 -> - {subset connect p g1 x <= connect p g2 x}. -Proof. by case=>g ->; apply: connectUnL. Qed. - -Lemma connect_avoid_graph g p x : - connect p g x =i - connect (predU p (predC (mem (dom g)))) g x. -Proof. -move=>y; apply/connectP/connectP; last first. -- case=>Dx; rewrite !inE negb_or Dx andbT=>Hx [xs][Px Ey Hxs]. - split=>//; exists xs; split=>// z /Hxs. - by rewrite !inE negb_or; case/andP. -case=>Dx Hx [xs][Px Ey Hxs]; split=>//. -- by rewrite !inE negb_or Hx Dx. -exists xs; split=>// z Z; rewrite !inE negb_or Hxs //= negbK. -by apply: path_dom Px _ Z. -Qed. - -Lemma connect_umfiltk g p x : - connect p (um_filterk (predC p) g) x =i connect p g x. -Proof. -case: (normalP g)=>[->|V y]; first by rewrite pfundef. -apply/connectP/connectP; case; rewrite dom_umfiltk inE /=. -- case/andP=>Hx Dx _ [xs][Px Ey Hxs]; split=>//. - exists xs; rewrite -(eq_in_path (@edge_umfiltkE A g p)) //. - by apply/allP/in_consE. -move=>Dx Hx [xs][Px Ey Hxs]; rewrite Dx Hx; split=>//. -exists xs; rewrite (eq_in_path (@edge_umfiltkE A g p)) //. -by apply/allP/in_consE. -Qed. - -Lemma connect_sub g (p1 p2 : pred node) x : - subpred p2 p1 -> - {subset connect p1 g x <= connect p2 g x}. -Proof. -move=>S y /connectP [Dx Hx][xs][Px Ey Hxs]. -apply/connectP; split=>//; first by apply: contra (S x) Hx. -by exists xs; split=>// z /Hxs; apply/contra/S. -Qed. - -Lemma connect_eq g p1 p2 x : - p1 =1 p2 -> - connect p1 g x =i connect p2 g x. -Proof. by move=>S y; apply/idP/idP; apply/connect_sub=>z; rewrite S. Qed. - -(* relativized versions of connect_sub and connect_eq *) -(* it suffices to only prove the precondition for nodes in g *) - -Lemma connect_in_sub g p1 p2 x : - {in dom g, subpred p2 p1} -> - {subset connect p1 g x <= connect p2 g x}. -Proof. -move=>S y; rewrite connect_avoid_graph=>C. -rewrite connect_avoid_graph. -apply: connect_sub C=>z; rewrite !inE -!(orbC (z \notin _)). -by case Dz: (z \in dom g)=>//=; apply: S Dz. -Qed. - -Lemma connect_in_eq g p1 p2 x : - {in dom g, p1 =1 p2} -> - connect p1 g x =i connect p2 g x. -Proof. -move=>S y; rewrite [LHS]connect_avoid_graph [RHS]connect_avoid_graph. -apply: connect_eq=>z; rewrite !inE -!(orbC (z \notin _)). -by case Dz: (z \in dom g)=>//=; rewrite S. -Qed. - -(* deconstructing connecting path from left *) -Lemma edge_connect p g x y : - y != x -> - y \in connect p g x -> - exists2 z, edge g x z & y \in connect p g z. -Proof. -move/eqP=>Nxy /connectP [Dx H][[|a xs]][/= Px Ey Hxs]; first by move/Nxy: Ey. -case/andP: Px=>Px Pxs; exists a=>//; apply/connectP; split. -- by rewrite (edgeD Px). -- by apply: Hxs; rewrite inE eqxx. -exists xs; split=>// z Z; apply: Hxs. -by rewrite inE Z orbT. -Qed. - -(* deconstructing connecting path from right *) -Lemma connect_edge p g x y : - y != x -> - y \in connect p g x -> - exists2 z, z \in connect p g x & edge g z y. -Proof. -move/eqP=>Nxy /connectP [Dx H][xs]. -case: {xs}(lastP xs)=>[|xs a][/= Px Ey Hxs]; first by move/Nxy: Ey. -rewrite last_rcons in Ey; rewrite -{a}Ey in Px Hxs. -rewrite rcons_path in Px; case/andP: Px=>Pxs Px. -exists (last x xs)=>//; apply/connectP; split=>//. -exists xs; split=>// z Z; apply: Hxs. -by rewrite mem_rcons inE Z orbT. -Qed. - -(* extending connecting path from left *) -Lemma edge_connectI p g x y z : - x \in dom g -> - x \notin p -> - edge g x y -> - z \in connect p g y -> - z \in connect p g x. -Proof. -move=>Dx Hx H /connectP [Dy Hy][ys][Py Ez Hys]. -apply/connectP; split=>//; exists (y::ys). -split=>[||w] //=; first by rewrite H Py. -by rewrite inE; case/orP=>[/eqP ->|/Hys]. -Qed. - -(* extending connecting path from right *) -Lemma connect_edgeI p g x y z : - y \in connect p g x -> - edge g y z -> - z \notin p -> - z \in connect p g x. -Proof. -case/connectP=>Dx Hx [xs][Px Ey Hxs] H Hz. -apply/connectP; split=>//. -exists (rcons xs z); split=>/=. -- by rewrite rcons_path Px -Ey H. -- by rewrite last_rcons. -by move=>w; rewrite mem_rcons inE; case/orP=>[/eqP ->|/Hxs]. -Qed. - -End ConnectLemmas. - -(***********************) -(* Connected component *) -(***********************) - -(* part of g reachable from x (avoiding nothing) *) -Definition subconnect A (g : pregraph A) x : pregraph A := - um_filterk (mem (connect pred0 g x)) g. - -Lemma edge_subconnect A (g : pregraph A) x y z : - edge (subconnect g x) y z = - [&& y \in connect pred0 g x, - z \in connect pred0 g x & - edge g y z]. -Proof. -rewrite /edge/children/links/oapp/= find_umfiltk /=. -case: ifP=>// _; case: (find y g)=>[ys|]; last by rewrite andbF. -by rewrite !mem_filter /= dom_umfiltk inE -andbA. -Qed. - -(* connectivity relation out of x in a connected subcomponent *) -(* is that same connectivity relation out of x in the graph *) -Lemma connect_subconnect A (g : pregraph A) x : - connect pred0 (subconnect g x) x =i - connect pred0 g x. -Proof. -move=>w; apply/iffE; split; case/connectP. -- rewrite dom_umfiltk=>/andP [Cx /= Dx] _ [xs][Px Ew _]. - apply/connectP; split=>//; exists xs; split=>//. - by apply/sub_path/Px=>y z; rewrite edge_subconnect=>/and3P []. -move=>Dx _ [xs][Px Ew _]; apply/connectP; split=>//. -- by rewrite dom_umfiltk inE topredE connect0 Dx. -exists xs; split=>//. -apply/(sub_in_path (P:=[in connect pred0 g x]))/Px. -- by move=>y z Cy Cz E; rewrite edge_subconnect Cy Cz E. -apply/allP=>z; rewrite inE; case/orP=>[/eqP ->|Hz]. -- by rewrite connect0 Dx. -case/splitPr: Hz Px=>xs1 xs2. -rewrite -cat1s catA cats1 cat_path=>/andP [Px _]. -apply/connectP; split=>//. -by exists (rcons xs1 z); rewrite last_rcons. -Qed. - -(****************************) -(* Avoidance sets (marking) *) -(****************************) - -Section MarkingLemmas. -Context {A : Type}. -Implicit Type g : pregraph A. - -(* deconstructing connecting path from left *) -(* strenghtening avoidance *) -Lemma edge_connectX p g (x y : node) : - y != x -> - y \in connect p g x -> - exists2 z, - edge g x z & y \in connect (predU (pred1 x) p) g z. -Proof. -move/eqP=>Nxy /connectX [Dx Hx][[|a xs]] /= [Px Ey Nx Hxs]. -- by move/Nxy: Ey. -case/andP: Px=>Exa Pa; exists a=>//; apply/connectP; split. -- by rewrite (edgeD Exa). -- rewrite !inE negb_or Hxs ?inE ?eqxx // andbT. - by case: (a =P x) Nx=>//= ->; rewrite inE eqxx. -exists xs; split=>// z Z. -rewrite !inE negb_or Hxs ?inE ?Z ?orbT // andbT. -by case: (z =P x) Nx=>// <-; rewrite inE Z orbT. -Qed. - -(* extending connecting path from left *) -(* weakening avoidance *) -Lemma edge_connectXI (p : pred node) g x y z : - x \in dom g -> - ~~ p x -> - edge g x y -> - z \in connect (predU (pred1 x) p) g y -> - z \in connect p g x. -Proof. -move=>Dx Px Ex /(connect_sub (p2:=p)) H. -apply: edge_connectI Ex (H _)=>//. -by move=>w W; rewrite inE /= W orbT. -Qed. - -(* the right-way lemmas can't be stated usefully *) - -(* equivalence variant *) -Lemma edge_connectXE p g x y : - y \in connect p g x <-> - [/\ x \in dom g, ~~ p x & y = x \/ exists2 z, - edge g x z & y \in connect (predU (pred1 x) p) g z]. -Proof. -split=>[C|[Dx Hx]]; last first. -- case=>[->|[z E C]]; first by rewrite connect0 Dx Hx. - by apply: edge_connectXI E C. -rewrite !(connectDx C); split=>//. -case: (y =P x)=>[->|/eqP Nxy]; first by left. -by right; apply: edge_connectX Nxy C. -Qed. - -(* if y is reachable from x, but not from x', then *) -(* y is reachable from x by a path that avoids whole *) -(* subcomponent of x' *) -Lemma connect_avoid p g x y x' : - y \notin connect p g x' -> - (y \in connect p g x) = - (y \in connect (predU p [in connect p g x']) g x). -Proof. -move=>Ny; apply/idP/idP; last first. -- by apply: connect_sub=>z Z; rewrite inE Z. -case/connectP=>Dx Hx [xs][Px Ey Hxs]. -have [Nx|Mx] := boolP (has [in connect p g x'] (x :: xs)); last first. -(* if path contains no nodes reachable from x', nothing to do *) -- move/hasPn: Mx=>Mx; apply/connectP; split=>//. - - by rewrite inE negb_or Hx Mx // inE eqxx. - by exists xs; split=>// z Z; rewrite inE negb_or Hxs // Mx // inE Z orbT. -(* if path contains node reachable from x' *) -(* let a be the last such node in the path *) -case: {-1} _ _ _ / (split_findlast Nx) (erefl (x::xs)). -move=>{Nx} a p1 p2=>Ca /hasPn/= Nx' X. -(* suffices to show that y is reachable from a *) -(* because then y is also reachable from x'; contradiction *) -suff Cy : y \in connect p g a. -- by move: (connect_trans Ca Cy) Ny; rewrite -!topredE /= =>->. -case: p1 X Ey Px Hxs=>[[<- _]|b p1 [_ ->]] /= Ey Px Hxs. -- by apply/connectP; split=>//; exists xs. -rewrite cat_path last_cat last_rcons rcons_path -andbA in Ey Px. -case/and3P: Px=>Px Ea P2; apply/connectP; rewrite !(connectDy Ca); split=>//. -by exists p2; split=>// z Z; rewrite Hxs // mem_cat Z orbT. -Qed. - -Lemma connect_avoid1 p g x y x' : - y \notin connect p g x' -> - (y \in connect p g x) = (y \in connect (predU (pred1 x') p) g x). -Proof. -move=>C; apply/idP/idP; last first. -- by apply: connect_sub=>z Z; rewrite inE Z orbT. -rewrite (connect_avoid _ C); apply/connect_in_sub=>z Dz. -by rewrite !inE=>/orP [/eqP <-|->//]; rewrite connect0 Dz orbN. -Qed. - -End MarkingLemmas. - -(******************) -(* Biconnectivity *) -(******************) - -(* symmetric closure of connectivity *) - -(* y is biconnected from x iff *) -(* x and y are mutually connected *) -Definition biconnect A p (g : pregraph A) x : pred node := - [pred y | (x \in connect p g y) && (y \in connect p g x)]. - -Section BiconnectLemmas. -Context {A : Type}. -Implicit Type g : pregraph A. - -Lemma biconnect0 p g x : - x \in dom g -> - ~~ p x -> - biconnect p g x x. -Proof. by move=>Dx Hp; rewrite /biconnect/= connect0 Dx Hp. Qed. - -Lemma biconnectUnL p g1 g2 x y : - valid (g1 \+ g2) -> - biconnect p g1 x y -> - biconnect p (g1 \+ g2) x y. -Proof. by move=>V /andP [Cx Cy]; apply/andP; split; apply/connectUnL. Qed. - -Lemma biconnectUnR p g1 g2 x y : - valid (g1 \+ g2) -> - biconnect p g2 x y -> - biconnect p (g1 \+ g2) x y. -Proof. by rewrite joinC; apply: biconnectUnL. Qed. - -Lemma biconnect_sub g (p1 p2 : pred node) (x : node) : - subpred p2 p1 -> - {subset biconnect p1 g x <= biconnect p2 g x}. -Proof. -by move=>S y; rewrite !inE=>/andP [Hx Hy]; rewrite !(connect_sub S). -Qed. - -Lemma biconnect_eq g (p1 p2 : pred node) (x : node) : - p1 =1 p2 -> - biconnect p1 g x =i biconnect p2 g x. -Proof. by move=>S y; rewrite !inE !(connect_eq g _ S). Qed. - -End BiconnectLemmas. - -(**********) -(* Cycles *) -(**********) - -Section CycleLemmas. -Context {A : Type}. -Implicit Type g : pregraph A. - -(* elements in a cycle are interconnected *) -(* avoiding everything outside the cycle *) - -Lemma connect_cycle g xs : - cycle (edge g) xs -> - {in xs &, forall x y, y \in connect [predC xs] g x}. -Proof. -move=>C x y /rot_to [i q Hr]; rewrite -(mem_rot i) Hr => Hy. -have Hx : x \in xs by rewrite -(mem_rot i) Hr inE eqxx. -have /= Hp1: cycle (edge g) (x :: q) by rewrite -Hr rot_cycle. -have Dx : x \in dom g. -- by move: Hp1; rewrite rcons_path=>/andP [_ /edgeD][]. -case/splitPl: Hy Hp1 Hr=>r s Ey. -rewrite rcons_cat cat_path=>/andP [Hxr]. -rewrite Ey rcons_path; case/andP=>Hlx /= Ex Hr. -apply/connectP; split=>//=; first by rewrite negbK. -exists r; split=>// z Z. -by rewrite negbK -(mem_rot i) Hr inE mem_cat Z orbT. -Qed. - -Lemma connect_cycle0 g xs : - cycle (edge g) xs -> - {in xs &, forall x y, y \in connect pred0 g x}. -Proof. by move=>C x y Hx /(connect_cycle C Hx); apply: connect_sub. Qed. - -(* elements in a cycle are bi-interconnected *) -(* avoiding everything outside the cycle *) - -Lemma biconnect_cycle g xs : - cycle (edge g) xs -> - {in xs &, forall x y, y \in biconnect [predC xs] g x}. -Proof. by move=>C x y Hx Hy; rewrite /biconnect inE !(connect_cycle C). Qed. - -Lemma biconnect_cycle0 g xs : - cycle (edge g) xs -> - {in xs &, forall x y, y \in biconnect pred0 g x}. -Proof. by move=>C x y Hx /(biconnect_cycle C Hx); apply: biconnect_sub. Qed. - -Lemma biconnect_cycle2P p g x y : - x != y -> - reflect (exists xs : seq node, - [/\ y \in xs, cycle (edge g) (x :: xs) & - {in x :: xs, forall z : node, ~~ p z}]) - (y \in biconnect p g x). -Proof. -move=>Nxy; apply/(iffP idP)=>[|[ys][Py Cy Hys]]; last first. -- apply: biconnect_sub (biconnect_cycle Cy _ _). - - by move=>z; apply/contraL/Hys. - - by rewrite inE eqxx. - by rewrite inE Py orbT. -case/andP=>/connectP [Dy Hy][ys]; elim/last_ind: ys Nxy=>[|ys a IH] Nxy. -- by move/eqP: Nxy=>Nxy [_ /Nxy]. -rewrite rcons_path last_rcons; case=>/[swap] <-{a} /andP [Py Ex Hys]. -case/connectP=>Dx Hx [xs][Px Ey Hxs]; exists (xs ++ ys); split=>/=. -- have := mem_last x xs. - by rewrite -Ey inE eq_sym (negbTE Nxy) /= mem_cat=>->. -- by rewrite rcons_path cat_path last_cat -Ey Px Py Ex. -by move=>z; rewrite -mem_rcons rcons_cat mem_cat=>/orP []; -[apply: Hxs|apply: Hys]. -Qed. - -Lemma biconnect_cycle2P0 g x y : - x != y -> - reflect (exists2 xs, y \in xs & cycle (edge g) (x :: xs)) - (y \in biconnect pred0 g x). -Proof. -move=>Nxy; apply/(iffP (biconnect_cycle2P pred0 g Nxy)). -- by case=>xs [H1 H2 H3]; exists xs. -by case=>xs H1 H2; exists xs. -Qed. - -End CycleLemmas. - -(**************) -(* Acyclicity *) -(**************) - -(* graph is preacyclic if only self-loops are biconnected *) -Definition preacyclic A (g : pregraph A) := - all2rel (fun x y => (y \in biconnect pred0 g x) ==> (x == y)) (dom g). - -(* graph is acyclic if it doesn't even have self-loops *) -Definition acyclic A (g : pregraph A) := - preacyclic g && all (fun x => ~~ edge g x x) (dom g). - -Section AcyclicityLemmas. -Context {A : Type}. -Implicit Type g : pregraph A. - -Lemma preacyclicP g : - reflect {in dom g &, forall x y, y \in biconnect pred0 g x -> x = y} - (preacyclic g). -Proof. -apply: (iffP idP)=>[|H]. -- by move/allrelP=>H x y Dx Dy C; apply/eqP/(implyP _ C)/H. -by apply/allrelP=>x y Dx Dy; apply/implyP=>K; apply/eqP/H. -Qed. - -Lemma preacyclicL g1 g2 : - valid (g1 \+ g2) -> - preacyclic (g1 \+ g2) -> - preacyclic g1. -Proof. -move=>V /preacyclicP H. -apply/preacyclicP=>x y Dx Dy C; apply: H. -- by rewrite domUn inE V Dx. -- by rewrite domUn inE V Dy. -by apply: biconnectUnL V C. -Qed. - -Lemma preacyclicR g1 g2 : - valid (g1 \+ g2) -> - preacyclic (g1 \+ g2) -> - preacyclic g2. -Proof. by rewrite joinC; apply: preacyclicL. Qed. - -Lemma acyclicL g1 g2 : - valid (g1 \+ g2) -> - acyclic (g1 \+ g2) -> - acyclic g1. -Proof. -move=>V /andP [Hp /allP Ha]. -apply/andP; split; first by apply: preacyclicL Hp. -apply/allP=>x Dx; apply: contra (edgeUnL V) (Ha x _). -by rewrite domUn inE V Dx. -Qed. - -Lemma acyclicR g1 g2 : - valid (g1 \+ g2) -> - acyclic (g1 \+ g2) -> - acyclic g2. -Proof. by rewrite joinC; apply: acyclicL. Qed. - -(* graph is acyclic = graph has no cycles *) -Lemma acyclic_cycleP g : - reflect (forall x xs, x \in dom g -> ~~ cycle (edge g) (x :: xs)) - (acyclic g). -Proof. -apply: (iffP idP)=>[|H]; last first. -- apply/andP; split; last first. - - by apply/allP=>x Dx; apply: contra (H _ [::] Dx); rewrite /= =>->. - apply/preacyclicP=>x y Dx Dy By; apply/eqP/(contraLR _ By)=>{By} Nxy. - by apply/(biconnect_cycle2P0 _ Nxy); case=>/= xs Hx; apply/negP/H/Dx. -case/andP=>/preacyclicP Ng /allP Ne x xs Dx /=. -rewrite rcons_path; apply/negP=>/andP []. -case: xs=>[_|y xs /= /andP [Exy Px]]; first by apply/negP/Ne. -have Dy : y \in dom g by rewrite (edgeD Exy). -have : y \notin biconnect pred0 g x. -- by apply: contraL Exy=>/(Ng x y Dx Dy) <-; apply: Ne Dx. -apply: contraNnot=>Ex; apply: (biconnect_cycle0 (xs:=x::y::xs))=>/=. -- by rewrite Exy rcons_path Px Ex. -- by rewrite inE eqxx. -by rewrite !inE eqxx orbT. -Qed. - -Lemma acyclic_links g x : - acyclic g -> - x \notin links g x. -Proof. -case/andP=>_ /allP H. -have [Dx|Nx] := boolP (x \in dom g); last by rewrite linksND. -by apply: contra (H _ Dx)=>Lx; rewrite edge_links Dx Lx. -Qed. - -End AcyclicityLemmas. - -(***************) -(* N-pregraphs *) -(***************) - -(* getting the n-th link *) -Definition get_nth A (g : pregraph A) x := nth null (links g x). - -Section NpregraphLemmas. -Context {A : Type}. -Implicit Type g : pregraph A. - -Lemma getnth_mem0 g x n : - (get_nth g x n == null) || - (get_nth g x n \in links g x). -Proof. -case: (ltnP n (size (links g x)))=>Hm; last first. -- by rewrite /get_nth nth_default. -by rewrite mem_nth // orbT. -Qed. - -Lemma getnth_mem g x n : - get_nth g x n != null -> - get_nth g x n \in links g x. -Proof. by move=>H; move: (getnth_mem0 g x n); rewrite (negbTE H). Qed. - -(* n-pregraph is pregraph with global bound n *) -(* for the number of links of a node *) -Definition n_pregraphb (n : nat) g := - all (fun x => size (links g x) == n) (dom g). -Definition n_pregraph_axiom (n : nat) g := - {in dom g, forall x, size (links g x) = n}. - -Lemma npregraphP (n : nat) g : - reflect (n_pregraph_axiom n g) (n_pregraphb n g). -Proof. -apply: (iffP allP)=>H; first by move=>x /H /eqP. -by move=>x D; apply/eqP/H. -Qed. - -Lemma npregraphE n g : - n_pregraph_axiom n g <-> - {in map snd (range g), forall xs, size xs = n}. -Proof. -split=>H x; last by move=>Dx; apply: H (range_links Dx). -case/mem_seqP/Mem_map_inv=>-[v xs][->] /In_rangeX [k X]. -by rewrite (In_graph X); apply: H (In_dom X). -Qed. - -Lemma npregraphUn n g1 g2 : - n_pregraph_axiom n g1 -> - n_pregraph_axiom n g2 -> - n_pregraph_axiom n (g1 \+ g2). -Proof. -move=>H1 H2 z; rewrite domUn inE; case/andP=>V. -case/orP=>D. -- by rewrite /links findUnL // D H1. -by rewrite /links findUnR // D H2. -Qed. - -Lemma npregraphUnL n g1 g2 : - valid (g1 \+ g2) -> - n_pregraph_axiom n (g1 \+ g2) -> - n_pregraph_axiom n g1. -Proof. -rewrite !npregraphE=>V H _ /mem_seqP/Mem_map_inv [[v xs]][-> X]. -by apply: H; apply/mem_seqP/Mem_map/In_rangeL/X. -Qed. - -Lemma npregraphUnR n g1 g2 : - valid (g1 \+ g2) -> - n_pregraph_axiom n (g1 \+ g2) -> - n_pregraph_axiom n g2. -Proof. by rewrite joinC; apply: npregraphUnL. Qed. - -Lemma npregraphF n g x : - n_pregraph_axiom n g -> - n_pregraph_axiom n (free g x). -Proof. -rewrite !npregraphE=>H z /mem_seqP/Mem_map_inv [[v xs]][->]. -by case/In_rangeF=>k' _ /In_range X; apply/H/mem_seqP/Mem_map. -Qed. - -Lemma links_nth n g x : - n_pregraph_axiom n g -> - x \in dom g -> - links g x = map (get_nth g x) (iota 0 n). -Proof. -move=>H Dx; apply/(eq_from_nth (x0:=null))=>[|i Hi]. -- by rewrite size_map size_iota H. -by rewrite map_nth_iota0 ?H // -(H _ Dx) take_size. -Qed. - -(* n_pregraph gives rise to seprel *) - -Definition n_pregraph n g1 g2 := - n_pregraphb n g1 && n_pregraphb n g2. - -Lemma npregraph_is_sep n : seprel_axiom (@n_pregraph n). -Proof. -split=>[|x y V|x y V|x y z V]; rewrite /n_pregraph. -- by rewrite /n_pregraphb dom0. -- by rewrite andbC. -- by rewrite /n_pregraphb dom0; case/andP=>->. -case/andP=>Hx Hy /andP [_ Hz]; rewrite Hx Hy Hz /=. -apply/allP=>w; rewrite domUn inE (validX V) /=. -case/orP=>Dw. -- by rewrite /links findUnL ?(validX V) // Dw (allP Hy). -by rewrite /links findUnR ?(validX V) // Dw (allP Hz). -Qed. - -HB.instance Definition _ n := - isSeprel.Build (pregraph A) (n_pregraph n) (npregraph_is_sep n). - -End NpregraphLemmas. - -(********************) -(********************) -(* Binary pregraphs *) -(********************) -(********************) - -(* notation for left/right node of x *) -Definition sel2 A (m : Side) (g : pregraph A) x := get_nth g x m. -Notation lft g x := (sel2 LL g x). -Notation rgh g x := (sel2 RR g x). - -(* updating binary pregraph at node x *) - -(* if v = None, updates links, keeping the label *) -Definition upd2 A (g : pregraph A) x (v : option A) (lft rgh : node) := - if find x (labels g) is Some v' then - if v is Some w then upd x (w, [:: lft; rgh]) g - else upd x (v', [:: lft; rgh]) g - else undef. - -Lemma upd2_is_binary A (g : pregraph A) x v lft rgh : - n_pregraph_axiom 2 g -> - n_pregraph_axiom 2 (@upd2 A g x v lft rgh). -Proof. -move=>H z; rewrite /upd2 find_omf /omfx /=; case: (find x g)=>[[m lnk]|//]. -case: v=>[w|]; rewrite domU inE /graph.links findU; -case: (x =P 0)=>//= /eqP Nx; case: (z =P x)=>[_ V|_ Dz]. -- by rewrite V. -- by rewrite H. -- by rewrite V. -by rewrite H. -Qed. - -(* updating just contents *) -Notation updC g x v := (upd2 g x (Some v) (lft g x) (rgh g x)). -(* updating just left link *) -Notation updL g x l := (upd2 g x None l (rgh g x)). -(* updating just right link *) -Notation updR g x r := (upd2 g x None (lft g x) r). -(* updating just contents and left link *) -Notation updCL g x v l := (upd2 g x (Some v) l (rgh g x)). -(* updating just contents and right link *) -Notation updCR g x v r := (upd2 g x (Some v) (lft g x) r). - -Lemma find_upd2 A (g : pregraph A) x p v lf rg : - find x (upd2 g p (Some v) lf rg) = - if p \in dom g then - if x == p then Some (v, [:: lf; rg]) - else find x g - else None. -Proof. -rewrite /upd2/labels find_omf /omfx/=. -case: (dom_find p)=>[|[k w]]; first by rewrite find_undef. -by move/In_find=>H _; rewrite findU (In_cond H) (In_valid H). -Qed. - -Lemma find_upd2N A (g : pregraph A) x p lf rg : - find x (upd2 g p None lf rg) = - if find p (labels g) is Some v then - if x == p then Some (v, [:: lf; rg]) - else find x g - else None. -Proof. -rewrite /upd2/labels find_omf /omfx /=. -case: (dom_find p)=>[|[k w]]; first by rewrite find_undef. -by move/In_find=>H _; rewrite findU (In_cond H) (In_valid H). -Qed. - -Lemma sel2U A (g : pregraph A) (x p : node) v lf rg i : - sel2 i (upd2 g p v lf rg) x = - if x == p then - if x \in dom g then nth null [:: lf; rg] i - else null - else - if (p \in dom g) && (x \in dom g) then sel2 i g x - else null. -Proof. -rewrite /sel2/get_nth/graph.links/upd2 find_omf/omfx /=. -case: (x =P p)=>[->|/eqP N]. -- case: (dom_find p)=>[|w /In_find H _]; first by rewrite nth_nil. - by case: v=>[v|] /=; rewrite findU eqxx /= (In_cond H) (In_valid H). -case: (dom_find p g)=>/=; first by rewrite nth_nil. -case=>k w /In_find H _; case: v=>[v|]; -rewrite /oapp findU (negbTE N) (In_cond H); -by case: (dom_find x g)=>//; rewrite nth_nil. -Qed. - -Lemma In_links2 A (g : pregraph A) (x : node) v lks : - n_pregraph_axiom 2 g -> - (x, (v, lks)) \In g -> - lks = [:: lft g x; rgh g x]. -Proof. -move=>H X; rewrite (In_graph X). -by rewrite (links_nth H (In_dom X)). -Qed. - -(* laying binary pregraph onto heap *) - -Definition lay2_k {A} (v : A * seq node) : dynamic id := - idyn (v.1, (nth null v.2 0, nth null v.2 1)). -Definition lay2 {A} (g : pregraph A) : heap := mapv lay2_k g. -HB.instance Definition _ A := OmapFun.copy (@lay2 A) (@lay2 A). - -Lemma In_layX A (g : pregraph A) x (c : A) lft rgh : - (x, idyn (c, (lft, rgh))) \In lay2 g <-> - exists lks, - [/\ lft = nth null lks 0, - rgh = nth null lks 1 & - (x, (c, lks)) \In g]. -Proof. -rewrite In_omfX; split; last first. -- by case=>lks [->-> H]; exists (c, lks). -case=>-[w lks] H /Some_inj/inj_pair2 /= [<-<-<-]. -by exists lks. -Qed. - -Lemma In_lay A (g : pregraph A) x (c : A) lks : - (x, (c, lks)) \In g -> - (x, idyn (c, (nth null lks 0, nth null lks 1))) \In lay2 g. -Proof. by rewrite In_layX=>H; exists lks. Qed. - -Lemma In_lay2X A (g : pregraph A) x (v : A) lft rgh : - n_pregraph_axiom 2 g -> - (x, idyn (v, (lft, rgh))) \In lay2 g <-> - (x, (v, [:: lft; rgh])) \In g. -Proof. -move=>H; split; last first. -- by move=>X; apply/In_layX; exists [:: lft; rgh]. -case/In_layX=>lks [L R X]. -rewrite -(_ : lks = [:: lft; rgh]) //. -rewrite (In_graph X) (links_nth H (In_dom X)) /=. -by rewrite /get_nth -(In_graph X) -L -R. -Qed. - -Lemma In_lay2 A (g : pregraph A) x (c : A) lft rgh : - n_pregraph_axiom 2 g -> - (x, (c, [:: lft; rgh])) \In g -> - (x, idyn (c, (lft, rgh))) \In lay2 g. -Proof. by move=>H /In_lay2X-/(_ H). Qed. - -Lemma lay2_eta A (g : pregraph A) x c pl pr : - n_pregraph_axiom 2 g -> - (x, (c, [:: pl; pr])) \In g -> - exists h, lay2 g = x :-> (c, (pl, pr)) \+ h. -Proof. by move=>H Gx; eexists _; apply/heap_eta2/In_find/In_lay2. Qed. - -(* lay with update that changes label *) -Lemma lay2CU A (g : pregraph A) x l r c : - lay2 (upd2 g x (Some c) l r) = - if x \in dom g then - upd x (idyn (c, (l, r))) (lay2 g) - else undef. -Proof. -rewrite /lay2/upd2 find_omf/omfx /=. -case: (dom_find x)=>[|v]; first by rewrite omap_undef. -move/In_find=>E _; rewrite (upd_eta x). -rewrite omapPtUn -(upd_eta x) validU. -rewrite (In_cond E) (In_valid E) /=. -rewrite (upd_eta x) !omap_free !omap_omap. -congr (_ \+ _); apply/eq_in_omap. -by case=>k w /= H; case: (k =P x). -Qed. - -(* lay with update that keeps label fixed *) -Lemma lay2U A (g : pregraph A) x l r : - lay2 (upd2 g x None l r) = - if find x (labels g) is Some c then - upd x (idyn (c, (l, r))) (lay2 g) - else undef. -Proof. -rewrite /olabel/lay2/upd2 find_omf /=. -case: (dom_find x g)=>[/In_findN D|v /In_find E _]. -- by rewrite omap_undef. -rewrite /omfx /= (upd_eta x). -rewrite omapPtUn -(upd_eta x) validU. -rewrite (In_cond E) (In_valid E) /=. -rewrite (upd_eta x) !omap_free !omap_omap. -congr (_ \+ _); apply/eq_in_omap. -by case=>k w /= H; case: (k =P x). -Qed. - -Lemma lay2PtUn A (g : pregraph A) x l r c : - lay2 (x &-> (c, [:: l; r]) \+ g) = - x :-> (c, (l, r)) \+ lay2 g. -Proof. -rewrite omfPtUn /=; case: ifP=>// V; set j := (_ \+ _). -case: (normalP j)=>[->//|]. -rewrite !validPtUn valid_omap dom_omf_some // in V *. -by rewrite V. -Qed. - -(**********) -(**********) -(* Graphs *) -(**********) -(**********) - -(* x is in_node if it's in the graph or is null *) -Definition in_node A (g : pregraph A) (x : node) := - (x == null) || (x \in dom g). - -(* x is out-node if no edge goes into it *) -Definition out_node A (g : pregraph A) (x : node) := - {in map snd (range g), forall xs, x \notin xs}. - -(* pregraph is graph if *) -(* all nodes in all adjacency lists are in-nodes *) -Definition graph_axiom A (g : pregraph A) := - forall xs x, xs \in map snd (range g) -> x \in xs -> in_node g x. - -HB.mixin Record isGraph A (g : pregraph A) := { - graph_subproof : graph_axiom g}. -#[short(type=graph)] -HB.structure Definition Graph A := {g of isGraph A g}. - -Section GraphLemmas. -Context {A : Type}. -Implicit Type g : pregraph A. - -(* removing out-node causes no dangling edges; *) -(* hence preserves graph axiom *) -Lemma graphF g x : - out_node g x -> - graph_axiom g -> - graph_axiom (free g x). -Proof. -move=>Hna Ha xs q /mem_seqP/MapP [[v vs]] /In_rangeF [k'] N -/In_range/(Mem_map snd)/mem_seqP Hs ->{xs} Hq. -move: (Ha vs q Hs Hq) (Hna _ Hs)=>{}Hs. -rewrite /in_node domF inE in Hs *. -by case: (x =P q) Hq=>//= ->->. -Qed. - -(* reachable component of a graph is a graph *) -Lemma graph_subconnect g x : - valid g -> - graph_axiom g -> - graph_axiom (subconnect g x). -Proof. -have E : g = subconnect g x \+ um_filterk - (negb \o connect pred0 g x) g by apply: umfilt_predC. -move=>V Ha xs /= n Hxs Hn; have {}Dn : in_node g n. -- case/mem_seqP/MapP: Hxs=>-[a b Hxs /= H]. - apply: Ha Hn; apply/mem_seqP/MapP; exists (a, b)=>//. - by rewrite E; apply/In_rangeL/Hxs; rewrite -E. -case/mem_seqP/MapP: Hxs=>-[v vs] /In_rangeX [k] /In_umfiltX [/= Ck]. -move/In_graph=>->{vs} Hf; rewrite /in_node in Dn *. -case/boolP: (n == null) Dn=>//= Hnn Dn. -case: (connectD Ck)=>_ Dk. -rewrite /subconnect dom_umfiltk inE /= Dn andbT. -apply: connect_trans Ck _; apply/connectP; split=>//. -by exists [:: n]; split=>//=; rewrite edge_links Dn -Hf Hn. -Qed. - -Lemma graph_links g x y : - graph_axiom g -> - y \in links g x -> - in_node g y. -Proof. by move=>H /[dup]/linksD Dx; apply/H/range_links. Qed. - -Lemma graph_children g x y : - graph_axiom g -> - (y \in children g x) = (y != null) && (y \in links g x). -Proof. -move=>H; rewrite mem_filter -!(andbC (y \in links _ _)). -case Ly : (y \in links g x)=>//=. -move/(graph_links H): Ly; rewrite /in_node. -by case: (y =P null)=>//= ->; rewrite cond_dom. -Qed. - -Lemma graph_getnth g x n : - graph_axiom g -> - in_node g (get_nth g x n). -Proof. -case: (ltnP n (size (links g x)))=>Hm; last first. -- by rewrite /get_nth nth_default. -by move/(graph_links (x:=x)); apply; rewrite mem_nth. -Qed. - -End GraphLemmas. - diff --git a/examples/schorr.v b/examples/schorr.v deleted file mode 100644 index 7fc5e32..0000000 --- a/examples/schorr.v +++ /dev/null @@ -1,264 +0,0 @@ -(* -Copyright 2024 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 HB Require Import structures. -From Coq Require Import ssreflect ssrfun. -From mathcomp Require Import eqtype ssrnat ssrbool seq path bigop. -From pcm Require Import options axioms pred prelude seqperm seqext. -From pcm Require Import pcm unionmap natmap heap autopcm automap. -From htt Require Import options model heapauto graph. - -(*****************) -(* helper lemmas *) -(*****************) - -Lemma eq_connect_aux A (g1 g2 : pregraph A) p: - valid (g1 \+ g2) -> - {subset dom g2 <= predC p} -> - um_filterk p (g1 \+ g2) = um_filterk p g1. -Proof. -move=>V /(umfiltk_subdom0 p (validR V)) S. -by rewrite umfiltUn // S unitR. -Qed. - -Lemma connectUn_eq A (g g1 g2 : pregraph A) (p : pred node) x : - valid (g \+ g1) -> - valid (g \+ g2) -> - {subset dom g1 <= p} -> - {subset dom g2 <= p} -> - connect p (g \+ g1) x =i connect p (g \+ g2) x. -Proof. -move=>V1 V2 S1 S2 z. -rewrite -connect_umfiltk eq_connect_aux //=; last first. -- by move=>y /S1; rewrite inE negbK. -rewrite -[RHS]connect_umfiltk eq_connect_aux //. -by move=>y /S2; rewrite !inE negbK. -Qed. - -(****************) -(* Schorr-Waite *) -(****************) - -(* type of markings *) -(* U = unmarked *) -(* M m = marked, but still exploring m (LL or RR) subgraph *) -(* MM = fully marked and explored both subgraphs *) -Inductive mark := U | M of Side | MM. -(* decidable equality on marks *) -Definition eq_mark l1 l2 : bool := - if (l1, l2) is ((U,U)|(M LL, M LL)|(M RR, M RR)|(MM, MM)) - then true else false. -Lemma eq_markP : Equality.axiom eq_mark. -Proof. -case; case=>//=; try by [constructor]; -by case=>[|[]|]; constructor. -Qed. -HB.instance Definition _ := hasDecEq.Build mark eq_markP. - -Notation pregraph := (pregraph mark). - -Definition marked (g : pregraph) : pregraph := - um_filterv (fun v => v.1 != U) g. -HB.instance Definition _ := OmapFun.copy marked marked. - -Definition omark (g : pregraph) x : option Side := - if olabel g x is Some (M m) then Some m else None. - -Lemma In_omark (g : pregraph) x c lks : - (x, (c, lks)) \In g -> - omark g x = if c is M m then Some m else None. -Proof. by rewrite /omark=>/In_olabel=>->. Qed. - - -(* old helpers kept for now, but of unclear utility *) - -(* -Definition ch (g : pregraph) x y : bool := - if omark g y is Some m then sel2 m g y == x else false. - -Lemma chN0 g x y : - ch g x y -> - y != null. -Proof. -rewrite /ch/omark/olabel find_omf /omfx /=. -by case: (dom_find y)=>[//|v] /In_find/In_cond. -Qed. - -Lemma ch_fun g a b x : - ch g a x -> - ch g b x -> - a = b. -Proof. by rewrite /ch; case: (omark g x)=>// v /eqP -> /eqP. Qed. - -Lemma ch_path g s x : - x \in s -> - path (ch g) null s -> - exists y, y \in belast null s /\ ch g y x. -Proof. exact: path_prev. Qed. - -Lemma ch_path_uniq g s : - path (ch g) null s -> - uniq (null :: s). -Proof. by apply: path_uniq; [apply: chN0|apply: ch_fun]. Qed. -*) - - -(* stack contains all and only nodes marked L or R *) -Definition stack_marked (g : pregraph) (s : seq node) := - forall x, x \in s = isSome (omark g x). - -(* consecutive stack nodes respect marking *) -Definition stack_consec (g : pregraph) (s : seq node) := - forall x y m, - (* if x is just under y in the stack *) - consec (null :: s) x y -> - (* and y is marked by m *) - omark g y = Some m -> - (* then x is y's child, left or right determined by m *) - x = sel2 m g y. - -(* graph g differs from g0 only in reversing edges on the stack *) -Definition graph_diff (g0 g : pregraph) (s : seq node) t := - [/\ (* graphs have equal nodes, and *) - dom g0 = dom g & forall x, - (* if x is marked m, then *) - if omark g x is Some m then - (* m-child of x in g0 is successor of x (or t, if x last) *) - [/\ consec (rcons s t) x (sel2 m g0 x) & - (* graphs agree on children of flipped marking *) - sel2 (flip m) g x = sel2 (flip m) g0 x] - (* otherwise, if x is unmarked or fully marked, then *) - (* graphs agree on children of x *) - else forall m, sel2 m g x = sel2 m g0 x]. - -(* each unmarked node is reachable either from t *) -(* or from stack's right spine, *) -(* in both cases by avoiding marked nodes *) -Definition reach (g : pregraph) (s : seq node) (t : node) := - forall x, - (* if x is unmarked node in g *) - (x, U) \In labels g -> - (* then x is reachable from t avoiding marked nodes *) - x \in connect [dom (marked g)] g t \/ - (* or x is reachable from some right child of a node in stack s *) - (* avoiding marked nodes *) - exists2 y, y \in s & - x \in connect [dom (marked g)] g (rgh g y). - -Definition shape (g0 g : pregraph) (r p t : node) := - fun h => exists s, - [/\ h = lay2 g, p = last null s, r = head t s, uniq (null :: s), - stack_marked g s, stack_consec g s, - graph_diff g0 g s t, reach g s t, - n_pregraph_axiom 2 g & graph_axiom g]. - -(* helper lemma of unclear utility *) - -(* -Lemma stack_path g s : - stack_marked g s -> - stack_consec g s -> - uniq (null :: s) -> - path (ch g) null s. -Proof. -move=>H1 H2 U; apply: consec_path=>//= x y Dx Dy C. -move: Dy; rewrite H1 /ch; case D : (omark g y)=>[a|//]. -by rewrite -(H2 x y). -Qed. -*) - - -Program Definition pop (p t : ptr): - STsep {g0 g r} (fun h => - [/\ shape g0 g r p t h, - (p, M RR) \In labels g & - t \in dom (marked g) \/ t = null], - [vfun res h => - shape g0 (updCR g p MM t) r res.1 res.2 h /\ - res = (rgh g p, p)]) := - Do ('(_, (l, r)) <-- read (mark*(node*node)) p; - p ::= (MM, (l, t));; - ret (r, p)). -Next Obligation. -move=>p t [g0][g][r][_][[s [->]]] Ep Er /= /andP [U1 U2] -Sm Sc Gd Rc G2 G Pm D /=; case/In_labelsX: Pm=>lks /[dup] /(In_links2 G2) -> Pm. -(* prepare for popping p from the end of s to obtain ss *) -move: (In_cond Pm)=>/= Np; case/(rcons_lastN Np): Ep=>ss ?; subst s. -rewrite headI /= in Er; rewrite mem_rcons inE negb_or eq_sym Np /= in U1. -rewrite rcons_uniq in U2; case/andP: U2=>U2 U3. -case: (lay2_eta G2 Pm)=>h /[dup] E ->; do !step; move=>V. -(* the new stack is ss *) -split=>//; exists ss; split=>//=. -- by rewrite lay2CU (In_dom Pm) upd_eta E freePtUn ?(validX V). -- rewrite (Sc (last null ss) p RR) ?(In_omark Pm) ?consec_rcons //=. - by rewrite mem_rcons inE negb_or eq_sym Np U1 rcons_uniq U2. -- by rewrite U1 U3. -- move=>x; move: (Sm x); rewrite mem_rcons inE. - rewrite /omark/olabel !find_omf find_upd2 (In_dom Pm) /omfx/=. - by case: (x =P p) U2=>// -> /negbTE ->. -- move=>x y m C; move: (Sm y); rewrite mem_rcons inE. - rewrite sel2U /omark/olabel !find_omf find_upd2 (In_dom Pm) /=. - case: (y =P p)=>// /eqP Ny; case: (dom_find y)=>//= [][][] //=. - move=>k v /In_find H _ S [<-]; apply: Sc (In_omark H). - rewrite -rcons_cons consec_ext ?inE ?S ?orbT //. - by rewrite rcons_uniq /= inE negb_or Np U1 U2 U3. - - -UP TO HERE - - - - - - -Program Definition swing (p t : ptr): - STsep {g0 g m r} (fun h => - [/\ shape g0 g m r p t h, - (p, L) \In m & - t \in dom m \/ t = null], - [vfun res h => - let: g' := upd p [:: t; gl g p] g in - let: m' := upd p R m in - shape g0 g' m' r res.1 res.2 h /\ - res = (t, gl g t) ]) := - Do (q <-- read (A:=node*node*mark) p; - p ::= (t, q.1.1, R);; - ret (p, q.1.2)). -Next Obligation. -move=>p t [g0][g][m][r][/= h][H P D]. -case: H=>stack [-> H1 H2 H3 H4 H5 H6 H7 H8 H9]. -move: (H1 p (In_dom P))=>/(glD H8) /(In_layX2 H8 H1) [h' ->]. -do 3!step; split=>//. -exists stack; split =>//. -admit. -Admitted. - -Program Definition push (p t : ptr): - STsep {g0 g m r} (fun h => - [/\ shape g0 g m r p t h, - t \in dom g & - t \notin dom m], - [vfun res h => - let: g' := upd t [:: p; gr g t] g in - let: m' := upd t L m in - shape g0 g' m' r res.1 res.2 h /\ - res = (t, gl g t) ]) := - Do (q <-- read (A:=node*node*mark) t; - t ::= (p, q.1.2, L);; - ret (t, q.1.1)). -Next Obligation. -move=>p t [g0][g][m][r][/= h][H P D]. -case: H=>stack [-> H1 H2 H3 H4 H5 H6 H7 H8 H9]. -admit. -Admitted. - diff --git a/pcm/automap.v b/pcm/automap.v deleted file mode 100644 index 2666050..0000000 --- a/pcm/automap.v +++ /dev/null @@ -1,764 +0,0 @@ -(* -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 eqtype seq. -From pcm Require Import options pred prelude. -From pcm Require Export auto. -From pcm Require Import pcm unionmap natmap. - -(**************************************************************************) -(**************************************************************************) -(* 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 *) -(* *) -(**************************************************************************) -(**************************************************************************) - -(* DEVCOMMENT *) -(* 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. *) -(* /DEVCOMMENT *) - -(* 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. *) -(* *) -(* The context of keys is thus seq K. The context of vars is seq U. *) - -Section ReflectionContexts. -Variables (K : ordType) (C : pred K) (T : Type) (U : union_map C T). - -Structure ctx := Context {keyx : seq K; varx : seq U}. - -Definition empx := Context [::] [::]. - -(* because contexts grow during computation, *) -(* we need a notion of sub-context *) - -Definition sub_ctx (i j : ctx) := - Prefix (keyx i) (keyx j) /\ Prefix (varx i) (varx j). - -Lemma sc_refl i : sub_ctx i i. -Proof. by []. Qed. - -Lemma sc_trans i j k : sub_ctx i j -> sub_ctx j k -> sub_ctx i k. -Proof. -by case=>K1 V1 [K2 V2]; split; [move: K2 | move: V2]; apply: Prefix_trans. -Qed. - -End ReflectionContexts. - -(* Keys and map variables are syntactified as De Bruijn indices in the context. *) -(* Disjoint union is syntactified as concatenation of lists. *) -(* *) -(* Pts n v : syntax for "key indexed the context under number n" \-> v *) -(* Var n : syntax for "expression indexed in the context under number n" *) - -(* now for reflection *) - -Section Reflection. -Variables (K : ordType) (C : pred K) (T : Type) (U : union_map C T). -Implicit Type i : ctx U. - -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 undef - | Var n => if onth (varx i) n is Some f then f else undef - end. - -(* main interpretation function *) -Notation fx i := (fun t f => interp' i t \+ f). -Definition interp i ts := foldr (fx i) Unit ts. - -Lemma fE i ts x : foldr (fx i) x ts = x \+ interp i ts. -Proof. by elim: ts x=>[|t ts IH] x; rewrite /= ?unitR // IH joinCA. Qed. - -Lemma interp_rev i ts : interp i (rev ts) = interp i ts. -Proof. -elim: ts=>[|t ts IH] //=; rewrite rev_cons -cats1. -by rewrite /interp foldr_cat fE /= unitR IH. -Qed. - -(* we also need a pretty-printer, which works the same as interp *) -(* but removes trailing Unit's *) - -Fixpoint pprint i ts := - if ts is t :: ts' then - if ts' is [::] then interp' i t else interp' i t \+ pprint i ts' - else Unit. - -Lemma pp_interp i ts : pprint i ts = interp i ts. -Proof. by elim: ts=>[|t ts /= <-] //; case: ts=>//; rewrite unitR. Qed. - -Definition key n t := if t is Pts m _ then n == m else false. -Definition var n t := if t is Var m then n == m else false. - -Definition kfree n t := rfilter (key n) t. -Definition vfree n t := rfilter (var n) t. -Arguments kfree /. -Arguments vfree /. - -Lemma keyN i n ts : ~~ has (key n) ts -> interp i (kfree n ts) = interp i ts. -Proof. by elim: ts=>[|t ts IH] //=; case: ifP=>//= _ /IH ->. Qed. - -Lemma varN i n ts : ~~ has (var n) ts -> interp i (vfree n ts) = interp i ts. -Proof. by elim: ts=>[|t ts IH] //=; case: ifP=>//= _ /IH ->. Qed. - -Lemma keyP i n k ts : - has (key n) ts -> onth (keyx i) n = Some k -> - exists v, interp i ts = pts k v \+ interp i (kfree n ts). -Proof. -elim: ts=>[|t ts IH] //=; case: ifP=>[|_ H]. -- by case: t=>//= _ v /eqP <- _ ->; exists v. -by case/(IH H)=>v ->; exists v; rewrite joinCA. -Qed. - -Lemma varP i n u ts : - has (var n) ts -> onth (varx i) n = Some u -> - interp i ts = u \+ interp i (vfree n ts). -Proof. -elim: ts=>[|t ts IH] //=; case: ifP=>[|_ H]. -- by case: t=>//= _ /eqP <- _ ->. -by move/(IH H)=>->; rewrite joinCA. -Qed. - -(* interpretation is invariant under context weakening *) -(* under assumption that the interpreted term is well-formed *) - -Definition wf i t := - match t with - Pts n _ => n < size (keyx i) - | Var n => n < size (varx i) - end. - -Lemma sc_wf i j ts : sub_ctx i j -> all (wf i) ts -> all (wf j) ts. -Proof. -case=>/Prefix_size H1 /Prefix_size H2; elim: ts=>[|t ts IH] //=. -case/andP=>H /IH ->; rewrite andbT. -by case: t H=>[n v|v] H; apply: leq_trans H _. -Qed. - -Lemma sc_interp i j ts : - sub_ctx i j -> all (wf i) ts -> interp i ts = interp j ts. -Proof. -case=>H1 H2; elim: ts=>[|t ts IH] //= /andP [H] /IH ->. -by case: t H=>[n v|n] /= /Prefix_onth <-. -Qed. - -Lemma valid_wf i ts : valid (interp i ts) -> all (wf i) ts. -Proof. -elim: ts=>[|t ts IH] //= V; rewrite (IH (validR V)) andbT. -case: t {V IH} (validL V)=>[n v|n] /=; -by case X : (onth _ _)=>[a|]; rewrite ?(onth_size X) // valid_undef. -Qed. - -Lemma wf_kfree i n ts : all (wf i) ts -> all (wf i) (kfree n ts). -Proof. by elim: ts=>//= t ts IH; case: ifP=>_ /andP [] //= -> /IH ->. Qed. - -Lemma wf_vfree i n ts : all (wf i) ts -> all (wf i) (vfree n ts). -Proof. by elim: ts=>//= t ts IH; case: ifP=>_ /andP [] //= -> /IH ->. Qed. - -(* sometimes we want to get keys in a list, not in a predicate *) - -Definition getkeys := - foldr (fun t ks => if t is Pts k _ then k :: ks else ks) [::]. - -Lemma has_getkeys ts n : n \in getkeys ts = has (key n) ts. -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 validX *) -(* and domeqX and invalidX respectively. *) -(**********************************************************************) -(**********************************************************************) - -(* subterm is purely functional version of validX *) - -Section Subterm. -Variables (K : ordType) (C : pred K) (T : Type) (U : union_map C T). -Implicit Types (i : ctx U) (ts : seq (term T)). - -Fixpoint subterm ts1 ts2 := - match ts1 with - Pts n _ :: tsx1 => - if has (key n) ts2 then subterm tsx1 (kfree n ts2) else false - | Var n :: tsx1 => - if has (var n) ts2 then subterm tsx1 (vfree n ts2) else false - | [::] => true - end. - -(* the procedure is sound for deciding wf subterms *) -Lemma subterm_sound i ts1 ts2 : - all (wf i) ts1 -> all (wf i) ts2 -> subterm ts1 ts2 -> - exists u, dom_eq (interp i ts1 \+ u) (interp i ts2). -Proof. -elim: ts1 ts2=>[|t ts1 IH] ts2 /= A1 A2. -- by exists (interp i ts2); rewrite unitL. -case/andP: A1; case: t=>[n v|n] /= /size_onth [k] X A1; -rewrite X; case: ifP=>Y //. -- case: (keyP Y X)=>w -> /(IH _ A1 (wf_kfree n A2)) [xs D]. - by exists xs; rewrite -joinA; apply: domeqUn D; apply: domeqPt. -move: (varP Y X)=>-> /(IH _ A1 (wf_vfree n A2)) [xs D]. -by exists xs; rewrite -joinA; apply: domeqUn D. -Qed. - -End Subterm. - -(* subtract is purely functional version of domeqX *) - -Section Subtract. -Variables (K : ordType) (C : pred K) (T : Type) (U : union_map C T). -Implicit Types (i : ctx U) (ts : seq (term T)). - -(* We need a subterm lemma that returns the uncancelled stuff from *) -(* both sides. xs accumulates uncancelled part of ts1, starting as nil *) - -Fixpoint subtract ts1 ts2 xs := - match ts1 with - Pts n v :: tsx1 => - if has (key n) ts2 then subtract tsx1 (kfree n ts2) xs - else subtract tsx1 ts2 (Pts n v :: xs) - | Var n :: tsx1 => - if has (var n) ts2 then subtract tsx1 (vfree n ts2) xs - else subtract tsx1 ts2 (Var T n :: xs) - | [::] => (xs, ts2) - end. - -(* below, the existentially quantified u is the cancelled part *) -Lemma subtract_sound i ts1 ts2 rs1 rs2 xs : - all (wf i) ts1 -> all (wf i) ts2 -> - subtract ts1 ts2 xs = (rs1, rs2) -> - exists u, dom_eq (interp i ts1 \+ interp i xs) (interp i rs1 \+ u) /\ - dom_eq (interp i ts2) (interp i rs2 \+ u). -Proof. -elim: ts1 ts2 xs=>[|t ts1 IH] ts2 xs /= A1 A2. -- by case=><-<-; exists Unit; rewrite unitL !unitR. -case/andP: A1; case: t=>[n v|n] /= /size_onth [x X] A1; rewrite X; case: ifP=>Y. -- case: (keyP Y X)=>w -> /(IH _ _ A1 (wf_kfree n A2)) [u][H1 H2]. - exists (pts x v \+ u); rewrite -joinA !(pull (pts x _)). - by split=>//; apply: domeqUn=>//; apply: domeqPt. -- by case/(IH _ _ A1 A2)=>u [/= H1 H2]; rewrite X joinCA joinA in H1; exists u. -- move: (varP Y X)=>-> /(IH _ _ A1 (wf_vfree n A2)) [u][H1 H2]. - by exists (x \+ u); rewrite -joinA !(pull x); split=>//; apply: domeqUn. -by case/(IH _ _ A1 A2)=>u [/= H1 H2]; rewrite X joinCA joinA in H1; exists u. -Qed. - -End Subtract. - -(* invalid is a purely functional test of invalidX *) - -Section Invalid. -Variables (K : ordType) (C : pred K) (T : Type) (U : union_map C T). -Implicit Types (i : ctx U) (t : term T) (ts : seq (term T)). - -Definition undefx i t := - if t is Var n then - if onth (varx i) n is Some x then undefb x else false - else false. - -Definition isundef i ts := ~~ uniq (getkeys ts) || has (undefx i) ts. - -Lemma isundef_sound i ts : - all (wf i) ts -> isundef i ts -> ~~ valid (interp i ts). -Proof. -elim: ts=>[|t ts IH] //= /andP [W A]. -case: t W=>[n v|n] /= /size_onth [k] X; rewrite /isundef /= X; last first. -- rewrite orbCA=>H; case: validUn=>// V; rewrite (negbTE (IH A _)) //. - 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_undef2 undef_join valid_undef. -Qed. - -End Invalid. - - -(********************************) -(********************************) -(* Canonical structure programs *) -(********************************) -(********************************) - -(* 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) (C : pred K) (T : Type) (U : union_map C T). -Implicit Types (i : ctx U) (ts : seq (term T)). - -(* a tagging structure to control the flow of computation *) -Structure tagged_map := Tag {untag : U}. - -Local Coercion untag : tagged_map >-> UMC.sort. - -(* in reversed order; first test for unions, then empty, pts and vars *) -Definition var_tag := Tag. -Definition key_tag := var_tag. -Definition empty_tag := key_tag. -Canonical Structure union_tag hc := empty_tag hc. - -(* Main structure *) -(* - i : input context *) -(* - j : output context *) -(* - ts : syntactification of map_of using context j *) - -Definition axiom i j ts (pivot : tagged_map) := - [/\ interp j ts = pivot :> U, sub_ctx i j & all (wf j) ts]. -Structure form i j ts := Form {pivot : tagged_map; _ : axiom i j ts pivot}. - -Local Coercion pivot : form >-> tagged_map. - -(* check for union *) - -Lemma union_pf i j k ts1 ts2 (f1 : form i j ts1) (f2 : form j k ts2) : - axiom i k (ts1 ++ ts2) (union_tag (untag f1 \+ untag f2)). -Proof. -case: f1 f2=>_ [<- S1 W1][_][<- S2 W2]; split. -- by rewrite /interp foldr_cat fE joinC -(sc_interp S2 W1). -- by apply: sc_trans S1 S2. -by rewrite all_cat (sc_wf S2 W1) W2. -Qed. - -Canonical union_form i j k ts1 ts2 f1 f2 := - Form (@union_pf i j k ts1 ts2 f1 f2). - -(* check if reached empty *) - -Lemma empty_pf i : axiom i i [::] (empty_tag Unit). -Proof. by []. Qed. - -Canonical empty_form i := Form (@empty_pf i). - -(* check for pts k v *) - -Lemma pts_pf vars keys1 keys2 k v (f : xfind keys1 keys2 k): - axiom (Context keys1 vars) (Context keys2 vars) - [:: Pts k v] (key_tag (pts (xuntag f) v)). -Proof. by case: f=>p [E H]; split=>//=; rewrite ?E ?unitR // (onth_size E). Qed. - -Canonical pts_form vars keys1 keys2 k v f := - Form (@pts_pf vars keys1 keys2 k v f). - -(* check for var *) - -Lemma var_pf keys vars1 vars2 n (f : xfind vars1 vars2 n) : - axiom (Context keys vars1) (Context keys vars2) - [:: Var T n] (var_tag (xuntag f)). -Proof. by case: f=>p [E H]; split=>//=; rewrite ?E ?unitR // (onth_size E). Qed. - -Canonical var_form keys vars1 vars2 v f := Form (@var_pf keys vars1 vars2 v f). - -End Syntactify. - -Module Exports. -Coercion untag : tagged_map >-> UMC.sort. -Coercion pivot : form >-> tagged_map. -Canonical union_tag. -Canonical union_form. -Canonical empty_form. -Canonical pts_form. -Canonical var_form. -End Exports. -End Syntactify. - -Export Syntactify.Exports. - -(*********************) -(* Automating validX *) -(*********************) - -(* validX is a refined lemma for subterm checking which *) -(* automatically discharges the spurious argument from above *) - -Module ValidX. -Section ValidX. -Variables (K : ordType) (C : pred K) (T : Type) (U : union_map C T). -Implicit Types (j : ctx U) (ts : seq (term T)). -Notation form := Syntactify.form. -Notation untag := Syntactify.untag. - -(* The rform structure has two important components: *) -(* *) -(* -- a packed/hoisted map m, which will be reified into the ts2 argument *) -(* of subterm ts2 ts1 *) -(* *) -(* -- a boolean b, which will be instantiated with true in the validX *) -(* lemma, and will be unified with subterm ts2 ts1 in the start *) -(* instance *) -(* *) -(* The other components of rform are j ts1 and pivot, which are forced by *) -(* needing to compose the proofs, but behave plainly in unification. *) - -Structure packed_map (m : U) := Pack {unpack : U}. -Canonical equate (m : U) := Pack m m. - -Definition raxiom j ts1 m (b : bool) (pivot : packed_map m) := - all (wf j) ts1 -> valid (interp j ts1) -> b -> valid (unpack pivot). - -Structure rform j ts1 m b := - RForm {pivot :> packed_map m; _ : raxiom j ts1 b pivot}. - -(* start instance: note how subterm ts2 ts1 is unified with *) -(* the boolean component of rform *) - -Lemma start_pf j k ts1 ts2 (f2 : form j k ts2) : - @raxiom j ts1 (untag f2) (subterm ts2 ts1) (equate f2). -Proof. -case: f2=>f2 [<- S A2] A1; rewrite (sc_interp S A1)=>V. -case/(subterm_sound A2 (sc_wf S A1))=>xs. -by case/domeqP; rewrite V=>/validL ->. -Qed. - -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=>/[swap]; exact: joinKx. Qed. - -Lemma joinxK' (x x1 x2 : U) : x \+ x1 = x \+ x2 -> valid (x \+ x1) -> x1 = x2. -Proof. by move=>/[swap]; exact: joinxK. 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) (C : pred K) (T : Type) (U : union_map C T). -Implicit Types (j : ctx U) (ts : seq (term T)). -Notation form := Syntactify.form. -Notation untag := Syntactify.untag. - -(* the main lemma; note how the boolean component of rform is set to true *) - -Lemma validX m j ts1 (f1 : form (empx U) j ts1) (g: rform j ts1 m true) : - valid (untag f1) -> valid (unpack (pivot g)). -Proof. by case: g f1; case=>pivot H [f1][<- Sc A] /(H A); apply. Qed. - -End Exports. - -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 domeqX *) -(*********************) - -Module DomeqX. -Section DomeqX. -Variables (K : ordType) (C : pred K) (T : Type) (U : union_map C T). -Implicit Types (j : ctx U) (ts : seq (term T)). -Notation form := Syntactify.form. -Notation untag := Syntactify.untag. - -Structure packed_map (m : U) := Pack {unpack : U}. -Canonical equate (m : U) := Pack m m. - -(* b is the pair of residual terms *) -Definition raxiom j k ts1 m b (pivot : packed_map m) := - all (wf j) ts1 -> sub_ctx j k /\ - (dom_eq (interp k b.1) (interp k b.2) -> - dom_eq (interp k ts1) (unpack pivot)). - -Structure rform j k ts1 m b := - RForm {pivot :> packed_map m; _ : raxiom j k ts1 b pivot}. - -(* start instance: note how subtract ts1 ts2 [::] is unified with *) -(* the b component of rform thus passing the residual terms *) - -Lemma start_pf j k ts1 ts2 (f2 : form j k ts2) : - @raxiom j k ts1 (untag f2) (subtract ts1 ts2 [::]) (equate f2). -Proof. -case: f2=>f2 [<- S A2]; case E : (subtract _ _ _)=>[rs1 rs2] A1; split=>//. -case/(subtract_sound (sc_wf S A1) A2): E=>ys [/= D1 D2 D]. -rewrite unitR in D1; apply: domeq_trans D1 _. -rewrite domeq_symE; apply: domeq_trans D2 _. -by rewrite domeq_symE; apply: domeqUn. -Qed. - -Canonical start j k ts1 ts2 f2 := RForm (@start_pf j k ts1 ts2 f2). - -End DomeqX. - -Module Exports. -Canonical equate. -Canonical start. - -Section Exports. -Variables (K : ordType) (C : pred K) (T : Type) (U : union_map C T). -Implicit Types (j : ctx U) (ts : seq (term T)). -Notation form := Syntactify.form. -Notation untag := Syntactify.untag. - -(* the main lemma; notice how residuals rs1, rs2 are passed to g to compute *) - -Lemma domeqX m j k rs1 rs2 ts1 (f1 : form (empx U) j ts1) - (g : rform j k ts1 m (rs1, rs2)) : - dom_eq (pprint k (rev rs1)) (pprint k rs2) -> - dom_eq (untag f1) (unpack (pivot g)). -Proof. -case: g f1; case=>pivot R [f1][<- _ A1] /=; case/(_ A1): R=>S D. -by rewrite !pp_interp interp_rev (sc_interp S A1). -Qed. - -End Exports. - -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). -Proof. apply: domeqX=>/=. Abort. - -End Exports. -End DomeqX. - -Export DomeqX.Exports. - -(***********************) -(* Automating invalidX *) -(***********************) - -Module InvalidX. -Section InvalidX. -Variables (K : ordType) (C : pred K) (T : Type) (U : union_map C T). -Implicit Types (i : ctx U) (ts : seq (term T)). -Notation form := Syntactify.form. -Notation untag := Syntactify.untag. - -Structure packed_map (m : U) := Pack {unpack : U}. -Canonical equate (m : U) := Pack m m. - -(* b is the boolean that we want to get out of isundef *) -Definition raxiom i ts m b (pivot : packed_map m) := - b -> valid (unpack pivot) = false. - -Structure rform i ts m b := - RForm {pivot :> packed_map m; _ : raxiom i ts b pivot}. - -(* start instance: note how isundef i ts is unified with *) -(* the b component of rform, which will be set to true by lemma statements *) - -Lemma start_pf i ts (f : form (empx U) i ts) : - @raxiom i ts (untag f) (isundef i ts) (equate f). -Proof. by case: f=>f [<- S A] /(isundef_sound A) /negbTE. Qed. - -Canonical start i ts f := RForm (@start_pf i ts f). - -End InvalidX. - -Module Exports. -Canonical equate. -Canonical start. - -Section Exports. -Variables (K : ordType) (C : pred K) (T : Type) (U : union_map 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) = undef. -Proof. by case: g; case=>pivot /= /(_ (erefl _))/negbT/invalidE. Qed. - -Lemma invalidX m i ts (g : rform i ts m true) : - valid (unpack (pivot g)) = false. -Proof. by rewrite undefX valid_undef. Qed. - -End Exports. - -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). -Proof. rewrite invalidX. Abort. - -End Exports. -End InvalidX. - -Export InvalidX.Exports. - -(************************************************) -(* Pushing an omap-like function inside a union *) -(************************************************) - -Module OmfX. -Section OmfX. - -Inductive syntx A (U : natmap A) := - | UnitOmf - | PtOmf of nat & A - | UnOmf of syntx U & syntx U - | OtherOmf of U. - -Fixpoint omf_interp A B (U : natmap A) (V : natmap B) (f : omap_fun U V) e : V := - match e with - UnitOmf => Unit - | PtOmf k v => if omf f (k, v) is Some w then pts k w else Unit - | UnOmf h1 h2 => omf_interp f h1 \+ omf_interp f h2 - | OtherOmf h => f h - end. - -Structure tagged_natmap A (U : natmap A) := Tag {untag : U}. -Definition unit_tag := Tag. -Definition pt_tag := unit_tag. -Definition union_tag := pt_tag. -Definition recurse_unguard_tag := union_tag. -Canonical recurse_guard_tag A (U : natmap A) f := @recurse_unguard_tag A U f. - -Definition guard_ax A B (U : natmap A) (V : natmap B) ts e (f : omap_fun U V) := - untag e = f (omf_interp idfun ts). - -Structure guarded_form A B (U : natmap A) (V : natmap B) ts := GForm { - gpivot :> tagged_natmap V; - guard_of : omap_fun U V; - _ : guard_ax ts gpivot guard_of}. - -Definition unguard_ax A (U : natmap A) (ts : syntx U) e := - untag e = omf_interp idfun ts. - -Structure unguarded_form A (U : natmap A) (ts : syntx U) := UForm { - upivot :> tagged_natmap U; - _ : unguard_ax ts upivot}. - -(* we first try to see if there's more function guards *) -Lemma recurse_guard_pf A B C (U : natmap A) (V : natmap B) (W : natmap C) - ts (f : omap_fun V W) - (g : @guarded_form A B U V ts) : - guard_ax ts (recurse_guard_tag (f (untag g))) - (f \o guard_of g). -Proof. by case: g=>e g pf /=; rewrite pf. Qed. -Canonical recurse_guard A B C U V W ts f g := GForm (@recurse_guard_pf A B C U V W ts f g). - -(* if not, we descend to unguarded form to syntactify the underlying expression *) -Lemma recurse_unguard_pf A (U : natmap A) ts (u : @unguarded_form A U ts) : - guard_ax ts (recurse_unguard_tag (untag u)) idfun. -Proof. by case: u. Qed. -Canonical recurse_unguard A U ts u := GForm (@recurse_unguard_pf A U ts u). - -(* syntactifying union recursively descends to both sides *) -Lemma unguard_union_pf A (U : natmap A) ts1 ts2 - (u1 : @unguarded_form A U ts1) (u2 : @unguarded_form A U ts2) : - unguard_ax (UnOmf ts1 ts2) (union_tag (untag u1 \+ untag u2)). -Proof. by case: u1 u2=>u1 pf1 [u2 pf2]; rewrite /= pf1 pf2. Qed. -Canonical unguard_union A U ts1 ts2 u1 u2 := UForm (@unguard_union_pf A U ts1 ts2 u1 u2). - -(* base case for syntactifying points-to *) -Lemma unguard_pts_pf A (U : natmap A) k (v : A) : - unguard_ax (PtOmf U k v) (pt_tag (pts k v)). -Proof. by []. Qed. -Canonical unguard_pts A U k v := UForm (@unguard_pts_pf A U k v). - -(* base case for syntactifying empty map Unit *) -Lemma unguard_unit_pf A (U : natmap A) : unguard_ax (UnitOmf U) (unit_tag Unit). -Proof. by []. Qed. -Canonical unguard_unit A U := UForm (@unguard_unit_pf A U). - -(* base case for syntactifying all other expressions *) -Lemma unguard_other_pf A (U : natmap A) (e : U) : unguard_ax (OtherOmf e) (Tag e). -Proof. by []. Qed. -Canonical unguard_other A U e := UForm (@unguard_other_pf A U e). - -End OmfX. - -Module Exports. -Canonical recurse_guard_tag. -Canonical recurse_guard. -Canonical recurse_unguard. -Canonical unguard_union. -Canonical unguard_pts. -Canonical unguard_unit. -Canonical unguard_other. - -(* main lemma *) - -Lemma omfX A B C (U : natmap A) (V : natmap B) (W : natmap C) ts - (f : omap_fun V W) - (g : @guarded_form A B U V ts) : - valid (omf_interp idfun ts) -> - f (untag g) = - omf_interp (f \o guard_of g) ts. -Proof. -case: g=>eq g /=; elim: ts eq=>[|s a|ts1 IH1 ts2 IH2|t] /= eq pf X. -- by rewrite pf !pfunit. -- rewrite validPt in X; rewrite pf omfPt // omf_comp /=. - by case: (omf g _)=>[x|]; rewrite ?omfPt ?pfunit. -- rewrite pf /= !omfUn //; last by rewrite -omfUn ?pfVE. - by rewrite IH1 ?(validL X) // IH2 ?(validR X). -by rewrite pf. -Qed. -End Exports. -End OmfX. - -Export OmfX.Exports. - diff --git a/pcm/autopcm.v b/pcm/autopcm.v deleted file mode 100644 index 8c4dea9..0000000 --- a/pcm/autopcm.v +++ /dev/null @@ -1,356 +0,0 @@ -(* -Copyright 2022 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. -From pcm Require Import options pred prelude pcm. -From pcm Require Export auto. - -(**************************************************************************) -(**************************************************************************) -(* Canonical structure lemmas for automating the following task: *) -(* *) -(* Splitting a PCM expression e into a combination of a input *) -(* subexpression e1 (modulo AC) and a residual expression obtained by *) -(* dropping all subterms of e1 from e. *) -(* *) -(* This allows us to formulate a generalized "pull"-style transformation, *) -(* which can factor out a given subexpression with one invocation. Note, *) -(* however, that this rewrites the goal into a right-biased form, losing *) -(* the associativity information. *) -(* *) -(**************************************************************************) -(**************************************************************************) - -(* Context structure for reflection of PCM expressions. An expression is *) -(* everything that are not recognized as a disjoint union, i.e., it can *) -(* be a variable or a literal. *) - -Section ReflectionContexts. -Variables (U : pcm). - -Structure ctx := Context {expx : seq U}. - -Definition empx := Context [::]. - -(* because contexts grow during computation, we need a notion of sub-context *) - -Definition sub_ctx (i j : ctx) := - Prefix (expx i) (expx j). - -Lemma sc_refl i : sub_ctx i i. -Proof. by []. Qed. - -Lemma sc_trans i j k : sub_ctx i j -> sub_ctx j k -> sub_ctx i k. -Proof. by apply: Prefix_trans. Qed. - -End ReflectionContexts. - -(* Expressions are syntactified as De Bruijn indices in the context. *) -(* Disjoint union is syntactified as concatenation of lists. *) - -(* now for reflection *) - -Section Reflection. -Variables (U : pcm). -Implicit Type i : ctx U. - -Variant term := Expr of nat. - -Definition nat_of (t : term) : nat := - let: Expr n := t in n. - -(* interpretation function for elements *) -Definition interp' (i : ctx U) (t : term) : option U := - let: Expr n := t in onth (expx i) n. - -(* main interpretation function *) -Notation fx i := (fun t f => interp' i t \+ f). -Definition interp (i : ctx U) (ts : seq term) : option U := - foldr (fx i) Unit ts. - -Lemma fE i ts x : foldr (fx i) x ts = x \+ interp i ts. -Proof. by elim: ts x=>[|t ts IH] x /=; rewrite ?unitR // IH /= joinCA. Qed. - -Lemma interp_rev i ts : interp i (rev ts) = interp i ts. -Proof. -elim: ts=>[|t ts IH] //=; rewrite rev_cons -cats1. -by rewrite /interp foldr_cat fE /= unitR IH. -Qed. - -(* we also need a pretty-printer, which works the same as interp *) -(* but removes trailing Unit's *) - -Fixpoint pprint i ts := - if ts is t :: ts' then - if ts' is [::] then interp' i t else interp' i t \+ pprint i ts' - else Unit. - -Lemma pp_interp i ts : pprint i ts = interp i ts. -Proof. by elim: ts=>[|t ts /= <-] //; case: ts=>//; rewrite unitR. Qed. - -Definition exp n t := let: Expr m := t in n == m. - -Definition efree n t := rfilter (exp n) t. -Arguments efree /. - -Lemma expP i n ts : - has (exp n) ts -> - interp i ts = onth (expx i) n \+ interp i (efree n ts). -Proof. -elim: ts=>[|t ts IH] //=; case: ifP=>/= [|_ H]. -- by case: t=>/= m /eqP->. -by rewrite (IH H) joinCA. -Qed. - -(* interpretation is invariant under context weakening *) -(* under assumption that the interpreted term is well-formed *) - -Definition wf i t := - let: Expr n := t in n < size (expx i). - -Lemma sc_wf i j ts : sub_ctx i j -> all (wf i) ts -> all (wf j) ts. -Proof. -move/Prefix_size=>H; elim: ts=>[|t ts IH] //=. -case/andP=>Hi /IH ->; rewrite andbT. -by case: t Hi=>v /= Hi; apply: leq_trans H. -Qed. - -Lemma sc_interp i j ts : - sub_ctx i j -> all (wf i) ts -> interp i ts = interp j ts. -Proof. -move=>H; elim: ts=>[|t ts IH] //= /andP [H0] /IH ->. -by case: t H0=>n /= /Prefix_onth <-. -Qed. - -Lemma wf_efree i n ts : all (wf i) ts -> all (wf i) (efree n ts). -Proof. by elim: ts=>//= t ts IH; case: ifP=>_ /andP [] //= -> /IH ->. Qed. - -End Reflection. - - -(************************************************************************) -(************************************************************************) -(* Purely functional decision procedure for the subtraction task. *) -(* Further below, it will be embedded into the canonical program pullX. *) -(************************************************************************) -(************************************************************************) - -(* subtract is purely functional version of pullX *) - -Section Subtract. -Variables (U : pcm). -Implicit Types (i : ctx U) (ts : seq term). - -(* We need a subterm lemma that returns the residiual of ts1. *) -(* xs accumulates uncancelled part of ts1, starting as nil *) - -Fixpoint subtract ts1 ts2 xs : option (seq term) := - match ts1 with - Expr n :: tsx1 => - if has (exp n) ts2 then subtract tsx1 (efree n ts2) xs - else subtract tsx1 ts2 (Expr n :: xs) - | [::] => if ts2 is [::] then Some xs else None - end. - -Lemma subtract_sound i ts1 ts2 rs1 xs : - all (wf i) ts1 -> all (wf i) ts2 -> - subtract ts1 ts2 xs = Some rs1 -> - interp i ts1 \+ interp i xs = interp i rs1 \+ interp i ts2. -Proof. -elim: ts1 ts2 xs=>[|t ts1 IH] ts2 xs /= =>[_|/andP [At A1]]. -- by case: ts2=>//=_; case=>->; rewrite unitR unitL. -case: t At=>n /= /size_onth [x X] A2; case: ifP=>Y. -- rewrite -joinA; move: (expP i Y)=>-> /(IH _ _ A1 (wf_efree n A2))->. - by rewrite joinCA. -by move/(IH _ _ A1 A2)=><-/=; rewrite joinCA joinA. -Qed. - -End Subtract. - - -(********************************) -(********************************) -(* Canonical structure programs *) -(********************************) -(********************************) - -Module Syntactify. -Section Syntactify. -Variables (U : pcm). -Implicit Types (i : ctx U) (ts : seq term). - -(* a tagging structure to control the flow of computation *) -Structure tagged_pcm := Tag {untag : U}. - -Local Coercion untag : tagged_pcm >-> PCM.sort. - -(* in reversed order; first test for unions, then empty, and exprs *) -Definition expr_tag := Tag. -Definition empty_tag := expr_tag. -Canonical Structure union_tag hc := empty_tag hc. - -(* Main structure *) -(* - i : input context *) -(* - j : output context *) -(* - ts : syntactification of pcm using context j *) - -Definition axiom i j ts (pivot : tagged_pcm) := - [/\ interp j ts = Some (untag pivot), sub_ctx i j & all (wf j) ts]. -Structure form i j ts := Form {pivot : tagged_pcm; _ : axiom i j ts pivot}. - -Local Coercion pivot : form >-> tagged_pcm. - -(* check for union *) - -Lemma union_pf i j k ts1 ts2 (f1 : form i j ts1) (f2 : form j k ts2) : - axiom i k (ts1 ++ ts2) (union_tag (untag f1 \+ untag f2)). -Proof. -case: f1 f2 =>[[u1]] /= [E1 S1 W1][[u2]][E2 S2 W2]; split=>/=. -- by rewrite /interp foldr_cat !fE unitL joinC -(sc_interp S2 W1) E1 E2. -- by apply: sc_trans S1 S2. -by rewrite all_cat (sc_wf S2 W1) W2. -Qed. - -Canonical union_form i j k ts1 ts2 (f1 : form i j ts1) (f2 : form j k ts2) := - Form (@union_pf i j k ts1 ts2 f1 f2). - -(* check if reached empty *) - -Lemma empty_pf i : axiom i i [::] (empty_tag Unit). -Proof. by split. Qed. - -Canonical empty_form i := - Form (@empty_pf i). - -(* check for expr *) - -Lemma expr_pf exprs1 exprs2 n (f : xfind exprs1 exprs2 n) : - axiom (Context exprs1) (Context exprs2) - [:: Expr n] (expr_tag (xuntag f)). -Proof. -case: f=>p [E H]; split=>//=; first by rewrite unitR. -by rewrite andbT (onth_size E). -Qed. - -Canonical expr_form exprs1 exprs2 n (f : xfind exprs1 exprs2 n) := - Form (@expr_pf exprs1 exprs2 n f). - -End Syntactify. - -Module Exports. -Coercion untag : tagged_pcm >-> PCM.sort. -Coercion pivot : form >-> tagged_pcm. -Canonical union_tag. -Canonical union_form. -Canonical empty_form. -Canonical expr_form. - -End Exports. -End Syntactify. - -Export Syntactify.Exports. - -(********************) -(* Automating pullX *) -(********************) - -Module PullX. -Section PullX. -Variables (U : pcm). -Implicit Types (j k : ctx U) (ts : seq term). -Notation form := Syntactify.form. -Notation untag := Syntactify.untag. - -Structure packed_pcm (m : U) := Pack {unpack : U}. - -Local Coercion unpack : packed_pcm >-> PCM.sort. - -Canonical equate (m : U) := Pack m m. - -(* j : input context *) -(* k : output context *) -(* ts : (input) the syntactification of subtractee in j *) -(* g : (aux) copy of the goal *) -(* r : (output) the optional residual term *) - -(* We have a choice of leaving the residual in a syntactified form or printing *) -(* it back into a PCM. Here we choose the second option since in HTT we need *) -(* to pass the residual to a partitioning autolemma which operates on heaps. *) - -Definition raxiom j k ts g (r : option U) (pivot : packed_pcm g) := - all (wf j) ts -> r -> sub_ctx j k /\ - Some (unpack pivot) = interp k ts \+ r. - -Structure rform j k ts g r := - RForm {pivot : packed_pcm g; _ : @raxiom j k ts g r pivot}. - -Local Coercion pivot : rform >-> packed_pcm. - -(* start instance: syntactify the goal (by triggering fg), *) -(* run the subtraction on quoted terms and print the residual. *) - -Lemma start_pf j k ts tg (fg : form j k tg) : - @raxiom j k ts (untag fg) - (obind (pprint k \o rev) (subtract tg ts [::])) - (equate fg). -Proof. -case: fg=>fg [Eg S Ag]; case E: (subtract _ _ _)=>[rs|] //= Am _; split=>//. -move/(subtract_sound Ag (sc_wf S Am)): E=>/=. -by rewrite unitR joinC Eg pp_interp interp_rev. -Qed. - -Canonical start j k ts tg (fg : form j k tg) := - RForm (@start_pf j k ts tg fg). - -End PullX. - -Module Exports. -Coercion unpack : packed_pcm >-> PCM.sort. -Coercion pivot : rform >-> packed_pcm. -Canonical equate. -Canonical start. - -Section Exports. -Variables (U : pcm). -Implicit Types (j : ctx U) (ts : seq term). -Notation form := Syntactify.form. -Notation untag := Syntactify.untag. - -(* we need to syntactify first the subtractee (fm), then the goal (fg) *) - -Lemma pullX' s j k ts g r (fs : form (empx U) j ts) - (fg : rform j k ts g (Some r)) : - untag fs = s -> - unpack (pivot fg) = s \+ r. -Proof. -move=><-; case: fg fs; case=>pivot R [fm][E _ A1] /=. -case/(_ A1 erefl): R=>S /=; rewrite -(sc_interp S A1) E. -by case. -Qed. - -End Exports. - -Arguments pullX' [U] s [j k ts g r fs fg] _. -Notation pullX s := (pullX' s erefl). - -Example ex0 (x y z : nat) : - 1 \+ x \+ 2 \+ y \+ 3 \+ z = 0. -Proof. -rewrite [LHS](pullX (3 \+ 2)) /=. -Abort. - -End Exports. -End PullX. - -Export PullX.Exports. diff --git a/pcm/heap.v b/pcm/heap.v deleted file mode 100644 index ea4c3c1..0000000 --- a/pcm/heap.v +++ /dev/null @@ -1,552 +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 heaps as possibly undefined finite maps from pointer *) -(* type to dynamic type. *) -(* Heaps are a special case of Partial Commutative Monoids (pcm) *) -(******************************************************************************) - -From HB Require Import structures. -From Coq Require Import ssreflect ssrbool ssrfun Eqdep. -From mathcomp Require Import ssrnat eqtype fintype tuple finfun seq path bigop. -From pcm Require Import options axioms prelude pred finmap. -From pcm Require Import pcm unionmap natmap. - -(************) -(* Pointers *) -(************) - -(* ptr/null is alternative name for nat/0 *) -Definition ptr := nat. -Definition null : ptr := 0. - -(* some pointer arithmetic: offsetting from a base *) -Definition ptr_offset (x : ptr) i := x + i. - -Notation "x .+ i" := (ptr_offset x i) - (at level 5, format "x .+ i"). - -Lemma ptr0 x : x.+0 = x. -Proof. by rewrite /ptr_offset addn0. Qed. - -Lemma ptr1 x : x .+ 1 = x.+1. -Proof. by rewrite /ptr_offset addn1. Qed. - -Lemma ptrA x i j : x.+i.+j = x.+(i+j). -Proof. by rewrite /ptr_offset addnA. Qed. - -Lemma ptrK x i j : (x.+i == x.+j) = (i == j). -Proof. by rewrite /ptr_offset eqn_add2l. Qed. - -Lemma ptr_null x m : (x.+m == null) = (x == null) && (m == 0). -Proof. by rewrite /ptr_offset addn_eq0. Qed. - -Lemma ptrT x y : {m : nat | (x == y.+m) || (y == x.+m)}. -Proof. -exists (if x <= y then (y - x) else (x - y)). -rewrite /ptr_offset leq_eqVlt /=. -by case: (ltngtP x y)=>/= E; rewrite subnKC ?(ltnW E) ?eq_refl ?orbT // E. -Qed. - -(*********) -(* Heaps *) -(*********) - -Inductive heap := - Undef | Def (finmap : {finMap ptr -> dynamic id}) of - null \notin supp finmap. - -Module Heap. -Section NullLemmas. -Variables (f g : {finMap ptr -> dynamic id}) (x : ptr) (v : dynamic id). - -Lemma upd_nullP : - x != null -> null \notin supp f -> null \notin supp (ins x v f). -Proof. by move=>H1 H2; rewrite supp_ins negb_or /= inE /= eq_sym H1. Qed. - -Lemma free_nullP : null \notin supp f -> null \notin supp (rem x f). -Proof. by move=>H; rewrite supp_rem negb_and /= H orbT. Qed. - -Lemma un_nullP : - null \notin supp f -> null \notin supp g -> - null \notin supp (fcat f g). -Proof. by move=>H1 H2; rewrite supp_fcat negb_or H1 H2. Qed. - -Lemma filt_nullP (q : pred ptr) : - null \notin supp f -> null \notin supp (kfilter q f). -Proof. by move=>H; rewrite supp_kfilt mem_filter negb_and H orbT. Qed. - -Lemma heap_base : null \notin supp f -> all (fun k => k != null) (supp f). -Proof. by move=>H; apply/allP=>k; case: eqP H=>// -> /negbTE ->. Qed. - -Lemma base_heap : all (fun k => k != null) (supp f) -> null \notin supp f. -Proof. by move/allP=>H; apply: (introN idP); move/H. Qed. - -Lemma heapE (h1 h2 : heap) : - h1 = h2 <-> - match h1, h2 with - Def f' pf, Def g' pg => f' = g' - | Undef, Undef => true - | _, _ => false - end. -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: eq_irrelevance. -Qed. - -End NullLemmas. - -(* methods *) - -Notation base := (@UM.base ptr (fun k => k != null) (dynamic id)). - -Definition def h := if h is Def _ _ then true else false. -Definition empty := @Def finmap.nil is_true_true. -Definition upd k v h := - if h is Def hs ns then - if decP (@idP (k != null)) is left pf then - Def (@upd_nullP _ _ v pf ns) - else Undef - else Undef. -Definition dom h : seq ptr := if h is Def f _ then supp f else [::]. -Definition assocs h : seq (ptr * 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 := - if h is Def hs _ then fnd x hs else None. -Definition union h1 h2 := - if (h1, h2) is (Def hs1 ns1, Def hs2 ns2) then - if disj hs1 hs2 then - Def (@un_nullP _ _ ns1 ns2) - else Undef - else Undef. -Definition pts (x : ptr) v := upd x v empty. -Definition empb h := if h is Def hs _ then supp hs == [::] else false. -Definition undefb h := if h is Undef then true else false. -Definition keys_of h : seq ptr := - if h is Def f _ then supp f else [::]. -Definition from (f : heap) : base := - if f is Def hs ns then UM.Def (heap_base ns) else UM.Undef _ _. -Definition to (b : base) : heap := - if b is UM.Def hs ns then Def (base_heap ns) else Undef. - -Module Exports. -(* heap has union map structure *) -Lemma heap_is_umc : union_map_axiom def empty Undef upd dom - assocs free find union empb - undefb pts from to. -Proof. -split; first by split=>[[]|[]] // f H; rewrite ?UM.umapE ?heapE. -split. -- split; first by split; [split=>[[]|]|rewrite heapE]. - split; last by case=>[|f H k] //; rewrite heapE. - split=>[|[]] //; split=>[k v [|h H]|[]] //=. - by case: decP=>// H1; rewrite heapE. -split=>[|k v]; last first. -- by rewrite /pts /UM.pts /UM.upd /=; case: decP=>// H; rewrite heapE. -split=>[|[]] //; split=>[|[]] //; split=>[k []|[|f1 H1][|f2 H2]] //. -by rewrite /union /UM.union /=; case: ifP=>D //; rewrite heapE. -Qed. - -HB.instance Definition _ := - isUnion_map.Build ptr (fun k => k != null) (dynamic id) heap heap_is_umc. -HB.instance Definition _ := isNatMap.Build (dynamic id) heap. -End Exports. -End Heap. -Export Heap.Exports. - -Notation "x :-> v" := (ptsT heap x (idyn v)) (at level 30). - -Canonical heap_PredType : PredType (nat * dynamic id) := um_PredType heap. -Coercion Pred_of_nmap (x : heap) : {Pred _} := [eta Mem_UmMap x]. - -(********************) -(* points-to lemmas *) -(********************) - -(* union_map pts lemmas combined with dyn_inj *) - -Section HeapPointsToLemmas. -Implicit Types (x : ptr) (h : heap). - -Lemma hcancelPtT A1 A2 x (v1 : A1) (v2 : A2) : - valid (x :-> v1) -> x :-> v1 = x :-> v2 -> A1 = A2. -Proof. by move=>V /(cancelPt V)/dyn_injT. Qed. - -Lemma hcancelPtT2 A1 A2 x1 x2 (v1 : A1) (v2 : A2) : - valid (x1 :-> v1) -> x1 :-> v1 = x2 :-> v2 -> (x1, A1) = (x2, A2). -Proof. by move=>V; case/(cancelPt2 V)=>-> E _; rewrite E. Qed. - -Lemma hcancelPtV A x (v1 v2 : A) : - valid (x :-> v1) -> x :-> v1 = x :-> v2 -> v1 = v2. -Proof. by move=>V; move/(cancelPt V)/dyn_inj. Qed. - -Lemma hcancelPtV2 A x1 x2 (v1 v2 : A) : - valid (x1 :-> v1) -> x1 :-> v1 = x2 :-> v2 -> (x1, v1) = (x2, v2). -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 h x. -Proof. by case/um_eta; case=>A v H; exists A, v. Qed. - -(* restatement of um_eta2, to avoid showing idyn's *) -Lemma heap_eta2 A x h (v : A) : - find x h = Some (idyn v) -> - h = x :-> v \+ free h x. -Proof. exact: um_eta2. Qed. - -Lemma hcancelT A1 A2 x (v1 : A1) (v2 : A2) h1 h2 : - valid (x :-> v1 \+ h1) -> - x :-> v1 \+ h1 = x :-> v2 \+ h2 -> A1 = A2. -Proof. by move=>V; case/(cancel V); move/dyn_injT. Qed. - -Lemma hcancelV A x (v1 v2 : A) h1 h2 : - valid (x :-> v1 \+ h1) -> - x :-> v1 \+ h1 = x :-> v2 \+ h2 -> [/\ v1 = v2, valid h1 & h1 = h2]. -Proof. by move=>V; case/(cancel V); move/dyn_inj. Qed. - -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 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. - -End HeapPointsToLemmas. - -Prenex Implicits heap_eta heap_eta2. - -(*******************************************) -(* Block update for reasoning about arrays *) -(*******************************************) - -Section BlockUpdate. -Variable A : Type. - -Fixpoint updi (x : ptr) (vs : seq A) {struct vs} : heap := - if vs is v'::vs' then x :-> v' \+ updi x.+1 vs' else Unit. - -Lemma updiS x v vs : updi x (v :: vs) = x :-> v \+ updi x.+1 vs. -Proof. by []. Qed. - -Lemma updi_last x v vs : - updi x (rcons vs v) = updi x vs \+ x.+(size vs) :-> v. -Proof. -elim: vs x v=>[|w vs IH] x v /=. -- by rewrite ptr0 unitR unitL. -by rewrite -(addn1 (size vs)) addnC -ptrA IH joinA ptr1. -Qed. - -Lemma updi_cat x vs1 vs2 : - updi x (vs1 ++ vs2) = updi x vs1 \+ updi x.+(size vs1) vs2. -Proof. -elim: vs1 x vs2=>[|v vs1 IH] x vs2 /=. -- by rewrite ptr0 unitL. -by rewrite -(addn1 (size vs1)) addnC -ptrA IH joinA ptr1. -Qed. - -Lemma updi_catI x y vs1 vs2 : - y = x + size vs1 -> updi x vs1 \+ updi y vs2 = updi x (vs1 ++ vs2). -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; first by rewrite dom0. -rewrite /= -addnS domPtUn inE /= negb_and negb_or -{4}(addn0 x). -by rewrite eqn_add2l -lt0n H IH // andbT orbT. -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 valid_unit orbC. -by rewrite /= validPtUn -addn1 updiVm' // orbF IH addn1 /= andbT. -Qed. - -Lemma updiVm x m xs : - x \in dom (updi x.+m xs) = [&& x != null, m == 0 & size xs > 0]. -Proof. -case: m=>[|m] /=; last first. -- by rewrite andbF; apply/negbTE/updiVm'. -case: xs=>[|v xs]; rewrite ptr0 ?andbF ?andbT ?dom0 //=. -by rewrite domPtUn inE /= eq_refl -updiS updiD orbF andbT /=. -Qed. - -Lemma updimV x m xs : - x.+m \in dom (updi x xs) = (x != null) && (m < size xs). -Proof. -case H: (x == null)=>/=. -- 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 ptr1 IH //. -by rewrite -updiS updiD H /= -{1}(ptr0 x) -ptr1 ptrA ptrK. -Qed. - -Lemma updiP x y xs : - reflect (y != null /\ exists m, x = y.+m /\ m < size xs) - (x \in dom (updi y xs)). -Proof. -case H: (y == null)=>/=. -- 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=>->. -- by rewrite updimV H /= => H1; split=>//; exists m. -rewrite updiVm; case/and3P=>H1; move/eqP=>-> H2. -by split=>//; exists 0; rewrite ptrA addn0 ptr0. -Qed. - -(* Invertibility *) -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. -- by case/esym/umap0E=>/unitbP; rewrite um_unitbPt. -- by case/umap0E=>/unitbP; rewrite um_unitbPt. -by case/(hcancelV D)=><- {}D /(IH _ _ D) <-. -Qed. - -Lemma updi_iinv x xs1 xs2 h1 h2 : - size xs1 = size xs2 -> valid (updi x xs1 \+ h1) -> - updi x xs1 \+ h1 = updi x xs2 \+ h2 -> 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; case/(hcancelV D)=><-{}D. -by case/(IH _ _ _ _ E D)=>->->. -Qed. - -Lemma updi_iota n (x : ptr) (f : nat -> A) : - updi x (map f (iota 0 n)) = - \big[join/Unit]_(i <- iota 0 n) x.+i :-> f i. -Proof. -elim: n x=>[|n IH] x; first by rewrite big_nil. -rewrite -addn1 iotaD add0n map_cat /= big_cat big_seq1 /=. -by rewrite updi_cat /= size_map size_iota /= unitR -IH. -Qed. - -Lemma updi_ord n (x : ptr) (f : 'I_n -> A) : - updi x (map f (enum 'I_n)) = - \big[join/Unit]_(i in 'I_n) x.+i :-> f i. -Proof. -case: n f=>[|n] f; first by rewrite -big_enum enum_ord0 big_nil. -set F := fun i => - if decP (b:=i < n.+1) idP is left pf then f (Ordinal pf) - else f (Ordinal (n:=n.+1) (m:=0) erefl). -have Ef i : f i = F i. -- rewrite /F; case: decP=>//. - by case: i=>i pf1 pf2; rewrite (pf_irr pf1 pf2). -set J := RHS. -have -> : J = \big[join/Unit]_(i in 'I_n.+1) x.+i :-> F i. -- by apply: eq_bigr=>i _; rewrite Ef. -rewrite -big_enum -(big_map _ predT (fun i=>x.+i :-> F i)). -rewrite val_enum_ord -updi_iota -val_enum_ord -map_comp. -by congr updi; apply: eq_map. -Qed. - -End BlockUpdate. - -Lemma updi_split {I : finType} T p k (f : {ffun I -> T}) : - updi p (fgraph f) = updi p (take (indx k) (fgraph f)) \+ - p.+(indx k) :-> f k \+ - updi (p.+(indx k).+1) (drop (indx k).+1 (fgraph f)). -Proof. -rewrite fgraph_codom /= codomE {1}(enum_split k) map_cat updi_cat /=. -rewrite map_take map_drop size_takel ?joinA; first by rewrite -ptr1 ptrA addn1. -by rewrite size_map index_size. -Qed. - -Lemma domeqUP A1 A2 x (xs1 : seq A1) (xs2 : seq A2) : - size xs1 = size xs2 -> - dom_eq (updi x xs1) (updi x xs2). -Proof. -move=>E; apply/domeqP; split; first by rewrite !updiD E. -apply/domE=>z; case: updiP=>[[H][m][->]|X]; first by rewrite updimV H E. -by case: updiP=>// [[H]][m][Ez S]; elim: X; split=>//; exists m; rewrite Ez E. -Qed. - -(*****************************) -(*****************************) -(* Automation of PtUn lemmas *) -(*****************************) -(*****************************) - -(* First, the mechanism for search-and-replace for the overloaded lemas, *) -(* pattern-matching on heap expressions. *) - -Structure tagged_heap := HeapTag {untag :> heap}. - -Definition right_tag := HeapTag. -Definition left_tag := right_tag. -Canonical found_tag i := left_tag i. - -Definition partition_axiom k r (h : tagged_heap) := untag h = k \+ r. - -Structure partition (k r : heap) := - Partition {heap_of :> tagged_heap; - _ : partition_axiom k r heap_of}. - -Lemma partitionE r k (f : partition k r) : untag f = k \+ r. -Proof. by case: f=>[[j]] /=; rewrite /partition_axiom /= => ->. Qed. - -Lemma found_pf k : partition_axiom k Unit (found_tag k). -Proof. by rewrite /partition_axiom unitR. Qed. - -Canonical found_struct k := Partition (found_pf k). - -Lemma left_pf h r (f : forall k, partition k r) k : - partition_axiom k (r \+ h) (left_tag (untag (f k) \+ h)). -Proof. by rewrite partitionE /partition_axiom /= joinA. Qed. - -Canonical left_struct h r (f : forall k, partition k r) k := - Partition (left_pf h f k). - -Lemma right_pf h r (f : forall k, partition k r) k : - partition_axiom k (h \+ r) (right_tag (h \+ f k)). -Proof. by rewrite partitionE /partition_axiom /= joinCA. Qed. - -Canonical right_struct h r (f : forall k, partition k r) k := - Partition (right_pf h f k). - -(* now the actual lemmas *) - -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}. - -Lemma defPt_nullO A h x (v : A) (f : partition (x :-> v) h) : - valid (untag f) -> x != null. -Proof. by rewrite partitionE; apply: validPtUn_cond. Qed. - -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 => /validR. Qed. - -Arguments defPt_defO [A][h] x [v][f] _. - -Lemma defPt_domO A h x (v : A) (f : partition (x :-> v) h) : - valid (untag f) -> x \notin dom h. -Proof. by rewrite partitionE; apply: validPtUnD. Qed. - -Arguments defPt_domO [A][h] x [v][f] _. - -Lemma domPtUnO A h x (v : A) (f : partition (x :-> v) h) : - dom (untag f) =i - [pred y | valid (untag f) & (x == y) || (y \in dom h)]. -Proof. by rewrite partitionE; apply: domPtUn. Qed. - -Arguments domPtUnO [A][h] x [v][f] _. - -Lemma lookPtUnO A h x (v : A) (f : partition (x :-> v) h) : - valid (untag f) -> find x (untag f) = Some (idyn v). -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 (untag f) x = h. -Proof. by rewrite partitionE; apply: freePtUn. Qed. - -Arguments freePtUnO [A h x v f] _. - -Lemma updPtUnO A1 A2 x i (v1 : A1) (v2 : A2) - (f : forall k, partition k i) : - upd x (idyn v2) (untag (f (x :-> v1))) = f (x :-> v2). -Proof. by rewrite !partitionE; apply: updPtUn. Qed. - -Arguments updPtUnO [A1 A2 x i v1 v2] f. - -Lemma cancelTO A1 A2 h1 h2 x (v1 : A1) (v2 : A2) - (f1 : partition (x :-> v1) h1) - (f2 : partition (x :-> v2) h2) : - valid (untag f1) -> f1 = f2 :> heap -> A1 = A2. -Proof. by rewrite !partitionE; apply: hcancelT. Qed. - -Arguments cancelTO [A1 A2 h1 h2] x [v1 v2 f1 f2] _ _. - -Lemma cancelO A h1 h2 x (v1 v2 : A) - (f1 : partition (x :-> v1) h1) - (f2 : partition (x :-> v2) h2) : - valid (untag f1) -> f1 = f2 :> heap -> - [/\ v1 = v2, valid h1 & h1 = h2]. -Proof. by rewrite !partitionE; apply: hcancelV. Qed. - -Arguments cancelO [A h1 h2] x [v1 v2 f1 f2] _ _. - -Lemma domPtUnXO A (v : A) x i (f : partition (x :-> v) i) : - valid (untag f) -> x \in dom (untag f). -Proof. by rewrite partitionE domPtUnE. Qed. - -(*******************************************************) -(*******************************************************) -(* Custom lemmas about singly-linked lists in the heap *) -(*******************************************************) -(*******************************************************) - -Fixpoint llist A p (xs : seq A) {struct xs} := - if xs is x::xt then - fun h => exists r h', h = p :-> (x, r) \+ h' /\ llist r xt h' - else fun h : heap => p = null /\ h = Unit. - -Lemma llist_prec A p (l1 l2 : seq A) h1 h2 g1 g2 : - valid (h1 \+ g1) -> - llist p l1 h1 -> llist p l2 h2 -> - h1 \+ g1 = h2 \+ g2 -> - [/\ l1 = l2, h1 = h2 & g1 = g2]. -Proof. -move=>V H1 H2 E; elim: l1 l2 h1 h2 p H1 H2 E V=>[|a1 l1 IH]. -- case=>[|a2 l2] h1 h2 p /= [->->]. - - by case=>_ ->; rewrite !unitL. - by case=>r [h'][-> _ ->] /validL /validPtUn_cond. -case=>[|a2 l2] h1 h2 p /= [p1][k1][-> H1]. -- by case=>->-> _ /validL /validPtUn_cond. -case=>p2 [k2][-> H2]; rewrite -!joinA => E V. -case: {E V} (hcancelV V E) H1 H2; case=>->-> V E H1 H2. -by case: (IH _ _ _ _ H1 H2 E V)=>->->->. -Qed. - -Lemma first_exists A p h (ls : seq A) : - p != null -> llist p ls h -> - exists x r h', - [/\ ls = x :: behead ls, h = p :-> (x, r) \+ h' & - llist r (behead ls) h']. -Proof. -case: ls=>[|a ls] /= H []; first by case: eqP H. -by move=>r [h'][-> H1]; eexists a, r, h'. -Qed. - -Lemma llist_null A (xs : seq A) h : - valid h -> llist null xs h -> xs = [::] /\ h = Unit. -Proof. -case: xs=>[|x xs] /= V H; first by case: H. -by case: H V=>p [h'][-> _] /validPtUn_cond. -Qed. - - diff --git a/pcm/invertible.v b/pcm/invertible.v deleted file mode 100644 index 89e6677..0000000 --- a/pcm/invertible.v +++ /dev/null @@ -1,124 +0,0 @@ -(* -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 pcm Require Import options axioms prelude. -From pcm Require Import pcm morphism. - -(* in the POPL21 paper, we use the notation - x D y D z =def= x D y /\ (x \+ y) D z -*) -Definition inv_rel (U : pcm) (D : rel U) := - 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 (U V : pcm) (f : pcm_morph U V) := - forall (a : U) (b1 b2 : V), - valid a -> sep f a Unit -> f a = b1 \+ b2 -> - exists a1 a2, [/\ a = a1 \+ a2, valid (a1 \+ a2), - sep f a1 a2, f a1 = b1 & f a2 = b2 ]. - -Definition inv_morph (U V : pcm) (f : pcm_morph U V) := - inv_rel (sep f) /\ inv_morph_axiom f. - - -(* trivial seprel in invertible *) -Lemma inv_sepT (U : pcm) : inv_rel (@relT U). Proof. by []. Qed. - -Lemma inv_idfun (U : pcm) : inv_morph (@idfun U). -Proof. by split=>// a b1 b2 V _ E; exists b1, b2; rewrite -E. Qed. - -(* composition of invertible morphisms is invertible *) - -Lemma inv_comp (U V W : pcm) - (f : pcm_morph U V) (g : pcm_morph V W) : - inv_morph f -> inv_morph g -> inv_morph (g \o f). -Proof. -case=>Sf Ff [Sg Fg]; split=>[a1 a2 a' V1|a b1 b2 V1] /andP [] /=. -- move/(Sf _ _ _ V1)=>H1 H2 /andP [/H1] /andP [D1 D2] /=. - rewrite /sepx/= !pfjoin ?(validLE3 V1,sepAxx V1 D1 D2) //= in H2 *. - by apply: Sg H2; apply: pfV3. -move=>D1; rewrite pfunit=>C1. -case/(Fg _ _ _ (pfV _ _ V1 D1) C1)=>c1 [c2][Ef Vc Xc <-<-]. -case/(Ff _ _ _ V1 D1): Ef Xc=>a1 [a2][-> Va Xc <-<- Xd]. -by exists a1, a2; rewrite /sepx/= Xc Xd. -Qed. - -(* morphism product of invertibles is invertible *) -Lemma inv_cprod (U1 U2 V1 V2 : pcm) - (f : pcm_morph U1 V1) (g : pcm_morph U2 V2) : - inv_morph f -> inv_morph g -> inv_morph (f \* g). -Proof. -case=>Sf Ff [Sg Fg]; split=>[a1 a2 a'|] /=. -- rewrite /rel_prod=>/andP [Va Vb] /andP [/(Sf _ _ _ Va) H1 /(Sg _ _ _ Vb) H2]. - by rewrite /sepx/=; case/andP=>/H1/andP [->->] /H2/andP [->->]. -move=>/= a1 b1 b2 /andP [Va Vb] /andP [H1 H2][]. -case/(Ff _ _ _ Va H1)=>c1 [c2][E1 Vc1 Y1 Fc1 Fc2]. -case/(Fg _ _ _ Vb H2)=>d1 [d2][E2 Vd1 Y2 Fd1 Fd2]. -exists (c1, d1), (c2, d2); rewrite /valid /= Vc1 Vd1. -rewrite /sepx/=/rel_prod/fprod /= Fc1 Fc2 Fd1 Fd2 Y1 Y2. -by rewrite pcmPE -E1 -E2 -!prod_eta. -Qed. - -(* ker requires only that sep f is invertible (not that f itself is) *) -Lemma inv_ker (U V : pcm) (f : pcm_morph U V) : - inv_rel (sep f) -> inv_rel (ker f). -Proof. -move=>Sf a1 a2 a' W /and3P [D1 /unitbP Eq1 _] /and3P [D2 /unitbP Eq2 /unitbP]. -case/andP: (Sf _ _ _ W D1 D2)=>{}D1 {}D2. -rewrite /kerx !pfjoin ?(validLE3 W) //; last by rewrite (sepAxx W). -by rewrite /kerx Eq1 Eq2 D1 D2 !unitL=>->; rewrite /sepU pcmE. -Qed. - -(* equalizer is invertible only for cancellative ranges *) -Lemma inv_eqlz (U : pcm) (V : eqcpcm) - (f : pcm_morph U V) (g : pcm_morph U V) : - inv_rel (sep f) -> inv_rel (sep g) -> inv_rel (eqlz f g). -Proof. -move=>Sf Sg x y z W; rewrite /eqlzx. -case/and4P=>X1 X2 Ex _ /and4P [Y1 Y2] Ey /eqP. -case/andP: (Sf _ _ _ W X1 Y1)=>E1 E1'. -case/andP: (Sg _ _ _ 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 => /joinxK ->; first by rewrite eqxx. -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 V : pcm) (S : subpcm_struct U V) - (W : pcm) (f : pcm_morph V W) : - subsep S =2 sep f -> - inv_morph f -> inv_morph (f \o pval S). -Proof. -move=>E [Sf Hf]; split=>[a1 a2 a' V1|a b1 b2 V1] /=; -rewrite /sepx/= !pfT /=. -- by rewrite !pfjoin ?(validLE3 V1) //; apply: Sf; apply: pfV3. -rewrite pfunit /= => H1 /(Hf _ _ _ (pfVI V1) H1). -case=>c1 [c2][Eq Vc H2 F1 F2]; exists (psub S c1), (psub S c2). -rewrite pfT /= -!pfjoin ?E //= -Eq psub_pval //=. -by rewrite !pval_psub ?(validL Vc,validR Vc) // E (sep0E Vc). -Qed. - - diff --git a/pcm/morphism.v b/pcm/morphism.v deleted file mode 100644 index 431fa91..0000000 --- a/pcm/morphism.v +++ /dev/null @@ -1,3341 +0,0 @@ -(* -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 HB Require Import structures. -From Coq Require Import ssreflect ssrbool ssrfun. -From mathcomp Require Import ssrnat eqtype fintype finfun. -From pcm Require Import options pred axioms prelude. -From pcm Require Import pcm. - -(*****************) -(*****************) -(* PCM Morphisms *) -(*****************) -(*****************) - -(* Orthogonality relations (in the POPL21 paper called separating relations) *) -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)]. - -(* Here, by separating relation we take a somewhat different definition *) -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)]. - -(* The two definitions differ in that seprel elide the validity requirement. *) -(* This is pragmatic distinction, because in practice we often have *) -(* the validity in context so we don't have to bother proving it. *) -(* 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 relation of a morphism f *) -(* we will write x \orth_f y iff valid (x \+ y) and sep f x y *) -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 sperels 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. - -HB.mixin Record isSeprel (U : pcm) (sep : rel U) := { - seprel_subproof : seprel_axiom sep}. - -#[short(type="seprel")] -HB.structure Definition Seprel (U : pcm) := {sep of isSeprel U sep}. - -Section Repack. -Variables (U : pcm) (sep : seprel U). - -Lemma sep00 : sep Unit Unit. -Proof. by case: (@seprel_subproof U sep). Qed. - -Lemma sepC x y : - valid (x \+ y) -> - sep x y = sep y x. -Proof. by case: (@seprel_subproof U sep)=>_ H _ _; apply: H. Qed. - -Lemma sepx0 x y : - valid (x \+ y) -> - sep x y -> - sep x Unit. -Proof. by case: (@seprel_subproof U sep)=>_ _ 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. -case: (@seprel_subproof U sep)=>_ _ _ H V R1 R2. -by 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; rewrite -!(@sepC y) ?unitR ?(validE2 V) //. -by 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 Repack. - - -(***********************************) -(* Sep relations and constructions *) -(***********************************) - -(* always-true relation *) - -Definition relT {U} : rel U := fun x y => true. - -Lemma relT_is_seprel {U : pcm} : seprel_axiom (@relT U). Proof. by []. Qed. -HB.instance Definition _ U := isSeprel.Build U relT relT_is_seprel. - - -(* always-unit relation *) -(* always-false relation is not seprel, because we need sep0 Unit Unit *) -(* i.e., sepU is really the smallest seprel we can have *) - -Definition sepU {U : pcm} : rel U := fun x y => unitb x && unitb y. - -Section SepU. -Variable U : pcm. - -Lemma sepU_is_seprel : seprel_axiom (@sepU U). -Proof. -rewrite /sepU; split=>[|x y|x y|x y z]. -- by rewrite pcmE. -- by rewrite andbC. -- by move=>V /andP [->]; rewrite pcmE. -move=>V /andP [/unitbP -> /unitbP ->] /andP [_ /unitbP ->]. -by rewrite unitL !pcmE. -Qed. - -HB.instance Definition _ := isSeprel.Build U sepU sepU_is_seprel. -End SepU. - - -(************************************) -(* Conjunction of seprels is seprel *) -(************************************) - -(* first conjunction of plain rels *) -Definition relI U (X Y : rel U) (x y : U) := X x y && Y x y. -(* and iterated variants *) -Definition rel3I U (X Y Z : rel U) (x y : U) := - [&& X x y, Y x y & Z x y]. -Definition rel4I U (X Y Z W : rel U) (x y : U) := - [&& X x y, Y x y, Z x y & W x y]. - -Arguments relI {U} X Y x y /. -Arguments rel3I {U} X Y Z x y /. -Arguments rel4I {U} X Y Z W x y /. - -Section SepI. -Variables (U : pcm) (X Y : seprel U). - -Lemma relI_is_seprel : seprel_axiom (relI X Y). -Proof. -split=>[|x y|x y|x y z] /=. -- by rewrite !sep00. -- by move=>V; rewrite !(sepC _ V). -- by move=>V /andP []; do ![move/(sepx0 V) ->]. -move=>V /andP [X1 Y1] /andP [X2 Y2]. -by rewrite !(sepAxx V X1 X2, sepAxx V Y1 Y2). -Qed. - -HB.instance Definition _ := isSeprel.Build U (relI X Y) relI_is_seprel. -End SepI. - -(* 3-way conjunction *) -Section Sep3I. -Variables (U : pcm) (X Y Z : seprel U). - -Lemma rel3I_is_seprel : seprel_axiom (rel3I X Y Z). -Proof. -split=>[|x y|x y|x y z] /=. -- by rewrite !sep00. -- by move=>V; rewrite !(sepC _ V). -- by move=>V /and3P []; do ![move/(sepx0 V) ->]. -move=>V /and3P [X1 Y1 Z1] /and3P [X2 Y2 Z2]. -by rewrite !(sepAxx V X1 X2, sepAxx V Y1 Y2, sepAxx V Z1 Z2). -Qed. - -HB.instance Definition _ := isSeprel.Build U (rel3I X Y Z) rel3I_is_seprel. -End Sep3I. - -(* 4-way conjunction *) -Section Sep4I. -Variables (U : pcm) (X Y Z W : seprel U). - -Lemma rel4I_is_seprel : seprel_axiom (rel4I X Y Z W). -Proof. -split=>[|x y|x y|x y z] /=. -- by rewrite !sep00. -- by move=>V; rewrite !(sepC _ V). -- by move=>V /and4P []; do ![move/(sepx0 V) ->]. -move=>V /and4P [X1 Y1 Z1 W1] /and4P [X2 Y2 Z2 W2]. -by rewrite !(sepAxx V X1 X2, sepAxx V Y1 Y2, sepAxx V Z1 Z2, sepAxx V W1 W2). -Qed. - -HB.instance Definition _ := isSeprel.Build U (rel4I X Y Z W) rel4I_is_seprel. -End Sep4I. - -(************************************) -(* projections and pairwise product *) -(************************************) - -Definition rel_fst U V (X : rel U) (x y : U * V) := X x.1 y.1. -Arguments rel_fst {U V} X x y /. -Definition rel_snd U V (Y : rel V) (x y : U * V) := Y x.2 y.2. -Arguments rel_snd {U V} Y x y /. - -Section SepFst. -Variables (U V : pcm) (X : seprel U). -Lemma relfst_is_seprel : seprel_axiom (@rel_fst U V X). -Proof. -split=>[|x y|x y|x y z] //=. -- by rewrite !sep00. -- by case/andP=>V1 _; rewrite sepC. -- by case/andP=>V1 _ /(sepx0 V1). -by case/andP=>V1 V2 X1 Y1; rewrite !(sepAxx V1 X1 Y1). -Qed. -HB.instance Definition _ := - isSeprel.Build (U * V)%type (rel_fst X) relfst_is_seprel. -End SepFst. - -Section SepSnd. -Variables (U V : pcm) (Y : seprel V). -Lemma relsnd_is_seprel : seprel_axiom (@rel_snd U V Y). -Proof. -split=>[|x y|x y|x y z] //=. -- by rewrite !sep00. -- by case/andP=>_ V2; rewrite sepC. -- by case/andP=>_ V2 /(sepx0 V2). -by case/andP=>V1 V2 X2 Y2; rewrite !(sepAxx V2 X2 Y2). -Qed. -HB.instance Definition _ := - isSeprel.Build (U * V)%type (rel_snd Y) relsnd_is_seprel. -End SepSnd. - -Definition rel_prod U V (X : rel U) (Y : rel V) := - relI (rel_fst X) (rel_snd Y). -Arguments rel_prod {U V} X Y /. - -HB.instance Definition _ (U V : pcm) (X : seprel U) (Y : seprel V) := - Seprel.copy (rel_prod X Y) (rel_prod X Y). - -(**********************) -(**********************) -(* Classes of seprels *) -(**********************) -(**********************) - -(* Non-symmetric seprels *) - -(* Often times, we build a seprel out of a non-symmetric relation *) -(* by taking the symmetric closure of the relation explicitly. *) -(* This is such a common approach, that its useful *) -(* to introduce a class of primitive non-symmetric seprels that *) -(* are made full seprels by closing up. *) - -Definition nseprel_axiom (U : pcm) (nsep : rel U) := - [/\ (* unit is separated from unit *) - nsep Unit Unit, - (* 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) -> nsep x y -> nsep x Unit, - (* if the first arguments is Unit, then nsep trivially holds *) - forall y, valid y -> nsep Unit y, - (* associativity 1 *) - forall x y z, valid (x \+ y \+ z) -> - nsep x y -> nsep (x \+ y) z -> nsep x (y \+ z) & - (* associativity 2 *) - forall x y z, valid (x \+ y \+ z) -> - nsep x (y \+ z) -> nsep y (x \+ z) -> nsep x y /\ nsep (x \+ y) z]. - -Lemma orth_nsep (U : pcm) (nsep : rel U) : - nseprel_axiom nsep -> - seprel_axiom (fun x y => nsep x y && nsep y x). -Proof. -case=>H1 H2 H3 H4 H5; split=>[|x y V|x y V|x y z V]. -- by rewrite H1. -- by rewrite andbC. -- by case/andP=>/(H2 _ _ V) -> _; rewrite (H3 _ (validL V)). -case/andP=>X1 X2 /andP [X3 X4]. -move: (H4 x y z V X1 X3)=>X5; rewrite X5 /=. -rewrite joinC in X3. -move: (H4 y x z); rewrite (validLE3 V)=>/(_ erefl X2 X3) X6. -rewrite joinC in X4; rewrite joinC in X6. -move: (H5 y z x); rewrite (validLE3 V)=>/(_ erefl X6 X4) [->->] /=. -by move: (H5 z y x); rewrite (validLE3 V)=>/(_ erefl X4 X6) [] ->. -Qed. - -(* Guarded non-symmetric seprels *) - -(* Further optimization, if the nsep has the form forall k, P k x -> Q k x y *) -(* for some P and Q *) - -Section GuardedSeprels. -Variable (U : pcm) (X : Type) (P : X -> U -> Prop) (Q : X -> U -> U -> Prop). -Variable (nsep : rel U). - -Hypothesis pf1 : forall x y, reflect (forall k, P k x -> Q k x y) (nsep x y). - -Definition gseprel_axiom := - [/\ (* P is disjunctive *) - forall k x y, valid (x \+ y) -> P k (x \+ y) <-> P k x \/ P k y, - (* unit is separated from unit *) - forall k, P k Unit -> Q k Unit Unit, - (* is x is in the domain (i.e., x is considered) *) - (* then it is separate from at least Unit *) - forall k x y, valid (x \+ y) -> P k x -> Q k x y -> Q k x Unit, - (* if the first argument is Unit, then nsep trivially holds *) - forall k y, valid y -> P k Unit -> Q k Unit y, - (* associativity 1 *) - forall k x y z, valid (x \+ y \+ z) -> - P k x -> Q k x y -> Q k (x \+ y) z -> Q k x (y \+ z) & - (* associativity 2 *) - forall k x y z, valid (x \+ y \+ z) -> - P k x -> Q k x (y \+ z) -> Q k x y /\ Q k (x \+ y) z]. - -Lemma orth_gsep : - gseprel_axiom -> - seprel_axiom (fun x y => nsep x y && nsep y x). -Proof. -case=>H1 H2 H3 H4 H5 H6; apply: orth_nsep. -split=>[|x y V|y V|x y z V|x y z V]. -- by apply/pf1. -- by move/pf1=>H; apply/pf1=>k /[dup] Y /H; apply: H3 V Y. -- by apply/pf1=>x; apply: H4. -- move/pf1=>X1 /pf1 X2; apply/pf1=>k H. - apply: H5=>//; first by apply: X1 H. - by apply: X2; apply/H1; [rewrite (validLE3 V)|left]. -move/pf1=>X1 /pf1 X2; split; apply/pf1=>k H. -- by case: (H6 k x y z V H (X1 _ H)). -case/H1: H; first by rewrite (validLE3 V). -- by move=>H; case: (H6 k x y z V H (X1 _ H)). -move=>H; rewrite joinC. -case: (H6 k y x z)=>//; first by rewrite (validLE3 V). -by apply: X2 H. -Qed. - -End GuardedSeprels. - -(* Further optimization *) -(* Most of the time, the guard is not only disjunctive, but *) -(* exclusively disjunctive. This allows to weaken the conditions. *) -(* We have to prove a couple of additional side-conditions however *) -(* so we'll see how useful this is. *) -(* I wasn't able to find a simpler formulation. *) - -Section ExclusiveGuardedSeprels. -Variable (U : pcm) (X : Type) (P : X -> U -> Prop) (Q : X -> U -> U -> Prop). -Variable (nsep : rel U). - -Hypothesis pf1 : forall x y, reflect (forall k, P k x -> Q k x y) (nsep x y). - -Definition xgseprel_axiom := - [/\ (* P is disjunctive *) - forall k x y, valid (x \+ y) -> P k (x \+ y) <-> P k x \/ P k y, - (* P is exclusive *) - forall k1 k2 x y, valid (x \+ y) -> P k1 x -> P k2 y -> k1 <> k2, - forall k1 k2 x y z, valid (x \+ y \+ z) -> - P k1 x -> Q k1 x y -> Q k2 (x \+ y) z -> k1 = k2, - forall k1 k2 x y z, valid (x \+ y \+ z) -> - P k1 x -> Q k1 x (y \+ z) -> P k2 y-> Q k2 y (x \+ z) -> k1 = k2, - (* unit is separated from unit *) - forall k, P k Unit -> Q k Unit Unit, - (* is x is in the domain (i.e., x is considered) *) - (* then it is separate from at least Unit *) - forall k x y, valid (x \+ y) -> P k x -> Q k x y -> Q k x Unit, - (* if the first argument is Unit, then nsep trivially holds *) - forall k y, valid y -> P k Unit -> Q k Unit y, - (* associativity 1 *) - forall k x y z, valid (x \+ y \+ z) -> - P k x -> (forall k', ~ P k' y) -> - Q k x y -> Q k (x \+ y) z -> Q k x (y \+ z) & - (* associativity 2 *) - forall k x y z, valid (x \+ y \+ z) -> - P k x -> (forall k', ~ P k' y) -> - Q k x (y \+ z) -> Q k x y /\ Q k (x \+ y) z]. - -Lemma xorth_gsep : - xgseprel_axiom -> - seprel_axiom (fun x y => nsep x y && nsep y x). -Proof. -case=>H1 Xp Xq1 Xq2 H2 H3 H4 H5 H6; apply: orth_nsep. -split=>[|x y V|y V|x y z V|x y z V]. -- by apply/pf1. -- by move/pf1=>H; apply/pf1=>k /[dup] Y /H; apply: H3 V Y. -- by apply/pf1=>x; apply: H4. -- have V' : valid (x \+ y) by rewrite (validLE3 V). - move/pf1=>X1 /pf1 X2; apply/pf1=>k Hx; apply: (H5)=>//; last first. - - by apply: X2; apply/H1=>//; left. - - by apply: X1. - move=>k' Hy; apply: Xp (V') (Hx) (Hy) _. - have Hxy : P k' (x \+ y) by rewrite H1 //; right. - by apply: Xq1 V (Hx) (X1 _ Hx) (X2 _ Hxy). -have V' : valid (x \+ y) by rewrite (validLE3 V). -move/pf1=>X1 /pf1 X2; split; apply/pf1=>k Hx. -- case: (H6 k x y z V Hx)=>//; last by apply: X1. - move=>k' Hy; apply: Xp V' (Hx) (Hy) _. - by apply: Xq2 (X1 _ Hx) _ (X2 _ Hy). -case/H1: Hx; first by rewrite (validLE3 V). -- move=>Hx; case: (H6 k x y z V Hx)=>//; last by apply: X1. - by move=>k' Hy; apply: Xp V' (Hx) (Hy) _; apply: Xq2 (X1 _ Hx) _ (X2 _ Hy). -move=>Hy; rewrite joinC. -case: (H6 k y x z)=>//; first by rewrite (validLE3 V). -- by move=>k' Hx; apply: Xp V' (Hx) (Hy) _; apply: Xq2 (X1 _ Hx) _ (X2 _ Hy). -by apply: X2. -Qed. - -End ExclusiveGuardedSeprels. - - -(*****************) -(*****************) -(* PCM morphisms *) -(*****************) -(*****************) - -(* morphism comes with a seprel on which it is valid *) - -Definition pcm_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]. - -HB.mixin Record isPCM_morphism (U V : pcm) (f : U -> V) := { - sep_op : seprel U; - pcm_morph_subproof : pcm_morph_axiom sep_op f}. - -#[short(type=pcm_morph)] -HB.structure Definition PCM_morphism (U V : pcm) := - {f of isPCM_morphism U V f}. - -Arguments sep_op {U V} _ /. -Arguments pcm_morph_subproof {U V}. - -(* When sep_op is applied to function f, the system infers *) -(* the pcm_morph structure F of f, as intended. *) -(* However, it then proceeds to display sep_op F instead of sep_op f, *) -(* which is often unreadable, as F is typically large. *) -(* The following arranges that f is printed instead of F. *) - -(* There are two options that differ in issues orthogonal to f/F. *) -(* 1. sepx is declared seprel *) -(* 2. sepx is declared rel, with canonical seprel attached *) -(* Option 1 automatically simplifies nested sepx's of composed functions. *) -(* Option 2 makes it easier to attach canonical structures to sepx, *) -(* but simplification requires explicit and often iterated unfolding *) -(* e.g. rewrite /sepx/=/sepx/=/sepx/=. *) -(* Here, choosing (2); if iterated unfolding of some seprel is *) -(* frequently needed, it should be exported as a lemma. *) - -(* -(* option 1 *) -Definition sepx U V (f : pcm_morph U V) - of phantom (U -> V) f := sep_op f. -Notation sep f := (sepx (Phantom (_ -> _) f)). -*) - -(* option 2 *) -Definition sepx U V (f : pcm_morph U V) - of phantom (U -> V) f : rel U := sep_op f. -Notation sep f := (sepx (Phantom (_ -> _) f)). -HB.instance Definition _ U V (f : pcm_morph U V) := - Seprel.on (sep f). - -(* we can similarly get uniform name for the structure itself *) -(* but we won't use this *) -Definition morphx (U V : pcm) (f : pcm_morph U V) - of phantom (U -> V) f := f. -Notation morph f := (morphx (Phantom (_ -> _) f)). - - -Section Laws. -Variables (U V : pcm) (f : pcm_morph U V). - -Lemma pfunit : f Unit = Unit. -Proof. by case: (pcm_morph_subproof f). Qed. - -Lemma pfjoinV (x y : U) : - valid (x \+ y) -> - sep f x y -> - valid (f x \+ f y) /\ f (x \+ y) = f x \+ f y. -Proof. by case: (pcm_morph_subproof f)=>_; apply. Qed. - -Lemma pfV2 x y : - valid (x \+ y) -> - sep f x y -> - valid (f x \+ f y). -Proof. by move=>H S; case: (pfjoinV H S). Qed. - -Lemma pfjoin x y : - valid (x \+ y) -> - sep f 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) -> - sep f x y -> - f x \+ f y = f (x \+ y). -Proof. by move=>*; rewrite pfjoin. Qed. - -Lemma pfV x : - valid x -> - sep f x Unit -> - valid (f x). -Proof. by rewrite -{1 3}[x]unitR=>W S; rewrite pfjoin // pfV2. Qed. - -Lemma pfV3 x y z : - valid (x \+ y \+ z) -> - sep f x y -> - sep f (x \+ y) z -> - valid (f x \+ f y \+ f z). -Proof. by move=>W D1 D2; rewrite -pfjoin ?(validL W) // pfV2. Qed. - -Lemma pfL (x y z : U) : - valid (x \+ z) -> - valid (y \+ z) -> - sep f x z -> - sep f y z -> - f y = f x -> - f (y \+ z) = f (x \+ z). -Proof. by move=>V1 V2 S1 S2 E; rewrite !pfjoin // E. Qed. - -Lemma pfR (x y z : U) : - valid (z \+ x) -> - valid (z \+ y) -> - sep f z x -> - sep f z y -> - f y = f x -> - f (z \+ y) = f (z \+ x). -Proof. by move=>V1 V2 S1 S2 E; rewrite !pfjoin // E. Qed. - -Lemma pfUnE (x1 y1 x2 y2 : U) : - valid (x1 \+ y1) -> - valid (x2 \+ y2) -> - sep f x1 y1 -> - sep f x2 y2 -> - f x2 = f x1 -> - f y2 = f y1 -> - f (x2 \+ y2) = f (x1 \+ y1). -Proof. -move=>V1 V2 S1 S2 E1 E2. -by rewrite !pfjoin // E1 E2. -Qed. - -Lemma pfunitL x : f Unit \+ x = x. -Proof. by rewrite pfunit unitL. Qed. - -Lemma pfunitR x : x \+ f Unit = x. -Proof. by rewrite pfunit unitR. Qed. - -End Laws. - -Arguments pfunit {U V}. -Arguments pfjoinV {U V}. -Arguments pfV2 {U V}. -Arguments pfjoin {U V}. -Arguments joinpf {U V}. -Arguments pfV {U V}. -Arguments pfV3 {U V}. -Arguments pfL {U V} f {x y}. -Arguments pfR {U V} f {x y}. -Arguments pfunitL {U V}. -Arguments pfunitR {U V}. - -(*********************) -(* Morphism equality *) -(*********************) - -(* morphisms are equal iff equal as functions and have equal seprels *) - -Definition pcm_morph_eq (U V : pcm) (f g : pcm_morph U V) := - f =1 g /\ sep f =s sep g. - -Notation "f =m g" := (pcm_morph_eq f g) (at level 55). - -Lemma pcm_morph_eq_refl {U V} : Reflexive (@pcm_morph_eq U V). -Proof. by []. Qed. - -Lemma pcm_morph_eq_sym {U V} : Symmetric (@pcm_morph_eq U V). -Proof. -by apply/sym_iff_pre_sym=>f1 f2 [H1 H2]; split=>// x y W; rewrite H2. -Qed. - -Lemma pcm_morph_eq_trans {U V} : Transitive (@pcm_morph_eq U V). -Proof. -move=>f2 f1 f3 [H12 S12][H23 S23]; split=>[x|x y W]. -- by rewrite H12. -by rewrite S12 // S23. -Qed. - -(***************************************) -(* Operations on morphisms and seprels *) -(***************************************) - -(* morphism preimage of seprel is seprel *) - -Definition preimx (U V : pcm) (f : pcm_morph U V) - of phantom (U -> V) f : rel V -> rel U := - fun R x y => sep f x y && R (f x) (f y). -Notation preim f := (preimx (Phantom (_ -> _) f)). -Arguments preimx {U V f} _ _ _ _ /. - -Section Preim. -Variables (U V : pcm) (f : pcm_morph U V) (R : seprel V). - -Lemma preim_is_seprel : seprel_axiom (preim f R). -Proof. -split=>[|x y|x y|x y z] /=. -- by rewrite !pfunit !sep00. -- move=>W; rewrite (sepC (sep f))=>//=; case H : (sep f y x)=>//. - by rewrite sepC // pfV2 // sepC. -- move=>H /andP [H1 H2]. - by rewrite !pfunit (sep0E H H1) (sep0E _ H2) // pfV2. -move=>W /andP [G1 F1] /andP [G2 F2]. -rewrite pfjoin ?(validLE3 W) // in F2. -rewrite pfjoin ?(validLE3 W) ?(sepAxx W G1 G2) //=. -move: (pfV2 _ _ _ W G2); rewrite pfjoin ?(validLE3 W) // => W2. -by rewrite !(sepAxx W2 F1 F2). -Qed. - -HB.instance Definition _ := isSeprel.Build U (preim f R) preim_is_seprel. -End Preim. - -(* kernel of morphism is seprel *) - -Definition kerx (U V : pcm) (f : pcm_morph U V) - of phantom (U -> V) f : rel U := - fun x y => sep f x y && sepU (f x) (f y). -Notation ker f := (kerx (Phantom (_ -> _) f)). -HB.instance Definition _ U V (f : pcm_morph U V) := - isSeprel.Build U (ker f) (preim_is_seprel f _). - -(* restriction of morphism by seprel is morphism *) -(* restriction has tpcm range because it returns undef results *) -Section Restriction. -Variables (U : pcm) (V : tpcm). - -Definition res (f : U -> V) (S : rel U) : U -> V := - fun x : U => if S x Unit then f x else undef. - -Variables (f : pcm_morph U V) (S : seprel U). - -Lemma res_is_morphism : pcm_morph_axiom (relI S (sep f)) (res f S). -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. - -HB.instance Definition _ := isPCM_morphism.Build U V (res f S) res_is_morphism. -End Restriction. - -(* equalizer of morphism is seprel *) - -Definition eqlzx (U : pcm) (V : eqpcm) (f1 f2 : pcm_morph U V) - of phantom (U -> V) f1 & phantom (U -> V) f2 : rel U := - fun x y => [&& sep f1 x y, sep f2 x y, f1 x == f2 x & f1 y == f2 y]. -Notation eqlz f1 f2 := (eqlzx (Phantom (_ -> _) f1) (Phantom (_ -> _) f2)). - -Section Equalizer. -Variables (U : pcm) (V : eqpcm) (f1 f2 : pcm_morph U V). - -Lemma eqlz_is_seprel : seprel_axiom (eqlz f1 f2). -Proof. -rewrite /eqlzx; split=>[|x y W|x y W|x y z W]. -- by rewrite !sep00 !pfunit eq_refl. -- by rewrite !andbA andbAC (sepC (sep f1)) // (sepC (sep f2)). -- 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. - -HB.instance Definition _ := isSeprel.Build U (eqlz f1 f2) eqlz_is_seprel. -End Equalizer. - -(* join of two morphism is a morphism with an appropriate seprel *) - -Definition join_relx (U V : pcm) (f1 f2 : pcm_morph U V) - of phantom (U -> V) f1 & phantom (U -> V) f2 : rel U := - fun x y => [&& sep f1 x y, sep f2 x y & - valid ((f1 x \+ f2 x) \+ (f1 y \+ f2 y))]. -Notation join_rel f1 f2 := - (join_relx (Phantom (_ -> _) f1) (Phantom (_ -> _) f2)). - -Definition join_fun (U V : pcm) (f1 f2 : U -> V) : U -> V := - fun x => f1 x \+ f2 x. - -Section JoinMorph. -Variables (U V : pcm) (f1 f2 : pcm_morph U V). - -Lemma join_rel_is_seprel : seprel_axiom (join_rel f1 f2). -Proof. -rewrite /join_relx; split=>[|x y W|x y W|x y z W]. -- by rewrite !sep00 ?unitL !pfunit ?unitL valid_unit. -- by rewrite (sepC (sep f1)) // (sepC (sep f2)) // (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. - -HB.instance Definition _ := - isSeprel.Build U (join_rel f1 f2) join_rel_is_seprel. - -Lemma join_fun_is_morph : - pcm_morph_axiom (join_rel f1 f2) (join_fun f1 f2). -Proof. -rewrite /join_fun; 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. - -HB.instance Definition _ := - isPCM_morphism.Build U V (join_fun f1 f2) join_fun_is_morph. -End JoinMorph. - -(*************************) -(* Categorical morphisms *) -(*************************) - -(* identity *) - -Lemma id_is_morphism (U : pcm) : pcm_morph_axiom relT (@idfun U). -Proof. by []. Qed. -HB.instance Definition _ (U : pcm) := - isPCM_morphism.Build U U idfun (id_is_morphism U). - -(* composition *) - -Section CompMorphism. -Variables U V W : pcm. -Variables (f : pcm_morph U V) (g : pcm_morph V W). -Lemma comp_is_morphism : pcm_morph_axiom (preim f (sep g)) (g \o f). -Proof. -split=>[|x y] /=; first by rewrite !pfunit. -by move=>D /andP [H1 H2]; rewrite !pfjoin // !pfV2. -Qed. -HB.instance Definition _ := - isPCM_morphism.Build U W (g \o f) comp_is_morphism. -End CompMorphism. - - -Section CategoricalLaws. -Variables (U V W X : pcm). -Variables (f : pcm_morph V U) (g : pcm_morph W V) (h : pcm_morph X W). - -Lemma morph0L : f \o idfun =m f. -Proof. by []. Qed. - -Lemma morph0R : idfun \o f =m f. -Proof. by split=>// x y; rewrite {1}/sepx /= andbT. Qed. - -Lemma morphA : f \o (g \o h) =m (f \o g) \o h. -Proof. by split=>//= x y H; rewrite /sepx/= andbA. Qed. -End CategoricalLaws. - -(***********************) -(* Cartesian morphisms *) -(***********************) - -(* always-unit function on valid elements *) -Section UnitFun. -Context {U : pcm} {V : tpcm}. - -Definition unit_fun (x : U) : V := - if valid x then Unit else undef. - -Lemma unitfun_is_pcm_morph : pcm_morph_axiom relT unit_fun. -Proof. -rewrite /unit_fun; split=>[|x y W _]. -- by rewrite valid_unit. -by rewrite W (validL W) (validR W) unitL. -Qed. - -HB.instance Definition _ := - isPCM_morphism.Build U V - unit_fun unitfun_is_pcm_morph. -End UnitFun. - -(* pairwise product of morphisms is a morphism *) -Section FProdMorph. -Variables U1 U2 V1 V2 : pcm. -Variables (f1 : pcm_morph U1 V1) (f2 : pcm_morph U2 V2). - -Lemma fprod_is_morph : - pcm_morph_axiom (rel_prod (sep f1) (sep f2)) (f1 \* f2). -Proof. -rewrite /fprod; split=>[|x y]; first by rewrite !pfunit. -by case/andP=>/= W1 W2 /andP [H1 H2]; rewrite pcmE /= !pfV2 //= !pfjoin. -Qed. - -HB.instance Definition _ := - isPCM_morphism.Build (U1 * U2)%type (V1 * V2)%type (f1 \* f2) fprod_is_morph. -End FProdMorph. - -(* projections are morphisms *) -Section ProjMorph. -Variables U1 U2 : pcm. - -Lemma fst_is_morph : pcm_morph_axiom relT (@fst U1 U2). -Proof. by split=>[|x y] // /andP []. Qed. - -HB.instance Definition _ := - isPCM_morphism.Build (U1 * U2)%type U1 fst fst_is_morph. - -Lemma snd_is_morph : pcm_morph_axiom relT (@snd U1 U2). -Proof. by split=>[|x y] // /andP []. Qed. - -HB.instance Definition _ := - isPCM_morphism.Build (U1 * U2)%type U2 snd snd_is_morph. -End ProjMorph. - -Section Proj3Morph. -Variables U1 U2 U3 : pcm. -Notation tp := (Prod3 U1 U2 U3). - -Lemma proj31_morph_ax : pcm_morph_axiom relT (proj31 : tp -> _). -Proof. by split=>[|x y] // /and3P []. Qed. -Lemma proj32_morph_ax : pcm_morph_axiom relT (proj32 : tp -> _). -Proof. by split=>[|x y] // /and3P []. Qed. -Lemma proj33_morph_ax : pcm_morph_axiom relT (proj33 : tp -> _). -Proof. by split=>[|x y] // /and3P []. Qed. - -HB.instance Definition _ := isPCM_morphism.Build tp U1 _ proj31_morph_ax. -HB.instance Definition _ := isPCM_morphism.Build tp U2 _ proj32_morph_ax. -HB.instance Definition _ := isPCM_morphism.Build tp U3 _ proj33_morph_ax. -End Proj3Morph. - -Section Proj4Morph. -Variables U1 U2 U3 U4 : pcm. -Notation tp := (Prod4 U1 U2 U3 U4). - -Lemma proj41_morph_ax : pcm_morph_axiom relT (proj41 : tp -> _). -Proof. by split=>[|x y] // /and4P []. Qed. -Lemma proj42_morph_ax : pcm_morph_axiom relT (proj42 : tp -> _). -Proof. by split=>[|x y] // /and4P []. Qed. -Lemma proj43_morph_ax : pcm_morph_axiom relT (proj43 : tp -> _). -Proof. by split=>[|x y] // /and4P []. Qed. -Lemma proj44_morph_ax : pcm_morph_axiom relT (proj44 : tp -> _). -Proof. by split=>[|x y] // /and4P []. Qed. - -HB.instance Definition _ := isPCM_morphism.Build tp U1 _ proj41_morph_ax. -HB.instance Definition _ := isPCM_morphism.Build tp U2 _ proj42_morph_ax. -HB.instance Definition _ := isPCM_morphism.Build tp U3 _ proj43_morph_ax. -HB.instance Definition _ := isPCM_morphism.Build tp U4 _ proj44_morph_ax. -End Proj4Morph. - -Section Proj5Morph. -Variables U1 U2 U3 U4 U5 : pcm. -Notation tp := (Prod5 U1 U2 U3 U4 U5). - -Lemma proj51_morph_ax : pcm_morph_axiom relT (proj51 : tp -> _). -Proof. by split=>[|x y] // /and5P []. Qed. -Lemma proj52_morph_ax : pcm_morph_axiom relT (proj52 : tp -> _). -Proof. by split=>[|x y] // /and5P []. Qed. -Lemma proj53_morph_ax : pcm_morph_axiom relT (proj53 : tp -> _). -Proof. by split=>[|x y] // /and5P []. Qed. -Lemma proj54_morph_ax : pcm_morph_axiom relT (proj54 : tp -> _). -Proof. by split=>[|x y] // /and5P []. Qed. -Lemma proj55_morph_ax : pcm_morph_axiom relT (proj55 : tp -> _). -Proof. by split=>[|x y] // /and5P []. Qed. - -HB.instance Definition _ := isPCM_morphism.Build tp U1 _ proj51_morph_ax. -HB.instance Definition _ := isPCM_morphism.Build tp U2 _ proj52_morph_ax. -HB.instance Definition _ := isPCM_morphism.Build tp U3 _ proj53_morph_ax. -HB.instance Definition _ := isPCM_morphism.Build tp U4 _ proj54_morph_ax. -HB.instance Definition _ := isPCM_morphism.Build tp U5 _ proj55_morph_ax. -End Proj5Morph. - -Section Proj6Morph. -Variables U1 U2 U3 U4 U5 U6 : pcm. -Notation tp := (Prod6 U1 U2 U3 U4 U5 U6). - -Lemma proj61_morph_ax : pcm_morph_axiom relT (proj61 : tp -> _). -Proof. by split=>[|x y] // /and6P []. Qed. -Lemma proj62_morph_ax : pcm_morph_axiom relT (proj62 : tp -> _). -Proof. by split=>[|x y] // /and6P []. Qed. -Lemma proj63_morph_ax : pcm_morph_axiom relT (proj63 : tp -> _). -Proof. by split=>[|x y] // /and6P []. Qed. -Lemma proj64_morph_ax : pcm_morph_axiom relT (proj64 : tp -> _). -Proof. by split=>[|x y] // /and6P []. Qed. -Lemma proj65_morph_ax : pcm_morph_axiom relT (proj65 : tp -> _). -Proof. by split=>[|x y] // /and6P []. Qed. -Lemma proj66_morph_ax : pcm_morph_axiom relT (proj66 : tp -> _). -Proof. by split=>[|x y] // /and6P []. Qed. - -HB.instance Definition _ := isPCM_morphism.Build tp U1 _ proj61_morph_ax. -HB.instance Definition _ := isPCM_morphism.Build tp U2 _ proj62_morph_ax. -HB.instance Definition _ := isPCM_morphism.Build tp U3 _ proj63_morph_ax. -HB.instance Definition _ := isPCM_morphism.Build tp U4 _ proj64_morph_ax. -HB.instance Definition _ := isPCM_morphism.Build tp U5 _ proj65_morph_ax. -HB.instance Definition _ := isPCM_morphism.Build tp U6 _ proj66_morph_ax. -End Proj6Morph. - -Section Proj7Morph. -Variables U1 U2 U3 U4 U5 U6 U7 : pcm. -Notation tp := (Prod7 U1 U2 U3 U4 U5 U6 U7). - -Lemma proj71_morph_ax : pcm_morph_axiom relT (proj71 : tp -> _). -Proof. by split=>[|x y] // /and7P []. Qed. -Lemma proj72_morph_ax : pcm_morph_axiom relT (proj72 : tp -> _). -Proof. by split=>[|x y] // /and7P []. Qed. -Lemma proj73_morph_ax : pcm_morph_axiom relT (proj73 : tp -> _). -Proof. by split=>[|x y] // /and7P []. Qed. -Lemma proj74_morph_ax : pcm_morph_axiom relT (proj74 : tp -> _). -Proof. by split=>[|x y] // /and7P []. Qed. -Lemma proj75_morph_ax : pcm_morph_axiom relT (proj75 : tp -> _). -Proof. by split=>[|x y] // /and7P []. Qed. -Lemma proj76_morph_ax : pcm_morph_axiom relT (proj76 : tp -> _). -Proof. by split=>[|x y] // /and7P []. Qed. -Lemma proj77_morph_ax : pcm_morph_axiom relT (proj77 : tp -> _). -Proof. by split=>[|x y] // /and7P []. Qed. - -HB.instance Definition _ := isPCM_morphism.Build tp U1 _ proj71_morph_ax. -HB.instance Definition _ := isPCM_morphism.Build tp U2 _ proj72_morph_ax. -HB.instance Definition _ := isPCM_morphism.Build tp U3 _ proj73_morph_ax. -HB.instance Definition _ := isPCM_morphism.Build tp U4 _ proj74_morph_ax. -HB.instance Definition _ := isPCM_morphism.Build tp U5 _ proj75_morph_ax. -HB.instance Definition _ := isPCM_morphism.Build tp U6 _ proj76_morph_ax. -HB.instance Definition _ := isPCM_morphism.Build tp U7 _ proj77_morph_ax. -End Proj7Morph. - -(* product morphism of morphisms is a morphism *) -Section PMorphMorph. -Variables U V1 V2 : pcm. -Variables (f1 : pcm_morph U V1) (f2 : pcm_morph U V2). - -Lemma pmorph_is_morph : pcm_morph_axiom (relI (sep f1) (sep f2)) (f1 \** f2). -Proof. -rewrite /pmorphism; split=>[|x y] /=; first by rewrite !pcmPE !pfunit. -by move=>V /andP [S1 S2]; rewrite pcmE /= !pfV2 // !pfjoin. -Qed. - -HB.instance Definition _ := - isPCM_morphism.Build U (V1 * V2)%type (f1 \** f2) pmorph_is_morph. -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 *) - -(* morphisms for finite products *) - -Section FinFunMorph. -Variables (T : finType) (Us : T -> pcm). -Implicit Types (f : {dffun forall t, Us t}) (t : T). - -Definition app f := [eta f]. - -Lemma app_is_morph : pcm_morph_axiom relT app. -Proof. -split=>[|x y]; first by apply: fext=>t; rewrite /app ffunE. -move/forallP=>V _; split; last by apply: fext=>t; rewrite /app ffunE. -by apply/forallP=>t; move: (V t); rewrite sel_fin. -Qed. - -(* -HB.instance Definition _ := - isPCM_morphism.Build {ffun forall t, Us t} (forall t, Us t) app app_is_morph. -*) - -(* above lemma isn't so useful in practice because can't rewrite *) -(* by pfjoin when sel is fully applied (which is always). *) -(* Thus, we need separate lemma for the case of applied sel. *) -(* This doesn't require function extensionality *) - -Lemma sel_is_morph t : pcm_morph_axiom relT (@sel T Us t). -Proof. by split=>[|x y]; [|move/forallP/(_ t)]; rewrite sel_fin. Qed. - -HB.instance Definition _ t := - isPCM_morphism.Build {dffun forall t, Us t} (Us t) (sel t) (sel_is_morph t). - -Lemma finfun_is_morph : pcm_morph_axiom relT (@finfun T Us : _ -> {dffun _}). -Proof. -split=>[|f1 f2]; first by apply/ffinP=>t; rewrite sel_fin. -move/forallP=>V _; split. -- by apply/forallP=>t; rewrite !sel_fin; apply: V. -by apply/ffunP=>t; rewrite !ffunE !sel_fin. -Qed. - -HB.instance Definition _ := - isPCM_morphism.Build (forall t, Us t) {dffun forall t, Us t} - (@finfun T Us : _ -> {dffun _}) - finfun_is_morph. - -(* Now we can prove pcm-like lemmas for splice *) - -Lemma valid_spliceUn t (x y : {dffun forall t, Us t}) (v : Us t) : - valid (x \+ y) -> - valid (v \+ sel t y) -> - valid (splice x v \+ y). -Proof. -move=>V V'; apply/forallP=>z; rewrite !sel_fin. -by case: eqP=>/= ?; [subst z; rewrite eqc | rewrite pfV2]. -Qed. - -Lemma spliceUn t (x y : {dffun forall t, Us t}) (v w : Us t) : - splice (x \+ y) (v \+ w) = splice x v \+ splice y w. -Proof. -apply/ffinP=>a; rewrite !sel_fin. -by case: eqP=>//= ?; subst a; rewrite !eqc. -Qed. - -End FinFunMorph. - -(****************************) -(****************************) -(* Classes of PCM morphisms *) -(****************************) -(****************************) - -(*********************************) -(* Normal and binormal morphisms *) -(*********************************) - -(* Normal morphisms: map invalids to invalids *) -(* contrapositively: valid images -> valid originals *) -(* basic healthiness property ensuring *) -(* that no invalid element is "rehabilitated" *) -(* by having the morphism map it to something valid *) - -(* Binormal morphism: map overlapping to overlapping *) -(* contrapositively: separated images -> separated originals *) -(* binary variant of normal -- no overlapping elements *) -(* are rehabilitated by having the morphism map them to *) -(* separated ones *) - -(* binormal -> normal but not vice-versa *) - -(* Both properties are needed, as there are useful morphisms *) -(* that are normal but not binormal. *) -(* Example is injection xval : U -> V in xsep construction below, *) -(* where new PCM U is created out of old one V by moding out V by *) -(* seprel D. This creates new notion of separatenss *) -(* by adding D to the separateness of V *) -(* But then, if two elements are separate in V (old notion) *) -(* it doesn't mean they should be separate in U (new notion) *) -(* as the whole point is that new separateness in U should be more *) -(* demaning. In other words, xval isn't binormal *) - -Definition norm_pcm_morph_axiom (U V : pcm) (f : pcm_morph U V) := - forall x : U, valid (f x) -> valid x /\ sep f x Unit. - -Definition binorm_pcm_morph_axiom (U V : pcm) (f : pcm_morph U V) := - forall x y, valid (f x \+ f y) -> valid (x \+ y) /\ sep f x y. - -(* structure definitions *) -HB.mixin Record isNorm_PCM_morphism (U V : pcm) (f : U -> V) - of @PCM_morphism U V f := { - norm_pcm_morph_subproof : norm_pcm_morph_axiom f}. - -#[short(type=norm_pcm_morph)] -HB.structure Definition Norm_PCM_morphism (U V : pcm) := - {f of isNorm_PCM_morphism U V f & @PCM_morphism U V f}. - -(* helper mixin to define binormal structure *) -HB.mixin Record isBinorm_PCM_morphism U V - f of @PCM_morphism U V f := { - binorm_subproof : binorm_pcm_morph_axiom f}. - -(* require both binorm_axiom and norm_axiom for sake of inheritance *) -(* we prove below that binorm_axiom implies norm_axiom *) -(* so we can actually build binorm_pcm_morph just out of binorm_axiom *) -#[short(type=binorm_pcm_morph)] -HB.structure Definition Binorm_PCM_morphism U V := - {f of isBinorm_PCM_morphism U V f & @Norm_PCM_morphism U V f}. - -(* Norm mixin follows from Binorm mixin *) -HB.builders Context U V f of isBinorm_PCM_morphism U V f. - -Lemma binorm_morph_is_norm_morph : norm_pcm_morph_axiom f. -Proof. -move=>x W; rewrite -{1}[x]unitR; apply: binorm_subproof. -by rewrite pfunit unitR. -Qed. -(* default instance of norm_pcm_morph *) -HB.instance Definition _ := - isNorm_PCM_morphism.Build U V f binorm_morph_is_norm_morph. -HB.end. - -(* Normality Lemmas *) -(* We use names that start with f to indicate that these lemmas *) -(* all have validity preconditions over f x *) - -Lemma fpVI (U V : pcm) (f : norm_pcm_morph U V) (x : U) : - valid (f x) -> valid x /\ sep f x Unit. -Proof. exact: norm_pcm_morph_subproof. Qed. - -Lemma fpVE (U V : pcm) (f : norm_pcm_morph U V) (x : U) : - valid (f x) = valid x && sep f x Unit. -Proof. by apply/idP/idP; [move/fpVI/andP|case/andP; apply: pfV]. Qed. - -Lemma fpV (U V : pcm) (f : norm_pcm_morph U V) (x : U) : - valid (f x) -> valid x. -Proof. by case/fpVI. Qed. - -Lemma fpVS (U V : pcm) (f : norm_pcm_morph U V) (x : U) : - valid (f x) -> sep f x Unit. -Proof. by case/fpVI. Qed. - -(* Binormality Lemmas *) - -Lemma fpV2I (U V : pcm) (f : binorm_pcm_morph U V) (x y : U) : - valid (f x \+ f y) -> valid (x \+ y) /\ sep f x y. -Proof. by apply: binorm_subproof. Qed. - -Lemma fpV2E (U V : pcm) (f : binorm_pcm_morph U V) (x y : U) : - valid (f x \+ f y) = valid (x \+ y) && sep f x y. -Proof. by apply/idP/idP; [move/fpV2I/andP|case/andP; apply: pfV2]. Qed. - -Lemma fpV2 (U V : pcm) (f : binorm_pcm_morph U V) (x y : U) : - valid (f x \+ f y) -> valid (x \+ y). -Proof. by case/fpV2I. Qed. - -Lemma fpV2S (U V : pcm) (f : binorm_pcm_morph U V) (x y : U) : - valid (f x \+ f y) -> sep f x y. -Proof. by case/fpV2I. Qed. - -Arguments fpVI {U V}. -Arguments fpVE {U V}. -Arguments fpV {U V}. -Arguments fpVS {U V}. -Arguments fpV2I {U V}. -Arguments fpV2E {U V}. -Arguments fpV2 {U V}. -Arguments fpV2S {U V}. - -(*******************************************) -(* Categorical morphisms and (bi)normality *) -(*******************************************) - -(* id is binormal *) -Lemma id_is_binorm_morphism {U : pcm} : binorm_pcm_morph_axiom (@idfun U). -Proof. by []. Qed. -HB.instance Definition _ (U : pcm) := - isBinorm_PCM_morphism.Build U U idfun id_is_binorm_morphism. - -(* composition preserves (bi)normality *) -Section CompMorphism. -Variables U V W : pcm. - -Section NormCompIsNormal. -Variables (f : norm_pcm_morph U V) (g : norm_pcm_morph V W). - -Lemma comp_is_norm_morphism : norm_pcm_morph_axiom (g \o f). -Proof. -move=>x /fpVI [] /fpVI [H1 H2] H3 /=. -by rewrite /sepx/= H1 H2 pfunit H3. -Qed. - -HB.instance Definition _ := - isNorm_PCM_morphism.Build U W (g \o f) comp_is_norm_morphism. - -End NormCompIsNormal. - -Section BinormCompIsBinorm. -Variables (f : binorm_pcm_morph U V) (g : binorm_pcm_morph V W). - -Lemma comp_is_binorm_morphism : binorm_pcm_morph_axiom (g \o f). -Proof. -move=>x y /fpV2I [/fpV2I][H1 H2 H3]. -by rewrite /sepx/= H1 H2 H3. -Qed. - -HB.instance Definition _ := - isBinorm_PCM_morphism.Build U W (g \o f) comp_is_binorm_morphism. -End BinormCompIsBinorm. -End CompMorphism. - -(***************************************) -(* Cartesian morphisms and binormality *) -(***************************************) - -(* unit_fun is normal, but not binormal *) -Section UnitFun. -Variables (U : pcm) (V : tpcm). - -Lemma unitfun_is_normal : norm_pcm_morph_axiom (@unit_fun U V). -Proof. -move=>x; case W : (valid x)=>//. -by move: (valid_undef V); rewrite /valid/=/unit_fun/= W =>->. -Qed. - -HB.instance Definition _ := - isNorm_PCM_morphism.Build U V - (@unit_fun U V) unitfun_is_normal. -End UnitFun. - -(* projections aren't (bi)normal, but product morphisms are *) - -Section FProdMorph. -Variables U1 U2 V1 V2 : pcm. - -Section Normal. -Variables (f1 : norm_pcm_morph U1 V1) (f2 : norm_pcm_morph U2 V2). - -Lemma fprod_is_norm_pcm_morph : norm_pcm_morph_axiom (f1 \* f2). -Proof. -move=>x /=; rewrite !pcmPV /sepx /=. -by case/andP=>/fpVI [->->] /fpVI [->->]. -Qed. - -HB.instance Definition _ := - isNorm_PCM_morphism.Build (U1 * U2)%type (V1 * V2)%type (f1 \* f2) - fprod_is_norm_pcm_morph. -End Normal. - -Section Binormal. -Variables (f1 : binorm_pcm_morph U1 V1) (f2 : binorm_pcm_morph U2 V2). - -Lemma fprod_is_binorm_pcm_morph : binorm_pcm_morph_axiom (f1 \* f2). -Proof. -move=>x y /=; rewrite !pcmPV /sepx/=. -by case/andP=>/fpV2I [->->] /fpV2I. -Qed. - -HB.instance Definition _ := -isBinorm_PCM_morphism.Build (U1 * U2)%type (V1 * V2)%type (f1 \* f2) - fprod_is_binorm_pcm_morph. -End Binormal. -End FProdMorph. - -Section PMorphMorph. -Variables U V1 V2 : pcm. - -Section Normal. -Variables (f1 : norm_pcm_morph U V1) (f2 : norm_pcm_morph U V2). - -Lemma pmorph_is_norm_pcm_morph : norm_pcm_morph_axiom (f1 \** f2). -Proof. -move=>x; rewrite /sepx /=. -by case/andP=>/fpVI [->->] /fpVI []. -Qed. - -HB.instance Definition _ := - isNorm_PCM_morphism.Build U (V1 * V2)%type (f1 \** f2) - pmorph_is_norm_pcm_morph. -End Normal. - -Section Binormal. -Variables (f1 : binorm_pcm_morph U V1) (f2 : binorm_pcm_morph U V2). - -Lemma pmorph_is_binorm_pcm_morph : binorm_pcm_morph_axiom (f1 \** f2). -Proof. -move=>x y /= /andP [] /fpV2I [-> H1] /fpV2I [_ H2]. -by rewrite /sepx/= H1 H2. -Qed. - -HB.instance Definition _ := - isBinorm_PCM_morphism.Build U (V1 * V2)%type (f1 \** f2) - pmorph_is_binorm_pcm_morph. -End Binormal. -End PMorphMorph. - -Section FinFunMorph. -Variables (T : finType) (Us : T -> pcm). - -Lemma finfun_is_binorm_pcm_morph : - binorm_pcm_morph_axiom (@finfun T Us : _ -> {dffun _}). -Proof. -move=>f1 f2 /forallP /= V; split=>//; apply/forallP=>t. -by move: (V t); rewrite !sel_fin. -Qed. - -HB.instance Definition _ := - isBinorm_PCM_morphism.Build (forall t, Us t) {dffun forall t, Us t} - (@finfun T Us : _ -> {dffun _}) finfun_is_binorm_pcm_morph. -End FinFunMorph. - - -(*****************) -(* Full morphism *) -(*****************) - -(* morphism with seprel relT (largest possible) *) -(* an alternative name would be "total" morphism *) -(* as this morphism is defined for all valid elements *) -(* but we keep "total" for functions that are definied on *) -(* *all* elements, valid or not *) - -Definition full_pcm_morph_axiom (U V : pcm) (f : pcm_morph U V) := - sep f =2 relT. - -HB.mixin Record isFull_PCM_morphism (U V : pcm) (f : U -> V) - of @PCM_morphism U V f := { - full_pcm_morphism_subproof : full_pcm_morph_axiom f}. - -#[short(type=full_pcm_morph)] -HB.structure Definition Full_PCM_morphism (U V : pcm) := - {f of isFull_PCM_morphism U V f & @PCM_morphism U V f}. - -#[short(type=full_norm_pcm_morph)] -HB.structure Definition Full_Norm_PCM_morphism (U V : pcm) := - {f of @Norm_PCM_morphism U V f & @Full_PCM_morphism U V f}. - -#[short(type=full_binorm_pcm_morph)] -HB.structure Definition Full_Binorm_PCM_morphism (U V : pcm) := - {f of @Binorm_PCM_morphism U V f & @Full_PCM_morphism U V f}. - -(* fullness lemmas *) - -Lemma pfSE (U V : pcm) (f : full_pcm_morph U V) : sep f =2 relT. -Proof. by apply: full_pcm_morphism_subproof. Qed. - -Lemma pfT (U V : pcm) (f : full_pcm_morph U V) x y : sep f x y. -Proof. by rewrite pfSE. Qed. - -#[export] Hint Resolve pfT : core. - -Lemma pfVI (U V : pcm) (f : full_pcm_morph U V) (x : U) : - valid x -> - valid (f x). -Proof. by move=>W; apply: pfV W _. Qed. - -Lemma pfV2I (U V : pcm) (f : full_pcm_morph U V) (x y : U) : - valid (x \+ y) -> - valid (f x \+ f y). -Proof. by move=>W; apply: pfV2 W _. Qed. - -Lemma pfV2IC (U V : pcm) (f : full_pcm_morph U V) (x y : U) : - valid (x \+ y) -> - valid (f y \+ f x). -Proof. by rewrite joinC=>/pfV2I. Qed. - -Lemma pfV3I (U V : pcm) (f : full_pcm_morph U V) (x y z : U) : - valid (x \+ y \+ z) -> - valid (f x \+ f y \+ f z). -Proof. by move=>W; rewrite -pfjoin ?(validL W) // pfV2. Qed. - -Lemma pfVE (U V : pcm) (f : full_norm_pcm_morph U V) (x : U) : - valid (f x) = valid x. -Proof. by rewrite fpVE pfT andbT. Qed. - -Lemma pfV2E (U V : pcm) (f : full_binorm_pcm_morph U V) (x y : U) : - valid (f x \+ f y) = valid (x \+ y). -Proof. by rewrite fpV2E pfT andbT. Qed. - -Lemma pfjoinT (U V : pcm) (f : full_pcm_morph U V) (x y : U) : - valid (x \+ y) -> - f (x \+ y) = f x \+ f y. -Proof. by move=>D; rewrite pfjoin. Qed. - -Lemma pfLT (U V : pcm) (f : full_pcm_morph U V) (x y z : U) : - valid (x \+ z) -> - valid (y \+ z) -> - f y = f x -> - f (y \+ z) = f (x \+ z). -Proof. by move=>V1 V2 E; rewrite !pfjoinT // E. Qed. - -Lemma pfRT (U V : pcm) (f : full_pcm_morph U V) (x y z : U) : - valid (z \+ x) -> - valid (z \+ y) -> - f y = f x -> - f (z \+ y) = f (z \+ x). -Proof. by move=>V1 V2 E; rewrite !pfjoinT // E. Qed. - -Lemma pfUnTE (U V : pcm) (f : full_pcm_morph U V) (x1 y1 x2 y2 : U) : - valid (x1 \+ y1) -> - valid (x2 \+ y2) -> - f x2 = f x1 -> - f y2 = f y1 -> - f (x2 \+ y2) = f (x1 \+ y1). -Proof. by move=>V1 V2 E1 E2; rewrite !pfjoinT // E1 E2. Qed. - -Arguments pfVI {U V f x}. -Arguments pfV2I {U V f x y}. -Arguments pfVE {U V f x}. -Arguments pfV2E {U V f x y}. -Arguments pfjoinT {U V f x}. -Arguments pfLT {U V} f {x y}. -Arguments pfRT {U V} f {x y}. - -(* Categorical morphisms and fullness *) - -(* id is full *) -Lemma id_is_full_morphism {U : pcm} : full_pcm_morph_axiom (@idfun U). -Proof. by []. Qed. -HB.instance Definition _ (U : pcm) := - isFull_PCM_morphism.Build U U idfun id_is_full_morphism. - -(* composition preserves fullness *) -Section CompMorphism. -Variables U V W : pcm. - -Section FullCompIsFull. -Variables (f : full_pcm_morph U V) (g : full_pcm_morph V W). - -Lemma comp_is_full_morphism : full_pcm_morph_axiom (g \o f). -Proof. by move=>x y /=; rewrite /sepx/= !pfT. Qed. - -HB.instance Definition _ := - isFull_PCM_morphism.Build U W (g \o f) comp_is_full_morphism. -End FullCompIsFull. - -(* instances for combinations must declare PCM_morphism.on *) -(* before declaring fullness structure. *) -(* DEVCOMMENT: *) -(* Alternative is to write the definition explicitly as below, *) -(* but that's too low level. *) -(* -HB.instance Definition _ (f : full_norm_pcm_morph U V) - (g : full_norm_pcm_morph V W) := - Full_Norm_PCM_morphism.copy (g \o f) - (Full_Norm_PCM_morphism.pack_ - (Norm_PCM_morphism.class (g \o f)) - (Full_PCM_morphism.class (g \o f))). *) -(* /DEVCOMMENT *) - -HB.instance Definition _ - (f : full_norm_pcm_morph U V) - (g : full_norm_pcm_morph V W) := - PCM_morphism.on (g \o f). -HB.instance Definition _ - (f : full_norm_pcm_morph U V) - (g : full_norm_pcm_morph V W) := - Full_Norm_PCM_morphism.on (g \o f). - -HB.instance Definition _ - (f : full_binorm_pcm_morph U V) - (g : full_binorm_pcm_morph V W) := - PCM_morphism.on (g \o f). -HB.instance Definition _ - (f : full_binorm_pcm_morph U V) - (g : full_binorm_pcm_morph V W) := - Full_Binorm_PCM_morphism.on (g \o f). - -End CompMorphism. - -(* Cartesian morphisms and fullness *) - -Section UnitFun. -Variables (U : pcm) (V : tpcm). - -Lemma unitfun_is_full_pcm_morph : full_pcm_morph_axiom (@unit_fun U V). -Proof. by []. Qed. - -HB.instance Definition _ := - isFull_PCM_morphism.Build U V - (@unit_fun U V) unitfun_is_full_pcm_morph. -End UnitFun. - - -Section Cartesians. -Notation pf := (fun _ _ => erefl _). - -Section FProdMorph. -Variables U1 U2 V1 V2 : pcm. -Variables (f1 : full_pcm_morph U1 V1) (f2 : full_pcm_morph U2 V2). - -Lemma fprod_is_full_pcm_morph : full_pcm_morph_axiom (f1 \* f2). -Proof. by move=>x y /=; rewrite /sepx/=!pfT. Qed. - -HB.instance Definition _ := - isFull_PCM_morphism.Build (U1 * U2)%type (V1 * V2)%type (f1 \* f2) - fprod_is_full_pcm_morph. -End FProdMorph. - -Section ProjMorph. -Variables U1 U2 : pcm. -HB.instance Definition _ := isFull_PCM_morphism.Build (U1 * U2)%type U1 fst pf. -HB.instance Definition _ := isFull_PCM_morphism.Build (U1 * U2)%type U2 snd pf. -End ProjMorph. - -Section Proj3Morph. -Variables U1 U2 U3 : pcm. -Notation tp := (Prod3 U1 U2 U3). -HB.instance Definition _ := isFull_PCM_morphism.Build tp U1 proj31 pf. -HB.instance Definition _ := isFull_PCM_morphism.Build tp U2 proj32 pf. -HB.instance Definition _ := isFull_PCM_morphism.Build tp U3 proj33 pf. -End Proj3Morph. - -Section Proj4Morph. -Variables U1 U2 U3 U4 : pcm. -Notation tp := (Prod4 U1 U2 U3 U4). -HB.instance Definition _ := isFull_PCM_morphism.Build tp U1 proj41 pf. -HB.instance Definition _ := isFull_PCM_morphism.Build tp U2 proj42 pf. -HB.instance Definition _ := isFull_PCM_morphism.Build tp U3 proj43 pf. -HB.instance Definition _ := isFull_PCM_morphism.Build tp U4 proj44 pf. -End Proj4Morph. - -Section Proj5Morph. -Variables U1 U2 U3 U4 U5 : pcm. -Notation tp := (Prod5 U1 U2 U3 U4 U5). -HB.instance Definition _ := isFull_PCM_morphism.Build tp U1 proj51 pf. -HB.instance Definition _ := isFull_PCM_morphism.Build tp U2 proj52 pf. -HB.instance Definition _ := isFull_PCM_morphism.Build tp U3 proj53 pf. -HB.instance Definition _ := isFull_PCM_morphism.Build tp U4 proj54 pf. -HB.instance Definition _ := isFull_PCM_morphism.Build tp U5 proj55 pf. -End Proj5Morph. - -Section Proj6Morph. -Variables U1 U2 U3 U4 U5 U6 : pcm. -Notation tp := (Prod6 U1 U2 U3 U4 U5 U6). -HB.instance Definition _ := isFull_PCM_morphism.Build tp U1 proj61 pf. -HB.instance Definition _ := isFull_PCM_morphism.Build tp U2 proj62 pf. -HB.instance Definition _ := isFull_PCM_morphism.Build tp U3 proj63 pf. -HB.instance Definition _ := isFull_PCM_morphism.Build tp U4 proj64 pf. -HB.instance Definition _ := isFull_PCM_morphism.Build tp U5 proj65 pf. -HB.instance Definition _ := isFull_PCM_morphism.Build tp U6 proj66 pf. -End Proj6Morph. - -Section Proj7Morph. -Variables U1 U2 U3 U4 U5 U6 U7 : pcm. -Notation tp := (Prod7 U1 U2 U3 U4 U5 U6 U7). -HB.instance Definition _ := isFull_PCM_morphism.Build tp U1 proj71 pf. -HB.instance Definition _ := isFull_PCM_morphism.Build tp U2 proj72 pf. -HB.instance Definition _ := isFull_PCM_morphism.Build tp U3 proj73 pf. -HB.instance Definition _ := isFull_PCM_morphism.Build tp U4 proj74 pf. -HB.instance Definition _ := isFull_PCM_morphism.Build tp U5 proj75 pf. -HB.instance Definition _ := isFull_PCM_morphism.Build tp U6 proj76 pf. -HB.instance Definition _ := isFull_PCM_morphism.Build tp U7 proj77 pf. -End Proj7Morph. - -Section PMorphMorph. -Variables U V1 V2 : pcm. -Variables (f1 : full_pcm_morph U V1) (f2 : full_pcm_morph U V2). - -Lemma pmorph_is_full_pcm_morph : full_pcm_morph_axiom (f1 \** f2). -Proof. by move=>x y /=; rewrite /sepx/= !pfT. Qed. - -HB.instance Definition _ := - isFull_PCM_morphism.Build U (V1 * V2)%type (f1 \** f2) - pmorph_is_full_pcm_morph. -End PMorphMorph. - -Section FinFunMorph. -HB.instance Definition _ (T : finType) (Us : T -> pcm) (t : T) := - isFull_PCM_morphism.Build {dffun forall t, Us t} (Us t) (sel t) pf. - -HB.instance Definition _ (T : finType) (Us : T -> pcm) := - isFull_PCM_morphism.Build (forall t, Us t) {dffun forall t, Us t} - (@finfun T Us : _ -> {dffun _}) pf. -End FinFunMorph. - -End Cartesians. - -(* fullness and other constructions *) - -Section RelApp. -Variables (U V : pcm) (S : seprel V) (f : full_pcm_morph U V). - -Lemma relapp_is_seprel : seprel_axiom (rel_app S f). -Proof. -split=>[|x y|x y|x y z] /=. -- by rewrite pfunit sep00. -- by move=>W; rewrite sepC // pfV2I. -- by move=>W; rewrite pfunit; move/sep0E=>-> //; rewrite pfV2I. -move=>W; rewrite !pfjoinT ?(validLE3 W) // => X Y. -by rewrite !(sepAxx _ X Y) // -!pfjoinT ?(validLE3 W) ?pfVI. -Qed. - -HB.instance Definition _ := - isSeprel.Build U (rel_app S f) relapp_is_seprel. -End RelApp. - -Section Pmorph. -Variables (U V1 V2 : pcm). -Variables (f1 : full_pcm_morph U V1) (f2 : full_pcm_morph U V2). - -Lemma pmorphism_is_full : full_pcm_morph_axiom (f1 \** f2). -Proof. by move=>x y /=; rewrite !pfT. Qed. - -HB.instance Definition _ := - isFull_PCM_morphism.Build U (V1 * V2)%type - (f1 \** f2) pmorphism_is_full. -End Pmorph. - -(******************) -(******************) -(* TPCM morphisms *) -(******************) -(******************) - -(* tpcm morphisms further preserve undef *) - -Definition tpcm_morph_axiom (U V : tpcm) (f : U -> V) := - f undef = undef. - -HB.mixin Record isTPCM_morphism (U V : tpcm) (f : U -> V) := { - tpcm_morphism_subproof : tpcm_morph_axiom f}. - -#[short(type=tpcm_morph)] -HB.structure Definition TPCM_morphism (U V : tpcm) := - {f of isTPCM_morphism U V f & @PCM_morphism U V f}. - -(* introduce names for the combinations of sub-properties *) - -#[short(type=norm_tpcm_morph)] -HB.structure Definition Norm_TPCM_morphism (U V : tpcm) := - {f of @Norm_PCM_morphism U V f & @TPCM_morphism U V f}. - -#[short(type=binorm_tpcm_morph)] -HB.structure Definition Binorm_TPCM_morphism (U V : tpcm) := - {f of @Binorm_PCM_morphism U V f & @TPCM_morphism U V f}. - -#[short(type=full_tpcm_morph)] -HB.structure Definition Full_TPCM_morphism (U V : tpcm) := - {f of @Full_PCM_morphism U V f & @TPCM_morphism U V f}. - -#[short(type=full_norm_tpcm_morph)] -HB.structure Definition Full_Norm_TPCM_morphism (U V : tpcm) := - {f of @Full_PCM_morphism U V f & @Norm_TPCM_morphism U V f}. - -#[short(type=full_binorm_tpcm_morph)] -HB.structure Definition Full_Binorm_TPCM_morphism (U V : tpcm) := - {f of @Full_PCM_morphism U V f & @Binorm_TPCM_morphism U V f}. - -(* TPCM lemmas *) - -Lemma pfundef {U V : tpcm} (f : tpcm_morph U V) : f undef = undef. -Proof. by apply: tpcm_morphism_subproof. Qed. - -(* full morphism on normal tpcm is normal *) - -Lemma fullmorph_is_norm (U : normal_tpcm) V (f : full_tpcm_morph U V) : - norm_pcm_morph_axiom f. -Proof. -move=>x; rewrite pfT; case: (normalP x)=>[->|] //=. -by rewrite pfundef valid_undef. -Qed. -Arguments fullmorph_is_norm {U V f}. - -(* Following lemma is a crutch whose use is discouraged. *) -(* Instead, declare f as normal, and use pfVE. *) -Lemma pfnVE (U : normal_tpcm) V (f : full_tpcm_morph U V) x : - valid (f x) = valid x. -Proof. by apply/idP/idP; [case/fullmorph_is_norm|apply/pfVI]. Qed. - -(* Categoricals are tpcm morphism *) - -HB.instance Definition _ (U : tpcm) := - isTPCM_morphism.Build U U (@idfun U) (erefl _). - -Section Comp. -Variables U V W : tpcm. - -Section CompTPCM. -Variables (f : tpcm_morph U V) (g : tpcm_morph V W). - -Lemma comp_is_tpcm_morphism : tpcm_morph_axiom (g \o f). -Proof. by rewrite /tpcm_morph_axiom /= !pfundef. Qed. - -HB.instance Definition _ := - isTPCM_morphism.Build U W (g \o f) comp_is_tpcm_morphism. -End CompTPCM. - -(* combinations declared explicitly *) - -HB.instance Definition _ - (f : norm_tpcm_morph U V) - (g : norm_tpcm_morph V W) := - PCM_morphism.on (g \o f). -HB.instance Definition _ - (f : norm_tpcm_morph U V) - (g : norm_tpcm_morph V W) := - Norm_TPCM_morphism.on (g \o f). - -HB.instance Definition _ - (f : binorm_tpcm_morph U V) - (g : binorm_tpcm_morph V W) := - PCM_morphism.on (g \o f). -HB.instance Definition _ - (f : binorm_tpcm_morph U V) - (g : binorm_tpcm_morph V W) := - Binorm_TPCM_morphism.on (g \o f). - -HB.instance Definition _ - (f : full_tpcm_morph U V) - (g : full_tpcm_morph V W) := - PCM_morphism.on (g \o f). -HB.instance Definition _ - (f : full_tpcm_morph U V) - (g : full_tpcm_morph V W) := - Full_TPCM_morphism.on (g \o f). - -HB.instance Definition _ - (f : full_norm_tpcm_morph U V) - (g : full_norm_tpcm_morph V W) := - PCM_morphism.on (g \o f). -HB.instance Definition _ - (f : full_norm_tpcm_morph U V) - (g : full_norm_tpcm_morph V W) := - Full_Norm_TPCM_morphism.on (g \o f). - -HB.instance Definition _ - (f : full_binorm_tpcm_morph U V) - (g : full_binorm_tpcm_morph V W) := - PCM_morphism.on (g \o f). -HB.instance Definition _ - (f : full_binorm_tpcm_morph U V) - (g : full_binorm_tpcm_morph V W) := - Full_Binorm_TPCM_morphism.on (g \o f). - -End Comp. - -(* Cartesians *) - -Section UnitFun. -Variables (U V : tpcm). - -Lemma unitfun_is_tpcm_morph : tpcm_morph_axiom (@unit_fun U V). -Proof. by rewrite /tpcm_morph_axiom/unit_fun valid_undef. Qed. - -HB.instance Definition _ := - isTPCM_morphism.Build U V - (@unit_fun U V) unitfun_is_tpcm_morph. -End UnitFun. - - -Section FProdMorph. -Variables U1 U2 V1 V2 : tpcm. -Variables (f1 : tpcm_morph U1 V1) (f2 : tpcm_morph U2 V2). - -Lemma fprod_is_tpcm_morph : tpcm_morph_axiom (f1 \* f2). -Proof. by rewrite /tpcm_morph_axiom /fprod !pfundef. Qed. - -HB.instance Definition _ := - isTPCM_morphism.Build (U1 * U2)%type (V1 * V2)%type (f1 \* f2) - fprod_is_tpcm_morph. -End FProdMorph. - -(* projections are tpcm morphisms *) -Section ProjMorph. -Variables U1 U2 : tpcm. - -Lemma fst_is_tpcm_morph : tpcm_morph_axiom (@fst U1 U2). -Proof. by rewrite /tpcm_morph_axiom /undef. Qed. - -HB.instance Definition _ := - isTPCM_morphism.Build (U1 * U2)%type U1 fst fst_is_tpcm_morph. - -Lemma snd_is_tpcm_morph : tpcm_morph_axiom (@snd U1 U2). -Proof. by rewrite /tpcm_morph_axiom /undef. Qed. - -HB.instance Definition _ := - isTPCM_morphism.Build (U1 * U2)%type U2 snd snd_is_tpcm_morph. -End ProjMorph. - -(* projections for iterated products are morphisms *) - -Section Proj3Morph. -Variables U1 U2 U3 : tpcm. -Notation tp := (Prod3 U1 U2 U3). - -Lemma proj31_is_tpcm_morph : tpcm_morph_axiom (proj31 : tp -> _). -Proof. by rewrite /tpcm_morph_axiom /undef. Qed. -Lemma proj32_is_tpcm_morph : tpcm_morph_axiom (proj32 : tp -> _). -Proof. by rewrite /tpcm_morph_axiom /undef. Qed. -Lemma proj33_is_tpcm_morph : tpcm_morph_axiom (proj33 : tp -> _). -Proof. by rewrite /tpcm_morph_axiom /undef. Qed. - -HB.instance Definition _ := isTPCM_morphism.Build tp U1 _ proj31_is_tpcm_morph. -HB.instance Definition _ := isTPCM_morphism.Build tp U2 _ proj32_is_tpcm_morph. -HB.instance Definition _ := isTPCM_morphism.Build tp U3 _ proj33_is_tpcm_morph. -End Proj3Morph. - -Section Proj4Morph. -Variables U1 U2 U3 U4 : tpcm. -Notation tp := (Prod4 U1 U2 U3 U4). - -Lemma proj41_is_tpcm_morph : tpcm_morph_axiom (proj41 : tp -> _). -Proof. by rewrite /tpcm_morph_axiom /undef. Qed. -Lemma proj42_is_tpcm_morph : tpcm_morph_axiom (proj42 : tp -> _). -Proof. by rewrite /tpcm_morph_axiom /undef. Qed. -Lemma proj43_is_tpcm_morph : tpcm_morph_axiom (proj43 : tp -> _). -Proof. by rewrite /tpcm_morph_axiom /undef. Qed. -Lemma proj44_is_tpcm_morph : tpcm_morph_axiom (proj44 : tp -> _). -Proof. by rewrite /tpcm_morph_axiom /undef. Qed. - -HB.instance Definition _ := isTPCM_morphism.Build tp U1 _ proj41_is_tpcm_morph. -HB.instance Definition _ := isTPCM_morphism.Build tp U2 _ proj42_is_tpcm_morph. -HB.instance Definition _ := isTPCM_morphism.Build tp U3 _ proj43_is_tpcm_morph. -HB.instance Definition _ := isTPCM_morphism.Build tp U4 _ proj44_is_tpcm_morph. -End Proj4Morph. - -Section Proj5Morph. -Variables U1 U2 U3 U4 U5 : tpcm. -Notation tp := (Prod5 U1 U2 U3 U4 U5). - -Lemma proj51_is_tpcm_morph : tpcm_morph_axiom (proj51 : tp -> _). -Proof. by rewrite /tpcm_morph_axiom /undef. Qed. -Lemma proj52_is_tpcm_morph : tpcm_morph_axiom (proj52 : tp -> _). -Proof. by rewrite /tpcm_morph_axiom /undef. Qed. -Lemma proj53_is_tpcm_morph : tpcm_morph_axiom (proj53 : tp -> _). -Proof. by rewrite /tpcm_morph_axiom /undef. Qed. -Lemma proj54_is_tpcm_morph : tpcm_morph_axiom (proj54 : tp -> _). -Proof. by rewrite /tpcm_morph_axiom /undef. Qed. -Lemma proj55_is_tpcm_morph : tpcm_morph_axiom (proj55 : tp -> _). -Proof. by rewrite /tpcm_morph_axiom /undef. Qed. - -HB.instance Definition _ := isTPCM_morphism.Build tp U1 _ proj51_is_tpcm_morph. -HB.instance Definition _ := isTPCM_morphism.Build tp U2 _ proj52_is_tpcm_morph. -HB.instance Definition _ := isTPCM_morphism.Build tp U3 _ proj53_is_tpcm_morph. -HB.instance Definition _ := isTPCM_morphism.Build tp U4 _ proj54_is_tpcm_morph. -HB.instance Definition _ := isTPCM_morphism.Build tp U5 _ proj55_is_tpcm_morph. -End Proj5Morph. - -Section Proj6Morph. -Variables U1 U2 U3 U4 U5 U6 : tpcm. -Notation tp := (Prod6 U1 U2 U3 U4 U5 U6). - -Lemma proj61_is_tpcm_morph : tpcm_morph_axiom (proj61 : tp -> _). -Proof. by rewrite /tpcm_morph_axiom /undef. Qed. -Lemma proj62_is_tpcm_morph : tpcm_morph_axiom (proj62 : tp -> _). -Proof. by rewrite /tpcm_morph_axiom /undef. Qed. -Lemma proj63_is_tpcm_morph : tpcm_morph_axiom (proj63 : tp -> _). -Proof. by rewrite /tpcm_morph_axiom /undef. Qed. -Lemma proj64_is_tpcm_morph : tpcm_morph_axiom (proj64 : tp -> _). -Proof. by rewrite /tpcm_morph_axiom /undef. Qed. -Lemma proj65_is_tpcm_morph : tpcm_morph_axiom (proj65 : tp -> _). -Proof. by rewrite /tpcm_morph_axiom /undef. Qed. -Lemma proj66_is_tpcm_morph : tpcm_morph_axiom (proj66 : tp -> _). -Proof. by rewrite /tpcm_morph_axiom /undef. Qed. - -HB.instance Definition _ := isTPCM_morphism.Build tp U1 _ proj61_is_tpcm_morph. -HB.instance Definition _ := isTPCM_morphism.Build tp U2 _ proj62_is_tpcm_morph. -HB.instance Definition _ := isTPCM_morphism.Build tp U3 _ proj63_is_tpcm_morph. -HB.instance Definition _ := isTPCM_morphism.Build tp U4 _ proj64_is_tpcm_morph. -HB.instance Definition _ := isTPCM_morphism.Build tp U5 _ proj65_is_tpcm_morph. -HB.instance Definition _ := isTPCM_morphism.Build tp U6 _ proj66_is_tpcm_morph. -End Proj6Morph. - -Section Proj7Morph. -Variables U1 U2 U3 U4 U5 U6 U7 : tpcm. -Notation tp := (Prod7 U1 U2 U3 U4 U5 U6 U7). - -Lemma proj71_is_tpcm_morph : tpcm_morph_axiom (proj71 : tp -> _). -Proof. by rewrite /tpcm_morph_axiom /undef. Qed. -Lemma proj72_is_tpcm_morph : tpcm_morph_axiom (proj72 : tp -> _). -Proof. by rewrite /tpcm_morph_axiom /undef. Qed. -Lemma proj73_is_tpcm_morph : tpcm_morph_axiom (proj73 : tp -> _). -Proof. by rewrite /tpcm_morph_axiom /undef. Qed. -Lemma proj74_is_tpcm_morph : tpcm_morph_axiom (proj74 : tp -> _). -Proof. by rewrite /tpcm_morph_axiom /undef. Qed. -Lemma proj75_is_tpcm_morph : tpcm_morph_axiom (proj75 : tp -> _). -Proof. by rewrite /tpcm_morph_axiom /undef. Qed. -Lemma proj76_is_tpcm_morph : tpcm_morph_axiom (proj76 : tp -> _). -Proof. by rewrite /tpcm_morph_axiom /undef. Qed. -Lemma proj77_is_tpcm_morph : tpcm_morph_axiom (proj77 : tp -> _). -Proof. by rewrite /tpcm_morph_axiom /undef. Qed. - -HB.instance Definition _ := isTPCM_morphism.Build tp U1 _ proj71_is_tpcm_morph. -HB.instance Definition _ := isTPCM_morphism.Build tp U2 _ proj72_is_tpcm_morph. -HB.instance Definition _ := isTPCM_morphism.Build tp U3 _ proj73_is_tpcm_morph. -HB.instance Definition _ := isTPCM_morphism.Build tp U4 _ proj74_is_tpcm_morph. -HB.instance Definition _ := isTPCM_morphism.Build tp U5 _ proj75_is_tpcm_morph. -HB.instance Definition _ := isTPCM_morphism.Build tp U6 _ proj76_is_tpcm_morph. -HB.instance Definition _ := isTPCM_morphism.Build tp U7 _ proj77_is_tpcm_morph. -End Proj7Morph. - -(* product morphism of tpcm morphisms is a tpcm morphism *) -Section PMorphMorph. -Variables U V1 V2 : tpcm. -Variables (f1 : tpcm_morph U V1) (f2 : tpcm_morph U V2). - -Lemma pmorph_is_tpcm_morph : tpcm_morph_axiom (f1 \** f2). -Proof. by rewrite /tpcm_morph_axiom/pmorphism !pfundef. Qed. - -HB.instance Definition _ := - isTPCM_morphism.Build U (V1 * V2)%type (f1 \** f2) pmorph_is_tpcm_morph. - -End PMorphMorph. - -(* combination of full and tpcm declared by hand *) -HB.instance Definition _ (U V1 V2 : tpcm) - (f1 : full_tpcm_morph U V1) (f2 : full_tpcm_morph U V2) := - Full_TPCM_morphism.copy (f1 \** f2) - (Full_TPCM_morphism.pack_ - (TPCM_morphism.class (f1 \** f2)) - (Full_PCM_morphism.class (f1 \** f2))). - -(* finite products *) -Section FinFunMorph. -Variables (T : ifinType) (Us : T -> tpcm). - -Lemma sel_is_tpcm_morph t : tpcm_morph_axiom (sel (Us:=Us) t). -Proof. by rewrite /tpcm_morph_axiom sel_fin. Qed. - -HB.instance Definition _ (t : T) := - isTPCM_morphism.Build {dffun forall t, Us t} (Us t) (sel t) - (sel_is_tpcm_morph t). - -Lemma finfun_is_tpcm_morph : - tpcm_morph_axiom (V:={dffun forall t, Us t}) finfun. -Proof. by []. Qed. - -HB.instance Definition _ := - isTPCM_morphism.Build (forall t, Us t) {dffun forall t, Us t} - (@finfun T Us : _ -> {dffun _}) finfun_is_tpcm_morph. -End FinFunMorph. - -(* restriction *) -Section RestrictionTPCM. -Variables (U V : tpcm) (f : tpcm_morph U V) (S : seprel U). - -Lemma res_is_tpcm_morph : tpcm_morph_axiom (res f S). -Proof. by rewrite /tpcm_morph_axiom /res pfundef if_same. Qed. - -HB.instance Definition _ := - isTPCM_morphism.Build U V (res f S) res_is_tpcm_morph. -End RestrictionTPCM. - -(* join *) -Section JoinFunTPCM. -Variables (U V : tpcm) (f1 f2 : tpcm_morph U V). - -Lemma joinfun_is_tpcm_morph : tpcm_morph_axiom (join_fun f1 f2). -Proof. by rewrite /tpcm_morph_axiom/join_fun !pfundef join_undef. Qed. - -HB.instance Definition _ := - isTPCM_morphism.Build U V (join_fun f1 f2) joinfun_is_tpcm_morph. -End JoinFunTPCM. - -(********************) -(* Option morphisms *) -(********************) - -Definition odfltu {A : pcm} (x : option A) := odflt Unit x. - -(* Some is full binormal morphism *) -Lemma some_is_morph {A : pcm} : pcm_morph_axiom relT (@Some A). -Proof. by []. Qed. -HB.instance Definition _ A := - isPCM_morphism.Build A (option A) Some some_is_morph. -Lemma some_is_binorm {A : pcm} : binorm_pcm_morph_axiom (@Some A). -Proof. by []. Qed. -HB.instance Definition _ A := - isBinorm_PCM_morphism.Build A (option A) Some some_is_binorm. -HB.instance Definition _ A := - isFull_PCM_morphism.Build A (option A) Some (fun _ _ => erefl _). - -(* odflt is full morphism, but doesn't preserve undef (hence, not normal) *) -Lemma odfltu_is_morph {A : pcm} : pcm_morph_axiom relT (@odfltu A). -Proof. by split=>//; case=>[x|][y|]. Qed. -HB.instance Definition _ (A : pcm) := - isPCM_morphism.Build (option A) A odfltu odfltu_is_morph. -HB.instance Definition _ (A : pcm) := - isFull_PCM_morphism.Build (option A) A odfltu (fun _ _ => erefl _). - -(* strenghtening of pfjoin for Some morphism (elides validity) *) -Lemma pfsome (A : pcm) (x y : A) : Some x \+ Some y = Some (x \+ y). -Proof. by []. Qed. - -Lemma odflt_some (A : pcm) (x : A) : odfltu (Some x) = x. -Proof. by []. Qed. - -Lemma odflt_someUn (A : pcm) (x y : A) : odfltu (Some x \+ Some y) = x \+ y. -Proof. by []. Qed. - -Lemma some_odflt (A : pcm) (x : option A) : valid x -> Some (odfltu x) = x. -Proof. by case: x. Qed. - -Lemma some_odfltUn (A : pcm) (x y : option A) : - valid (x \+ y) -> Some (odfltu x \+ odfltu y) = x \+ y. -Proof. by case: x; case: y. Qed. - -(* constructions for option nat, specifically *) - -Lemma option_nat_is_normal : normal_tpcm_axiom (option nat). -Proof. by apply: option_is_normal. Qed. -HB.instance Definition _ := - isNormal_TPCM.Build (option nat) option_nat_is_normal. - -Lemma onat0E (x y : option nat) : x \+ y = Unit -> (x = Unit) * (y = Unit). -Proof. -case: x=>[x|]; case: y=>[y|] // [] /eqP. -by rewrite addn_eq0=>/andP [/eqP -> /eqP ->]. -Qed. - - -(**********) -(**********) -(* SubPCM *) -(**********) -(**********) - -(* Subpcm structure between U and V consists of two morphisms *) -(* pval : U -> V and psub : V -> U with special properties. *) -(* The definition thus resembles Galois connections and adjunctions *) -(* which also postulate existence of morphisms in the opposite directions. *) -(* This formulation allows reasoning about composition *) -(* of subpcm structures *) - -(* We thus define the structure collecting the two morphisms *) -(* (with identity and composition) *) - -Record sub_struct (U V : Type) := SubStruct { - pval : U -> V; - psub : V -> U}. - -Arguments pval : simpl never. -Arguments psub : simpl never. - -Definition id_sub {U} := @SubStruct U U idfun idfun. - -Definition comp_sub U V W (X : sub_struct U V) (Y : sub_struct V W) := - SubStruct (pval Y \o pval X) (psub X \o psub Y). - -(* explicit rewrite rules *) -Lemma pval_id U : pval (@id_sub U) = idfun. Proof. by []. Qed. -Lemma psub_id U : psub (@id_sub U) = idfun. Proof. by []. Qed. -Lemma pval_comp U V W X Y : - pval (@comp_sub U V W X Y) = pval Y \o pval X. -Proof. by []. Qed. -Lemma psub_comp U V W X Y : - psub (@comp_sub U V W X Y) = psub X \o psub Y. -Proof. by []. Qed. - -(* subpcm axoms *) -(* pval - normality needed so that transitions of subresources *) -(* don't require side condition on validity *) -(* pval - fullness is required for simplicity *) -(* psub - binormality needed so that xsep (defined below) is pcm *) -(* remaining two axioms are expected for injection/retraction pair *) - -Definition subpcm_struct_axiom' (U V : pcm) - (pval : full_norm_pcm_morph U V) - (psub : binorm_pcm_morph V U) := - [/\ (* inject then retract is id *) - forall u, psub (pval u) = u & - (* retract then inject is id on valid elements *) - forall v, valid v -> sep psub v Unit -> pval (psub v) = v]. - -Notation subpcm_struct_axiom S := - (subpcm_struct_axiom' (pval S) (psub S)). - -HB.mixin Record isSubPCM_struct (U V : pcm) (S : sub_struct U V) := { - pval_submix : @Full_Norm_PCM_morphism U V (pval S); - psub_submix : @Binorm_PCM_morphism V U (psub S); - subpcm_struct_subproof : - subpcm_struct_axiom' - (Full_Norm_PCM_morphism.Pack pval_submix) - (Binorm_PCM_morphism.Pack psub_submix)}. - -#[short(type=subpcm_struct)] -HB.structure Definition SubPCM_struct U V := {S of isSubPCM_struct U V S}. - -Arguments subpcm_struct_subproof {U V}. - -(* declare pval/psub to be pcm morphisms *) -HB.instance Definition _ (U V : pcm) (S : subpcm_struct U V) := - Full_Norm_PCM_morphism.copy (pval S) - (Full_Norm_PCM_morphism.Pack pval_submix). -HB.instance Definition _ (U V : pcm) (S : subpcm_struct U V) := - Binorm_PCM_morphism.copy (psub S) - (Binorm_PCM_morphism.Pack psub_submix). - -Notation subsep S := (sep (psub S)). - -Section Repack. -Variables (U V : pcm) (S : subpcm_struct U V). - -Lemma psub_pval (x : U) : psub S (pval S x) = x. -Proof. by case: (subpcm_struct_subproof S)=>H ?; apply: H. Qed. - -Lemma pval_psub (x : V) : - valid x -> subsep S x Unit -> pval S (psub S x) = x. -Proof. by case: (subpcm_struct_subproof S)=>?; apply. Qed. - -End Repack. - -Arguments psub_pval {U V}. -Arguments pval_psub {U V}. - -Lemma sep_pval (U V : pcm) (S : subpcm_struct U V) x y : sep (pval S) x y. -Proof. exact: pfT. Qed. - -#[export] Hint Resolve sep_pval : core. - -Section DerivedLemmas. -Variables (U V : pcm) (S : subpcm_struct U V). - -(* unary lemmas *) - -Lemma valid_sepE (x : U) : - valid x = valid (pval S x) && subsep S (pval S x) Unit. -Proof. by rewrite -fpVE /= psub_pval. Qed. - -Lemma valid_sep (x : U) : - valid x -> valid (pval S x) /\ subsep S (pval S x) Unit. -Proof. by rewrite valid_sepE=>/andP. Qed. - -Lemma valid_pval (x : U) : valid x -> valid (pval S x). -Proof. by case/valid_sep. Qed. - -Lemma valid_sepS (x : U) : valid x -> subsep S (pval S x) Unit. -Proof. by case/valid_sep. Qed. - -(* pvalE is full and normal *) -Lemma valid_pvalE (x : U) : valid (pval S x) = valid x. -Proof. by rewrite pfVE. Qed. - -(* iff variant for convenience *) -Lemma valid_pvalEP (x : U) : valid (pval S x) <-> valid x. -Proof. by rewrite valid_pvalE. Qed. - -Lemma valid_pvalS (x : U) : valid (pval S x) -> subsep S (pval S x) Unit. -Proof. by move/valid_pvalEP/valid_sepS. Qed. - -(* we can collect the two conditions of pval_psub into one *) -Lemma valid_psubS (x : V) : valid (psub S x) -> pval S (psub S x) = x. -Proof. by case/fpVI; apply: pval_psub. Qed. - - -(* binary lemmas *) - -Lemma valid_sepUnE (x y : U) : - valid (x \+ y) = - valid (pval S x \+ pval S y) && subsep S (pval S x) (pval S y). -Proof. by rewrite -fpV2E /= !psub_pval. Qed. - -Lemma valid_sepUn (x y : U) : - valid (x \+ y) -> - valid (pval S x \+ pval S y) /\ subsep S (pval S x) (pval S y). -Proof. by rewrite valid_sepUnE=>/andP. Qed. - -Lemma valid_pvalUn (x y : U) : - valid (x \+ y) -> valid (pval S x \+ pval S y). -Proof. by case/valid_sepUn. Qed. - -Lemma valid_sepUnS (x y : U) : valid (x \+ y) -> subsep S (pval S x) (pval S y). -Proof. by rewrite valid_sepUnE=>/andP []. Qed. - -(* binary versions of valid_pvalE, valid_pvalS only hold if *) -(* subsep operates on arguments independently *) -Lemma valid_pvalUnE (x y : U) : - subsep S (pval S x) (pval S y) = - subsep S (pval S x) Unit && subsep S (pval S y) Unit -> - valid (pval S x \+ pval S y) = valid (x \+ y). -Proof. -rewrite valid_sepUnE=>->; case W : (valid _)=>//=. -by rewrite !valid_pvalS ?(validL W, validR W). -Qed. - -Lemma valid_pvalUnS (x y : U) : - subsep S (pval S x) (pval S y) = - subsep S (pval S x) Unit && subsep S (pval S y) Unit -> - valid (pval S x \+ pval S y) -> subsep S (pval S x) (pval S y). -Proof. by move/valid_pvalUnE=>-> /valid_sepUnS. Qed. - - -(* ternary lemmas (occasionally useful) *) - -Lemma valid_sep3E (x y z : U) : - valid (x \+ y \+ z) = - [&& valid (pval S x \+ pval S y \+ pval S z), - subsep S (pval S x) (pval S y) & - subsep S (pval S x \+ pval S y) (pval S z)]. -Proof. -apply/idP/idP=>[/[dup] W /validL D|/and3P [W S1 S2]]. -- by rewrite -pfjoin // andbCA -valid_sepUnE W andbT valid_sepUnS. -by rewrite valid_sepUnE pfjoin ?W // valid_sepUnE (validL W) S1. -Qed. - -Lemma valid_sep3 (x y z : U) : - valid (x \+ y \+ z) -> - [/\ valid (pval S x \+ pval S y \+ pval S z), - subsep S (pval S x) (pval S y) & - subsep S (pval S x \+ pval S y) (pval S z)]. -Proof. by rewrite valid_sep3E=>/and3P. Qed. - - -(* lemmas for zig-zag interaction of psub and pval *) - -Lemma valid_psubUnX (x : V) (y : U) : - valid (psub S x \+ y) = valid (x \+ pval S y) && subsep S x (pval S y). -Proof. by rewrite -{1}(psub_pval S y) fpV2E. Qed. - -Lemma valid_psubXUn (x : U) (y : V) : - valid (x \+ psub S y) = valid (pval S x \+ y) && subsep S (pval S x) y. -Proof. by rewrite -{1}(psub_pval S x) fpV2E. Qed. - -Lemma psubUnX (x : V) (y : U) : - valid (psub S x \+ y) -> psub S x \+ y = psub S (x \+ pval S y). -Proof. by rewrite valid_psubUnX=>/andP [W H]; rewrite pfjoin //= psub_pval. Qed. - -Lemma psubXUn (x : U) (y : V) : - valid (x \+ psub S y) -> x \+ psub S y = psub S (pval S x \+ y). -Proof. by rewrite valid_psubXUn=>/andP [W H]; rewrite pfjoin //= psub_pval. Qed. - -Lemma pvalXUn (x : V) (y : U) : - valid (psub S x \+ y) -> x \+ pval S y = pval S (psub S x \+ y). -Proof. -by move/[dup]=>W /validL/fpVI [/= V1 V2]; rewrite pfjoin //= pval_psub. -Qed. - -Lemma pvalUnX (x : U) (y : V) : - valid (x \+ psub S y) -> pval S x \+ y = pval S (x \+ psub S y). -Proof. by rewrite joinC=>/pvalXUn <-; rewrite joinC. Qed. - - -(* injectivity *) - -Lemma pval_inj : injective (pval S). -Proof. by move=>x y E; rewrite -(psub_pval S x) E psub_pval. Qed. - -Lemma psub_inj (x y : V) : valid (psub S x) -> psub S x = psub S y -> x = y. -Proof. -move/[swap]=>E /[dup]; rewrite {2}E. -case/fpVI=>/= W1 H1 /fpVI [/= W2 H2]. -by rewrite -(pval_psub S _ W1 H1) -(pval_psub S _ W2 H2) E. -Qed. - -End DerivedLemmas. - -Prenex Implicits valid_sepE valid_pvalE valid_pvalEP valid_pvalS valid_psubS -valid_sepUnE valid_pvalUnE valid_pvalUnS valid_sep3E valid_psubUnX valid_psubXUn -psubUnX psubXUn pvalXUn pvalUnX pval_inj psub_inj. - - -(* properties of V propagate to U *) -Section SubPCMPropagation. -Variable U : pcm. - -(* if V is tpcm, psub undef isn't valid *) -Lemma psub_undef (V : tpcm) (S : subpcm_struct U V) : ~~ valid (psub S undef). -Proof. by rewrite fpVE /= valid_undef. Qed. - -(* if V is cancellative, so is U *) -Lemma subpcm_is_cpcm (V : cpcm) (S : subpcm_struct U V) : cpcm_axiom U. -Proof. -move=>x1 x2 x W E; move: (W) (W). -rewrite {1}E !(valid_sepUnE S)=>/andP [W2 D2] /andP [W1 D1]. -move: E; rewrite -(psub_pval S x1) -(psub_pval S x2) -(psub_pval S x). -rewrite -pfjoin // -[R in _ = R]pfjoin //; move/psub_inj. -by rewrite fpVE W1 sepU0 // => /(_ (erefl _)) /(joinKx W1) ->. -Qed. - -(* no canonical projection to latch onto, so no generic inheritance *) -(* but subpcm_is_cpcm can be used for any individual type U *) -(* -HB.instance Definition _ (V : cpcm) (S : subpcm_struct U V) := - isCPCM.Build (PCM.pack_ (PCM.class U)) (subpcm_is_cpcm S). -*) -End SubPCMPropagation. - - -(* identity subpcm structure *) -Lemma id_sub_is_subpcm_struct {U : pcm} : subpcm_struct_axiom (@id_sub U). -Proof. by []. Qed. -HB.instance Definition _ U := - isSubPCM_struct.Build U U id_sub id_sub_is_subpcm_struct. -Lemma subsep_id (U : pcm) : subsep (@id_sub U) = relT. Proof. by []. Qed. - -(* Composition of subpcm structures *) -Section SubPCMStructCompose. -Variables U V W : pcm. -Variables (X : subpcm_struct U V) (Y : subpcm_struct V W). - -Lemma comp_is_subpcm_struct : subpcm_struct_axiom (comp_sub X Y). -Proof. -split=>x/=; first by rewrite !psub_pval. -rewrite /sepx/= pfunit => D /andP [/= H1 H2]. -by rewrite !pval_psub // fpVE /= D H1. -Qed. - -HB.instance Definition _ := - isSubPCM_struct.Build U W (comp_sub X Y) comp_is_subpcm_struct. - -Lemma subsep_comp : - subsep (comp_sub X Y) = - fun y1 y2 => subsep Y y1 y2 && subsep X (psub Y y1) (psub Y y2). -Proof. by []. Qed. - -End SubPCMStructCompose. - -(***********) -(* SubTPCM *) -(***********) - -(* subtpcm structure is a subpcm struct on tpcms *) -(* where pval is further required to preserve undef *) -(* (i.e., pval is tpcm morphism) *) -(* it then follows that psub preserves undef as well *) - -Definition subtpcm_struct_axiom (U V : tpcm) (S : subpcm_struct U V) := - pval S undef = undef. - -HB.mixin Record isSubTPCM_struct (U V : tpcm) - S of @SubPCM_struct U V S := { - subtpcm_struct_subproof : subtpcm_struct_axiom S}. - -#[short(type=subtpcm_struct)] -HB.structure Definition SubTPCM_struct (U V : tpcm) := - {S of isSubTPCM_struct U V S & @SubPCM_struct U V S}. - -Arguments subtpcm_struct_subproof {U V}. - -(* psub preserves undef *) -Lemma psub_is_tpcm_morph (U V : tpcm) (S : subtpcm_struct U V) : - tpcm_morph_axiom (psub S). -Proof. -by rewrite /tpcm_morph_axiom -(subtpcm_struct_subproof S) psub_pval. -Qed. - -(* declare pval/psub to be tpcm morphisms *) -HB.instance Definition _ (U V : tpcm) (S : subtpcm_struct U V) := - isTPCM_morphism.Build U V (pval S) (subtpcm_struct_subproof S). -HB.instance Definition _ (U V : tpcm) (S : subtpcm_struct U V) := - isTPCM_morphism.Build V U (psub S) (psub_is_tpcm_morph S). - -(* identity subtpcm structure *) -HB.instance Definition _ (U : tpcm) := - isSubTPCM_struct.Build U U id_sub (erefl undef). - -(* composition of subtpcm structures *) -Section SubTPCMStructCompose. -Variables U V W : tpcm. -Variables (X : subtpcm_struct U V) (Y : subtpcm_struct V W). -Lemma comp_is_subtpcm_struct : subtpcm_struct_axiom (comp_sub X Y). -Proof. by rewrite /subtpcm_struct_axiom/pval /= !pfundef. Qed. -HB.instance Definition _ := - isSubTPCM_struct.Build U W (comp_sub X Y) comp_is_subtpcm_struct. -End SubTPCMStructCompose. - - -(*****************************) -(*****************************) -(* SubPCM/TPCM constructions *) -(*****************************) -(*****************************) - - -(**********************************) -(* Modding out TPCM V by seprel D *) -(**********************************) - -(* requires TPCM, because morphisms need undef *) -(* to return when D not satisfied *) - -(* xsep is the type of the subpcm *) -(* xsub is the subpcm structure *) -Definition xsepP (V : tpcm) (D : seprel V) (x : V) := - valid x && D x Unit \/ x = undef. -Inductive xsep (V : tpcm) (D : seprel V) := - ex_sep x of xsepP D x. - -(* makes defs opaque to avoid performance penalty *) - -Module Type XSepSig. -Parameter valx : forall (V : tpcm) (D : seprel V), xsep D -> V. -Parameter subx : forall (V : tpcm) (D : seprel V), V -> xsep D. -Definition xsub V D := SubStruct (@valx V D) (@subx V D). - -(* pcm operation *) -Parameter xsep_valid : forall (V : tpcm) (D : seprel V), xsep D -> bool. -Parameter xsep_unit : forall (V : tpcm) (D : seprel V), xsep D. -Parameter xsep_unitb : forall (V : tpcm) (D : seprel V), xsep D -> bool. -Parameter xsep_join' : forall (V : tpcm) (D : seprel V), - xsep D -> xsep D -> V. -Parameter xsep_join : forall (V : tpcm) (D : seprel V), - xsep D -> xsep D -> xsep D. -(* tpcm operations *) -Parameter xsep_undef : forall (V : tpcm) (D : seprel V), xsep D. -Parameter xsep_undefb : forall (V : tpcm) (D : seprel V), xsep D -> bool. - -Parameter valxE : forall (V : tpcm) (D : seprel V), - @valx V D = fun (x : xsep D) => let: ex_sep v _ := x in v. -Parameter subxE : forall (V : tpcm) (D : seprel V), - @subx V D = fun x => - if decP (valid x && D x Unit =P true) is left pf - then ex_sep (or_introl pf) - else ex_sep (or_intror (erefl undef)). -Parameter xsep_validE : forall (V : tpcm) (D : seprel V), - @xsep_valid V D = fun x => valid (valx x) && D (valx x) Unit. -Parameter xsep_unitP : forall (V : tpcm) (D : seprel V), - xsepP D (Unit : V). -Parameter xsep_unitE : forall (V : tpcm) (D : seprel V), - @xsep_unit V D = ex_sep (@xsep_unitP V D). -Parameter xsep_unitbE : forall (V : tpcm) (D : seprel V), - @xsep_unitb V D = fun x => unitb (valx x). -Parameter xsep_joinE' : forall (V : tpcm) (D : seprel V), - @xsep_join' V D = fun x y => - if valid (valx x \+ valx y) && D (valx x) (valx y) - then valx x \+ valx y else undef. -Parameter xsep_joinP : forall (V : tpcm) (D : seprel V) x y, - xsepP D (@xsep_join' V D x y). -Parameter xsep_joinE : forall (V : tpcm) (D : seprel V), - @xsep_join V D = fun x y => ex_sep (xsep_joinP x y). -Parameter xsep_undefP : forall (V : tpcm) (D : seprel V), - xsepP D undef. -Parameter xsep_undefE : forall (V : tpcm) (D : seprel V), - @xsep_undef V D = ex_sep (@xsep_undefP V D). -Parameter xsep_undefbE : forall (V : tpcm) (D : seprel V), - @xsep_undefb V D = fun x => undefb (valx x). -End XSepSig. - -Module XSepSubPCMDef : XSepSig. -Section XSepSubPCMDef. -Variables (V : tpcm) (D : seprel V). - -Definition valx (x : xsep D) := let: ex_sep v _ := x in v. -Definition subx (x : V) := - if decP (valid x && D x Unit =P true) is left pf - then ex_sep (or_introl pf) - else ex_sep (or_intror (erefl undef)). -Definition xsep_valid x := valid (valx x) && D (valx x) Unit. -Lemma xsep_unitP : xsepP D (Unit : V). -Proof. by rewrite /xsepP valid_unit sep00; left. Qed. -Definition xsep_unit := ex_sep xsep_unitP. -Definition xsep_unitb x := unitb (valx x). -Definition xsep_join' x y := - if valid (valx x \+ valx y) && D (valx x) (valx y) - then valx x \+ valx y else undef. -Lemma xsep_joinP x y : xsepP D (xsep_join' x y). -Proof. -rewrite /xsepP /xsep_join'; case: ifP; last by right. -by case/andP=>W /(sepU0 W) ->; rewrite W; left. -Qed. -Definition xsep_join x y := ex_sep (xsep_joinP x y). -Lemma xsep_undefP : xsepP D undef. Proof. by right. Qed. -Definition xsep_undef := ex_sep xsep_undefP. -Definition xsep_undefb (x : xsep D) := undefb (valx x). - -Definition valxE := erefl valx. -Definition subxE := erefl subx. -Definition xsub := SubStruct valx subx. -Definition xsep_validE := erefl xsep_valid. -Definition xsep_unitE := erefl xsep_unit. -Definition xsep_unitbE := erefl xsep_unitb. -Definition xsep_joinE' := erefl xsep_join'. -Definition xsep_joinE := erefl xsep_join. -Definition xsep_undefE := erefl xsep_undef. -Definition xsep_undefbE := erefl xsep_undefb. -End XSepSubPCMDef. -End XSepSubPCMDef. - -Export XSepSubPCMDef. - -Section XSepSubPCM. -Variables (V : tpcm) (D : seprel V). - -(* helper lemma *) -Lemma valx_inj (x y : xsep D) : - valx x = valx y -> - x = y. -Proof. -case: x y=>x Hx [y Hy]; rewrite !valxE => E. -by subst y; rewrite (pf_irr Hx). -Qed. - -(* unary and binary orthogonality relations *) -Notation orth1 x := (valid x && D x Unit). -Notation orth2 x y := (valid (x \+ y) && D x y). - -Notation xsep_valid := (@xsep_valid V D). -Notation xsep_join := (@xsep_join V D). -Notation xsep_unit := (@xsep_unit V D). -Notation xsep_unitb := (@xsep_unitb V D). -Notation xsep_undef := (@xsep_undef V D). -Notation xsep_undefb := (@xsep_undefb V D). - -(* xsep is pcm *) -Lemma xsep_is_pcm : pcm_axiom xsep_valid xsep_join xsep_unit xsep_unitb. -Proof. -have joinC : commutative xsep_join. -- case=>x Hx [y Hy]; apply: valx_inj; rewrite valxE xsep_joinE xsep_joinE'. - by rewrite joinC; case W: (valid _)=>//=; rewrite -sepC. -split=>[//||[x Hx]|x y||x]. -- suff joinAC : right_commutative xsep_join. - - by move=>a b c; rewrite !(joinC a) joinAC. - case=>a Ha [b Hb][c Hc]; apply: valx_inj; rewrite valxE. - rewrite xsep_joinE; do ![rewrite {1}xsep_joinE' !valxE /=]. - case Sab: (orth2 a b); case Sac: (orth2 a c); rewrite ?tpcmE //=; last first. - - case/andP: Sac=>_ Sac; case: andP=>//; case=>W Sacb. - by rewrite (sepAxx W Sac Sacb) andbT (validLE3 W) in Sab. - - case/andP: Sab=>_ Sab; case: andP=>//; case=>W Sabc. - by rewrite (sepAxx W Sab Sabc) andbT (validLE3 W) in Sac. - case/andP: Sab=>_ Sab; case/andP: Sac=>_ Sac. - case Sabc: (orth2 (a \+ b) c). - - case/andP: Sabc=>W Sabc; rewrite sepC (joinAC a c b) W //=. - by rewrite (sepAxx W Sab Sabc). - case Sacb: (orth2 (a \+ c) b)=>//. - case/andP: Sacb=>W Sacb; rewrite sepC (joinAC a b c) W // in Sabc. - by rewrite (sepAxx W Sac Sacb) in Sabc. -- apply: valx_inj; rewrite !valxE /=. - rewrite xsep_joinE xsep_unitE xsep_joinE' !{1}valxE. - rewrite unitL; case: Hx=>[|->]; last by rewrite tpcmE. - by case/andP=>W E; rewrite sepC ?unitL // W E. -- rewrite xsep_validE xsep_joinE valxE xsep_joinE' !{1}valxE /=. - case: ifP; last by rewrite tpcmE. - by case/andP=>W E; rewrite !(validE2 W) (sepx0 W E). -- by rewrite xsep_validE xsep_unitE valxE valid_unit sep00. -rewrite xsep_unitbE xsep_unitE; case: x=>x H /=; rewrite valxE. -case: unitbP=>X; constructor; last by case=>/X. -by rewrite X in H *; rewrite (pf_irr H (@xsep_unitP V D)). -Qed. - -HB.instance Definition _ : isPCM (xsep D) := - isPCM.Build (xsep D) xsep_is_pcm. - -(* xsep is tpcm *) -Lemma xsep_is_tpcm : tpcm_axiom xsep_undef xsep_undefb. -Proof. -split=>[/= x||/= x]. -- rewrite xsep_undefbE xsep_undefE valxE; case: x=>x H /=. - case: undefbP=>X; constructor; last by case=>/X. - by rewrite X in H *; rewrite (pf_irr H (xsep_undefP D)). -- by rewrite pcmE /= xsep_validE xsep_undefE valxE tpcmE. -apply: valx_inj; rewrite xsep_undefE !valxE. -by rewrite /join/= xsep_joinE xsep_joinE' valxE /= !tpcmE. -Qed. - -HB.instance Definition _ : isTPCM (xsep D) := - isTPCM.Build (xsep D) xsep_is_tpcm. - -(* xsep is normal tpcm *) -Lemma xsep_is_normal : normal_tpcm_axiom (xsep D). -Proof. -case=>x [] H; [left|right]. -- by rewrite /valid/= xsep_validE valxE. -by apply/valx_inj; rewrite !valxE /undef /= xsep_undefE. -Qed. - -HB.instance Definition _ : isNormal_TPCM (xsep D) := - isNormal_TPCM.Build (xsep D) xsep_is_normal. - -(* Next, morphisms *) - -(* valx is morphism *) -Lemma valx_is_morph : pcm_morph_axiom relT (@valx V D). -Proof. -split=>[|x y]; first by rewrite valxE /Unit /= xsep_unitE. -rewrite {1}/valid /= xsep_validE !valxE /=. -rewrite pcm_joinE /= xsep_joinE xsep_joinE' !valxE. -by case: ifP; rewrite ?tpcmE //; case/andP. -Qed. - -HB.instance Definition _ := - isPCM_morphism.Build (xsep D) V (@valx V D) valx_is_morph. - -(* valx is full morphism *) -Lemma valx_is_full_morph : full_pcm_morph_axiom (@valx V D). -Proof. by []. Qed. - -HB.instance Definition _ := - isFull_PCM_morphism.Build (xsep D) V (@valx V D) valx_is_full_morph. - -(* valx is normal morphism *) -Lemma valx_is_norm_morph : norm_pcm_morph_axiom (@valx V D). -Proof. -move=>x; rewrite {2}/valid/= xsep_validE. -rewrite !{1}valxE /= => W; split=>//. -by case: x W=>x /= [/andP [->->]|] // ->; rewrite tpcmE. -Qed. - -HB.instance Definition _ := - isNorm_PCM_morphism.Build (xsep D) V (@valx V D) valx_is_norm_morph. - -(* NOTE: valx is *not* binormal morphism as it doesn't preserve separation. *) -(* The subpcm gives new notion of separation that is more stringent *) -(* compared to super-pcm. If valx were binormal, that would imply that *) -(* the new notion isn't more stringent, and thus defeat the purpose *) -(* of the construction. *) -Lemma valx_is_binorm_morph : binorm_pcm_morph_axiom (@valx V D). -Proof. -case=>x Hx [y Hy] /=; rewrite !{1}valxE => W. -rewrite /valid/=/sepx/= xsep_validE /= !{1}valxE /=. -rewrite pcm_joinE /= xsep_joinE /= xsep_joinE' !{1}valxE. -case H : (D x y); first by rewrite !W (sepU0 W H). -rewrite W /=. -Abort. - -(* subx is morphism *) -Lemma subx_is_morph : pcm_morph_axiom D (@subx V D). -Proof. -rewrite subxE; split=>[|x y W E]. -- apply: valx_inj; rewrite !valxE /Unit /= xsep_unitE; case: eqP=>//=. - by rewrite valid_unit /= sep00. -case: eqP=>Hx /=; last by rewrite (sep0E W E) (validE2 W) in Hx. -case: eqP=>Hy /=; last by rewrite (sep0E W E) (validE2 W) in Hy. -case: eqP=>H /=; last by rewrite W (sepU0 W E) in H. -rewrite /valid/= xsep_validE pcm_joinE valxE /= xsep_joinE /=. -do ![rewrite {1}xsep_joinE' valxE]. -rewrite {1 2}W {1 2}E {1}W {1}(sepU0 W E) /=. -split=>//; apply: valx_inj; rewrite valxE /=. -by rewrite xsep_joinE' valxE W E. -Qed. - -HB.instance Definition _ := - isPCM_morphism.Build V (xsep D) (subx D) subx_is_morph. - -(* subx is binormal morphism *) -Lemma subx_is_binorm_morph : binorm_pcm_morph_axiom (subx D). -Proof. -move=>x y /=. -rewrite /sepx/= subxE {1}/valid/= xsep_validE valxE !pcm_joinE /= xsep_joinE. -case: eqP=>Hx; case: eqP=>Hy; rewrite xsep_joinE' valxE; -case H: (valid (x \+ y) && D x y) Hx Hy; -by rewrite /= ?tpcmE //=; case/andP: H. -Qed. - -HB.instance Definition _ := - isBinorm_PCM_morphism.Build V (xsep D) (subx D) subx_is_binorm_morph. - -(* Next, substructures *) - -(* xsub is subpcm *) -Definition xsub := SubStruct (@valx V D) (subx D). - -Lemma xsub_is_subpcm_struct : subpcm_struct_axiom xsub. -Proof. -split=>[x|x] //=; last first. -- rewrite /sepx/= subxE valxE => W H. - by case: eqP=>//=; rewrite W H. -apply: valx_inj; rewrite valxE subxE /=. -by case: eqP; case: x=>// x []. -Qed. - -HB.instance Definition _ := - isSubPCM_struct.Build (xsep D) V xsub xsub_is_subpcm_struct. - -Lemma xsub_is_subtpcm_struct : subtpcm_struct_axiom xsub. -Proof. -rewrite /subtpcm_struct_axiom. -by rewrite /pval/= /undef/= xsep_undefE valxE. -Qed. - -HB.instance Definition _ := - isSubTPCM_struct.Build (xsep D) V xsub xsub_is_subtpcm_struct. -End XSepSubPCM. - -Lemma psub_undefN (V : tpcm) (D : seprel V) (x : V) : - ~~ D x Unit -> - psub (xsub D) x = undef. -Proof. -move=>X; apply: valx_inj. -rewrite /undef/= xsep_undefE valxE /psub/= subxE /=. -by case: decP=>//; rewrite (negbTE X) andbF. -Qed. - -(*****************************************) -(* Normalize = mod out trivially by relT *) -(*****************************************) - -(* removes non-valid elements != undef *) - -Definition normalize (U : tpcm) := xsep (@relT U). -Definition norm_sub {U : tpcm} := xsub (@relT U). - -(* redeclare structures for normalize, to save on unfolding *) -HB.instance Definition _ (U : tpcm) := - Normal_TPCM.on (normalize U). -HB.instance Definition _ (U : tpcm) := - SubTPCM_struct.on (@norm_sub U). -(* psub also becomes full morphism *) -HB.instance Definition _ (U : tpcm) := - isFull_PCM_morphism.Build U (normalize U) (psub norm_sub) - (fun _ _ => erefl _). - -(***************************) -(* Normal product of TPCMs *) -(***************************) - -(* product which is immediately normalized *) -(* to remove spurious invalid elements *) -(* that arose out of pairing invalids of one TPCM *) -(* with valids of the other. *) - -Section NormalProd. -Variables U V : tpcm. - -Definition nprod := normalize (prod U V). -Definition nprod_sub := @norm_sub (prod U V). - -Definition nfst : nprod -> U := fst \o pval nprod_sub. -Definition nsnd : nprod -> V := snd \o pval nprod_sub. -Definition npair : U * V -> nprod := psub nprod_sub. - -End NormalProd. - -Arguments nprod_sub {U V}. -Prenex Implicits nfst nsnd npair nprod_sub. - -(* redeclare again *) -HB.instance Definition _ (U V : tpcm) := - Normal_TPCM.on (nprod U V). -HB.instance Definition _ (U V : tpcm) := - SubTPCM_struct.on (@nprod_sub U V). - -(* redeclare morphisms *) -HB.instance Definition _ (U V : tpcm) := - Full_Binorm_TPCM_morphism.on (@npair U V). - -(* nfst and nsnd are full by inheritance *) -HB.instance Definition _ (U V : tpcm) := - Full_TPCM_morphism.on (@nfst U V). -HB.instance Definition _ (U V : tpcm) := - Full_TPCM_morphism.on (@snd U V). - -(* but now they are also normal *) -(* though not binormal (as they shouldn't) *) -Lemma nfst_is_norm_pcm_morphism U V : norm_pcm_morph_axiom (@nfst U V). -Proof. -move=>/= x W; split=>//; case: (normalP x) W=>// ->. -by rewrite pfundef valid_undef. -Qed. - -HB.instance Definition _ (U V : tpcm) := - isNorm_PCM_morphism.Build (nprod U V) U (@nfst U V) - (@nfst_is_norm_pcm_morphism U V). - -Lemma nsnd_is_norm_pcm_morphism U V : norm_pcm_morph_axiom (@nsnd U V). -Proof. -move=>/= x W; split=>//; case: (normalP x) W=>// ->. -by rewrite pfundef valid_undef. -Qed. - -HB.instance Definition _ (U V : tpcm) := - isNorm_PCM_morphism.Build (nprod U V) V (@nsnd U V) - (@nsnd_is_norm_pcm_morphism U V). - -(* lemmas for normal products *) - -Lemma nprod_eta (U V : tpcm) (x : nprod U V) : - x = npair (nfst x, nsnd x). -Proof. by rewrite -prod_eta /npair psub_pval. Qed. - -Lemma nfst_pair (U V : tpcm) (x : U * V) : - valid x -> nfst (npair x) = x.1. -Proof. by move=>W; rewrite /nfst /= pval_psub. Qed. - -Lemma nsnd_pair (U V : tpcm) (x : U * V) : - valid x -> nsnd (npair x) = x.2. -Proof. by move=>W; rewrite /nsnd /= pval_psub. Qed. - -Lemma nprodV (U V : tpcm) (x : nprod U V) : - valid x = valid (nfst x) && valid (nsnd x). -Proof. by rewrite {1}[x]nprod_eta pfVE. Qed. - -Lemma nprodUnV (U V : tpcm) (x y : nprod U V) : - valid (x \+ y) = valid (nfst x \+ nfst y) && - valid (nsnd x \+ @nsnd U V y). -Proof. by rewrite {1}[x]nprod_eta {1}[y]nprod_eta pfV2E. Qed. - - -(***********************) -(* Relativized seprels *) -(***********************) - -(* Relation S is seprel-relative-to-morphism-f *) -(* if it becomes seprel once composed with f and paired with sep f. *) -(* Useful when sequencing subpcm constructions *) - -Definition seprel_on_axiom U V (f : pcm_morph U V) (S : rel V) := - [/\ S Unit Unit, - forall x y, valid (x \+ y) -> sep f x y -> - S (f x) (f y) = S (f y) (f x), - forall x y, valid (x \+ y) -> sep f x y -> - S (f x) (f y) -> S (f x) Unit & - forall x y z, valid (x \+ y \+ z) -> - sep f x y -> sep f (x \+ y) z -> - S (f x) (f y) -> S (f x \+ f y) (f z) -> - S (f y) (f z) && S (f x) (f y \+ f z)]. - -(* S is seprel-on-f iff *) -(* relI (sep f) (rel_app S f) is seprel on U *) -Lemma seprel_on_app U V (f : pcm_morph U V) S : - seprel_on_axiom f S <-> - seprel_axiom (relI (sep f) (rel_app S f)). -Proof. -split=>[[H1 H2 H3 H4]|[/andP [H1 H2 H3 H4 H5]]]. -- split=>[|x y W|x y W|x y z W] /=. - - by rewrite pfunit sep00 H1. - - rewrite (sepC _ W) /=. - by case X: (sep f y x)=>//; rewrite H2 // sepC. - - by case/andP=>X1 X2; rewrite pfunit (sep0E W) // (H3 _ _ W). - case/andP=>X1 X2 /andP [X3 X4]. - rewrite pfjoin ?(sepAxx W X1 X3) ?(validLE3 W) //=. - by rewrite H4 // -pfjoin // (validL W). -rewrite /= !pfunit in H2 H4; split=>[//|x y|x y|x y z] W. -- rewrite (sepC (sep f) W) => X. - by move: (H3 x y W); rewrite /= (sepC (sep f) W) /= X. -- by move=>X; move: (H4 x y W); rewrite X (sep0E W). -move=>X1 X2 X3 X4; case: (sepAx W X1 X2)=>/= X5 X6. -move: (H5 x y z W); rewrite /= X1 X2 X5 X6 !pfjoin ?(validLE3 W) //=. -by apply. -Qed. - -HB.mixin Record isSeprelOn U V (f : pcm_morph U V) (S : rel V) := { - seprel_on_subproof : seprel_on_axiom f S}. - -#[short(type=seprel_on)] -HB.structure Definition SeprelOn U V f := {S of @isSeprelOn U V f S}. - -Section Repack. -Variables (U V : pcm) (f : pcm_morph U V) (S : seprel_on f). - -Lemma sepon00 : S Unit Unit. -Proof. by case: (@seprel_on_subproof U V f S). Qed. - -Lemma seponC x y : - valid (x \+ y) -> - sep f x y -> - S (f x) (f y) = S (f y) (f x). -Proof. by case: (@seprel_on_subproof U V f S)=>_ H _ _; apply: H. Qed. - -Lemma seponx0 x y : - valid (x \+ y) -> - sep f x y -> - S (f x) (f y) -> - S (f x) Unit. -Proof. by case: (@seprel_on_subproof U V f S)=>_ _ H _; apply: H. Qed. - -Lemma seponAx x y z : - valid (x \+ y \+ z) -> - sep f x y -> - sep f (x \+ y) z -> - S (f x) (f y) -> - S (f x \+ f y) (f z) -> - S (f y) (f z) * S (f x) (f y \+ f z). -Proof. -case: (@seprel_on_subproof U V f S)=>_ _ _ H W R1 R2 R3 R4. -by rewrite !(andX (H _ _ _ W R1 R2 R3 R4)). -Qed. - -(* derived lemmas *) - -Lemma sepon0x x y : - valid (x \+ y) -> - sep f x y -> - S (f x) (f y) -> - S Unit (f y). -Proof. -move=>W sf; rewrite seponC // => Sf. -rewrite sepC //= in sf; rewrite joinC in W. -rewrite -(pfunit f) -(@seponC y) ?pfunit ?unitR ?(validE2 W) //. -- by apply: seponx0 W sf Sf. -by apply: sepx0 W sf. -Qed. - -Lemma sepon0E x y : - valid (x \+ y) -> - sep f x y -> - S (f x) (f y) -> - S (f x) Unit * S (f y) Unit. -Proof. -move=>W sf Sf; rewrite (seponx0 W sf Sf). -rewrite -(pfunit f) seponC ?pfunit ?unitR ?(validE2 W) //. -- by rewrite (sepon0x W sf Sf). -by rewrite (sep0E W sf). -Qed. - -Lemma seponAx23_helper x y z : - valid (x \+ y \+ z) -> - sep f x y -> - sep f (x \+ y) z -> - S (f x) (f y) -> - S (f x \+ f y) (f z) -> - ((S (f x) (f z) * S (f z) (f x)) * (S (f y) (f z) * S (f z) (f y))) * - ((S (f x) (f y \+ f z) * S (f x) (f z \+ f y)) * - (S (f y) (f x \+ f z) * S (f y) (f z \+ f x))). -Proof. -move=>W s1 s2 S1 S2. -rewrite !(@seponC z) ?(validLE3 W) ?(sepAxx W s1 s2) ?(joinC (f z)) //. -rewrite !(seponAx W s1 s2 S1 S2). -rewrite seponC ?(validLE3 W) // in S1. -rewrite sepC ?(validLE3 W) //= in s1. -rewrite (joinC x) in W s2; rewrite (joinC (f x)) in S2. -by rewrite !(seponAx W s1 s2 S1 S2). -Qed. - -Lemma seponxA x y z : - valid (x \+ (y \+ z)) -> - sep f y z -> - sep f x (y \+ z) -> - S (f y) (f z) -> - S (f x) (f y \+ f z) -> - S (f x \+ f y) (f z) * S (f x) (f y). -Proof. -move=>W s1 s2 S1 S2; have /= R := (sepxxA W s1 s2, validRE3 W). -rewrite -pfjoin ?R // seponC ?R // pfjoin ?R // in S2. -rewrite sepC ?R //= in s2. -rewrite -pfjoin ?R // seponC ?R // pfjoin ?R //. -by rewrite !(seponAx23_helper _ s1 s2 S1 S2) ?R. -Qed. - -Lemma seponAxx x y z : - valid (x \+ y \+ z) -> - sep f x y -> - sep f (x \+ y) z -> - S (f x) (f y) -> - S (f x \+ f y) (f z) -> - (((S (f x) (f y \+ f z) * S (f x) (f z \+ f y)) * - (S (f y) (f x \+ f z) * S (f y) (f z \+ f x))) * - ((S (f z) (f x \+ f y) * S (f z) (f y \+ f x)) * - (S (f y \+ f z) (f x) * S (f z \+ f y) (f x)))) * - (((S (f x \+ f z) (f y) * S (f z \+ f x) (f y)) * - (S (f x \+ f y) (f z) * S (f y \+ f x) (f z))) * - (((S (f x) (f y) * S (f y) (f x)) * - (S (f x) (f z) * S (f z) (f x))))) * - (S (f y) (f z) * S (f z) (f y)). -Proof. -move=>W s1 s2 S1 S2; have /= R := (sepAxx W s1 s2, validLE3 W). -move: (seponAx23_helper W s1 s2 S1 S2)=>E; rewrite S1 S2 !E. -rewrite -!pfjoin ?R // -!(@seponC x) ?R // !pfjoin ?R // S1 !E. -rewrite -!pfjoin ?R // !(@seponC z) ?R // !pfjoin ?R //. -rewrite (joinC (f y)) S2. -rewrite -!pfjoin ?R // -!(@seponC y) ?R // !pfjoin ?R //. -by rewrite (joinC (f z)) E. -Qed. - -Lemma seponxxA x y z : - valid (x \+ (y \+ z)) -> - sep f y z -> - sep f x (y \+ z) -> - S (f y) (f z) -> - S (f x) (f y \+ f z) -> - (((S (f x) (f y \+ f z) * S (f x) (f z \+ f y)) * - (S (f y) (f x \+ f z) * S (f y) (f z \+ f x))) * - ((S (f z) (f x \+ f y) * S (f z) (f y \+ f x)) * - (S (f y \+ f z) (f x) * S (f z \+ f y) (f x)))) * - (((S (f x \+ f z) (f y) * S (f z \+ f x) (f y)) * - (S (f x \+ f y) (f z) * S (f y \+ f x) (f z))) * - (((S (f x) (f y) * S (f y) (f x)) * - (S (f x) (f z) * S (f z) (f x))))) * - (S (f y) (f z) * S (f z) (f y)). -Proof. -move=>W s1 s2 S1 S2; have /= R := (sepxxA W s1 s2, validRE3 W). -rewrite -pfjoin ?R // seponC // pfjoin ?R // in S2. -rewrite sepC //= in s2; rewrite joinC in W. -rewrite !(seponAx23_helper W s1 s2 S1 S2). -by rewrite !(seponAxx W s1 s2 S1 S2). -Qed. - -Lemma seponU0 x y : - valid (x \+ y) -> - sep f x y -> - S (f x) (f y) -> - S (f x \+ f y) Unit. -Proof. -move=>W s1 S1; rewrite -(pfunit f). -rewrite seponAxx ?pfunit ?unitL //. -- by rewrite sepC ?unitL ?(validL W) ?(sepx0 W s1). -rewrite seponC // in S1; rewrite sepC //= in s1; rewrite joinC in W. -by apply: (@sepon0x y). -Qed. - -Lemma sepon0U x y : - valid (x \+ y) -> - sep f x y -> - S (f x) (f y) -> - S Unit (f x \+ f y). -Proof. -move=>W s1 S1. -rewrite -(pfunit f) -pfjoin // seponC ?unitL ?(sep0U W s1) //. -by rewrite pfjoin ?pfunit ?seponU0. -Qed. - -End Repack. - - -(* if f is full tpcm morphism *) -(* S is seprel-on-f iff (S o f) is seprel on U *) -Lemma seprel_on_full U V (f : full_pcm_morph U V) S : - seprel_on_axiom f S <-> - seprel_axiom (rel_app S f). -Proof. -rewrite seprel_on_app; apply/sepXEQ=>x y W. -by rewrite /relI /= pfSE. -Qed. - -Lemma seprel_on_id (U : pcm) (S : rel U) : - seprel_on_axiom idfun S <-> - seprel_axiom S. -Proof. by rewrite seprel_on_full; apply/sepXEQ. Qed. - -(* S is seprel-on-f iff (S o f) is seprel on subpcm U/(sep f) *) -Lemma seprel_on_pval (U : tpcm) V (f : pcm_morph U V) S : - seprel_on_axiom f S <-> - seprel_axiom (rel_app S (f \o pval (xsub (sep f)))). -Proof. -set xf := (xsub (sep f)). -split; case=>/= H1 H2 H3 H4. -- split=>[|x y W|x y W|x y z W /= X1 X2] /=. - - by rewrite !pfunit. - - by case/(valid_sepUn xf): W=>W /(H2 _ _ W). - - by rewrite !pfunit; case/(valid_sepUn xf): W=>W /(H3 _ _ W). - case: (valid_sepUn xf W)=>Wxyz Sxyz. - case: (valid_sepUn xf (validL W))=>Wxy Sxy. - case: (valid_sepUn xf (validAR W))=>Wyz Syz. - rewrite !pfjoin ?(validLE3 W) //=. - rewrite H4 ?pfV3 // -!pfjoin ?(validLE3 W) //=. - by rewrite pfjoin ?(validLE3 W). -rewrite pfunit in H1 H3. -split=>[|x y W X|x y W X|x y z W X1 X2]. -- by rewrite pfunit in H1. -- rewrite -(pval_psub xf x) ?(sep0E W X,validL W) //. - rewrite -(pval_psub xf y) ?(sep0E W X,validR W) //. - by apply/H2/pfV2. -- rewrite -(pval_psub xf x) ?(sep0E W X,validL W) //. - rewrite -(pval_psub xf y) ?(sep0E W X,validR W) //. - by move/H3; rewrite pfunit; apply; apply: (pfV2 (psub xf)). -rewrite -!pfjoin ?(sepAxx W X1 X2,validLE3 W) //. -rewrite -(pval_psub xf x) ?(sep0E _ X1,validLE3 W) //. -rewrite -(pval_psub xf y) ?(sep0E _ X1,validLE3 W) //. -rewrite -(pval_psub xf z) ?(sep0E _ X2,validLE3 W) //. -rewrite -!(pfjoin (pval xf)) ?(pfV2,sepAxx W X1 X2,validLE3 W) //. -by apply/H4/pfV3. -Qed. - -(* when the morphism is restriction of idfun with seprel K *) -(* then S is seprel alongside K *) -Lemma seprel_with_on (U : tpcm) (K : seprel U) (S : rel U) : - seprel_on_axiom (res idfun K) S <-> - seprel_axiom (relI K S). -Proof. -rewrite seprel_on_app; apply/sepXEQ=>x y V. -rewrite /sepx/= /res/= andbT. -by case X: (K x y)=>//=; rewrite !(sep0E _ X). -Qed. - -Notation "'seprel_with' K" := (seprel_on (res idfun K)) - (at level 1, format "'seprel_with' K"). - -(* Every seprel is trivially seprel_on idfun morphism *) -Section SeprelOnIdCoercion. -Variables (U : pcm) (S : seprel U). -(* dummy name to enable inheritance *) -Definition eta_rel : rel U := fun x y => S x y. -Lemma seprel_is_seprelonid : seprel_on_axiom idfun eta_rel. -Proof. by apply/seprel_on_id/seprel_subproof. Qed. -HB.instance Definition _ := - isSeprelOn.Build U U idfun eta_rel seprel_is_seprelonid. -Definition seprel_on_idfun : seprel_on idfun := eta_rel. -End SeprelOnIdCoercion. - -Coercion seprel_on_idfun : seprel >-> seprel_on. - -(* naming the completion seprel for specific case *) -(* of seprel_on relations where the morphism is full *) -Section SeprelOnFull. -Variables (U V : pcm) (f : full_pcm_morph U V) (S : seprel_on f). -Definition onsep := rel_app S f. -Lemma onsep_is_seprel : seprel_axiom onsep. -Proof. by apply/seprel_on_full/seprel_on_subproof. Qed. -HB.instance Definition _ := - isSeprel.Build U onsep onsep_is_seprel. -End SeprelOnFull. - -(* trivially true seprel is seprel_on for any morphism *) - -Lemma relT_is_on U V (f : pcm_morph U V) : seprel_on_axiom f relT. -Proof. by []. Qed. - -HB.instance Definition _ U V (f : pcm_morph U V) := - isSeprelOn.Build U V f relT (relT_is_on f). - -(* conjunction of seprel-on's is seprel-on *) -Section SepOnI. -Variables (U V : pcm) (f : pcm_morph U V) (X Y : seprel_on f). - -Lemma relI_is_seprelon : seprel_on_axiom f (relI X Y). -Proof. -split=>[|x y W s|x y W s|x y z W s1 s2] /=. -- by rewrite !sepon00. -- by rewrite !(seponC _ W s). -- by case/andP; do ![move/(seponx0 W s) ->]. -case/andP=>X1 Y1 /andP [X2 Y2]. -by rewrite !(seponAxx W s1 s2 X1 X2) !(seponAxx W s1 s2 Y1 Y2). -Qed. - -HB.instance Definition _ := - isSeprelOn.Build U V f (relI X Y) relI_is_seprelon. -End SepOnI. - -(* 3-way conjunction *) -Section SepOn3I. -Variables (U V : pcm) (f : pcm_morph U V) (X Y Z : seprel_on f). - -Lemma rel3I_is_seprelon : seprel_on_axiom f (rel3I X Y Z). -Proof. -split=>[|x y W s|x y W s|x y z W s1 s2] /=. -- by rewrite !sepon00. -- by rewrite !(seponC _ W s). -- by case/and3P; do ![move/(seponx0 W s) ->]. -case/and3P=>X1 Y1 Z1 /and3P [X2 Y2 Z2]. -by rewrite !(seponAxx W s1 s2 X1 X2, seponAxx W s1 s2 Y1 Y2, - seponAxx W s1 s2 Z1 Z2). -Qed. - -HB.instance Definition _ := - isSeprelOn.Build U V f (rel3I X Y Z) rel3I_is_seprelon. -End SepOn3I. - -(* 4-way conjunction *) -Section SepOn4I. -Variables (U V : pcm) (f : pcm_morph U V) (X Y Z W : seprel_on f). - -Lemma rel4I_is_seprelon : seprel_on_axiom f (rel4I X Y Z W). -Proof. -split=>[|x y VV s|x y VV s|x y z VV s1 s2] /=. -- by rewrite !sepon00. -- by rewrite !(seponC _ VV s). -- by case/and4P; do ![move/(seponx0 VV s) ->]. -case/and4P=>X1 Y1 Z1 W1 /and4P [X2 Y2 Z2 W2]. -by rewrite !(seponAxx VV s1 s2 X1 X2, seponAxx VV s1 s2 Y1 Y2, - seponAxx VV s1 s2 Z1 Z2, seponAxx VV s1 s2 W1 W2). -Qed. - -HB.instance Definition _ := - isSeprelOn.Build U V f (rel4I X Y Z W) rel4I_is_seprelon. -End SepOn4I. - -(*******************) -(* Local functions *) -(*******************) - -(* Unlike morphisms which operate over a single PCM element *) -(* local functions operate over a pair (self, other) *) -(* and then have to support a form of framing, i.e., moving from other *) -(* to self. I haven't quite yet found a need to develop their theory *) -(* but, they could be something we can call "bimorphisms" *) - -(* Transitions should fit into this category of local functions *) - -Definition local_fun (U : pcm) (f : U -> U -> option U) := - forall p x y r, valid (x \+ (p \+ y)) -> f x (p \+ y) = Some r -> - valid (r \+ p \+ y) /\ f (x \+ p) y = Some (r \+ p). - -Lemma localV U f x y r : - @local_fun U f -> valid (x \+ y) -> f x y = Some r -> valid (r \+ y). -Proof. by move=>H V E; move: (H Unit x y r); rewrite unitL !unitR; case. Qed. - -Lemma idL (U : pcm) : @local_fun U (fun x y => Some x). -Proof. by move=>p x y _ V [<-]; rewrite -joinA. Qed. - -(********************) -(* Global functions *) -(********************) - -(* given function over two arguments, make it global *) -(* by making it operate over the join *) - -Definition glob (U : pcm) R (f : U -> U -> R) := - fun x y => f (x \+ y) Unit. -Arguments glob {U R} f x y /. - -(*******************) -(* Local relations *) -(*******************) - -(* Local relations are needed at some places *) -(* but are weaker than separating relations *) -(* For example, separating relation would allow moving p from y to x *) -(* only if R p y; this is the associativity property *) -(* of seprels, and is essential for the subPCM construction *) -(* But here we don't require that property, because we won't be *) -(* modding out U by a local rel to obtain a subPCM *) -(* Also, we don't require any special behavior wrt unit. *) -(* And no commutativity (for now) *) -(* Also, its sometimes useful to have a condition under *) -(* which the relation is local *) -(* In practice, it frequently happens that some relation is a seprel *) -(* only if some other seprel already holds. Thus, it makes sense to *) -(* consider locality under a binary condition too. *) - -(* Local rel is the main concept *) -Definition local_rel (U : pcm) (R : U -> U -> Prop) (cond : U -> Prop) := - forall p x y, cond (x \+ p \+ y) -> R x (p \+ y) -> R (x \+ p) y. - - diff --git a/pcm/mutex.v b/pcm/mutex.v deleted file mode 100644 index d7af5e1..0000000 --- a/pcm/mutex.v +++ /dev/null @@ -1,583 +0,0 @@ -(* -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 defines the generalized mutex type. *) -(* We need a notion of mutex where the options are not just nown/own but *) -(* an arbitrary number of elements in-between. This reflects the possible *) -(* stages of acquiring a lock. Once a thread embarks on the first stage, it *) -(* excludes others, but it may require more steps to fully acquire the lock. *) -(******************************************************************************) - -From HB Require Import structures. -From Coq Require Import ssreflect ssrbool ssrfun. -From Coq Require Setoid. -From mathcomp Require Import ssrnat eqtype seq. -From pcm Require Import options prelude pred. -From pcm Require Import pcm morphism natmap. - -(***********************) -(* Generalized mutexes *) -(***********************) - -(* T encodes mutex stages, excluding undef and unit *) - -Inductive mutex T := mx_undef | nown | mx of T. - -Arguments mx_undef {T}. -Arguments nown {T}. -Prenex Implicits mx_undef nown. - -Section GeneralizedMutex. -Variable T : Type. -Definition def_mx_valid := [fun x : mutex T => - if x is mx_undef then false else true]. -Definition def_mx_join := [fun x y : mutex T => - match x, y with - mx_undef, _ => mx_undef - | _, mx_undef => mx_undef - | nown, x => x - | x, nown => x - | mx _, mx _ => mx_undef - end]. -Definition def_mx_unitb := [fun x : mutex T => - if x is nown then true else false]. - -Lemma mutex_is_pcm : pcm_axiom def_mx_valid def_mx_join nown def_mx_unitb. -Proof. -by split=>[[||x][||y]|[||x][||y][||z]|[]|[]||[||x]] //; constructor. -Qed. -HB.instance Definition _ : isPCM (mutex T) := - isPCM.Build (mutex T) mutex_is_pcm. - -(* cancellativity *) -Lemma mutex_is_cpcm : cpcm_axiom (mutex T). -Proof. by case=>[||m][||m1][||m2]; rewrite !pcmE. Qed. -HB.instance Definition _ : isCPCM (mutex T) := - isCPCM.Build (mutex T) mutex_is_cpcm. - -(* topped structure *) -Definition def_mx_undefb (x : mutex T) : bool := - if x is mx_undef then true else false. - -Lemma mutex_is_tpcm : tpcm_axiom mx_undef def_mx_undefb. -Proof. by split=>//; case=>[||x]; constructor. Qed. -HB.instance Definition _ : isTPCM (mutex T) := - isTPCM.Build (mutex T) mutex_is_tpcm. - -(* normality *) -Lemma mutex_is_normal : normal_tpcm_axiom (mutex T). -Proof. by case; [right|rewrite valid_unit; left|move=>t; left]. Qed. -HB.instance Definition _ : isNormal_TPCM (mutex T) := - isNormal_TPCM.Build (mutex T) mutex_is_normal. -End GeneralizedMutex. - -(* if T is eqType, so is mutex T *) -Section Equality. -Variable T : eqType. - -Definition mutex_eq (x y : mutex T) := - match x, y with - mx_undef, mx_undef => true - | nown, nown => true - | mx x', mx y' => x' == y' - | _, _ => false - end. - -Lemma mutex_eqP : Equality.axiom mutex_eq. -Proof. -case=>[||x]; case=>[||y] /=; try by constructor. -by case: eqP=>[->|H]; constructor=>//; case=>/H. -Qed. - -HB.instance Definition _ := hasDecEq.Build (mutex T) mutex_eqP. -End Equality. - -(* 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)). - -(* some lemmas for generalized mutexes *) - -Section MutexLemmas. -Variable T : Type. -Implicit Types (t : T) (x y : mutex T). - -Variant mutex_spec x : mutex T -> Type := -| mutex_undef of x = undef : mutex_spec x undef -| mutex_nown of x = Unit : mutex_spec x Unit -| 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. - -Lemma mxE0 x y : x \+ y = Unit -> (x = Unit) * (y = Unit). -Proof. by case: x; case: y. 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. - -Lemma cancelxM t1 t2 x : (x \+ mx t1 = mx t2) <-> (t1 = t2) * (x = Unit). -Proof. by rewrite joinC cancelMx. Qed. - -(* DEVCOMMENT *) -(* the next batch of lemmas about validity should be automated *) -(* /DEVCOMMENT *) -Lemma mxMx t x : valid (mx t \+ x) -> x = Unit. -Proof. by case: x. Qed. - -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 joinA=>/mxxM. Qed. - -Lemma mxMxy t x y : valid (mx t \+ x \+ y) -> x \+ y = Unit. -Proof. by rewrite -joinA=>/mxMx. Qed. - -Lemma mxxMy t x y : valid (x \+ (mx t \+ y)) -> x \+ y = Unit. -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. - -(* 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 = Unit & y = Unit : mxjoin_spec x y Unit Unit Unit -| leftmx t of x = mx t & y = Unit : mxjoin_spec x y (mx t) (mx t) Unit -| rightmx t of x = Unit & y = mx t : mxjoin_spec x y (mx t) Unit (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. - -(* DEVCOMMENT *) -(* and the same for own; do we need to repeat? *) -(* /DEVCOMMENT *) -Section OwnMutex. -Variables T : Type. -Implicit Types x y : mtx T. - -Lemma mxOx x : valid (own \+ x) -> x = Unit. -Proof. by apply: mxMx. Qed. - -Lemma mxxO x : valid (x \+ own) -> x = Unit. -Proof. by apply: mxxM. Qed. - -Lemma mxxyO x y : valid (x \+ (y \+ own)) -> x \+ y = Unit. -Proof. by apply: mxxyM. Qed. - -Lemma mxOxy x y : valid (own \+ x \+ y) -> x \+ y = Unit. -Proof. by apply: mxMxy. Qed. - -Lemma mxxOy x y : valid (x \+ (own \+ y)) -> x \+ y = Unit. -Proof. by apply: mxxMy. Qed. - -Lemma mxyOx x y : valid (y \+ own \+ x) -> y \+ x = Unit. -Proof. by apply: mxyMx. Qed. -End OwnMutex. - -Prenex Implicits mxOx mxxO mxxyO mxOxy mxxOy mxyOx. - -(* specific lemmas for binary mutexes *) - -(* DEVCOMMENT *) -(* these also are a bit of a featuritis *) -(* /DEVCOMMENT *) -Lemma mxON (x : mtx2) : valid x -> x != own -> x = Unit. -Proof. by case: x=>//; case. Qed. - -Lemma mxNN (x : mtx2) : valid x -> x != Unit -> x = own. -Proof. by case: x=>//; case=>//; case. Qed. - -(* the next batch of lemmas is for automatic simplification *) - -Section MutexRewriting. -Variable T : eqType. -Implicit Types (t : T) (x : mutex T). - -Lemma mxE1 : (((undef == Unit :> mutex T) = false) * - ((Unit == undef :> mutex T) = false)). -Proof. by []. Qed. - -Lemma mxE2 t : (((mx t == Unit) = false) * - ((Unit == mx t) = false)) * - (((mx t == undef) = false) * - ((undef == mx t) = false)). -Proof. by []. Qed. - -Lemma mxE3 t x : ((((mx t \+ x == Unit) = false) * - ((x \+ mx t == Unit) = false)) * - (((Unit == mx t \+ x) = false) * - ((Unit == 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))). -Proof. -have L : forall t1 t2 x, (mx t1 \+ x == mx t2) = (t1 == t2) && (x == Unit). -- 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_injE t1 t2 : (mx t1 == mx t2) = (t1 == t2). -Proof. by []. Qed. - -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))). - -End MutexRewriting. - -(* function mapping all mx's to own *) -Definition mxown {T} (x : mutex T) : mtx2 := - match x with - mx_undef => undef - | nown => Unit - | _ => own - end. - -(* this is a morphism under full domain *) -Lemma mxown_is_pcm_morph {T} : pcm_morph_axiom relT (@mxown T). -Proof. by split=>[|x y] //; case: x; case: y. Qed. -HB.instance Definition _ T := - isPCM_morphism.Build (mutex T) mtx2 mxown mxown_is_pcm_morph. -HB.instance Definition _ T := - isFull_PCM_morphism.Build (mutex T) mtx2 (@mxown T) (fun _ _ => erefl _). -HB.instance Definition _ T := - isTPCM_morphism.Build (mutex T) mtx2 mxown (erefl undef). -Lemma mxown_is_binormal {T} : binorm_pcm_morph_axiom (@mxown T). -Proof. by case=>[||x][||y]. Qed. -HB.instance Definition _ T := - isBinorm_PCM_morphism.Build (mutex T) mtx2 mxown mxown_is_binormal. - -(* 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 = Unit -> x = Unit. -Proof. by case: x. Qed. - -Definition mxundef {T} (x : mtx T) : mtx2 := - match x with - | nown => Unit - | 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 sep_mxundef_is_seprel {T} : seprel_axiom (@sep_mxundef T). -Proof. by split=>//; case=>[||x] // [||y] // [||[]] //=; case: y. Qed. -HB.instance Definition _ T := - isSeprel.Build (mtx T) sep_mxundef sep_mxundef_is_seprel. - -Lemma mxundef_is_pcm_morph {T} : pcm_morph_axiom (@sep_mxundef T) mxundef. -Proof. by split=>//; case=>[||x] // [||[]] //; case: x. Qed. -HB.instance Definition _ T := - isPCM_morphism.Build (mtx T) mtx2 mxundef mxundef_is_pcm_morph. -HB.instance Definition _ T := - isTPCM_morphism.Build (mtx T) mtx2 mxundef (erefl undef). -Lemma mxundef_is_binormal {T} : binorm_pcm_morph_axiom (@mxundef T). -Proof. by case=>[||[x|]][||[y|]]. Qed. -HB.instance Definition _ T := - isBinorm_PCM_morphism.Build (mtx T) mtx2 mxundef mxundef_is_binormal. - -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 Unit Unit -| 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 Unit else own. - -Variant nxown_spec n : mtx2 -> Type := -| nxZ of n = 0 : nxown_spec n Unit -| 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 Unit -| 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. - -(**********************************) -(* 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. -Variable U : natmap (bool * bool). - -Definition omg_sep := fun x y : U => - [&& last_atval false x ==> (last_key y < last_key x) & - last_atval false y ==> (last_key x < last_key y)]. - -Lemma omg_is_seprel : seprel_axiom omg_sep. -Proof. -rewrite /omg_sep; split=>[|x y|x y|x y z] /=. -- by rewrite lastatval0. -- by rewrite andbC. -- move=>V /andP [H _]; rewrite lastkey0 lastatval0 /=. - by case: (x in x ==> _) H=>// /(leq_trans _) ->. -move=>V /andP [X Y] /andP []. -rewrite !lastkeyUn !lastatvalUn !(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. - -HB.instance Definition _ := isSeprel.Build U omg_sep omg_is_seprel. - -Definition omg (x : U) : mtx2 := - if undefb x then undef else - if last_atval false x then own else nown. - -(* DEVCOMMENT *) -(* omg isn't tpcm morphism because it doesn't preserve undef *) -(* this makes it less useful than it might be *) -(* but we wait with fixing the definition until it becomes necessary *) -(* (the definition should branch on x being undef *) -(* /DEVCOMMENT *) - -Lemma omg_is_pcm_morph : pcm_morph_axiom omg_sep omg. -Proof. -rewrite /omg; split=>[|x y V /andP [X Y]] /=. -- by rewrite undefb0 lastatval0. -rewrite !undefNV !(validE2 V) /= lastatvalUn V; case: ltngtP X Y=>H X Y; -by rewrite ?(negbTE X) ?(negbTE Y) //; case: ifP. -Qed. -HB.instance Definition _ := isPCM_morphism.Build U mtx2 omg omg_is_pcm_morph. - -Lemma omg_is_tpcm_morph : tpcm_morph_axiom omg. -Proof. by rewrite /tpcm_morph_axiom/omg undefb_undef. Qed. -HB.instance Definition _ := isTPCM_morphism.Build U mtx2 omg omg_is_tpcm_morph. - -Lemma omg_is_norm : norm_pcm_morph_axiom omg. -Proof. -move=>/= x; rewrite /sepx/=/omg/omg_sep lastkey0 lastatval0. -case: (normalP x)=>// Vx; case: ifP=>H _ //=. -by case: lastkeyP (lastatval_indomb H). -Qed. -HB.instance Definition _ := isNorm_PCM_morphism.Build U mtx2 omg omg_is_norm. - -(* transfer lemmas *) - -(* omg isn't full, but admits valid h -> valid (omg h) *) -Lemma valid_omg h : valid (omg h) = valid h. -Proof. -apply/idP/idP=>[/fpV//|V]. -rewrite pfV //= /sepx/= /omg_sep /= lastatval0 andbT lastkey0. -case N: (last_key h); last by apply/implyP. -by rewrite /last_atval /atval /= cond_find // N. -Qed. - -Lemma omgPos (V : pcm) (v : V) (ht : V -> U) : - last_atval false (ht v) = (omg (ht v) == own). -Proof. -rewrite /omg; case: normalP=>[->//|]; last by case: ifP. -by rewrite lastatval_undef. -Qed. - -Lemma omgPosMorph (V : pcm) (v1 v2 : V) (ht : pcm_morph V U): - valid (v1 \+ v2) -> - preim ht omg_sep v1 v2 -> - last_atval false (ht v1 \+ ht v2) = - (omg (ht v1) \+ omg (ht v2) == own). -Proof. -move=>W /andP [G] /andP []; rewrite /omg /= in G *. -rewrite !undefNV !(validE2 (pfV2 _ _ _ W G)) lastatvalUn pfV2 //=. -by case: ltngtP=>H H1 H2; rewrite ?(negbTE H1) ?(negbTE H2) //; case: ifP. -Qed. - -Lemma omgNeg' (V : pcm) (v : V) (ht : V -> U) : - ~~ last_atval false (ht v) = - (omg (ht v) == nown) || undefb (omg (ht v)). -Proof. by rewrite omgPos; case: (omg _)=>//; case. Qed. - -Lemma omgNeg (V : pcm) (v : V) (ht : V -> U) : - valid (ht v) -> - ~~ last_atval false (ht v) = (omg (ht v) == nown). -Proof. by move=>W; rewrite omgNeg' undefNV valid_omg W; case: (omg _). Qed. - -Lemma omgNegMorph (V : pcm) (v1 v2 : V) (ht : pcm_morph V U) : - valid (v1 \+ v2) -> preim ht omg_sep v1 v2 -> - ~~ last_atval false (ht v1 \+ ht v2) = - (omg (ht v1) \+ omg (ht v2) == nown). -Proof. -move=>W /andP [H1 H2]. -have W' : valid (ht (v1 \+ v2)) by rewrite pfV // (sepU0 W H1). -by rewrite -!pfjoin //= omgNeg. -Qed. - -Lemma omgidPos (v : U) : last_atval false v = (omg v == own). -Proof. by rewrite (omgPos _ id). Qed. - -Lemma omgidPosMorph (v1 v2 : U) : - valid (v1 \+ v2) -> omg_sep v1 v2 -> - last_atval false (v1 \+ v2) = (omg v1 \+ omg v2 == own). -Proof. by move=>W S; rewrite (omgPosMorph (ht:=idfun)). Qed. - -Lemma omgidNeg (v : U) : - valid v -> - ~~ last_atval false v = (omg v == Unit). -Proof. exact: (@omgNeg _ _ id). Qed. - -Lemma omgidNegMorph (v1 v2 : U) : - valid (v1 \+ v2) -> omg_sep v1 v2 -> - ~~ last_atval false (v1 \+ v2) = - (omg v1 \+ omg v2 == Unit). -Proof. by move=>W S; rewrite (omgNegMorph (ht:=idfun)). Qed. - -Definition omgP := - (valid_omg, - (omgidNegMorph, omgidPosMorph, omgPosMorph, omgNegMorph), - (omgidPos, omgidNeg, omgPos, omgNeg)). - -Lemma omg_fresh_val (V : pcm) (v1 v2 : V) (ht : pcm_morph V U) : - valid (v1 \+ v2) -> - sep ht v1 v2 -> - (omg (pts (fresh (ht v1 \+ ht v2)) (false, true) \+ ht v1) = own) * - (omg (pts (fresh (ht v1 \+ ht v2)) (true, false) \+ ht v1) = nown) * - (omg (pts (fresh (ht v1 \+ ht v2)) (false, true) \+ ht v2) = own) * - (omg (pts (fresh (ht v1 \+ ht v2)) (true, false) \+ ht v2) = nown). -Proof. -move=>A O; have Vh : valid (ht v1 \+ ht v2) by rewrite pfV2. -rewrite /omg !undefNV !validPtUn /= !(validE2 Vh) !negbK /=. -by rewrite !lastatval_freshUn // !(negbTE (fresh_dom _)) ?(freshUnL,freshUnR). -Qed. - -Lemma omg_fresh_sep (V : pcm) (v1 v2 : V) (ht : pcm_morph V U) op : - valid (v1 \+ v2) -> - sep ht v1 v2 -> - (omg (ht v2) = Unit -> - omg_sep (pts (fresh (ht v1 \+ ht v2)) op \+ ht v1) (ht v2)) * - (omg (ht v1) = Unit -> - omg_sep (ht v1) (pts (fresh (ht v1 \+ ht v2)) op \+ ht v2)). -Proof. -move=>A O; have Vh : valid (ht v1 \+ ht v2) by rewrite pfV2. -rewrite /omg_sep lastatval_freshUn //. -rewrite !lastkeyPtUn ?(validE2 Vh,freshUnL,freshUnR) //. -by split=>N; rewrite implybT omgP N. -Qed. - -Definition omg_fresh := (omg_fresh_val, omg_fresh_sep). - -(* DEVCOMMENT *) -(* some extra properties the need for which appeared later *) -(* TODO eventually organize into their proper places *) -(* /DEVCOMMENT *) -Lemma omg_eta (h : U): - valid h -> omg h = own -> - exists h' v, [/\ h' = free h (last_key h), - h = pts (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. -move=>V; rewrite /omg /=; rewrite undefNV V /=. -case: ifP=>// N _; set k := last_key h. -have D : k \in dom h. -- rewrite /last_atval /atval /oapp in N. - by case: dom_find N. -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_atval /atval Ef /= in N; rewrite -N. -have: last_key h' <= k. -- by rewrite /k Eh lastkeyPtUnE -Eh V leq_maxr. -rewrite leq_eqVlt; case/orP=>// /eqP E. -by rewrite -E in Nd K; case: lastkeyP K Nd. -Qed. - -(* specialize to alternating histories *) -Lemma omg_eta_all (h : U) : - valid h -> omg h = own -> um_all (fun k v => v.2 = ~~ v.1) h -> - exists h', [/\ h' = free h (last_key h), - h = pts (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_sep x y -> - last_key (x \+ y) = last_key x) * - (omg y = own -> valid (x \+ y) -> omg_sep x y -> - last_key (x \+ y) = last_key y). -Proof. -rewrite /omg_sep /omg !undefNV. -split=>L V S; rewrite (validE2 V) /= in L; 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. - -Arguments omg {U}. -Arguments omg_sep {U}. - diff --git a/pcm/natmap.v b/pcm/natmap.v deleted file mode 100644 index a8bcb40..0000000 --- a/pcm/natmap.v +++ /dev/null @@ -1,2860 +0,0 @@ -(* -Copyright 2015 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 maps over non-zero natural *) -(* numbers, pcm instance for natmap, gapless natmaps, natmaps with binary *) -(* range, several sorts of continuous natmaps. *) -(* Histories are a special case of natmaps. *) -(******************************************************************************) - -From HB Require Import structures. -From Coq Require Import ssreflect ssrbool ssrfun. -From mathcomp Require Import ssrnat eqtype seq path interval. -From pcm Require Import axioms options prelude pred finmap rtc seqext. -From pcm Require Export useqord uslice uconsec pcm morphism unionmap. -Import order.Order.NatOrder. (* listed last to avoid notation clash *) -Local Open Scope order_scope. - -(************************) -(* Maps over non-0 nats *) -(************************) - -Definition null := 0. -Notation nat_pred := (fun x => x != 0). - -(* natmap is union map of non-0 nat keys *) -HB.mixin Record isNatMap V U of UMC nat (fun x => x != 0) V U. -#[short(type="natmap")] -HB.structure Definition NatMap V := { - U of UMC nat (fun x => x != 0) V U & isNatMap V U}. - -(* pred structure (to write x \In h) is copied from union_map *) -Canonical natmap_PredType A (U : natmap A) : PredType (nat * A) := - um_PredType U. -Coercion Pred_of_natmap A (U : natmap A) (x : U) : {Pred _} := - [eta Mem_UmMap x]. - -(* canonical natmap is histories *) - -Record history A := Hist {hist_base : @UM.base nat nat_pred A}. - -Section HistoryUMC. -Variables A : Type. -Implicit Type f : history A. -Local Coercion hist_base : history >-> UM.base. -Let ht_valid f := @UM.valid nat nat_pred A f. -Let ht_empty := Hist (@UM.empty nat nat_pred A). -Let ht_undef := Hist (@UM.Undef nat nat_pred A). -Let ht_upd k v f := Hist (@UM.upd nat nat_pred A k v f). -Let ht_dom f := @UM.dom nat nat_pred A f. -Let ht_assocs f := @UM.assocs nat nat_pred A f. -Let ht_free f k := Hist (@UM.free nat nat_pred A f k). -Let ht_find k f := @UM.find nat nat_pred A k f. -Let ht_union f1 f2 := Hist (@UM.union nat nat_pred A f1 f2). -Let ht_empb f := @UM.empb nat nat_pred A f. -Let ht_undefb f := @UM.undefb nat nat_pred A f. -Let ht_from (f : history A) : UM.base _ _ := f. -Let ht_to (b : @UM.base nat nat_pred A) : history A := Hist b. -Let ht_pts k v := Hist (@UM.pts nat nat_pred A k v). - -Lemma history_is_umc : - union_map_axiom ht_valid ht_empty ht_undef ht_upd ht_dom - ht_assocs ht_free ht_find ht_union ht_empb - ht_undefb ht_pts ht_from ht_to. -Proof. by split=>//; split=>[|[]]. Qed. - -HB.instance Definition _ := - isUnion_map.Build nat nat_pred A (history A) history_is_umc. -End HistoryUMC. - -HB.instance Definition _ A := isNatMap.Build A (history A). -HB.instance Definition _ (A : eqType) := - hasDecEq.Build (history A) (@union_map_eqP nat _ A (history A)). -Canonical history_PredType A : PredType (nat * A) := - um_PredType (history A). -Coercion Pred_of_history A (x : history A) : {Pred _} := - [eta Mem_UmMap x]. - -Notation "x \-> v" := (ptsT (history _) x v) (at level 30). - -(* DEVCOMMENT *) -(* tests *) -Lemma xx : 1 \-> true = null \-> false. -Abort. - -Lemma xx : ((1 \-> true) \+ (2 \-> false)) == (1 \-> false). -rewrite joinC. -Abort. - -Lemma xx (x : history nat) : x \+ x == x \+ x. -Abort. - -Lemma xx : 1 \-> (1 \-> 3) = 2 \-> (7 \-> 3). -Abort. - -Lemma xx : (1, 3) \In (1 \-> 3). -Abort. -(* /DEVCOMMENT *) - -(* Sometimes it's useful to not think of natmaps as recording *) -(* times (i.e., being histories), but just as ordinary maps *) -(* over nats. For that occassion, use the nmap type *) - -Record nmap A := NMap {nmap_base : @UM.base nat nat_pred A}. - -Section NMapUMC. -Variables A : Type. -Implicit Type f : nmap A. -Local Coercion nmap_base : nmap >-> UM.base. -Let nm_valid f := @UM.valid nat nat_pred A f. -Let nm_empty := NMap (@UM.empty nat nat_pred A). -Let nm_undef := NMap (@UM.Undef nat nat_pred A). -Let nm_upd k v f := NMap (@UM.upd nat nat_pred A k v f). -Let nm_dom f := @UM.dom nat nat_pred A f. -Let nm_assocs f := @UM.assocs nat nat_pred A f. -Let nm_free f k := NMap (@UM.free nat nat_pred A f k). -Let nm_find k f := @UM.find nat nat_pred A k f. -Let nm_union f1 f2 := NMap (@UM.union nat nat_pred A f1 f2). -Let nm_empb f := @UM.empb nat nat_pred A f. -Let nm_undefb f := @UM.undefb nat nat_pred A f. -Let nm_from (f : nmap A) : UM.base _ _ := f. -Let nm_to (b : @UM.base nat nat_pred A) : nmap A := NMap b. -Let nm_pts k v := NMap (@UM.pts nat nat_pred A k v). - -Lemma nmap_is_umc : - union_map_axiom nm_valid nm_empty nm_undef nm_upd nm_dom - nm_assocs nm_free nm_find nm_union nm_empb - nm_undefb nm_pts nm_from nm_to. -Proof. by split=>//; split=>[|[]]. Qed. - -HB.instance Definition _ := - isUnion_map.Build nat nat_pred A (nmap A) nmap_is_umc. -End NMapUMC. - -HB.instance Definition _ A := isNatMap.Build A (nmap A). -HB.instance Definition _ (A : eqType) := - hasDecEq.Build (nmap A) (@union_map_eqP nat _ A (nmap A)). -Canonical nmap_PredType A : PredType (nat * A) := - um_PredType (nmap A). -Coercion Pred_of_nmap A (x : nmap A) : {Pred _} := - [eta Mem_UmMap x]. - -(* no points-to notation for nmaps *) - -(*************) -(* Main defs *) -(*************) - -Definition last_key {A} {U : natmap A} (h : U) := last 0 (dom h). -Definition fresh {A} {U : natmap A} (h : U) := (last_key h).+1. -Definition last_val {A} {U : natmap A} (h : U) := find (last_key h) h. - -(**********************) -(* generic properties *) -(**********************) - -(* alternative to In_cond *) -Lemma In_neq A (U : natmap A) (h : U) k : - k \In h -> - k.1 > 0. -Proof. by move/In_cond; rewrite lt0n. Qed. - -Lemma mem_domN0 A (U : natmap A) {h : U} : 0 \notin dom h. -Proof. by rewrite cond_dom. Qed. - -Lemma mem_dom0 A (U : natmap A) {h : U} : 0 \in dom h = false. -Proof. by rewrite cond_dom. Qed. - -Lemma uniq_dom0 A (U : natmap A) {h : U} : uniq (0 :: dom h). -Proof. by rewrite /= uniq_dom cond_dom. Qed. - -Lemma all_leq0 A (U : natmap A) {h : U} : all (leq 0) (dom h). -Proof. by apply/allP=>x /dom_cond; case: ltngtP. Qed. - -Lemma all_ltn0 A (U : natmap A) {h : U} : all (ltn 0) (dom h). -Proof. by apply/allP=>x /dom_cond; case: ltngtP. Qed. - -#[export] Hint Resolve uniq_dom0 all_leq0 all_ltn0 mem_domN0 : core. - -Lemma In_dom0 {A} {U : natmap A} (h : U) k e : - (k, e) \In h -> - k \in 0 :: dom h. -Proof. by move=>H; rewrite inE (In_dom H) orbT. Qed. - -Lemma sorted_leq_dom A {U : natmap A} {h : U} : sorted leq (dom h). -Proof. by rewrite -(eq_sorted nat_oleq) sorted_dom_oleq. Qed. - -Lemma sorted_ltn_dom A {U : natmap A} {h : U} : sorted ltn (dom h). -Proof. by rewrite sorted_dom. Qed. - -Lemma path_leq_dom A {U : natmap A} {h : U} : path leq 0 (dom h). -Proof. by rewrite path_min_sorted // sorted_leq_dom. Qed. - -Lemma path_ltn_dom A {U : natmap A} {h : U} : path ltn 0 (dom h). -Proof. by rewrite path_min_sorted. Qed. - -#[export] Hint Resolve sorted_leq_dom sorted_ltn_dom - path_leq_dom path_ltn_dom : core. - -(* form of totality of key order *) -Lemma umfiltT A (U : natmap A) k1 k2 (h : U) : - k1 \in dom h -> k1 \notin dom (um_filterk (ltn^~ k2) h) -> - k2 \in dom h -> k2 \notin dom (um_filterk (ltn^~ k1) h) -> - k1 = k2. -Proof. -case/In_domX=>v1 D1 S1 /In_domX [v2 D2] S2. -case: (ltngtP k1 k2)=>// N. -- by move/(In_umfiltk (ltn^~ k2))/(_ N)/In_dom: D1; rewrite (negbTE S1). -by move/(In_umfiltk (ltn^~ k1))/(_ N)/In_dom: D2; rewrite (negbTE S2). -Qed. - -Lemma umfiltT0 A (U : natmap A) k1 k2 (h : U) : - k1 \in dom h -> um_filterk (ltn^~ k1) h = Unit -> - k2 \in dom h -> um_filterk (ltn^~ k2) h = Unit -> - k1 = k2. -Proof. -by move=>D1 U1 D2 U2; apply: umfiltT D1 _ D2 _; rewrite ?U1 ?U2 dom0. -Qed. - -Lemma omf_subdom0 {A1 A2} {U1 : natmap A1} {U2 : natmap A2} - {f : omap_fun U1 U2} {h : U1} : - {subset 0::dom (f h) <= 0::dom h}. -Proof. -move=>x; rewrite !inE; case/orP=>[|/omf_subdom] -> //. -by rewrite orbT. -Qed. - -(**************) -(* Base cases *) -(**************) - -Lemma lastkey_undef A (U : natmap A) : last_key (undef : U) = 0. -Proof. by rewrite /last_key dom_undef. Qed. - -Lemma lastkey0 A (U : natmap A) : last_key (Unit : U) = 0. -Proof. by rewrite /last_key dom0. Qed. - -Lemma lastkeyPt A (U : natmap A) (x : nat) (v : A) : - last_key (pts x v : U) = x. -Proof. by rewrite /last_key domPtK; case: (x =P 0). Qed. - -Lemma lastval0 A (U : natmap A) : last_val (Unit : U) = None. -Proof. by rewrite /last_val lastkey0 find0E. Qed. - -Lemma lastval_undef A (U : natmap A) : last_val (undef : U) = None. -Proof. by rewrite /last_val lastkey_undef find_undef. Qed. - -Lemma lastvalPt A (U : natmap A) x v : - last_val (pts x v : U) = if x != 0 then Some v else None. -Proof. by rewrite /last_val lastkeyPt findPt. Qed. - -(* basic transfer lemma *) -Lemma In_last A (U : natmap A) (h : U) v : - (last_key h, v) \In h <-> last_val h = Some v. -Proof. exact: In_find. Qed. - -(* for inlined rewriting *) -Lemma In_lastE A (U : natmap A) (h : U) v : - (last_key h, v) \In h -> last_val h = Some v. -Proof. exact: In_findE. Qed. - -(* fresh and 0 *) - -Lemma fresh_undef A (U : natmap A) : fresh (undef : U) = 1. -Proof. by rewrite /fresh lastkey_undef. Qed. - -Lemma fresh0 A (U : natmap A) : fresh (Unit : U) = 1. -Proof. by rewrite /fresh lastkey0. Qed. - -Lemma fresh_lt0 A (U : natmap A) (h : U) x : fresh h <= x -> 0 < x. -Proof. by case: x. Qed. - -Lemma freshN0 A (U : natmap A) (h : U) x : fresh h <= x -> x != 0. -Proof. by case: x. Qed. - -(*************) -(* Main view *) -(*************) - -(* captures that following properties are equivalent: *) -(* (undefb && unitb) = (last_key != 0) = (last_key > 0) = ... *) -(* ... = (last_key \in dom) = (dom <> [::]) = lastval = Some *) - -Variant lastkey_dom_spec A (U : natmap A) (h : U) : - bool -> bool -> bool -> nat -> seq nat -> - bool -> bool -> bool -> option A -> Type := -| lastkey_dom_undef of h = undef : - lastkey_dom_spec h false true false 0 [::] true false false None -| lastkey_dom_unit of h = Unit : - lastkey_dom_spec h true false true 0 [::] true false false None -| lastkey_dom_valid v of (last_key h, v) \In h : - lastkey_dom_spec h true false false (last_key h) - (dom h) false true true (Some v). - -Lemma lastkeyP A (U : natmap A) (h : U) : - lastkey_dom_spec h (valid h) (undefb h) (unitb h) - (last_key h) (dom h) (last_key h == 0) - (last_key h > 0) (last_key h \in dom h) - (last_val h). -Proof. -have L (x : U) : valid x -> last_key x \notin dom x -> x = Unit. -- rewrite /last_key !umEX /UM.valid/UM.dom/UM.empty -{4}[x]tfE. - case: (UMC_from x)=>//=; case=>s H H1 _ /seq_last_in. - by rewrite eqE UM.umapE /supp fmapE /= {H H1}; elim: s. -case: (normalP0 h)=>[->|->|]. -- by rewrite lastkey_undef dom_undef lastval_undef; apply: lastkey_dom_undef. -- by rewrite lastkey0 dom0 lastval0; apply: lastkey_dom_unit. -move=>V H. -have H1 : last_key h \in dom h by apply: contraT=>/(L _ V)/H. -have H2 : last_key h != 0 by apply: dom_cond H1. -have [v H4] : {v | last_val h = Some v}. -- by rewrite /last_val; case: dom_find H1=>// v; exists v. -by rewrite H1 lt0n (negbTE H2) H4; apply/lastkey_dom_valid/In_last. -Qed. - -(* main membership invariant of last_key *) -Lemma lastkey_mem0 A (U : natmap A) (h : U) : - last_key h \in 0 :: dom h. -Proof. by rewrite inE; case: lastkeyP. Qed. - -(* view for last_val can simplify obligations from last_keyP *) -(* by joining the cases for unit and undef *) -Variant lastval_spec A (U : natmap A) (h : U) : option A -> Type := -| lastval_none of last_val h = None & (forall kv, kv \Notin h) : - lastval_spec h None -| lastval_some r kv of last_val h = Some r & kv \In h : - lastval_spec h (Some r). - -Lemma lastvalP A (U : natmap A) (h : U) : lastval_spec h (last_val h). -Proof. -case: lastkeyP=>[->|->|v]. -- by apply: lastval_none; [rewrite lastval_undef|move=>kv /In_undef]. -- by apply: lastval_none; [rewrite lastval0|move=>kv /In0]. -by move=>H; apply/lastval_some/H/In_findE. -Qed. - -(************) -(* last_key *) -(************) - -Section LastkeyLemmas. -Variables (A : Type) (U : natmap A). -Implicit Type h : U. - -(* last_key is upper bound *) -Lemma dom_lastkey h k : - k \in dom h -> - k <= last_key h. -Proof. -rewrite /last_key !umEX /UM.dom; case: (UMC_from h)=>//; case=>s H _ /=. -rewrite /supp/ord /= (leq_eqVlt k) orbC. -by apply: sorted_last_key_maxR otrans (sorted_oleq H). -Qed. - -(* useful variants with 0 :: dom and contrapositives *) -Lemma dom0_lastkey h k : k \in 0 :: dom h -> k <= last_key h. -Proof. by rewrite inE; case/orP=>[/eqP -> //|]; apply: dom_lastkey. Qed. -Lemma lastkey_dom h k : last_key h < k -> k \notin dom h. -Proof. by apply: contraL; rewrite -leqNgt; apply: dom_lastkey. Qed. -Lemma lastkey_dom0 h k : last_key h < k -> k \notin 0 :: dom h. -Proof. by apply: contraL; rewrite -leqNgt; apply: dom0_lastkey. Qed. - -(* last_key is least upper bound *) -Lemma lastkey_lub h k : - {in dom h, forall x, x <= k} -> - last_key h <= k. -Proof. by case lastkeyP=>// v /In_dom H; apply. Qed. - -(* variant with contrapositive side-condition *) -Lemma lastkey_memk h k : - (forall x, k < x -> x \notin dom h) -> - last_key h <= k. -Proof. -move=>H; apply: lastkey_lub=>x; rewrite leqNgt. -by apply/contraL/H. -Qed. - -(* useful variant with equality *) -Lemma lub_lastkey h k : - k \in dom h -> - (forall x, k < x -> x \notin dom h) -> - k = last_key h. -Proof. -case: lastkeyP=>// v H /dom_lastkey Dx X. -rewrite (@anti_leq (last_key _) k) // Dx andbT. -apply: lastkey_lub=>x; apply: contraTT. -by rewrite -ltnNge; apply: X. -Qed. - -(* no need to quantify over x in lastkey_memk *) -(* it suffices to use x := last_key h *) -Lemma lastkey_memkX h k : - (k < last_key h -> last_key h \notin dom h) -> - last_key h <= k. -Proof. by case: lastkeyP=>//= v _; case: ltngtP=>// _; apply. Qed. - -(* no need to quantify over x in lub_lastkey *) -(* it suffices to use x := last_key h *) -Lemma lub_lastkeyX h k : - k \in dom h -> - (k < last_key h -> last_key h \notin dom h) -> - k = last_key h. -Proof. -move/dom_lastkey=>D /lastkey_memkX H. -by rewrite (@anti_leq (last_key _) k) // H D. -Qed. - -(* equivalences *) - -Lemma lastkey_leq h k : - (last_key h <= k) = all (fun x => x <= k) (dom h). -Proof. -apply/idP/allP=>[H x /dom_lastkey D|]; -by [apply: leq_trans D H | apply: lastkey_lub]. -Qed. - -Lemma leq_lastkey h k : - (k <= last_key h) = has (fun x => k <= x) (0 :: dom h). -Proof. -apply/idP/hasP=>[|[x] D H]; last by apply: leq_trans H (dom0_lastkey D). -by exists (last_key h)=>//; apply: lastkey_mem0. -Qed. - -Lemma ltn_lastkey h k : - (k < last_key h) = has (fun x => k < x) (dom h). -Proof. -apply/idP/hasP=>[|[x] D H]; last by apply: leq_trans H (dom_lastkey D). -by exists (last_key h)=>//; case: lastkeyP H. -Qed. - -Lemma lastkey_ltn h k : - (last_key h < k) = all (fun x => x < k) (0 :: dom h). -Proof. -apply/idP/allP=>[H x D|]; last by apply; apply: lastkey_mem0. -by apply: leq_ltn_trans (dom0_lastkey D) H. -Qed. - -(* unfolding equivalences into implications (forward) *) -Lemma lastkey_leq_dom h x k : last_key h <= k -> x \in dom h -> x <= k. -Proof. by rewrite lastkey_leq=>/allP; apply. Qed. -Lemma lastkey_leq_dom0 h x k : last_key h <= k -> x \in 0 :: dom h -> x <= k. -Proof. by move=>H; case/orP=>[/eqP ->//|]; apply: lastkey_leq_dom H. Qed. -Lemma lastkey_ltn_dom0 h x k : last_key h < k -> x \in 0 :: dom h -> x < k. -Proof. by rewrite lastkey_ltn=>/allP; apply. Qed. -Lemma lastkey_ltn_dom h x k : last_key h < k -> x \in dom h -> x < k. -Proof. by move=>H D; apply: lastkey_ltn_dom0 H _; rewrite inE D orbT. Qed. - -(* unfolding equivalences into implications (backward) *) -Lemma dom0_leq_lastkey h x k : x \in 0 :: dom h -> k <= x -> k <= last_key h. -Proof. by move=>D H; rewrite leq_lastkey; apply/hasP; exists x. Qed. -Lemma dom_leq_lastkey h x k : x \in dom h -> k <= x -> k <= last_key h. -Proof. by move=>D; apply: dom0_leq_lastkey; rewrite inE D orbT. Qed. -Lemma dom_ltn_lastkey h x k : x \in dom h -> k < x -> k < last_key h. -Proof. by move=>D H; rewrite ltn_lastkey; apply/hasP; exists x. Qed. -Lemma dom0_ltn_lastkey h x k : x \in 0 :: dom h -> k < x -> k < last_key h. -Proof. by case/orP=>[/eqP ->//|]; apply: dom_ltn_lastkey. Qed. -End LastkeyLemmas. - -(* interaction of two lastkeys *) - -Lemma lastkey_mono A1 A2 (U1 : natmap A1) (U2 : natmap A2) (h1 : U1) (h2 : U2) : - {subset dom h1 <= dom h2} -> - last_key h1 <= last_key h2. -Proof. -rewrite leq_eqVlt orbC; apply: (seq_last_monoR orefl otrans); -by rewrite (eq_path nat_oleq). -Qed. - -Lemma lastkey_subdom A1 A2 (U1 : natmap A1) (U2 : natmap A2) (h1 : U1) (h2 : U2) : - {subset dom h1 <= dom h2} -> - last_key h2 \in dom h1 -> - last_key h1 = last_key h2. -Proof. -by move=>S D; apply/eqP; rewrite eqn_leq (dom_lastkey D) (lastkey_mono S). -Qed. - -Lemma lastkeyUnL A (U : natmap A) (h1 h2 : U) : - valid (h1 \+ h2) -> - last_key h1 <= last_key (h1 \+ h2). -Proof. by move=>V; apply: lastkey_mono=>x; rewrite domUn inE V => ->. Qed. - -Lemma lastkeyUnLT A (U : natmap A) (h1 h2 : U) k : - valid (h1 \+ h2) -> - last_key (h1 \+ h2) < k -> - last_key h1 < k. -Proof. by move=>V; apply/leq_ltn_trans/lastkeyUnL/V. Qed. - -Lemma lastkeyUnR A (U : natmap A) (h1 h2 : U) : - valid (h1 \+ h2) -> - last_key h2 <= last_key (h1 \+ h2). -Proof. by rewrite joinC; apply: lastkeyUnL. Qed. - -Lemma lastkeyUnRT A (U : natmap A) (h1 h2 : U) k : - valid (h1 \+ h2) -> - last_key (h1 \+ h2) < k -> - last_key h2 < k. -Proof. by rewrite joinC; apply: lastkeyUnLT. Qed. - -(* disjoint maps with equal last keys are both empty *) -Lemma lastkeyUn0T A1 A2 (U1 : natmap A1) (U2 : natmap A2) (h1 : U1) (h2 : U2) : - valid h1 -> valid h2 -> - (forall x, x \in dom h1 -> x \in dom h2 -> False) -> - last_key h1 = last_key h2 -> - (h1 = Unit) * (h2 = Unit). -Proof. -case: (lastkeyP h1)=>[//|-> _ V _ /esym/eqP|]; first by case: (lastkeyP h2) V. -move=>v1 H1 _ V2 /(_ _ (In_dom H1)) /[swap]/eqP /=. -case: (lastkeyP h2) V2=>//; first by rewrite (negbTE (In_cond H1)). -by move=>v2 H2 _ /eqP -> /(_ (In_dom H2)). -Qed. - -(* specialization to equally-typed maps *) -Lemma lastkeyUn0 A (U : natmap A) (h1 h2 : U) : - valid (h1 \+ h2) -> - last_key h1 = last_key h2 -> - (h1 = Unit) * (h2 = Unit). -Proof. by case: validUn=>// ?? D _; apply: lastkeyUn0T=>// x /D/negbTE ->. Qed. - -(* variation to avoid validity *) -Lemma lastkey00 A1 A2 (U1 : natmap A1) (U2 : natmap A2) (h1 : U1) (h2 : U2) : - (forall x, x \in dom h1 -> x \in dom h2 -> False) -> - last_key h1 = last_key h2 -> - (last_key h1 = 0) * (last_key h2 = 0). -Proof. -case: (normalP h1)=>[->|V1] D; first by rewrite lastkey_undef. -case: (normalP h2) D=>[->|V2] D; first by rewrite lastkey_undef. -by move=>H; rewrite !(lastkeyUn0T V1 V2 D H) !lastkey0. -Qed. - -(* variation with subdoms *) -Lemma lastkey_subdom00 A1 A2 (U1 : natmap A1) (U2 : natmap A2) - (h1 : U1) (h2 : U2) (s1 s2 : pred nat) : - (forall x, x \in s1 -> x \in s2 -> False) -> - {subset dom h1 <= s1} -> - {subset dom h2 <= s2} -> - last_key h1 = last_key h2 -> - (last_key h1 = 0) * (last_key h2 = 0). -Proof. by move=>D S1 S2; apply: lastkey00=>x /S1 H /S2; apply: D H. Qed. - -(* variation with subdoms and indexing *) -Lemma lastkey_subdom00E A1 A2 (U1 : natmap A1) (U2 : natmap A2) - (g : nat -> pred nat) (h1 : U1) (h2 : U2) (t1 t2 : nat) : - (forall x, x \in g t1 -> x \in g t2 -> t1 = t2) -> - {subset dom h1 <= g t1} -> - {subset dom h2 <= g t2} -> - last_key h1 != 0 -> - last_key h1 = last_key h2 -> - t1 = t2. -Proof. -move=>H S1 S2 X E; apply: contraNeq X=>/eqP X. -move/(lastkey_subdom00 _ S1 S2): E=>-> //. -by move=>x D1 /(H x D1)/X. -Qed. - -(* interaction with constructors *) - -Section LastkeyConstructors. -Variables (A : Type) (U : natmap A). -Implicit Type h : U. - -Lemma lastkey_find h k : last_key h < k -> find k h = None. -Proof. by move/lastkey_dom; case: dom_find. Qed. - -(* monotonicity gives <=; removing last_key gives strict < *) -Lemma lastkeyF h : - last_key h != 0 -> - last_key (free h (last_key h)) < last_key h. -Proof. -move=>D; have {}D : last_key h \in dom h by case: lastkeyP D. -case: (um_eta D)=>v [Ef Eh]. -have N : last_key h \notin dom (free h (last_key h)). -- by rewrite domF inE eqxx. -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 : last_key h > 0 by move/dom_cond: D; case: (last_key h). -by case: lastkeyP N. -Qed. - -(* lastkey of union is max of lastkeys *) -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 : U) : valid (k1 \+ k2) -> - last_key k1 < last_key k2 -> last_key (k1 \+ k2) = last_key k2. -- move=>V H; apply/eqP; rewrite eqn_leq lastkeyUnR // andbT lastkey_leq. - apply/allP=>/= x; rewrite domUn inE V /=. - by case/orP=>/dom_lastkey // T; apply: leq_trans T (ltnW H). -case: (normalP (h1 \+ h2))=>[->|V]; first by rewrite lastkey_undef. -case: (ltngtP (last_key h2) (last_key h1)). -- by rewrite joinC in V *; apply: H. -- by apply: H. -by move/esym/(lastkeyUn0 V)=>E; rewrite !E unitL. -Qed. - -Lemma lastkeyPtUnV h k v : last_key h < k -> valid (pts k v \+ h) = valid h. -Proof. -move=>N; rewrite validPtUn -lt0n (leq_ltn_trans _ N) //= andbC. -by case D: (k \in dom h)=>//; move/dom_lastkey: D; case: leqP N. -Qed. - -Lemma lastkey_domPtUn h k v : - valid h -> last_key h < k -> dom (pts k v \+ h) = rcons (dom h) k. -Proof. -move=>V N; rewrite joinC domUnPtK //. -- by rewrite joinC lastkeyPtUnV. -by apply/allP=>x /dom_lastkey D; apply: leq_ltn_trans D N. -Qed. - -Lemma lastkey_rangePtUn h k v : - valid h -> last_key h < k -> range (pts k v \+ h) = rcons (range h) v. -Proof. -move=>V K; rewrite joinC rangeUnPtK //; last first. -- by apply/allP=>x /dom_lastkey H; apply: leq_ltn_trans H K. -by rewrite joinC lastkeyPtUnV. -Qed. - -Lemma lastkeyPtUnE h k v : - last_key (pts k v \+ h) = - if valid (pts k v \+ h) then maxn k (last_key h) else 0. -Proof. by rewrite lastkeyUn lastkeyPt. Qed. - -Lemma lastkeyPtUn h k v : - valid h -> last_key h < k -> last_key (pts k v \+ h) = k. -Proof. -move=>V N; rewrite lastkeyPtUnE (lastkeyPtUnV _ N) V. -by apply: maxn_idPl (ltnW N). -Qed. - -Lemma lastkeyPtUnEX v2 v1 h k : - last_key (pts k v1 \+ h) = last_key (pts k v2 \+ h). -Proof. by rewrite !lastkeyPtUnE !validPtUn. Qed. - -Lemma lastkeyUnPtEX v2 v1 h k : - last_key (h \+ pts k v1) = last_key (h \+ pts k v2). -Proof. by rewrite !(joinC h) (lastkeyPtUnEX v2). Qed. - -Lemma lastkeyU h k v : k \in dom h -> last_key (upd k v h) = last_key h. -Proof. by case/um_eta=>u [_ ->]; rewrite updPtUn (lastkeyPtUnEX u). Qed. - -Lemma lastkeyUE h k v : - last_key (upd k v h) = - if k \in dom h then last_key h else - if valid (upd k v h) then maxn k (last_key h) else 0. -Proof. -case: ifP=>[|H]; first by apply: lastkeyU. -rewrite validU; case: eqVneq=>[->|N] /=. -- by rewrite upd_condN // lastkey_undef. -case: (normalP h)=>[->|V]; first by rewrite upd_undef lastkey_undef. -by rewrite upd_eta2 ?H // lastkeyUn -upd_eta2 ?H // validU N V lastkeyPt. -Qed. - -(* backward induction on valid natmaps *) -Lemma valid_indb (P : U -> Prop) : - P Unit -> - (forall v, P (pts 1 v)) -> - (forall k v h, - P h -> last_key h < k -> - valid (pts k v \+ h) -> P (pts k v \+ h)) -> - forall h, valid h -> P h. -Proof. -move=>H1 H2 H3; elim/um_indb=>[||k v f H V K _]; rewrite ?valid_undef //. -rewrite joinC in V; move: (validR V)=>Vf; case E: (unitb f); last first. -- by apply: H3 (H Vf) (K _ _) V; case: lastkeyP Vf E. -move/unitbP: {K H} E V (H3 k v Unit H1)=>->. -by rewrite unitR validPt lastkey0 lt0n=>V; apply. -Qed. - -(* forward induction on valid natmaps *) -Lemma valid_indf (P : U -> Prop) : - P Unit -> - (forall k v h, P h -> - {in dom h, forall x, k < x} -> - valid (pts k v \+ h) -> P (pts k v \+ h)) -> - forall h, valid h -> P h. -Proof. -move=>H1 H2; elim/um_indf=>[||k v f H V K _]; rewrite ?valid_undef //. -by apply: H2 (H (validR V)) _ V; apply/allP/(order_path_min _ K). -Qed. - -(* membership *) - -Lemma In_lastkey h k v : (k, v) \In h -> k <= last_key h. -Proof. by move/In_dom/dom_lastkey. Qed. - -Lemma In_lastkeyPtUn h x k v w : - last_key h < k -> (x, w) \In h -> (x, w) \In pts k v \+ h. -Proof. by move=>N H; apply: InR=>//; rewrite lastkeyPtUnV ?(In_valid H). Qed. - -Lemma In_lastkeyPtUn_inv h x k v w : - x <= last_key h -> last_key h < k -> - (x, w) \In pts k v \+ h -> (x, w) \In h. -Proof. -move=>N1 N2 H; case/(InPtUnE _ (In_valid H)): H N2=>//. -by case=><- _; case: ltngtP N1. -Qed. - -End LastkeyConstructors. - -(* last_key and omap_fun -- compositions with omf_subdom/omf_subdom0 *) - -Section LastKeyOmapFun. -Variables (A1 A2 : Type) (U1 : natmap A1) (U2 : natmap A2). -Implicit Type f : omap_fun U1 U2. - -(* upper bound lemmas *) -Lemma odom_lastkey f h k : k \in dom (f h) -> k <= last_key h. -Proof. by move/omf_subdom/dom_lastkey. Qed. -Lemma odom0_lastkey f h k : k \in 0 :: dom (f h) -> k <= last_key h. -Proof. by move/omf_subdom0/dom0_lastkey. Qed. -Lemma lastkey_odom f h k : last_key h < k -> k \notin dom (f h). -Proof. by move/lastkey_dom; apply/(subsetC omf_subdom). Qed. -Lemma lastkey_odom0 f h k : last_key h < k -> k \notin 0 :: dom (f h). -Proof. move/lastkey_dom0; apply/(subsetC omf_subdom0). Qed. - -Lemma olub_lastkeyX f h k v : - (k, v) \In h -> omf f (k, v) -> - (forall w, k < last_key (f h) -> - (last_key (f h), w) \In h -> - omf f (last_key (f h), w) -> False) -> - k = last_key (f h). -Proof. -move=>D1 D2 H; apply: lub_lastkeyX=>[|X]. -- by apply/In_dom_omfX; exists v. -by apply/In_dom_omfX; case=>w []; apply: H X. -Qed. - -(* lub lemmas: combine lastkey_lub/lub_lastkey with In_dom_omfX *) -Lemma lastkey_olub f h k : - (forall x v, (x, v) \In h -> omf f (x, v) -> x <= k) -> - last_key (f h) <= k. -Proof. -move=>H; apply: lastkey_lub=>x /In_dom_omfX [v][X E]. -by apply: H X _; rewrite E. -Qed. - -Lemma olub_lastkey f h k : - k \in dom (f h) -> - (forall x v, k < x -> (x, v) \In h -> omf f (x, v) -> False) -> - k = last_key (f h). -Proof. -move=>D H; apply: lub_lastkey=>// x K; apply/In_dom_omfX; case=>y [X E]. -by apply: H K X _; rewrite E. -Qed. - -(* equivalence lemmas (forward) *) -Lemma lastkey_leq_odom f h x k : last_key h <= k -> x \in dom (f h) -> x <= k. -Proof. by move/lastkey_leq_dom=>H /omf_subdom /H. Qed. -Lemma lastkey_leq_odom0 f h x k : last_key h <= k -> x \in 0 :: dom (f h) -> x <= k. -Proof. by move/lastkey_leq_dom0=>H /omf_subdom0 /H. Qed. -Lemma lastkey_ltn_odom0 f h x k : last_key h < k -> x \in 0 :: dom (f h) -> x < k. -Proof. by move/lastkey_ltn_dom0=>H /omf_subdom0 /H. Qed. -Lemma lastkey_ltn_odom f h x k : last_key h < k -> x \in dom (f h) -> x < k. -Proof. by move/lastkey_ltn_dom=>H /omf_subdom /H. Qed. - -(* equivalence lemms (backward) *) -Lemma odom0_leq_lastkey f h x k : x \in 0 :: dom (f h) -> k <= x -> k <= last_key h. -Proof. by move/omf_subdom0; apply/dom0_leq_lastkey. Qed. -Lemma odom_leq_lastkey f h x k : x \in dom (f h) -> k <= x -> k <= last_key h. -Proof. by move/omf_subdom; apply/dom_leq_lastkey. Qed. -Lemma odom_ltn_lastkey f h x k : x \in dom (f h) -> k < x -> k < last_key h. -Proof. by move/omf_subdom; apply: dom_ltn_lastkey. Qed. -Lemma odom0_ltn_lastkey f h x k : x \in 0 :: dom (f h) -> k < x -> k < last_key h. -Proof. by move/omf_subdom0; apply: dom0_ltn_lastkey. Qed. - -(* other lemmas *) - -Lemma lastkey_omf f h : last_key (f h) <= last_key h. -Proof. by apply: lastkey_mono; apply: omf_subdom. Qed. - -Lemma lastkey_omfT f h k : last_key h <= k -> last_key (f h) <= k. -Proof. by apply/leq_trans/lastkey_omf. Qed. - -Lemma lastkey_omfT' f h k : last_key h < k -> last_key (f h) < k. -Proof. by apply/leq_ltn_trans/lastkey_omf. Qed. - -Lemma lastkey_omfUn f h1 h2 : - valid (h1 \+ h2) -> last_key (f h1 \+ f h2) <= last_key (h1 \+ h2). -Proof. by move=>V; rewrite -omfUn // lastkey_omf. Qed. - -Lemma lastkey_omf_some f h : - (forall x, x \In h -> oapp predT false (omf f x)) -> - last_key (f h) = last_key h. -Proof. by rewrite /last_key=>/dom_omf_some ->. Qed. - -Lemma lastkey_omfE f h : last_key h \in dom (f h) -> last_key (f h) = last_key h. -Proof. by apply: lastkey_subdom omf_subdom. Qed. - -Lemma lastkey_omfPtUn f k v h : - last_key h < k -> - f (pts k v \+ h) = - if omf f (k, v) is Some w then pts k w \+ f h else f h. -Proof. -case: (normalP h)=>[->|V N]; last by rewrite omfPtUn lastkeyPtUnV // V. -by case: (omf _ _)=>[a|]; rewrite pfundef !join_undef pfundef. -Qed. - -Lemma omf_sublastkey f k v h : - last_key h < k -> - {subset dom (f h) <= dom (f (pts k v \+ h))}. -Proof. -move=>N x H; rewrite lastkey_omfPtUn //; case: (omf _ _)=>[a|//]. -by rewrite domPtUn inE H orbT lastkeyPtUnV ?(dom_valid H,lastkey_omfT'). -Qed. - -End LastKeyOmapFun. - -(* lastkey_subdom00 for application of omap functions *) -Lemma lastkey00E A A1 A2 (U : natmap A) (U1 : natmap A1) (U2 : natmap A2) - (f1 : omap_fun U U1) (f2 : omap_fun U U2) (h1 : U) (h2 : U) : - (forall x, x \in dom h1 -> x \in dom h2 -> False) -> - last_key (f1 h1) = last_key (f2 h2) -> - (last_key (f1 h1) = 0) * (last_key (f2 h2) = 0). -Proof. by move=>H; apply: lastkey_subdom00 H omf_subdom omf_subdom. Qed. - -(* formulation with application and indexing *) -(* NOTE: not a consequence of lastkey_subdom00E *) -Lemma lastkeyFE T A A1 A2 (U : natmap A) (g : nat -> T -> U) - (U1 : natmap A1) (f1 : omap_fun U U1) - (U2 : natmap A2) (f2 : omap_fun U U2) - (h1 h2 : T) t1 t2 : - (forall x, x \in dom (g t1 h1) -> x \in dom (g t2 h2) -> t1 = t2) -> - last_key (f1 (g t1 h1)) != 0 -> - last_key (f1 (g t1 h1)) = last_key (f2 (g t2 h2)) -> - t1 = t2. -Proof. -move=>X N1 E; apply: contraNeq N1=>/eqP N2. -by move/lastkey00E: E=>-> // x D1 /(X x D1)/N2. -Qed. - -(* formulation with application of composition and with indexing *) -Lemma lastkeyFFE T A A1 A2 A1' A2' (U : natmap A) (g : nat -> T -> U) - (U1 : natmap A1) (f1 : omap_fun U U1) - (U2 : natmap A2) (f2 : omap_fun U U2) - (U1' : natmap A1') (f1' : omap_fun U1 U1') - (U2' : natmap A2') (f2' : omap_fun U2 U2') - (h1 h2 : T) t1 t2 : - (forall x, x \in dom (g t1 h1) -> x \in dom (g t2 h2) -> t1 = t2) -> - last_key (f1' (f1 (g t1 h1))) != 0 -> - last_key (f1' (f1 (g t1 h1))) = last_key (f2' (f2 (g t2 h2))) -> - t1 = t2. -Proof. by rewrite (compE f1') (compE f2'); apply: lastkeyFE. Qed. - -(* last_key and filter *) - -Lemma lastkey_umfilt_le A (U : natmap A) (h : U) k : - last_key h <= k -> um_filterk [pred x | x <= k] h = h. -Proof. -move=>N; rewrite -[RHS]umfilt_predT -eq_in_umfilt. -by case=>x v /In_dom/dom_lastkey/leq_trans/(_ N). -Qed. - -Lemma lastkey_umfilt_lt A (U : natmap A) (h : U) k : - last_key h < k -> um_filterk [pred x | x < k] h = h. -Proof. -move=>N; rewrite -[RHS]umfilt_predT -eq_in_umfilt. -by case=>x v /In_dom/dom_lastkey/leq_ltn_trans/(_ N). -Qed. - -Lemma lastkey_umfiltPtUn A (U : natmap A) (p : pred (nat * A)) (h : U) k v : - last_key h < k -> - um_filter p (pts k v \+ h) = - if p (k, v) then pts k v \+ um_filter p h else um_filter p h. -Proof. by move=>V; rewrite lastkey_omfPtUn ?omf_omap //=; case: ifP. Qed. - -(* last_key and side *) - -Lemma lastkey_sidePtUn (T : eqType) (Us : T -> Type) - (U : natmap (sigT Us)) (Ut : forall t, natmap (Us t)) - (h : U) t k v : - fresh h <= k -> - side_map Ut t (pts k v \+ h) = - if decP (t =P tag v) is left pf then - pts k (cast Us pf (tagged v)) \+ side_map Ut t h - else side_map Ut t h. -Proof. by case: v=>tx vx N; rewrite lastkey_omfPtUn //= /omfx/=; case: eqP. Qed. - -Lemma lastkey_dom_sidePtUn (T : eqType) (Us : T -> Type) - (U : natmap (sigT Us)) (Ut : forall t, natmap (Us t)) - (h : U) t k v : - last_key h < k -> - dom (side_map Ut t (pts k v \+ h)) = - if valid h then - if t == tag v then rcons (dom (side_map Ut t h)) k - else dom (side_map Ut t h) - else [::]. -Proof. -case: (normalP h)=>[->|V N]; first by rewrite join_undef pfundef dom_undef. -rewrite lastkey_sidePtUn //; case: eqP=>//= ?; subst t. -by rewrite eqc lastkey_domPtUn ?(pfV (side_map Ut _)) ?lastkey_omfT'. -Qed. - -(* last_key and non-omap morphisms *) -Lemma plastkey (A : Type) (U1 : pcm) (U2 : natmap A) - (f : pcm_morph U1 U2) k (x y : U1) : - valid (x \+ y) -> - sep f x y -> - last_key (f (x \+ y)) < k -> - last_key (f x) < k. -Proof. -move=>V D; apply: leq_ltn_trans. -by rewrite pfjoin // lastkeyUnL // pfV2. -Qed. - -Lemma plastkeyT A (U1 : pcm) (U2 : natmap A) - (f : full_pcm_morph U1 U2) k (x y : U1) : - valid (x \+ y) -> - last_key (f (x \+ y)) < k -> - last_key (f x) < k. -Proof. by move=>V; apply/(plastkey V)/pfT. Qed. - -(* monotonicity for natmap nat *) -Lemma lastkey_monoPtUn (U : natmap nat) k w (h : U) : - last_key h < k -> - (forall k v, (k, v) \In h -> v < w) -> - um_mono (pts k w \+ h) = um_mono h. -Proof. -move=>N H; case: (normalP h)=>[->|V]; first by rewrite join_undef. -apply: ummonoPtUn; first by rewrite lastkeyPtUnV. -move=>x v X; split; last by apply: H X. -by apply: leq_ltn_trans (dom_lastkey (In_dom X)) N. -Qed. - -(* last_key and pcm ordering *) -Lemma lastkey_fun A (U1 : pcm) (U2 : natmap A) - (f : full_pcm_morph U1 U2) (x y : U1) : - [pcm y <= x] -> - valid x -> - last_key (f y) <= last_key (f x). -Proof. by case=>z -> V; rewrite pfjoinT // lastkeyUnL // pfV2I. Qed. - -(* special case of lastkey_fun when f = id *) -Lemma lastkey_umpleq A (U : natmap A) (h1 h2 : U) : - [pcm h1 <= h2] -> - valid h2 -> - last_key h1 <= last_key h2. -Proof. by apply: (lastkey_fun idfun). Qed. - -Lemma lastkeyL A (U1 : pcm) (U2 : natmap A) - (f : full_pcm_morph U1 U2) (x y : U1) : - valid (x \+ y) -> - last_key (f x) <= last_key (f (x \+ y)). -Proof. by apply: lastkey_fun. Qed. - -Lemma lastkeyR A (U1 : pcm) (U2 : natmap A) - (f : full_pcm_morph U1 U2) (x y : U1) : - valid (x \+ y) -> - last_key (f y) <= last_key (f (x \+ y)). -Proof. by apply: lastkey_fun. Qed. - -(*********) -(* fresh *) -(*********) - -(* mostly new names for lastkey lemmas *) -(* but there are several new lemmas also *) -(* to be systematic, we list all *) - -Section FreshLemmas. -Variables (A : Type) (U : natmap A). -Implicit Type h : U. - -(* fresh is strict upper bound *) -Lemma dom_fresh h k : k \in dom h -> k < fresh h. -Proof. exact: dom_lastkey. Qed. - -(* variants with 0 :: dom and contrapositives *) -Lemma dom0_fresh h k : k \in 0 :: dom h -> k < fresh h. -Proof. exact: dom0_lastkey. Qed. -Lemma fresh_dom h k : fresh h <= k -> k \notin dom h. -Proof. exact: lastkey_dom. Qed. -Lemma fresh_dom0 h k : fresh h <= k -> k \notin 0 :: dom h. -Proof. exact: lastkey_dom0. Qed. - -(* equivalences (similar to lastkey, but with x.+1) *) -Lemma fresh_leq h k : (fresh h <= k) = all (fun x => x.+1 <= k) (0 :: dom h). -Proof. exact: lastkey_ltn. Qed. - -Lemma leq_fresh h k : (k <= fresh h) = has (fun x => k <= x.+1) (0 :: dom h). -Proof. -apply/idP/hasP=>[|[x] D H]; last by apply: leq_trans H (dom0_fresh D). -by exists (last_key h)=>//; apply: lastkey_mem0. -Qed. - -Lemma ltn_fresh h k : (k < fresh h) = has (fun x => k < x.+1) (0 :: dom h). -Proof. by rewrite ltnS leq_lastkey. Qed. - -Lemma fresh_ltn h k : (fresh h < k) = all (fun x => x.+1 < k) (0 :: dom h). -Proof. -apply/idP/allP=>[H x D|]; last by apply; apply: lastkey_mem0. -by apply: leq_ltn_trans (dom0_fresh D) H. -Qed. - -(* unfolding equivalences into implications (forward) *) -Lemma fresh_leq_dom0 h x k : fresh h <= k -> x \in 0 :: dom h -> x < k. -Proof. exact: lastkey_ltn_dom0. Qed. -Lemma fresh_leq_dom h x k : fresh h <= k -> x \in dom h -> x < k. -Proof. exact: lastkey_ltn_dom. Qed. -Lemma fresh_ltn_dom0 h x k : fresh h < k -> x \in 0 :: dom h -> x.+1 < k. -Proof. by rewrite fresh_ltn=>/allP; apply. Qed. -Lemma fresh_ltn_dom h x k : fresh h < k -> x \in dom h -> x.+1 < k. -Proof. by move=>H D; apply: fresh_ltn_dom0 H _; rewrite inE D orbT. Qed. - -(* unfolding equivalences into implications (backward) *) -Lemma dom0_leq_fresh h x k : x \in 0 :: dom h -> k <= x.+1 -> k <= fresh h. -Proof. by move=>D H; rewrite leq_fresh; apply/hasP; exists x. Qed. -Lemma dom_leq_fresh h x k : x \in dom h -> k <= x.+1 -> k <= fresh h. -Proof. by move=>D; apply: dom0_leq_fresh; rewrite inE D orbT. Qed. -Lemma dom0_ltn_fresh h x k : x \in 0 :: dom h -> k <= x -> k < fresh h. -Proof. exact: dom0_leq_lastkey. Qed. -Lemma dom_ltn_fresh h x k : x \in dom h -> k <= x -> k < fresh h. -Proof. exact: dom_leq_lastkey. Qed. -End FreshLemmas. - -(* interaction of two fresh's *) - -Lemma fresh_mono A1 A2 (U1 : natmap A1) (U2 : natmap A2) - (h1 : U1) (h2 : U2) : - {subset dom h1 <= dom h2} -> - fresh h1 <= fresh h2. -Proof. exact: lastkey_mono. Qed. - -Lemma fresh_subdom A1 A2 (U1 : natmap A1) (U2 : natmap A2) - (h1 : U1) (h2 : U2) : - {subset dom h1 <= dom h2} -> - fresh h2 \notin dom h1. -Proof. by move/fresh_mono/fresh_dom. Qed. - -Lemma fresh_subdom_lt A1 A2 (U1 : natmap A1) (U2 : natmap A2) - (h1 : U1) (h2 : U2) x : - {subset dom h1 <= dom h2} -> - x \in dom h1 -> - x < fresh h2. -Proof. by move=>H /H /dom_fresh. Qed. - -Lemma fresh_subdomN A1 A2 (U1 : natmap A1) (U2 : natmap A2) - (h1 : U1) (h2 : U2) x : - {subset dom h1 <= dom h2} -> - x \in dom h1 -> - x != fresh h2. -Proof. by move=>S /(fresh_subdom_lt S); case: ltngtP. Qed. - -Lemma freshUnL A (U : natmap A) (h1 h2 : U) : - valid (h1 \+ h2) -> - fresh h1 <= fresh (h1 \+ h2). -Proof. exact: lastkeyUnL. Qed. - -Lemma freshUnLT A (U : natmap A) (h1 h2 : U) k : - valid (h1 \+ h2) -> - fresh (h1 \+ h2) <= k -> - fresh h1 <= k. -Proof. exact: lastkeyUnLT. Qed. - -Lemma freshUnR A (U : natmap A) (h1 h2 : U) : - valid (h1 \+ h2) -> - fresh h2 <= fresh (h1 \+ h2). -Proof. exact: lastkeyUnR. Qed. - -Lemma freshUnRT A (U : natmap A) (h1 h2 : U) k : - valid (h1 \+ h2) -> - fresh (h1 \+ h2) <= k -> - fresh h2 <= k. -Proof. exact: lastkeyUnRT. Qed. - -(* interaction with constructors *) - -Section FreshConstructors. -Variables (A : Type) (U : natmap A). -Implicit Type h : U. - -Lemma fresh_find h k : - fresh h <= k -> find k h = None. -Proof. exact: lastkey_find. Qed. - -Lemma freshF h : - last_key h != 0 -> - fresh (free h (last_key h)) < fresh h. -Proof. exact: lastkeyF. Qed. - -Lemma freshUn h1 h2 : - fresh (h1 \+ h2) = - if valid (h1 \+ h2) then maxn (fresh h1) (fresh h2) else 1. -Proof. by rewrite /fresh lastkeyUn maxnSS; case: ifP. Qed. - -Lemma freshPtUnV h k v : - fresh h <= k -> - valid (pts k v \+ h) = valid h. -Proof. exact: lastkeyPtUnV. Qed. - -(* easier name to remember *) -Definition valid_fresh := freshPtUnV. - -Lemma fresh_domPtUn h k v : - valid h -> - fresh h <= k -> - dom (pts k v \+ h) = rcons (dom h) k. -Proof. exact: lastkey_domPtUn. Qed. - -Lemma fresh_rangePtUn h k v : - valid h -> - fresh h <= k -> - range (pts k v \+ h) = rcons (range h) v. -Proof. exact: lastkey_rangePtUn. Qed. - -Lemma freshPtUnE h k v : - fresh (pts k v \+ h) = - if valid (pts k v \+ h) then maxn k.+1 (fresh h) else 1. -Proof. by rewrite /fresh lastkeyPtUnE maxnSS; case: ifP. Qed. - -Lemma freshPtUn h k v : - valid h -> fresh h <= k -> fresh (pts k v \+ h) = k.+1. -Proof. by move=>V N; rewrite /fresh lastkeyPtUn. Qed. - -Lemma freshPtUnEX v2 v1 h k : fresh (pts k v1 \+ h) = fresh (pts k v2 \+ h). -Proof. by congr (_.+1); apply: lastkeyPtUnEX. Qed. - -Lemma freshUnPtEX v2 v1 h k : fresh (h \+ pts k v1) = fresh (h \+ pts k v2). -Proof. by rewrite !(joinC h) (freshPtUnEX v2). Qed. - -Lemma freshU h k v : k \in dom h -> fresh (upd k v h) = fresh h. -Proof. by move=>D; congr _.+1; apply: lastkeyU D. Qed. - -Lemma freshUE h k v : - fresh (upd k v h) = - if k \in dom h then fresh h else - if valid (upd k v h) then maxn k.+1 (fresh h) else 1. -Proof. by rewrite /fresh lastkeyUE maxnSS; case: ifP=>//; case: ifP. Qed. - -(* membership *) - -Lemma In_fresh h k v : (k, v) \In h -> k < fresh h. -Proof. exact: In_lastkey. Qed. - -Lemma In_freshPtUn h x k v w : - fresh h <= k -> (x, w) \In h -> (x, w) \In pts k v \+ h. -Proof. exact: In_lastkeyPtUn. Qed. - -Lemma In_freshPtUn_inv h x k v w : - x < fresh h -> fresh h <= k -> - (x, w) \In pts k v \+ h -> (x, w) \In h. -Proof. exact: In_lastkeyPtUn_inv. Qed. - -End FreshConstructors. - -(* fresh and omap_fun -- compositions with omf_subdom/omf_subdom0 *) -Section FreshOmapFun. -Variables (A1 A2 : Type) (U1 : natmap A1) (U2 : natmap A2). -Implicit Types (f : omap_fun U1 U2) (h : U1). - -(* strict upper bound lemmas *) -Lemma odom_fresh f h k : k \in dom (f h) -> k < fresh h. -Proof. exact: odom_lastkey. Qed. -Lemma odom0_fresh f h k : k \in 0 :: dom (f h) -> k < fresh h. -Proof. exact: odom0_lastkey. Qed. -Lemma fresh_odom f h k : fresh h <= k -> k \notin dom (f h). -Proof. exact: lastkey_odom. Qed. -Lemma fresh_odom0 f h k : fresh h <= k -> k \notin 0 :: dom (f h). -Proof. exact: lastkey_odom0. Qed. - -(* equivalence lemmas (forward) *) -Lemma fresh_leq_odom0 f h x k : fresh h <= k -> x \in 0 :: dom (f h) -> x < k. -Proof. exact: lastkey_ltn_odom0. Qed. -Lemma fresh_leq_odom f h x k : fresh h <= k -> x \in dom (f h) -> x < k. -Proof. exact: lastkey_ltn_odom. Qed. -Lemma fresh_ltn_odom0 f h x k : fresh h < k -> x \in 0 :: dom (f h) -> x.+1 < k. -Proof. by move/fresh_ltn_dom0=>H /omf_subdom0 /H. Qed. -Lemma fresh_ltn_odom f h x k : fresh h < k -> x \in dom (f h) -> x.+1 < k. -Proof. by move/fresh_ltn_dom=>H /omf_subdom /H. Qed. - -(* equivalence lemmas (backward) *) -Lemma odom0_leq_fresh f h x k : x \in 0 :: dom (f h) -> k <= x.+1 -> k <= fresh h. -Proof. by move/omf_subdom0; apply/dom0_leq_fresh. Qed. -Lemma odom_leq_fresh f h x k : x \in dom (f h) -> k <= x.+1 -> k <= fresh h. -Proof. by move/omf_subdom; apply/dom_leq_fresh. Qed. -Lemma odom0_ltn_fresh f h x k : x \in 0 :: dom (f h) -> k <= x -> k < fresh h. -Proof. exact: odom0_leq_lastkey. Qed. -Lemma odom_ltn_fresh f h x k : x \in dom (f h) -> k <= x -> k < fresh h. -Proof. exact: odom_leq_lastkey. Qed. - -(* other lemmas *) - -Lemma fresh_omf f h : fresh (f h) <= fresh h. -Proof. exact: lastkey_omf. Qed. - -Lemma fresh_omfT f k h : - fresh h <= k -> - fresh (f h) <= k. -Proof. exact: lastkey_omfT'. Qed. - -Lemma fresh_omfUn f h1 h2 : - valid (h1 \+ h2) -> - fresh (f h1 \+ f h2) <= fresh (h1 \+ h2). -Proof. exact: lastkey_omfUn. Qed. - -Lemma fresh_omf_some f h : - (forall x, x \In h -> oapp predT false (omf f x)) -> - fresh (f h) = fresh h. -Proof. by move=>H; congr _.+1; apply: lastkey_omf_some H. Qed. - -Lemma fresh_omfE f h : - last_key h \in dom (f h) -> - fresh (f h) = fresh h. -Proof. by move=>H; congr _.+1; apply: lastkey_omfE. Qed. - -Lemma fresh_omfPtUn f k v h : - fresh h <= k -> - f (pts k v \+ h) = - if omf f (k, v) is Some w then pts k w \+ f h else f h. -Proof. exact: lastkey_omfPtUn. Qed. - -Lemma omf_subfresh f k v h : - fresh h <= k -> - {subset dom (f h) <= dom (f (pts k v \+ h))}. -Proof. exact: omf_sublastkey. Qed. - -End FreshOmapFun. - -(* fresh and filter *) - -Lemma fresh_umfilt_lt A (U : natmap A) (h : U) k : - fresh h <= k -> - um_filterk [pred x | x < k] h = h. -Proof. exact: lastkey_umfilt_lt. Qed. - -Lemma fresh_umfiltPtUn A (U : natmap A) (p : pred (nat * A)) (h : U) k v : - fresh h <= k -> - um_filter p (pts k v \+ h) = - if p (k, v) then pts k v \+ um_filter p h else um_filter p h. -Proof. exact: lastkey_umfiltPtUn. Qed. - -(* fresh and side *) - -Lemma fresh_sidePtUn (T : eqType) (Us : T -> Type) - (U : natmap (sigT Us)) (Ut : forall t, natmap (Us t)) - (h : U) t k v : - fresh h <= k -> - side_map Ut t (pts k v \+ h) = - if decP (t =P tag v) is left pf then - pts k (cast Us pf (tagged v)) \+ side_map Ut t h - else side_map Ut t h. -Proof. exact: lastkey_sidePtUn. Qed. - -Lemma fresh_dom_sidePtUn (T : eqType) (Us : T -> Type) - (U : natmap (sigT Us)) (Ut : forall t, natmap (Us t)) - (h : U) t k v : - fresh h <= k -> - dom (side_map Ut t (pts k v \+ h)) = - if valid h then - if t == tag v then rcons (dom (side_map Ut t h)) k - else dom (side_map Ut t h) - else [::]. -Proof. exact: lastkey_dom_sidePtUn. Qed. - -(* fresh and non-omap morphisms *) -Lemma pfresh (A : Type) (U1 : pcm) (U2 : natmap A) - (f : pcm_morph U1 U2) k (x y : U1) : - valid (x \+ y) -> - sep f x y -> - fresh (f (x \+ y)) <= k -> - fresh (f x) <= k. -Proof. exact: plastkey. Qed. - -Lemma pfreshT (A : Type) (U1 : pcm) (U2 : natmap A) - (f : full_pcm_morph U1 U2) k (x y : U1) : - valid (x \+ y) -> - fresh (f (x \+ y)) <= k -> - fresh (f x) <= k. -Proof. exact: plastkeyT. Qed. - -(* monotonicity for natmap nat *) -Lemma fresh_monoPtUn (U : natmap nat) k w (h : U) : - fresh h <= k -> - (forall k v, (k, v) \In h -> v < w) -> - um_mono (pts k w \+ h) = um_mono h. -Proof. exact: lastkey_monoPtUn. Qed. - -(* fresh and pcm ordering *) -Lemma fresh_fun A (U1 : pcm) (U2 : natmap A) - (f : full_pcm_morph U1 U2) (x y : U1) : - [pcm y <= x] -> - valid x -> - fresh (f y) <= fresh (f x). -Proof. exact: lastkey_fun. Qed. - -(* special case of fresh_fun when f = id *) -Lemma fresh_umpleq A (U : natmap A) (h1 h2 : U) : - [pcm h1 <= h2] -> - valid h2 -> - fresh h1 <= fresh h2. -Proof. exact: lastkey_umpleq. Qed. - -Lemma freshL A (U1 : pcm) (U2 : natmap A) - (f : full_pcm_morph U1 U2) (x y : U1) : - valid (x \+ y) -> - fresh (f x) <= fresh (f (x \+ y)). -Proof. exact: lastkeyL. Qed. - -Lemma freshR A (U1 : pcm) (U2 : natmap A) - (f : full_pcm_morph U1 U2) (x y : U1) : - valid (x \+ y) -> - fresh (f y) <= fresh (f (x \+ y)). -Proof. by apply: lastkeyR. Qed. - -(***********************************************) -(***********************************************) -(* Executing up to (and including/excluding) t *) -(***********************************************) -(***********************************************) - -Definition oexec_le V (U : natmap V) R (a : R -> V -> R) ks t (h : U) z0 := - oevalv a (&=ks `]-oo, t]) h z0. -Definition oexec_lt V (U : natmap V) R (a : R -> V -> R) ks t (h : U) z0 := - oevalv a (&=ks `]-oo, t[) h z0. - -(* explicit rewrites to facilitate folding *) -Lemma oexleP V U R a ks t h z0 : - oevalv a (&=ks `]-oo, t]) h z0 = @oexec_le V U R a ks t h z0. -Proof. by []. Qed. -Lemma oexltP V U R a ks t h z0 : - oevalv a (&=ks `]-oo, t[) h z0 = @oexec_lt V U R a ks t h z0. -Proof. by []. Qed. - -(* Main transfer lemmas *) - -Lemma oexleE V (U : natmap V) R a t v (h : U) ks (z0 : R) : - 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=>K H; rewrite /oexec_le/oexec_lt eqsl_uxR K (oev_rconsP _ _ _ H). -Qed. - -Arguments oexleE [V U R a t v h ks z0]. - -Lemma oexleNE V (U : natmap V) R a t (h : U) ks (z0 : R) : - (t \notin ks) || (t \notin dom h) -> - oexec_le a ks t h z0 = oexec_lt a ks t h z0. -Proof. -rewrite /oexec_le/oexec_lt; case K: (t \in ks)=>/= H; last first. -- rewrite (eqsl_uoxx (t1:=t) (t2:=t)); last by exact: sle_refl. - by rewrite eqsl_kk1 /= K cats0. -rewrite [LHS]oevFK [RHS]oevFK; congr oeval. -by rewrite eqsl_uxR K filter_rcons (negbTE H). -Qed. - -Arguments oexleNE [V U R a t h ks z0]. - -(**********************************************************) -(* Interaction of oexec_lt and oexec_le with constructors *) -(**********************************************************) - -(* DEVCOMMENT *) -(* -Lemma oexlt0 V (U : natmap V) R a ks (h : U) (z0 : R) : oexec_lt a ks 0 h z0 = z0. -Proof. by rewrite /oexec_lt squo0. Qed. - -Lemma oexle0 V R a ks (h : natmap V) (z0 : R) : oexec_le a ks 0 h z0 = z0. -Proof. -rewrite /oexec_le squx0; case: ifP=>//= _. -set xs := filter _ _; rewrite oevFK; set ys := filter _ _. -rewrite (_ : ys = [::]) //. -rewrite -[RHS](filter_pred0 xs); apply: eq_in_filter. -by move=>x; rewrite mem_filter=>/andP [/eqP ->]; rewrite cond_dom. -Qed. -*) -(* /DEVCOMMENT *) - -Lemma oexlt_notin V (U : natmap V) R a ks t (h : U) (z0 : R) : - t \notin ks -> - oexec_lt a ks t h z0 = oevalv a ks h z0. -Proof. by move=>K; rewrite /oexec_lt eqsl_uL_notinE. Qed. - -Arguments oexlt_notin [V U R a ks t h z0]. - -Lemma oexle_notin V (U : natmap V) R a ks t (h : U) (z0 : R) : - t \notin ks -> - oexec_le a ks t h z0 = oevalv a ks h z0. -Proof. by move=>K; rewrite /oexec_le eqsl_uL_notinE. Qed. - -Arguments oexle_notin [V U R a ks t h z0]. - -Lemma oexlt_cat V (U : natmap V) R a ks1 ks2 t (h : U) (z0 : R) : - t \notin ks1 -> - 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. -move=>T; rewrite /oexec_lt eqsl_uL_catE (negbTE T). -by rewrite oev_cat (eqsl_uL_notinE (s:=ks1)). -Qed. - -Arguments oexlt_cat [V U R a ks1 ks2 t h z0]. - -Lemma oexle_cat V (U : natmap V) R a ks1 ks2 t (h : U) (z0 : R) : - t \notin ks1 -> - 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. -move=>T; rewrite /oexec_le eqsl_uL_catE (negbTE T). -by rewrite oev_cat (eqsl_uL_notinE (s:=ks1)). -Qed. - -Arguments oexle_cat [V U R a ks1 ks2 t h z0]. - -Lemma oexlt_cons V (U : natmap V) R a ks k t v (h : U) (z0 : R) : - (k, v) \In h -> t != k -> - oexec_lt a (k :: ks) t h z0 = oexec_lt a ks t h (a z0 v). -Proof. -move=>H Ntk; rewrite /oexec_lt eqsl_uL_consE (negbTE Ntk) /=. -by move/In_find: H=>->. -Qed. - -Arguments oexlt_cons [V U R a ks k t v h z0]. - -Lemma oexle_cons V (U : natmap V) R a ks k t v (h : U) (z0 : R) : - (k, v) \In h -> - oexec_le a (k :: ks) t h z0 = - if t == k then a z0 v else oexec_le a ks t h (a z0 v). -Proof. -move=>H; rewrite /oexec_le eqsl_uL_consE. -by case: eqP=>/=; move/In_find: H=>->. -Qed. - -Arguments oexle_cons [V U R a ks k t v h z0]. - -(* for oexlt, we need to cover the case when k == t *) -Lemma oexlt_cons_same V (U : natmap V) R a ks k (h : U) (z0 : R) : - oexec_lt a (k :: ks) k h z0 = z0. -Proof. by rewrite /oexec_lt eqsl_uL_consL. Qed. - -Arguments oexlt_cons_same {V U R a ks k h z0}. - -Lemma oexlt_cons_notin V (U : natmap V) R a ks k t (h : U) (z0 : R) : - k \notin dom h -> t != k -> - oexec_lt a (k :: ks) t h z0 = oexec_lt a ks t h z0. -Proof. -move=>H Ntk; rewrite /oexec_lt eqsl_uL_consE (negbTE Ntk) /=. -by case: dom_find H=>//= -> _. -Qed. - -Arguments oexlt_cons_notin [V U R a ks k t h z0]. - -(* for oexle, we have two "notin" lemmas *) -Lemma oexle_cons_same_notin V (U : natmap V) R a ks k (h : U) (z0 : R) : - k \notin dom h -> oexec_le a (k :: ks) k h z0 = z0. -Proof. -move=>H. -by rewrite /oexec_le [in LHS]oevFK eqsl_uL_consL /= (negbTE H). -Qed. - -Arguments oexle_cons_same_notin [V U R a ks k h z0]. - -Lemma oexle_cons_notin V (U : natmap V) R a ks k t (h : U) (z0 : R) : - k \notin dom h -> t != k -> - oexec_le a (k :: ks) t h z0 = oexec_le a ks t h z0. -Proof. -move=>H Ntk; rewrite /oexec_le [in LHS]oevFK [in RHS]oevFK. -by rewrite eqsl_uL_consE (negbTE Ntk) /= (negbTE H). -Qed. - -Arguments oexle_cons_notin [V U R a ks k t h z0]. - -Lemma oexlt_rcons V (U : natmap V) R a ks k t (h : U) (z0 : R) : - t \in ks -> - oexec_lt a (rcons ks k) t h z0 = oexec_lt a ks t h z0. -Proof. by move=>T; rewrite /oexec_lt eqsl_uL_rconsE T. Qed. - -Arguments oexlt_rcons [V U R a ks k t h z0]. - -Lemma oexle_rcons V (U : natmap V) R a ks k t (h : U) (z0 : R) : - t \in ks -> - oexec_le a (rcons ks k) t h z0 = oexec_le a ks t h z0. -Proof. by move=>T; rewrite /oexec_le eqsl_uL_rconsE T. Qed. - -Arguments oexle_rcons [V U R a ks k t h z0]. - -Lemma oexlt_rcons_notin V (U : natmap V) R a ks k t v (h : U) (z0 : R) : - (k, v) \In h -> t \notin ks -> t != k -> - oexec_lt a (rcons ks k) t h z0 = - a (oexec_lt a ks t h z0) v. -Proof. -move=>H T Ntk; rewrite /oexec_lt eqsl_uL_rconsE /= (negbTE T) Ntk. -by rewrite oev_rconsE eqsl_uL_notinE //; move/In_find: H=>->. -Qed. - -Arguments oexlt_rcons_notin [V U R a ks k t v h z0]. - -Lemma oexle_rcons_notin V (U : natmap V) R a ks k t (h : U) (z0 : R) : - t \notin ks -> - oexec_le a (rcons ks k) t h z0 = oevalv a (rcons ks k) h z0. -Proof. by move=>T; rewrite /oexec_le eqsl_uL_rconsE (negbTE T). Qed. - -Arguments oexle_rcons_notin [V U R a ks k t h z0]. - -(* in case of oexlt we must cover the case when t == k *) -Lemma oexlt_rcons_same V (U : natmap V) R a ks k (h : U) (z0 : R) : - k \notin ks -> - oexec_lt a (rcons ks k) k h z0 = oevalv a ks h z0. -Proof. by move=>N; rewrite /oexec_lt eqsl_uL_rconsE eqxx /= (negbTE N). Qed. - -Arguments oexlt_rcons_same [V U R a ks k h z0]. - -(* in case of oexle, case t == k can be optimized *) -(* DEVCOMMENT *) -(* TODO doesn't simplify anything now, remove? *) -(* /DEVCOMMENT *) -Lemma oexle_rcons_same V (U : natmap V) R a ks k (h : U) (z0 : R) : - k \notin ks -> - oexec_le a (rcons ks k) k h z0 = oevalv a (rcons ks k) h z0. -Proof. exact: oexle_rcons_notin. Qed. - -Arguments oexle_rcons_same [V U R a ks k h z0]. - -(* in case of oexle, when we know the value at k *) -(* we can further expand the right-hand side *) -Lemma oexle_rcons_notin_in V (U : natmap V) R a ks k t v (h : U) (z0 : R) : - (k, v) \In h -> t \notin ks -> - oexec_le a (rcons ks k) t h z0 = a (oexec_le a ks t h z0) v. -Proof. -move=>H T; rewrite oexle_rcons_notin // oev_rconsE. -by move/In_find: (H)=>->; rewrite oexleNE ?T // /oexec_lt eqsl_uL_notinE. -Qed. - -Arguments oexle_rcons_notin_in [V U R a ks k t v h z0]. - -(* oexlt/oexle filter lemmas *) - -Lemma oexlt_umfilt V (U : natmap V) R a ks p t (h : U) (z0 : R) : - oexec_lt a ks t (um_filterk p h) z0 = - oevalv a (filter p (&=ks `]-oo, t[)) h z0. -Proof. by rewrite /oexec_lt oev_filter. Qed. - -Arguments oexlt_umfilt {V U R a ks p t h z0}. - -Lemma oexle_umfilt V (U : natmap V) R a ks p t (h : U) (z0 : R) : - oexec_le a ks t (um_filterk p h) z0 = - oevalv a (filter p (&=ks `]-oo, t])) h z0. -Proof. by rewrite /oexec_le oev_filter. Qed. - -Arguments oexle_umfilt {V U R a ks p t h z0}. - -(* two somewhat simpler rewrites under a condition *) -Lemma oexlt_umfiltN V (U : natmap V) R a ks p t (h : U) (z0 : R) : - (t \notin ks) || (p t) -> - oexec_lt a ks t (um_filterk p h) z0 = - oexec_lt a (filter p ks) t h z0. -Proof. -move=>H; rewrite /oexec_lt oev_umfilt eqsl_filterL //=. -apply: eq_in_oevK; rewrite -!filter_predI; apply: eq_in_filter=>k Ks /=. -by case D : (k \in dom h)=>//=; case/In_domX: D=>v /In_find ->. -Qed. - -Arguments oexlt_umfiltN [V U R a ks p t h z0]. - -Lemma oexle_umfiltN V (U : natmap V) R a ks p t (h : U) (z0 : R) : - (t \notin ks) || (p t) -> - oexec_le a ks t (um_filterk p h) z0 = - oexec_le a (filter p ks) t h z0. -Proof. -move=>H; rewrite /oexec_le oev_umfilt eqsl_filterL //=. -apply: eq_in_oevK; rewrite -!filter_predI; apply: eq_in_filter=>k Ks /=. -by case D : (k \in dom h)=>//=; case/In_domX: D=>v /In_find ->. -Qed. - -Arguments oexle_umfiltN [V U R a ks p t h z0]. - -(* restating the last two lemmas for the other direction *) -(* DEVCOMMENT *) -(* TODO why not just state them like this initially? *) -(* /DEVCOMMENT *) -Lemma oexlt_filter V (U : natmap V) R a ks p t (h : U) (z0 : R) : - (t \notin ks) || (p t) -> - oexec_lt a (filter p ks) t h z0 = - oexec_lt a ks t (um_filterk p h) z0. -Proof. by move=>H; rewrite (oexlt_umfiltN H). Qed. - -Lemma oexle_filter V (U : natmap V) R a ks p t (h : U) (z0 : R) : - (t \notin ks) || (p t) -> - oexec_le a (filter p ks) t h z0 = - oexec_le a ks t (um_filterk p h) z0. -Proof. by move=>H; rewrite (oexle_umfiltN H). Qed. - -(* interaction with the map *) - -Lemma oexlt_filter_dom V (U : natmap V) R a ks t (h : U) (z0 : R) : - (t \notin ks) || (t \in dom h) -> - oexec_lt a ks t h z0 = - oexec_lt a (filter (mem (dom h)) ks) t h z0. -Proof. by move=>H; rewrite oexlt_filter // umfiltk_dom'. Qed. - -Lemma oexle_filter_dom V (U : natmap V) R a ks t (h : U) (z0 : R) : - (t \notin ks) || (t \in dom h) -> - oexec_le a ks t h z0 = - oexec_le a (filter (mem (dom h)) ks) t h z0. -Proof. by move=>H; rewrite oexle_filter // umfiltk_dom'. Qed. - -(* interaction of oexlt, oexle and last *) - -Lemma oexlt_oexle_last V (U : natmap V) R a k ks t (h : U) (z0 : R) : - uniq (k::ks) -> t != k -> - oexec_lt a (k::ks) t h z0 = - oexec_le a (k::ks) (last k (&=ks `]-oo, t[)) h z0. -Proof. -move=>Uq Ntk; rewrite /oexec_lt/oexec_le [LHS]oevFK [RHS]oevFK. -by rewrite eqsl_uox_last. -Qed. - -Lemma oev_oexle_last V (U : natmap V) R a k ks (h : U) (z0 : R) : - uniq (k::ks) -> - oevalv a (k::ks) h z0 = - oexec_le a (k::ks) (last k ks) h z0. -Proof. -case/andP=>U1 U2; rewrite /oexec_le eqsl_uL_consE. -case: eqP=>/= [/last_nochange|/eqP/last_change H]. -- by rewrite (negbTE U1)=>/eqP->. -by congr oeval; apply/eqsl_lastR_uniq. -Qed. - -(*******************************************************) -(* oexlt/oexle on endpoints of intervals of invariance *) -(*******************************************************) - -(* these can be called "induction" lemmas over the interval *) - -Definition oex_inv V (U : natmap V) R (P : R -> Prop) a ks (h : U) (z0 : R) := - forall k v, (k, v) \In h -> - let z := oexec_lt a ks k h z0 in P z -> P (a z v). - -(* The naming schema in these lemmas is based on backward reasoning *) -(* E.g., the suffix ox means we're proving an lt_fact (o) from le_fact (x) *) - -(* two-sided intervals *) - -Lemma oex_ox V (U : natmap V) R (P : R -> Prop) a ks (h : U) t1 t2 (z0 : R) : - uniq ks -> t1 <[ks] t2 -> - {in &=ks `]t1, t2[, oex_inv P a ks h z0} -> - P (oexec_le a ks t1 h z0) -> P (oexec_lt a ks t2 h z0). -Proof. -move=>Uq T IH P0; rewrite /oexec_lt (eqsl_uxoo T) oev_cat. -apply: oev_indX=>// k ks1 ks2 v _ Kh Eh ->; rewrite -oev_cat. -have K : k \in &=ks `]t1, t2[ by rewrite Eh mem_cat inE eqxx orbT. -have Nk : k \notin ks1. -- move/(eqslice_uniq `]t1, t2[): Uq. - by rewrite Eh cat_uniq /= negb_or -andbA; case/and5P. -case/mem_oo: (K)=>// [_ T1K T2K]. -suff {IH Uq K}-> : ks1 = &=ks `]t1, k[ by rewrite -eqsl_uxoo //; apply: IH. -move: Eh; rewrite (eqslice_split (b:=true) (x:=k)) /=; last first. -- rewrite !lteBSide /= leEnat -seqlt_unlockE -seqle_unlock. - by rewrite T1K (sltW T2K). -rewrite eqsl_xoL T2K /= => Eh; rewrite (cat_cancel _ _ Eh) //. -by apply: eqsliceRO_notin. -Qed. - -(* one can try to derive the next lemma from the previous *) -(* but we need to reason about successor of t1 in ks *) -(* which we don't have yet. Hence, we prove the lemma directly *) -(* using the same approach as in the previous lemma. *) - -Lemma oex_oo V (U : natmap V) R (P : R -> Prop) a ks (h : U) t1 t2 (z0 : R) : - uniq ks -> t1 <=[ks] t2 -> - {in &=ks `[t1, t2[, oex_inv P a ks h z0} -> - P (oexec_lt a ks t1 h z0) -> P (oexec_lt a ks t2 h z0). -Proof. -move=>Uq T IH P0; rewrite /oexec_lt (eqsl_uoxo T) oev_cat. -apply: oev_indX=>// k ks1 ks2 v _ Kh Eh ->; rewrite -oev_cat. -have K : k \in &=ks `[t1, t2[ by rewrite Eh mem_cat inE eqxx orbT. -have Nk : k \notin ks1. -- move/(eqslice_uniq `[t1, t2[): Uq. - by rewrite Eh cat_uniq /= negb_or -andbA; case/and5P. -case/mem_xo: (K)=>// [_ T1K T2K]. -suff {IH Uq K}-> : ks1 = &=ks `[t1, k[ by rewrite -eqsl_uoxo //; apply: IH. -move: Eh; rewrite (eqslice_split (b:=true) (x:=k)) /=; last first. -- by rewrite !lteBSide /= !leEnat -!seqle_unlock T1K (sltW T2K). -rewrite (eqsl_xoL k) T2K /= => Eh; rewrite (cat_cancel _ _ Eh) //. -by apply: eqsliceRO_notin. -Qed. - -Lemma oex_xo V (U : natmap V) R (P : R -> Prop) a ks (h : U) t1 t2 (z0 : R) : - uniq ks -> t1 <=[ks] t2 -> - {in &=ks `[t1, t2], oex_inv P a ks h z0} -> - P (oexec_lt a ks t1 h z0) -> P (oexec_le a ks t2 h z0). -Proof. -move=>Uq T IH P0; rewrite /oexec_le (eqsl_uoxx T) oev_cat. -apply: oev_indX=>// k ks1 ks2 v _ Kh Eh ->; rewrite -oev_cat. -have K : k \in &=ks `[t1, t2] by rewrite Eh mem_cat inE eqxx orbT. -have Nk : k \notin ks1. -- move/(eqslice_uniq `[t1, t2]): Uq. - by rewrite Eh cat_uniq /= negb_or -andbA; case/and5P. -case/mem_xx: (K)=>// [Ks T1K T2K]. -suff {IH Uq K}-> : ks1 = &=ks `[t1, k[ by rewrite -eqsl_uoxo //; apply: IH. -move: Eh; rewrite (eqslice_split (b:=true) (x:=k)) /=; last first. -- by rewrite /order.Order.le/=/order.Order.le/= -!seqle_unlock T1K T2K. -rewrite eqsl_xxL T2K Ks /= => Eh; rewrite (cat_cancel _ _ Eh) //. -by apply: eqsliceRO_notin. -Qed. - -Lemma oex_xx V (U : natmap V) R (P : R -> Prop) a ks (h : U) t1 t2 (z0 : R) : - uniq ks -> t1 <=[ks] t2 -> - {in &=ks `]t1, t2], oex_inv P a ks h z0} -> - P (oexec_le a ks t1 h z0) -> P (oexec_le a ks t2 h z0). -Proof. -move=>Uq T IH P0; rewrite /oexec_le (eqsl_uxox T) oev_cat. -apply: oev_indX=>// k ks1 ks2 v _ Kh Eh ->; rewrite -oev_cat. -have K : k \in &=ks `]t1, t2] by rewrite Eh mem_cat inE eqxx orbT. -have Nk : k \notin ks1. -- move/(eqslice_uniq `]t1, t2]): Uq. - by rewrite Eh cat_uniq /= negb_or -andbA; case/and5P. -case/mem_ox: (K)=>// [Ks T1K T2K]. -suff {IH Uq K}-> : ks1 = &=ks `]t1, k[ by rewrite -eqsl_uxoo //; apply: IH. -move: Eh; rewrite (eqslice_split (b:=true) (x:=k)) /=; last first. -- rewrite /order.Order.le/=/order.Order.le/=/order.Order.lt /=. - by rewrite -seqlt_unlock -seqle_unlock T1K T2K. -rewrite eqsl_xxL T2K Ks /= => Eh; rewrite (cat_cancel _ _ Eh) //. -by apply: eqsliceRO_notin. -Qed. - -Arguments oex_ox [V U R] P [a ks h t1 t2 z0]. -Arguments oex_oo [V U R] P [a ks h t1 t2 z0]. -Arguments oex_xo [V U R] P [a ks h t1 t2 z0]. -Arguments oex_xx [V U R] P [a ks h t1 t2 z0]. - -(* one-sided intervals *) - -Lemma oex_ux V (U : natmap V) R (P : R -> Prop) a ks (h : U) t (z0 : R) : - uniq ks -> - {in &=ks `]t,+oo[, oex_inv P a ks h z0} -> - P (oexec_le a ks t h z0) -> P (oevalv a ks h z0). -Proof. -move=>Uq IH P0; rewrite (eqsl_uxou ks t) oev_cat. -apply: oev_indX=>// k ks1 ks2 v _ Kh Eh ->; rewrite -oev_cat. -have K : k \in &=ks `]t, +oo[ by rewrite Eh mem_cat inE eqxx orbT. -have Nk : k \notin ks1. -- move/(eqslice_uniq `]t, +oo[): Uq. - by rewrite Eh cat_uniq /= negb_or -andbA; case/and5P. -case/mem_ou: (K)=>// [Ks TK]. -suff {IH Uq K}-> : ks1 = &=ks `]t, k[ by rewrite -eqsl_uxoo //; apply: IH. -move: Eh; rewrite (eqslice_split (b:=true) (x:=k)) /=; last first. -- by rewrite /order.Order.le/=/order.Order.lt/= -seqlt_unlock TK. -rewrite eqsl_xuL Ks /= => Eh; rewrite (cat_cancel _ _ Eh) //. -by apply: eqsliceRO_notin. -Qed. - -Lemma oex_uo V (U : natmap V) R (P : R -> Prop) a ks (h : U) t (z0 : R) : - uniq ks -> - {in &=ks `[t,+oo[, oex_inv P a ks h z0} -> - P (oexec_lt a ks t h z0) -> P (oevalv a ks h z0). -Proof. -move=>Uq IH P0; rewrite (eqsl_uoxu ks t) oev_cat. -apply: oev_indX=>// k ks1 ks2 v _ Kh Eh ->; rewrite -oev_cat. -have K : k \in &=ks `[t, +oo[ by rewrite Eh mem_cat inE eqxx orbT. -have Nk : k \notin ks1. -- move/(eqslice_uniq `[t, +oo[): Uq. - by rewrite Eh cat_uniq /= negb_or -andbA; case/and5P. -case/mem_xu: (K)=>// [Ks TK]. -suff {IH Uq K}-> : ks1 = &=ks `[t, k[ by rewrite -eqsl_uoxo //; apply: IH. -move: Eh; rewrite (eqslice_split (b:=true) (x:=k)) /=; last first. -- by rewrite /order.Order.le/=/order.Order.le/= -seqle_unlock TK. -rewrite eqsl_xuL Ks => Eh; rewrite (cat_cancel _ _ Eh) //. -by apply: eqsliceRO_notin. -Qed. - -Lemma oex_xu V (U : natmap V) R (P : R -> Prop) a ks (h : U) t (z0 : R) : - uniq ks -> - {in &=ks `]-oo, t], oex_inv P a ks h z0} -> - P z0 -> P (oexec_le a ks t h z0). -Proof. -move=>Uq IH P0; apply: oev_indX=>// k ks1 ks2 v _ Kh Eh ->. -have K : k \in &=ks `]-oo, t] by rewrite Eh mem_cat inE eqxx orbT. -have Nk : k \notin ks1. -- move/(eqslice_uniq `]-oo, t]): Uq. - by rewrite Eh cat_uniq /= negb_or -andbA; case/and5P. -case/mem_ux: (K)=>// [Ks TK]. -suff {IH Uq K}-> : ks1 = &=ks `]-oo, k[ by apply: IH. -move: Eh; rewrite (eqslice_split (b:=true) (x:=k)) //=; last first. -- by rewrite /order.Order.le/=/order.Order.le/= -seqle_unlock. -rewrite eqsl_xxL TK Ks /= => Eh; rewrite (cat_cancel _ _ Eh) //. -by apply: eqsliceRO_notin. -Qed. - -Lemma oex_ou V (U : natmap V) R (P : R -> Prop) a ks (h : U) t (z0 : R) : - uniq ks -> - {in &=ks `]-oo, t[, oex_inv P a ks h z0} -> - P z0 -> P (oexec_lt a ks t h z0). -Proof. -move=>Uq IH P0; apply: oev_indX=>// k ks1 ks2 v _ Kh Eh ->. -have K : k \in &=ks `]-oo, t[ by rewrite Eh mem_cat inE eqxx orbT. -have Nk : k \notin ks1. -- move/(eqslice_uniq `]-oo, t[): Uq. - by rewrite Eh cat_uniq /= negb_or -andbA; case/and5P. -case/mem_uo: (K)=>// [Ks TK]. -suff {IH Uq K}-> : ks1 = &=ks `]-oo, k[ by apply: IH. -move: Eh; rewrite (eqslice_split (b:=true) (x:=k)) /=; last first. -- by rewrite lteBSide /= leEnat -seqle_unlock (sltW TK). -rewrite eqsl_xoL TK => Eh; rewrite (cat_cancel _ _ Eh) //. -by apply: eqsliceRO_notin. -Qed. - -Arguments oex_ux [V U R] P [a ks h t z0]. -Arguments oex_uo [V U R] P [a ks h t z0]. -Arguments oex_xu [V U R] P [a ks h t z0]. -Arguments oex_ou [V U R] P [a ks h t z0]. - -(* whole list *) - -Lemma oex_uu V (U : natmap V) R (P : R -> Prop) a ks (h : U) (z0 : R) : - uniq ks -> - {in ks, oex_inv P a ks h z0} -> - P z0 -> P (oevalv a ks h z0). -Proof. -move=>Uq IH P0; apply: oev_indX=>// k ks1 ks2 v _ Kh Eh ->. -have K : k \in ks by rewrite Eh mem_cat inE eqxx orbT. -have Nk : k \notin ks1. -- move: Uq. - by rewrite Eh cat_uniq /= negb_or -andbA; case/and5P. -suff -> : ks1 = &=ks `]-oo, k[ by apply: IH. -move: Eh; rewrite {1}(eqsl_uxou ks k) eqsl_uxR K -cats1 -catA /= => Eh. -by rewrite (cat_cancel _ _ Eh) //=; apply: eqsliceRO_notin. -Qed. - -Arguments oex_uu [V U R] P [a ks h z0]. - -(**********************************************) -(* functional versions of the interval lemmas *) -(**********************************************) - -Definition oexF_inv V (U : natmap V) R X (f : R -> X) a ks (h : U) (z0 : R) := - forall k v, (k, v) \In h -> - let z := oexec_lt a ks k h z0 in f (a z v) = f z. - -(* two-sided intervals *) - -Lemma oexF_ox V (U : natmap V) R X (f : R -> X) a ks (h : U) t1 t2 (z0 : R) : - uniq ks -> t1 <[ks] t2 -> - {in &=ks `]t1, t2[, oexF_inv f a ks h z0} -> - f (oexec_lt a ks t2 h z0) = f (oexec_le a ks t1 h z0). -Proof. -move=>Uq T H; pose P := fun x => f x = f (oexec_le a ks t1 h z0). -by apply: (oex_ox P) T _ _=>// x S v K E ; rewrite /P -E; apply: H. -Qed. - -Lemma oexF_oo V (U : natmap V) R X (f : R -> X) a ks (h : U) t1 t2 (z0 : R) : - uniq ks -> t1 <=[ks] t2 -> - {in &=ks `[t1, t2[, oexF_inv f a ks h z0} -> - f (oexec_lt a ks t2 h z0) = f (oexec_lt a ks t1 h z0). -Proof. -move=>Uq T H; pose P := fun x => f x = f (oexec_lt a ks t1 h z0). -by apply: (oex_oo P) T _ _=>// x S v K E; rewrite /P -E; apply: H. -Qed. - -Lemma oexF_xo V (U : natmap V) R X (f : R -> X) a ks (h : U) t1 t2 (z0 : R) : - uniq ks -> t1 <=[ks] t2 -> - {in &=ks `[t1, t2], oexF_inv f a ks h z0} -> - f (oexec_le a ks t2 h z0) = f (oexec_lt a ks t1 h z0). -Proof. -move=>Uq T H; pose P := fun x => f x = f (oexec_lt a ks t1 h z0). -by apply: (oex_xo P) T _ _=>// x S v K E; rewrite /P -E; apply: H. -Qed. - -Lemma oexF_xx V (U : natmap V) R X (f : R -> X) a ks (h : U) t1 t2 (z0 : R) : - uniq ks -> t1 <=[ks] t2 -> - {in &=ks `]t1, t2], oexF_inv f a ks h z0} -> - f (oexec_le a ks t2 h z0) = f (oexec_le a ks t1 h z0). -Proof. -move=>Uq T H; pose P := fun x => f x = f (oexec_le a ks t1 h z0). -by apply: (oex_xx P) T _ _=>// x S v K E; rewrite /P -E; apply: H. -Qed. - -Arguments oexF_ox [V U R X] f [a ks h t1 t2 z0]. -Arguments oexF_oo [V U R X] f [a ks h t1 t2 z0]. -Arguments oexF_xo [V U R X] f [a ks h t1 t2 z0]. -Arguments oexF_xx [V U R X] f [a ks h t1 t2 z0]. - -(* one-sided intervals *) - -Lemma oexF_ux V (U : natmap V) R X (f : R -> X) a ks (h : U) t (z0 : R) : - uniq ks -> - {in &=ks `]t, +oo[, oexF_inv f a ks h z0} -> - f (oevalv a ks h z0) = f (oexec_le a ks t h z0). -Proof. -move=>Uq H; pose P := fun x => f x = f (oexec_le a ks t h z0). -by apply: (oex_ux P)=>// x S v K E; rewrite /P -E; apply: H. -Qed. - -Lemma oexF_uo V (U : natmap V) R X (f : R -> X) a ks (h : U) t (z0 : R) : - uniq ks -> - {in &=ks `[t, +oo[, oexF_inv f a ks h z0} -> - f (oevalv a ks h z0) = f (oexec_lt a ks t h z0). -Proof. -move=>Uq H; pose P := fun x => f x = f (oexec_lt a ks t h z0). -by apply: (oex_uo P)=>// x S v K E; rewrite /P -E; apply: H. -Qed. - -Lemma oexF_xu V (U : natmap V) R X (f : R -> X) a ks (h : U) t (z0 : R) : - uniq ks -> - {in &=ks `]-oo, t], oexF_inv f a ks h z0} -> - f (oexec_le a ks t h z0) = f z0. -Proof. -move=>Uq H; pose P := fun x => f x = f z0. -by apply: (oex_xu P)=>// x S v K E; rewrite /P -E; apply: H. -Qed. - -Lemma oexF_ou V (U : natmap V) R X (f : R -> X) a ks (h : U) t (z0 : R) : - uniq ks -> - {in &=ks `]-oo, t[, oexF_inv f a ks h z0} -> - f (oexec_lt a ks t h z0) = f z0. -Proof. -move=>Uq H; pose P := fun x => f x = f z0. -by apply: (oex_ou P)=>// x S v K E; rewrite /P -E; apply: H. -Qed. - -Arguments oexF_ux [V U R X] f [a ks h t z0]. -Arguments oexF_uo [V U R X] f [a ks h t z0]. -Arguments oexF_xu [V U R X] f [a ks h t z0]. -Arguments oexF_ou [V U R X] f [a ks h t z0]. - -(* whole list *) - -Lemma oexF_uu V (U : natmap V) R X (f : R -> X) a ks (h : U) (z0 : R) : - uniq ks -> - {in ks, oexF_inv f a ks h z0} -> - f (oevalv a ks h z0) = f z0. -Proof. -move=>Uq H; pose P := fun x => f x = f z0. -by apply: (oex_uu P)=>// x S v K E; rewrite /P -E; apply: H. -Qed. - -Arguments oexF_uu [V U R X] f [a ks h z0]. - -(* we can now combine the split lemmas with *) -(* cons properties of oeval, to obtain a point-split *) -(* when the middle point is in the map *) - -Lemma oev2_split V (U : natmap V) R a t1 v (h : U) ks (z0 : R) : - t1 \in ks -> (t1, v) \In h -> - oevalv a ks h z0 = - oevalv a (&=ks `]t1, +oo[) h (a (oexec_lt a ks t1 h z0) v). -Proof. -move=>D H; rewrite {1}(eqsl_uoxu ks t1) oev_cat. -by rewrite eqsl_xuL D /=; move/In_find: H=>->. -Qed. - -Arguments oev2_split [V U R a t1 v h ks z0] _ _. - -Lemma oex2_split V (U : natmap V) R a t1 t2 v (h : U) ks (z0 : R) : - t1 <[ks] t2 -> (t1, v) \In h -> - oexec_lt a ks t2 h z0 = - oevalv a (&=ks `]t1, t2[) h (a (oexec_lt a ks t1 h z0) v). -Proof. -move=>T H; rewrite /oexec_lt. -rewrite (eqsl_uoxo (sltW T)) oev_cat eqsl_xoL T /=. -by move/In_find: H=>->. -Qed. - -Arguments oex2_split [V U R a t1 t2 v h ks z0] _ _. - -(* we frequently iterate oex2_split, so the following saves verbiage *) -Lemma oex3_split V (U : natmap V) R a t1 t2 t3 v1 v2 (h : U) ks (z0 : R) : - t1 <[ks] t2 -> (t1, v1) \In h -> - t2 <[ks] t3 -> (t2, v2) \In h -> - oexec_lt a ks t3 h z0 = - oevalv a (&=ks `]t2, t3[) h - (a (oevalv a (&=ks `]t1, t2[) h - (a (oexec_lt a ks t1 h z0) v1)) - v2). -Proof. -move=>T1 H1 T2 H2. -by rewrite (oex2_split T2 H2) (oex2_split T1 H1). -Qed. - -Arguments oex3_split [V U R a t1 t2 t3 v1 v2 h ks z0] _ _ _ _. - -(******************************************) -(* Interaction of consec with oexlt/oexle *) -(******************************************) - -Lemma oexlt_consec V (U : natmap V) R a ks t1 t2 (h : U) (z0 : R) : - uniq ks -> - consec ks t1 t2 -> - oexec_lt a ks t2 h z0 = oexec_le a ks t1 h z0. -Proof. -move=>Uq C; apply: (oexF_ox id)=>//; first by apply: consec_slt. -by move=>x; rewrite consec_oo. -Qed. - -Arguments oexlt_consec [V U R a ks t1 t2 h z0]. - -(* the following is direct *) -Lemma oexlt_split V (U : natmap V) R a x1 t1 t2 x2 (h : U) (z0 : R) : - uniq (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=>Uq; apply: oexlt_consec=>//; apply/consecP_split=>//. -- by rewrite mem_cat !inE eqxx !orbT. -by exists x1, x2. -Qed. - -Lemma oexlt_consec_in V (U : natmap V) R a t1 t2 v (h : U) ks (z0 : R) : - uniq ks -> - (t1, v) \In h -> - consec ks t1 t2 -> - oexec_lt a ks t2 h z0 = a (oexec_lt a ks t1 h z0) v. -Proof. -move=>Uq H C; move/slt_memE: (consec_slt C)=>T. -by rewrite (oexlt_consec Uq C) (oexleE T H). -Qed. - -Lemma oexlt_consec_fst V (U : natmap V) R a t (h : U) k ks (z0 : R) : - uniq (k::ks) -> k \notin dom h -> t \in k::ks -> - consec (k::ks) k t -> - oexec_lt a (k::ks) t h z0 = z0. -Proof. -move=>Uq K T C; case/(consecP_split _ Uq T): C=>xs1 [xs2] X. -case: xs1 X Uq {T}=>[|x xs1]; last first. -- by case=>->-> /=; rewrite mem_cat inE eqxx !orbT. -case=>->{ks}; rewrite /= inE negb_or -andbA; case/and4P=>U1 U2 U3 U4. -rewrite oexlt_cons_notin ?inE 1?negb_or ?(eq_sym t) ?U1 ?U2 //. -by rewrite oexlt_cons_same. -Qed. - -Lemma oexlt_consec_find V (U : natmap V) R a t1 t2 (h : U) ks (z0 : R) : - uniq ks -> - consec ks t1 t2 -> - oexec_lt a ks t2 h z0 = - if find t1 h is Some e - then a (oexec_lt a ks t1 h z0) e - else oexec_lt a ks t1 h z0. -Proof. -move=>Uq C; rewrite (oexlt_consec Uq C). -case E : (find t1 h)=>[e|]; first by move/In_find: E=>/(oexleE (consec_mem C)) ->. -by rewrite oexleNE // orbC; move/In_findN: E=>->. -Qed. - -Arguments oexlt_consec_find [V U R a t1 t2 h ks z0]. - -Lemma oexlt_last V (U : natmap V) R a t2 (h : U) ks (z0 : R) : - uniq ks -> - t2 \notin ks -> - oexec_lt a ks t2 h z0 = - if ks is k :: ks' then - if last k ks' \in dom h then - oexec_le a ks (last k ks') h z0 - else oexec_lt a ks (last k ks') h z0 - else z0. -Proof. -case: ks=>[|k ks] Uq //= T2. -have C : consec (k :: ks) (last k ks) t2. -- by rewrite (consecP_lastE k) // mem_last /=. -rewrite (oexlt_consec_find Uq C). -case: dom_find=>// v /In_find E _. -by rewrite -oexleE // mem_last. -Qed. - -Lemma oexlt_last0 V (U : natmap V) R a t2 (h : U) ks (z0 : R) : - uniq (0::ks) -> - t2 \notin 0::ks -> - oexec_lt a ks t2 h z0 = - if last 0 ks \in dom h then - oexec_le a ks (last 0 ks) h z0 - else oexec_lt a ks (last 0 ks) h z0. -Proof. -move=>Uq T2. -have D : 0 \notin dom h by rewrite cond_dom. -have N : t2 != 0 by case: eqP T2=>// ->. -rewrite -(oexlt_cons_notin D N) oexlt_last //. -case: ifP=>L. -- by rewrite oexle_cons_notin // (dom_cond L). -case: ks Uq {N T2 D L}=>[|k ks] //= Uq. -rewrite oexlt_cons_notin ?cond_dom //. -apply/negP=>/eqP N. -by rewrite -N (mem_last k ks) in Uq. -Qed. - -(* The lemmas past this point are currently used for some examples, *) -(* but will be deprecated and removed in future releases *) - -(*******************) -(*******************) -(* Gapless natmaps *) -(*******************) -(*******************) - -Section Gapless. -Variables (A : Type) (U : natmap A). -Implicit Type h : U. - -Definition gapless h := forall k, 0 < k <= last_key h -> k \in dom h. - -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. -Proof. by move=>k; rewrite lastkey0; case: k. Qed. - -Lemma gp_fresh h u : gapless (pts (fresh h) u \+ h) <-> gapless h. -Proof. -case V : (valid h); last first. -- by move/invalidE: (negbT V)=>->; rewrite join_undef. -split=>H k; move: (V); rewrite -(freshPtUnV u (leqnn _))=>V'; last first. -- rewrite lastkeyPtUn // domPtUn inE V' /= (leq_eqVlt k) eq_sym. - by move: (H k); rewrite -(ltnS k); case: (ltngtP k (last_key h).+1). -rewrite -(ltnS k) -/(fresh h); case/andP=>Z N. -move: (H k); rewrite lastkeyPtUn // domPtUn inE V' Z /= leq_eqVlt eq_sym. -by case: ltngtP N=>//= _ _; apply. -Qed. - -Lemma gpPtUn h k u : - valid (pts k u \+ h) -> - gapless (pts k u \+ h) -> last_key h < k -> k = fresh h. -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 (pts k u \+ h) = k. -- move: (C (fresh h)); rewrite Z (leq_eqVlt _ k) Y orbT andbT; move/(_ (erefl _)). - by rewrite domPtUn inE (negbTE X) /=; case/andP=>_ /dom_fresh; rewrite ltnn. -by rewrite lastkeyPtUn // (validR V). -Qed. - -Lemma gaplessPtUnE u2 u1 h k : - gapless (pts k u1 \+ h) <-> gapless (pts k u2 \+ h). -Proof. -rewrite /gapless (lastkeyPtUnEX 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 U][h] u. - -(*****************************) -(*****************************) -(* Natmaps with binary range *) -(*****************************) -(*****************************) - -Section AtVal. -Variables (A : Type) (U : natmap (A * A)). -Implicit Type h : U. - -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; case/(gapless_pleq G V): {V} H (V)=>h [->{h2} H] V. -by apply: atvalUn. -Qed. - -Definition last_atval v h := atval v (last_key h) h. - -Lemma lastatval0 v : last_atval v Unit = v. -Proof. by rewrite /last_atval /atval lastkey0 find0E. Qed. - -Lemma lastatval_undef v : last_atval v undef = v. -Proof. by rewrite /last_atval /atval lastkey_undef find_undef. Qed. - -Lemma lastatvalPt v p x : p != 0 -> last_atval v (pts p x) = x.2. -Proof. -by move=>V; rewrite /last_atval /atval lastkeyPt // findPt /= V. -Qed. - -Lemma lastatval_fresh v x h : - valid h -> last_atval v (pts (fresh h) x \+ h) = x.2. -Proof. -by move=>V; rewrite /last_atval /atval lastkeyPtUn // findPtUn // freshPtUnV. -Qed. - -Lemma lastatvalUn v h1 h2 : - last_atval v (h1 \+ h2) = - if valid (h1 \+ h2) then - if last_key h1 < last_key h2 then last_atval v h2 else last_atval v h1 - else v. -Proof. -rewrite /last_atval /atval lastkeyUn maxnC /maxn. -case: ifP; last by move/negbT/invalidE=>->; rewrite find_undef. -case: (ltngtP (last_key h1) (last_key h2)) => N V. -- by rewrite findUnR //; case: lastkeyP N. -- by rewrite findUnL //; case: lastkeyP N. -by rewrite !(lastkeyUn0 V N) unitL lastkey0 find0E. -Qed. - -Lemma lastatval_freshUn v x h1 h2 : - valid h1 -> [pcm h2 <= h1] -> - last_atval v (pts (fresh h1) x \+ h2) = x.2. -Proof. -move=>V H; move: (pleqV H V)=>V2; move: (fresh_mono (umpleq_subdom V H))=>N. -by rewrite /last_atval /atval lastkeyPtUn // findPtUn // freshPtUnV. -Qed. - -Lemma lastatval_indom v h : - last_atval v h <> v -> last_key h \in dom h. -Proof. by rewrite /last_atval /atval; case: dom_find=>// ->. Qed. - -End AtVal. - -(* in case A = bool *) -Lemma lastatval_indomb (U : natmap (bool*bool)) (h : U) : - last_atval false h -> last_key h \in dom h. -Proof. by move=>H; apply: (lastatval_indom (v:=false)); rewrite H. Qed. - -(* Continuous maps with binary range *) - -Section Continuous. -Variables (A : Type) (U : natmap (A * A)). -Implicit Type h : U. - -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 undef. -Proof. by move=>x w; rewrite find_undef. Qed. - -Lemma cn0 v : continuous v Unit. -Proof. by move=>x w; rewrite find0E. Qed. - -Lemma cn_fresh v h x : - valid h -> - continuous v (pts (fresh h) x \+ h) <-> - continuous v h /\ last_atval v h = x.1. -Proof. -rewrite -(freshPtUnV x (leqnn _))=>V; split; last first. -- case=>C H k y; rewrite !findPtUn2 // eqSS; case: ltngtP=>N. - - by rewrite ltn_eqF; [apply: C | apply: (ltn_trans N _)]. - - by move/find_some/dom_fresh/(ltn_trans N); rewrite ltnn. - by case=><-; rewrite N ltn_eqF. -move=>C; split; last first. -- move: (C (last_key h) x). - by rewrite !findPtUn2 // eqxx ltn_eqF //; apply. -move=>k w; case: (ltnP k (last_key h))=>N; last first. -- by move/find_some/dom_fresh/(leq_ltn_trans N); rewrite ltnn. -by move: (C k w); rewrite !findPtUn2 // eqSS !ltn_eqF // (ltn_trans N _). -Qed. - -Lemma cn_sub v h x y k : - valid (pts k.+1 (x, y) \+ h) -> continuous v (pts k.+1 (x, y) \+ h) -> - oapp snd v (find k h) = x. -Proof. -by move=>V /(_ k (x, y)); rewrite !findPtUn2 // eqxx ltn_eqF //; apply. -Qed. - -End Continuous. - -Arguments cn_fresh [A U][v][h][x] _. - -(* Complete natmaps with binary range *) - -Section Complete. -Variables (A : Type) (U : natmap (A * A)). -Implicit Type h : U. - -Definition complete v0 h vn := - [/\ valid h, gapless h, continuous v0 h & last_atval v0 h = vn]. - -Lemma cm_valid v0 h vn : complete v0 h vn -> valid h. -Proof. by case. Qed. - -Lemma cm0 v : complete v Unit v. -Proof. by split=>//; [apply: gp0|apply: cn0|rewrite lastatval0]. Qed. - -Lemma cm_fresh v0 vn h x : - complete v0 (pts (fresh h) x \+ h) vn <-> vn = x.2 /\ complete v0 h x.1. -Proof. -split. -- by case=>/validR V /gp_fresh G /(cn_fresh V) []; rewrite lastatval_fresh. -case=>-> [V] /(gp_fresh x) G C E; split=>//; -by [rewrite freshPtUnV|apply/(cn_fresh V)| rewrite lastatval_fresh]. -Qed. - -Lemma cmPtUn v0 vn h k x : - complete v0 (pts k x \+ h) vn -> last_key h < k -> k = fresh h. -Proof. by case=>V /(gpPtUn V). Qed. - -Lemma cmPt v0 vn k x : complete v0 (pts k x) vn -> k = 1 /\ x = (v0, vn). -Proof. -case; rewrite validPt; case: k=>//= k _ /(_ 1). -rewrite lastkeyPt //= domPt inE /= => /(_ (erefl _))/eqP ->. -move/(_ 0 x); rewrite findPt findPt2 /= => -> //. -by rewrite /last_atval lastkeyPt // /atval findPt /= => <-; case: x. -Qed. - -Lemma cmCons v0 vn k x h : - complete v0 (pts k x \+ h) vn -> last_key h < k -> - [/\ k = fresh h, vn = x.2 & complete v0 h x.1]. -Proof. by move=>C H; move: {C} (cmPtUn C H) (C)=>-> /cm_fresh []. Qed. - -End Complete. - -Prenex Implicits cm_valid cmPt. - -(************************) -(************************) -(************************) -(* Surgery on histories *) -(* using leq filtering *) -(************************) -(************************) -(************************) - -Notation le t := (fun '(k, _) => k <= t). -Notation lt t := (fun '(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 - (fun kv : (nat * V) => - let '(k, v) := kv in 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 - (fun kv : nat * V => - let '(k, v) := kv in t1 < k < t2). -Proof. by move=>T [k v] /=; rewrite -ltnNge. Qed. - -Lemma lastkey_umfilt_leq A (U : natmap A) (h : U) t : - last_key (um_filterk (leq^~ t) h) <= t. -Proof. -case: (normalP h)=>[->|V]. -- by rewrite pfundef lastkey_undef. -set j := um_filterk _ _. -have V' : valid j by rewrite pfV. -case E : (unitb j); first by move/unitbP: E=>->; rewrite lastkey0. -have : last_key j \in dom j by case: lastkeyP V' E. -by case/In_dom_umfilt. -Qed. - -Lemma lastkey_umfilt_dom A (U : natmap A) (h : U) t : - last_key (um_filterk (leq^~ t) h) <= last_key h. -Proof. by apply: lastkey_mono; apply: omf_subdom. Qed. - -Lemma umfilt_le_last A (U : natmap A) (h : U) t : - last_key h <= t -> - um_filter (le t) h = h. -Proof. -case: (normalP h)=>[->|V]; first by rewrite pfundef. -move=>N; rewrite -{2}[h]umfilt_predT; apply/eq_in_umfilt. -by case=>k v /In_dom/dom_lastkey/leq_trans; apply. -Qed. - -Lemma umfilt_last A (U : natmap A) (h : U) : - um_filter (le (last_key h)) h = h. -Proof. by apply: umfilt_le_last. Qed. - -Lemma umfilt_le_fresh A (U : natmap A) (h : U) : - um_filter (le (fresh h)) h = h. -Proof. by apply: umfilt_le_last; apply: ltnW. Qed. - -Lemma umfilt_le0 A (U : natmap A) (h : U) 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 eqxx. -Qed. - -Lemma umfilt_le_split A (U : natmap A) (h : U) t1 t2 : - t1 <= t2 -> - um_filter (le t2) h = - um_filter (le t1) h \+ um_filter (fun '(k, _) => t1 < k <= 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 (U : natmap A) (h : U) t1 t2 k : - t1 <= k <= t2 -> - um_filter (fun '(x, _)=>t1 < x <= t2) h = - um_filter (fun '(x, _)=>t1 < x <= k) h \+ - um_filter (fun '(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 (U : natmap A) (h : U) k v : - (k, v) \In h -> - um_filter (fun '(x, _)=>k.-1 < x <= k) h = pts 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 (U : natmap A) (h1 h2 : U) 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 (U : natmap A) (h1 h2 : U) 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 (U : natmap A) (h : U) 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 (U : natmap A) R a (h1 h2 : U) 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 (U : natmap A) R a (h1 h2 : U) 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 (U : natmap A) R a (h : U) (z0 : R) : eval a (le 0) h z0 = z0. -Proof. -case: (normalP h)=>[->|V]; first by 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 (U : natmap A) R a (h : U) t1 t2 (z0 : R) : - t1 <= t2 -> - eval a (le t2) h z0 = - eval a (fun '(k, _)=>t1 < k <= t2) h (eval a (le t1) h z0). -Proof. -move=>T; case: (normalP h)=>[->|V]; first by rewrite !eval_undef. -rewrite eval_umfilt (umfilt_predD h (pts_sub T)) evalUn; last first. -- move=>x y /In_dom_umfilt [vx X _] /In_dom_umfilt [wy /= /andP][]. - by rewrite /= -ltnNge; move/(leq_ltn_trans X). -- by rewrite -(umfilt_predD h (pts_sub T)) pfV. -by rewrite -!eval_umfilt; apply: eq_in_eval; case=>k v _; apply: ptsD. -Qed. - -Lemma eval_lt_split A (U : natmap A) R a (h : U) t1 t2 (z0 : R) : - t1 < t2 -> - eval a (lt t2) h z0 = - eval a (fun '(k, _)=>t1 < k < t2) h (eval a (le t1) h z0). -Proof. -move=>T; case: (normalP h)=>[->|V]; first by rewrite !eval_undef. -rewrite eval_umfilt (umfilt_predD h (pts_sub_lt T)) evalUn; last first. -- move=>x y /In_dom_umfilt [vx X _] /In_dom_umfilt [wy /= /andP][]. - by rewrite /= -ltnNge; move/(leq_ltn_trans X). -- by rewrite -(umfilt_predD h (pts_sub_lt T)) pfV. -by rewrite -!eval_umfilt; apply: eq_in_eval; case=>k v _; apply: ptsD_lt. -Qed. - -Lemma eval_le_lt_split A (U : natmap A) R a (h : U) t (z0 : R) : - eval a (le t) h z0 = - eval a (fun '(k, _)=>k == t) h (eval a (lt t) h z0). -Proof. -case: (normalP h)=>[->|V]; first by 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 /In_dom_umfilt [vx X _] /In_dom_umfilt [wy /= /andP][]. - by rewrite /= -ltnNge; move/(leq_ltn_trans X). -- by rewrite -(umfilt_predD h D) pfV. -rewrite -!eval_umfilt; apply: eq_in_eval; case=>k v _ /=. -by case: ltngtP. -Qed. - -Lemma eval_eq A (U : natmap A) R a (h : U) t v (z0 : R) : - (t, v) \In h -> - eval a (fun '(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 (fun '(k, _)=>k == t) h = pts t v by rewrite evalPt /= N. -apply/umem_eq=>[||[k w]]; first by rewrite pfV; case: H. -- by rewrite validPt. -rewrite In_umfiltX; split; last by move/InPt =>[[->->]]. -by move=>[/eqP -> H1]; rewrite (In_fun H H1); apply: In_condPt. -Qed. - -Lemma eval_le_last A (U : natmap A) R a (h : U) 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/dom_lastkey/leq_trans; apply. -Qed. - -Lemma eval_fresh_le A (U : natmap A) R a (h : U) t v (z0 : R) : - eval a (le t) (pts (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: (normalP h)=>[->|V]. -- by rewrite join_undef !eval_undef; case: ifP. -case: ifP=>H. -- by rewrite eval_umfilt umfiltPtUn freshPtUnV // V ltnNge H -eval_umfilt. -rewrite joinC evalUnPt; last first. -- by apply/allP=>x; apply: dom_lastkey. -- by rewrite joinC freshPtUnV. -rewrite ltnNge H; congr a; apply: eq_in_eval. -case=>k w /In_dom/dom_lastkey /=. -by case: ltngtP H=>// /ltnW H _ /leq_trans; apply. -Qed. - -Lemma eval_fresh_lt A (U : natmap A) R a (h : U) t v (z0 : R) : - eval a (lt t) (pts (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: (normalP h)=>[->|V]. -- by 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: dom_lastkey. -- by rewrite joinC valid_fresh. -rewrite ltnNge H; congr a; apply: eq_in_eval. -case=>k w /In_dom/dom_lastkey /=; rewrite /fresh -ltnNge in H. -by case: ltngtP H=>// /ltnW H _ /leq_ltn_trans; apply. -Qed. - -Lemma eval_le_fresh A (U : natmap A) R a (h : U) t v (z0 : R) : - t <= last_key h -> - eval a (le t) (pts (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 (U : natmap A) R a (h : U) t v (z0 : R) : - t <= fresh h -> - eval a (lt t) (pts (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 (U : natmap A) R a (P : R -> Prop) (h : U) 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 (U : natmap A) R a (P : R -> Prop) (h : U) 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 (U : natmap A) R X a (f : R -> X) (h : U) 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 (U : natmap A) R X a (f : R -> X) (h : U) 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 (U : natmap A) p (h : U) t1 t2 : - t1 <= t2 -> - um_count (predI p (le t2)) h = - um_count (predI p (le t1)) h + - um_count (predI p (fun '(k, _)=>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 (U : natmap A) p (h : U) 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 (U : natmap A) p (h1 h2 : U) 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 (U : natmap A) p (h1 h2 : U) 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 (U : natmap A) p (h : U) 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 (U : natmap A) p (h : U) t v : - um_count (predI p (le t)) (pts (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: (normalP h)=>[->|V]. -- by rewrite join_undef !umcnt_undef 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 (U : natmap A) p (h : U) t v : - t <= last_key h -> - um_count (predI p (le t)) (pts (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). - -(*********************************) -(* 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) (U : natmap V). - -Lemma exec0 a (h : U) (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) (U : natmap V) (X : ordType) (a : R -> V -> R) (f : R -> X). -Implicit Types p : pred (nat*V). - -Definition growing (h : U) := - 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 (pts k v \+ h) -> - valid (pts k v \+ h) -> - all (ord k) (dom h) -> - p (k, v) -> - f (evalv a p (pts 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 pts 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 \+ (pts k v \+ h2)) -> - valid (h1 \+ (pts k v \+ h2)) -> - {in dom h1, forall x, x < k} -> - all (ord k) (dom h2) -> - p (k, v) -> - f (evalv a p (h1 \+ (pts 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 ?W // in E1; last first. -- move=>x y /D1 X1; rewrite domPtUn inE (validR 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_umfiltX [_ /G]. -- by rewrite -(umfilt_le_split h N) pfV. -by move=>??/In_dom_umfilt[? /leq_ltn_trans Y _]/In_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 : U := um_filter (le t1) h. -pose h1 : U := um_filter (fun '(x, _)=>t1 < x <= k.-1) h. -pose h2 : U := um_filter (fun '(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 \+ (pts 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 \+ (pts k v \+ h2))) by rewrite -Eh pfV. -rewrite eval_umfilt (umfilt_le_split h K2) evalUn ?(validAL W1) //; last first. -- by move=>??/In_dom_umfilt[?/leq_ltn_trans Y] _ /In_dom_umfilt[?] /andP [/Y]. -rewrite -(eval_umfilt (le t1)); apply: helper2 (validR W1) _ _ _ _ =>//. -- by apply: growR W1 _; rewrite -Eh=>k1 v1 z1 /In_umfiltX [] _ /G. -- by move=>x /In_dom_umfilt; rewrite (subX t1 x) //; case=>v0 /andP []. -- by apply/allP=>x /In_dom_umfilt; case=>v0 /andP []. -rewrite eval_umfilt Eh evalUn ?(validX W1) -?eval_umfilt // in E. -move=>x y /In_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) =>/In_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) (U : natmap V) (X : ordType) (a : R -> V -> R) (f : R -> nat). -Implicit Types p : pred (nat*V). - -Lemma growIn (h : U) 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 : U) 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. diff --git a/pcm/pcm.v b/pcm/pcm.v deleted file mode 100644 index ba2189d..0000000 --- a/pcm/pcm.v +++ /dev/null @@ -1,1660 +0,0 @@ -(* -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 defines pcm -- a type equipped with partial commutative *) -(* monoid structure, several subclasses of PCMs, and important *) -(* concrete instances. *) -(******************************************************************************) - -From HB Require Import structures. -From Coq Require Import ssreflect ssrbool ssrfun Setoid. -From mathcomp Require Import ssrnat eqtype seq bigop fintype finset finfun. -From pcm Require Import options axioms prelude seqperm pred seqext. - -Declare Scope pcm_scope. -Delimit Scope pcm_scope with pcm. -Open Scope pcm_scope. - -(*******************************) -(* Partial Commutative Monoids *) -(*******************************) - -Definition pcm_axiom T (valid : T -> bool) (join : T -> T -> T) - (unit : T) (unitb : T -> bool) := - Prod6 (commutative join) - (associative join) - (left_id unit join) - (* monotonicity of valid *) - (forall x y, valid (join x y) -> valid x) - (* unit is valid *) - (valid unit) - (* (as a matter of principle, we make points decidable) *) - (forall x, reflect (x = unit) (unitb x)). - -HB.mixin Record isPCM T := { - valid : T -> bool; - join : T -> T -> T; - Unit : T; - unitb : T -> bool; - pcm_subproof : pcm_axiom valid join Unit unitb}. - -#[short(type="pcm")] -HB.structure Definition PCM := {T of isPCM T}. - -Infix "\+" := join (at level 43, left associativity) : pcm_scope. - -Arguments Unit {s}. -Arguments valid {s} : simpl never. -Prenex Implicits join Unit. - -Lemma joinC {U} : commutative (@join U). Proof. by case: (@pcm_subproof U). Qed. -Lemma joinA {U} : associative (@join U). Proof. by case: (@pcm_subproof U). Qed. -Lemma unitL {U} : left_id Unit (@join U). Proof. by case: (@pcm_subproof U). Qed. -Lemma valid_unit {U} : valid (@Unit U). Proof. by case: (@pcm_subproof U). Qed. -Lemma validL {U : pcm} {x y : U} : valid (x \+ y) -> valid x. -Proof. by case: (@pcm_subproof U) x y. Qed. -Lemma unitbP {U : pcm} {x : U} : reflect (x = Unit) (unitb x). -Proof. by case: (@pcm_subproof U). Qed. - -Section DerivedLemmas. -Variables U V : pcm. - -Lemma joinAC (x y z : U) : x \+ y \+ z = x \+ z \+ y. -Proof. by rewrite -joinA (joinC y) joinA. Qed. - -Lemma joinCA (x y z : U) : x \+ (y \+ z) = y \+ (x \+ z). -Proof. by rewrite joinA (joinC x) -joinA. Qed. - -Lemma validR (x y : U) : valid (x \+ y) -> valid y. -Proof. by rewrite joinC; apply: validL. 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 unitR (x : U) : x \+ Unit = x. -Proof. by rewrite joinC unitL. Qed. - -(* some helper lemmas for easier extraction of validity claims *) -Lemma validAR (x y z : U) : valid (x \+ y \+ z) -> valid (y \+ z). -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 DerivedLemmas. - -Arguments unitL {U}. -Arguments unitR {U}. - -#[export] -Hint Resolve valid_unit : core. - -Section UnfoldingRules. -Variable U : pcm. - -Lemma pcm_joinE (x y : U) : x \+ y = isPCM.join (PCM.class U) x y. -Proof. by []. Qed. - -Lemma pcm_validE (x : U) : valid x = isPCM.valid (PCM.class U) x. -Proof. by []. Qed. - -Lemma pcm_unitE : Unit = isPCM.Unit (PCM.class U). -Proof. by []. Qed. - -Lemma unitb0 : unitb (Unit : U). -Proof. by case: unitbP. Qed. - -Definition pcmE := (pcm_joinE, pcm_validE, pcm_unitE, unitb0). - -(* also a simple rearrangment equation *) -Definition pull (x y : U) := (joinC y x, joinCA y x). - -End UnfoldingRules. - - -(*********************) -(* Cancellative PCMs *) -(*********************) - -(* predicate precision for arbitrary PCM U *) - -Definition precise (U : pcm) (P : U -> Prop) := - forall s1 s2 t1 t2, - valid (s1 \+ t1) -> P s1 -> P s2 -> - s1 \+ t1 = s2 \+ t2 -> s1 = s2. - -Definition cpcm_axiom (U : pcm) := - forall x1 x2 x : U, - valid (x1 \+ x) -> x1 \+ x = x2 \+ x -> x1 = x2. - -HB.mixin Record isCPCM U of PCM U := { - cpcm_subproof : cpcm_axiom U}. - -#[short(type="cpcm")] -HB.structure Definition CPCM := {U of PCM U & isCPCM U}. - -Lemma joinKx (U : cpcm) {x1 x2 x : U} : - valid (x1 \+ x) -> x1 \+ x = x2 \+ x -> x1 = x2. -Proof. by apply: cpcm_subproof. Qed. - - -Section CPCMLemmas. -Variable U : cpcm. - -Lemma joinxK (x x1 x2 : U) : valid (x \+ x1) -> x \+ x1 = x \+ x2 -> x1 = x2. -Proof. by rewrite !(joinC x); apply: joinKx. Qed. - -Lemma cancPL (P : U -> Prop) s1 s2 t1 t2 : - precise P -> valid (s1 \+ t1) -> P s1 -> P s2 -> - s1 \+ t1 = s2 \+ t2 -> (s1 = s2) * (t1 = t2). -Proof. -move=>H V H1 H2 E; move: (H _ _ _ _ V H1 H2 E)=>X. -by rewrite -X in E *; rewrite (joinxK V E). -Qed. - -Lemma cancPR (P : U -> Prop) s1 s2 t1 t2 : - precise P -> valid (s1 \+ t1) -> P t1 -> P t2 -> - s1 \+ t1 = s2 \+ t2 -> (s1 = s2) * (t1 = t2). -Proof. -move=>H V H1 H2 E; rewrite (joinC s1) (joinC s2) in E V. -by rewrite !(cancPL H V H1 H2 E). -Qed. - -End CPCMLemmas. - - -(***************) -(* Topped PCMs *) -(***************) - -(* PCM with an explicit undef value *) -(* we and undef elemeent and function undefb *) -(* to test decidably if an element is undef *) - -(* DEVCOMMENT *) -(* obsoleted conditions *) -(* _ : 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)]; *) -(* /DEVCOMMENT *) - -Definition tpcm_axiom (U : pcm) (undef : U) - (undefb : U -> bool) := - Prod3 (forall x : U, reflect (x = undef) (undefb x)) - (~~ valid undef) - (forall x : U, undef \+ x = undef). - -HB.mixin Record isTPCM U of PCM U := { - undef : U; - undefb : U -> bool; - tpcm_subproof : tpcm_axiom undef undefb}. - -#[short(type="tpcm")] -HB.structure Definition TPCM := {U of PCM U & isTPCM U}. - -Lemma undefbP (U : tpcm) {x : U} : reflect (x = undef) (undefb x). -Proof. by case: (@tpcm_subproof U). Qed. - -Lemma valid_undefN {U} : ~~ valid (@undef U). -Proof. by case: (@tpcm_subproof U). Qed. - -Lemma undef_join (U : tpcm) (x : U) : undef \+ x = undef. -Proof. by case: (@tpcm_subproof U). Qed. - -Section TPCMLemmas. -Variable U : tpcm. - -Lemma undefbE (f : U) : f = undef <-> undefb f. -Proof. by case: undefbP. Qed. - -Lemma undefb_undef : undefb (undef : U). -Proof. by case: undefbP. Qed. - -Lemma valid_undef : valid (undef : U) = false. -Proof. by rewrite (negbTE valid_undefN). Qed. - -Lemma join_undef (x : U) : x \+ undef = undef. -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. - -Lemma unitb_undef : unitb (undef : U) = false. -Proof. by case: unitbP =>// /undef0. Qed. - -Lemma undefD (x : U) : decidable (x = undef). -Proof. -case D: (undefb x). -- by left; move/undefbP: D. -by right; move/undefbP: D. -Qed. - -End TPCMLemmas. - -Definition tpcmE := (undef_join, join_undef, valid_undef, - unitb_undef, undefb_undef). - - -(***************) -(* Normal TPCM *) -(***************) - -(* TPCM whose only invalid element is undef *) - -Definition normal_tpcm_axiom (U : tpcm) := forall x : U, valid x \/ x = undef. - -HB.mixin Record isNormal_TPCM U of TPCM U := { - normal_tpcm_subproof : normal_tpcm_axiom U}. - -#[short(type="normal_tpcm")] -HB.structure Definition Normal_TPCM := {U of TPCM U & isNormal_TPCM U}. - -(* branching on valid x or x = undef *) -Variant normal_spec {U : normal_tpcm} (x : U) : - bool -> bool -> bool -> Type := -| normal_undef of x = undef : normal_spec x false true false -| normal_valid of valid x : normal_spec x true false (unitb x). - -Lemma normalP {U : normal_tpcm} (x : U) : - normal_spec x (valid x) (undefb x) (unitb x). -Proof. -case: undefbP=>[->|N]. -- by rewrite valid_undef unitb_undef; apply: normal_undef. -have V : (valid x) by case: (normal_tpcm_subproof x) N. -by rewrite V; apply: normal_valid. -Qed. - -Lemma undefNV {U : normal_tpcm} (x : U) : undefb x = ~~ valid x. -Proof. by case: normalP. Qed. - -Lemma undefVN {U : normal_tpcm} (x : U) : valid x = ~~ undefb x. -Proof. by case: normalP. Qed. - -(* branching on valid x or x = undef or x = unit *) -Variant normal0_spec {U : normal_tpcm} (x : U) : - bool -> bool -> bool -> Type := -| normal0_undef of x = undef : normal0_spec x false true false -| normal0_unit of x = Unit : normal0_spec x true false true -| normal0_valid of valid x & x <> Unit : normal0_spec x true false false. - -Lemma normalP0 {U : normal_tpcm} (x : U) : - normal0_spec x (valid x) (undefb x) (unitb x). -Proof. -case: (normalP x)=>[->|V]; first by apply: normal0_undef. -case E: (unitb x); move/unitbP: E; -by [apply: normal0_unit|apply: normal0_valid]. -Qed. - - -(***************************************) -(* PCMs with combination of properties *) -(***************************************) - -(* pcm with decidable equality *) -#[short(type="eqpcm")] -HB.structure Definition EQPCM := {U of Equality U & PCM U}. - -(* cancellative pcm with decidable equality *) -#[short(type="eqcpcm")] -HB.structure Definition EQCPCM := {U of CPCM U & EQPCM U}. - -(* tpcm with decidable equality *) -#[short(type="eqtpcm")] -HB.structure Definition EQTPCM := {U of TPCM U & EQPCM U}. - -(* normal tpcm with decidable equality *) -#[short(type="normal_eqtpcm")] -HB.structure Definition Normal_EQTPCM := {U of Normal_TPCM U & EQPCM U}. - -(* cancellative tpcm *) -#[short(type="ctpcm")] -HB.structure Definition CTPCM := {U of CPCM U & TPCM U}. - -(* cancellative tpcm with decidable equality *) -#[short(type="eqctpcm")] -HB.structure Definition EQCTPCM := {U of Equality U & CTPCM U}. - -(* normal cancellative tpcm *) -#[short(type="normal_ctpcm")] -HB.structure Definition Normal_CTPCM := {U of Normal_TPCM U & CPCM U}. - -(* normal cancellative tpcm with decidable equality *) -#[short(type="normal_eqctpcm")] -HB.structure Definition Normal_EQCTPCM := {U of Normal_TPCM U & EQCPCM U}. - - -(***************************************) -(* Support for big operators over PCMs *) -(***************************************) - -(* U has the laws of commutative monoids from bigop *) -HB.instance Definition _ (U : pcm) := - Monoid.isComLaw.Build U Unit join joinA joinC unitL. - -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). - -Lemma big_pmorph I r (P : pred I) F : - valid (\big[join/Unit]_(i <- r | P i) F i) -> - f (\big[join/Unit]_(i <- r | P i) F i) = - \big[op1/id1]_(i <- r | P i) f (F i). -Proof. -rewrite unlock; elim: r=>[|x r IH] //=; case: ifP=>// H V. -by rewrite f_op // IH //; apply: validR V. -Qed. - -End BigPartialMorph. - - -(*********************) -(* PCM constructions *) -(*********************) - -(* nats with addition are a pcm *) - -(* we use isPCM to build a PCM over a type *) -(* if the type has other characteristics, e.g. equality *) -(* the appropriate structure is built automatically *) -(* e.g., the following builds both the PCM and the EQPCM *) -(* eqpcm instance is automatically inferred *) -Lemma nat_is_pcm : pcm_axiom xpredT addn 0 (eq_op^~ 0). -Proof. by split=>//; [apply:addnC|apply:addnA|apply:(@eqP _^~_)]. Qed. -HB.instance Definition natPCM : isPCM nat := isPCM.Build nat nat_is_pcm. -(* Check (PCM.clone nat _). *) - -(* nats are pcm with multiplication too *) -(* but the instance isn't declared canonical as natPCM already is *) -(* DEVCOMMENT *) -(* To have both, we must redo PCM def so that it keys on join op *) -(* (as in bigops), and not on type. But that is drastic and of unclear *) -(* utility in this setting (e.g., we can't have uniform notation \+). *) -(* /DEVCOMMENT *) -Lemma nat_is_mulpcm : pcm_axiom xpredT mult 1 (eq_op^~ 1). -Proof. by split=>//; [apply:mulnC|apply:mulnA|apply:mul1n|apply:(@eqP _^~_)]. Qed. -HB.instance Definition nat_mulPCM : isPCM nat := isPCM.Build nat nat_is_mulpcm. - -(* nats with max too *) -Lemma nat_is_maxpcm : pcm_axiom xpredT maxn 0 (eq_op^~ 0). -Proof. by split=>//; [apply:maxnC|apply:maxnA|apply:max0n|apply:(@eqP _^~_)]. Qed. -HB.instance Definition nat_maxPCM : isPCM nat := isPCM.Build nat nat_is_maxpcm. - -(* bools with conjunction *) -(* eqpcm automatically inferred *) -Lemma bool_is_pcm : pcm_axiom xpredT andb true (eq_op^~ true). -Proof. by split=>//; [apply:andbC|apply:andbA|apply:(@eqP _^~_)]. Qed. -HB.instance Definition boolPCM : isPCM bool := isPCM.Build bool bool_is_pcm. -(* Check (EQPCM.clone bool _). *) - -(* bools with disjunction *) -(* but the instance isn't declared canonical as boolPCM alread is *) -Lemma bool_is_disjpcm : pcm_axiom xpredT orb false (eq_op^~ false). -Proof. by split=>//; [apply:orbC|apply:orbA|apply:(@eqP _^~_)]. Qed. -HB.instance Definition bool_disjPCM : isPCM bool := - isPCM.Build bool bool_is_disjpcm. - -(* positive nats under max *) -Section PosNatMax. -Definition posNat := sig (fun x => x > 0). - -Definition pos_valid := [fun x : posNat => true]. -Definition pos_unit : posNat := Sub 1 (leqnn 1). -Definition pos_unitb := [fun x : posNat => val x == 1]. -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_join := [fun x y : posNat => - Sub (maxn (val x) (val y)) (max_pos x y) : posNat]. - -Lemma posnat_is_pcm : pcm_axiom pos_valid pos_join pos_unit pos_unitb. -Proof. -split=>[x y|x y z|x|||x] //. -- by apply/eqP; rewrite -val_eqE /= maxnC. -- by apply/eqP; rewrite -val_eqE /= maxnA. -- by apply/eqP; rewrite -val_eqE; apply/eqP/maxn_idPr; case: x. -by apply/eqP. -Qed. - -HB.instance Definition _ : isPCM posNat := isPCM.Build posNat posnat_is_pcm. -(* inherit equality type from sig and nat *) -HB.instance Definition _ := Equality.copy posNat _. -End PosNatMax. - -Arguments pos_unit /. - -(* unit is a PCM, but not a TPCM, as there is no undefined element. *) -(* we'll have to lift with option types for that *) -Section UnitPCM. -Definition unit_valid := [fun x : unit => true]. -Definition unit_join := [fun x y : unit => tt]. -Definition unit_unit := tt. -Definition unit_unitb := [fun x : unit => true]. - -Lemma unit_is_pcm : pcm_axiom unit_valid unit_join unit_unit unit_unitb. -Proof. by split=>//; case=>//; apply/eqP. Qed. - -HB.instance Definition _ : isPCM unit := isPCM.Build unit unit_is_pcm. -End UnitPCM. - -Arguments unit_unit /. - -(***********************) -(* Option PCM and TPCM *) -(***********************) - -Section OptionTPCM. -Variables U : pcm. - -Definition ovalid := [fun x : option U => - if x is Some a then valid a else false]. -Definition ojoin := [fun x y : option U => - if x is Some a then - if y is Some b then Some (a \+ b) else None - else None]. -Definition ounit : option U := Some Unit. -Definition ounitb := [fun x : option U => - if x is Some v then unitb v else false]. -Definition oundef : option U := None. -Definition oundefb := [fun x : option U => - if x is Some a then false else true]. - -Lemma option_is_pcm : pcm_axiom ovalid ojoin ounit ounitb. -Proof. -split=>[x y|x y z|x|x y||x]. -- by case: x; case: y=>//=b a; rewrite joinC. -- by case: x; case: y; case: z=>//=c b a; rewrite joinA. -- by case: x=>//=a; rewrite unitL. -- by case x=>//=a; case: y=>//=b /validL. -- by rewrite /= valid_unit. -case: x=>[a|] /=; last by constructor. -by case: unitbP=>[->|H]; constructor=>//; case. -Qed. - -HB.instance Definition _ : isPCM (option U) := - isPCM.Build (option U) option_is_pcm. - -Lemma option_is_tpcm : tpcm_axiom oundef oundefb. -Proof. by split=>//; case=>[a|]; constructor. Qed. - -HB.instance Definition _ : isTPCM (option U) := - isTPCM.Build (option U) option_is_tpcm. -End OptionTPCM. - -Arguments ounit /. -Arguments oundef /. - -(* option preserves equality *) -HB.instance Definition _ (U : eqpcm) := PCM.copy (option U) _. -(* Check (EQPCM.clone (option nat) _). *) - -(* option preserves cancellativity *) -Lemma option_is_cpcm (U : cpcm) : cpcm_axiom (option U). -Proof. by case=>[x|][y|][z|] // V [] /(joinKx V)=>->. Qed. -HB.instance Definition _ (U : cpcm) := - isCPCM.Build (option U) (@option_is_cpcm U). - -(* option is normal if U is all valid *) -Lemma option_is_normal (U : pcm) : - @valid U =1 xpredT -> normal_tpcm_axiom (option U). -Proof. by move=>E [x|]; [left; rewrite /valid /= E|right]. Qed. - -(* case analysis infrastructure *) -CoInductive option_pcm_spec (A : pcm) (x y : option A) : - option A -> Type := -| some_case x' y' of x = some x' & y = some y' : - option_pcm_spec x y (some (x' \+ y')) -| none_case of ~~ valid (x \+ y) : - option_pcm_spec x y None. - -Lemma optionP (A : pcm) (x y : option A) : option_pcm_spec x y (x \+ y). -Proof. -case: x=>[x|]; case: y=>[y|]=>/=; first by [apply: some_case]; -by apply: none_case. -Qed. - -(*****************************) -(* Cartesian product of PCMs *) -(*****************************) - -Section ProdPCM. -Variables U V : pcm. -Local Notation tp := (U * V)%type. - -Definition valid2 := [fun x : tp => valid x.1 && valid x.2]. -Definition join2 := [fun x1 x2 : tp => (x1.1 \+ x2.1, x1.2 \+ x2.2)]. -Definition unit2 : tp := (Unit, Unit). -Definition unitb2 := [fun x : U * V => unitb x.1 && unitb x.2]. - -Lemma prod_is_pcm : pcm_axiom valid2 join2 unit2 unitb2. -Proof. -split=>[x y|x y z|x|x y||x]. -- by rewrite /= (joinC x.1) (joinC x.2). -- by rewrite /= !joinA. -- by rewrite /= !unitL -prod_eta. -- by rewrite /valid2 /=; case/andP=>/validL -> /validL ->. -- by rewrite /valid2 /= !valid_unit. -case: x=>x1 x2 /=; case: andP=>H; constructor. -- by case: H=>/unitbP -> /unitbP ->. -by case=>X1 X2; elim: H; rewrite X1 X2 !pcmE. -Qed. - -HB.instance Definition _ : isPCM (U * V)%type := - isPCM.Build (U * V)%type prod_is_pcm. -End ProdPCM. - -(* explicitly extend to eqpcms *) -HB.instance Definition _ (U V : eqpcm) := PCM.copy (U * V)%type _. - -Arguments unit2 /. - -(* product simplification *) -Section Simplification. -Variable U V : pcm. - -Lemma pcmPJ (x1 y1 : U) (x2 y2 : V) : - (x1, x2) \+ (y1, y2) = (x1 \+ y1, x2 \+ y2). -Proof. by rewrite pcmE. 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 : U * V) : valid x = valid x.1 && valid x.2. -Proof. by []. Qed. - -Lemma pcmPU : Unit = (Unit, Unit) :> U * V. -Proof. by rewrite pcmE. Qed. - -Definition pcmPE := (pcmPJ, pcmFJ, pcmSJ, pcmPV, pcmPU). - -End Simplification. - -(* We often iterate PCM-products so *) -(* we provide primitives for small numbers *) - -Section Prod3PCM. -Variables U1 U2 U3 : pcm. -Notation tp := (Prod3 U1 U2 U3). -Definition valid3 := [fun x : tp => - [&& valid (proj31 x), - valid (proj32 x) & - valid (proj33 x)]]. -Definition join3 := [fun x y : tp => - mk3 (proj31 x \+ proj31 y) - (proj32 x \+ proj32 y) - (proj33 x \+ proj33 y)]. -Definition unit3 : tp := mk3 Unit Unit Unit. -Definition unitb3 := [fun x : tp => - [&& unitb (proj31 x), - unitb (proj32 x) & - unitb (proj33 x)]]. - -Lemma prod3_is_pcm : pcm_axiom valid3 join3 unit3 unitb3. -Proof. -split=>[x y|x y z||x y||x] /=. -- by congr mk3; rewrite joinC. -- by congr mk3; rewrite joinA. -- by case=>*; rewrite /= !unitL. -- by case/and3P; do ![move/validL=>->]. -- by rewrite !valid_unit. -case: x=>x1 x2 x3 /=. -do ![case: unitbP=>[->|H]; last by constructor; case]; -by constructor. -Qed. - -HB.instance Definition _ : isPCM (Prod3 U1 U2 U3) := - isPCM.Build (Prod3 U1 U2 U3) prod3_is_pcm. -End Prod3PCM. - -HB.instance Definition _ (U1 U2 U3 : eqpcm) := - PCM.copy (Prod3 U1 U2 U3) _. - -Arguments unit3 /. - -Section Prod4PCM. -Variables U1 U2 U3 U4 : pcm. -Notation tp := (Prod4 U1 U2 U3 U4). -Definition valid4 := [fun x : tp => - [&& valid (proj41 x), - valid (proj42 x), - valid (proj43 x) & - valid (proj44 x)]]. -Definition join4 := [fun x y : tp => - mk4 (proj41 x \+ proj41 y) - (proj42 x \+ proj42 y) - (proj43 x \+ proj43 y) - (proj44 x \+ proj44 y)]. -Definition unit4 : tp := mk4 Unit Unit Unit Unit. -Definition unitb4 := [fun x : tp => - [&& unitb (proj41 x), - unitb (proj42 x), - unitb (proj43 x) & - unitb (proj44 x)]]. - -Lemma prod4_is_pcm : pcm_axiom valid4 join4 unit4 unitb4. -Proof. -split=>[x y|x y z||x y||x] /=. -- by congr mk4; rewrite joinC. -- by congr mk4; rewrite joinA. -- by case=>*; rewrite /= !unitL. -- by case/and4P; do ![move/validL=>->]. -- by rewrite !valid_unit. -case: x=>x1 x2 x3 x4 /=. -do ![case: unitbP=>[->|H]; last by constructor; case]; -by constructor. -Qed. - -HB.instance Definition _ : isPCM (Prod4 U1 U2 U3 U4) := - isPCM.Build (Prod4 U1 U2 U3 U4) prod4_is_pcm. -End Prod4PCM. - -HB.instance Definition _ (U1 U2 U3 U4 : eqpcm) := - PCM.copy (Prod4 U1 U2 U3 U4) _. - -Arguments unit4 /. - -Section Prod5PCM. -Variables U1 U2 U3 U4 U5 : pcm. -Notation tp := (Prod5 U1 U2 U3 U4 U5). - -Definition valid5 := [fun x : tp => - [&& valid (proj51 x), - valid (proj52 x), - valid (proj53 x), - valid (proj54 x) & - valid (proj55 x)]]. -Definition join5 := [fun x y : tp => - mk5 (proj51 x \+ proj51 y) - (proj52 x \+ proj52 y) - (proj53 x \+ proj53 y) - (proj54 x \+ proj54 y) - (proj55 x \+ proj55 y)]. -Definition unit5 : tp := mk5 Unit Unit Unit Unit Unit. -Definition unitb5 := [fun x : tp => - [&& unitb (proj51 x), - unitb (proj52 x), - unitb (proj53 x), - unitb (proj54 x) & - unitb (proj55 x)]]. - -Lemma prod5_is_pcm : pcm_axiom valid5 join5 unit5 unitb5. -Proof. -split=>[x y|x y z||x y||x] /=. -- by congr mk5; rewrite joinC. -- by congr mk5; rewrite joinA. -- by case=>*; rewrite /= !unitL. -- by case/and5P; do ![move/validL=>->]. -- by rewrite !valid_unit. -case: x=>x1 x2 x3 x4 x5 /=. -do ![case: unitbP=>[->|H]; last by constructor; case]; -by constructor. -Qed. - -HB.instance Definition _ : isPCM (Prod5 U1 U2 U3 U4 U5) := - isPCM.Build (Prod5 U1 U2 U3 U4 U5) prod5_is_pcm. -End Prod5PCM. - -HB.instance Definition _ (U1 U2 U3 U4 U5 : eqpcm) := - PCM.copy (Prod5 U1 U2 U3 U4 U5) _. - -Arguments unit5 /. - -Section Prod6PCM. -Variables U1 U2 U3 U4 U5 U6 : pcm. -Notation tp := (Prod6 U1 U2 U3 U4 U5 U6). - -Definition valid6 := [fun x : tp => - [&& valid (proj61 x), - valid (proj62 x), - valid (proj63 x), - valid (proj64 x), - valid (proj65 x) & - valid (proj66 x)]]. -Definition join6 := [fun x y : tp => - mk6 (proj61 x \+ proj61 y) - (proj62 x \+ proj62 y) - (proj63 x \+ proj63 y) - (proj64 x \+ proj64 y) - (proj65 x \+ proj65 y) - (proj66 x \+ proj66 y)]. -Definition unit6 : tp := mk6 Unit Unit Unit Unit Unit Unit. -Definition unitb6 := [fun x : tp => - [&& unitb (proj61 x), - unitb (proj62 x), - unitb (proj63 x), - unitb (proj64 x), - unitb (proj65 x) & - unitb (proj66 x)]]. - -Lemma prod6_is_pcm : pcm_axiom valid6 join6 unit6 unitb6. -Proof. -split=>[x y|x y z||x y||x] /=. -- by congr mk6; rewrite joinC. -- by congr mk6; rewrite joinA. -- by case=>*; rewrite /= !unitL. -- by case/and6P; do ![move/validL=>->]. -- by rewrite !valid_unit. -case: x=>x1 x2 x3 x4 x5 x6 /=. -do ![case: unitbP=>[->|H]; last by constructor; case]; -by constructor. -Qed. - -HB.instance Definition _ : isPCM (Prod6 U1 U2 U3 U4 U5 U6) := - isPCM.Build (Prod6 U1 U2 U3 U4 U5 U6) prod6_is_pcm. -End Prod6PCM. - -HB.instance Definition _ (U1 U2 U3 U4 U5 U6 : eqpcm) := - PCM.copy (Prod6 U1 U2 U3 U4 U5 U6) _. - -Arguments unit6 /. - -Section Prod7PCM. -Variables U1 U2 U3 U4 U5 U6 U7 : pcm. -Notation tp := (Prod7 U1 U2 U3 U4 U5 U6 U7). - -Definition valid7 := [fun x : tp => - [&& valid (proj71 x), - valid (proj72 x), - valid (proj73 x), - valid (proj74 x), - valid (proj75 x), - valid (proj76 x) & - valid (proj77 x)]]. -Definition join7 := [fun x y : tp => - mk7 (proj71 x \+ proj71 y) - (proj72 x \+ proj72 y) - (proj73 x \+ proj73 y) - (proj74 x \+ proj74 y) - (proj75 x \+ proj75 y) - (proj76 x \+ proj76 y) - (proj77 x \+ proj77 y)]. -Definition unit7 : tp := mk7 Unit Unit Unit Unit Unit Unit Unit. -Definition unitb7 := [fun x : tp => - [&& unitb (proj71 x), - unitb (proj72 x), - unitb (proj73 x), - unitb (proj74 x), - unitb (proj75 x), - unitb (proj76 x) & - unitb (proj77 x)]]. - -Lemma prod7_is_pcm : pcm_axiom valid7 join7 unit7 unitb7. -Proof. -split=>[x y|x y z||x y||x] /=. -- by congr mk7; rewrite joinC. -- by congr mk7; rewrite joinA. -- by case=>*; rewrite /= !unitL. -- by case/and7P; do ![move/validL=>->]. -- by rewrite !valid_unit. -case: x=>x1 x2 x3 x4 x5 x6 x7 /=. -do ![case: unitbP=>[->|H]; last by constructor; case]; -by constructor. -Qed. - -HB.instance Definition _ : isPCM (Prod7 U1 U2 U3 U4 U5 U6 U7) := - isPCM.Build (Prod7 U1 U2 U3 U4 U5 U6 U7) prod7_is_pcm. -End Prod7PCM. - -HB.instance Definition _ (U1 U2 U3 U4 U5 U6 U7 : eqpcm) := - PCM.copy (Prod7 U1 U2 U3 U4 U5 U6 U7) _. - -Arguments unit7 /. - -(* Finite products of PCMs as functions *) -Section FunPCM. -Variables (T : finType) (Us : T -> pcm). -Notation tp := (forall t, Us t). - -Definition fun_valid := [fun f : tp => [forall t, valid (f t)]]. -Definition fun_join := [fun f1 f2 : tp => fun t => f1 t \+ f2 t]. -Definition fun_unit : tp := fun t => Unit. -Definition fun_unitb := [fun f : tp => [forall t, unitb (f t)]]. - -Lemma fun_is_pcm : pcm_axiom fun_valid fun_join fun_unit fun_unitb. -Proof. -split=>[f g|f g h|f|f g||f] /=. -- by apply: fext=>t; rewrite joinC. -- by apply: fext=>t; rewrite joinA. -- by apply: fext=>t; rewrite unitL. -- by move/forallP=>V; apply/forallP=>t; apply: validL (V t). -- by apply/forallP=>t; apply: valid_unit. -case: forallP=>H; constructor. -- by apply: fext=>t; apply/unitbP; apply: H. -by move=>H1; elim: H=>t; apply/unitbP; rewrite H1. -Qed. - -HB.instance Definition _ : isPCM (forall t, Us t) := - isPCM.Build (forall t, Us t) fun_is_pcm. -End FunPCM. - -Arguments fun_unit /. - -(* we won't use fun types for any examples *) -(* so we don't need equality structure on them *) - - -(* Finite products of PCMs as finfuns *) -(* We will work with this second structure, though *) -(* the one above is needed to state that sel and finfun are morphisms *) -(* (between FunPCM and FinPCM) *) -(* dffun used for inheritance (see finfun.v) *) -Section FinPCM. -Variables (T : finType) (Us : T -> pcm). -Notation tp := {dffun forall t, Us t}. - -Definition fin_valid := [fun f : tp => [forall t, valid (sel t f)]]. -Definition fin_join := [fun f g : tp => [ffun t => sel t f \+ sel t g]]. -Definition fin_unit : tp := [ffun => Unit]. -Definition fin_unitb := [fun f : tp => [forall t, unitb (sel t f)]]. - -Lemma ffinprod_is_pcm : pcm_axiom fin_valid fin_join fin_unit fin_unitb. -Proof. -split=>[x y|x y z|x|x y||x] /=. -- by apply/ffunP=>t; rewrite !ffunE joinC. -- by apply/ffunP=>t; rewrite /sel !ffunE joinA. -- by apply/ffunP=>t; rewrite /sel !ffunE unitL. -- move/forallP=>H; apply/forallP=>t; move: (H t). - by rewrite /sel !ffunE=>/validL. -- by apply/forallP=>t; rewrite /sel ffunE valid_unit. -case H: [forall t, _]; constructor. -- move/forallP: H=>H; apply/ffinP=>t. - by rewrite sel_fin; move/unitbP: (H t). -move=>E; move/negP: H; apply; apply/forallP=>t. -by rewrite E sel_fin unitb0. -Qed. - -HB.instance Definition _ := - isPCM.Build {dffun forall t, Us t} ffinprod_is_pcm. -End FinPCM. - -HB.instance Definition _ (T : finType) (Us : T -> eqpcm) := - PCM.copy {dffun forall t, Us t} _. - -Arguments fin_unit /. - -(* product of TPCMs is a TPCM *) - -(* With TPCMs we could remove the pairs where *) -(* one element is valid and the other is not. *) -(* We will do that later with normalization procedure. *) -(* That will give us a new construction for normalized products *) -(* but we first need to introduce PCM morphisms *) -(* Here, we give non-normalized ones. *) -Section ProdTPCM. -Variables U V : tpcm. -Definition undef2 := (@undef U, @undef V). -Definition undefb2 := [fun x : U * V => undefb x.1 && undefb x.2]. - -Lemma prod_is_tpcm : tpcm_axiom undef2 undefb2. -Proof. -split=>/=. -- case=>x1 x2; case: andP=>/= H; constructor. - - by case: H=>/undefbP -> /undefbP ->. - by case=>X1 X2; elim: H; rewrite X1 X2 !tpcmE. -- by rewrite /valid /= !valid_undef. -by case=>x1 x2; rewrite pcmPJ !undef_join. -Qed. - -HB.instance Definition _ := isTPCM.Build (U * V)%type prod_is_tpcm. -End ProdTPCM. - -Arguments undef2 /. - -(* iterated TPCMs *) - -Section Prod3TPCM. -Variables U1 U2 U3 : tpcm. -Notation tp := (Prod3 U1 U2 U3). -Definition undef3 : tp := mk3 undef undef undef. -Definition undefb3 := [fun x : tp => - [&& undefb (proj31 x), - undefb (proj32 x) & - undefb (proj33 x)]]. - -Lemma prod3_is_tpcm : tpcm_axiom undef3 undefb3. -Proof. -split=>[x||x]. -- case: x=>x1 x2 x3 /=. - do ![case: undefbP=>[->|H]; last by constructor; case]. - by constructor. -- by rewrite /valid /= !valid_undef. -by rewrite pcmE /= !undef_join. -Qed. - -HB.instance Definition _ : isTPCM (Prod3 U1 U2 U3) := - isTPCM.Build (Prod3 U1 U2 U3) prod3_is_tpcm. -End Prod3TPCM. - -Arguments undef3 /. - -Section Prod4TPCM. -Variables U1 U2 U3 U4 : tpcm. -Notation tp := (Prod4 U1 U2 U3 U4). -Definition undef4 : tp := mk4 undef undef undef undef. -Definition undefb4 := [fun x : tp => - [&& undefb (proj41 x), - undefb (proj42 x), - undefb (proj43 x) & - undefb (proj44 x)]]. - -Lemma prod4_is_tpcm : tpcm_axiom undef4 undefb4. -Proof. -split=>[x||x] /=. -- case: x=>x1 x2 x3 x4 /=. - do ![case: undefbP=>[->|H]; last by constructor; case]. - by constructor. -- by rewrite /valid /= !valid_undef. -by rewrite pcmE /= !undef_join. -Qed. - -HB.instance Definition _ : isTPCM (Prod4 U1 U2 U3 U4) := - isTPCM.Build (Prod4 U1 U2 U3 U4) prod4_is_tpcm. -End Prod4TPCM. - -Arguments undef4 /. - -Section Prod5TPCM. -Variables U1 U2 U3 U4 U5 : tpcm. -Notation tp := (Prod5 U1 U2 U3 U4 U5). -Definition undef5 : tp := mk5 undef undef undef undef undef. -Definition undefb5 := [fun x : tp => - [&& undefb (proj51 x), - undefb (proj52 x), - undefb (proj53 x), - undefb (proj54 x) & - undefb (proj55 x)]]. - -Lemma prod5_is_tpcm : tpcm_axiom undef5 undefb5. -Proof. -split=>[x||x] /=. -- case: x=>x1 x2 x3 x4 x5 /=. - do ![case: undefbP=>[->|H]; last by constructor; case]. - by constructor. -- by rewrite /valid /= !valid_undef. -by rewrite pcmE /= !undef_join. -Qed. - -HB.instance Definition _ : isTPCM (Prod5 U1 U2 U3 U4 U5) := - isTPCM.Build (Prod5 U1 U2 U3 U4 U5) prod5_is_tpcm. -End Prod5TPCM. - -Arguments undef5 /. - - -Section Prod6TPCM. -Variables U1 U2 U3 U4 U5 U6 : tpcm. -Notation tp := (Prod6 U1 U2 U3 U4 U5 U6). -Definition undef6 : tp := mk6 undef undef undef undef undef undef. -Definition undefb6 := [fun x : tp => - [&& undefb (proj61 x), - undefb (proj62 x), - undefb (proj63 x), - undefb (proj64 x), - undefb (proj65 x) & - undefb (proj66 x)]]. - -Lemma prod6_is_tpcm : tpcm_axiom undef6 undefb6. -Proof. -split=>[x||x] /=. -- case: x=>x1 x2 x3 x4 x5 x6 /=. - do ![case: undefbP=>[->|H]; last by constructor; case]. - by constructor. -- by rewrite /valid /= !valid_undef. -by rewrite pcmE /= !undef_join. -Qed. - -HB.instance Definition _ : isTPCM (Prod6 U1 U2 U3 U4 U5 U6) := - isTPCM.Build (Prod6 U1 U2 U3 U4 U5 U6) prod6_is_tpcm. -End Prod6TPCM. - -Arguments undef6 /. - -Section Prod7TPCM. -Variables U1 U2 U3 U4 U5 U6 U7 : tpcm. -Notation tp := (Prod7 U1 U2 U3 U4 U5 U6 U7). -Definition undef7 : tp := mk7 undef undef undef undef undef undef undef. -Definition undefb7 := [fun x : tp => - [&& undefb (proj71 x), - undefb (proj72 x), - undefb (proj73 x), - undefb (proj74 x), - undefb (proj75 x), - undefb (proj76 x) & - undefb (proj77 x)]]. - -Lemma prod7_is_tpcm : tpcm_axiom undef7 undefb7. -Proof. -split=>[x||x]. -- case: x=>x1 x2 x3 x4 x5 x6 x7 /=. - do ![case: undefbP=>[->|H]; last by constructor; case]. - by constructor. -- by rewrite /valid /= !valid_undef. -by rewrite pcmE /= !undef_join. -Qed. - -HB.instance Definition _ : isTPCM (Prod7 U1 U2 U3 U4 U5 U6 U7) := - isTPCM.Build (Prod7 U1 U2 U3 U4 U5 U6 U7) prod7_is_tpcm. -End Prod7TPCM. - -Arguments undef7 /. - -(* TPCM proofs use function extensionality *) -(* it's TPCM only if T inhabited finite type *) -(* (otherwise valid undef) *) - -Definition fun_undef T (Us : T -> tpcm) : forall t, Us t - := fun t => undef. -Arguments fun_undef {T Us} /. - -Definition fun_undefb (T : finType) (Us : T -> tpcm) := - fun f : forall t, Us t => [forall t, undefb (f t)]. -Arguments fun_undefb {T Us} f /. - -Section FunTPCM. -Variables (T : ifinType) (Us : T -> tpcm). -Notation tp := (forall t, Us t). - -Lemma fun_is_tpcm : tpcm_axiom fun_undef (fun_undefb (Us:=Us)). -Proof. -split=>[f||f] /=. -- case: forallP=>H; constructor. - - by apply: fext=>t; apply/undefbP; apply: H. - by move=>H1; elim: H=>t; apply/undefbP; rewrite H1. -- by apply/negP=>/forallP/(_ inhab); rewrite valid_undef. -by apply: fext=>t; rewrite /join /= undef_join. -Qed. - -HB.instance Definition _ : isTPCM (forall t, Us t) := - isTPCM.Build (forall t, Us t) fun_is_tpcm. - -End FunTPCM. - - -(* TPCM only if T inhabited finite type *) -(* (otherwise valid undef) *) -Definition fin_undef (T : finType) (Us : T -> tpcm) - : {dffun forall t, Us t} := [ffun t => undef]. -Arguments fin_undef {T Us} /. - -Definition fin_undefb (T : finType) (Us : T -> tpcm) - (x : {dffun forall t, Us t}) := - [forall t, undefb (sel t x)]. -Arguments fin_undefb {T Us} x /. - -Section FinTPCM. -Variables (T : ifinType) (Us : T -> tpcm). -Notation tp := {dffun forall t, Us t}. - -Lemma finprod_is_tpcm : tpcm_axiom fin_undef (fin_undefb (Us:=Us)). -Proof. -split=>[x||x] /=. -- case H : [forall t, _]; constructor. - - move/forallP: H=>H; apply/ffinP=>t. - by rewrite sel_fin; move/undefbP: (H t). - move=>E; move/negP: H; apply; apply/forallP=>t. - by rewrite E sel_fin undefb_undef. -- apply/negP=>/forallP/(_ inhab). - by rewrite sel_fin valid_undef. -by apply/ffinP=>t; rewrite !sel_fin undef_join. -Qed. - -HB.instance Definition _ : isTPCM tp := - isTPCM.Build tp finprod_is_tpcm. -End FinTPCM. - -(*************************) -(* PCM-induced pre-order *) -(*************************) - -Definition pcm_preord (U : pcm) (x y : U) := exists z, y = x \+ z. - -Notation "[ 'pcm' x '<=' y ]" := (@pcm_preord _ x y) - (at level 0, x, y at level 69, - format "[ '[hv' 'pcm' x '/ ' <= y ']' ]") : type_scope. - -Section PleqLemmas. -Variable U : pcm. -Implicit Types x y z : U. - -Lemma pleq_unit x : [pcm Unit <= x]. -Proof. by exists x; rewrite unitL. Qed. - -(* preorder lemmas *) - -(* 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]. -Proof. by exists Unit; rewrite unitR. Qed. - -Lemma pleq_trans x y z : [pcm x <= y] -> [pcm y <= z] -> [pcm x <= z]. -Proof. by case=>a -> [b ->]; exists (a \+ b); rewrite joinA. Qed. - -(* monotonicity lemmas *) - -Lemma pleq_join2r x x1 x2 : [pcm x1 <= x2] -> [pcm x1 \+ x <= x2 \+ x]. -Proof. by case=>a ->; exists a; rewrite joinAC. Qed. - -Lemma pleq_join2l x x1 x2 : [pcm x1 <= x2] -> [pcm x \+ x1 <= x \+ x2]. -Proof. by rewrite !(joinC x); apply: pleq_join2r. Qed. - -Lemma pleq_joinr {x1 x2} : [pcm x1 <= x1 \+ x2]. -Proof. by exists x2. Qed. - -Lemma pleq_joinl {x1 x2} : [pcm x2 <= x1 \+ x2]. -Proof. by rewrite joinC; apply: pleq_joinr. Qed. - -(* validity lemmas *) - -Lemma pleqV (x1 x2 : U) : [pcm x1 <= x2] -> valid x2 -> valid x1. -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 -> V; rewrite (validRE3 V). Qed. - -Lemma pleq_validR (x x1 x2 : U) : - [pcm x1 <= x2] -> valid (x2 \+ x) -> valid (x1 \+ x). -Proof. by rewrite -!(joinC x); apply: pleq_validL. Qed. - -(* the next two lemmas only hold for cancellative PCMs *) - -Lemma pleq_joinxK (V : cpcm) (x x1 x2 : V) : - valid (x2 \+ x) -> [pcm x1 \+ x <= x2 \+ x] -> [pcm x1 <= x2]. -Proof. by move=>W [a]; rewrite joinAC=>/(joinKx W) ->; exists a. Qed. - -Lemma pleq_joinKx (V : cpcm) (x x1 x2 : V) : - valid (x \+ x2) -> [pcm x \+ x1 <= x \+ x2] -> [pcm x1 <= x2]. -Proof. by rewrite !(joinC x); apply: pleq_joinxK. Qed. - -End PleqLemmas. - -#[export] -Hint Resolve pleq_unit pleq_refl pleq_joinr pleq_joinl : core. -Prenex Implicits pleq_refl pleq_joinl pleq_joinr. - -(* shorter names *) -Notation pcmR := pleq_refl. -Notation pcmS := pleq_joinr. -Notation pcmO := pleq_joinl. - -Lemma pleq_undef (U : tpcm) (x : U) : [pcm x <= undef]. -Proof. by exists undef; rewrite join_undef. Qed. - -#[export] -Hint Resolve pleq_undef : core. - -(********************) -(* PCMs and folding *) -(********************) - -(* folding functions that are morphisms in the first argument *) - -Section PCMfoldLeft. -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. -rewrite foldl_cat /= foldl_cat; congr foldl. -elim: {s2} s1 z0=>[|k ks IH] 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 PCMfoldLeft. - -Section PCMfoldRight. -Variables (A : Type) (R : pcm) (a : A -> R -> R). -Hypothesis H : (forall k x y, a k (x \+ y) = a k x \+ y). - -Lemma foldr_helper (s1 s2 : seq A) (z0 : R) x : - foldr a z0 (s1 ++ x :: s2) = foldr a z0 (s1 ++ rcons s2 x). -Proof. -rewrite -!foldl_rev rev_cat rev_cons cat_rcons. -rewrite foldl_helper; last by move=>*; rewrite H. -by rewrite -[x :: _ ++ _]revK rev_cons rev_cat !revK rcons_cat. -Qed. - -Lemma foldr_perm (s1 s2 : seq A) (z0 : R) : - perm s1 s2 -> - foldr a z0 s1 = foldr a z0 s2. -Proof. -move=>P; rewrite -!foldl_rev; apply: foldl_perm. -- by move=>*; rewrite H. -apply/(pperm_trans _ (pperm_rev _))/pperm_sym. -by apply/(pperm_trans _ (pperm_rev _))/pperm_sym. -Qed. - -Lemma foldr_init (s : seq A) (x y : R) : - foldr a (x \+ y) s = foldr a x s \+ y. -Proof. by rewrite -!foldl_rev foldl_init. Qed. - -End PCMfoldRight. - -(* fold-join lemmas *) - -Lemma foldl_join A (U : pcm) h (s : seq A) (a : A -> U): - foldl (fun h t => a t \+ h) h s = - foldl (fun h t => a t \+ h) Unit s \+ h. -Proof. -set g := fun x => x \+ h. -apply/esym/(fusion_foldl (g:=g)); last by rewrite /g unitL. -by move=>x y; rewrite joinA. -Qed. - -Lemma foldr_join A (U : pcm) h (s : seq A) (a : A -> U): - foldr (fun t h => h \+ a t) h s = - h \+ foldr (fun t h => h \+ a t) Unit s. -Proof. -apply/esym/fusion_foldr; last by rewrite unitR. -by move=>x y; rewrite joinA. -Qed. - -Section BigOps. -Variables (U : pcm). -Variables (I : Type) (f : I -> U). - -Lemma big_validV (xs : seq I) i : - valid (\big[join/Unit]_(i <- xs) f i) -> - i \In xs -> valid (f i). -Proof. -elim: xs=>[|x xs IH] in i * => //. -rewrite big_cons InE=>V [->|]; first by apply: (validL V). -by apply: IH; rewrite (validR V). -Qed. - -Lemma big_validVL (xs : seq I) z i : - valid (f z \+ \big[join/Unit]_(i <- xs) f i) -> - i \In xs -> i <> z -> valid (f z \+ f i). -Proof. -elim: xs=>[|x xs IH] in i * => //. -rewrite big_cons InE =>V [->_ |]. -- by move: V; rewrite joinA; apply: validL. -by apply: IH; move: V; rewrite joinCA; apply: validR. -Qed. - -Lemma big_validV2 (xs : seq I) : - valid (\big[join/Unit]_(i <- xs) f i) -> - forall i j, i \In xs -> j \In xs -> i <> j -> valid (f i \+ f j). -Proof. -elim: xs=>[|x xs IH] //=; rewrite big_cons. -move=>V i j; rewrite !InE; case=>[->|Xi][->|Xj] N //; last first. -- by apply: IH=>//; apply: (validR V). -- by rewrite joinC; apply: (big_validVL V). -by apply: (big_validVL V)=>// /esym. -Qed. - -End BigOps. - -(***********************************) -(* separating conjunction aka star *) -(***********************************) - -Definition star {U : pcm} (p1 p2 : Pred U) : Pred U := - [Pred h | exists h1 h2, [ /\ h = h1 \+ h2, h1 \In p1 & h2 \In p2] ]. -Definition emp {U : pcm} : Pred U := eq^~ Unit. -Definition top {U : pcm} : Pred U := PredT. - -Notation "p1 '#' p2" := (star p1 p2) - (at level 57, right associativity) : rel_scope. - -Add Parametric Morphism U : (@star U) with signature - @Eq_Pred _ _ ==> @Eq_Pred _ _ ==> @Eq_Pred _ _ as star_morph. -Proof. -move=>x y E x1 y1 E1 m /=; split; case=>h1 [h2][-> H1 H2]; -by exists h1, h2; split=>//; [apply/E|apply/E1]. -Qed. - -(* monoid laws for star hold under <~>, not = *) -(* thus, require fun/prop extensionality to make hold under = *) -(* working with <~> entails use of setoids *) -(* and prevents use of monoids from bigops *) -(* the latter isn't a big deal, however, as most lemmas *) -(* in bigops are about decidable/finite sets *) -(* and won't work with Prop anyway *) - -Section StarMonoid. -Context {U : pcm}. -Implicit Type x y z : Pred U. - -Lemma starA x y z : x # y # z <~> (x # y) # z. -Proof. -move=>m /=; split; case=>h1 [h2][H1 H2]. -- case=>h3 [h4][H3 H4 H5]; subst h2. - eexists (h1 \+ h3), h4; rewrite -joinA; split=>//. - by exists h1, h3. -case: H2=>h3 [h4][? H2 H3 H4]; subst h1. -exists h3, (h4 \+ h2); rewrite joinA; split=>//. -by exists h4, h2. -Qed. - -Lemma starC x y : x # y <~> y # x. -Proof. -move=>m; split. -- by case=>h1 [h2][H H1 H2]; exists h2, h1; rewrite joinC. -by case=>h1 [h2][H H1 H2]; exists h2, h1; rewrite joinC. -Qed. - -Lemma starL x : emp # x <~> x. -Proof. -move=>m; split=>[|H]; first by case=>h1 [h2][->->]; rewrite unitL. -by exists Unit, m; rewrite unitL. -Qed. - -Lemma starR x : x # emp <~> x. -Proof. by rewrite starC starL. Qed. - -Lemma starAC x y z : x # y # z <~> x # z # y. -Proof. by rewrite (starC y). Qed. - -Lemma starCA x y z : x # (y # z) <~> y # (x # z). -Proof. by rewrite starA (starC x) -starA. Qed. - -(* -HB.instance Definition _ := - Monoid.isComLaw.Build (Pred U) emp star starA starC starL. -*) - -End StarMonoid. - - -(* star iterated over sequence *) - -Section IterStarSeq. -Context {U : pcm} {A : Type}. - -Notation seq_join hs := (\big[join/Unit]_(i <- hs) i). - -(* definition is locked to prevent automation from going inside *) -Definition sepit_seq (s : seq A) (f : A -> Pred U) : Pred U := - locked - [Pred h | exists hs : seq U, - [/\ size hs = size s, h = seq_join hs & - All (fun '(a, h) => h \In f a) (zip s hs)]]. - -Lemma sepit_seq_unlock s f : - sepit_seq s f = - [Pred h | exists hs : seq U, - [/\ size hs = size s, h = seq_join hs & - All (fun '(a, h) => h \In f a) (zip s hs)]]. -Proof. by rewrite /sepit_seq -lock. Qed. - -Lemma sepit_seq0 f : sepit_seq [::] f <~> emp. -Proof. -rewrite sepit_seq_unlock=>h; split. -- by case=>/= hs [/size0nil -> -> _]; rewrite !big_nil. -by move=>->; exists [::]; rewrite !big_nil. -Qed. - -Lemma sepitseq_cons x s f : - sepit_seq (x::s) f <~> f x # sepit_seq s f. -Proof. -rewrite !sepit_seq_unlock=>h; split. -- case=>/=; case=>[|h0 hs]; case=>//= /eqP; rewrite eqSS =>/eqP Hs. - rewrite big_cons =>->[H0 H1]. - by exists h0, (seq_join hs); do!split=>//; exists hs. -case=>h1[_][{h}-> H1][hs][E -> H]. -by exists (h1 :: hs); rewrite /= E !big_cons. -Qed. - -(* alternative def *) -(* appears simpler, but no much gain from \big ops *) -(* as star/emp are only monoids under <~>, not = *) -Lemma sepit_seqE s f : - sepit_seq s f <~> \big[star/emp]_(i <- s) f i. -Proof. -elim: s=>[|x s IH]; first by rewrite sepit_seq0 big_nil. -by rewrite sepitseq_cons big_cons IH. -Qed. - -Lemma sepitseq_cat s1 s2 f : - sepit_seq (s1 ++ s2) f <~> - sepit_seq s1 f # sepit_seq s2 f. -Proof. -rewrite sepit_seq_unlock. -elim: s1 s2=>[|x s1 IH] s2 h /=; split. -- case=>hs [E {h}-> H2]. - exists Unit, (seq_join hs); rewrite unitL. - split=>//; first by rewrite sepit_seq0. - by rewrite sepit_seq_unlock; exists hs. -- case=>h1[h2][{h}->]; rewrite sepit_seq0=>->. - by rewrite sepit_seq_unlock unitL. -- case=>/=; case=>[|h0 hs]; case=>//= /eqP; rewrite eqSS=>/eqP E. - rewrite !big_cons /= =>->[H0 HS]. - case: (IH s2 (seq_join hs)).1; first by exists hs. - move=>h1[h2][HJ H1 H2]. - exists (h0 \+ h1), h2; rewrite HJ joinA; split=>//. - by rewrite sepitseq_cons; exists h0, h1. -case=>h1 [h2][{h} ->]; rewrite sepit_seq_unlock. -case; case=>[|h0 hs1]; case=>//= /eqP; rewrite eqSS=>/eqP E1. -rewrite !big_cons /= =>{h1}->[H0 H1]. -rewrite sepit_seq_unlock; case=>hs2[E2 {h2}-> H2]. -exists (h0 :: hs1 ++ hs2); rewrite big_cons big_cat joinA; split=>//=. -- by rewrite !size_cat E1 E2. -rewrite zip_cat //=; split=>//. -by apply/All_cat. -Qed. - -Lemma sepitseq_perm s1 s2 (f : A -> Pred U) : - perm s1 s2 -> - sepit_seq s1 f <~> sepit_seq s2 f. -Proof. -elim: s1 s2 =>[|x s1 IH] s2 /=; first by move/pperm_nil=>->. -move=>H; have Hx: x \In s2 by apply/(pperm_in H)/In_cons; left. -case: (In_split Hx)=>s21[s22] E; rewrite {s2 Hx}E in H *. -move/pperm_cons_cat_cons: H=>/IH {}IH. -rewrite sepitseq_cons sepitseq_cat=>h; rewrite !toPredE; split. -- case=>h1 [h2][->{h} H1]; rewrite IH sepitseq_cat 2!sepit_seq_unlock. - case=>_ [_][->{h2}][hs3][E3 -> H3][hs4][E4 -> H4]; rewrite joinCA. - exists (seq_join hs3), (h1 \+ seq_join hs4); split=>//; first by exists hs3. - rewrite sepitseq_cons sepit_seq_unlock; exists h1, (seq_join hs4). - by split=>//; exists hs4. -rewrite sepitseq_cons sepit_seq_unlock. -case=>_ [h2][{h}-> [hs1][Hs1 -> H1]]. -rewrite sepit_seq_unlock. -case=>h3 [_][->{h2} H3 [hs2][Hs2 -> H2]]; rewrite joinCA. -exists h3, (seq_join hs1 \+ seq_join hs2); split=>//. -rewrite IH sepit_seq_unlock; exists (hs1 ++ hs2); split. -- by rewrite !size_cat Hs1 Hs2. -- by rewrite big_cat. -by rewrite zip_cat //; apply/All_cat. -Qed. - -Lemma sepit_seqF s (f1 f2 : A -> Pred U) : - (forall x, x \In s -> f1 x <~> f2 x) -> - sepit_seq s f1 <~> sepit_seq s f2. -Proof. -elim: s=>[|x s IH] H h; rewrite !toPredE. -- by rewrite !sepit_seq0. -have /IH {IH}H': forall x : A, x \In s -> f1 x =p f2 x. - by move=>? H0; apply: H; apply/In_cons; right. -have Hx : x \In x :: s by apply/In_cons; left. -by rewrite !sepitseq_cons; split; case=>h1[h2][{h}-> H1 H2]; exists h1, h2; -split=>//; [rewrite -H | rewrite -H' | rewrite H | rewrite H']. -Qed. - -Lemma big_sepitseq (s : seq A) (f : A -> U) m : - m = \big[join/Unit]_(i <- s) f i <-> - m \In sepit_seq s (fun i h => h = f i). -Proof. -rewrite sepit_seqE; elim: s m=>[|x xs IH] /= m; first by rewrite !big_nil. -rewrite !big_cons; split=>[E|]; last first. -- by case=>h1 [h2][Em H1 H2]; rewrite Em H1; congr (_ \+ _); rewrite IH. -by rewrite InE; exists (f x), (\big[join/Unit]_(j <- xs) f j); rewrite -IH. -Qed. - -Lemma eq_sepitseqF s (f1 f2 : A -> Pred U) : - (forall x, x \In s -> f1 x <~> f2 x) -> - sepit_seq s f1 <~> sepit_seq s f2. -Proof. by move=>H; apply: sepit_seqF=>x Hx; apply: H. Qed. - -Lemma sepitseq_emp (s : seq A) (f : A -> Pred U) : - (forall x, x \In s -> f x <~> emp (U:=U)) -> - sepit_seq s f <~> emp. -Proof. -move=>H; rewrite sepit_seqE. -elim: s H=>[|a xs IH] H; first by rewrite big_nil. -rewrite big_cons H ?InE; last by left. -by rewrite starL IH // => x X; apply: H; rewrite InE; right. -Qed. - -End IterStarSeq. - -(* iterated star on eqType *) - -Section IterStarEq. -Context {U : pcm} {A : eqType}. - -Lemma sepit_seqP x s (f : A -> Pred U) : - uniq s -> - sepit_seq s f <~> - if x \in s then f x # sepit_seq (filter (predC1 x) s) f - else sepit_seq s f. -Proof. -(* can't use big library for this proof, because all *) -(* necessary lemmas require AC wrt =; thus, direct proof below *) -case E: (x \in s)=>//. -elim: s E=>[|y s IH] //= /[swap]; case/andP=>Hy Hu. -- rewrite sepitseq_cons inE; case/orP. - by move/eqP=>->; rewrite eq_refl filter_predC1. -move=>Hx h0. -have ->: (y != x) by apply/eqP=>Hxy; rewrite Hxy Hx in Hy. -by split; case=>ha[h1][{h0}-> Ha]; [rewrite (IH Hx Hu) | rewrite sepitseq_cons]; -case=>hb[h][{h1}-> Hb H]; rewrite joinCA; exists hb, (ha \+ h); split=>//; -[rewrite sepitseq_cons | rewrite (IH Hx Hu)]; exists ha, h. -Qed. - -Lemma perm_sepitseq s1 s2 (f : A -> Pred U) : - perm_eq s1 s2 -> - sepit_seq s1 f <~> sepit_seq s2 f. -Proof. by move/perm_eq_perm; exact: sepitseq_perm. Qed. - -End IterStarEq. - -(* iterated star on finsets *) - -Section FinIterStar. -Context {U : pcm} {I : finType}. - -Definition sepit (s : {set I}) (Ps : I -> Pred U) := - sepit_seq (enum s) Ps. - -Lemma sepitE (s : {set I}) f : - sepit s f <~> \big[star/emp]_(i in s) f i. -Proof. by rewrite -big_filter /sepit sepit_seqE. Qed. - -Lemma sepit0 f : sepit set0 f <~> emp. -Proof. by rewrite sepitE big_pred0 // => x; rewrite in_set0. Qed. - -Lemma sepitF (s : {set I}) f1 f2 : - (forall x, x \in s -> f1 x <~> f2 x) -> - sepit s f1 <~> sepit s f2. -Proof. -move=>H; apply: eq_sepitseqF=>x /mem_seqP. -by rewrite mem_enum; apply: H. -Qed. - -Lemma eq_sepit (s1 s2 : {set I}) f : - s1 =i s2 -> - sepit s1 f <~> sepit s2 f. -Proof. by move/eq_enum=>H; apply: perm_sepitseq; rewrite H. Qed. - -Lemma sepitS x (s : {set I}) f : - sepit s f <~> if x \in s then f x # sepit (s :\ x) f - else sepit s f. -Proof. -rewrite /sepit {1}(sepit_seqP x f (enum_uniq s)) mem_enum. -case: (x \in s)=>//; rewrite !sepit_seqE !big_filter big_filter_cond. -rewrite (eq_bigl (mem (s :\ x))) //. -by move=>z; rewrite !inE andbC. -Qed. - -Lemma sepitT1 x f : sepit setT f <~> f x # sepit (setT :\ x) f. -Proof. by rewrite (sepitS x) in_setT. Qed. - -Lemma big_sepit (s : {set I}) (f : I -> U) m : - m = \big[join/Unit]_(i in s) (f i) <-> - m \In sepit s (fun i h => h = f i). -Proof. by rewrite /sepit /= -big_sepitseq -big_enum. Qed. - -Lemma big_sepitT (f : I -> U) m : - m = \big[join/Unit]_i (f i) <-> - m \In sepit setT (fun i h => h = f i). -Proof. -rewrite -big_sepit; apply: iff_sym. -by rewrite -big_enum /= enum_T big_enum. -Qed. - -End FinIterStar. diff --git a/pcm/unionmap.v b/pcm/unionmap.v deleted file mode 100644 index ae16212..0000000 --- a/pcm/unionmap.v +++ /dev/null @@ -1,7028 +0,0 @@ -(* -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 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 HB Require Import structures. -From Coq Require Import ssreflect ssrbool ssrfun. -From mathcomp Require Import ssrnat eqtype seq path bigop. -From pcm Require Import options axioms prelude finmap seqperm pred seqext. -From pcm Require Export ordtype. -From pcm Require Import pcm morphism. - -(****************************) -(****************************) -(* Reference Implementation *) -(****************************) -(****************************) - -(* UM.base type is reference implementation of union_maps. *) -(* Type has union_map structure if it exhibits isomorphism with UM.base. *) -(* Tying to reference implementation is alternative to axiomatizing *) -(* union maps, which seems unwieldy. *) -(* Thus, the underlying type of all union_maps instances will be *) -(* (copy of) UM.base. Different copies will *) -(* correspond to different canonical structures. *) -(* Copying further enables using different types for different *) -(* instances, as long as there is isomorphism between them. *) -(* This will allow mixing large and small types and *) -(* storing union maps into union maps (e.g., histories into heaps) *) -(* without universe errors. *) - -Module UM. -Section UM. -Variables (K : ordType) (C : pred K) (V : Type). - -Inductive base := - Undef | Def (f : finMap K V) of all C (supp f). - -Section FormationLemmas. -Variable (f g : finMap K V). - -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 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 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 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' - | 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: eq_irrelevance. -Qed. - -Definition valid f := if f is Def _ _ then true else false. - -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 (C k)) is left pf then - Def (all_supp_insP v pf fpf) - else Undef - else Undef. - -Definition dom f : seq K := - if f is Def fs _ then supp fs else [::]. - -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:=K) k fs else None. - -Definition union f1 f2 := - if (f1, f2) is (Def fs1 pf1, Def fs2 pf2) then - if disj fs1 fs2 then - Def (all_supp_fcatP pf1 pf2) - else Undef - 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. - -Definition pts k v := upd k v empty. - -(* forward induction *) -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) -> - P (union (pts k v) f)) -> - forall f, P f. -Proof. -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] : 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)). -- rewrite /valid /union /pts /upd /=. - case: decP; last by rewrite T1. - by move=>T; case: ifP=>//; rewrite E disjC disj_ins N disj_nil. -move/(H3 k v _ (IH T2)). -rewrite (_ : union (pts k v) (Def T2) = Def H). -- by apply. -rewrite umapE /union /pts /upd /=. -case: decP=>// T; rewrite /disj /= N /=. -by rewrite E fcat_inss // fcat0s. -Qed. - -(* backward induction *) -Lemma base_indb (P : base -> Prop) : - P Undef -> P empty -> - (forall k v f, P f -> valid (union f (pts k v)) -> - (forall x, x \in dom f -> ord x k) -> - P (union (pts k v) f)) -> - forall f, P f. -Proof. -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] : 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 (Def T2) (pts k v)). -- rewrite /valid /union /pts /upd /=. - case: decP; last by rewrite T1. - 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 /=. -case: decP=>// T; rewrite /disj /= N /=. -by rewrite E fcat_inss // fcat0s. -Qed. - -End UM. -End UM. - - -(***********************) -(***********************) -(* Union Map structure *) -(***********************) -(***********************) - -(* type T has union_map structure is it's cancellative TPCM *) -(* with an isomorphism to UM.base *) - -Definition union_map_axiom (K : ordType) (C : pred K) V T - (defined : T -> bool) (empty : T) (undef : T) (upd : K -> V -> T -> T) - (dom : T -> seq K) (assocs : T -> seq (K * V)) (free : T -> K -> T) - (find : K -> T -> option V) (union : T -> T -> T) - (empb : T -> bool) (undefb : T -> bool) (pts : K -> V -> T) - (from : T -> UM.base C V) (to : UM.base C V -> T) := - (((forall b, from (to b) = b) * - (forall f, to (from f) = f)) * - (((forall f, defined f = UM.valid (from f)) * - (undef = to (UM.Undef C V)) * - (empty = to (UM.empty C V))) * - ((forall k v f, upd k v f = to (UM.upd k v (from f))) * - (forall f, dom f = UM.dom (from f)) * - (forall f, assocs f = UM.assocs (from f)) * - (forall f k, free f k = to (UM.free (from f) k))) * - ((forall k f, find k f = UM.find k (from f)) * - (forall f1 f2, union f1 f2 = to (UM.union (from f1) (from f2))) * - (forall f, empb f = UM.empb (from f)) * - (forall f, undefb f = UM.undefb (from f)) * - (forall k v, pts k v = to (UM.pts C k v)))))%type. - -HB.mixin Record isUnionMap_helper (K : ordType) (C : pred K) V T of - Normal_CTPCM T := { - upd : K -> V -> T -> T; - dom : T -> seq K; - assocs : T -> seq (K * V); - free : T -> K -> T; - find : K -> T -> option V; - pts_op : K -> V -> T; - UMC_from : T -> UM.base C V; - UMC_to : UM.base C V -> T; - union_map_helper_subproof : - union_map_axiom valid Unit undef upd dom assocs - free find join unitb undefb pts_op UMC_from UMC_to}. - -(* embed cancellative topped PCM structure to enable inheritance *) -(* we prove later that this isn't needed *) -#[short(type="union_map")] -HB.structure Definition UMC K C V := - {T of @isUnionMap_helper K C V T & Normal_TPCM T & CTPCM T}. - -Lemma ftE K C V (U : @union_map K C V) b : UMC_from (UMC_to b : U) = b. -Proof. by rewrite union_map_helper_subproof. Qed. - -Lemma tfE K C V (U : @union_map K C V) b : UMC_to (UMC_from b) = b :> U. -Proof. by rewrite union_map_helper_subproof. Qed. - -Lemma eqE K C V (U : @union_map K C V) (b1 b2 : UM.base C V) : - UMC_to b1 = UMC_to b2 :> U <-> b1 = b2. -Proof. by split=>[E|->//]; rewrite -(ftE U b1) -(ftE U b2) E. Qed. - -Definition umEX K C V (U : @union_map K C V) := - (@union_map_helper_subproof K C V U, eqE, UM.umapE). - -(* Build CTCPM structure from the rest *) -HB.factory Record isUnion_map (K : ordType) (C : pred K) V T := { - um_defined : T -> bool; - um_empty : T; um_undef : T; - um_upd : K -> V -> T -> T; - um_dom : T -> seq K; - um_assocs : T -> seq (K * V); - um_free : T -> K -> T; - um_find : K -> T -> option V; - um_union : T -> T -> T; - um_empb : T -> bool; - um_undefb : T -> bool; - um_pts : K -> V -> T; - um_from : T -> UM.base C V; - um_to : UM.base C V -> T; - union_map_subproof : - union_map_axiom um_defined um_empty um_undef um_upd um_dom um_assocs - um_free um_find um_union um_empb um_undefb um_pts - um_from um_to}. - -HB.builders Context K C V T of isUnion_map K C V T. - -Definition umEX := snd union_map_subproof. - -Lemma ftE b : um_from (um_to b) = b. -Proof. by rewrite union_map_subproof. Qed. - -Lemma tfE b : um_to (um_from b) = b. -Proof. by rewrite union_map_subproof. Qed. - -Lemma eqE (b1 b2 : UM.base C V) : - um_to b1 = um_to b2 :> T <-> b1 = b2. -Proof. by split=>[E|->//]; rewrite -[b1]ftE -[b2]ftE E. Qed. - -Lemma union_map_is_pcm : pcm_axiom um_defined um_union um_empty um_empb. -Proof. -have joinC f1 f2 : um_union f1 f2 = um_union f2 f1. -- rewrite !umEX !eqE UM.umapE /UM.union. - case: (um_from f1)=>[|f1' H1]; case: (um_from f2)=>[|f2' H2] //. - by case: ifP=>E; rewrite disjC E // fcatC. -have joinCA f1 f2 f3 : um_union f1 (um_union f2 f3) = - um_union f2 (um_union f1 f3). -- rewrite !umEX !eqE !ftE UM.umapE /UM.union. - case: (um_from f1) (um_from f2) (um_from f3)=>[|f1' H1][|f2' H2][|f3' H3] //. - case E1: (disj f2' f3'); last first. - - by case E2: (disj f1' f3')=>//; rewrite disj_fcat E1 andbF. - rewrite disj_fcat andbC. - case E2: (disj f1' f3')=>//; rewrite disj_fcat (disjC f2') E1 /= andbT. - by case E3: (disj f1' f2')=>//; rewrite fcatAC // E1 E2 E3. -split=>[//|f1 f2 f3|f|f1 f2||f]. -- by rewrite (joinC f2) joinCA joinC. -- rewrite !umEX !ftE -[RHS]tfE eqE UM.umapE /UM.union /UM.empty. - by case: (um_from f)=>[//|f' H]; rewrite disjC disj_nil fcat0s. -- by rewrite !umEX !ftE; case: (um_from f1). -- by rewrite !umEX ftE. -rewrite !umEX /= /UM.empty /UM.empb -{1}[f]tfE. -case: (um_from f)=>[|f' F]; first by apply: ReflectF; rewrite eqE. -case: eqP=>E; constructor; rewrite eqE UM.umapE. -- by move/supp_nilE: E=>->. -by move=>H; rewrite H in E. -Qed. - -HB.instance Definition _ := isPCM.Build T union_map_is_pcm. - -Lemma union_map_is_cpcm : cpcm_axiom T. -Proof. -move=>f1 f2 f. -rewrite -{3}[f1]tfE -{2}[f2]tfE !pcmE /= !umEX /= !ftE !eqE. -rewrite !UM.umapE /UM.valid /UM.union. -case: (um_from f) (um_from f1) (um_from f2)=> -[|f' H]; case=>[|f1' H1]; case=>[|f2' H2] //; -case: ifP=>//; case: ifP=>// D1 D2 _ E. -by apply: fcatsK E; rewrite D1 D2. -Qed. - -HB.instance Definition _ := isCPCM.Build T union_map_is_cpcm. - -Lemma union_map_is_tpcm : tpcm_axiom um_undef um_undefb. -Proof. -split=>[/= x||/= x]. -- rewrite !umEX /= /UM.undefb -{1}[x]tfE. - case: (um_from x)=>[|f' F]; first by apply: ReflectT. - by apply: ReflectF; rewrite eqE. -- by rewrite pcmE /= !umEX ftE. -by rewrite pcmE /= !umEX ftE. -Qed. - -HB.instance Definition _ := isTPCM.Build T union_map_is_tpcm. - -Lemma union_map_is_normal : @normal_tpcm_axiom T. -Proof. -move=>/= x; rewrite pcmE /= /undef /= -{2}[x]tfE !umEX /UM.valid. -by case: (um_from x)=>[|f' H]; [right|left]. -Qed. - -HB.instance Definition _ := isNormal_TPCM.Build T union_map_is_normal. -HB.instance Definition _ := isUnionMap_helper.Build K C V T union_map_subproof. -HB.end. - -(* Making pts infer union_map structure *) -Definition ptsx K C V (U : union_map C V) k v of phant U : U := - @pts_op K C V U k v. - -(* use ptsT to pass map type explicitly *) -(* use pts when type inferrable or for printing *) -Notation ptsT U k v := (ptsx k v (Phant U)) (only parsing). -Notation pts k v := (ptsT _ k v). - -(*************************************) -(* Union map with decidable equality *) -(*************************************) - -(* union_map has decidable equality if V does *) - -Section UnionMapEq. -Variables (K : ordType) (C : pred K) (V : eqType). -Variable (U : union_map C V). - -Definition union_map_eq (f1 f2 : U) := - match (UMC_from f1), (UMC_from f2) with - | UM.Undef, UM.Undef => true - | UM.Def f1' _, UM.Def f2' _ => f1' == f2' - | _, _ => false - end. - -Lemma union_map_eqP : Equality.axiom union_map_eq. -Proof. -move=>x y; rewrite -{1}[x]tfE -{1}[y]tfE /union_map_eq. -case: (UMC_from x)=>[|f1 H1]; case: (UMC_from y)=>[|f2 H2] /=. -- by constructor. -- by constructor=>/eqE. -- by constructor=>/eqE. -case: eqP=>E; constructor; rewrite eqE. -- by rewrite UM.umapE. -by case. -Qed. - -(* no canonical projection to latch onto, so no generic inheritance *) -(* but unionmap_eqP can be used for any individual U *) -(* -HB.instance Definition _ := hasDecEq.Build U union_map_eqP. -*) -End UnionMapEq. - - -(*********************************) -(* Instance of union maps *) -(* with trivially true condition *) -(*********************************) - -Record umap K V := UMap {umap_base : @UM.base K xpredT V}. - -Section UmapUMC. -Variables (K : ordType) (V : Type). -Implicit Type f : umap K V. -Local Coercion umap_base : umap >-> UM.base. - -Let um_valid f := @UM.valid K xpredT V f. -Let um_empty := UMap (@UM.empty K xpredT V). -Let um_undef := UMap (@UM.Undef K xpredT V). -Let um_upd k v f := UMap (@UM.upd K xpredT V k v f). -Let um_dom f := @UM.dom K xpredT V f. -Let um_assocs f := @UM.assocs K xpredT V f. -Let um_free f k := UMap (@UM.free K xpredT V f k). -Let um_find k f := @UM.find K xpredT V k f. -Let um_union f1 f2 := UMap (@UM.union K xpredT V f1 f2). -Let um_empb f := @UM.empb K xpredT V f. -Let um_undefb f := @UM.undefb K xpredT V f. -Let um_from (f : umap K V) : UM.base _ _ := f. -Let um_to (b : @UM.base K xpredT V) : umap K V := UMap b. -Let um_pts k v := UMap (@UM.pts K xpredT V k v). - -Lemma umap_is_umc : - union_map_axiom um_valid um_empty um_undef um_upd um_dom - um_assocs um_free um_find um_union um_empb - um_undefb um_pts um_from um_to. -Proof. by split=>//; split=>[|[]]. Qed. - -HB.instance Definition _ := isUnion_map.Build K xpredT V (umap K V) umap_is_umc. -End UmapUMC. - -(* if V is eqtype so is umap K V *) -HB.instance Definition _ (K : ordType) (V : eqType) := - hasDecEq.Build (umap K V) (@union_map_eqP K xpredT V (umap K V)). - -Notation "x \\-> v" := (ptsT (umap _ _) x v) (at level 30). - -(* DEVCOMMENT *) -(* remove these "tests" for release *) -(* 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). -rewrite joinC. -Abort. - -(* can we use the base type? *) -Lemma xx (x : umap nat nat) : x \+ x == x \+ x. -Abort. - -(* can maps be stored into maps without universe inconsistencies? *) -(* yes, the idea of the class works *) -Lemma xx : 1 \\-> (1 \\-> 3) = 2 \\-> (7 \\-> 3). -Abort. - -(* /DEVCOMMENT *) - -(***************) -(* Finite sets *) -(***************) - -(* like umaps with unit range *) -(* but we give them their own copy of the type *) - -Record fset K := FSet {fset_base : @UM.base K xpredT unit}. - -Section FsetUMC. -Variable K : ordType. -Implicit Type f : fset K. -Local Coercion fset_base : fset >-> UM.base. - -Let fs_valid f := @UM.valid K xpredT unit f. -Let fs_empty := FSet (@UM.empty K xpredT unit). -Let fs_undef := FSet (@UM.Undef K xpredT unit). -Let fs_upd k v f := FSet (@UM.upd K xpredT unit k v f). -Let fs_dom f := @UM.dom K xpredT unit f. -Let fs_assocs f := @UM.assocs K xpredT unit f. -Let fs_free f k := FSet (@UM.free K xpredT unit f k). -Let fs_find k f := @UM.find K xpredT unit k f. -Let fs_union f1 f2 := FSet (@UM.union K xpredT unit f1 f2). -Let fs_empb f := @UM.empb K xpredT unit f. -Let fs_undefb f := @UM.undefb K xpredT unit f. -Let fs_from (f : fset K) : UM.base _ _ := f. -Let fs_to (b : @UM.base K xpredT unit) : fset K := FSet b. -Let fs_pts k v := FSet (@UM.pts K xpredT unit k v). - -Lemma fset_is_umc : - union_map_axiom fs_valid fs_empty fs_undef fs_upd fs_dom - fs_assocs fs_free fs_find fs_union fs_empb - fs_undefb fs_pts fs_from fs_to. -Proof. by split=>//; split=>[|[]]. Qed. - -HB.instance Definition _ := - isUnion_map.Build K xpredT unit (fset K) fset_is_umc. -HB.instance Definition _ := - hasDecEq.Build (fset K) (@union_map_eqP K xpredT unit (fset K)). -End FsetUMC. - -Notation "# x" := (ptsT (fset _) x tt) (at level 30, format "# x"). - -(* DEVCOMMENT *) -(* test *) -(* /DEVCOMMENT *) -Lemma xx : #1 \+ #2 = Unit. -Abort. - -(*******) -(* dom *) -(*******) - -Notation "[ 'dom' f ]" := [mem (dom f)] : fun_scope. - -Section DomLemmas. -Variables (K : ordType) (C : pred K) (V : Type) (U : union_map C V). -Implicit Types (k : K) (v : V) (f g : U). - -Lemma dom_undef : dom (undef : U) = [::]. -Proof. by rewrite !umEX. Qed. - -Lemma dom0 : dom (Unit : U) = [::]. -Proof. by rewrite !umEX. Qed. - -Lemma dom0E f : - valid f -> - dom f =i pred0 -> - f = Unit. -Proof. -rewrite !umEX /UM.valid /UM.dom /UM.empty -{3}[f]tfE !eqE UM.umapE. -case: (UMC_from f)=>[|f' S] //= _; rewrite fmapE /= {S}. -by case: f'; case=>[|kv s] //= P /= /(_ kv.1); rewrite inE eq_refl. -Qed. - -Lemma domNE f : reflect (dom f = [::]) (unitb f || undefb f). -Proof. -case: (normalP0 f)=>[->|->|W N]; constructor; rewrite ?dom0 ?dom_undef //. -by move=>M; apply: N; apply: dom0E W _; rewrite M. -Qed. - -Lemma dom0NP f : reflect (exists k, k \in dom f) (dom f != [::]). -Proof. -case: has_nilP=>X; constructor. -- by case/hasP: X=>x; exists x. -by case=>k D; apply: X; apply/hasP; exists k. -Qed. - -Lemma domU k v f : - dom (upd k v f) =i - [pred x | C k & if x == k then valid f else x \in dom f]. -Proof. -rewrite !umEX /UM.dom /UM.upd /UM.valid /=. -case: (UMC_from f)=>[|f' H] /= x. -- by rewrite !inE andbC; case: ifP. -case: decP=>[->|/(introF idP)->//]. -by rewrite supp_ins. -Qed. - -Lemma domF k f : - 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 /=; -by case: ifP=>// E; rewrite supp_rem inE /= eq_sym E. -Qed. - -Lemma subdomF k f : {subset dom (free f k) <= dom f}. -Proof. by move=>x; rewrite domF inE; case: eqP. Qed. - -Lemma domUn f1 f2 : - dom (f1 \+ f2) =i - [pred x | valid (f1 \+ f2) & (x \in dom f1) || (x \in dom f2)]. -Proof. -rewrite !umEX /UM.dom /UM.valid /UM.union. -case: (UMC_from f1) (UMC_from f2)=>[|f1' H1] // [|f2' H2] // x. -by case: ifP=>E //; rewrite supp_fcat. -Qed. - -(* bidirectional version of domUn *) -Lemma domUnE f1 f2 x : - valid (f1 \+ f2) -> - x \in dom (f1 \+ f2) = (x \in dom f1) || (x \in dom f2). -Proof. by move=>W; rewrite domUn inE W. Qed. - -Lemma dom_valid k f : k \in dom f -> valid f. -Proof. by rewrite !umEX; case: (UMC_from f). Qed. - -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. - -Lemma dom_inIR k f1 f2 : - valid (f1 \+ f2) -> k \in dom f2 -> k \in dom (f1 \+ f2). -Proof. by rewrite joinC; apply: dom_inIL. Qed. - -Lemma dom_NNL k f1 f2 : - valid (f1 \+ f2) -> k \notin dom (f1 \+ f2) -> k \notin dom f1. -Proof. by move=> D; apply/contra/dom_inIL. Qed. - -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 f k = f. -Proof. -rewrite -{3}[f]tfE !umEX /UM.dom /UM.free /=. -by case: (UMC_from f)=>[|f' H] //; apply: rem_supp. -Qed. - -Variant dom_find_spec k f : option V -> bool -> Type := -| dom_find_none'' of find k f = None : dom_find_spec k f None false -| dom_find_some'' v of find k f = Some v & - f = upd k v (free f k) : dom_find_spec k f (Some v) true. - -Lemma dom_find k f : dom_find_spec k f (find k f) (k \in dom f). -Proof. -rewrite !umEX /UM.dom -{1}[f]tfE. -case: (UMC_from f)=>[|f' H]. -- by apply: dom_find_none''; rewrite !umEX. -case: suppP (allP H k)=>[v|] /[dup] H1 /= -> I; last first. -- by apply: dom_find_none''; rewrite !umEX. -apply: (dom_find_some'' (v:=v)); rewrite !umEX // /UM.upd /UM.free. -case: decP=>H2; last by rewrite I in H2. -apply/fmapP=>k'; rewrite fnd_ins. -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. - -Lemma find_none k f : find k f = None <-> k \notin dom f. -Proof. by case: dom_find=>// v ->. 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) -> - f1 \+ g1 = f2 \+ g2 -> - dom f1 =i dom f2 -> f1 = f2. -Proof. -rewrite -{4}[f1]tfE -{3}[f2]tfE /= !umEX. -rewrite /UM.valid /UM.union /UM.dom; move=>D E. -case: (UMC_from f1) (UMC_from f2) (UMC_from g1) (UMC_from g2) E D=> -[|f1' F1][|f2' F2][|g1' G1][|g2' G2] //; -case: ifP=>// D1; case: ifP=>// D2 E _ E1; apply/fmapP=>k. -move/(_ k): E1=>E1. -case E2: (k \in supp f2') in E1; last first. -- by move/negbT/fnd_supp: E1=>->; move/negbT/fnd_supp: E2=>->. -suff L: forall m s, k \in supp m -> disj m s -> - fnd k m = fnd k (fcat m s) :> option V. -- by rewrite (L _ _ E1 D1) (L _ _ E2 D2) E. -by move=>m s S; case: disjP=>//; move/(_ _ S)/negbTE; rewrite fnd_fcat=>->. -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. -move=>Vh; apply: uniq_perm; last 1 first. -- by move=>x; rewrite mem_cat domUn inE Vh. -- by apply: uniq_dom. -rewrite cat_uniq !uniq_dom /= andbT; apply/hasPn=>x. -rewrite !umEX /UM.valid /UM.union /UM.dom in Vh *. -case: (UMC_from f1) (UMC_from f2) Vh=>// f1' H1 [//|f2' H2]. -by case: disjP=>// H _; apply: contraL (H x). -Qed. - -Lemma size_domUn f1 f2 : - valid (f1 \+ f2) -> - size (dom (f1 \+ f2)) = size (dom f1) + size (dom f2). -Proof. by move/perm_domUn/seq.perm_size; rewrite size_cat. Qed. - -Lemma size_domF k f : - k \in dom f -> size (dom (free f k)) = (size (dom f)).-1. -Proof. -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 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. -rewrite !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. - -End DomLemmas. - -Arguments subdomF {K C V U k f}. - -#[export] -Hint Resolve sorted_dom uniq_dom all_dom : core. -Prenex Implicits find_some find_none subdomF. - -(* lemmas for comparing doms of two differently-typed maps *) -Section DomLemmas2. -Variables (K : ordType) (C1 C2 : pred K) (V1 V2 : Type). -Variables (U1 : union_map C1 V1) (U2 : union_map C2 V2). - -Lemma domE (f1 : U1) (f2 : U2) : dom f1 =i dom f2 <-> dom f1 = dom f2. -Proof. by split=>[|->] //; apply/ord_sorted_eq/sorted_dom/sorted_dom. Qed. - -Lemma domKi (f1 g1 : U1) (f2 g2 : U2) : - valid (f1 \+ g1) -> valid (f2 \+ g2) -> - dom (f1 \+ g1) =i dom (f2 \+ g2) -> - dom f1 =i dom f2 -> dom g1 =i dom g2. -Proof. -rewrite !umEX /UM.valid /UM.union /UM.dom. -case: (UMC_from f1) (UMC_from f2) (UMC_from g1) (UMC_from g2)=> -[|f1' F1][|f2' F2][|g1' G1][|g2' G2] //. -case: disjP=>// D1 _; case: disjP=>// D2 _ E1 E2 x. -move: {E1 E2} (E2 x) (E1 x); rewrite !supp_fcat !inE /= => E. -move: {D1 D2} E (D1 x) (D2 x)=><- /implyP D1 /implyP D2. -case _ : (x \in supp f1') => //= in D1 D2 *. -by move/negbTE: D1=>->; move/negbTE: D2=>->. -Qed. - -Lemma domK (f1 g1 : U1) (f2 g2 : U2) : - valid (f1 \+ g1) -> valid (f2 \+ g2) -> - dom (f1 \+ g1) = dom (f2 \+ g2) -> - dom f1 = dom f2 -> dom g1 = dom g2. -Proof. by move=>D1 D2 /domE H1 /domE H2; apply/domE; apply: domKi H2. Qed. - -Lemma domKUni (f1 g1 : U1) (f2 g2 : U2) : - valid (f1 \+ g1) = valid (f2 \+ g2) -> - dom f1 =i dom f2 -> dom g1 =i dom g2 -> - dom (f1 \+ g1) =i dom (f2 \+ g2). -Proof. by move=>E D1 D2 x; rewrite !domUn !inE E D1 D2. Qed. - -Lemma domKUn (f1 g1 : U1) (f2 g2 : U2) : - valid (f1 \+ g1) = valid (f2 \+ g2) -> - dom f1 = dom f2 -> dom g1 = dom g2 -> - dom (f1 \+ g1) = dom (f2 \+ g2). -Proof. by move=>E /domE D1 /domE D2; apply/domE/domKUni. Qed. - -Lemma subdom_filter (f1 : U1) (f2 : U2) : - {subset dom f1 <= dom f2} -> - dom f1 = filter [dom f1] (dom f2). -Proof. -have Tx : transitive (@ord K) by apply: trans. -move=>S; apply: (sorted_eq (leT := @ord K))=>//. -- by apply: antisym. -- by apply: sorted_dom. -- by rewrite sorted_filter // sorted_dom. -apply: uniq_perm; rewrite ?filter_uniq ?uniq_dom // => y. -rewrite mem_filter /=; case E: (y \in dom _)=>//=. -by move/S: E=>->. -Qed. - -End DomLemmas2. - -(*********) -(* valid *) -(*********) - -Section ValidLemmas. -Variables (K : ordType) (C : pred K) (V : Type) (U : union_map C V). -Implicit Types (k : K) (v : V) (f g : U). - -Lemma invalidE f : ~~ valid f <-> f = undef. -Proof. by rewrite undefVN negbK; case: undefbP. Qed. - -Lemma validU k v f : valid (upd k v f) = C k && valid f. -Proof. -rewrite !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 f k) = valid f. -Proof. by rewrite !umEX; case: (UMC_from f). Qed. - -Variant validUn_spec f1 f2 : bool -> Type := -| valid_false1 of ~~ valid f1 : validUn_spec f1 f2 false -| valid_false2 of ~~ valid f2 : validUn_spec f1 f2 false -| valid_false3 k of k \in dom f1 & k \in dom f2 : validUn_spec f1 f2 false -| valid_true of valid f1 & valid f2 & - (forall x, x \in dom f1 -> x \notin dom f2) : validUn_spec f1 f2 true. - -Lemma validUn f1 f2 : validUn_spec f1 f2 (valid (f1 \+ f2)). -Proof. -rewrite !umEX -{1}[f1]tfE -{1}[f2]tfE. -rewrite /UM.valid /UM.union /=. -case: (UMC_from f1) (UMC_from f2)=>[|f1' F1][|f2' F2] /=. -- by apply: valid_false1; rewrite !umEX. -- by apply: valid_false1; rewrite !umEX. -- by apply: valid_false2; rewrite !umEX. -case: ifP=>E. -- apply: valid_true; try by rewrite !umEX. - by move=>k; rewrite !umEX => H; case: disjP E=>//; move/(_ _ H). -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 f1 k \+ f2). -Proof. -case: validUn=>// V1 V2 H _; case: validUn=>//. -- by rewrite validF V1. -- by rewrite V2. -by move=>k'; rewrite domF inE; case: eqP=>// _ /H/negbTE ->. -Qed. - -Lemma validUnF k f1 f2 : - valid (f1 \+ f2) -> valid (f1 \+ free f2 k). -Proof. by rewrite !(joinC f1); apply: validFUn. Qed. - -Lemma validUnU k v f1 f2 : - k \in dom f2 -> valid (f1 \+ upd k v f2) = valid (f1 \+ f2). -Proof. -move=>D; apply/esym; move: D; case: validUn. -- by move=>V' _; apply/esym/negbTE; apply: contra V'; move/validL. -- move=>V' _; apply/esym/negbTE; apply: contra V'; move/validR. - by rewrite validU; case/andP. -- move=>k' H1 H2 _; case: validUn=>//; rewrite validU => D1 /andP [/= H D2]. - by move/(_ k' H1); rewrite domU !inE H D2 H2; case: ifP. -move=>V1 V2 H1 H2; case: validUn=>//. -- by rewrite V1. -- by rewrite validU V2 (dom_cond H2). -move=>k'; rewrite domU (dom_cond H2) inE /= V2; move/H1=>H3. -by rewrite (negbTE H3); case: ifP H2 H3=>// /eqP ->->. -Qed. - -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 dom_inNL k f1 f2 : - valid (f1 \+ f2) -> k \in dom f1 -> k \notin dom f2. -Proof. by case: validUn=>//_ _ H _; apply: H. Qed. - -Lemma dom_inNR k f1 f2 : - valid (f1 \+ f2) -> k \in dom f2 -> k \notin dom f1. -Proof. by rewrite joinC; apply: dom_inNL. Qed. - -Lemma dom_inNLX k f1 f2 : - valid (f1 \+ f2) -> k \in dom f1 -> ~ k \in dom f2. -Proof. by move=>W /(dom_inNL W)/negbTE ->. Qed. - -Lemma dom_inNRX k f1 f2 : - valid (f1 \+ f2) -> k \in dom f2 -> ~ k \in dom f1. -Proof. by move=>W /(dom_inNR W)/negbTE ->. Qed. - -(* a more symmetric version of the above *) -Lemma dom_inN k1 k2 f1 f2 : - valid (f1 \+ f2) -> - k1 \in dom f1 -> k2 \in dom f2 -> k1 != k2. -Proof. by case: (k1 =P k2) =>// <- W H /(dom_inNLX W H). Qed. - -Lemma dom_inNX k1 k2 f1 f2 : - valid (f1 \+ f2) -> - k1 \in dom f1 -> k2 \in dom f2 -> k1 <> k2. -Proof. by move=>W K1 K2; apply/eqP; apply: dom_inN K1 K2. Qed. - -End ValidLemmas. - - -(************) -(* um_unitb *) -(************) - -(* AKA empb *) - -Section UnitbLemmas. -Variables (K : ordType) (C : pred K) (V : Type) (U : union_map C V). -Implicit Types (k : K) (v : V) (f : U). - -Lemma um_unitbU k v f : unitb (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 um_unitbUn f1 f2 : unitb (f1 \+ f2) = unitb f1 && unitb f2. -Proof. -rewrite !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 umap0E f1 f2 : f1 \+ f2 = Unit -> f1 = Unit /\ f2 = Unit. -Proof. -by move/unitbP; rewrite um_unitbUn=>/andP [H1 H2]; split; apply/unitbP. -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 !umEX /UM.valid /UM.dom /UM.empty -{3}[f]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/unitbP: E=>->; rewrite unitL valid_unit. -case: validUn=>// W _ H; move/validEb: E; elim; split=>// k. -by apply/negP=>/[dup]/H/negbTE ->. -Qed. - -End UnitbLemmas. - - -(**********) -(* undefb *) -(**********) - -Section UndefbLemmas. -Variables (K : ordType) (C : pred K) (V : Type) (U : union_map C V). -Implicit Types (k : K) (v : V) (f : U). - -Lemma undefb_undef : undefb (undef : U). -Proof. by rewrite !umEX. Qed. - -Lemma undefbP f : reflect (f = undef) (undefb f). -Proof. -rewrite !umEX /UM.undefb -{1}[f]tfE. -by case: (UMC_from f)=>[|f' F]; constructor; rewrite !umEX. -Qed. - -Lemma undefbE f : f = undef <-> undefb f. -Proof. by case: undefbP. Qed. - -Lemma undefb0 : undefb (Unit : U) = false. -Proof. by case: undefbP=>// /esym/undef0. Qed. - -Lemma valid_undefbI f : valid f -> undefb f = false. -Proof. -move=>W; apply/negP=>/undefbP E. -by move: E W=>->; rewrite valid_undef. -Qed. - -End UndefbLemmas. - - -(**********) -(* dom_eq *) -(**********) - -Section DomEqDef. -Variables (K : ordType) (C1 C2 : pred K) (V1 V2 : Type). -Variable U1 : union_map C1 V1. -Variable U2 : union_map C2 V2. - -Definition dom_eq (f1 : U1) (f2 : U2) := - (valid f1 == valid f2) && (dom f1 == dom f2). - -End DomEqDef. - -Section DomEqLemmas. -Variables (K : ordType) (C1 C2 C3 : pred K) (V1 V2 V3 : Type). -Variable U1 : union_map C1 V1. -Variable U2 : union_map C2 V2. -Variable U3 : union_map C3 V3. - -Lemma domeqP (f1 : U1) (f2 : U2) : - reflect (valid f1 = valid f2 /\ dom f1 = dom f2) (dom_eq f1 f2). -Proof. by rewrite /dom_eq; apply/andPP/eqP/eqP. Qed. - -Lemma domeqVE (f1 : U1) (f2 : U2) : dom_eq f1 f2 -> valid f1 = valid f2. -Proof. by case/domeqP. Qed. - -Lemma domeqDE (f1 : U1) (f2 : U2) : dom_eq f1 f2 -> dom f1 = dom f2. -Proof. by case/domeqP. Qed. - -Lemma domeqE (f1 : U1) (f2 : U2) : - dom_eq f1 f2 -> (valid f1 = valid f2) * (dom f1 = dom f2). -Proof. by move=>H; rewrite (domeqVE H) (domeqDE H). Qed. - -Lemma domeq0E (f : U1) : dom_eq f (Unit : U1) -> f = Unit. -Proof. -case/andP=>/eqP; rewrite valid_unit dom0=>H1 H2. -by apply: dom0E=>//; rewrite (eqP H2). -Qed. - -Lemma domeq_unit : dom_eq (Unit : U1) (Unit : U2). -Proof. by rewrite /dom_eq !valid_unit !dom0. Qed. - -Lemma domeq_undef : dom_eq (undef : U1) (undef : U2). -Proof. by rewrite /dom_eq !valid_undef !dom_undef. Qed. - -Lemma domeq_refl (f : U1) : dom_eq f f. -Proof. by rewrite /dom_eq !eqxx. Qed. - -Lemma domeq_sym (f1 : U1) (f2 : U2) : dom_eq f1 f2 -> dom_eq f2 f1. -Proof. by case/andP=>V D; apply/andP; rewrite (eqP V) (eqP D). Qed. - -Lemma domeq_trans (f1 : U1) (f2 : U2) (f3 : U3) : - dom_eq f1 f2 -> dom_eq f2 f3 -> dom_eq f1 f3. -Proof. by rewrite /dom_eq=>/andP [/eqP -> /eqP ->]. Qed. - -Lemma domeqVUnE (f1 f1' : U1) (f2 f2' : U2) : - dom_eq f1 f2 -> dom_eq f1' f2' -> - valid (f1 \+ f1') = valid (f2 \+ f2'). -Proof. -suff L (C1' C2' : pred K) V1' V2' (U1' : union_map C1' V1') - (U2' : union_map C2' V2') (g1 g1' : U1') (g2 g2' : U2') : - dom_eq g1 g2 -> dom_eq g1' g2' -> - valid (g1 \+ g1') -> valid (g2 \+ g2'). -- by move=>D D'; apply/idP/idP; apply: L=>//; apply: domeq_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. - -Lemma domeqVUn (f1 f1' : U1) (f2 f2' : U2) : - dom_eq f1 f2 -> dom_eq f1' f2' -> - valid (f1 \+ f1') -> valid (f2 \+ f2'). -Proof. by move=>D D'; rewrite -(domeqVUnE D D'). Qed. - -Lemma domeqUn (f1 f1' : U1) (f2 f2' : U2) : - dom_eq f1 f2 -> dom_eq f1' f2' -> - dom_eq (f1 \+ f1') (f2 \+ f2'). -Proof. -move/[dup]=>D /domeqP [V E] /[dup] D' /domeqP [V' E']. -move: (domeqVUnE D D')=>Ex; apply/domeqP=>//. -by rewrite Ex (domKUn Ex E E'). -Qed. - -Lemma domeqK (f1 f1' : U1) (f2 f2' : U2) : - valid (f1 \+ f1') -> - dom_eq (f1 \+ f1') (f2 \+ f2') -> - dom_eq f1 f2 -> dom_eq f1' f2'. -Proof. -move=>D1 /domeqP [E1 H1] /domeqP [E2 H2]; move: (D1); rewrite E1=>D2. -apply/domeqP; split; first by rewrite (validE2 D1) (validE2 D2). -by apply: domK D1 D2 H1 H2. -Qed. - -Lemma big_domeqUn A (xs : seq A) (f1 : A -> U1) (f2 : A -> U2) : - (forall x, x \In xs -> dom_eq (f1 x) (f2 x)) -> - dom_eq (\big[join/Unit]_(i <- xs) (f1 i)) - (\big[join/Unit]_(i <- xs) (f2 i)). -Proof. -elim: xs=>[|x xs IH] D; first by rewrite !big_nil domeq_unit. -rewrite !big_cons; apply: domeqUn. -- by apply: D; rewrite InE; left. -by apply: IH=>z Z; apply: D; rewrite InE; right. -Qed. - -End DomEqLemmas. - -#[export] -Hint Resolve domeq_refl : core. - -Section DomEq1Lemmas. -Variables (K : ordType) (C : pred K) (V : Type) (U : union_map C V). - -Lemma domeqfU k v (f : U) : k \in dom f -> dom_eq f (upd k v f). -Proof. -move=>D; rewrite /dom_eq validU (dom_cond D) (dom_valid D) /=. -apply/eqP/domE=>x; rewrite domU inE (dom_cond D) (dom_valid D). -by case: eqP=>// ->. -Qed. - -Lemma domeqUf k v (f : U) : k \in dom f -> dom_eq (upd k v f) f. -Proof. by move=>D; apply/domeq_sym/domeqfU. Qed. - -Lemma domeqfUn (f f1 f2 f1' f2' : U) : - dom_eq (f \+ f1) f2 -> - dom_eq f1' (f \+ f2') -> - dom_eq (f1 \+ f1') (f2 \+ f2'). -Proof. -move=>D1 D2; apply: (domeq_trans (f2:=f1 \+ f \+ f2')). -- by rewrite -joinA; apply: domeqUn. -by rewrite -joinA joinCA joinA; apply: domeqUn. -Qed. - -Lemma domeqUnf (f f1 f2 f1' f2' : U) : - dom_eq f1 (f \+ f2) -> dom_eq (f \+ f1') f2' -> - dom_eq (f1 \+ f1') (f2 \+ f2'). -Proof. -by move=>D1 D2; rewrite (joinC f1) (joinC f2); apply: domeqfUn D2 D1. -Qed. - -End DomEq1Lemmas. - -Section DomEq2Lemmas. -Variables (K : ordType) (C1 C2 : pred K) (V1 V2 : Type). -Variables (U1 : union_map C1 V1) (U2 : union_map C2 V2). - -Lemma domeq_symE (f1 : U1) (f2 : U2) : dom_eq f1 f2 = dom_eq f2 f1. -Proof. by apply/idP/idP; apply: domeq_sym. Qed. - -End DomEq2Lemmas. - - -(**********) -(* update *) -(**********) - -Section UpdateLemmas. -Variable (K : ordType) (C : pred K) (V : Type) (U : union_map C V). -Implicit Types (k : K) (v : V) (f : U). - -Lemma upd_undef k v : upd k v undef = undef :> U. -Proof. by rewrite !umEX. Qed. - -Lemma upd_condN k v f : ~~ C k -> upd k v f = undef. -Proof. -rewrite !umEX /UM.upd. -case: (UMC_from f)=>[|f' H']=>//. -by case: decP=>//= ->. -Qed. - -Lemma upd_inj k v1 v2 f : - valid f -> C k -> upd k v1 f = upd k v2 f -> v1 = v2. -Proof. -rewrite !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. -Qed. - -Lemma updU k1 k2 v1 v2 f : - upd k1 v1 (upd k2 v2 f) = - if k1 == k2 then upd k1 v1 f else upd k2 v2 (upd k1 v1 f). -Proof. -rewrite !umEX /UM.upd. -case: (UMC_from f)=>[|f' H']; case: ifP=>// E; -case: decP=>H1; case: decP=>H2 //; rewrite !umEX. -- by rewrite ins_ins E. -- by rewrite (eqP E) in H2 *. -by rewrite ins_ins E. -Qed. - -Lemma updF k1 k2 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; -by case: decP=>H2 //; rewrite !umEX ins_rem H1. -Qed. - -Lemma updUnL k v f1 f2 : - upd k v (f1 \+ f2) = - if k \in dom f1 then upd k v f1 \+ f2 else f1 \+ upd k v f2. -Proof. -rewrite !umEX /UM.upd /UM.union /UM.dom. -case: (UMC_from f1) (UMC_from f2)=>[|f1' F1][|f2' F2] //=. -- by case: ifP=>//; case: decP. -case: ifP=>// D; case: ifP=>// H1; case: decP=>// H2. -- rewrite disjC disj_ins disjC D !umEX; case: disjP D=>// H _. - by rewrite (H _ H1) /= fcat_inss // (H _ H1). -- by rewrite disj_ins H1 D /= !umEX fcat_sins. -- by rewrite disjC disj_ins disjC D andbF. -by rewrite disj_ins D andbF. -Qed. - -Lemma updUnR k v f1 f2 : - upd 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. - -Lemma domUE k e f : k \in dom f -> dom (upd k e f) = dom f. -Proof. -move=>H; apply/domE=>x; rewrite domU inE (dom_cond H) /=. -by case: eqP=>// ->; rewrite H (dom_valid H). -Qed. - -End UpdateLemmas. - - -(********) -(* free *) -(********) - -Section FreeLemmas. -Variables (K : ordType) (C : pred K) (V : Type) (U : union_map C V). -Implicit Types (k : K) (v : V) (f : U). - -Lemma free0 k : free Unit k = Unit :> U. -Proof. by rewrite !umEX /UM.free /UM.empty rem_empty. Qed. - -Lemma free_undef k : free undef k = undef :> U. -Proof. by rewrite !umEX. Qed. - -Lemma freeU k1 k2 v f : - free (upd k1 v f) k2 = - if k1 == k2 then - if C k1 then free f k2 else undef - else upd k1 v (free f k2). -Proof. -rewrite !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 eq_sym H1. -- by case: ifP H2. -by rewrite !umEX rem_ins eq_sym H1. -Qed. - -Lemma freeF k1 k2 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 (f1 \+ f2) k = - if k \in dom (f1 \+ f2) then free f1 k \+ free f2 k - else f1 \+ f2. -Proof. -rewrite !umEX /UM.free /UM.union /UM.dom. -case: (UMC_from f1) (UMC_from f2)=>[|f1' F1][|f2' F2] //. -case: ifP=>// E1; rewrite supp_fcat inE /=. -case: ifP=>E2; last by rewrite !umEX rem_supp // supp_fcat inE E2. -rewrite disj_rem; last by rewrite disjC disj_rem // disjC. -rewrite !umEX; case/orP: E2=>E2. -- suff E3: k \notin supp f2' by rewrite -fcat_rems // (rem_supp E3). - by case: disjP E1 E2=>// H _; move/H. -suff E3: k \notin supp f1' by rewrite -fcat_srem // (rem_supp E3). -by case: disjP E1 E2=>// H _; move/contra: (H k); rewrite negbK. -Qed. - -Lemma freeUnV k f1 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 freeUnR 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 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 => ->. -move=>x H3; move: (H _ H3); rewrite domF inE /=. -by case: ifP H3 V1=>[|_ _ _]; [move/eqP=><- -> | move/negbTE=>->]. -Qed. - -Lemma freeUnL k f1 f2 : - k \notin dom f2 -> - free (f1 \+ f2) k = free f1 k \+ f2. -Proof. by move=>H; rewrite joinC freeUnR // joinC. Qed. - -(* positive re-statement of freeUnL, but requires validity *) -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 freeUnL. Qed. - -(* positive re-statement of freeUnR, but requires validity *) -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. - -Lemma freeND k f : k \notin dom f -> free f k = f. -Proof. by move=>W; rewrite -[f]unitR freeUnR // free0. Qed. - -End FreeLemmas. - - -(********) -(* find *) -(********) - -Section FindLemmas. -Variables (K : ordType) (C : pred K) (V : Type) (U : union_map C V). -Implicit Types (k : K) (v : V) (f : U). - -Lemma find0E k : find k (Unit : U) = None. -Proof. by rewrite !umEX /UM.find /= fnd_empty. Qed. - -Lemma find_undef k : find k (undef : U) = None. -Proof. by rewrite !umEX /UM.find. Qed. - -Lemma find_cond k f : ~~ C k -> find k f = None. -Proof. -rewrite !umEX /UM.find; case: (UMC_from f)=>[|f' F] // H. -suff: k \notin supp f' by case: suppP. -by apply: contra H; case: allP F=>// F _ /F. -Qed. - -Lemma findU k1 k2 v f : - find k1 (upd k2 v f) = - if k1 == k2 then - if C k2 && valid f then Some v else None - else if C k2 then find k1 f else None. -Proof. -rewrite !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 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. -by rewrite fnd_rem. -Qed. - -Lemma findUnL k f1 f2 : - valid (f1 \+ f2) -> - find k (f1 \+ f2) = if k \in dom f1 then find k f1 else find k f2. -Proof. -rewrite !umEX /UM.valid /UM.find /UM.union /UM.dom. -case: (UMC_from f1) (UMC_from f2)=>[|f1' F1][|f2' F2] //. -case: ifP=>// D _; case: ifP=>E1; last first. -- rewrite fnd_fcat; case: ifP=>// E2. - by rewrite fnd_supp ?E1 // fnd_supp ?E2. -suff E2: k \notin supp f2' by rewrite fnd_fcat (negbTE E2). -by case: disjP D E1=>// H _; apply: H. -Qed. - -Lemma findUnR k f1 f2 : - valid (f1 \+ f2) -> - find k (f1 \+ f2) = if k \in dom f2 then find k f2 else find k f1. -Proof. by rewrite joinC=>D; rewrite findUnL. Qed. - -Lemma find_eta f1 f2 : - valid f1 -> valid f2 -> - (forall k, find k f1 = find k f2) -> f1 = f2. -Proof. -rewrite !umEX /UM.valid /UM.find -{2 3}[f1]tfE -{2 3}[f2]tfE. -case: (UMC_from f1) (UMC_from f2)=>[|f1' F1][|f2' F2] // _ _ H. -by rewrite !umEX; apply/fmapP=>k; move: (H k); rewrite !umEX. -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) (C : pred K) (U : union_map C unit) (f1 f2 : U) : - dom_eq f1 f2 -> f1 = f2. -Proof. -case/domeqP=>V E; case V1 : (valid f1); last first. -- by move: V1 (V1); rewrite {1}V; do 2![move/negbT/invalidE=>->]. -apply: find_eta=>//; first by rewrite -V. -move=>k; case D : (k \in dom f1); move: D (D); rewrite {1}E. -- by do 2![case: dom_find=>// [[_ _]]]. -by do 2![case: dom_find=>// _]. -Qed. - -(**********) -(* assocs *) -(**********) - -Section AssocsLemmas. -Variables (K : ordType) (C : pred K) (V : Type) (U : union_map C V). -Implicit Type f : U. - -Lemma assocs_valid k f : k \In assocs f -> valid f. -Proof. by rewrite !umEX; case: (UMC_from f). Qed. - -Lemma assocs0 : assocs (Unit : U) = [::]. -Proof. by rewrite !umEX. Qed. - -Lemma assocs_undef : assocs (undef : U) = [::]. -Proof. by rewrite !umEX. Qed. - -Lemma assocsF f x : - assocs (free f x) = filter (fun kv => kv.1 != x) (assocs f). -Proof. by rewrite !umEX /UM.assocs; case: (UMC_from f)=>//=; case. Qed. - -Lemma assocs_perm f1 f2 : - valid (f1 \+ f2) -> perm (assocs (f1 \+ f2)) (assocs f1 ++ assocs f2). -Proof. -rewrite !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 assocs_dom f : dom f = map fst (assocs f). -Proof. by rewrite !umEX; case: (UMC_from f). Qed. - -Lemma size_assocs f : size (assocs f) = size (dom f). -Proof. by rewrite assocs_dom size_map. Qed. - -End AssocsLemmas. - -Lemma uniq_assocs K C (V : eqType) (U : @union_map K C V) (f : U) : - uniq (assocs f). -Proof. -rewrite !umEX /UM.assocs /=; case: (UMC_from f)=>[|[s H _]] //=. -by move/(sorted_uniq (@trans K) (@irr K)): H; apply: map_uniq. -Qed. - -Lemma assocs_map K C (V : Type) (U : @union_map K C V) (f : U) k v1 v2 : - (k, v1) \In assocs f -> (k, v2) \In assocs f -> v1 = v2. -Proof. -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. - -(*********************************) -(* Interaction of subset and dom *) -(*********************************) - -Section Interaction. -Variables (K : ordType) (C : pred K) (A : Type) (U : union_map C A). -Implicit Types (x y a b : U) (p q : pred K). - -(* Some equivalent forms for subset with dom *) - -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 *) - -Lemma subset_disjE p q : {subset p <= predC q} <-> [predI p & q] =1 pred0. -Proof. -split=>H a /=; first by case X: (a \in p)=>//; move/H/negbTE: X. -by move=>D; move: (H a); rewrite inE /= D; move/negbT. -Qed. - -Lemma valid_subdom x y : valid (x \+ y) -> {subset dom x <= [predC dom y]}. -Proof. by case: validUn. Qed. - -Lemma subdom_valid x y : - {subset dom x <= [predC dom 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_validL x y a : - valid a -> - valid (x \+ y) -> {subset dom a <= dom x} -> valid (a \+ y). -Proof. -rewrite !validUnAE=>-> /and3P [_ -> D] S. -by apply: sub_all D=>z /(contra (S z)). -Qed. - -End Interaction. - - -(****************************) -(* Generic points-to lemmas *) -(****************************) - -Section PointsToLemmas. -Variables (K : ordType) (C : pred K) (V : Type) (U : union_map 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 /ptsx !umEX /UM.pts /UM.upd; case: decP. Qed. - -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 /ptsx !umEX /= /UM.dom /supp /UM.pts /UM.upd /UM.empty /=. -by case D : (decP _)=>[a|a] /=; rewrite ?a ?(introF idP a). -Qed. - -(* a weaker legacy version of 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) = C k. -Proof. by rewrite ptsU validU valid_unit andbT. Qed. - -Lemma domPtV k v : valid (pts k v : U) -> k \in dom (pts k v : U). -Proof. by move=>D; rewrite domPtK -(validPt k v) D inE. Qed. - -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) && 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 & 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 : - C k1 -> - free (pts k1 v) k2 = if k1 == k2 then Unit else pts k1 v :> U. -Proof. by move=>N; rewrite ptsU freeU free0 N. Qed. - -Lemma freePt k v : - C k -> - free (pts k v : U) k = Unit. -Proof. by move=>N; rewrite freePt2 // eq_refl. Qed. - -Lemma cancelPt k v1 v2 : - valid (pts k v1 : U) -> - pts k v1 = pts k v2 :> U -> - v1 = v2. -Proof. by rewrite validPt !ptsU; apply: upd_inj. Qed. - -Lemma cancelPt2 k1 k2 v1 v2 : - valid (pts k1 v1 : U) -> - pts k1 v1 = pts k2 v2 :> U -> - (k1, v1) = (k2, v2). -Proof. -move=>H E; have : dom (pts k1 v1 : U) = dom (pts k2 v2 : U) by rewrite E. -rewrite !domPtK -(validPt _ v1) -(validPt _ v2) -E H. -by case=>E'; rewrite -{k2}E' in E *; rewrite (cancelPt H E). -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 um_unitbPt k v : unitb (pts k v : U) = false. -Proof. by rewrite ptsU um_unitbU. Qed. - -(* assocs *) - -Lemma assocsPt k v : - assocs (pts k v : U) = - if C k then [:: (k, v)] else [::]. -Proof. -rewrite /ptsx !umEX /UM.assocs/UM.pts /=. -by case E: decP=>[a|a] /=; [rewrite a | case: ifP]. -Qed. - -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. -rewrite /ptsx !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 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 /ptsx !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. - -(* valid *) - -Lemma validPtUn k v 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) //. - by rewrite domPtK H1 inE. -- by rewrite validPt; move/negbTE=>->. -- by move/negbTE=>->; rewrite andbF. -rewrite domPtK=>x; case: ifP=>// _. -by rewrite inE=>/eqP ->->; rewrite andbF. -Qed. - -Lemma validUnPt k v 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. - -Lemma validPt2 k1 k2 v1 v2 : - valid (pts k1 v1 \+ pts k2 v2 : U) = - [&& C k1, C k2 & k1 != k2]. -Proof. -rewrite validPtUn validPt domPt inE negb_and. -by case: (C k2)=>//=; rewrite eq_sym. -Qed. - -(* the projections from validPtUn are often useful *) - -Lemma validPtUnI v1 k f : - valid (pts k v1 \+ f) -> [&& C k, valid f & k \notin dom f]. -Proof. by rewrite validPtUn. 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 validPtUn_cond k v f : valid (pts k v \+ f) -> C k. -Proof. by rewrite validPtUn; case/and3P. 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. - -Lemma validUnPtD k v f : valid (f \+ pts k v) -> k \notin dom f. -Proof. by rewrite joinC; apply: validPtUnD. Qed. - -(* dom *) - -Lemma domPtUn k v f : - dom (pts k v \+ f) =i - [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: (C k). -Qed. - -Lemma domUnPt k v f : - dom (f \+ pts k v) =i - [pred x | valid (f \+ pts k v) & (k == x) || (x \in dom f)]. -Proof. by rewrite joinC; apply: domPtUn. Qed. - -Lemma domPtUnE k v f : k \in dom (pts k v \+ f) = valid (pts k v \+ f). -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 domPt2 k1 k2 v1 v2 x : - x \in dom (pts k1 v1 \+ pts k2 v2 : U) = - [&& C k1, C k2, k1 != k2 & x \in pred2 k1 k2]. -Proof. -rewrite domPtUn !inE validPt2 domPt inE !(eq_sym x). -by case: (C k1)=>//=; case: (C k2). -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: ord_sorted_eq=>//=. -- 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: ord_sorted_eq=>//=. -- 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. - -(* find *) - -Lemma findPtUn k v f : - valid (pts k v \+ f) -> find k (pts k v \+ f) = Some v. -Proof. -move=>V'; rewrite findUnL // domPt inE. -by rewrite ptsU findU eq_refl valid_unit (validPtUn_cond V'). -Qed. - -Lemma findUnPt k v f : - valid (f \+ pts k v) -> find k (f \+ pts k v) = Some v. -Proof. by rewrite joinC; apply: findPtUn. Qed. - -Lemma findPtUn2 k1 k2 v f : - valid (pts k2 v \+ f) -> - find k1 (pts k2 v \+ f) = - if k1 == k2 then Some v else find k1 f. -Proof. -move=>V'; rewrite findUnL // domPt inE eq_sym. -by rewrite findPt2 (validPtUn_cond V') andbT /=; case: ifP=>// ->. -Qed. - -Lemma findUnPt2 k1 k2 v f : - valid (f \+ pts k2 v) -> - find k1 (f \+ pts k2 v) = - if k1 == k2 then Some v else find k1 f. -Proof. by rewrite joinC; apply: findPtUn2. Qed. - -(* free *) - -Lemma freePtUn2 k1 k2 v f : - valid (pts k1 v \+ f) -> - free (pts k1 v \+ f) k2 = - if k1 == k2 then f else pts k1 v \+ free f k2. -Proof. -move=>V'; rewrite freeUn domPtUn inE V' /=. -rewrite ptsU freeU (validPtUn_cond V') free0. -case: (k1 =P k2)=>/= E. -- by rewrite -E unitL (dom_free (validPtUnD V')). -by case: ifP=>// N; rewrite dom_free // N. -Qed. - -Lemma freePtUn k v 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 (f \+ pts k v) k = f. -Proof. by rewrite joinC; apply: freePtUn. Qed. - -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. -Proof. -case V1 : (valid (pts k v2 \+ f)). -- by rewrite updUnL domPt inE eq_refl updPt (validPtUn_cond V1). -have V2 : valid (pts k v1 \+ f) = false. -- by rewrite !validPtUn in V1 *. -move/negbT/invalidE: V1=>->; move/negbT/invalidE: V2=>->. -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. - -Lemma upd_eta k v f : upd k v f = pts k v \+ free f k. -Proof. -case: (normalP f)=>[->|W]; first by 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. - -Lemma upd_eta2 k v f : - k \notin dom f -> - upd k v f = pts k v \+ f. -Proof. by move=>D; rewrite upd_eta freeND. Qed. - -(* um_unitb *) - -Lemma um_unitbPtUn k v f : unitb (pts k v \+ f) = false. -Proof. by rewrite um_unitbUn um_unitbPt. Qed. - -Lemma um_unitbUnPt k v f : unitb (f \+ pts k v) = false. -Proof. by rewrite joinC; apply: um_unitbPtUn. Qed. - -(* undef *) - -Lemma pts_undef k v : pts k v = undef :> U <-> ~~ C k. -Proof. by rewrite -(validPt k v) invalidE. Qed. - -Lemma pts_undef2 k v1 v2 : pts k v1 \+ pts k v2 = undef :> U. -Proof. -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 f k. -Proof. -case: dom_find=>// v E1 E2 _; exists v; split=>//. -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 f k. -Proof. by move=>E; case/um_eta: (find_some E)=>v' []; rewrite E; case=><-. Qed. - -Lemma cancel k v1 v2 f1 f2 : - valid (pts k v1 \+ f1) -> - pts k v1 \+ f1 = pts k v2 \+ f2 -> [/\ v1 = v2, valid f1 & f1 = f2]. -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) (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 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 f1 k2. -- move: (erefl (free (pts k1 v1 \+ f1) k2)). - rewrite {2}E freeUn E domPtUn inE V2 eq_refl /=. - by rewrite ptsU freeU X free0 -ptsU freePtUn. -suff E'' : free f1 k2 = free f2 k1. -- by rewrite -E''; rewrite E' joinCA in E; case/(cancel V1): E. -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 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 /ptsx !umEX => H1 H2 H3 f; rewrite -[f]tfE. -apply: (UM.base_indf (P := (fun b => P (UMC_to b))))=>//. -move=>k v g H V1; move: (H3 k v _ H); rewrite !umEX. -by apply. -Qed. - -(* backward progressing over keys *) -Lemma um_indb (P : U -> Prop) : - 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 /ptsx !umEX => H1 H2 H3 f; rewrite -[f]tfE. -apply: (UM.base_indb (P := (fun b => P (UMC_to b))))=>//. -move=>k v g H V1; move: (H3 k v _ H); rewrite !umEX. -by apply. -Qed. - -(* validity holds pairwise *) -Lemma um_valid3 f1 f2 f3 : - valid (f1 \+ f2 \+ f3) = - [&& valid (f1 \+ f2), valid (f2 \+ f3) & valid (f1 \+ f3)]. -Proof. -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 : - 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 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)). -- by rewrite -(validPt _ v) -E -joinA in X. -rewrite -[pts k v]unitR -joinA in E. -move/(cancel2 W): E; case: ifP. -- by move/eqP=>-> [->] /umap0E [->->]; rewrite unitR; left. -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. - -Prenex Implicits validPtUn_cond findPt_inv um_eta2 um_contra. -Prenex Implicits validPtUnD validUnPtD cancelPt cancelPt2 um_prime. - -(* dom_eq *) - -Section DomeqPtLemmas. -Variables (K : ordType) (C1 C2 : pred K) (V1 V2 : Type). -Variables (U1 : union_map C1 V1) (U2 : union_map C2 V2). - -Lemma domeqPt (k : K) (v1 : V1) (v2 : V2) : - C1 k = C2 k -> - dom_eq (pts k v1 : U1) (pts k v2 : U2). -Proof. by move=>E; apply/domeqP; rewrite !validPt !domPtK E. Qed. - -Lemma domeqPtUn k v1 v2 (f1 : U1) (f2 : U2) : - C1 k = C2 k -> - dom_eq f1 f2 -> dom_eq (pts k v1 \+ f1) (pts k v2 \+ f2). -Proof. by move=>E; apply: domeqUn=>//; apply: domeqPt. Qed. - -Lemma domeqUnPt k v1 v2 (f1 : U1) (f2 : U2) : - C1 k = C2 k -> - dom_eq f1 f2 -> dom_eq (f1 \+ pts k v1) (f2 \+ pts k v2). -Proof. by rewrite (joinC f1) (joinC f2); apply: domeqPtUn. Qed. - -End DomeqPtLemmas. - -#[export] -Hint Resolve domeqPt domeqPtUn domeqUnPt : core. - -(* decidable equality for umaps *) - -Section EqPtLemmas. -Variables (K : ordType) (C : pred K) (V : eqType). -Variables (U : union_map C V). -Notation Ue := - (Equality.pack_ (Equality.Mixin (union_map_eqP (U:=U)))). - -Lemma umPtPtE (k1 k2 : K) (v1 v2 : V) : - (pts k1 v1 == pts k2 v2 :> Ue) = - if C k1 then [&& C k2, k1 == k2 & v1 == v2] else ~~ C k2. -Proof. -case D1 : (C k1); last first. -- case D2 : (C k2); rewrite pts_condN ?D1 //. - - by apply/eqP/nesym; rewrite pts_undef D2. - by rewrite pts_condN ?D2 // eq_refl. -rewrite -(validPt U k1 v1) -(validPt U k2 v2) in D1 *. -apply/idP/idP. -- by case/eqP/(cancelPt2 D1)=><-<-; rewrite D1 !eq_refl. -by case/and3P=>_ /eqP -> /eqP ->. -Qed. - -Lemma umPtPtEq (k1 k2 : K) (v1 v2 : V) : - (pts k1 v1 == pts k2 v2 :> Ue) = - [&& C k1 == C k2, k1 == k2 & v1 == v2] || - (~~ C k1 && ~~ C k2). -Proof. -rewrite umPtPtE. -case D1 : (C k1); case D2 : (C k2)=>//=. -- by rewrite orbF. -by rewrite orbT. -Qed. - -Lemma umPt0E (k : K) (v : V) : (pts k v == Unit :> Ue) = false. -Proof. by apply/eqP=>/unitbP; rewrite um_unitbPt. Qed. - -Lemma um0PtE (k : K) (v : V) : (Unit == pts k v :> Ue) = false. -Proof. by rewrite eq_sym umPt0E. Qed. - -Lemma umPtUndefE (k : K) (v : V) : (pts k v == undef :> Ue) = ~~ C k. -Proof. by apply/idP/idP=>[|H]; [move/eqP/pts_undef|apply/eqP/pts_undef]. Qed. - -Lemma umUndefPtE (k : K) (v : V) : (undef == pts k v :> Ue) = ~~ C k. -Proof. by rewrite eq_sym umPtUndefE. Qed. - -Lemma umUndef0E : (undef == Unit :> Ue) = false. -Proof. by apply/eqP/undef0. Qed. - -Lemma um0UndefE : (Unit == undef :> Ue) = false. -Proof. by rewrite eq_sym umUndef0E. Qed. - -Lemma umPtUE (k : K) (v : V) f : (pts k v \+ f == Unit :> Ue) = false. -Proof. by apply/eqP=>/umap0E/proj1/unitbP; rewrite um_unitbPt. Qed. - -Lemma umUPtE (k : K) (v : V) f : (f \+ pts k v == Unit :> Ue) = false. -Proof. by rewrite joinC umPtUE. Qed. - -Lemma umPtUPtE (k1 k2 : K) (v1 v2 : V) f : - pts k1 v1 \+ f == pts k2 v2 :> Ue = - if C k1 then - if C k2 then [&& k1 == k2, v1 == v2 & unitb f] - else ~~ valid (pts k1 v1 \+ f) - else ~~ C k2. -Proof. -case D1 : (C k1); last first. -- rewrite pts_condN ?D1 //= undef_join eq_sym. - case D2 : (C k2). - - by apply/eqP=>/=; rewrite pts_undef D2. - by rewrite pts_condN ?D2 // eq_refl. -case D2 : (C k2); last first. -- rewrite (pts_condN U v2) ?D2 //=; apply/idP/idP. - - by move/eqP=>/undefbP; rewrite undefNV. - by move=>D; apply/eqP/undefbP; rewrite undefNV. -apply/idP/idP; last first. -- by case/and3P=>/eqP -> /eqP -> /unitbP ->; rewrite unitR. -case/eqP/(um_prime D2); case. -- move/(cancelPt2 _); rewrite validPt=>/(_ D1) [->->->]. - by rewrite !eq_refl unitb0. -by move/unitbP; rewrite um_unitbPt. -Qed. - -Lemma umPtPtUE (k1 k2 : K) (v1 v2 : V) f : - pts k1 v1 == pts k2 v2 \+ f :> Ue = - if C k2 then - if C k1 then [&& k1 == k2, v1 == v2 & unitb f] - else ~~ valid (pts k2 v2 \+ f) - else ~~ C k1. -Proof. by rewrite eq_sym umPtUPtE (eq_sym k1) (eq_sym v1). Qed. - -Lemma umUPtPtE (k1 k2 : K) (v1 v2 : V) f : - f \+ pts k1 v1 == pts k2 v2 :> Ue = - if C k1 then - if C k2 then [&& k1 == k2, v1 == v2 & unitb f] - else ~~ valid (pts k1 v1 \+ f) - else ~~ C k2. -Proof. by rewrite joinC umPtUPtE. Qed. - -Lemma umPtUPt2E (k1 k2 : K) (v1 v2 : V) f : - pts k1 v1 == f \+ pts k2 v2 :> Ue = - if C k2 then - if C k1 then [&& k1 == k2, v1 == v2 & unitb f] - else ~~ valid (pts k2 v2 \+ f) - else ~~ C k1. -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 U, @unitR U), (validPt, @valid_unit U))), - (((eq_refl, unitb0), - (um_unitbPt, undef_join)), (join_undef, unitb_undef))). -End EqPtLemmas. - - -(*****************) -(* Other notions *) -(*****************) - -(************) -(* um_preim *) -(************) - -Section PreimDefLemmas. -Variables (K : ordType) (C : pred K) (V : Type) (U : union_map C V). -Implicit Types (k : K) (v : V) (f : U) (p : pred V). - -Definition um_preim p f k := oapp p false (find k f). - -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. - -Lemma umpreim_cond p f k : um_preim p f k -> C k. -Proof. -rewrite /um_preim; case E: (find k f)=>[v|] //= _. -by move/find_some/dom_cond: E. -Qed. - -Lemma umpreimPt p k v : - C k -> - um_preim p (pts k v) =1 [pred x | (x == k) && p v]. -Proof. -move=>Hk x; rewrite /um_preim /= findPt2. -by case: eqP=>//= _; rewrite Hk. -Qed. - -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. - -Lemma umpreim_pred0 f : um_preim pred0 f =1 pred0. -Proof. by move=>x; rewrite /um_preim; by case: (find x f). Qed. - -Lemma umpreim_dom f : um_preim predT f =1 [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. - -End PreimDefLemmas. - - -(****************************) -(* map membership predicate *) -(****************************) - -Section MapMembership. -Variables (K : ordType) (C : pred K) (V : Type) (U : union_map C V). -Implicit Type f : U. - -Definition Mem_UmMap f := - fun x : K * V => valid f /\ [pcm pts x.1 x.2 <= f]. -Canonical um_PredType := PropPredType Mem_UmMap. - -Lemma In_undef x : x \In (undef : U) -> False. -Proof. by case; rewrite valid_undef. Qed. - -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. - -(* for inlined rewriting *) -Lemma In_findE k v f : (k, v) \In f -> find k f = Some v. -Proof. by move/In_find. 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. -case: x=>a b /In_find; rewrite findPt2. -by 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 InPtE k v f : - C k -> - f = pts k v -> (k, v) \In f. -Proof. by move=>W ->; apply: In_condPt W. 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_valid x f : x \In f -> valid f. -Proof. by case. Qed. - -Lemma In_domN x f : x \In f -> dom f <> [::]. -Proof. by move/In_dom=>H; apply/eqP/dom0NP; exists x.1. Qed. - -Lemma In_unitN x f : x \In f -> f <> Unit. -Proof. by move/[swap]=>-> /In0. Qed. - -Lemma In_inj_fun k1 k2 v f : - {in dom f, injective (fun x => find x f)} -> - (k1, v) \In f -> (k2, v) \In f -> k1 = k2. -Proof. -move=>H H1 H2; apply: H. -- by move/In_dom: H1. -by move/In_find: H1 H2=>-> /In_find ->. -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. - -(* this is just find_none stated as reflection *) -Lemma In_findN k f : reflect (find k f = None) (k \notin dom f). -Proof. -case: In_domX=>H; constructor. -- by case: H=>x /In_find ->. -by apply/find_none/negP=>/In_domX [v X]; apply: H; exists v. -Qed. - -Lemma In_findNE k f : k \notin dom f -> find k f = None. -Proof. by move/In_findN. 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. - -(* a frequently used pattern *) -Lemma InPtUnEL k v f f' : - valid f -> f = pts k v \+ f' -> (k, v) \In f. -Proof. by move/[swap] ->; apply: InPtUnL. 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 -> - [/\ [&& C k, valid f & k \notin dom 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). -rewrite (validPtUnI (W v)); 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 In_assocs 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. - -(* if we have equality of domains, we can get rid of one direction *) -(* in the hypothesis in umem_eq *) - -Lemma umem_eqD f1 f2 : - valid f1 -> valid f2 -> - dom f1 =i dom f2 -> - (forall x, x \In f1 -> x \In f2) -> f1 = f2. -Proof. -move=>V1 V2 E H; apply: umem_eq=>//; case=>k v; split; first by apply: H. -move=>H2; move: (In_dom H2); rewrite -E /= =>/In_domX [w H1]. -by move/H/(In_fun H2): (H1)=>->. -Qed. - -Lemma umem_eq_assocs f1 f2 : - valid f1 -> valid f2 -> - assocs f1 = assocs f2 -> - f1 = f2. -Proof. -move=>V1 V2 E; apply: umem_eqD V1 V2 _ _. -- by rewrite !assocs_dom E. -by move=>x /In_assocs; rewrite E=>/In_assocs. -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. - -(* different format that avoids conditionals *) -Lemma In_U x k v f : - x \In upd k v f <-> - valid (upd k v f) /\ - [\/ x.1 == k /\ x.2 = v | - x.1 != k /\ x \In f]. -Proof. -rewrite InU; split; case=>W H; split=>//. -- by case: eqP H; [left | right]. -by case: H=>[[->]//|[/negbTE ->]]. -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 InPtUnEL In_findE. - -(* Umap and fset are special cases of union_map but are *) -(* defined before membership structure is declared. *) -(* Their membership structures aren't directly inheritted from that *) -(* of union_map, but must be declared separately *) -Canonical umap_PredType (K : ordType) V : PredType (K * V) := - um_PredType (umap K V). -Coercion Pred_of_umap K V (x : umap K V) : {Pred _} := - [eta Mem_UmMap x]. -Canonical fset_PredType (K : ordType) : PredType (K * unit) := - um_PredType (fset K). -Coercion Pred_of_fset K (x : fset K) : {Pred _} := - [eta Mem_UmMap x]. - -Section MorphMembership. -Variables (K : ordType) (C : pred K) (V : Type). -Variables (U1 : pcm) (U2 : union_map C V). - -Section PlainMorph. -Variable f : pcm_morph U1 U2. - -Lemma InpfL e (x y : U1) : - valid (x \+ y) -> - sep f x y -> - e \In f x -> - e \In f (x \+ y). -Proof. by move=>W H1 H2; rewrite pfjoin //=; apply: InL=>//; rewrite pfV2. Qed. - -Lemma InpfR e (x y : U1) : - valid (x \+ y) -> - sep f x y -> - e \In f y -> - e \In f (x \+ y). -Proof. by move=>W H1 H2; rewrite pfjoin //=; apply: InR=>//; rewrite pfV2. Qed. - -Lemma InpfUn e (x y : U1) : - valid (x \+ y) -> - sep f x y -> - e \In f (x \+ y) <-> e \In f x \/ e \In f y. -Proof. -move=>W H; rewrite pfjoin //; split; first by case/InUn; eauto. -by case; [apply: InL|apply: InR]; rewrite pfV2. -Qed. - -End PlainMorph. - -Section FullMorph. -Variable f : full_pcm_morph U1 U2. - -Lemma InpfLT e (x y : U1) : - valid (x \+ y) -> - e \In f x -> - e \In f (x \+ y). -Proof. by move=>W; apply: InpfL. Qed. - -Lemma InpfRT e (x y : U1) : - valid (x \+ y) -> - e \In f y -> - e \In f (x \+ y). -Proof. by move=>W; apply: InpfR. Qed. - -Lemma InpfUnT e (x y : U1) : - valid (x \+ y) -> - e \In f (x \+ y) <-> e \In f x \/ e \In f y. -Proof. by move=>W; apply: InpfUn. Qed. -End FullMorph. - -End MorphMembership. - -(*********) -(* range *) -(*********) - -Section Range. -Variables (K : ordType) (C : pred K) (V : Type) (U : union_map 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. - -Lemma rangePtK k v : C k -> range (pts k v) = [:: v]. -Proof. by move=>C'; rewrite /range assocsPt C'. 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 In_rangeF k v f : - v \In range (free f k) <-> - exists2 k', k' != k & (k', v) \In f. -Proof. -rewrite In_rangeX; split; first by case=>x /InF []; exists x. -case=>k' Nk H; exists k'; apply/InF; split=>//. -by rewrite validF (In_valid H). -Qed. - -End Range. - -Prenex Implicits In_range_valid In_range In_rangeUn In_rangeF. - -(* decidable versions, when V is eqtype *) - -Section DecidableRange. -Variables (K : ordType) (C : pred K) (V : eqType) (U : union_map 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: (normalP f)=>[->|W]. -- by rewrite range_undef; 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/In_assocs H1 /mem_seqP/In_assocs 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/In_assocs/mem_seqP: H1=>->; move/In_assocs/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 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 rangeF k (f : U) : {subset range (free f k) <= range f}. -Proof. -case: (normalP f)=>[->|W]; first by rewrite free_undef !range_undef. -case D: (k \in dom f); last by move/negbT/dom_free: D=>->. -case: (um_eta D) W=>x [_] {1 3}-> Vf p. -by rewrite rangePtUn inE Vf=>->; rewrite orbT. -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: (normalP f)=>[->|W]; first by 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 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/In_assocs 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/In_assocs/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 k1 k2 v f : - um_mono f -> (k1, v) \In f -> (k2, v) \In f -> k1 = k2. -Proof. -move/ummonoP=>M H1 H2; case: (ordP k1 k2). -- by move/(M _ _ _ _ H1 H2); rewrite irr. -- by move/eqP. -by move/(M _ _ _ _ H2 H1); rewrite irr. -Qed. - -Lemma In_mono_range v f1 f2 : - valid (f1 \+ f2) -> um_mono (f1 \+ f2) -> - v \in range f1 -> v \in range f2 -> false. -Proof. -move=>W M /mem_seqP/In_rangeX [k1 H1] /mem_seqP/In_rangeX [k2 H2]. -have H1' : (k1, v) \In f1 \+ f2 by apply/InL. -have H2' : (k2, v) \In f1 \+ f2 by apply/InR. -rewrite -{k2 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 ummono_inj_find f : - um_mono f -> {in dom f & predT, injective (fun x => find x f)}. -Proof. -move/ummono_leP=>H k1 k2 /In_domX [x1 F1] _ E. -have /In_domX [x2 F2] : k2 \in dom f. -- by case: (dom_find k2) F1 E=>// _ /In_find ->. -move/In_find: (F1) E=>->; move/In_find: (F2)=>-> [?]; subst x2. -move: (H _ _ _ _ F1 F2) (H _ _ _ _ F2 F1); rewrite orefl=>{H} H1 H2. -case: (equivP idP H1) (@oantisym K k1 k2)=>// _. -by case: (equivP idP H2)=>// _; apply. -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. - -Lemma ummonoF f x : um_mono f -> um_mono (free f x). -Proof. -move/ummonoP=>X; apply/ummonoP=>k k' v v'. -by case/InF=>_ _ F /InF [_ _]; apply: X F. -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 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 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. - -Lemma In_umfoldl z0 d (f : U) (k : K) (v : V) : - (k, v) \In f -> [pcm a Unit k v <= um_foldl a z0 d f]. -Proof. -move=>H; move: (H); rewrite (In_eta H); case=>W _. -by rewrite umfoldlUn_frame // umfoldlPt (validPtUn_cond W) /= -joinA; eauto. -Qed. - -End PCMFold. - -(* Fold when the function produces a map *) - -Section FoldMap. -Variables (K : ordType) (C : pred K) (V : Type) (U : union_map C V). -Variable (a : U -> K -> V -> U). -Hypothesis frame : forall x y k v, a (x \+ y) k v = a x k v \+ y. - -Lemma In_umfoldlMX (f : U) (k : K) (v : V) : - (k, v) \In um_foldl a Unit undef f -> - exists k1 (v1 : V), (k, v) \In a Unit k1 v1 /\ (k1, v1) \In f. -Proof. -elim/um_indf: f. -- by rewrite umfoldl_undef=>/In_undef. -- by rewrite umfoldl0=>/In0. -move=>k1 v1 f IH W P. -rewrite umfoldlUn_frame // umfoldlPt (validPtUn_cond W) unitR. -case/InUn=>[H|/IH [k2][v2][H2 R2]]; first by exists k1, v1. -by exists k2, v2; split=>//; apply/InPtUnE=>//; right. -Qed. - -Lemma In_umfoldlM (f : U) (k : K) (v : V) k1 v1 : - valid (um_foldl a Unit undef f) -> - (k, v) \In a Unit k1 v1 -> (k1, v1) \In f -> - (k, v) \In um_foldl a Unit undef f. -Proof. -move=>W H /(In_umfoldl frame Unit undef) [x E]. -by move: E W=>-> W; apply/InL. -Qed. - -End FoldMap. - -(**************) -(* Map prefix *) -(**************) - -Section MapPrefix. -Variables (K : ordType) (C : pred K) (V : Type). -Variables (U : union_map C V). - -Definition um_prefix (h1 h2 : U) := Prefix (assocs h1) (assocs h2). - -Lemma umpfx_undef h : um_prefix undef h. -Proof. -by rewrite /um_prefix assocs_undef; apply/PrefixE; exists (assocs h). -Qed. - -Lemma umpfx0 h : um_prefix Unit h. -Proof. -by rewrite /um_prefix assocs0; apply/PrefixE; exists (assocs h). -Qed. - -Lemma umpfxD h1 h2 : um_prefix h1 h2 -> Prefix (dom h1) (dom h2). -Proof. -by case/PrefixE=>x E; rewrite !assocs_dom E map_cat; apply: Prefix_cat. -Qed. - -(* we next need a helper lemma, which should be in assocs section *) -(* but couldn't be proved there, because um_indf appears later *) -Lemma assocsUn (h1 h2 : U) : - valid (h1 \+ h2) -> - (forall k1 k2, k1 \in dom h1 -> k2 \in dom h2 -> ord k1 k2) -> - assocs (h1 \+ h2) = assocs h1 ++ assocs h2. -Proof. -move: h1 h2; apply: um_indf=>[h1|h2 W H| - k v f IH W1 /(order_path_min (@trans _)) P h2 W2 H]. -- by rewrite undef_join valid_undef. -- by rewrite assocs0 unitL. -rewrite -joinA in W2; rewrite -joinA !assocsPtUn //= ?IH //. -- by rewrite (validR W2). -- by move=>k1 k2 K1 K2; apply: H=>//; rewrite domPtUn inE K1 orbT W1. -apply/allP=>x; rewrite domUn inE (validR W2) /=. -by case/orP=>[/(allP P)//|]; apply: H; rewrite domPtUnE. -Qed. - -Lemma umpfxI h1 h2 : - valid (h1 \+ h2) -> - (forall x y, x \in dom h1 -> y \in dom h2 -> ord x y) -> - um_prefix h1 (h1 \+ h2). -Proof. by move=>W X; rewrite /um_prefix assocsUn //; apply: Prefix_cat. Qed. - -Lemma umpfxE h1 h : - valid h1 -> - um_prefix h1 h -> - exists2 h2, h = h1 \+ h2 & - forall x y, x \in dom h1 -> y \in dom h2 -> ord x y. -Proof. -move=>V1; case: (normalP h)=>[->|W]. -- by exists undef; rewrite ?join_undef ?dom_undef. -move: h1 h V1 W; apply: um_indf=>[h|h _ W| -k v h1 IH W1 /(order_path_min (@trans _)) P h _ W]. -- by rewrite valid_undef. -- by exists h; [rewrite unitL | rewrite dom0]. -case/PrefixE=>h2'; rewrite assocsPtUn //= => Eh. -set h' := free h k; have W' : valid h' by rewrite validF. -have : um_prefix h1 h'. -- rewrite /um_prefix assocsF Eh /= eq_refl filter_cat /=. - have : forall kv, kv \In assocs h1 -> kv.1 != k. - - move=>kv /In_assocs/In_dom H; move/allP/(_ _ H): P. - by case: eqP=>// ->; rewrite irr. - by move/eq_In_filter=>->; rewrite filter_predT; apply: Prefix_cat. -case/(IH _ (validR W1) W')=>h2 Eh' H; exists h2. -- rewrite -joinA -Eh'; apply/um_eta2/In_find/In_assocs. - by rewrite Eh In_cons; left. -move=>x y; rewrite domPtUn inE W1 /=; case/orP; last by apply: H. -move/eqP=><-{x} Dy. -have : y \in dom h' by rewrite Eh' domUn inE -Eh' W' /= Dy orbT. -rewrite domF inE; case: (k =P y)=>// /eqP N. -rewrite assocs_dom Eh /= inE eq_sym (negbTE N) /=. -case/mem_seqP/MapP; case=>a b X -> /=. -have {}P : path ord k (map fst (assocs h1 ++ h2')). -- by move: (sorted_dom h); rewrite assocs_dom Eh /=. -suff {X} : forall x, x \In assocs h1 ++ h2' -> ord k x.1 by move/(_ _ X). -apply/allPIn/(order_path_min (x:=(k,v)) (leT:=fun k x=>ord k.1 x.1)). -- by move=>???; apply: trans. -by rewrite -path_map. -Qed. - -End MapPrefix. - - -(********) -(* omap *) -(********) - -(* Combines map and filter by having filtering function return option. *) -(* This is very common form, so we also build structure for it. *) - -(* Definition and basic properties *) -Section OmapDefs. -Variables (K : ordType) (C : pred K) (V V' : Type). -Variables (U : union_map C V) (U' : union_map C V'). -Variable f : K * V -> option V'. - -Definition omap (x : U) : U' := - if valid x then - foldl (fun z kv => if f kv is Some v' then z \+ pts kv.1 v' else z) - Unit (assocs x) - else undef. - -Lemma omap0 : omap Unit = Unit. -Proof. by rewrite /omap valid_unit assocs0. Qed. - -Lemma omap_undef : omap undef = undef. -Proof. by rewrite /omap valid_undef. Qed. - -Lemma omapPt k v : - omap (pts k v) = - if C k then - if f (k, v) is Some w then pts k w else Unit - else undef. -Proof. -rewrite /omap validPt assocsPt; case D : (C k)=>//=. -by case: (f _)=>// w; rewrite unitL. -Qed. - -Lemma omapUn x1 x2 : - valid (x1 \+ x2) -> - omap (x1 \+ x2) = omap x1 \+ omap x2. -Proof. -move=>W; rewrite /omap W (validL W) (validR W); set o := fun z kv => _. -have H (x y : U') kv : o (x \+ y) kv = o x kv \+ y. -- by rewrite /o; case: (f kv)=>// w; rewrite joinAC. -rewrite (foldl_perm H _ (assocs_perm W)) foldl_cat. -by rewrite joinC -(foldl_init H) unitL. -Qed. - -Lemma omapVUn x1 x2 : - omap (x1 \+ x2) = - if valid (x1 \+ x2) then omap x1 \+ omap x2 else undef. -Proof. -case: (normalP (x1 \+ x2))=>[->|W]; -by [rewrite omap_undef|apply: omapUn]. -Qed. - -Lemma omapPtUn k v x : - omap (pts k v \+ x) = - if valid (pts k v \+ x) then - if f (k, v) is Some v' then pts k v' \+ omap x else omap x - else undef. -Proof. -case: ifP=>D; last first. -- by move/invalidE: (negbT D)=>->; rewrite omap_undef. -rewrite omapUn // omapPt (validPtUn_cond D). -by case: (f _)=>//; rewrite unitL. -Qed. - -Lemma omap_subdom x : {subset dom (omap x) <= dom x}. -Proof. -elim/um_indf: x=>[||k v x IH W P] t. -- by rewrite omap_undef dom_undef. -- by rewrite omap0 dom0. -rewrite omapPtUn W domPtUn W !inE /=. -case E : (f _)=>[b|]; last by move/IH=>->; rewrite orbT. -rewrite domPtUn inE; case/andP=>W2. -by case/orP=>[->//|/IH ->]; rewrite orbT. -Qed. - -Lemma valid_omap x : valid (omap x) = valid x. -Proof. -elim/um_indf: x=>[||k v x IH W P]. -- by rewrite omap_undef !valid_undef. -- by rewrite omap0 !valid_unit. -rewrite omapPtUn W; case E : (f _)=>[b|]; last by rewrite IH (validR W). -rewrite validPtUn (validPtUn_cond W) IH (validR W). -by apply: contra (notin_path P); apply: omap_subdom. -Qed. - -End OmapDefs. - -Arguments omap {K C V V' U U'}. - -(* Structure definition *) - -Definition omap_fun_axiom (K : ordType) (C : pred K) (V V' : Type) - (U : union_map C V) (U' : union_map C V') - (f : U -> U') (omf : K * V -> option V') := - f =1 omap omf. - -(* factory to use if full/norm/tpcm morphism property already proved *) -(* (omap_fun isn't binormal as it can drop timestamps) *) -HB.mixin Record isOmapFun_morph (K : ordType) (C : pred K) (V V' : Type) - (U : union_map C V) (U' : union_map C V') (f : U -> U') of - @Full_Norm_TPCM_morphism U U' f := { - omf_op : K * V -> option V'; - omfE_op : omap_fun_axiom f omf_op}. - -#[short(type=omap_fun)] -HB.structure Definition OmapFun (K : ordType) (C : pred K) (V V' : Type) - (U : union_map C V) (U' : union_map C V') := - {f of isOmapFun_morph K C V V' U U' f &}. - -Arguments omfE_op {K C V V' U U'}. -Arguments omf_op {K C V V' U U'} _ _ /. - -(* factory to use when full/norm/tpcm morphism property *) -(* hasn't been established *) -HB.factory Record isOmapFun (K : ordType) (C : pred K) (V V' : Type) - (U : union_map C V) (U' : union_map C V') (f : U -> U') := { - omf : K * V -> option V'; - omfE : omap_fun_axiom f omf}. - -HB.builders Context K C V V' U U' f of isOmapFun K C V V' U U' f. - -Lemma omap_fun_is_pcm_morph : pcm_morph_axiom relT f. -Proof. -split=>[|x y W _]; first by rewrite omfE omap0. -by rewrite !omfE -omapUn // valid_omap. -Qed. - -HB.instance Definition _ := - isPCM_morphism.Build U U' f omap_fun_is_pcm_morph. - -Lemma omap_fun_is_tpcm_morph : tpcm_morph_axiom f. -Proof. by rewrite /tpcm_morph_axiom omfE omap_undef. Qed. - -HB.instance Definition _ := - isTPCM_morphism.Build U U' f omap_fun_is_tpcm_morph. - -Lemma omap_fun_is_full_morph : full_pcm_morph_axiom f. -Proof. by []. Qed. - -HB.instance Definition _ := - isFull_PCM_morphism.Build U U' f omap_fun_is_full_morph. - -Lemma omap_fun_is_norm_morph : norm_pcm_morph_axiom f. -Proof. by move=>x /=; rewrite omfE valid_omap. Qed. - -HB.instance Definition _ := - isNorm_PCM_morphism.Build U U' f omap_fun_is_norm_morph. -HB.instance Definition _ := - isOmapFun_morph.Build K C V V' U U' f omfE. -HB.end. - -(* notation to hide the structure when projecting omf *) -Section OmapFunNotation. -Variables (K : ordType) (C : pred K) (V V' : Type). -Variables (U : union_map C V) (U' : union_map C V'). - -Definition omfx (f : omap_fun U U') of phantom (U -> U') f : - K * V -> option V' := omf_op f. - -Notation omf f := (omfx (Phantom (_ -> _) f)). - -Lemma omfE (f : omap_fun U U') : f =1 omap (omf f). -Proof. exact: omfE_op. Qed. -End OmapFunNotation. - -Notation omf f := (omfx (Phantom (_ -> _) f)). - -(* omap is omap_fun *) -Section OmapOmapFun. -Variables (K : ordType) (C : pred K) (V V' : Type). -Variables (U : union_map C V) (U' : union_map C V'). - -HB.instance Definition _ f := - isOmapFun.Build K C V V' U U' (omap f) (fun => erefl _). - -Lemma omf_omap f : omf (omap f) = f. Proof. by []. Qed. -End OmapOmapFun. - - -(* general omap_fun lemmas *) - -Section OmapFunLemmas. -Variables (K : ordType) (C : pred K) (V V' : Type). -Variables (U : union_map C V) (U' : union_map C V'). -Implicit Types (f : omap_fun U U') (x : U). - -(* different name for pfjoin on full morphisms *) -Lemma omfUn f x1 x2 : valid (x1 \+ x2) -> f (x1 \+ x2) = f x1 \+ f x2. -Proof. exact: pfjoinT. Qed. - -Lemma omfPtE f k v : - f (pts k v) = - if C k then - if omf f (k, v) is Some w then pts k w else Unit - else undef. -Proof. by rewrite omfE omapPt. Qed. - -Lemma omfPt f k v : - C k -> - f (pts k v) = if omf f (k, v) is Some w then pts k w else Unit. -Proof. by rewrite omfPtE=>->. Qed. - -Lemma omfPtUn f k v x : - f (pts k v \+ x) = - if valid (pts k v \+ x) then - if omf f (k, v) is Some v' then pts k v' \+ f x else f x - else undef. -Proof. -case: ifP=>X; last first. -- by move/invalidE: (negbT X)=>->; rewrite pfundef. -rewrite omfUn // omfPtE (validPtUn_cond X). -by case: (omf f _)=>//; rewrite unitL. -Qed. - -Lemma omfUnPt f k v x : - f (x \+ pts k v) = - if valid (x \+ pts k v) then - if omf f (k, v) is Some v' then f x \+ pts k v' else f x - else undef. -Proof. -rewrite joinC omfPtUn; case: ifP=>// W. -by case: (omf f _)=>// w; rewrite joinC. -Qed. - -(* membership *) -Lemma In_omfX f k v x : - (k, v) \In f x <-> - exists2 w, (k, w) \In x & omf f (k, w) = Some v. -Proof. -rewrite omfE; elim/um_indf: x k v=>[||k' v' x IH W P] k v. -- by rewrite pfundef; split=>[/In_undef|[w][]] //; rewrite valid_undef. -- by rewrite pfunit; split=>[/In0|[w] /In0]. -split=>[H|[w] H E]. -- move: (dom_valid (In_dom H)); rewrite omapPtUn W in H *. - case E : (omf f _) H=>[z|] H W1; last first. - - by case/IH: H=>w1; exists w1=>//; apply/InR. - move: H; rewrite InPtUnE //. - case=>[[->->]|]; first by exists v'. - by case/IH=>w; exists w=>//; apply/InR. -have : valid (omap (U':=U') (omf f) (pts k' v' \+ x)) by rewrite pfV. -rewrite omapPtUn W; move/(InPtUnE _ W): H E; case. -- by case=>->->-> V1; apply/InPtUnE=>//; left. -case: (omf _ (k', v'))=>[b|] H E V1; last by apply/IH; exists w. -by rewrite InPtUnE //; right; apply/IH; exists w. -Qed. - -(* one side of the In_omfX can destruct the existential *) -Lemma In_omf f k v w x : - (k, w) \In x -> omf f (k, w) = Some v -> (k, v) \In f x. -Proof. by move=>H1 H2; apply/In_omfX; exists w. Qed. - -(* dom *) -Lemma omf_subdom f x : {subset dom (f x) <= dom x}. -Proof. rewrite omfE; apply: omap_subdom. Qed. - -Arguments omf_subdom {f x}. - -Lemma In_odom f x k : k \In f x -> k.1 \in dom x. -Proof. by move/In_dom/omf_subdom. Qed. - -Lemma omf_cond f x k : k \in dom (f x) -> C k. -Proof. by move/omf_subdom/dom_cond. Qed. - -Lemma omf_sorted f x : sorted ord (dom (f x)). -Proof. by apply: sorted_dom. Qed. - -Lemma path_omf f x k : path ord k (dom x) -> path ord k (dom (f x)). -Proof. -apply: subseq_path; first by apply: trans. -apply: (sorted_subset_subseq (ltT := ord)); last by apply: omf_subdom. -- by apply: irr. -- by apply: trans. -- by apply: sorted_dom. -by apply: sorted_dom. -Qed. - -Lemma In_dom_omfX f x k : - reflect (exists v, (k, v) \In x /\ omf f (k, v)) - (k \in dom (f x)). -Proof. -case: In_domX=>H; constructor. -- by case: H=>v /In_omfX [w H1 H2]; exists w; rewrite H2. -case=>v [X]; case Y: (omf _ _)=>[w|] // _. -by elim: H; exists w; apply/In_omfX; exists v. -Qed. - -Lemma dom_omfUn f x1 x2 : - dom (f (x1 \+ x2)) =i - [pred k | valid (x1 \+ x2) & (k \in dom (f x1)) || (k \in dom (f x2))]. -Proof. -move=>k; rewrite !(omfE f) inE. -case: (normalP (x1 \+ x2))=>[->|W]; first by rewrite pfundef dom_undef. -by rewrite pfjoinT //= domUn inE -pfjoinT //= pfV //= W. -Qed. - -Lemma dom_omfPtUn f k v x : - dom (f (pts k v \+ x)) =i - [pred t | valid (pts k v \+ x) & - (omf f (k, v) && (t == k)) || (t \in dom (f x))]. -Proof. -move=>t; rewrite dom_omfUn !inE; case D : (valid _)=>//=. -rewrite omfPtE (validPtUn_cond D); case E: (omf _ _)=>[a|] /=. -- by rewrite domPt inE (validPtUn_cond D) eq_sym. -by rewrite dom0. -Qed. - -(* validity *) - -Lemma In_omfV f kw x : - kw \In f x -> - valid x. -Proof. by case/In_omfX=>v /In_valid. Qed. - -(* stronger than morphism property as uses two different f's *) -Lemma valid_omfUn f1 f2 x1 x2 : - valid (x1 \+ x2) -> - valid (f1 x1 \+ f2 x2). -Proof. -move=>W; rewrite (omfE f1) (omfE f2) validUnAE !pfVE !(validE2 W) /=. -by apply/allP=>x /omf_subdom D2; apply/negP=>/omf_subdom/(dom_inNRX W D2). -Qed. - -Lemma valid_omfPtUn f k v x : - [&& C k, valid x & k \notin dom x] -> - valid (pts k v \+ f x). -Proof. -case/and3P=>H1 H2 H3; rewrite validPtUn H1 pfVE H2 /=. -by apply: contra (omf_subdom _) H3. -Qed. - -Lemma valid_omfUnPt f k v x : - [&& C k, valid x & k \notin dom x] -> - valid (f x \+ pts k v). -Proof. by move=>H; rewrite joinC; apply: valid_omfPtUn. Qed. - -(* range *) - -Lemma In_range_omapUn f x1 x2 v : - v \In range (f (x1 \+ x2)) <-> - valid (x1 \+ x2) /\ (v \In range (f x1) \/ v \In range (f x2)). -Proof. -split=>[|[D]]. -- case: (normalP (x1 \+ x2))=>[->|D]; last first. - - by rewrite pfjoinT // => /In_rangeUn. - by rewrite pfundef range_undef. -rewrite pfjoinT //; case. -- by apply/In_rangeL/pfV2I. -by apply/In_rangeR/pfV2I. -Qed. - -(* other interactions *) - -Lemma find_omf f k x : - find k (f x) = - if find k x is Some v then omf f (k, v) else None. -Proof. -case E1 : (find k _)=>[b|]. -- by move/In_find: E1=>/In_omfX [w] /In_find ->. -move/find_none/negP: E1=>E1. -case E2 : (find k x)=>[c|] //; move/In_find: E2=>E2. -case E3 : (omf f (k, c))=>[d|] //; elim: E1. -by apply/In_dom_omfX; exists c; rewrite E3. -Qed. - -Lemma omfF f k x : f (free x k) = free (f x) k. -Proof. -case: (normalP x)=>[->|D]; first by rewrite pfundef !free_undef pfundef. -apply: umem_eq. -- by rewrite pfVE validF. -- by rewrite validF pfVE. -case=>t v; split. -- case/In_omfX=>w /InF /= [_ N M E]. - apply/InF; rewrite validF pfVE D N; split=>//. - by apply/In_omfX; exists w. -case/InF=>/= _ N /In_omfX [v'] M E. -by apply/In_omfX; exists v'=>//; apply/InF; rewrite validF. -Qed. - -Lemma omfUE f k (v : V) (x : U) : - f (upd k v x) = - if C k then - if omf f (k, v) is Some v' then upd k v' (f x) - else free (f x) k - else undef. -Proof. -case: ifP=>// D; last by rewrite (upd_condN v x (negbT D)) pfundef. -case: (normalP x)=>[->|H]. -- by case: (omf _ _)=>[a|]; rewrite ?(upd_undef,free_undef,pfundef). -rewrite upd_eta omfPtUn validPtUn D validF H domF inE eq_refl /=. -case: (omf _ _)=>[xa|]; last by rewrite omfF. -by rewrite upd_eta omfF. -Qed. - -Lemma omfU f k (v : V) (x : U) : - C k -> - f (upd k v x) = - if omf f (k, v) is Some v' then upd k v' (f x) - else free (f x) k. -Proof. by move=>D; rewrite omfUE D. Qed. - -(* when mapped functions are equal *) - -Lemma eq_in_omf f1 f2 x : - f1 x = f2 x <-> - (forall kv, kv \In x -> omf f1 kv = omf f2 kv). -Proof. -have L h : (forall kv, kv \In h -> omf f1 kv = omf f2 kv) -> - f1 h = f2 h. -- elim/um_indf: h=>[||k v h IH W P H]; rewrite ?pfundef ?pfunit //. - by rewrite !omfPtUn W (H _ (InPtUnL W)) -IH // => kv /(InR W)/H. -split; last by apply: L. -elim/um_indf: x=>[_ kv /In_undef|_ kv /In0|k v x IH W P H kv] //. -have E t : find t (f1 (pts k v \+ x)) = find t (f2 (pts k v \+ x)). -- by rewrite H. -case/(InPtUnE _ W)=>[->|]. -- by move: (E k); rewrite !find_omf findPtUn. -apply: {kv} IH; apply: L; case=>k1 v1 X. -move: (E k1); rewrite !find_omf findPtUn2 //. -case: (k1 =P k) X =>[-> /In_dom|_ /In_find ->] //=. -by move/negbTE: (validPtUnD W)=>->. -Qed. - -(* if mapped function is total, we have some stronger lemmas *) - -Lemma dom_omf_some f x : - (forall kv, kv \In x -> omf f kv) -> dom (f x) = dom x. -Proof. -move=>H; apply/domE=>k; apply/idP/idP=>/In_domX [w]. -- by case/In_omfX=>v /In_dom. -by move/[dup]/H=>E X; apply/In_dom_omfX; exists w. -Qed. - -Lemma domeq_omf_some f x : - (forall kv, kv \In x -> omf f kv) -> dom_eq (f x) x. -Proof. by move/dom_omf_some=>E; rewrite /dom_eq E pfVE !eq_refl. Qed. - -(* if mapped function is total on x1, x2 *) -(* we don't need validity condition for union *) -Lemma omfUn_some f x1 x2 : - (forall k, k \In x1 -> omf f k) -> - (forall k, k \In x2 -> omf f k) -> - f (x1 \+ x2) = f x1 \+ f x2. -Proof. -move=>H1 H2; have Ev : valid (f x1 \+ f x2) = valid (x1 \+ x2). -- by rewrite !validUnAE !pfVE !dom_omf_some. -case: (normalP (x1 \+ x2)) Ev=>[->|D _]; last by apply: pfjoinT. -by rewrite pfundef=>/negbT/invalidE. -Qed. - -Lemma omf_predU f1 f2 x : - (forall kv, omf f1 kv = None \/ omf f2 kv = None) -> - f1 x \+ f2 x = - omap (fun kv => if omf f1 kv is Some v - then Some v else omf f2 kv) x. -Proof. -move=>E; rewrite (omfE f1) (omfE f2). -elim/um_indf: x=>[||k v x IH W /(order_path_min (@trans K)) P]. -- by rewrite !pfundef undef_join. -- by rewrite !pfunit unitL. -rewrite !omapPtUn W -IH. -case E1 : (omf f1 (k, v))=>[b1|]; case E2 : (omf f2 (k, v))=>[b2|] //. -- by move: (E (k, v)); rewrite E1 E2; case. -- by rewrite joinA. -by rewrite joinCA. -Qed. - -Lemma omf_unit f x : - reflect (valid x /\ forall kv, kv \In x -> omf f kv = None) - (unitb (f x)). -Proof. -case: unitbP=>H; constructor; last first. -- case=>D; elim/um_indf: x D H=>[||k v x IH W P _]; - rewrite ?valid_undef ?pfunit ?omfPtUn //= W. - case E : (omf f _)=>[a|]; first by move=>_ /(_ _ (InPtUnL W)); rewrite E. - by move=>H1 H2; apply: (IH (validR W) H1)=>kv H; apply: H2 (InR W H). -split; first by rewrite -(pfVE (f:=f)) /= H valid_unit. -elim/um_indf: x H=>[H kv /In_undef|H kv /In0|k v x IH W P /[swap] kv] //. -rewrite omfPtUn W; case E : (omf f _)=>[a|]. -- by move/unitbP; rewrite um_unitbPtUn. -by move=>H /(InPtUnE _ W) [->//|]; apply: IH H _. -Qed. - -Lemma omf_none f x : - valid x -> - (forall kv, kv \In x -> omf f kv = None) -> - f x = Unit. -Proof. by move=>D H; apply/unitbP/omf_unit. Qed. - -Lemma omf_noneR f x : - f x = Unit -> forall kv, kv \In x -> omf f kv = None. -Proof. by case/unitbP/omf_unit. Qed. - -End OmapFunLemmas. - -Arguments omf_subdom {K C V V' U U' f x}. -Arguments In_omf {K C V V' U U'} _ {k v w x}. -Arguments In_odom {K C V V' U U' f x k} . - -(* omap_fun's with different ranges *) - -Section OmapFun2. -Variables (K : ordType) (C : pred K) (V V1 V2 : Type). -Variables (U : union_map C V) (U1 : union_map C V1) (U2 : union_map C V2). -Variables (f1 : omap_fun U U1) (f2 : omap_fun U U2). - -(* when one mapped function is included in other *) - -Lemma omf_dom_subseq (x : U) : - (forall kv, kv \In x -> omf f1 kv -> omf f2 kv) -> - subseq (dom (f1 x)) (dom (f2 x)). -Proof. -elim/um_indf: x=>[||k v x IH W P] H; -rewrite ?pfundef ?pfunit ?dom_undef ?dom0 //. -move: (W); rewrite validPtUn=>W'; rewrite !omfPtUn W. -have T : transitive (@ord K) by apply: trans. -have B : (k, v) \In pts k v \+ x by apply: InPtUnL. -case E1 : (omf f1 (k, v))=>[x1|]. -- have /(H _ B): (omf f1 (k, v)) by rewrite E1. - case: (omf f2 (k, v))=>// x2 _. - rewrite !domPtUnK //=; last first. - - by apply/allP=>? /In_dom_omfX [?][] /In_dom Y _; apply: path_mem Y. - - by rewrite valid_omfPtUn. - - by apply/allP=>? /In_dom_omfX [?][] /In_dom Y _; apply: path_mem Y. - - by rewrite valid_omfPtUn. - by rewrite eq_refl; apply: IH=>kx X; apply: H (InR _ _). -case E2 : (omf f2 (k, v))=>[x2|]; last by apply: IH=>kx X; apply: H (InR _ _). -rewrite domPtUnK /=; last first. -- by apply/allP=>? /In_dom_omfX [?][] /In_dom Y _; apply: path_mem Y. -- by rewrite valid_omfPtUn. -case D : (dom (f1 x))=>[//|t ts]. -case: eqP D=>[-> D|_ <-]; last by apply: IH=>kv X; apply: H (InR _ _). -have : k \in dom (f1 x) by rewrite D inE eq_refl. -case/In_dom_omfX=>w1 [] /In_dom /= D1. -by rewrite validPtUn D1 !andbF in W. -Qed. - -End OmapFun2. - -Section OmapFun2Eq. -Variables (K : ordType) (C : pred K) (V V1 V2 : Type). -Variables (U : union_map C V) (U1 : union_map C V1) (U2 : union_map C V2). -Variables (f1 : omap_fun U U1) (f2 : omap_fun U U2). - -Lemma omf_dom_eq (x : U) : - (forall kv, kv \In x -> omf f1 kv <-> omf f2 kv) -> - dom (f1 x) = dom (f2 x). -Proof. -move=>H; apply: subseq_anti. -by rewrite !omf_dom_subseq // => kv /H ->. -Qed. - -End OmapFun2Eq. - -Section OmapMembershipExtra. -Variables (K : ordType) (C : pred K) (V1 V2 : Type). -Variables (U1 : union_map C V1) (U2 : union_map C V2). -Variables f : omap_fun U1 U2. - -Lemma omfL e (x y : U1) : - e.1 \in dom x -> - e \In f (x \+ y) -> - e \In f x. -Proof. -move=>D /[dup] /In_valid/fpV V; rewrite pfjoinT //. -by case/InUn=>// /In_dom/omf_subdom/(dom_inNLX V D). -Qed. - -Lemma omfR e (x y : U1) : - e.1 \in dom y -> - e \In f (x \+ y) -> - e \In f y. -Proof. -move=>D /[dup] /In_valid/fpV V; rewrite pfjoinT //. -by case/InUn=>// /In_dom/omf_subdom/(dom_inNRX V D). -Qed. - -Lemma omfDL k (x y : U1) : - k \in dom x -> - k \in dom (f (x \+ y)) -> - k \in dom (f x). -Proof. by move=>D /In_domX [v] /omfL-/(_ D)/In_dom. Qed. - -Lemma omfDR k (x y : U1) : - k \in dom y -> - k \in dom (f (x \+ y)) -> - k \in dom (f y). -Proof. by move=>D /In_domX [v] /omfR-/(_ D)/In_dom. Qed. - -Lemma InpfLTD k (x y : U1) : - valid (x \+ y) -> - k \in dom (f x) -> - k \in dom (f (x \+ y)). -Proof. by move=>V /In_domX [v] /(InpfLT V)/In_dom. Qed. - -Lemma InpfRTD k (x y : U1) : - valid (x \+ y) -> - k \in dom (f y) -> - k \in dom (f (x \+ y)). -Proof. by move=>V /In_domX [v] /(InpfRT V)/In_dom. Qed. -End OmapMembershipExtra. - -Section OmapMembershipExtra2. -Variables (K : ordType) (C : pred K) (V1 V2 : Type). -Variables (U1 : union_map C V1) (U2 : union_map C V2). -Variables f : omap_fun U1 U2. - -Lemma omfDL00 k a (x y : U1) : - k \in a::dom x -> - k \in a::dom (f (x \+ y)) -> - k \in a::dom (f x). -Proof. by rewrite !inE; case: (k =P a)=>//= _; apply: omfDL. Qed. - -Lemma omfDL0 k a (x y : U1) : - k \in dom x -> - k \in a::dom (f (x \+ y)) -> - k \in a::dom (f x). -Proof. by rewrite !inE; case: (k =P a)=>//= _; apply: omfDL. Qed. - -Lemma omfDR00 k a (x y : U1) : - k \in a::dom y -> - k \in a::dom (f (x \+ y)) -> - k \in a::dom (f y). -Proof. by rewrite !inE; case: (k =P a)=>//= _; apply: omfDR. Qed. - -Lemma omfDR0 k a (x y : U1) : - k \in dom y -> - k \in a::dom (f (x \+ y)) -> - k \in a::dom (f y). -Proof. by rewrite !inE; case: (k =P a)=>//= _; apply: omfDR. Qed. - -Lemma InpfLTD0 k a (x y : U1) : - valid (x \+ y) -> - k \in a::dom (f x) -> - k \in a::dom (f (x \+ y)). -Proof. by move=>V; rewrite !inE; case: (k =P a)=>//= _; apply: InpfLTD. Qed. - -Lemma InpfRTD0 k a (x y : U1) : - valid (x \+ y) -> - k \in a::dom (f y) -> - k \in a::dom (f (x \+ y)). -Proof. by move=>V; rewrite !inE; case: (k =P a)=>//= _; apply: InpfRTD. Qed. -End OmapMembershipExtra2. - -(* categoricals *) - -Section OmapFunId. -Variables (K : ordType) (C : pred K) (V : Type). -Variables (U : union_map C V). - -Lemma omf_some (f : omap_fun U U) x : - (forall kv, kv \In x -> omf f kv = Some kv.2) -> - f x = x. -Proof. -elim/um_indf: x=>[||k v x IH W P] H; rewrite ?pfundef ?pfunit //. -by rewrite omfPtUn W H //= IH // => kv D; apply: H (InR _ D). -Qed. - -Lemma id_is_omap_fun : omap_fun_axiom (@idfun U) (Some \o snd). -Proof. by move=>x; rewrite omf_some. Qed. - -(* we already have that id is full/norm/tpcm morph *) -(* so we use the factory for that situation *) -HB.instance Definition _ := - isOmapFun_morph.Build K C V V U U idfun id_is_omap_fun. - -Lemma omf_id : omf idfun = Some \o snd. Proof. by []. Qed. - -Lemma valid_omfUnL (f : omap_fun U U) x1 x2 : - valid (x1 \+ x2) -> valid (f x1 \+ x2). -Proof. by rewrite {2}(_ : x2 = @idfun U x2) //; apply: valid_omfUn. Qed. - -Lemma valid_omfUnR (f : omap_fun U U) x1 x2 : - valid (x1 \+ x2) -> valid (x1 \+ f x2). -Proof. by rewrite !(joinC x1); apply: valid_omfUnL. Qed. - -End OmapFunId. - - -Section OmapFunComp. -Variables (K : ordType) (C : pred K) (V1 V2 V3 : Type). -Variables (U1 : union_map C V1) (U2 : union_map C V2) (U3 : union_map C V3). -Variables (f1 : omap_fun U1 U2) (f2 : omap_fun U2 U3). - -Lemma comp_is_omap_fun : - omap_fun_axiom (f2 \o f1) - (fun x => obind (fun v => omf f2 (x.1, v)) (omf f1 x)). -Proof. -elim/um_indf=>[||k v x /= IH W P]; rewrite ?pfundef ?pfunit //=. -rewrite !omapPtUn !omfPtUn W /=; case: (omf f1 (k, v))=>[w|//]. -rewrite omfPtUn validPtUn (validPtUn_cond W) pfVE (validR W) /= IH. -by rewrite (contra _ (validPtUnD W)) //; apply: omf_subdom. -Qed. - -HB.instance Definition _ := - isOmapFun_morph.Build K C V1 V3 U1 U3 (f2 \o f1) comp_is_omap_fun. - -Lemma omf_comp : - omf (f2 \o f1) = - fun x => obind (fun v => omf f2 (x.1, v)) (omf f1 x). -Proof. by []. Qed. - -End OmapFunComp. - - -(*********************) -(* omap specifically *) -(*********************) - -(* 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 OmapId. -Variables (K : ordType) (C : pred K) (V : Type) (U : union_map C V). - -Lemma omapv_id : omapv Some =1 @id U. -Proof. by move=>x; apply: omf_some. Qed. - -Lemma mapv_id : mapv id =1 @id U. -Proof. by apply: omapv_id. Qed. - -End OmapId. - -Section OmapComp. -Variables (K : ordType) (C : pred K) (V1 V2 V3 : Type). -Variables (U1 : union_map C V1) (U2 : union_map C V2) (U3 : union_map C V3). - -Lemma omap_omap f1 f2 (x : U1) : - omap f2 (omap f1 x : U2) = - omap (fun kv => obind (fun v => f2 (kv.1, v)) (f1 kv)) x :> U3. -Proof. by rewrite (compE (omap f2)) omfE. Qed. - -Lemma omapv_omapv f1 f2 (x : U1) : - omapv f2 (omapv f1 x : U2) = omapv (obind f2 \o f1) x :> U3. -Proof. by rewrite omap_omap. Qed. - -Lemma omapv_mapv f1 f2 (x : U1) : - omapv f2 (mapv f1 x : U2) = omapv (f2 \o f1) x :> U3. -Proof. by rewrite omapv_omapv. Qed. - -(* RHS is just filtering; will restate with um_filter below *) -Lemma mapv_omapvE f g (x : U1) : - ocancel f g -> - mapv g (omapv f x : U2) = - omapv (fun v => if f v then Some v else None) x :> U1. -Proof. -move=>O; rewrite (compE (mapv g)) eq_in_omf omf_omap omf_comp !omf_omap /=. -by case=>k v H; case X: (f v)=>[a|] //=; rewrite -(O v) X. -Qed. - -End OmapComp. - - -Section OmapMembership. -Variables (K : ordType) (C : pred K) (V V' : Type). -Variables (U : union_map C V) (U' : union_map C V'). - -Lemma In_omapX f k v (x : U) : - (k, v) \In (omap f x : U') <-> - exists2 w, (k, w) \In x & f (k, w) = Some v. -Proof. exact: In_omfX. Qed. - -Lemma In_omap f k v w (x : U) : - (k, w) \In x -> f (k, w) = Some v -> (k, v) \In (omap f x : U'). -Proof. by move=>H1 H2; apply: In_omf H1 _; apply: H2. Qed. - -Lemma In_dom_omapX (f : K * V -> option V') k (x : U) : - reflect (exists w, (k, w) \In x /\ f (k, w)) - (k \in dom (omap f x : U')). -Proof. by case: In_dom_omfX=>H; constructor; exact: H. Qed. - -Lemma In_omapv f k v w (x : U) : - (k, w) \In x -> f w = Some v -> (k, v) \In (omapv f x : U'). -Proof. by move=>H1 H2; apply: In_omf H1 _; rewrite omf_omap. Qed. - -Lemma In_mapv f k v (x : U) : - injective f -> (k, f v) \In (mapv f x : U') <-> (k, v) \In x. -Proof. by move=>I; rewrite In_omfX; split; [case=>w H [] /I <-|exists v]. Qed. - -End OmapMembership. - -Arguments In_omapv {K C V V' U U' f k v w}. - -Section OmapMembership2. -Variables (K : ordType) (C : pred K) (V V' : Type). -Variables (U : union_map C V) (U' : union_map C V'). - -Lemma In_omapV f g k v (x : U) : - ocancel f g -> pcancel g f -> - (k, v) \In (omapv f x : U') <-> (k, g v) \In x. -Proof. -move=>O P; rewrite -(In_mapv U _ _ _ (pcan_inj P)). -rewrite mapv_omapvE // In_omapX /=. -split; first by case=>w H; case: (f w)=>[a|] //= [<-]. -by exists (g v)=>//=; rewrite P. -Qed. - -Lemma In_rangev f g v (x : U) : - ocancel f g -> pcancel g f -> - v \In range (omapv f x : U') <-> g v \In range x. -Proof. -move=>O P; rewrite !In_rangeX; split; case=>k H; exists k. -- by rewrite -(In_omapV _ _ _ O P). -by rewrite (In_omapV _ _ _ O P). -Qed. - -End OmapMembership2. - -Section OmapMembership3. -Variables (K : ordType) (C : pred K) (V V' : eqType). -Variables (U : union_map C V) (U' : union_map C V'). - -(* decidable variant of In_rangev *) -Lemma mem_rangev f g v (x : U) : - ocancel f g -> pcancel g f -> - v \in range (omapv f x : U') = (g v \in range x). -Proof. -by move=>O P; apply/idP/idP; move/mem_seqP/(In_rangev _ _ _ O P)/mem_seqP. -Qed. - -End OmapMembership3. - - -(* freeing k representable as omap *) -Section OmapFree. -Variables (K : ordType) (C : pred K) (V V' : Type). -Variables (U : union_map C V) (U' : union_map C V'). - -Lemma omap_free k x : - free x k = - omap (fun kv => if kv.1 == k then None else Some kv.2) x :> U. -Proof. -case: (normalP x)=>[->|D]; first by rewrite free_undef pfundef. -apply: umem_eq. -- by rewrite validF. -- by rewrite pfVE. -case=>t w; rewrite InF In_omapX /= validF D; split. -- by case=>_ /negbTE ->; exists w. -by case=>w' H; case: eqP=>// N [<-]. -Qed. - -End OmapFree. - -(* eq_in_omap *) -Section EqInOmap. -Variables (K : ordType) (C : pred K) (V V' : Type). -Variables (U : union_map C V) (U' : union_map C V'). - -Lemma eq_in_omap a1 a2 (h : U) : - (forall kv, kv \In h -> a1 kv = a2 kv) <-> - omap a1 h = omap a2 h :> U'. -Proof. by rewrite eq_in_omf. Qed. -End EqInOmap. - - -(*************) -(* 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 C V). -Implicit Types (f : U) (p q : pred (K * V)). - -Definition um_filter p f : U := - omap (fun kv => if p kv then Some kv.2 else None) f. - -(* um_filter is omap_fun *) -HB.instance Definition _ p := OmapFun.copy (um_filter p) (um_filter p). - -Lemma umfiltPtE 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 omfPtE omf_omap; case: (ifP (p _)). Qed. - -Lemma umfiltPt p k v : - C k -> - um_filter p (pts k v) = if p (k, v) then pts k v else Unit. -Proof. by move=>D; rewrite umfiltPtE D. Qed. - -Lemma umfiltUn p f1 f2 : - valid (f1 \+ f2) -> - um_filter p (f1 \+ f2) = um_filter p f1 \+ um_filter p f2. -Proof. exact: pfjoinT. 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 omfPtUn omf_omap; case: (ifP (p _)). 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 omfUnPt omf_omap; case: (ifP (p _)). Qed. - -Lemma In_umfiltX x p f : x \In um_filter p f <-> p x /\ x \In f. -Proof. -case: x => k v; rewrite In_omfX omf_omap; split=>[[w H]|[I H]] /=. -- by case E: (p (k, w))=>//=; case=><-. -by exists v=>//; rewrite I. -Qed. - -(* In_umfilt is really only good in one direction. *) -(* For the other direction, we need the following one *) -Lemma In_umfilt p x f : x \In f -> p x -> x \In um_filter p f. -Proof. by move=>X1 X2; apply/In_umfiltX. Qed. - -Lemma In_dom_umfilt p f k : - reflect (exists2 v, p (k, v) & (k, v) \In f) - (k \in dom (um_filter p f)). -Proof. -apply: (iffP (In_domX _ _)). -- by case=>v /In_umfiltX []; exists v. -by case=>v Hp Hf; exists v; apply/In_umfilt. -Qed. - -Lemma dom_omf_umfilt V' (U' : union_map C V') (f : omap_fun U U') x : - dom (f x) = dom (um_filter (isSome \o omf f) x). -Proof. -apply/domE=>k; apply/idP/idP. -- case/In_dom_omfX=>//= v [H1 H2]. - by apply/In_dom_umfilt; exists v. -case/In_dom_umfilt=>w /=. -case E: (omf f (k, w))=>[a|] // _ H. -by move/In_dom: (In_omf _ H E). -Qed. - -Lemma dom_omap_umfilt V' (U' : union_map C V') a x : - dom (omap a x : U') = dom (um_filter (isSome \o a) x). -Proof. exact: dom_omf_umfilt. 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: ord_sorted_eq=>//=. -- by apply: sorted_filter; [apply: trans | apply: sorted_dom]. -move=>k; rewrite mem_filter; apply/idP/idP. -- by case/In_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/In_dom_umfilt; exists v. -Qed. - -Lemma umfilt_pred0 f : valid f -> um_filter pred0 f = Unit. -Proof. by move=>W; apply: omf_none. Qed. - -Lemma umfilt_predT f : um_filter predT f = f. -Proof. by apply: omf_some. Qed. - -Lemma umfilt_predI p1 p2 f : - um_filter (predI p1 p2) f = um_filter p1 (um_filter p2 f). -Proof. -rewrite [RHS]compE eq_in_omf omf_comp !omf_omap /=. -by case=>k v H /=; case: (ifP (p2 _))=>//=; 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. -rewrite eq_in_omf !omf_omap /=. -split=>E [k v] H; first by rewrite E. -by case: (p1 _) (E _ H); case: (p2 _). -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 omf_predU=>[|kv]. -- by rewrite eq_in_omf !omf_omap /= => kv; case: (p1 _). -by rewrite !omf_omap /=; case: (p1 _)=>/=; [right|left]. -Qed. - -(* we put localization back In for xor *) -(* DEVCOMMENT *) -(* TODO: this should be done for all umfilt_?pred? lemmas *) -(* /DEVCOMMENT *) -Lemma umfilt_predX f p q : - (forall kv, kv \In f -> p kv (+) q kv) -> - f = um_filter p f \+ um_filter q f. -Proof. -move=>D; rewrite -{1}[f]umfilt_predT. -have : forall kv, kv \In f -> predT kv = (predU p q) kv. -- by move=>kv /D /=; case: (p kv). -move/eq_in_umfilt=>->; rewrite umfilt_predU; congr (_ \+ _). -by apply/eq_in_umfilt=>kv /D /=; case: (p kv)=>// /negbTE ->. -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. - -Corollary umfilt_predC f p : f = um_filter p f \+ um_filter (predC p) f. -Proof. -rewrite -umfilt_dpredU; last by move=>? /=; rewrite negbK. -rewrite -[LHS]umfilt_predT; apply: eq_in_umfiltE=>kv /=. -by rewrite orbN. -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: omf_subdom. - move=>D; move: (D); rewrite -{1}E domUn inE W' /=. - by case/orP=>// /omf_subdom; 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 umfiltUE 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 omfUE omf_omap; case: (p _). Qed. - -Lemma umfiltU p k v f : - C k -> - um_filter p (upd k v f) = - if p (k, v) then upd k v (um_filter p f) - else free (um_filter p f) k. -Proof. by move=>H; rewrite umfiltUE H. Qed. - -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 !pfundef !umfoldl_undef. -- by rewrite !pfunit !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 /In_dom_umfilt [w _] /In_dom. - by apply: path_mem (@trans K) P. -have : valid (um_filter p (pts k v \+ f)) by rewrite pfVE 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. by move/omf_noneR=>H k v /H; rewrite omf_omap /=; case: (p _). Qed. - -Lemma umfilt_mem0X p f k v : - (k, v) \In f -> um_filter p f = Unit -> ~ p (k, v). -Proof. by move=>H /umfilt_mem0/(_ _ _ H)/negP. Qed. - -Lemma umfilt_mem0L p f : - valid f -> (forall k v, (k, v) \In f -> ~~ p (k, v)) -> - um_filter p f = Unit. -Proof. by move=>W H; rewrite omf_none // omf_omap; case=>k v /H/negbTE ->. 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 pfundef assocs_undef. -- by rewrite pfunit 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 pfVE. -by apply/allP=>x; move/allP: P=>P; move/omf_subdom/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. by rewrite find_omf. Qed. - -Lemma unitb_umfilt p f : unitb f -> unitb (um_filter p f). -Proof. by move/unitbP=>->; rewrite pfunit 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. - -Lemma dom_umfilt2 p1 p2 f x : - x \in dom (um_filter p1 (um_filter p2 f)) = - (x \in dom (um_filter p1 f)) && (x \in dom (um_filter p2 f)). -Proof. -rewrite -umfilt_predI; apply/idP/idP. -- case/In_dom_umfilt=>v /andP [X1 X2] H. - by apply/andP; split; apply/In_dom_umfilt; exists v. -case/andP=>/In_dom_umfilt [v1 X1 H1] /In_dom_umfilt [v2 X2 H2]. -move: (In_fun H1 H2)=>E; rewrite -{v2}E in X2 H2 *. -by apply/In_dom_umfilt; exists v1=>//; apply/andP. -Qed. - -End FilterDefLemmas. - -#[export] -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). - -Arguments In_umfilt [K C V U] p x f _ _. - -Section FilterKLemmas. -Variables (K : ordType) (C : pred K) (V : Type) (U : union_map C V). -Implicit Type f : U. -Implicit Type p q : pred K. - -Lemma dom_umfiltkE p f : dom (um_filterk p f) = filter p (dom f). -Proof. -apply: ord_sorted_eq=>//=. -- by apply: sorted_filter; [apply: trans | apply: sorted_dom]. -move=>k; rewrite mem_filter; apply/idP/idP. -- by case/In_dom_umfilt=>v /= -> /In_dom. -by case/andP=>H1 /In_domX [v H2]; apply/In_dom_umfilt; exists v. -Qed. - -Lemma valid_umfiltkUn p1 p2 f : - valid f -> - {in dom f, forall x, p1 x -> p2 x -> false} -> - valid (um_filterk p1 f \+ um_filterk p2 f). -Proof. -move=>W H; rewrite validUnAE !pfVE W /=. -apply/allP=>x; case/In_dom_umfilt=>v1 /= H2 F1. -apply/negP; case/In_dom_umfilt=>v2 /= H1 F2. -by move: (H x (In_dom F1) H1 H2). -Qed. - -Lemma dom_umfiltk p f : dom (um_filterk p f) =i predI p (mem (dom f)). -Proof. by move=>k; rewrite dom_umfiltkE mem_filter. Qed. - -(* DEVCOMMENT *) -(* this also holds for invalid f1, as the corollary shows *) -(* /DEVCOMMENT *) -Lemma umfiltk_dom f1 f2 : - valid (f1 \+ f2) -> - um_filterk [dom f1] (f1 \+ f2) = f1. -Proof. -move=>W; apply/umem_eq; first by rewrite pfVE. -- by rewrite (validL W). -case=>k v; rewrite In_umfiltX; split=>[|H]. -- by case=>H /InUn [] // /In_dom; rewrite (negbTE (dom_inNL W H)). -by split; [apply: In_dom H | apply: InL]. -Qed. - -Corollary umfiltk_dom' f : um_filterk [dom f] f = f. -Proof. -case: (normalP f)=>[->|H]; first by rewrite pfundef. -by rewrite -{2}[f]unitR umfiltk_dom // unitR. -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 umfiltk_subdom p f : {subset dom (um_filterk p f) <= p}. -Proof. by move=>a; rewrite dom_umfiltk; case/andP. Qed. - -Lemma umfiltk_subdomE 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 [dom f] f = f. -Proof. by apply/umfiltk_subdomE. Qed. - -Lemma find_umfiltk k (p : pred K) f : - find k (um_filterk p f) = if p k then find k f else None. -Proof. by rewrite find_umfilt /=; case: (find _ _)=>[a|]; case: ifP. Qed. - -Lemma umfiltk_subdom0 p f : - valid f -> - {subset dom f <= predC p} <-> um_filterk p f = Unit. -Proof. -move=>W; split=>[H|H k X]. -- rewrite (eq_in_umfiltk (p2:=pred0)) ?umfilt_pred0 //. - by move=>a /H /negbTE ->. -case: dom_find X H=>// v _ -> _; rewrite omfUE !inE omf_omap /=. -case: (ifP (p k))=>// _ /unitbP. -by rewrite fun_if um_unitbU unitb_undef if_same. -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 umfiltUE pfunit 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: (normalP (pts k v \+ f))=>[->|W]; first by rewrite pfundef. -rewrite pfjoinT //= umfiltPtE (validPtUn_cond W) /=. -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. - -(* filters commute with omap_fun *) - -Lemma umfiltk_omf V' (U' : union_map C V') (f : omap_fun U U') p x : - f (um_filterk p x) = um_filterk p (f x). -Proof. -rewrite (compE f) [RHS]compE eq_in_omf !omf_comp !omf_omap /=. -by rewrite /obind/oapp /=; case=>k v; case: ifP; case: (omf f _). -Qed. - -Lemma umfiltk_dom_omf V' (U' : union_map C V') (f : omap_fun U U') x : - um_filterk [dom x] (f x) = f x. -Proof. by rewrite -umfiltk_omf umfiltk_dom'. Qed. - -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 umfiltUE /=; case: ifP; case: ifP=>//= Hp Hc. -- by rewrite dom_free // 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 omfF; case: ifP=>//= N. -by rewrite dom_free // dom_umfiltk inE N. -Qed. - -Lemma In_umfiltk p x f : x \In f -> p x.1 -> x \In um_filterk p f. -Proof. by apply: In_umfilt. Qed. - -End FilterKLemmas. - -Arguments In_umfiltk {K C V U} p {x f} _ _. - -Section FilterVLemmas. -Variables (K : ordType) (C : pred K) (V : Type) (U : union_map 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. - -Lemma In_umfiltv p x f : x \In f -> p x.2 -> x \In um_filterv p f. -Proof. by apply: In_umfilt. Qed. - -End FilterVLemmas. - -Arguments In_umfiltv {K C V U} p {x f} _ _. - -Section OmapMembershipLemmas. -Variables (K : ordType) (C : pred K) (V V' : Type). -Variables (U : union_map C V) (U' : union_map C V'). - -Lemma mapv_omapv f g (x : U) : - ocancel f g -> - mapv g (omapv f x : U') = um_filterv (isSome \o f) x. -Proof. exact: mapv_omapvE. Qed. - -End OmapMembershipLemmas. - -(************************) -(* PCM-induced ordering *) -(************************) - -(* Union maps and PCM ordering. *) - -Section UmpleqLemmas. -Variables (K : ordType) (C : pred K) (V : Type) (U : union_map C V). -Implicit Types (x y a b : U) (p : pred K). - -(* PCM-induced preorder is 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/umap0E=>->; 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: (normalP y)=>[->|]. -- by rewrite pfundef; apply: pleq_undef. -by case: H=>z -> D; rewrite pfjoinT. -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. -move=>D /(umfilt_pleq_mono (p \o fst)). -by 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 [dom a] x = a. -Proof. by move=>W; split=>[|<-] // H; case: H W=>b ->; 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: (normalP f)=>[->???|Df Dab]; first by apply: pleq_undef. -move/(umpleqk_pleqE _ Df)=> <- /(umpleqk_pleqE _ Df) <-. -by rewrite -umfilt_dpredU //; case=>/= k _; apply: dom_inNL Dab. -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/subsetR: (valid_subdom W). -by move/(umfiltk_subdom0 _ (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 move/(dom_inNLX V12)=>X /X. -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 *) -(********************) - -(* DEVCOMMENT *) -(* naturally belongs to dom section, but proofs use lemmas *) -(* that haven't been proved before the dom section *) -(* /DEVCOMMENT *) -Section Precision. -Variables (K : ordType) (C : pred K) (V : Type) (U : union_map 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. - by rewrite E domUn inE -E V1 /= (negbTE (dom_inNL V12 Hs)) orbF. -move=>Hs; case: validUn V12=>//. -- by rewrite (validE2 V1). -- by rewrite E in V1; rewrite (validE2 V1). -by rewrite E in V1; move=>k /Hs /(dom_inNL V1)/negbTE ->. -Qed. - -Lemma prec_filtV x1 x2 y1 y2 : - valid (x1 \+ y1) -> x1 \+ y1 = x2 \+ y2 -> - reflect (x1 = um_filterk [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 *) -(****************) - -(* version of eval where the user provides new order of evaluation *) -(* as list of timestamps which are then read in-order. *) - -Section OrdEvalDefLemmas. -Variables (K : ordType) (C : pred K) (V R : Type) (U : union_map 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: (normalP f)=>[-> _|D]; first by apply: oev_undef. -by move/(dom0E D)=>->; apply: oev0. -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. -elim: ks z0=>[|k' ks IH] z0 //=. -by 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. -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_umfiltX /=; 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) -> k \notin ks -> - 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 right. -by case=>// [][??]; subst k0; rewrite H in S. -Qed. - -(* a somewhat different version; makes side conditions easier to discharge *) -Lemma oevPtUn_sub 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 F; apply: oevPtUn=>//; apply/negP=>/F. -by rewrite (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. - -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. - -(* a somewhat stronger lemma making clear that z0' isn't *) -(* arbitrary but always obtained by evaluation starting from z0 *) -Lemma oev_indX {P : R -> Prop} f ks a z0 : - P z0 -> - (forall k ks1 ks2 v z0', (k, v) \In f -> ks = ks1 ++ k :: ks2 -> - z0' = oeval a ks1 f z0 -> 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. -- move=>k' ks1 ks2 v' z' X D E. - by apply: (H k' (k::ks1) ks2 v' z' X _ _)=>//; rewrite D. -case E: (find k f)=>[b|] //; move/In_find: E=>E. -by apply: (H k [::] _ b z0). -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 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 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 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 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_umfiltX. -split=>[|[]//]; split=>//; case/In_dom_umfilt: H=>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 umfiltPtE. -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 pfV. -rewrite /eval umfiltPtUn W. -case: (p (k, v))=>W'; last first. -- rewrite oevPtUn //; apply/negP=>/omf_subdom. - by rewrite (negbTE (validPtUnD W)). -rewrite domPtUnK //=; last by apply/allP=>x /omf_subdom /H. -by rewrite findPtUn // oevPtUn // (validPtUnD W'). -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 pfV. -rewrite /eval umfiltUnPt W. -case: (p (k, v))=>W'; last first. -- rewrite joinC oevPtUn //; first by rewrite joinC. - by apply/negP=>/omf_subdom; rewrite (negbTE (validUnPtD W)). -rewrite domUnPtK //=; last by apply/allP=>x /omf_subdom /H. -rewrite (oev_rconsP _ (v:=v)) // joinC oevPtUn //; first by rewrite joinC. -by apply/negP=>/omf_subdom; rewrite (negbTE (validUnPtD W)). -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: (normalP f)=>[->|D]; first by 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 C V) (U' : union_map 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 pfundef !eval_undef. -- by rewrite pfunit !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 /omf_subdom /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 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). -Variables (V1 V2 R1 R2 : Type). -Variables (U1 : union_map C1 V1) (U2 : union_map C2 V2). -Variables (U : union_map 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 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_sub=>//; apply: omf_subdom. -rewrite oevUn // -(oev_sub_filter (p:=mem [:: k])) ?(domPtK,Ck) //. -rewrite -dom_umfiltkE umfiltPtUn /= valid_omfUnR // inE eq_refl. -rewrite umfilt_mem0L ?(inE,pfV,validR W) //=; first last. -- by move=>?? /In_umfiltX [] _ /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_umfiltkE umfiltPtUn /= valid_omfUnR // (negbTE D). -congr (a (oeval a (dom _) f z0) k v). -by rewrite -umfiltk_subdomE; apply: omf_subdom. -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 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 pfunit dom0. Qed. - -Lemma umcnt_undef p : um_count p undef = 0. -Proof. by rewrite /um_count pfundef 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 umfiltPtE; 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. -by move=>W; rewrite /um_count umfiltUn // size_domUn // -umfiltUn // pfV. -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/In_dom_umfilt: E=>w H1 H2. - rewrite -{w H2}(In_fun H H2) in H1 *. - rewrite H1 omfF size_domF ?subn1 //=. - by apply/In_dom_umfilt; exists v. -rewrite omfF; case: ifP=>P; last by rewrite dom_free // E. -by move/negP: E; elim; apply/In_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: (normalP f)=>[->|W]; first by rewrite !umcnt_undef. -rewrite /um_count (umfilt_predD _ S) size_domUn ?leq_addr //. -by rewrite -umfilt_predD // pfV. -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: (normalP f)=>[->|D]; first by 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: (normalP f)=>[->|W]; first by rewrite !umcnt_undef. -rewrite /um_count umfilt_predU size_domUn //. -case: validUn=>//; rewrite ?(pfV,W) //. -move=>k /In_dom_umfilt [v1 P1 H1] /In_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. by rewrite umcnt_umfiltC umcnt_predU !(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 ?pfV ?D. -Qed. - -Lemma umcnt_eval0 R a p f (z0 : R) : um_count p f = 0 -> eval a p f z0 = z0. -Proof. -case: (normalP f)=>[->|W]; first by 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: (normalP f)=>[->|W]. -- by 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. - - -(*************************************) -(* Filtering maps of tagged elements *) -(*************************************) - -Section SideFilter. -Variables (T : eqType) (Us : T -> Type). - -(* could also be defined as *) -(* Definition side_m t : sigT Us -> option (Us t) := *) -(* fun '(Tag tx ux) => *) -(* if t =P tx is ReflectT pf then Some (cast Us pf ux) *) -(* else None. *) -(* but that doesn't reduce to then/else clause *) -(* if t == tx and t != tx, respectively *) -(* The following definition gets that reduction *) - -Definition side_m t : sigT Us -> option (Us t) := - fun '(Tag tx ux) => - if decP (t =P tx) is left pf then Some (cast Us pf ux) - else None. - -Lemma side_ocancel t : ocancel (side_m t) (Tag t). -Proof. by case=>tx vx /=; case: eqP=>//= pf; subst tx; rewrite eqc. Qed. - -End SideFilter. - -#[export] -Hint Extern 0 (ocancel (side_m _) (Tag _)) => - (apply: side_ocancel) : core. - -(* select all tags in h that equal t *) -(* differs from tags t in that it drops t *) -(* from the result map *) -(* whereas tags keeps the entries in result map unmodified *) - -Section Side. -Variables (K : ordType) (C : pred K) (T : eqType) (Us : T -> Type). -Variables (U : union_map C (sigT Us)). -Variables (Ut : forall t, union_map C (Us t)). -Variables t : T. - -Definition side_map : U -> Ut t := locked (omapv (side_m t)). -Lemma side_unlock : side_map = omapv (side_m t). -Proof. by rewrite /side_map -lock. Qed. -Lemma sidemap_is_omf : omap_fun_axiom side_map (side_m t \o snd). -Proof. by rewrite side_unlock; apply: omfE. Qed. -HB.instance Definition _ := - isOmapFun.Build K C (sigT Us) (Us t) U (Ut t) side_map sidemap_is_omf. - -(* explicit name for validity of side *) -Lemma valid_side (h : U) : valid (side_map h) = valid h. -Proof. exact: pfVE. Qed. - -Lemma In_side x (v : Us t) (h : U) : - (x, v) \In side_map h <-> (x, Tag t v) \In h. -Proof. -rewrite side_unlock In_omapX; split=>[|H]; last first. -- by exists (Tag t v)=>//=; case: eqP=>//= ?; rewrite eqc. -case; case=>t' v' /= H; case: eqP=>//= ?; subst t'. -by rewrite eqc; case=><-. -Qed. - -Lemma side_umfilt p q (h : U) : - (forall k v, - (k, Tag t v) \In h -> p (k, Tag t v) = q (k, v)) -> - side_map (um_filter p h) = um_filter q (side_map h). -Proof. -move=>H; rewrite side_unlock /um_filter !omap_omap eq_in_omf !omf_omap /=. -rewrite /side_m/obind/oapp/=; case=>k; case=>t' v X /=. -by case P : (p _); case: eqP=>//= ?; subst t'; rewrite eqc -H // P. -Qed. - -(* if p can only inspect keys *) -Lemma side_umfiltk (p : pred K) h : - side_map (um_filterk p h) = um_filterk p (side_map h). -Proof. by apply: side_umfilt. Qed. - -Lemma In_side_umfiltX x (v : Us t) p (h : U) : - (x, v) \In side_map (um_filter p h) <-> - p (x, Tag t v) /\ (x, Tag t v) \In h. -Proof. by rewrite In_side In_umfiltX. Qed. - -Lemma In_side_umfilt x (v : Us t) (p : pred (K * sigT Us)) (h : U) : - p (x, Tag t v) -> (x, Tag t v) \In h -> - (x, v) \In side_map (um_filter p h). -Proof. by move=>H1 H2; apply/In_side_umfiltX. Qed. - -Lemma In_dom_sideX x (h : U) : - reflect (exists a, (x, Tag t a) \In h) - (x \in dom (side_map h)). -Proof. -case: In_dom_omfX=>/= H; constructor. -- case: H=>[[t' v']][H]; rewrite /omfx/=. - by case: eqP=>// ?; subst t'; exists v'. -case=>v H'; elim: H. -exists (Tag t v); split=>//; rewrite /omfx/=. -by case: eqP=>// ?; rewrite eqc. -Qed. - -Lemma In_dom_side x a (h : U) : - (x, Tag t a) \In h -> x \in dom (side_map h). -Proof. by move/In_side/In_dom. Qed. - -Lemma sidePtE x e : - side_map (pts x e) = - if C x then - if decP (t =P tag e) is left pf then - pts x (cast Us pf (tagged e)) else Unit - else undef. -Proof. by case: e=>k v; rewrite omfPtE /omfx/=; case: eqP. Qed. - -Lemma dom_sidePt x e : - dom (side_map (pts x e)) = - if C x && (t == tag e) then [:: x] else [::]. -Proof. -rewrite sidePtE; case H : (C x)=>//=; last by rewrite dom_undef. -case: eqP=>[pf|] /=; last by rewrite dom0. -by rewrite domPtK H. -Qed. - -Lemma dom_sidePtUn k e h : - dom (side_map (pts k e \+ h)) =i - [pred x | valid (pts k e \+ h) & - (x == k) && (t == tag e) || (x \in dom (side_map h))]. -Proof. -move=>x; rewrite dom_omfPtUn !inE /omfx/= (andbC (x == k)). -by case: e=>t' v /=; case: eqP. -Qed. - -End Side. - -Arguments side_map {K C T Us U}. -Arguments In_side {K C T Us U} Ut {t x v h}. -Arguments In_dom_sideX {K C T Us U} Ut {t x h}. -Arguments In_dom_side {K C T Us U} Ut {t x a h}. - - -(* lemmas about two different sides *) - -Section Side2. -Variables (K : ordType) (C : pred K) (T : eqType) (Us : T -> Type). -Variables (U : union_map C (sigT Us)). -Variables (Ut : forall t, union_map C (Us t)). -Variables t1 t2 : T. - -Lemma sideN_all_predC (h : U) : - t1 <> t2 -> - all [predC dom (side_map Ut t1 h)] - (dom (side_map Ut t2 h)). -Proof. -move=>N; apply/allP=>x /In_domX [v1] /In_side H /=. -apply/negP=>/In_domX [v2] /In_side /(In_fun H). -by case=>X; rewrite X in N. -Qed. - -Lemma In_side_fun k (v1 : Us t1) (v2 : Us t2) (h : U) : - (k, v1) \In side_map Ut t1 h -> - (k, v2) \In side_map Ut t2 h -> - t1 = t2 /\ jmeq Us v1 v2. -Proof. -move/In_side=>H /In_side/(In_fun H) [?]; subst t2. -by move/inj_pair2=>->. -Qed. - -Lemma dom_sideE k (h : U) : - k \in dom (side_map Ut t1 h) -> - k \in dom (side_map Ut t2 h) -> - t1 = t2. -Proof. by case/In_domX=>v1 H1 /In_domX [v2] /(In_side_fun H1) []. Qed. - -Lemma dom_sideEX k (h : U) : - k \in dom (side_map Ut t1 h) -> - k \in dom (side_map Ut t2 h) = (t1 == t2). -Proof. -case/In_dom_sideX=>v H; case: (t1 =P t2)=>[?|N]. -- by subst t2; apply/In_dom_sideX; exists v. -by apply/In_dom_sideX; case=>w /(In_fun H) [] /N. -Qed. - -Lemma dom_sideN k1 k2 (h : U) : - k1 \in dom (side_map Ut t1 h) -> - k2 \in dom (side_map Ut t2 h) -> - t1 != t2 -> k1 != k2. -Proof. by move=>D1 D2; apply: contra=>E; rewrite -(dom_sideEX D1) (eqP E). Qed. - -End Side2. - -(* iterated tagging = when sigT Ts is used as tag *) - -Definition sliceT T (Ts : T -> Type) Us t := - {x : Ts t & Us (Tag t x)}. -Arguments sliceT {T Ts}. - -Section IterTag. -Variables (T : Type) (Ts : T -> Type) (Us : sigT Ts -> Type). - -(* tag re-association *) - -(* split (i.e., slice) tag in two *) -Definition slice_m : sigT Us -> sigT (sliceT Us) := - fun '(Tag (Tag t k) v) => Tag t (Tag k v). -(* conjoin (i.e. gather) first and second tag *) -Definition gather_m : sigT (sliceT Us) -> sigT Us := - fun '(Tag t (Tag k v)) => Tag (Tag t k) v. - -Variables (K : ordType) (C : pred K). -Variables (U : union_map C (sigT Us)). -Variables (S : union_map C (sigT (sliceT Us))). - -Definition slice : U -> S := mapv slice_m. -HB.instance Definition _ := OmapFun.copy slice slice. -Definition gather : S -> U := mapv gather_m. -HB.instance Definition _ := OmapFun.copy gather gather. - -Lemma In_slice x t (k : Ts t) (v : Us (Tag t k)) h : - (x, Tag t (Tag k v)) \In slice h <-> - (x, Tag (Tag t k) v) \In h. -Proof. -rewrite In_omfX; split=>[|H]; last first. -- by exists (Tag (Tag t k) v). -case; case; case=>t' k' v' H /=. -case=>?; subst t'=>/inj_pair2 ?; subst k'. -by move/inj_pair2/inj_pair2=><-. -Qed. - -Lemma In_gather x t (k : Ts t) (v : Us (Tag t k)) h : - (x, Tag (Tag t k) v) \In gather h <-> - (x, Tag t (Tag k v)) \In h. -Proof. -rewrite In_omfX; split=>[|H]; last first. -- by exists (Tag t (Tag k v)). -case; case=>t' [k' v'] H /=. -case=>?; subst t'=>/inj_pair2 ?; subst k'. -by move/inj_pair2=><-. -Qed. - -Lemma gather_slice h : gather (slice h) = h. -Proof. -case: (normalP h)=>[->|V]; first by rewrite !pfundef. -apply: umem_eq=>//; first by rewrite !pfV. -by case=>x [[t k v]]; rewrite In_gather In_slice. -Qed. - -Lemma slice_gather h : slice (gather h) = h. -Proof. -case: (normalP h)=>[->|V]; first by rewrite !pfundef. -apply: umem_eq=>//; first by rewrite !pfV. -by case=>x [t][k v]; rewrite In_slice In_gather. -Qed. - -End IterTag. - -Arguments slice {T Ts Us K C U}. -Arguments gather {T Ts Us K C U}. - - -(* grafting *) -(* clear side t from h and replace it with ht *) -(* side+graft are for union_maps what sel+splice are for finfuns *) - -Section GraftDef. -Variables (K : ordType) (C : pred K) (T : eqType) (Us : T -> Type). -Variables (U : union_map C (sigT Us)). -Variables (Ut : forall t, union_map C (Us t)). -Variables (h : U) (t : T). - -Definition graft (ht : Ut t) : U := - um_filterv (fun v => t != tag v) h \+ mapv (Tag t) ht. -End GraftDef. - -Arguments graft {K C T Us U Ut}. - -Section GraftLemmas. -Variables (K : ordType) (C : pred K) (T : eqType) (Us : T -> Type). -Variables (U : union_map C (sigT Us)) (Ut : forall t, union_map C (Us t)). -Implicit Type h : U. - -(* ht has disjoint domain with other sides of h *) -Definition disjoint_graft h t (ht : Ut t) := - forall tx x, - x \in dom (side_map Ut tx h) -> - x \in dom ht -> - tx = t. - -Lemma valid_graft h (t : T) (ht : Ut t) : - [/\ valid h, valid ht & disjoint_graft h ht] -> - valid (graft h t ht). -Proof. -case=>Vh V D. -rewrite validUnAE !pfV //=; apply/allP=>x /In_dom_omapX []. -move=>t1 [/In_dom /= H _]; apply/In_dom_umfilt. -case; case=>tx vx /= /eqP /= N. -by move/(In_side Ut)/In_dom/D/(_ H)/esym/N. -Qed. - -Lemma valid_graftUn h1 h2 (t : T) (ht : Ut t) : - [/\ valid (h1 \+ h2), valid ht, - disjoint_graft h1 ht & - forall x, x \in dom h2 -> x \in dom ht -> false] -> - valid (graft h1 t ht \+ h2). -Proof. -case=>Vh V D1 D2. -rewrite validUnAE valid_graft ?(validL Vh, validR Vh) //=. -apply/allP=>x D; apply/In_domX; case; case=>tx vx. -case/InUn; first by case/In_umfiltX=>_ /In_dom /(dom_inNLX Vh). -case/In_omapX=>w /In_dom /= H [?]; subst tx=>/inj_pair2 ?; subst w. -by move: (D2 _ D H). -Qed. - -(* if ht has disjoint domain with other sides of h *) -(* then side t (graft h t ht) = ht *) -Lemma side_graftE h t (ht : Ut t) : - valid h -> - disjoint_graft h ht -> - side_map Ut t (graft h t ht) = ht. -Proof. -move=>Vh D; case: (normalP ht)=>[->|V]. -- by rewrite /graft pfundef join_undef pfundef. -have W : valid (graft h t ht) by apply: valid_graft. -rewrite /graft pfjoin //=; apply/umem_eq=>//=; first by rewrite pfV2. -case=>k v; split=>[|H]. -- case/InUn; first by case/In_side/In_umfiltX; rewrite /= eqxx. - by case/In_side/In_omapX=>w H [] /inj_pair2 <-. -apply: InR; first by rewrite pfV2. -by apply/In_side/(In_omap _ H). -Qed. - -Lemma side_graftN h (t1 t2 : T) (ht2 : Ut t2) : - valid ht2 -> - disjoint_graft h ht2 -> - t1 != t2 -> - side_map Ut t1 (graft h t2 ht2) = side_map Ut t1 h. -Proof. -move=>V D N; case: (normalP h)=>[->|Vh]. -- by rewrite /graft pfundef undef_join. -have W : valid (graft h t2 ht2) by apply: valid_graft. -rewrite /graft pfjoin //=; apply/umem_eq=>/=. -- by rewrite pfV2. -- by rewrite pfV. -case=>k v; split. -- case/InUn; first by case/In_side/In_umfiltX=>_ /(In_side Ut). - case/In_side/In_omapX=>w /In_dom /= H [?]; subst t2. - by rewrite eqxx in N. -move/In_side=>H; apply/InL; first by rewrite pfV2. -by apply/In_side/In_umfiltX; rewrite /= eq_sym N. -Qed. - -End GraftLemmas. - - -(***********) -(* dom_map *) -(***********) - -(* Viewing domain as finite set instead of sequence. *) -(* Because sequences don't have PCM structure, *) -(* dom_maps facilitate proving disjointness of sequences *) -(* by enabling the pcm lemmas. *) - -Section DomMap. -Variables (K : ordType) (C : pred K) (V : Type). -Variables (U : union_map C V) (U' : union_map C unit). - -Definition dom_map (x : U) : U' := omap (fun => Some tt) x. - -HB.instance Definition _ := OmapFun.copy dom_map dom_map. - -Lemma dom_mapE x : dom (dom_map x) = dom x. -Proof. by apply: dom_omf_some. Qed. - -(* like pfjoinT but without validity condition *) -Lemma dom_mapUn (x y : U) : - dom_map (x \+ y) = dom_map x \+ dom_map y. -Proof. -case W : (valid (x \+ y)); first exact: pfjoinT. -move/invalidE: (negbT W)=>->; rewrite pfundef. -case: validUn W=>//. -- by move/invalidE=>->; rewrite pfundef undef_join. -- by move/invalidE=>->; rewrite pfundef 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 domeq_dom_mapL x : dom_eq (dom_map x) x. -Proof. by rewrite /dom_eq pfVE dom_mapE !eq_refl. Qed. - -Lemma domeq_dom_mapR x : dom_eq x (dom_map x). -Proof. by apply: domeq_sym; apply: domeq_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 C V). -Variables (U' : union_map C unit) (U'' : union_map C unit). - -Lemma dom_map_idemp (x : U) : dom_map (dom_map x : U') = dom_map x :> U''. -Proof. by rewrite /dom_map omap_omap. Qed. - -End DomMapIdempotent. - - -(*****************) -(* Map inversion *) -(*****************) - -Section MapInvert. -Variables (K V : ordType) (C : pred K) (C' : pred V). -Variables (U : union_map C V) (U' : union_map C' K). - -(* inverting f = flipping each pts k v in f into pts v k *) -Definition invert (f : U) : U' := - um_foldl (fun r k v => r \+ pts v k) Unit undef f. - -(* invert isn't morphism as pfV doesn't hold *) -(* but invert has several morphism properties *) - -Lemma invert_undef : invert undef = undef. -Proof. by rewrite /invert umfoldl_undef. Qed. - -Lemma invert0 : invert Unit = Unit. -Proof. by rewrite /invert umfoldl0. Qed. - -Lemma invertUn f1 f2 : - valid (f1 \+ f2) -> - invert (f1 \+ f2) = invert f1 \+ invert f2. -Proof. -move=>W; rewrite /invert umfoldlUn_frame ?unitR //. -by move=>*; rewrite joinAC. -Qed. - -Lemma invertPt k v : C k -> invert (pts k v) = pts v k. -Proof. by move=>H; rewrite /invert umfoldlPt unitL H. Qed. - -Lemma invertPtUn k v f : - valid (pts k v \+ f) -> - invert (pts k v \+ f) = pts v k \+ invert f. -Proof. by move=>W; rewrite invertUn // invertPt // (validPtUn_cond W). Qed. - -Lemma dom_invert f : valid (invert f) -> dom (invert f) =i range f. -Proof. -rewrite /invert/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_invert f : - valid (invert 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 invert_undef !valid_undef. -- by rewrite invert0 !valid_unit range0. -rewrite invertPtUn 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_invert ->. -Qed. - -(* The next few lemmas that depend on valid (invert f) *) -(* can be strenghtened to require just uniq (range f) and *) -(* all C' (range f). *) -(* However, the formulation with the hypothesis valid (invert f) *) -(* is packaged somewhat better, and is what's encountered in practice. *) - -Lemma range_invert f : - valid (invert f : U') -> - range (invert f) =i dom f. -Proof. -elim/um_indf: f=>[||k v f IH W P]. -- by rewrite invert_undef valid_undef. -- by rewrite invert0 range0 dom0. -rewrite invertPtUn // => W' x. -rewrite rangePtUn inE W' domPtUn inE W /=. -by case: eqP=>//= _; rewrite IH // (validR W'). -Qed. - -Lemma In_invert k v f : - valid (invert f) -> - (k, v) \In f <-> (v, k) \In invert f. -Proof. -elim/um_indf: f k v=>[||x w f IH W /(order_path_min (@trans K)) P] k v. -- by rewrite invert_undef valid_undef. -- by rewrite invert0; split=>/In0. -move=>W'; rewrite invertPtUn // !InPtUnE //; last by rewrite -invertPtUn. -rewrite IH; first by split; case=>[[->->]|]; auto. -rewrite !valid_invert rangePtUnK // (validR W) in W' *. -by case/and3P: W'=>_ /= /andP [_ ->] /andP [_ ->]. -Qed. - -Lemma uniq_range_invert f : uniq (range (invert f)). -Proof. -case: (normalP (invert f))=>[->|W]; first by 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/In_assocs/(In_invert _ _ W): H1=>H1. -move/mem_seqP/In_assocs/(In_invert _ _ W): H2=>H2. -by rewrite (In_fun H1 H2). -Qed. - -Lemma all_range_invert f : all C (range (invert f)). -Proof. -apply: (@sub_all _ (fun k => k \in dom f)); first by move=>x /dom_cond. -by apply/allP=>x H; rewrite -range_invert //; apply: range_valid H. -Qed. - -Lemma sorted_range_invert f : - valid (invert f) -> - sorted ord (range f) -> - sorted ord (range (invert f)). -Proof. -move=>W /ummonoP X; apply/ummonoP=>v v' k k'. -move/(In_invert _ _ W)=>H1 /(In_invert _ _ 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 MapInvert. - -Arguments invert {K V C C' U U'}. - - -Section InvertLaws. -Variables (K V : ordType) (C : pred K) (C' : pred V). -Variables (U : union_map C V) (U' : union_map C' K). - -Lemma valid_invert_idemp (f : U) : - valid (invert (invert f : U') : U) = valid (invert f : U'). -Proof. by rewrite valid_invert uniq_range_invert all_range_invert !andbT. Qed. - -Lemma invert_idemp (f : U) : - valid (invert f : U') -> - invert (invert f : U') = f. -Proof. -move=>W; apply/umem_eq. -- by rewrite valid_invert_idemp. -- by move: W; rewrite valid_invert; case/and3P. -move=>x; split=>H. -- by apply/(In_invert _ _ W); apply/(@In_invert _ _ _ _ _ U) =>//; case: H. -case: x H=>k v /(In_invert _ _ W)/In_invert; apply. -by rewrite valid_invert_idemp. -Qed. - -End InvertLaws. - -Arguments In_invert {K V C C' U U' k v f}. - - -(***************************) -(* Composition of two maps *) -(***************************) - -(* composing maps as functions, rather than PCMs *) - -Section MapComposition. -Variables (K1 K2 : ordType) (C1 : pred K1) (C2 : pred K2) (V : Type). -Variables (Uf : union_map C1 K2) (Ug : union_map C2 V) (U : union_map 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 umcomp_subdom 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: umcomp_subdom. -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 umcompfPtE 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 umcompfPt g k v : - C1 k -> - um_comp g (pts k v) = - if find v g is Some w then pts k w else Unit. -Proof. by move=>H; rewrite umcompfPtE H. 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 // umcompfPtE (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 pfunit. - by case: ifP=>//; rewrite valid_unit. -rewrite umcompfUn // umcompfPtE (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 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). -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) -> 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. - -#[export] -Hint Resolve umall_undef umall0 : core. - - -(***********) -(* um_allb *) -(***********) - -(* decidable um_all *) - -Section MapAllDecidable. -Variables (K : ordType) (C : pred K) (V : Type) (U : union_map C V). -Implicit Types (f : U) (p : pred (K*V)). - -Definition um_allb p f := um_count p f == size (dom f). - -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 umallb_is_pcm_morph p : pcm_morph_axiom relT (um_allb p). -Proof. -rewrite /um_allb; split=>[|x y W _]; first by rewrite umcnt0 dom0. -split=>//; apply/idP/idP. -- by move/umallbP=>H; apply/andP; split; apply/umallbP=>k X; - apply: H; [apply: InL | apply: InR]. -case/andP=>/umallbP X1 /umallbP X2; apply/umallbP=>k. -by case/InUn; [apply: X1 | apply: X2]. -Qed. - -HB.instance Definition _ p := - isPCM_morphism.Build U bool (um_allb p) (umallb_is_pcm_morph p). -HB.instance Definition _ p := - isFull_PCM_morphism.Build U bool (um_allb p) (fun x y => erefl). - -Lemma umallb0 p : um_allb p Unit. -Proof. exact: pfunit. Qed. - -Lemma umallbUn p f1 f2 : - valid (f1 \+ f2) -> - um_allb p (f1 \+ f2) = um_allb p f1 && um_allb p f2. -Proof. exact: pfjoinT. Qed. - -(* bool isn't tpcm, so we can't declare umallb tpcm morphism *) -(* however, we can prove some properties about under *) - -Lemma umallb_undef p : um_allb p undef. -Proof. by rewrite /um_allb umcnt_undef dom_undef. 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 C V). -Variables (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]; first 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. - -(* big join and union maps *) - -Section BigOpsUM. -Variables (K : ordType) (C : pred K) (T : Type). -Variables (U : union_map C T). -Variables (I : Type) (f : I -> U). - -Lemma big_domUn (xs : seq I) : - dom (\big[join/Unit]_(i <- xs) f i) =i - [pred x | valid (\big[join/Unit]_(i <- xs) f i) & - has (fun i => x \in dom (f i)) xs]. -Proof. -elim: xs=>[|x xs IH] i; first by rewrite big_nil inE /= dom0 valid_unit. -rewrite big_cons /= inE domUn inE IH inE /=. -by case V : (valid _)=>//=; rewrite (validR V) /=. -Qed. - -Lemma big_domUnE (xs : seq I) a : - valid (\big[join/Unit]_(i <- xs) f i) -> - a \in dom (\big[join/Unit]_(i <- xs) f i) = - has (fun i => a \in dom (f i)) xs. -Proof. by move=>V; rewrite big_domUn inE V. Qed. - -(* we can construct a big validity, if we know that *) -(* the list contains unique elements *) -Lemma big_validV2I (xs : seq I) : - Uniq xs -> - (forall i, i \In xs -> valid (f i)) -> - (forall i j, i \In xs -> j \In xs -> i <> j -> valid (f i \+ f j)) -> - valid (\big[join/Unit]_(i <- xs) f i). -Proof. -elim: xs=>[|x xs IH] /=; first by rewrite big_nil valid_unit. -case=>X Uq H1 H2; rewrite big_cons validUnAE. -rewrite H1 /=; last by rewrite InE; left. -rewrite IH //=; last first. -- by move=>i j Xi Xj; apply: H2; rewrite InE; right. -- by move=>i Xi; apply: H1; rewrite InE; right. -apply/allP=>a /=; apply: contraL=>Dx; apply/negP. -rewrite big_domUn inE=>/andP [V] /hasPIn [y Y Dy]. -have : x <> y by move=>N; rewrite N in X; move/X: Y. -move/(H2 x y (or_introl erefl) (or_intror Y)). -by case: validUn=>// _ _ /(_ a Dx); rewrite Dy. -Qed. - -Lemma big_size_dom (xs : seq I) : - valid (\big[join/Unit]_(i <- xs) f i) -> - size (dom (\big[join/Unit]_(i <- xs) f i)) = - \sum_(i <- xs) (size (dom (f i))). -Proof. -elim: xs=>[|x xs IH] /=; first by rewrite !big_nil dom0. -rewrite !big_cons=>V; rewrite size_domUn //; congr (size _ + _). -by apply/IH/(validR V). -Qed. - -Lemma big_find_some (xs : seq I) a i v : - valid (\big[join/Unit]_(i <- xs) f i) -> - i \In xs -> - find a (f i) = some v -> - find a (\big[join/Unit]_(i <- xs) f i) = some v. -Proof. -elim: xs=>[|x xs IH /[swap]] //; rewrite big_cons InE. -case=>[<-{x}|Xi] V E; first by rewrite findUnL // (find_some E). -rewrite findUnR // big_domUnE ?(validR V) //=. -rewrite ifT; first by apply: IH (validR V) Xi E. -by apply/hasPIn; exists i=>//; apply: find_some E. -Qed. - -Lemma big_find_someD (xs : seq I) a i v : - i \In xs -> - a \in dom (f i) -> - find a (\big[join/Unit]_(x <- xs) f x) = Some v -> - find a (f i) = Some v. -Proof. -elim: xs v=>[|y xs IH] v //=; rewrite big_cons InE. -case=>[->|Xi] Da /[dup]/In_find/In_valid V; first by rewrite findUnL // Da. -rewrite findUnR // big_domUnE ?(validR V) //=. -by rewrite ifT; [apply: IH | apply/hasPIn; exists i]. -Qed. - -Lemma big_find_someX (xs : seq I) a v : - find a (\big[join/Unit]_(i <- xs) f i) = Some v -> - exists2 i, i \In xs & find a (f i) = Some v. -Proof. -elim: xs=>[|x xs IH]; first by rewrite big_nil find0E. -move/[dup]=>E1; rewrite big_cons=>/[dup] E /find_some. -rewrite domUn inE=>/andP [V] /orP [] D; last first. -- move: E; rewrite findUnR // D=>/IH [j Dj Xj]. - by exists j=>//; rewrite InE; right. -case/In_domX: D=>w /In_find /[dup] X. -move/(big_find_some (dom_valid (find_some E1)) (or_introl erefl)). -by rewrite E1; case=>->; exists x=>//; rewrite InE; left. -Qed. - -Lemma big_find_someXP (xs : seq I) P a v : - find a (\big[join/Unit]_(i <- xs | P i) f i) = Some v -> - exists i, [/\ i \In xs, P i & find a (f i) = Some v]. -Proof. -by rewrite -big_filter=>/big_find_someX [i] /Mem_filter [H1 H2 H3]; exists i. -Qed. - -Lemma bigIn (xs : seq I) a i v : - valid (\big[join/Unit]_(i <- xs) f i) -> - i \In xs -> - (a, v) \In f i -> - (a, v) \In \big[join/Unit]_(i <- xs) f i. -Proof. by move=>V H /In_find/(big_find_some V H)/In_find. Qed. - -Lemma bigInD (xs : seq I) a x v : - x \In xs -> - a \in dom (f x) -> - (a, v) \In \big[join/Unit]_(i <- xs) f i -> - (a, v) \In f x. -Proof. by move=>X D /In_find/(big_find_someD X D)/In_find. Qed. - -Lemma bigInX (xs : seq I) a v : - (a, v) \In \big[join/Unit]_(i <- xs) f i -> - exists2 i, i \In xs & (a, v) \In f i. -Proof. by case/In_find/big_find_someX=>x X /In_find; exists x. Qed. - -Lemma bigInXP (xs : seq I) P a v : - (a, v) \In \big[join/Unit]_(i <- xs | P i) f i -> - exists i, [/\ i \In xs, P i & (a, v) \In f i]. -Proof. by case/In_find/big_find_someXP=>x [X1 X2 /In_find]; exists x. Qed. - -Lemma bigF (xs : seq I) k : - valid (\big[join/Unit]_(i <- xs) f i) -> - free (\big[join/Unit]_(i <- xs) f i) k = - \big[join/Unit]_(i <- xs) free (f i) k. -Proof. -elim: xs=>[|x xs IH]; first by rewrite !big_nil free0. -rewrite !big_cons => V; rewrite freeUn. -rewrite -IH ?(validR V) //; case: ifP=>// D. -rewrite !freeND //; apply: contraFN D=>D; -by rewrite domUn inE D V //= orbT. -Qed. - -(* if xs is domain of some map *) - -Lemma bigD1FE (k : K) (F : K -> U) (g : U) : - \big[join/Unit]_(i <- dom g) F i = - if k \in dom g then - F k \+ \big[join/Unit]_(i <- dom (free g k)) F i - else \big[join/Unit]_(i <- dom (free g k)) F i. -Proof. -case: ifP=>D; last by rewrite freeND // (negbT D). -rewrite (bigD1_seq k) //= -big_filter (_ : filter _ _ = dom (free g k)) //. -apply/ord_sorted_eq; first by rewrite sorted_filter. -- by rewrite sorted_dom. -by move=>z; rewrite mem_filter domF inE eq_sym; case: (k =P z). -Qed. - -Lemma bigPtUnE (x : K) (a : T) (g : U) (F : K -> U) : - \big[join/Unit]_(i <- dom (pts x a \+ g)) F i = - if valid (pts x a \+ g) then - F x \+ \big[join/Unit]_(i <- dom g) F i - else Unit. -Proof. -rewrite (bigD1FE x) domPtUnE. -case: (normalP (pts x a \+ g))=>[->|V]. -- by rewrite free_undef dom_undef big_nil. -by rewrite freePtUn. -Qed. - -Lemma bigPtUn (x : K) (a : T) (g : U) (F : K -> U) : - valid (pts x a \+ g) -> - \big[join/Unit]_(i <- dom (pts x a \+ g)) F i = - F x \+ \big[join/Unit]_(i <- dom g) F i. -Proof. by move=>V; rewrite bigPtUnE V. Qed. - -Lemma bigDUE (x : K) (a : T) (g : U) (F : K -> U) : - \big[join/Unit]_(i <- dom (upd x a g)) F i = - if C x && valid g then - F x \+ \big[join/Unit]_(i <- dom (free g x)) F i - else Unit. -Proof. -by rewrite (upd_eta x) bigPtUnE validPtUn validF domF inE eqxx andbT. -Qed. - -End BigOpsUM. - -Prenex Implicits big_find_some big_find_someD. -Prenex Implicits big_find_someX big_find_someXP bigIn bigInD bigInX bigInXP. - -(* if the index type is eqtype, we can have a somewhat better *) -(* big validity lemma *) - -Section BigOpsUMEq. -Variables (K : ordType) (C : pred K) (T : Type) (U : union_map C T). -Variables (I : eqType) (f : I -> U). - -Lemma big_validP (xs : seq I) : - reflect - [/\ forall i, i \in xs -> valid (f i), - forall i k, i \in xs -> k \in dom (f i) -> count_mem i xs = 1 & - forall i j, i \in xs -> j \in xs -> i <> j -> valid (f i \+ f j)] - (valid (\big[join/Unit]_(i <- xs) f i)). -Proof. -elim: xs=>[|x xs IH] /=; first by rewrite big_nil valid_unit; constructor. -rewrite big_cons validUnAE. -case H1 : (valid (f x))=>/=; last first. -- by constructor; case=>/(_ x); rewrite inE eqxx H1=>/(_ erefl). -rewrite andbC; case: allP=>/= H2; last first. -- constructor; case=>H3 H4 H5; apply: H2=>k; rewrite big_domUn inE. - case/andP=>V /hasP [y]; case : (x =P y)=>[<- /[swap]|N X D]. - - by move/(H4 x); rewrite inE eqxx add1n=>/(_ erefl) []/count_memPn/negbTE ->. - by apply: dom_inNR (H5 _ _ _ _ N) D; rewrite inE ?eqxx ?X ?orbT. -case: {2 3}(valid _) / IH (erefl (valid (\big[join/Unit]_(j <- xs) f j))) -=> H V; constructor; last first. -- case=>H3 H4 H5; apply: H; split; last 1 first. - - by move=>i k X1 X2; apply: H5; rewrite inE ?X1 ?X2 orbT. - - by move=>i X; apply: H3; rewrite inE X orbT. - move=>i k X1 X2; move: (H4 i k); rewrite inE X1 orbT=>/(_ erefl X2). - case: (x =P i) X1 X2=>[<-{i} X1 X2|]; last by rewrite add0n. - by rewrite add1n; case=>/count_memPn; rewrite X1. -case: H=>H3 H4 H5; split; first by move=>i; rewrite inE=>/orP [/eqP ->|/H3]. -- move=>i k; rewrite inE eq_sym; case: (x =P i)=>[<- _|X] /= D; last first. - - by rewrite add0n; apply: H4 D. - rewrite add1n; congr (_.+1); apply/count_memPn. - apply: contraL (D)=>X; apply: H2; rewrite big_domUn V inE. - by apply/hasP; exists x. -move=>i j; rewrite !inE=>/orP [/eqP ->{i}| Xi] /orP [/eqP ->{j}|Xj] // N. -- rewrite validUnAE H1 H3 //=; apply/allP=>k D; apply: H2. - by rewrite big_domUn V inE; apply/hasP; exists j. -- rewrite validUnAE H3 ?H1 //=; apply/allP=>k; apply: contraL=>D; apply: H2. - by rewrite big_domUn V inE; apply/hasP; exists i. -by apply: H5 Xi Xj N. -Qed. - -End BigOpsUMEq. - -Section BigCatSeqUM. -Variables (K : ordType) (C : pred K) (T : Type) (U : union_map C T). -Variables (I : eqType) (g : I -> U). - -Lemma big_valid_dom_seq (xs : seq I) : - (valid (\big[join/Unit]_(i <- xs) g i) = - all (fun i => valid (g i)) xs && - uniq (\big[cat/[::]]_(i <- xs) dom (g i))) * - (dom (\big[join/Unit]_(i <- xs) g i) =i - if valid (\big[join/Unit]_(i <- xs) g i) then - \big[cat/[::]]_(i <- xs) dom (g i) - else [::]). -Proof. -elim: xs=>[|x xs [IH1 IH2]]; first by rewrite !big_nil valid_unit /= dom0. -rewrite !big_cons cat_uniq /=; split=>[|z]; last first. -- rewrite domUn inE; case V : (valid _)=>//=. - by rewrite mem_cat IH2 (validR V). -rewrite validUnAE IH1 -all_predC -!andbA uniq_dom /=. -case V : (valid (g x))=>//=; case A : (all _)=>//=. -rewrite [RHS]andbC; case Uq : (uniq _)=>//=. -by apply: eq_all_r=>i; rewrite IH2 IH1 A Uq. -Qed. - -Lemma big_valid_seq (xs : seq I) : - valid (\big[join/Unit]_(i <- xs) g i) = - all (fun i => valid (g i)) xs && - uniq (\big[cat/[::]]_(i <- xs) dom (g i)). -Proof. by rewrite big_valid_dom_seq. Qed. - -Lemma big_dom_seq (xs : seq I) : - dom (\big[join/Unit]_(i <- xs) g i) =i - if valid (\big[join/Unit]_(i <- xs) g i) then - \big[cat/[::]]_(i <- xs) dom (g i) - else [::]. -Proof. by move=>x; rewrite big_valid_dom_seq. Qed. - -End BigCatSeqUM. - -Lemma bigPt_valid (K : ordType) (C : pred K) T (U : union_map C T) - (g : U) (F : K -> T) : - valid (\big[join/Unit]_(i <- dom g) pts i (F i) : U). -Proof. -rewrite big_valid_seq; apply/andP; split. -- by apply/allP=>z Dz; rewrite validPt (dom_cond Dz). -apply/uniq_big_catEX=>//; split=>[i Di|i j k Di Dj] //. -by rewrite !domPt !inE (dom_cond Di) (dom_cond Dj)=>/eqP -> /eqP. -Qed. - -Section OMapBig. -Variables (K : ordType) (C : pred K) (T T' : Type). -Variables (U : union_map C T) (U' : union_map C T'). -Variables (I : Type) (f : I -> U). - -Lemma big_omapVUn (a : K * T -> option T') ts : - omap a (\big[join/Unit]_(x <- ts) f x) = - if valid (\big[join/Unit]_(x <- ts) f x) - then \big[join/Unit]_(x <- ts) omap a (f x) - else undef : U'. -Proof. -elim: ts=>[|t ts IH]; first by rewrite !big_nil omap0 valid_unit. -by rewrite !big_cons omapVUn IH; case: ifP=>// /validR ->. -Qed. - -End OMapBig. - - -(* DEVCOMMENT: *) -(* remove "tests" for release *) -Lemma xx (f : umap nat nat) : 3 \in dom (free f 2). -rewrite domF -domF. -Abort. -(* /DEVCOMMENT *) From 0c83825cedcf45670a2432ed7a6f28f0179ca489 Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Mon, 23 Sep 2024 10:39:43 +0200 Subject: [PATCH 81/93] added Marcos to the list of authors --- meta.yml | 2 ++ 1 file changed, 2 insertions(+) diff --git a/meta.yml b/meta.yml index f36d9a1..04c4461 100644 --- a/meta.yml +++ b/meta.yml @@ -58,6 +58,8 @@ authors: initial: false - name: Alexander Gryzlov initial: false +- name: Marcos Grandury + initial: false publications: - pub_url: https://software.imdea.org/~aleks/papers/reflect/reflect.pdf From ae6f24f74d93c03c797a015c556445ebdfdf5f0f Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Mon, 23 Sep 2024 10:40:57 +0200 Subject: [PATCH 82/93] regenerated .opam files from meta.yml --- .github/workflows/docker-action.yml | 5 +++-- README.md | 8 ++++---- coq-htt-examples.opam | 8 ++++---- coq-htt.opam | 10 +++++----- dune-project | 4 ++-- 5 files changed, 18 insertions(+), 17 deletions(-) diff --git a/.github/workflows/docker-action.yml b/.github/workflows/docker-action.yml index a774d92..f703150 100644 --- a/.github/workflows/docker-action.yml +++ b/.github/workflows/docker-action.yml @@ -16,8 +16,9 @@ jobs: strategy: matrix: image: - - 'mathcomp/mathcomp:1.17.0-coq-8.15' - - 'mathcomp/mathcomp:1.17.0-coq-8.18' + - 'mathcomp/mathcomp:2.2.0-coq-8.19' + - 'mathcomp/mathcomp:2.2.0-coq-8.20' + - 'mathcomp/mathcomp:latest-coq-dev' fail-fast: false steps: - uses: actions/checkout@v2 diff --git a/README.md b/README.md index 0fe5f74..e9a1336 100644 --- a/README.md +++ b/README.md @@ -37,13 +37,13 @@ that HTT implements Separation logic as a shallow embedding in Coq. - Aleksandar Nanevski (initial) - Germán Andrés Delbianco - Alexander Gryzlov + - Marcos Grandury - License: [Apache-2.0](LICENSE) -- Compatible Coq versions: Coq 8.15 to 8.17 +- Compatible Coq versions: Coq 8.19 to 8.20 - Additional dependencies: - - [MathComp ssreflect 1.17](https://math-comp.github.io) + - [MathComp ssreflect 2.2](https://math-comp.github.io) - [MathComp algebra](https://math-comp.github.io) - - [MathComp fingroup](https://math-comp.github.io) - - [FCSL-PCM 1.8](https://github.com/imdea-software/fcsl-pcm) + - [FCSL-PCM 2.0](https://github.com/imdea-software/fcsl-pcm) - [Dune](https://dune.build) 2.5 or later - Coq namespace: `htt` - Related publication(s): diff --git a/coq-htt-examples.opam b/coq-htt-examples.opam index caeded0..dd2c9e9 100644 --- a/coq-htt-examples.opam +++ b/coq-htt-examples.opam @@ -33,11 +33,10 @@ that HTT implements Separation logic as a shallow embedding in Coq.""" build: [make "-j%{jobs}%"] install: [make "install"] depends: [ - "coq" { (>= "8.15" & < "8.19~") | (= "dev") } - "coq-mathcomp-ssreflect" { (>= "1.17.0" & < "1.18~") | (= "dev") } + "coq" { (>= "8.19" & < "8.21~") | (= "dev") } + "coq-mathcomp-ssreflect" { (>= "2.2.0" & < "2.3~") | (= "dev") } "coq-mathcomp-algebra" - "coq-mathcomp-fingroup" - "coq-fcsl-pcm" { (>= "1.8.0" & < "1.9~") | (= "dev") } + "coq-fcsl-pcm" { (>= "2.0.0" & < "2.1~") | (= "dev") } ] tags: [ @@ -50,4 +49,5 @@ authors: [ "Aleksandar Nanevski" "Germán Andrés Delbianco" "Alexander Gryzlov" + "Marcos Grandury" ] diff --git a/coq-htt.opam b/coq-htt.opam index 1517652..95ee94a 100644 --- a/coq-htt.opam +++ b/coq-htt.opam @@ -33,12 +33,11 @@ that HTT implements Separation logic as a shallow embedding in Coq.""" build: ["dune" "build" "-p" name "-j" jobs] depends: [ - "dune" {>= "2.5"} - "coq" { (>= "8.15" & < "8.19~") | (= "dev") } - "coq-mathcomp-ssreflect" { (>= "1.17.0" & < "1.18~") | (= "dev") } + "dune" {>= "3.6"} + "coq" { (>= "8.19" & < "8.21~") | (= "dev") } + "coq-mathcomp-ssreflect" { (>= "2.2.0" & < "2.3~") | (= "dev") } "coq-mathcomp-algebra" - "coq-mathcomp-fingroup" - "coq-fcsl-pcm" { (>= "1.8.0" & < "1.9~") | (= "dev") } + "coq-fcsl-pcm" { (>= "2.0.0" & < "2.1~") | (= "dev") } ] tags: [ @@ -51,4 +50,5 @@ authors: [ "Aleksandar Nanevski" "Germán Andrés Delbianco" "Alexander Gryzlov" + "Marcos Grandury" ] diff --git a/dune-project b/dune-project index bf8c8f3..4d97b3b 100644 --- a/dune-project +++ b/dune-project @@ -1,3 +1,3 @@ -(lang dune 2.5) -(using coq 0.2) +(lang dune 3.6) +(using coq 0.6) (name htt) From 7ecce21bba6f6d095fc621ab2817e4c1ee43ca52 Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Mon, 23 Sep 2024 10:49:30 +0200 Subject: [PATCH 83/93] removed devcomments --- examples/hashtab.v | 3 --- htt/domain.v | 6 ------ htt/dune | 11 ----------- 3 files changed, 20 deletions(-) delete mode 100644 htt/dune diff --git a/examples/hashtab.v b/examples/hashtab.v index 4c8fcca..bc5e7c4 100644 --- a/examples/hashtab.v +++ b/examples/hashtab.v @@ -22,9 +22,6 @@ From htt Require Import array kvmaps. (* hash table is array of buckets, i.e. KV maps *) (* bucket indices are provided by the hash function *) (* using dynaming kv-maps for buckets *) -(* DEVCOMMENT: *) -(* it's possible to develop this with static buckers *) -(* /DEVCOMMENT *) Module Type Hashtab_sig. Parameter root : forall {K : ordType} {V : Type} (buckets : dkvm K V) diff --git a/htt/domain.v b/htt/domain.v index cbb13ef..8ee1d97 100644 --- a/htt/domain.v +++ b/htt/domain.v @@ -1028,10 +1028,4 @@ HB.instance Definition _ := isContFun.Build (S1*S2)%type (T1*T2)%type (f1 \* f2) prod_is_cont. End ProdCont. -(* DEVCOMMENT *) -(* TODO: *) -(* 1. limit of a chain of continuous functions is a continuous function *) -(* 2. CPO of continuous functions should be the subCPO of functions *) -(* 3. kleene_lfp is a continuous operation *) -(* /DEVCOMMENT *) diff --git a/htt/dune b/htt/dune deleted file mode 100644 index 501fc38..0000000 --- a/htt/dune +++ /dev/null @@ -1,11 +0,0 @@ -; This file was generated from `meta.yml`, please do not edit manually. - -(coq.theory - (name htt) - (package coq-htt) - (synopsis "Hoare Type Theory") - (flags :standard - -w -notation-overridden - -w -local-declaration - -w -redundant-canonical-projection - -w -projection-no-head-constant)) From 5113278451652cd400227255f3d5d528f042e89f Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Mon, 23 Sep 2024 12:06:50 +0200 Subject: [PATCH 84/93] removed some lemmas that have recently been included into mathcomp --- examples/quicksort.v | 25 ------------------------- 1 file changed, 25 deletions(-) diff --git a/examples/quicksort.v b/examples/quicksort.v index 8f88cd7..d405474 100644 --- a/examples/quicksort.v +++ b/examples/quicksort.v @@ -28,31 +28,6 @@ Local Open Scope order_scope. (* but the development is repeated here to make the files *) (* self-contained *) -Lemma perm_onC {T : finType} (S1 S2 : {set T}) (u1 u2 : {perm T}) : - perm_on S1 u1 -> - perm_on S2 u2 -> - [disjoint S1 & S2] -> - commute u1 u2. -Proof. -move=> pS1 pS2 S12; apply/permP => t; rewrite !permM. -case/boolP : (t \in S1) => tS1. -- have /[!disjoint_subset] /subsetP {}S12 := S12. - by rewrite !(out_perm pS2) //; apply: S12; rewrite // perm_closed. -case/boolP : (t \in S2) => tS2. -- have /[1!disjoint_sym] /[!disjoint_subset] /subsetP {}S12 := S12. - by rewrite !(out_perm pS1) //; apply: S12; rewrite // perm_closed. -by rewrite (out_perm pS1) // (out_perm pS2) // (out_perm pS1). -Qed. - -Lemma tperm_on {T : finType} (x y : T) : - perm_on [set x; y] (tperm x y). -Proof. -apply/subsetP => z /[!inE]; case: tpermP => [->|->|]; -by rewrite eqxx // orbT. -Qed. - -(****) - Lemma leq_choose a b : a < b -> sumbool (a.+1 == b) (a.+1 < b). From 03220f8b99daf238b708b86cc37a71bbec7443cf Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Tue, 24 Sep 2024 10:10:38 +0200 Subject: [PATCH 85/93] regenerated htt/dune file --- htt/dune | 11 +++++++++++ regenerate.sh | 2 +- 2 files changed, 12 insertions(+), 1 deletion(-) create mode 100644 htt/dune diff --git a/htt/dune b/htt/dune new file mode 100644 index 0000000..501fc38 --- /dev/null +++ b/htt/dune @@ -0,0 +1,11 @@ +; This file was generated from `meta.yml`, please do not edit manually. + +(coq.theory + (name htt) + (package coq-htt) + (synopsis "Hoare Type Theory") + (flags :standard + -w -notation-overridden + -w -local-declaration + -w -redundant-canonical-projection + -w -projection-no-head-constant)) diff --git a/regenerate.sh b/regenerate.sh index 1e1d9f5..c6520b8 100755 --- a/regenerate.sh +++ b/regenerate.sh @@ -24,7 +24,7 @@ for f in "$srcdir"/*.mustache; do mustache='{{ dune }}' bool=$(get_yaml meta.yml <<<"$mustache") if [ -n "$bool" ] && [ "$bool" != false ]; then - mkdir -p -v theories && target="theories/$target" + mkdir -p -v htt && target="htt/$target" else continue fi From be7aaa895fa2bad1ad7275aa734b7f74fe85513b Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Fri, 27 Sep 2024 15:53:18 +0200 Subject: [PATCH 86/93] playing with dune --- coq-htt-examples.opam | 53 ------------------- coq-htt.opam | 5 +- htt/dune | 2 +- meta.yml | 2 +- regenerate.sh | 2 +- .../coq-htt-examples.opam.mustache | 50 ----------------- templates-extra/dune.mustache | 2 +- 7 files changed, 6 insertions(+), 110 deletions(-) delete mode 100644 coq-htt-examples.opam delete mode 100644 templates-extra/coq-htt-examples.opam.mustache diff --git a/coq-htt-examples.opam b/coq-htt-examples.opam deleted file mode 100644 index dd2c9e9..0000000 --- a/coq-htt-examples.opam +++ /dev/null @@ -1,53 +0,0 @@ -# This file was generated from `meta.yml`, please do not edit manually. - -opam-version: "2.0" -maintainer: "fcsl@software.imdea.org" -version: "dev" - -homepage: "https://github.com/imdea-software/htt" -dev-repo: "git+https://github.com/imdea-software/htt.git" -bug-reports: "https://github.com/imdea-software/htt/issues" -license: "Apache-2.0" - -synopsis: "Hoare Type Theory" -description: """ -Hoare Type Theory (HTT) is a verification system for reasoning about sequential heap-manipulating -programs based on Separation logic. - -HTT incorporates Hoare-style specifications via preconditions and postconditions into types. A -Hoare type `ST P (fun x : A => Q)` denotes computations with a precondition `P` and postcondition -`Q`, returning a value `x` of type `A`. Hoare types are a dependently typed version of monads, -as used in the programming language Haskell. Monads hygienically combine the language features -for pure functional programming, with those for imperative programming, such as state or -exceptions. In this sense, HTT establishes a formal connection in the style of Curry-Howard -isomorphism between monads and (functional programming variant of) Separation logic. Every -effectful command in HTT has a type that corresponds to the appropriate non-structural inference -rule in Separation logic, and vice versa, every non-structural inference rule corresponds to a -command in HTT that has that rule as the type. The type for monadic bind is the Hoare rule for -sequential composition, and the type for monadic unit combines the Hoare rules for the idle -program (in a small-footprint variant) and for variable assignment (adapted for functional -variables). The connection reconciles dependent types with effects of state and exceptions and -establishes Separation logic as a type theory for such effects. In implementation terms, it means -that HTT implements Separation logic as a shallow embedding in Coq.""" - -build: [make "-j%{jobs}%"] -install: [make "install"] -depends: [ - "coq" { (>= "8.19" & < "8.21~") | (= "dev") } - "coq-mathcomp-ssreflect" { (>= "2.2.0" & < "2.3~") | (= "dev") } - "coq-mathcomp-algebra" - "coq-fcsl-pcm" { (>= "2.0.0" & < "2.1~") | (= "dev") } -] - -tags: [ - "category:Computer Science/Data Types and Data Structures" - "keyword:partial commutative monoids" - "keyword:separation logic" - "logpath:htt" -] -authors: [ - "Aleksandar Nanevski" - "Germán Andrés Delbianco" - "Alexander Gryzlov" - "Marcos Grandury" -] diff --git a/coq-htt.opam b/coq-htt.opam index 95ee94a..dd2c9e9 100644 --- a/coq-htt.opam +++ b/coq-htt.opam @@ -1,5 +1,4 @@ # This file was generated from `meta.yml`, please do not edit manually. -# Follow the instructions on https://github.com/coq-community/templates to regenerate. opam-version: "2.0" maintainer: "fcsl@software.imdea.org" @@ -31,9 +30,9 @@ variables). The connection reconciles dependent types with effects of state and establishes Separation logic as a type theory for such effects. In implementation terms, it means that HTT implements Separation logic as a shallow embedding in Coq.""" -build: ["dune" "build" "-p" name "-j" jobs] +build: [make "-j%{jobs}%"] +install: [make "install"] depends: [ - "dune" {>= "3.6"} "coq" { (>= "8.19" & < "8.21~") | (= "dev") } "coq-mathcomp-ssreflect" { (>= "2.2.0" & < "2.3~") | (= "dev") } "coq-mathcomp-algebra" diff --git a/htt/dune b/htt/dune index 501fc38..a6eeaa3 100644 --- a/htt/dune +++ b/htt/dune @@ -2,7 +2,7 @@ (coq.theory (name htt) - (package coq-htt) + (package coq-htt-core) (synopsis "Hoare Type Theory") (flags :standard -w -notation-overridden diff --git a/meta.yml b/meta.yml index 04c4461..9459fe1 100644 --- a/meta.yml +++ b/meta.yml @@ -1,7 +1,7 @@ fullname: Hoare Type Theory shortname: htt organization: imdea-software -opam_name: coq-htt +opam_name: coq-htt-core community: false action: true dune: true diff --git a/regenerate.sh b/regenerate.sh index c6520b8..47ca0d5 100755 --- a/regenerate.sh +++ b/regenerate.sh @@ -1,7 +1,7 @@ #!/usr/bin/env bash TMP=$(mktemp -d); git clone https://github.com/coq-community/templates.git $TMP -$TMP/generate.sh README.md coq-htt.opam dune-project +$TMP/generate.sh README.md coq-htt-core.opam dune-project echo "Proceeding with customized generation..." diff --git a/templates-extra/coq-htt-examples.opam.mustache b/templates-extra/coq-htt-examples.opam.mustache deleted file mode 100644 index 2d53c31..0000000 --- a/templates-extra/coq-htt-examples.opam.mustache +++ /dev/null @@ -1,50 +0,0 @@ -# This file was generated from `meta.yml`, please do not edit manually. - -opam-version: "2.0" -maintainer: "{{& opam-file-maintainer }}{{^ opam-file-maintainer }}palmskog@gmail.com{{/ opam-file-maintainer }}" -version: "{{ opam-file-version }}{{^ opam-file-version }}dev{{/ opam-file-version }}" - -homepage: "https://github.com/{{ organization }}/{{ shortname }}" -dev-repo: "git+https://github.com/{{ organization }}/{{ shortname }}.git" -bug-reports: "https://github.com/{{ organization }}/{{ shortname }}/issues"{{# coqdoc }} -doc: "https://{{ organization }}.github.io/{{ shortname }}/"{{/ coqdoc }} -{{# license }}license: "{{ identifier }}"{{/ license }} - -synopsis: "{{& synopsis }}" -description: """ -{{& description }}""" - -build: [make "-j%{jobs}%"{{# make_target }} "{{ make_target }}"{{/ make_target }}] -{{# test_target }} -run-test: [make "-j%{jobs}%" "{{ test_target }}"] -{{/ test_target }} -install: [make "install"{{# install_flag }} "{{ install_flag }}"{{/ install_flag }}] -depends: [ -{{# supported_ocaml_versions }} - "ocaml" {{& opam }} -{{/ supported_ocaml_versions }} - "coq" {{& supported_coq_versions.opam }} -{{# dependencies }} -{{# opam }} - "{{ name }}" {{& version }} -{{/ opam }} -{{/ dependencies }} -] - -tags: [ -{{# categories }} - "category:{{ name }}" -{{/ categories }} -{{# keywords }} - "keyword:{{& name }}" -{{/ keywords }} -{{# date }} - "date:{{ date }}" -{{/ date }} - "logpath:{{ namespace }}" -] -authors: [ -{{# authors }} - "{{& name }}{{# email }} <{{& email }}>{{/ email }}" -{{/ authors }} -] diff --git a/templates-extra/dune.mustache b/templates-extra/dune.mustache index 0d8e5fa..351cc62 100644 --- a/templates-extra/dune.mustache +++ b/templates-extra/dune.mustache @@ -8,4 +8,4 @@ -w -notation-overridden -w -local-declaration -w -redundant-canonical-projection - -w -projection-no-head-constant)) \ No newline at end of file + -w -projection-no-head-constant)) From 3dfe5dc7c4dbbd6c9dc96895563ed157d2567954 Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Fri, 27 Sep 2024 16:10:06 +0200 Subject: [PATCH 87/93] playing with dune --- dune-project | 2 +- htt/dune | 2 +- 2 files changed, 2 insertions(+), 2 deletions(-) diff --git a/dune-project b/dune-project index 4d97b3b..ac59e7a 100644 --- a/dune-project +++ b/dune-project @@ -1,3 +1,3 @@ (lang dune 3.6) (using coq 0.6) -(name htt) +(name coq-htt-core) diff --git a/htt/dune b/htt/dune index a6eeaa3..501fc38 100644 --- a/htt/dune +++ b/htt/dune @@ -2,7 +2,7 @@ (coq.theory (name htt) - (package coq-htt-core) + (package coq-htt) (synopsis "Hoare Type Theory") (flags :standard -w -notation-overridden From 48f3a50d5bfab699fbcfa82bc9cb3313d1a72b2e Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Fri, 27 Sep 2024 16:10:32 +0200 Subject: [PATCH 88/93] forgot some files --- coq-htt-core.opam | 54 +++++++++++++++++++++++++++ templates-extra/coq-htt.opam.mustache | 50 +++++++++++++++++++++++++ 2 files changed, 104 insertions(+) create mode 100644 coq-htt-core.opam create mode 100644 templates-extra/coq-htt.opam.mustache diff --git a/coq-htt-core.opam b/coq-htt-core.opam new file mode 100644 index 0000000..95ee94a --- /dev/null +++ b/coq-htt-core.opam @@ -0,0 +1,54 @@ +# This file was generated from `meta.yml`, please do not edit manually. +# Follow the instructions on https://github.com/coq-community/templates to regenerate. + +opam-version: "2.0" +maintainer: "fcsl@software.imdea.org" +version: "dev" + +homepage: "https://github.com/imdea-software/htt" +dev-repo: "git+https://github.com/imdea-software/htt.git" +bug-reports: "https://github.com/imdea-software/htt/issues" +license: "Apache-2.0" + +synopsis: "Hoare Type Theory" +description: """ +Hoare Type Theory (HTT) is a verification system for reasoning about sequential heap-manipulating +programs based on Separation logic. + +HTT incorporates Hoare-style specifications via preconditions and postconditions into types. A +Hoare type `ST P (fun x : A => Q)` denotes computations with a precondition `P` and postcondition +`Q`, returning a value `x` of type `A`. Hoare types are a dependently typed version of monads, +as used in the programming language Haskell. Monads hygienically combine the language features +for pure functional programming, with those for imperative programming, such as state or +exceptions. In this sense, HTT establishes a formal connection in the style of Curry-Howard +isomorphism between monads and (functional programming variant of) Separation logic. Every +effectful command in HTT has a type that corresponds to the appropriate non-structural inference +rule in Separation logic, and vice versa, every non-structural inference rule corresponds to a +command in HTT that has that rule as the type. The type for monadic bind is the Hoare rule for +sequential composition, and the type for monadic unit combines the Hoare rules for the idle +program (in a small-footprint variant) and for variable assignment (adapted for functional +variables). The connection reconciles dependent types with effects of state and exceptions and +establishes Separation logic as a type theory for such effects. In implementation terms, it means +that HTT implements Separation logic as a shallow embedding in Coq.""" + +build: ["dune" "build" "-p" name "-j" jobs] +depends: [ + "dune" {>= "3.6"} + "coq" { (>= "8.19" & < "8.21~") | (= "dev") } + "coq-mathcomp-ssreflect" { (>= "2.2.0" & < "2.3~") | (= "dev") } + "coq-mathcomp-algebra" + "coq-fcsl-pcm" { (>= "2.0.0" & < "2.1~") | (= "dev") } +] + +tags: [ + "category:Computer Science/Data Types and Data Structures" + "keyword:partial commutative monoids" + "keyword:separation logic" + "logpath:htt" +] +authors: [ + "Aleksandar Nanevski" + "Germán Andrés Delbianco" + "Alexander Gryzlov" + "Marcos Grandury" +] diff --git a/templates-extra/coq-htt.opam.mustache b/templates-extra/coq-htt.opam.mustache new file mode 100644 index 0000000..2d53c31 --- /dev/null +++ b/templates-extra/coq-htt.opam.mustache @@ -0,0 +1,50 @@ +# This file was generated from `meta.yml`, please do not edit manually. + +opam-version: "2.0" +maintainer: "{{& opam-file-maintainer }}{{^ opam-file-maintainer }}palmskog@gmail.com{{/ opam-file-maintainer }}" +version: "{{ opam-file-version }}{{^ opam-file-version }}dev{{/ opam-file-version }}" + +homepage: "https://github.com/{{ organization }}/{{ shortname }}" +dev-repo: "git+https://github.com/{{ organization }}/{{ shortname }}.git" +bug-reports: "https://github.com/{{ organization }}/{{ shortname }}/issues"{{# coqdoc }} +doc: "https://{{ organization }}.github.io/{{ shortname }}/"{{/ coqdoc }} +{{# license }}license: "{{ identifier }}"{{/ license }} + +synopsis: "{{& synopsis }}" +description: """ +{{& description }}""" + +build: [make "-j%{jobs}%"{{# make_target }} "{{ make_target }}"{{/ make_target }}] +{{# test_target }} +run-test: [make "-j%{jobs}%" "{{ test_target }}"] +{{/ test_target }} +install: [make "install"{{# install_flag }} "{{ install_flag }}"{{/ install_flag }}] +depends: [ +{{# supported_ocaml_versions }} + "ocaml" {{& opam }} +{{/ supported_ocaml_versions }} + "coq" {{& supported_coq_versions.opam }} +{{# dependencies }} +{{# opam }} + "{{ name }}" {{& version }} +{{/ opam }} +{{/ dependencies }} +] + +tags: [ +{{# categories }} + "category:{{ name }}" +{{/ categories }} +{{# keywords }} + "keyword:{{& name }}" +{{/ keywords }} +{{# date }} + "date:{{ date }}" +{{/ date }} + "logpath:{{ namespace }}" +] +authors: [ +{{# authors }} + "{{& name }}{{# email }} <{{& email }}>{{/ email }}" +{{/ authors }} +] From e0539feac94ef7be63cbc35067dcf2d01cf4bdff Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Fri, 27 Sep 2024 16:31:29 +0200 Subject: [PATCH 89/93] blah --- dune-project | 2 +- htt/dune | 2 +- 2 files changed, 2 insertions(+), 2 deletions(-) diff --git a/dune-project b/dune-project index ac59e7a..4d97b3b 100644 --- a/dune-project +++ b/dune-project @@ -1,3 +1,3 @@ (lang dune 3.6) (using coq 0.6) -(name coq-htt-core) +(name htt) diff --git a/htt/dune b/htt/dune index 501fc38..a6eeaa3 100644 --- a/htt/dune +++ b/htt/dune @@ -2,7 +2,7 @@ (coq.theory (name htt) - (package coq-htt) + (package coq-htt-core) (synopsis "Hoare Type Theory") (flags :standard -w -notation-overridden From 83baf96c7341f8aec6fa25ab8ae565ca33619ced Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Fri, 27 Sep 2024 16:35:11 +0200 Subject: [PATCH 90/93] blah --- templates-extra/coq-htt.opam.mustache | 2 -- 1 file changed, 2 deletions(-) diff --git a/templates-extra/coq-htt.opam.mustache b/templates-extra/coq-htt.opam.mustache index 2d53c31..e53aaba 100644 --- a/templates-extra/coq-htt.opam.mustache +++ b/templates-extra/coq-htt.opam.mustache @@ -1,5 +1,3 @@ -# This file was generated from `meta.yml`, please do not edit manually. - opam-version: "2.0" maintainer: "{{& opam-file-maintainer }}{{^ opam-file-maintainer }}palmskog@gmail.com{{/ opam-file-maintainer }}" version: "{{ opam-file-version }}{{^ opam-file-version }}dev{{/ opam-file-version }}" From 9a80b1516aec94b1011e350ade70a3ac3d21bea8 Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Fri, 27 Sep 2024 16:36:00 +0200 Subject: [PATCH 91/93] blah --- templates-extra/coq-htt.opam.mustache | 1 + 1 file changed, 1 insertion(+) diff --git a/templates-extra/coq-htt.opam.mustache b/templates-extra/coq-htt.opam.mustache index e53aaba..f03cfbf 100644 --- a/templates-extra/coq-htt.opam.mustache +++ b/templates-extra/coq-htt.opam.mustache @@ -1,3 +1,4 @@ +# This file was generated from `meta.yml`, please do not edit manually. opam-version: "2.0" maintainer: "{{& opam-file-maintainer }}{{^ opam-file-maintainer }}palmskog@gmail.com{{/ opam-file-maintainer }}" version: "{{ opam-file-version }}{{^ opam-file-version }}dev{{/ opam-file-version }}" From d725c09fc58ad0ea4c35c8b46a957998f98ce9fe Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Fri, 27 Sep 2024 16:36:36 +0200 Subject: [PATCH 92/93] blah --- .github/workflows/docker-action.yml | 2 +- coq-htt.opam | 1 - templates-extra/docker-action.yml.mustache | 2 +- 3 files changed, 2 insertions(+), 3 deletions(-) diff --git a/.github/workflows/docker-action.yml b/.github/workflows/docker-action.yml index f703150..761fe01 100644 --- a/.github/workflows/docker-action.yml +++ b/.github/workflows/docker-action.yml @@ -24,7 +24,7 @@ jobs: - uses: actions/checkout@v2 - uses: coq-community/docker-coq-action@v1 with: - opam_file: 'coq-htt-examples.opam' + opam_file: 'coq-htt.opam' custom_image: ${{ matrix.image }} diff --git a/coq-htt.opam b/coq-htt.opam index dd2c9e9..7086e27 100644 --- a/coq-htt.opam +++ b/coq-htt.opam @@ -1,5 +1,4 @@ # This file was generated from `meta.yml`, please do not edit manually. - opam-version: "2.0" maintainer: "fcsl@software.imdea.org" version: "dev" diff --git a/templates-extra/docker-action.yml.mustache b/templates-extra/docker-action.yml.mustache index 06e35a6..ca8b1a1 100644 --- a/templates-extra/docker-action.yml.mustache +++ b/templates-extra/docker-action.yml.mustache @@ -32,7 +32,7 @@ jobs: {{/ submodule }} - uses: coq-community/docker-coq-action@v1 with: - opam_file: 'coq-htt-examples.opam' + opam_file: 'coq-htt.opam' {{! Change delimiters to avoid the next line being parsed as mustache syntax. }} {{=<% %>=}} custom_image: ${{ matrix.image }} From 90688078144cf8c2bd94df4c4888635a36b60619 Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Fri, 27 Sep 2024 16:50:16 +0200 Subject: [PATCH 93/93] changed mustache files --- templates-extra/coq-htt.opam.mustache | 3 ++- templates-extra/docker-action.yml.mustache | 3 ++- 2 files changed, 4 insertions(+), 2 deletions(-) diff --git a/templates-extra/coq-htt.opam.mustache b/templates-extra/coq-htt.opam.mustache index f03cfbf..91e12ca 100644 --- a/templates-extra/coq-htt.opam.mustache +++ b/templates-extra/coq-htt.opam.mustache @@ -1,4 +1,5 @@ -# This file was generated from `meta.yml`, please do not edit manually. +; This file was generated from `meta.yml`, please do not edit manually. + opam-version: "2.0" maintainer: "{{& opam-file-maintainer }}{{^ opam-file-maintainer }}palmskog@gmail.com{{/ opam-file-maintainer }}" version: "{{ opam-file-version }}{{^ opam-file-version }}dev{{/ opam-file-version }}" diff --git a/templates-extra/docker-action.yml.mustache b/templates-extra/docker-action.yml.mustache index ca8b1a1..89cf25b 100644 --- a/templates-extra/docker-action.yml.mustache +++ b/templates-extra/docker-action.yml.mustache @@ -1,4 +1,5 @@ -# This file was generated from `meta.yml`, please do not edit manually. +; This file was generated from `meta.yml`, please do not edit manually. + name: Docker CI on: