From 497a4c38cded3cf028c9d286297a56936d15ddd2 Mon Sep 17 00:00:00 2001 From: Anton Trunov Date: Fri, 20 Apr 2018 16:08:54 +0200 Subject: [PATCH] Initial commit --- .gitignore | 7 + CREDITS | 14 + CoqMakefile.local | 2 + LICENSE | 202 ++++ Make | 15 + Makefile | 19 + README.md | 33 + _CoqProject | 2 + descr | 11 + finmap/finmap.v | 1281 +++++++++++++++++++++++++ finmap/ordtype.v | 331 +++++++ opam | 24 + pcm/automap.v | 834 ++++++++++++++++ pcm/axioms.v | 124 +++ pcm/heap.v | 764 +++++++++++++++ pcm/lift.v | 241 +++++ pcm/mutex.v | 304 ++++++ pcm/natmap.v | 607 ++++++++++++ pcm/pcm.v | 640 +++++++++++++ pcm/pred.v | 562 +++++++++++ pcm/prelude.v | 247 +++++ pcm/unionmap.v | 2303 +++++++++++++++++++++++++++++++++++++++++++++ 22 files changed, 8567 insertions(+) create mode 100644 .gitignore create mode 100644 CREDITS create mode 100644 CoqMakefile.local create mode 100644 LICENSE create mode 100644 Make create mode 100644 Makefile create mode 100644 README.md create mode 100644 _CoqProject create mode 100644 descr create mode 100644 finmap/finmap.v create mode 100644 finmap/ordtype.v create mode 100644 opam create mode 100644 pcm/automap.v create mode 100644 pcm/axioms.v create mode 100644 pcm/heap.v create mode 100644 pcm/lift.v create mode 100644 pcm/mutex.v create mode 100644 pcm/natmap.v create mode 100644 pcm/pcm.v create mode 100644 pcm/pred.v create mode 100644 pcm/prelude.v create mode 100644 pcm/unionmap.v diff --git a/.gitignore b/.gitignore new file mode 100644 index 0000000..9640d6f --- /dev/null +++ b/.gitignore @@ -0,0 +1,7 @@ +*.vo +*.glob +*.aux +*.v.d +.coqdeps.d +CoqMakefile +CoqMakefile.conf diff --git a/CREDITS b/CREDITS new file mode 100644 index 0000000..4968d19 --- /dev/null +++ b/CREDITS @@ -0,0 +1,14 @@ +GENERAL INFORMATION +------------------- + +This code has been developed as an IMDEA Software Institute project. + + +CONTRIBUTORS +------------ + +Aleks Nanevski (2009-now) +Ruy Ley-Wild (2010-2012) +Ilya Sergey (2013-2015) +German Delbianco (2014-2017) +Anton Trunov (2017-now) diff --git a/CoqMakefile.local b/CoqMakefile.local new file mode 100644 index 0000000..9ac1e9d --- /dev/null +++ b/CoqMakefile.local @@ -0,0 +1,2 @@ +clean:: + $(HIDE)rm -f CoqMakefile CoqMakefile.conf diff --git a/LICENSE b/LICENSE new file mode 100644 index 0000000..d645695 --- /dev/null +++ b/LICENSE @@ -0,0 +1,202 @@ + + Apache License + Version 2.0, January 2004 + http://www.apache.org/licenses/ + + TERMS AND CONDITIONS FOR USE, REPRODUCTION, AND DISTRIBUTION + + 1. Definitions. + + "License" shall mean the terms and conditions for use, reproduction, + and distribution as defined by Sections 1 through 9 of this document. + + "Licensor" shall mean the copyright owner or entity authorized by + the copyright owner that is granting the License. + + "Legal Entity" shall mean the union of the acting entity and all + other entities that control, are controlled by, or are under common + control with that entity. For the purposes of this definition, + "control" means (i) the power, direct or indirect, to cause the + direction or management of such entity, whether by contract or + otherwise, or (ii) ownership of fifty percent (50%) or more of the + outstanding shares, or (iii) beneficial ownership of such entity. + + "You" (or "Your") shall mean an individual or Legal Entity + exercising permissions granted by this License. + + "Source" form shall mean the preferred form for making modifications, + including but not limited to software source code, documentation + source, and configuration files. + + "Object" form shall mean any form resulting from mechanical + transformation or translation of a Source form, including but + not limited to compiled object code, generated documentation, + and conversions to other media types. + + "Work" shall mean the work of authorship, whether in Source or + Object form, made available under the License, as indicated by a + copyright notice that is included in or attached to the work + (an example is provided in the Appendix below). + + "Derivative Works" shall mean any work, whether in Source or Object + form, that is based on (or derived from) the Work and for which the + editorial revisions, annotations, elaborations, or other modifications + represent, as a whole, an original work of authorship. For the purposes + of this License, Derivative Works shall not include works that remain + separable from, or merely link (or bind by name) to the interfaces of, + the Work and Derivative Works thereof. + + "Contribution" shall mean any work of authorship, including + the original version of the Work and any modifications or additions + to that Work or Derivative Works thereof, that is intentionally + submitted to Licensor for inclusion in the Work by the copyright owner + or by an individual or Legal Entity authorized to submit on behalf of + the copyright owner. For the purposes of this definition, "submitted" + means any form of electronic, verbal, or written communication sent + to the Licensor or its representatives, including but not limited to + communication on electronic mailing lists, source code control systems, + and issue tracking systems that are managed by, or on behalf of, the + Licensor for the purpose of discussing and improving the Work, but + excluding communication that is conspicuously marked or otherwise + designated in writing by the copyright owner as "Not a Contribution." + + "Contributor" shall mean Licensor and any individual or Legal Entity + on behalf of whom a Contribution has been received by Licensor and + subsequently incorporated within the Work. + + 2. Grant of Copyright License. Subject to the terms and conditions of + this License, each Contributor hereby grants to You a perpetual, + worldwide, non-exclusive, no-charge, royalty-free, irrevocable + copyright license to reproduce, prepare Derivative Works of, + publicly display, publicly perform, sublicense, and distribute the + Work and such Derivative Works in Source or Object form. + + 3. Grant of Patent License. Subject to the terms and conditions of + this License, each Contributor hereby grants to You a perpetual, + worldwide, non-exclusive, no-charge, royalty-free, irrevocable + (except as stated in this section) patent license to make, have made, + use, offer to sell, sell, import, and otherwise transfer the Work, + where such license applies only to those patent claims licensable + by such Contributor that are necessarily infringed by their + Contribution(s) alone or by combination of their Contribution(s) + with the Work to which such Contribution(s) was submitted. If You + institute patent litigation against any entity (including a + cross-claim or counterclaim in a lawsuit) alleging that the Work + or a Contribution incorporated within the Work constitutes direct + or contributory patent infringement, then any patent licenses + granted to You under this License for that Work shall terminate + as of the date such litigation is filed. + + 4. Redistribution. You may reproduce and distribute copies of the + Work or Derivative Works thereof in any medium, with or without + modifications, and in Source or Object form, provided that You + meet the following conditions: + + (a) You must give any other recipients of the Work or + Derivative Works a copy of this License; and + + (b) You must cause any modified files to carry prominent notices + stating that You changed the files; and + + (c) You must retain, in the Source form of any Derivative Works + that You distribute, all copyright, patent, trademark, and + attribution notices from the Source form of the Work, + excluding those notices that do not pertain to any part of + the Derivative Works; and + + (d) If the Work includes a "NOTICE" text file as part of its + distribution, then any Derivative Works that You distribute must + include a readable copy of the attribution notices contained + within such NOTICE file, excluding those notices that do not + pertain to any part of the Derivative Works, in at least one + of the following places: within a NOTICE text file distributed + as part of the Derivative Works; within the Source form or + documentation, if provided along with the Derivative Works; or, + within a display generated by the Derivative Works, if and + wherever such third-party notices normally appear. The contents + of the NOTICE file are for informational purposes only and + do not modify the License. You may add Your own attribution + notices within Derivative Works that You distribute, alongside + or as an addendum to the NOTICE text from the Work, provided + that such additional attribution notices cannot be construed + as modifying the License. + + You may add Your own copyright statement to Your modifications and + may provide additional or different license terms and conditions + for use, reproduction, or distribution of Your modifications, or + for any such Derivative Works as a whole, provided Your use, + reproduction, and distribution of the Work otherwise complies with + the conditions stated in this License. + + 5. Submission of Contributions. Unless You explicitly state otherwise, + any Contribution intentionally submitted for inclusion in the Work + by You to the Licensor shall be under the terms and conditions of + this License, without any additional terms or conditions. + Notwithstanding the above, nothing herein shall supersede or modify + the terms of any separate license agreement you may have executed + with Licensor regarding such Contributions. + + 6. Trademarks. This License does not grant permission to use the trade + names, trademarks, service marks, or product names of the Licensor, + except as required for reasonable and customary use in describing the + origin of the Work and reproducing the content of the NOTICE file. + + 7. Disclaimer of Warranty. Unless required by applicable law or + agreed to in writing, Licensor provides the Work (and each + Contributor provides its Contributions) on an "AS IS" BASIS, + WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or + implied, including, without limitation, any warranties or conditions + of TITLE, NON-INFRINGEMENT, MERCHANTABILITY, or FITNESS FOR A + PARTICULAR PURPOSE. You are solely responsible for determining the + appropriateness of using or redistributing the Work and assume any + risks associated with Your exercise of permissions under this License. + + 8. Limitation of Liability. In no event and under no legal theory, + whether in tort (including negligence), contract, or otherwise, + unless required by applicable law (such as deliberate and grossly + negligent acts) or agreed to in writing, shall any Contributor be + liable to You for damages, including any direct, indirect, special, + incidental, or consequential damages of any character arising as a + result of this License or out of the use or inability to use the + Work (including but not limited to damages for loss of goodwill, + work stoppage, computer failure or malfunction, or any and all + other commercial damages or losses), even if such Contributor + has been advised of the possibility of such damages. + + 9. Accepting Warranty or Additional Liability. While redistributing + the Work or Derivative Works thereof, You may choose to offer, + and charge a fee for, acceptance of support, warranty, indemnity, + or other liability obligations and/or rights consistent with this + License. However, in accepting such obligations, You may act only + on Your own behalf and on Your sole responsibility, not on behalf + of any other Contributor, and only if You agree to indemnify, + defend, and hold each Contributor harmless for any liability + incurred by, or claims asserted against, such Contributor by reason + of your accepting any such warranty or additional liability. + + END OF TERMS AND CONDITIONS + + APPENDIX: How to apply the Apache License to your work. + + To apply the Apache License to your work, attach the following + boilerplate notice, with the fields enclosed by brackets "[]" + replaced with your own identifying information. (Don't include + the brackets!) The text should be enclosed in the appropriate + comment syntax for the file format. We also recommend that a + file or class name and description of purpose be included on the + same "printed page" as the copyright notice for easier + identification within third-party archives. + + Copyright [yyyy] [name of copyright owner] + + 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. diff --git a/Make b/Make new file mode 100644 index 0000000..53ec127 --- /dev/null +++ b/Make @@ -0,0 +1,15 @@ +finmap/finmap.v +finmap/ordtype.v +pcm/automap.v +pcm/axioms.v +pcm/heap.v +pcm/mutex.v +pcm/lift.v +pcm/natmap.v +pcm/pcm.v +pcm/pred.v +pcm/prelude.v +pcm/unionmap.v + +-R . fcsl +-arg "-w -notation-overridden,-redundant-canonical-projection" diff --git a/Makefile b/Makefile new file mode 100644 index 0000000..a8cb8cc --- /dev/null +++ b/Makefile @@ -0,0 +1,19 @@ +# KNOWNTARGETS will not be passed along to CoqMakefile +KNOWNTARGETS := CoqMakefile +# KNOWNFILES will not get implicit targets from the final rule, and so depending on them won’t invoke the submake +# Warning: These files get declared as PHONY, so any targets depending on them always get rebuilt +KNOWNFILES := Makefile Make + +.DEFAULT_GOAL := invoke-coqmakefile + +CoqMakefile: Makefile Make + $(COQBIN)coq_makefile -f Make -o CoqMakefile + +invoke-coqmakefile: CoqMakefile + $(MAKE) --no-print-directory -f CoqMakefile $(filter-out $(KNOWNTARGETS),$(MAKECMDGOALS)) + +.PHONY: invoke-coqmakefile $(KNOWNFILES) + +# This should be the last rule, to handle any targets not declared above +%: invoke-coqmakefile + @true diff --git a/README.md b/README.md new file mode 100644 index 0000000..208f375 --- /dev/null +++ b/README.md @@ -0,0 +1,33 @@ +# The PCM library + +The PCM library provides a formalisation of Partial Commutative Monoids (PCMs), +a common algebraic structure used in separation logic for verification of +pointer-manipulating sequential and concurrent programs. + +The library provides lemmas for mechanised and automated reasoning about PCMs +in the abstract, but also supports concrete common PCM instances, such as heaps, +histories and mutexes. + +It is based on the [Coq](http://coq.inria.fr) proof assistant, +[SSReflect](https://coq.inria.fr/distrib/current/refman/ssreflect.html) +proof language, and [Mathcomp](https://github.com/math-comp/math-comp) library. + +## Installation + +The PCM library can be installed via [OPAM](https://opam.ocaml.org) package manager: + +``` +opam repo add coq-released https://coq.inria.fr/opam/released +opam install coq-fcsl-pcm +``` + +## Getting help + +If you need assistance or would like to report a bug, drop us an email: + +or open and [issue](https://github.com/imdea-software/fcsl-pcm/issues). + +## More Information + +More information can be obtained via the [FCSL web page](http://software.imdea.org/fcsl/). +An earlier version of this library was developed as part of [Hoare type theory](http://software.imdea.org/~aleks/htt/). diff --git a/_CoqProject b/_CoqProject new file mode 100644 index 0000000..9f2d724 --- /dev/null +++ b/_CoqProject @@ -0,0 +1,2 @@ +-R . fcsl +-arg "-w -notation-overridden,-projection-no-head-constant,-redundant-canonical-projection" diff --git a/descr b/descr new file mode 100644 index 0000000..cb1d34b --- /dev/null +++ b/descr @@ -0,0 +1,11 @@ +Partial Commutative Monoids + +The PCM library provides a formalisation of Partial Commutative Monoids (PCMs), +a common algebraic structure used in separation logic for verification of +pointer-manipulating sequential and concurrent programs. +The library provides lemmas for mechanised and automated reasoning about PCMs +in the abstract, but also supports concrete common PCM instances, such as heaps, +histories and mutexes. + +This library relies on extensionality axioms: propositional and +functional extentionality. diff --git a/finmap/finmap.v b/finmap/finmap.v new file mode 100644 index 0000000..ddcd1ce --- /dev/null +++ b/finmap/finmap.v @@ -0,0 +1,1281 @@ +(* +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 Coq Require Import ssreflect ssrbool ssrfun. +From mathcomp Require Import ssrnat eqtype seq path. +From fcsl Require Import ordtype. +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +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 (bool_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: (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: (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'). + +Definition supp s := map key (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). + +Lemma ord_path (x y : K) s : ord x y -> path ord y s -> path ord x s. +Proof. +elim: s x y=>[|k s IH] x y //=. +by move=>H1; case/andP=>H2 ->; rewrite (trans H1 H2). +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: totalP=>// 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: totalP 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 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} /= 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: (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. + +CoInductive 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: (totalP k1 k2)=>/= O H1 H2. +- move/(_ k1); rewrite !inE eq_refl /= eq_sym. + case: totalP 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: totalP 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: (totalP k2 k3)=>H1 /=; rewrite ?irr ?eq_refl //. + case: (totalP k2 k3) H1=>H2 _ //. + by rewrite IH //; apply: path_sorted H. +elim: s H k1 k2 H1 v1 v2=>[|[k3 v3] s IH] H k1 k2 H1 v1 v2 /=. +- rewrite H1 eq_sym H1. + by case: (totalP k1 k2) H1=>H2 H1. +case: (totalP k2 k3)=>H2 /=. +- case: (totalP k1 k2) (H1)=>H3 _ //=; last first. + - by case: (totalP k1 k3)=>//= H4; rewrite ?H2 ?H3. + case: (totalP k1 k3)=>H4 /=. + - case: (totalP k2 k1) H3=>//= H3. + by case: (totalP k2 k3) H2=>//=. + - rewrite (eqP H4) in H3. + by case: (totalP k2 k3) H2 H3. + by case: (totalP k1 k3) (trans H3 H2) H4. +- rewrite -(eqP H2) {H2} (H1). + case: (totalP k1 k2) (H1)=>//= H2 _; rewrite ?irr ?eq_refl //. + rewrite eq_sym H1. + by case: (totalP k2 k1) H1 H2. +case: (totalP k1 k3)=>H3 /=. +- rewrite eq_sym H1. + case: (totalP k2 k1) H1 (trans H3 H2)=>//. + by case: (totalP k2 k3) H2=>//=. +- rewrite (eqP H3). + by case: (totalP k2 k3) H2. +case: (totalP 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: (totalP k3 k2)=>H1 /=; rewrite ?eq_refl //=; + case: (totalP k3 k2) H1=>//= H1 _. + by rewrite IH //; apply: path_sorted H. +elim: s k1 k2 H1 H=>[|[k3 v3] s IH] k1 k2 H1 /= H; first by rewrite eq_sym H1. +case: (totalP k2 k3)=>H2 /=. +- rewrite eq_sym H1 /=. + case: (totalP k3 k1)=>H3 /=; case: (totalP k2 k3) (H2)=>//=. + rewrite -(eqP H3) in H1 *. + rewrite -IH //; last by apply: path_sorted H. + rewrite last_ins' /= 1?eq_sym ?H1 //. + by apply: ord_path H. +- by move: H1; rewrite (eqP H2) /= eq_sym => -> /=; rewrite irr eq_refl. +case: (totalP k3 k1)=>H3 /=. +- case: (totalP k2 k3) H2=>//= H2 _. + by rewrite IH //; apply: path_sorted H. +- rewrite -(eqP H3) in H1 *. + by rewrite IH //; apply: path_sorted H. +case: (totalP k2 k3) H2=>//= H2 _. +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 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: (totalP 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: (totalP k2 k1)=>//= H2. +- have H: path ord k2 (supp (ins k1 v1 s1)). + - by apply: (path_supp_ins _ H2); apply: ord_path H1. + by rewrite {1}/fcat seqof_ins //= fcat_ins' ?notin_path. +- by rewrite IH ins_ins H2 IH ins_ins H2. +have H: path ord k1 (supp (ins k2 v2 s1)) by apply: (path_supp_ins _ H2). +rewrite ins_ins. +case: (totalP k2 k1) H2 => // H2 _. +rewrite {1}/fcat seqof_ins //= fcat_ins' ?notin_path // IH ?notin_path //. +rewrite ins_ins; case: (totalP 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: totalP=>/= 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 {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 {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). + +CoInductive 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. +move/(nth_find k); move: H; rewrite has_find=>/(mem_nth k). +by 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. + +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. + +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. + +Canonical Structure fmap_eqMixin := EqMixin feqP. +Canonical Structure fmap_eqType := EqType (finMap K V) fmap_eqMixin. +End EqType. + +(* mapping a function over a contents of a finite map *) + +Section Map. +Variables (K : ordType) (U V : Type) (f : U -> V). + +Definition mapf' (m : seq (K * U)) : seq (K * V) := + map (fun kv => (key kv, f (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. +move=>y; rewrite map_key_mapf. +by apply/allP; apply: order_path_min H; apply: trans. +Qed. + +Definition mapf (m : finMap K U) : finMap K V := + let: FinMap _ pf := m in FinMap (sorted_map pf). + +Lemma mapf_ins k v s : mapf (ins k v s) = ins k (f v) (mapf s). +Proof. +case: s=>s H; apply/fmapE=>/=. +elim: s k v H=>[|[k1 v1] s IH] //= k v H. +rewrite eq_sym; case: totalP=>O //=. +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 map_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 map_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 : V -> V). +Variable (comm : commutative zip_f). +Variable (assoc : forall x y z, + obind (zip_f x) (zip_f y z) = obind (zip_f^~ z) (zip_f x y)). +Variable (unitL : forall x, zip_f (unit_f x) x = Some x). +Variable (unitE : forall x y, + (exists z, zip_f x y = Some z) <-> unit_f x = unit_f y). + +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, exists 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=>[_ [<-]|_]; first by exists v1, 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 v) (zip_unit f). +Proof. +case: f=>s H; apply/fmapE=>/=; rewrite /zip_unit'. +elim: s k v H=>[|[k1 v1] s IH] //= k v H. +rewrite eq_sym; case: totalP=>//= O. +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/finmap/ordtype.v b/finmap/ordtype.v new file mode 100644 index 0000000..f2f1327 --- /dev/null +++ b/finmap/ordtype.v @@ -0,0 +1,331 @@ +(* +Copyright 2010 IMDEA Software Institute +Licensed under the Apache License, Version 2.0 (the "License"); +you may not use this file except in compliance with the License. +You may obtain a copy of the License at + http://www.apache.org/licenses/LICENSE-2.0 +Unless required by applicable law or agreed to in writing, software +distributed under the License is distributed on an "AS IS" BASIS, +WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. +See the License for the specific language governing permissions and +limitations under the License. +*) + +(******************************************************************************) +(* This file defines ordType - the structure for types with a decidable *) +(* (strict) order relation. *) +(* ordType is a subclass of mathcomp's eqType *) +(* This file also defines some important instances of ordType *) +(******************************************************************************) + +From Coq Require Import ssreflect ssrbool ssrfun. +From mathcomp Require Import ssrnat eqtype seq path fintype. +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Module Ordered. + +Section RawMixin. + +Structure mixin_of (T : eqType) := + Mixin {ordering : rel T; + _ : irreflexive ordering; + _ : transitive ordering; + _ : forall x y, [|| ordering x y, x == y | ordering y x]}. + +End RawMixin. + +(* the class takes a naked type T and returns all the *) +(* relatex mixins; the inherited ones and the added ones *) +Section ClassDef. + +Record class_of (T : Type) := Class { + base : Equality.class_of T; + mixin : mixin_of (Equality.Pack base T)}. + +Local Coercion base : class_of >-> Equality.class_of. + +Structure type : Type := Pack {sort : Type; _ : class_of sort; _ : Type}. +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 T. + +(* produce an ordered type out of the inherited mixins *) +(* equalize m0 and m by means of a phantom; will be exploited *) +(* further down in the definition of OrdType *) +Definition pack b (m0 : mixin_of (EqType T b)) := + fun m & phant_id m0 m => Pack (@Class T b m) T. + +Definition eqType := Equality.Pack class cT. + +End ClassDef. + +Module Exports. +Coercion sort : type >-> Sortclass. +Coercion eqType : type >-> Equality.type. +Canonical Structure eqType. +Notation ordType := Ordered.type. +Notation OrdMixin := Mixin. +Notation OrdType T m := (@pack T _ m _ id). +Definition ord T : rel (sort T) := (ordering (mixin (class T))). +Notation "[ 'ordType' 'of' T 'for' cT ]" := (@clone T cT _ id) + (at level 0, format "[ 'ordType' 'of' T 'for' cT ]") : form_scope. +Notation "[ 'ordType' 'of' T ]" := (@clone T _ _ id) + (at level 0, format "[ 'ordType' 'of' T ]") : form_scope. +End Exports. +End Ordered. +Export Ordered.Exports. + +Definition oleq (T : ordType) (t1 t2 : T) := ord t1 t2 || (t1 == t2). + +Prenex Implicits ord oleq. + +Section Lemmas. +Variable T : ordType. + +Lemma irr : irreflexive (@ord T). +Proof. by case: T=>s [b [m]]. Qed. + +Lemma trans : transitive (@ord T). +Proof. by case: T=>s [b [m]]. Qed. + +Lemma total (x y : T) : [|| ord x y, x == y | ord y x]. +Proof. by case: T x y=>s [b [m]]. Qed. + +Lemma nsym (x y : T) : ord x y -> ord y x -> False. +Proof. by move=>E1 E2; move: (trans E1 E2); rewrite irr. Qed. + +Lemma otrans : transitive (@oleq T). +Proof. +move=>x y z /=; case/orP; last by move/eqP=>->. +rewrite /oleq; move=>T1; case/orP; first by move/(trans T1)=>->. +by move/eqP=><-; rewrite T1. +Qed. + +Lemma sorted_oleq s : sorted (@ord T) s -> sorted (@oleq T) s. +Proof. by elim: s=>[|x s IH] //=; apply: sub_path=>z y; rewrite /oleq=>->. Qed. + +End Lemmas. + +Section Totality. +Variable K : ordType. + +CoInductive total_spec (x y : K) : bool -> bool -> bool -> Type := +| total_spec_lt of ord x y : total_spec x y true false false +| total_spec_eq of x == y : total_spec x y false true false +| total_spec_gt of ord y x : total_spec x y false false true. + +Lemma totalP x y : total_spec x y (ord x y) (x == y) (ord y x). +Proof. +case H1: (x == y). +- by rewrite (eqP H1) irr; apply: total_spec_eq. +case H2: (ord x y); case H3: (ord y x). +- by case: (nsym H2 H3). +- by apply: total_spec_lt H2. +- by apply: total_spec_gt H3. +by move: (total x y); rewrite H1 H2 H3. +Qed. +End Totality. + + +(* Monotone (i.e. strictly increasing) functions for Ord Types *) +Section Mono. +Variables (A B :ordType). + +Definition strictly_increasing f x y := @ord A x y -> @ord B (f x) (f y). + +Structure mono : Type := Mono + {fun_of: A -> B; _: forall x y, strictly_increasing fun_of x y}. + +End Mono. +Arguments strictly_increasing {A B} f x y. +Arguments Mono {A B _} _. + +Section NatOrd. +Lemma irr_ltn_nat : irreflexive ltn. Proof. by move=>x; rewrite /= ltnn. Qed. +Lemma trans_ltn_nat : transitive ltn. Proof. by apply: ltn_trans. Qed. +Lemma total_ltn_nat x y : [|| x < y, x == y | y < x]. +Proof. by case: ltngtP. Qed. + +Definition nat_ordMixin := OrdMixin irr_ltn_nat trans_ltn_nat total_ltn_nat. +Canonical Structure nat_ordType := OrdType nat nat_ordMixin. +End NatOrd. + +Section ProdOrd. +Variables K T : ordType. + +(* lexicographic ordering *) +Definition lex : rel (K * T) := + fun x y => if x.1 == y.1 then ord x.2 y.2 else ord x.1 y.1. + +Lemma irr_lex : irreflexive lex. +Proof. by move=>x; rewrite /lex eq_refl irr. Qed. + +Lemma trans_lex : transitive lex. +Proof. +move=>[x1 x2][y1 y2][z1 z2]; rewrite /lex /=. +case: ifP=>H1; first by rewrite (eqP H1); case: eqP=>// _; apply: trans. +case: ifP=>H2; first by rewrite (eqP H2) in H1 *; rewrite H1. +case: ifP=>H3; last by apply: trans. +by rewrite (eqP H3)=>R1; move/(nsym R1). +Qed. + +Lemma total_lex : forall x y, [|| lex x y, x == y | lex y x]. +Proof. +move=>[x1 x2][y1 y2]; rewrite /lex /=. +case: ifP=>H1. +- rewrite (eqP H1) eq_refl -pair_eqE /= eq_refl /=; exact: total. +rewrite (eq_sym y1) -pair_eqE /= H1 /=. +by move: (total x1 y1); rewrite H1. +Qed. + +Definition prod_ordMixin := OrdMixin irr_lex trans_lex total_lex. +Canonical Structure prod_ordType := Eval hnf in OrdType (K * T) prod_ordMixin. +End ProdOrd. + +Section FinTypeOrd. +Variable T : finType. + +Definition ordf : rel T := + fun x y => index x (enum T) < index y (enum T). + +Lemma irr_ordf : irreflexive ordf. +Proof. by move=>x; rewrite /ordf ltnn. Qed. + +Lemma trans_ordf : transitive ordf. +Proof. by move=>x y z; rewrite /ordf; apply: ltn_trans. Qed. + +Lemma total_ordf x y : [|| ordf x y, x == y | ordf y x]. +Proof. +rewrite /ordf; case: ltngtP=>//= H; rewrite ?orbT ?orbF //. +have [H1 H2]: x \in enum T /\ y \in enum T by rewrite !mem_enum. +by rewrite -(nth_index x H1) -(nth_index x H2) H eq_refl. +Qed. + +Definition fin_ordMixin := OrdMixin irr_ordf trans_ordf total_ordf. +End FinTypeOrd. + +(* notation to let us write I_n instead of (ordinal_finType n) *) +Notation "[ 'fin_ordMixin' 'of' T ]" := + (fin_ordMixin _ : Ordered.mixin_of [eqType of T]) (at level 0). + +Definition ordinal_ordMixin n := [fin_ordMixin of 'I_n]. +Canonical Structure ordinal_ordType n := OrdType 'I_n (ordinal_ordMixin n). + +Section SeqOrd. +Variable (T : ordType). + +Fixpoint ords x : pred (seq T) := + fun y => match x , y with + | [::] , [::] => false + | [::] , t :: ts => true + | x :: xs , y :: ys => if x == y then ords xs ys + else ord x y + | _ :: _ , [::] => false + end. + +Lemma irr_ords : irreflexive ords. +Proof. by elim=>//= a l ->; rewrite irr; case:eqP=> //=. Qed. + +Lemma trans_ords : transitive ords. +Proof. +elim=>[|y ys IHy][|x xs][|z zs]//=. +case:eqP=>//[->|H0];case:eqP=>//H; first by move/IHy; apply. +- by case:eqP=>//; rewrite -H; first (by move/H0). +case:eqP=>//[->|H1] H2; first by move/(nsym H2). +by move/(trans H2). +Qed. + +Lemma total_ords : forall x y, [|| ords x y, x == y | ords y x]. +Proof. +elim=>[|x xs IH][|y ys]//=; case:eqP=>//[->|H1]; + (case:eqP=>//= H; first (by rewrite orbT //=)). +- by case:eqP=>//H3 ; case: (or3P (IH ys))=> [-> | /eqP H0 | ->]; + [ rewrite orTb // | apply: False_ind; apply: H; rewrite H0 | rewrite orbT //]. +case:eqP; first by move/(esym)/H1. +by move=>_ ;case: (or3P (total x y))=>[-> //| /eqP /H1 //| -> //]; rewrite orbT. +Qed. + +Definition seq_ordMixin := OrdMixin irr_ords trans_ords total_ords. +Canonical Structure seq_ordType := Eval hnf in OrdType (seq T) seq_ordMixin. +End SeqOrd. + +(* A trivial total ordering for Unit *) +Section unitOrd. +Let ordtt (x y : unit ) := false. + +Lemma irr_tt : irreflexive ordtt. +Proof. by []. Qed. + +Lemma trans_tt : transitive ordtt. +Proof. by []. Qed. + +Lemma total_tt x y : [|| ordtt x y, x == y | ordtt y x ]. +Proof. by []. Qed. + +Let unit_ordMixin := OrdMixin irr_tt trans_tt total_tt. +Canonical Structure unit_ordType := Eval hnf in OrdType unit unit_ordMixin. +End unitOrd. + + +(* ordering with path, seq and last *) + +Lemma seq_last_in (A : eqType) (s : seq A) x : + last x s \notin s -> s = [::]. +Proof. +case: (lastP s)=>// {s} s y; case: negP=>//; elim; rewrite last_rcons. +by elim: s=>[|y' s IH]; rewrite /= inE // IH orbT. +Qed. + +Lemma path_last (A : ordType) (s : seq A) x : + path oleq x s -> oleq x (last x s). +Proof. +elim: s x=>[|y s IH] /= x; first by rewrite /oleq eq_refl orbT. +case/andP=>H1 /IH; case/orP=>H2; rewrite /oleq. +- case/orP: H1=>H1; first by rewrite (trans H1 H2). + by rewrite (eqP H1) H2. +by rewrite -(eqP H2); case/orP: H1=>-> //=; rewrite orbT. +Qed. + +(* in a sorted list, the last element is maximal *) +(* and the maximal element is last *) + +Lemma sorted_last_key_max (A : ordType) (s : seq A) x y : + sorted oleq s -> x \in s -> oleq x (last y s). +Proof. +elim: s x y=>[|z s IH] //= x y H; rewrite inE /=. +case: eqP=>[->|] /= _; first by apply: path_last. +by apply: IH (path_sorted H). +Qed. + +Lemma sorted_max_key_last (A : ordType) (s : seq A) x y : + sorted oleq s -> x \in s -> + (forall z, z \in s -> oleq z x) -> last y s = x. +Proof. +elim: s x y => [|w s IH] //= x y; rewrite inE /=. +case: eqP=>[<- /= H1 _ H2 | _ H /= H1 H2]; last first. +- apply: IH (path_sorted H) H1 _ => z H3; apply: H2. + by rewrite inE /= H3 orbT. +apply/eqP; move: (H2 (last x s)) (path_last H1); rewrite inE /= /oleq eq_sym. +case: totalP=>//=; case E: (last x s \in s)=>//. +by move/negbT/seq_last_in: E=>->; rewrite irr. +Qed. + +Lemma seq_last_mono (A : ordType) (s1 s2 : seq A) x : + path oleq x s1 -> path oleq x s2 -> + {subset s1 <= s2} -> + oleq (last x s1) (last x s2). +Proof. +case: s1=>/= [_ H1 _|a s]; first by apply: path_last H1. +case/andP=>H1 H2 H3 H; apply: sorted_last_key_max (path_sorted H3) _. +apply: {x s2 H1 H3} H; rewrite inE orbC -implyNb. +by case E: (_ \notin _) (@seq_last_in _ s a)=>//= ->. +Qed. + + + + + diff --git a/opam b/opam new file mode 100644 index 0000000..9fc6a49 --- /dev/null +++ b/opam @@ -0,0 +1,24 @@ +opam-version: "1.2" +name: "coq-fcsl-pcm" +version: "1.0.0" +maintainer: "FCSL " + +homepage: "http://software.imdea.org/fcsl/" +bug-reports: "https://github.com/imdea-software/fcsl-pcm/issues" +license: "Apache 2.0" + +build: [ make "-j%{jobs}%" ] +install: [ make "install" ] +remove: [ "sh" "-c" "rm -rf '%{lib}%/coq/user-contrib/fcsl'" ] +depends: [ + "coq" {>= "8.7" & < "8.8"} + "coq-mathcomp-ssreflect" {>= "1.6.2" & < "1.7"} +] + +tags: [ + "keyword:separation logic" + "keyword:partial commutative monoid" +] +authors: [ + "Aleksandar Nanevski" +] diff --git a/pcm/automap.v b/pcm/automap.v new file mode 100644 index 0000000..6fca5f9 --- /dev/null +++ b/pcm/automap.v @@ -0,0 +1,834 @@ +(* +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 fcsl Require Import pred ordtype pcm unionmap. +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +(******************************************************************************) +(* This file implementations of Canonical structure lemmas (a.k.a. overloaded *) +(* lemmas) for automating three tasks: *) +(* *) +(* 1. checking if implications of the form valid e1 -> valid e2 hold by *) +(* deciding if the terms in e2 are all contained in e1 *) +(* *) +(* 2. checking if dom_eq e1 e2 holds by cancelling the common terms, to *) +(* obtain residuals rs1 and rs2, and then issuing a subgoal dom_eq rs1 *) +(* rs2. *) +(* *) +(* 3. checking if the union e is undef, because it contains duplicate *) +(* pointers or an undef *) +(******************************************************************************) + +(* For each task, we have two implementations: a naive and a *) +(* sophisticated one. The lemmas validO, domeqO, invalidO are the naive *) +(* ones, and validX, domeqX, invalidX are the sophisticated ones. I keep *) +(* both O/X versions for now, for experimentation purposes, but *) +(* eventually should retain only validX and domeqX. *) + +(* First, two general helper concepts for searching in sequences. They will *) +(* be useful in syntactifying the expressions e1 and e2 for both tasks. *) +(* The concepts are: *) +(* *) +(* - onth: returns None to signal that an element is not found prefix: *) +(* - will be used for growing interpretation contexts *) + +Section Helpers. +Variable A : Type. + +Fixpoint onth (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 s1 s2 := + forall n x, onth s1 n = some x -> onth s2 n = some x. + +(* Lemmas *) + +Lemma size_onth s n : n < size s -> exists x, onth s n = Some x. +Proof. +elim: s n=>[//|a s IH] [|n] /=; first by exists a. +by rewrite -(addn1 n) -(addn1 (size s)) ltn_add2r; apply: IH. +Qed. + +Lemma onth_size s n x : onth s n = Some x -> n < size s. +Proof. by elim: s n=>[//|a s IH] [//|n]; apply: IH. 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_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. + +End Helpers. + +Hint Resolve prefix_refl. + +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. + +(* 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) (T : Type) (U : union_map_class K 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 := 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 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" *) + +(* In the cancellation algorithms, 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). + +Fixpoint rfilter ts : seq A := + if ts is t :: ts' then if p t then ts' else t :: rfilter ts' else [::]. +End OneShotFilter. + +(* now for reflection *) + +Section Reflection. +Variables (K : ordType) (T : Type) (U : union_map_class K 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 um_undef + | Var n => if onth (varx i) n is Some f then f else um_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. + +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 validO/validX *) +(* and domeqO/domeqX and invalidO/invalidX respectively. *) +(**************************************************************************) +(**************************************************************************) + +(* subterm is purely functional version of validO/validX *) + +Section Subterm. +Variables (K : ordType) (T : Type) (U : union_map_class K 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. +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 domeqO/domeqX *) + +Section Subtract. +Variables (K : ordType) (T : Type) (U : union_map_class K 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. +- 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 invalidO/invalidX *) + +Section Invalid. +Variables (K : ordType) (T : Type) (U : union_map_class K 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_undef join_undefL valid_undef. +Qed. + +End Invalid. + + +(********************************) +(********************************) +(* Canonical structure programs *) +(********************************) +(********************************) + +(* first a helper program for searching and inserting in a list *) + +Section XFind. +Variable A : Type. + +Structure tagged_elem := XTag {xuntag :> A}. + +(* Local Coercion untag : tagged_elem >-> A. *) + +Definition extend_tag := XTag. +Definition recurse_tag := extend_tag. +Canonical found_tag x := recurse_tag x. + +(* 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. + +(* Next, we syntactify a unionmap into a seq term as follows. *) +(* *) +(* - if the map is f1 \+ f2, then recurse over both and concat. the results *) +(* - if the map is the empty map, return [::] *) +(* - if the map is k \-> v then add k to the context, and return [Pts x v], *) +(* where x is the index for l in the context *) +(* if the map is whatever else, add the map to the context and return *) +(* [Var n], where n is the index for the map in the context *) + +Module Syntactify. +Section Syntactify. +Variables (K : ordType) (T : Type) (U : union_map_class K 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 validO/validX *) +(****************************) + +(* validO is a naive lemma that automates subterm checking *) +(* it leaves a goal of the form subterm ts2 ts1 *) +(* this one evaluates to a boolean, so it's easy to discharge *) +(* but it takes the space for the lemma's argument *) +(* which, thus, can't be picked from the goal *) + +Section ValidO. +Variables (K : ordType) (T : Type) (U : union_map_class K T). +Implicit Types (i : ctx U) (ts : seq (term T)). +Notation form := Syntactify.form. +Notation untag := Syntactify.untag. + +Lemma validO j k ts1 ts2 (f1 : form (@empx K T U) j ts1) + (f2 : form j k ts2) : + valid (untag f1) -> subterm ts2 ts1 -> valid (untag f2). +Proof. +case: f1 f2=>f1 [<- _ A1][f2][<- S2 A2] /= V; rewrite (sc_interp S2 A1) in V. +by case/(subterm_sound A2 (sc_wf S2 A1))=>xs /domeqP []; rewrite V=>/validL ->. +Qed. + +End ValidO. + +Arguments validO [K T U j k ts1 ts2 f1 f2] _ _. + +Example ex0 (x y z : nat) (v1 v2 : nat) h: + valid (Unit \+ y \\-> v1 \+ h \+ x \\-> v1) -> + valid (x \\-> v2 \+ Unit). +Proof. by move=>V; rewrite (validO V). Abort. + +(* validX is a more refined lemma for subterm checking which *) +(* automatically discharges the spurious argument from above *) + +Module ValidX. +Section ValidX. +Variables (K : ordType) (T : Type) (U : union_map_class K T). +Implicit Types (j : ctx U) (ts : seq (term T)). +Notation form := Syntactify.form. +Notation untag := Syntactify.untag. + +(* 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. + +Module Exports. +Canonical equate. +Canonical start. + +Section Exports. +Variables (K : ordType) (T : Type) (U : union_map_class K T). +Implicit Types (j : ctx U) (ts : seq (term T)). +Notation form := Syntactify.form. +Notation untag := Syntactify.untag. + +(* 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 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. + +End Exports. +End ValidX. + +Export ValidX.Exports. + +(****************************) +(* Automating domeqO/domeqX *) +(****************************) + +(* naive lemma first *) + +Section DomeqO. +Variables (K : ordType) (T : Type) (U : union_map_class K T). +Implicit Types (j : ctx U) (ts : seq (term T)). +Notation form := Syntactify.form. +Notation untag := Syntactify.untag. + +Lemma domeqO j k rs1 rs2 ts1 ts2 + (f1 : form (empx U) j ts1) (f2 : form j k ts2) : + subtract ts1 ts2 [::] = (rs1, rs2) -> + dom_eq (pprint k (rev rs1)) (pprint k rs2) -> + dom_eq (untag f1) (untag f2). +Proof. +case: f1 f2=>f1 [<- _ A1][f2][<- S A2]. +case/(subtract_sound (sc_wf S A1) A2)=>// ys [/= D1 D2]. +rewrite unitR in D1; rewrite (sc_interp S A1). +rewrite !pp_interp interp_rev => D; apply: domeq_trans D1 _. +rewrite domeq_sym; apply: domeq_trans D2 _. +by rewrite domeq_sym; apply: domeqUn. +Qed. + +End DomeqO. + +Example ex0 (x y z : nat) (v1 v2 : nat) h: + dom_eq (Unit \+ y \\-> v1 \+ h \+ x \\-> v1) (x \\-> v2 \+ Unit). +Proof. apply: domeqO=>//=. Abort. + +Module DomeqX. +Section DomeqX. +Variables (K : ordType) (T : Type) (U : union_map_class K T). +Implicit Types (j : ctx U) (ts : seq (term T)). +Notation form := Syntactify.form. +Notation untag := Syntactify.untag. + +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_sym; apply: domeq_trans D2 _. +by rewrite domeq_sym; 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) (T : Type) (U : union_map_class K T). +Implicit Types (j : ctx U) (ts : seq (term T)). +Notation form := Syntactify.form. +Notation untag := Syntactify.untag. + +(* 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 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 invalidO/invalidX *) +(********************************) + +(* naive lemma first *) + +Section InvalidO. +Variables (K : ordType) (T : Type) (U : union_map_class K T). +Implicit Types (i : ctx U) (ts : seq (term T)). +Notation form := Syntactify.form. +Notation untag := Syntactify.untag. + +Lemma undefO i ts (f : form (empx U) i ts) : + isundef i ts -> untag f = um_undef. +Proof. by case: f=>f [<- _ A] /(isundef_sound A)/invalidE. Qed. + +Lemma invalidO i ts (f : form (empx U) i ts) : + isundef i ts -> valid (untag f) = false. +Proof. by move/undefO=>->; rewrite valid_undef. Qed. + +End InvalidO. + +Example ex0 (x y z : nat) (v1 v2 : nat) h: + (Unit \+ y \\-> v1 \+ h \+ y \\-> v1) = um_undef. +Proof. by apply: undefO. Abort. + +(* now the sophisticated one *) + +Module InvalidX. +Section InvalidX. +Variables (K : ordType) (T : Type) (U : union_map_class K T). +Implicit Types (i : ctx U) (ts : seq (term T)). +Notation form := Syntactify.form. +Notation untag := Syntactify.untag. + +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) (T : Type) (U : union_map_class K T). +Implicit Types (i : ctx U) (ts : seq (term T)). +Notation form := Syntactify.form. +Notation untag := Syntactify.untag. + +(* the main lemmas *) + +Lemma undefX m i ts (g : rform i ts m true) : unpack (pivot g) = um_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 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. + + + + + diff --git a/pcm/axioms.v b/pcm/axioms.v new file mode 100644 index 0000000..10bd890 --- /dev/null +++ b/pcm/axioms.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 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. +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +(*****************************) +(* Axioms and extensionality *) +(*****************************) + +(* 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 eta A (B : A -> Type) (f : forall x, B x) : f = [eta f]. +Proof. by apply: fext. Qed. + +Lemma sval_inj A P : injective (@sval A P). +Proof. +move=>[x Hx][y Hy] /= H; move: Hx Hy; rewrite H=>*. +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 *) +(*****************************************) + +Section Cast. +Variable (C : Type) (interp : C -> 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) := forall pf, v = cast pf w. + +Lemma jmrefl A (v : interp A) : jmeq v v. +Proof. by move=>pf; rewrite eqc. Qed. + +Lemma jmsym A B (v : interp A) (w : interp B) : jmeq v w -> jmeq w v. +Proof. +move=> H pf; rewrite (H (esym pf)). +by move: (pf); rewrite pf in w H * => {pf} pf; rewrite !eqc. +Qed. + +Lemma jmE A (v w : interp A) : jmeq v w <-> v = w. +Proof. by split=>[/(_ erefl) //|->]; apply: jmrefl. Qed. + +Lemma castE A B (pf1 pf2 : A = B) (v1 v2 : interp B) : + v1 = v2 <-> cast pf1 v1 = cast pf2 v2. +Proof. by move: (pf1) pf2; rewrite pf1 =>*; rewrite !eqc. Qed. + +End Cast. + +Arguments cast {C} interp [A][B] pf v. +Arguments jmeq {C} interp [A][B] v w. +Hint Resolve jmrefl. +(* 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). + +(* 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 {C} interp {A} _ : rename. +Notation idyn v := (@dyn _ id _ v). + +Lemma dynE (A B : Type) interp (pf : A = B) (v : interp A) (w : interp B) : + jmeq interp v w <-> dyn interp v = dyn interp w. +Proof. by rewrite -pf in w *; rewrite jmE; split => [->|/dyn_inj]. Qed. + diff --git a/pcm/heap.v b/pcm/heap.v new file mode 100644 index 0000000..82b8d51 --- /dev/null +++ b/pcm/heap.v @@ -0,0 +1,764 @@ +(* +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 Coq Require Import ssreflect ssrbool ssrfun Eqdep. +From mathcomp Require Import ssrnat eqtype seq path. +From fcsl Require Import axioms ordtype pcm finmap unionmap. +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +(*************) +(* Locations *) +(*************) + +Inductive ptr := ptr_nat of nat. + +Definition null := ptr_nat 0. + +Definition nat_ptr (x : ptr) := let: ptr_nat y := x in y. + +Definition eq_ptr (x y : ptr) : bool := + match x, y with ptr_nat m, ptr_nat n => m == n end. + +Lemma eq_ptrP : Equality.axiom eq_ptr. +Proof. by case=>x [y] /=; case: eqP=>[->|*]; constructor=>//; case. Qed. + +Definition ptr_eqMixin := EqMixin eq_ptrP. +Canonical ptr_eqType := EqType ptr ptr_eqMixin. + +(* some pointer arithmetic: offsetting from a base *) + +Definition ptr_offset x i := ptr_nat (nat_ptr x + i). + +Notation "x .+ i" := (ptr_offset x i) + (at level 3, format "x .+ i"). + +Lemma ptrE x y : (x == y) = (nat_ptr x == nat_ptr y). +Proof. by move: x y=>[x][y]. Qed. + +Lemma ptr0 x : x.+0 = x. +Proof. by case: x=>x; rewrite /ptr_offset addn0. Qed. + +Lemma ptrA x i j : x.+i.+j = x.+(i+j). +Proof. by case: x=>x; rewrite /ptr_offset addnA. Qed. + +Lemma ptrK x i j : (x.+i == x.+j) = (i == j). +Proof. by case: x=>x; rewrite ptrE eqn_add2l. Qed. + +Lemma ptr_null x m : (x.+m == null) = (x == null) && (m == 0). +Proof. by case: x=>x; rewrite !ptrE addn_eq0. Qed. + +Lemma ptrT x y : {m : nat | (x == y.+m) || (y == x.+m)}. +Proof. +case: x y=>x [y]; exists (if x <= y then (y - x) else (x - y)). +rewrite !ptrE leq_eqVlt /=. +by case: (ltngtP x y)=>/= E; rewrite subnKC ?(ltnW E) ?eq_refl ?orbT // E. +Qed. + +Definition ltn_ptr (x y : ptr) := + match x, y with ptr_nat m, ptr_nat n => m < n end. + +Lemma ltn_ptr_irr : irreflexive ltn_ptr. +Proof. by case=>x /=; rewrite ltnn. Qed. + +Lemma ltn_ptr_trans : transitive ltn_ptr. +Proof. by case=>x [y][z]; apply: ltn_trans. Qed. + +Lemma ltn_ptr_total : forall x y : ptr, [|| ltn_ptr x y, x == y | ltn_ptr y x]. +Proof. by case=>x [y]; rewrite ptrE /=; case: ltngtP. Qed. + +Definition ptr_ordMixin := OrdMixin ltn_ptr_irr ltn_ptr_trans ltn_ptr_total. +Canonical ptr_ordType := OrdType ptr ptr_ordMixin. + +(*********) +(* Heaps *) +(*********) + +Module Heap. + +Inductive heap := + Undef | Def (finmap : {finMap ptr -> dynamic id}) of + null \notin supp finmap. + +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: bool_irrelevance. +Qed. + +End NullLemmas. + + +(* methods *) + +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 dom_eq h1 h2 := + match h1, h2 with + Def f1 _, Def f2 _ => supp f1 == supp f2 + | Undef, Undef => true + | _, _ => false + end. + +Definition free x h := + 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 um_filter q f := + if f is Def fs pf then Def (@filt_nullP fs q pf) else Undef. + +Definition pts (x : ptr) v := upd x v empty. + +Definition empb h := if h is Def hs _ then supp hs == [::] else false. + +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 [::]. + +Local Notation base := + (@UM.base ptr_ordType (dynamic id) (fun k => k != null)). + +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. + +Lemma ftE b : from (to b) = b. +Proof. by case: b=>// f H; rewrite UM.umapE. Qed. + +Lemma tfE f : to (from f) = f. +Proof. by case: f=>// f H; rewrite heapE. Qed. + +Lemma undefE : Undef = to (@UM.Undef _ _ _). +Proof. by []. Qed. + +Lemma defE f : def f = UM.valid (from f). +Proof. by case: f. Qed. + +Lemma emptyE : empty = to (@UM.empty _ _ _). +Proof. by rewrite heapE. Qed. + +Lemma updE k v f : upd k v f = to (UM.upd k v (from f)). +Proof. by case: f=>[|f H] //=; case: decP=>// H1; rewrite heapE. Qed. + +Lemma domE f : dom f = UM.dom (from f). +Proof. by case: f. Qed. + +Lemma dom_eqE f1 f2 : dom_eq f1 f2 = UM.dom_eq (from f1) (from f2). +Proof. by case: f1 f2=>[|f1 H1][|f2 H2]. Qed. + +Lemma freeE k f : free k f = to (UM.free k (from f)). +Proof. by case: f=>[|f H] //; rewrite heapE. Qed. + +Lemma findE k f : find k f = UM.find k (from f). +Proof. by case: f. Qed. + +Lemma unionE f1 f2 : union f1 f2 = to (UM.union (from f1) (from f2)). +Proof. +case: f1 f2=>[|f1 H1][|f2 H2] //; rewrite /union /UM.union /=. +by case: ifP=>D //; rewrite heapE. +Qed. + +Lemma umfiltE q f : um_filter q f = to (UM.um_filter q (from f)). +Proof. by case: f=>[|f H] //; rewrite heapE. Qed. + +Lemma empbE f : empb f = UM.empb (from f). +Proof. by case: f. Qed. + +Lemma undefbE f : undefb f = UM.undefb (from f). +Proof. by case: f. Qed. + +Lemma ptsE k (v : dynamic id) : pts k v = to (@UM.pts _ _ _ k v). +Proof. +by rewrite /pts /UM.pts /UM.upd /=; case: decP=>// H; rewrite heapE. +Qed. + +Module Exports. + +(* the inheritance from PCMs *) + +Definition heapUMCMix := + UMCMixin ftE tfE defE undefE emptyE updE domE dom_eqE freeE + findE unionE umfiltE empbE undefbE ptsE. +Canonical heapUMC := Eval hnf in UMC heap heapUMCMix. + +Definition heapPCMMix := union_map_classPCMMix heapUMC. +Canonical heapPCM := Eval hnf in PCM heap heapPCMMix. + +Definition heapCPCMMix := union_map_classCPCMMix heapUMC. +Canonical heapCPCM := Eval hnf in CPCM heap heapCPCMMix. + +Definition heapTPCMMix := union_map_classTPCMMix heapUMC. +Canonical heapTPCM := Eval hnf in TPCM heap heapTPCMMix. + +End Exports. + +End Heap. + +Export Heap.Exports. + +Notation heap := Heap.heap. + +Definition heap_pts A (x : ptr) (v : A) := + @UMC.pts _ _ heapUMC x (idyn v). + +Notation "x :-> v" := (@heap_pts _ x v) (at level 30). + +(*****************************) +(* Specific points-to lemmas *) +(*****************************) + +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 x h. +Proof. by case/um_eta; case=>A v H; exists A, v. Qed. + +Lemma heap_eta2 A x h (v : A) : + find x h = Some (idyn v) -> h = x :-> v \+ free x h. +Proof. +move=>E; case: (heap_eta (find_some E))=>B [w][]. +rewrite {}E; case=>E; rewrite -E in w *. +by move/(@inj_pair2 _ _ _ _ _)=>->. +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 x1 h2 = free x2 h1, + h1 = x2 :-> v2 \+ free x1 h2 & + h2 = x1 :-> v1 \+ free x2 h1]. +Proof. by move=>V /(cancel2 V); case: ifP=>// _ [/dyn_inj]. Qed. + +End HeapPointsToLemmas. + +Prenex Implicits heap_eta2. + +(******************************************) +(* additional stuff, on top of union maps *) +(* mostly random stuff, kept for legacy *) +(* should be removed/redone eventually *) +(******************************************) + +Definition fresh (h : heap) := + (if h is Heap.Def hs _ then last null (supp hs) else null) .+ 1. + +Definition pick (h : heap) := + if h is Heap.Def hs _ then head null (supp hs) else null. + +(*********) +(* fresh *) +(*********) + +Lemma path_last n s x : path ord x s -> ord x (last x s).+(n+1). +Proof. +move: n s x. +suff L: forall s x, path ord x s -> ord x (last x s).+(1). +- elim=>[|n IH] // s x; move/IH=>E; apply: trans E _. + by case: (last x s)=>m; rewrite /ord /= addSn (addnS m). +elim=>[|y s IH x] /=; first by case=>x; rewrite /ord /= addn1. +by case/andP=>H1; move/IH; apply: trans H1. +Qed. + +Lemma path_filter (A : ordType) (s : seq A) (p : pred A) x : + path ord x s -> path ord x (filter p s). +Proof. +elim: s x=>[|y s IH] x //=. +case/andP=>H1 H2. +case: ifP=>E; first by rewrite /= H1 IH. +apply: IH; elim: s H2=>[|z s IH] //=. +by case/andP=>H2 H3; rewrite (@trans _ y). +Qed. + +Lemma dom_fresh h n : (fresh h).+n \notin dom h. +Proof. +suff L2: forall h x, x \in dom h -> ord x (fresh h). +- by apply: (contra (L2 _ _)); rewrite -leqNgt leq_addr. +case=>[|[s H1]] //; rewrite /supp => /= H2 x. +rewrite /dom /fresh /supp /=. +elim: s H1 null H2 x=>[|[y d] s IH] //= H1 x. +rewrite inE negb_or; case/andP=>H3 H4 z; rewrite inE. +case/orP; first by move/eqP=>->{z}; apply: (path_last 0). +by apply: IH; [apply: path_sorted H1 | apply: notin_path H1]. +Qed. + +Lemma fresh_null h : fresh h != null. +Proof. by rewrite -lt0n addn1. Qed. + +Opaque fresh. + +Hint Resolve dom_fresh fresh_null. + +(********) +(* pick *) +(********) + +Lemma emp_pick (h : heap) : (pick h == null) = (~~ valid h || empb h). +Proof. +rewrite /empb; case: h=>[|h] //=; case: (supp h)=>[|x xs] //=. +by rewrite inE negb_or eq_sym; case/andP; move/negbTE=>->. +Qed. + +Lemma pickP h : valid h && ~~ empb h = (pick h \in dom h). +Proof. +rewrite /dom /empb; case: h=>[|h] //=. +by case: (supp h)=>// *; rewrite inE eq_refl. +Qed. + + +(***********************) +(* Some derived lemmas *) +(***********************) + +Lemma domPtUnX A (v : A) x i : valid (x :-> v \+ i) -> x \in dom (x :-> v \+ i). +Proof. by move=>D; rewrite domPtUn inE /= D eq_refl. Qed. + +Lemma domPtX A (v : A) x : valid (x :-> v) -> x \in dom (x :-> v). +Proof. by move=>D; rewrite -(unitR (x :-> v)) domPtUnX // unitR. Qed. + + +(*******************************************) +(* Block update for reasoning about arrays *) +(*******************************************) + +Section BlockUpdate. +Variable (A : Type). + +Fixpoint updi x (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. +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. +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. +rewrite ptrA domPtUn inE /= negb_and negb_or -{4}(ptr0 x) ptrK -lt0n H /=. +by rewrite orbC IH // addn1. +Qed. + +Lemma updiD x xs : valid (updi x xs) = (x != null) || (size xs == 0). +Proof. +elim: xs x=>[|v xs IH] x //=; first by rewrite orbC. +by rewrite validPtUn updiVm' // orbF IH ptr_null andbF andbC. +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 //=. +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)=>/=. +- by case: xs=>// a s; rewrite (eqP H). +elim: xs x m H=>[|v vs IH] x m H //; case: m=>[|m]. +- by rewrite ptr0 /= domPtUn inE /= eq_refl andbT -updiS updiD H. +rewrite -addn1 addnC -ptrA updiS domPtUn inE /= IH; last first. +- by rewrite ptrE /= addn1. +by rewrite -updiS updiD H /= -{1}(ptr0 x) 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)=>/=. +- by rewrite (eqP H); elim: xs=>[|z xs IH] //=; 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; +[move/esym| |]; try by rewrite empbE empbUn empbPt. +by case/(hcancelV D)=><- {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} 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. +move=>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. + +(*************************************) +(* the automation of the PtUn lemmas *) +(*************************************) + +(* First, the mechanism for search-and-replace for the overloaded lemas, *) +(* pattern-matching on heap expressions. *) + +Structure tagged_heap := Tag {untag :> heap}. + +Definition right_tag := Tag. +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; apply: validPtUnV. 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 x (untag f) = 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; apply: domPtUnX. 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 ?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 case: r1; case=>r1 H1; case: r2; case=>r2 H2 /= ->. Qed. + +End Array. + diff --git a/pcm/lift.v b/pcm/lift.v new file mode 100644 index 0000000..6ee129a --- /dev/null +++ b/pcm/lift.v @@ -0,0 +1,241 @@ +(* +Copyright 2018 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. +From fcsl Require Import prelude pcm. +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +(******************************************************************************) +(* This file contains an implementation of lift type adding an undef element *) +(* a structure. *) +(* Lifting turns an unlifted structure into a PCM and preserves equality. *) +(******************************************************************************) + +Module Unlifted. + +Record mixin_of (T : Type) := Mixin { + ounit_op : T; + ojoin_op : T -> T -> option T; + ojoinC_op : forall x y, ojoin_op x y = ojoin_op y x; + ojoinA_op : forall x y z, + obind (ojoin_op x) (ojoin_op y z) = obind (ojoin_op^~ z) (ojoin_op x y); + ounitL_op : forall x, ojoin_op ounit_op x = Some x}. + +Section ClassDef. + +Notation class_of := mixin_of (only parsing). + +Structure type : Type := Pack {sort : Type; _ : class_of sort; _ : Type}. +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 pack c := @Pack T c T. +Definition clone := fun c & cT -> T & phant_id (pack c) cT => pack c. + +Definition ounit := ounit_op class. +Definition ojoin := ojoin_op class. + +End ClassDef. + +Module Exports. +Coercion sort : type >-> Sortclass. +Notation unlifted := type. +Notation UnliftedMixin := Mixin. +Notation Unlifted T m := (@pack T m). + +Notation "[ 'unliftedMixin' 'of' T ]" := (class _ : mixin_of T) + (at level 0, format "[ 'unliftedMixin' 'of' T ]") : form_scope. +Notation "[ 'unlifted' 'of' T 'for' C ]" := (@clone T C _ idfun id) + (at level 0, format "[ 'unlifted' 'of' T 'for' C ]") : form_scope. +Notation "[ 'unlifted' 'of' T ]" := (@clone T _ _ id id) + (at level 0, format "[ 'unlifted' 'of' T ]") : form_scope. + +Notation ounit := ounit. +Notation ojoin := ojoin. + +Arguments ounit [cT]. + +Lemma ojoinC (U : unlifted) (x y : U) : ojoin x y = ojoin y x. +Proof. by case: U x y=>T [ou oj ojC]. Qed. + +Lemma ojoinA (U : unlifted) (x y z : U) : + obind (ojoin x) (ojoin y z) = obind (@ojoin U^~ z) (ojoin x y). +Proof. by case: U x y z=>T [ou oj ojC ojA]. Qed. + +Lemma ounitL (U : unlifted) (x : U) : ojoin ounit x = Some x. +Proof. by case: U x=>T [ou oj ojC ojA ojL]. Qed. + +End Exports. + +End Unlifted. + +Export Unlifted.Exports. + +(**************************************************) +(* Lifting turns an unlifted structure into a PCM *) +(**************************************************) + +Inductive lift A := down | up of A. + +Module Lift. +Section Lift. +Variable A : unlifted. + +Let unit := up (@ounit A). + +Let valid := + [fun x : lift A => if x is up _ then true else false]. + +Let join := + [fun x y : lift A => + if (x, y) is (up v, up w) then + if ojoin v w is Some u then up u + else @down A + else @down A]. + +Lemma joinC (x y : lift A) : join x y = join y x. +Proof. by case: x y=>[|x][|y] //=; rewrite ojoinC. Qed. + +Lemma joinA (x y z : lift A) : join x (join y z) = join (join x y) z. +Proof. +case: x y z =>[|x][|y][|z] //=; first by case: (ojoin x y). +case E1: (ojoin y z)=>[v1|]. +- case E2: (ojoin x y)=>[v2|]; + by move: (ojoinA x y z); rewrite E1 E2 /=; move=>->. +case E2: (ojoin x y)=>[v2|] //. +by move: (ojoinA x y z); rewrite E1 E2 /= =><-. +Qed. + +Lemma unitL x : join unit x = x. +Proof. by case: x=>[|x] //=; rewrite ounitL. Qed. + +Lemma validL x y : valid (join x y) -> valid x. +Proof. by case: x y=>[|x][|y]. Qed. + +Lemma validU : valid unit. +Proof. by []. Qed. + +End Lift. + +(* Lifting preserves equality types *) +Section LiftEqType. +Variable A : eqType. + +Definition lift_eq (u v : lift A) := + match u, v with + up a, up b => a == b + | down, down => true + | _, _ => false + end. + +Lemma lift_eqP : Equality.axiom lift_eq. +Proof. +case=>[|x][|y] /=; try by constructor. +by apply: (iffP eqP)=>[->|[]]. +Qed. + +Definition liftEqMixin := EqMixin LiftEqType.lift_eqP. +Canonical liftEqType := Eval hnf in EqType _ liftEqMixin. + +End LiftEqType. + +Module Exports. +Arguments down [A]. +Arguments up [A]. +Canonical liftEqType. + +Section Exports. +(* View for pattern-matching lifted pcm's *) + +CoInductive lift_spec A (x : lift A) : lift A -> Type := +| up_spec n of x = up n : lift_spec x (up n) +| undef_spec of x = down : lift_spec x down. + +Lemma liftP A (x : lift A) : lift_spec x x. +Proof. by case: x=>[|a]; [apply: undef_spec | apply: up_spec]. Qed. + +Variable A : unlifted. + +Definition liftPCMMixin := + PCMMixin (@Lift.joinC A) (@Lift.joinA A) + (@Lift.unitL A) (@Lift.validL A) (@Lift.validU A). +Canonical liftPCM := Eval hnf in PCM (lift A) liftPCMMixin. + +Lemma upE (a1 a2 : A) : + up a1 \+ up a2 = if ojoin a1 a2 is Some a then up a else down. +Proof. by []. Qed. + +Lemma valid_down : valid (@down A) = false. +Proof. by []. Qed. + +Lemma down_join (x : lift A) : down \+ x = down. +Proof. by []. Qed. + +Lemma join_down (x : lift A) : x \+ down = down. +Proof. by case: x. Qed. + +Definition downE := (down_join, join_undef, valid_down). + +CoInductive liftjoin_spec (x y : lift A) : _ -> _ -> _ -> Type := +| upcase1 n1 n2 of x = up n1 & y = up n2 : + liftjoin_spec x y (if ojoin n1 n2 is Some u then up u else down) x y +| invalid1 of ~~ valid (x \+ y) : liftjoin_spec x y down x y. + +Lemma liftPJ (x y : lift A) : liftjoin_spec x y (x \+ y) x y. +Proof. by case: x y=>[|x][|y]; rewrite ?downE; constructor. Qed. + +End Exports. +End Exports. +End Lift. + +Export Lift.Exports. + + +Module NatUnlift. + +Local Definition ojoin (x y : nat) := Some (x + y). +Local Definition ounit := 0. + +Lemma ojoinC x y : ojoin x y = ojoin y x. +Proof. by rewrite /ojoin addnC. Qed. + +Lemma ojoinA x y z : + obind (ojoin x) (ojoin y z) = obind (ojoin^~ z) (ojoin x y). +Proof. by rewrite /obind/ojoin /= addnA. Qed. + +Lemma ounitL x : ojoin ounit x = Some x. +Proof. by case: x. Qed. + +End NatUnlift. + +Definition natUnliftedMix := + UnliftedMixin NatUnlift.ojoinC NatUnlift.ojoinA NatUnlift.ounitL. +Canonical natUnlifted := Eval hnf in Unlifted nat natUnliftedMix. + +(* some lemmas for lifted nats *) + +Lemma nxV (m1 m2 : lift natUnlifted) : + valid (m1 \+ m2) -> exists n1 n2, m1 = up n1 /\ m2 = up n2. +Proof. by case: m1=>// n1; case: m2=>// n2; exists n1, n2. Qed. + + +Lemma nxE0 (n1 n2 : lift natUnlifted) : + n1 \+ n2 = up 0 -> (n1 = up 0) * (n2 = up 0). +Proof. +case: n1 n2=>[|n1][|n2] //; rewrite upE /ojoin /=. +by case=>/eqP; rewrite addn_eq0=>/andP [/eqP -> /eqP ->]. +Qed. diff --git a/pcm/mutex.v b/pcm/mutex.v new file mode 100644 index 0000000..3618986 --- /dev/null +++ b/pcm/mutex.v @@ -0,0 +1,304 @@ +(* +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 Coq Require Import ssreflect ssrbool ssrfun. +From mathcomp Require Import ssrnat eqtype. +From fcsl Require Import prelude pcm. +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +(***********************) +(* Generalized mutexes *) +(***********************) + +Module GeneralizedMutex. +Section GeneralizedMutex. +Variable T : Type. (* the mutex stages, excluding undef and unit *) + +Inductive mutex := undef | nown | own | mx of T. + +Definition join x y := + match x, y with + undef, _ => undef + | _, undef => undef + | nown, x => x + | x, nown => x + | own, _ => undef + | _, own => undef + | mx _, mx _ => undef + end. + +Definition valid x := if x is undef then false else true. + +Lemma joinC : commutative join. +Proof. by case=>[|||x]; case=>[|||y]. Qed. + +Lemma joinA : associative join. +Proof. by case=>[|||x]; case=>[|||y]; case=>[|||z]. Qed. + +Lemma unitL : left_id nown join. +Proof. by case. Qed. + +Lemma validL x y : valid (join x y) -> valid x. +Proof. by case: x. Qed. + +Lemma valid_unit : valid nown. +Proof. by []. Qed. + +Definition mutexPCMMix := PCMMixin joinC joinA unitL validL valid_unit. +Canonical mutexPCM := Eval hnf in PCM mutex mutexPCMMix. + +(* cancelativity *) + +Lemma joinmK m1 m2 m : valid (m1 \+ m) -> m1 \+ m = m2 \+ m -> m1 = m2. +Proof. by case: m m1 m2=>[|||m][|||m1][|||m2]. Qed. + +Definition mutexCPCMMix := CPCMMixin joinmK. +Canonical mutexCPCM := Eval hnf in CPCM mutex mutexCPCMMix. + +(* topped structure *) + +Lemma unitb (x : mutex) : + reflect (x = Unit) (if x is nown then true else false). +Proof. by case: x; constructor. Qed. + +Lemma join0E (x y : mutex) : x \+ y = Unit -> x = Unit /\ y = Unit. +Proof. by case: x y=>[|||x][|||y]. Qed. + +Lemma valid3 (x y z : mutex) : valid (x \+ y \+ z) = + [&& valid (x \+ y), valid (y \+ z) & valid (x \+ z)]. +Proof. by case: x y z=>[|||x][|||y][|||z]. Qed. + +Lemma valid_undef : ~~ valid undef. +Proof. by []. Qed. + +Lemma undef_join x : undef \+ x = undef. +Proof. by []. Qed. + +Definition mutexTPCMMix := TPCMMixin unitb join0E valid_undef undef_join. +Canonical mutexTPCM := Eval hnf in TPCM mutex mutexTPCMMix. + +End GeneralizedMutex. + +Section Equality. +Variable T : eqType. + +Definition mutex_eq (x y : mutex T) := + match x, y with + undef, undef => true + | nown, nown => true + | own, own => true + | mx x', mx y' => x' == y' + | _, _ => false + end. + +Lemma mutex_eqP : Equality.axiom mutex_eq. +Proof. +case=>[|||x]; case=>[|||y] /=; try by constructor. +by case: eqP=>[->|H]; constructor=>//; case=>/H. +Qed. + +Definition mutexEqMix := EqMixin mutex_eqP. +Canonical mutexEqType := Eval hnf in EqType (mutex T) mutexEqMix. + +End Equality. + +Module Exports. +Canonical mutexPCM. +Canonical mutexCPCM. +Canonical mutexEqType. +Canonical mutexTPCM. +Notation mutex := mutex. +Notation mx := mx. +Notation mx_undef := undef. +Notation nown := nown. +Notation own := own. +Notation mutexPCM := mutexPCM. +Notation mutexTPCM := mutexTPCM. +Arguments mx_undef [T]. +Arguments nown [T]. +Arguments own [T]. + +(* binary mutexes can be obtained with T = Empty_set *) +Notation mtx := (mutex Empty_set). + +End Exports. +End GeneralizedMutex. + +Export GeneralizedMutex.Exports. + +(* some lemmas for generalized mutexes *) + +Section MutexLemmas. +Variable T : Type. +Implicit Types (t : T) (x y : mutex T). + +CoInductive mutex_spec x : mutex T -> Type := +| mutex_undef of x = mx_undef : mutex_spec x mx_undef +| mutex_nown of x = Unit : mutex_spec x Unit +| mutex_own of x = own : mutex_spec x own +| mutex_mx t of x = mx t : mutex_spec x (mx t). + +Lemma mxP x : mutex_spec x x. +Proof. by case: x=>[|||t]; constructor. Qed. + +Lemma mxE0 x y : x \+ y = Unit -> (x = Unit) * (y = Unit). +Proof. by case: x=>[|||t1]; case: y=>[|||t2]. 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. + +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 joinC joinAC=>/mxMx; rewrite joinC. 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. + +(* and the same for own *) + +Lemma mxOx x : valid (own \+ x) -> x = Unit. +Proof. by case: x. Qed. + +Lemma mxxO x : valid (x \+ own) -> x = Unit. +Proof. by rewrite joinC=>/mxOx. Qed. + +Lemma mxxyO x y : valid (x \+ (y \+ own)) -> x \+ y = Unit. +Proof. by rewrite joinC joinAC=>/mxOx; rewrite joinC. Qed. + +Lemma mxOxy x y : valid (own \+ x \+ y) -> x \+ y = Unit. +Proof. by rewrite -joinA=>/mxOx. Qed. + +Lemma mxxOy x y : valid (x \+ (own \+ y)) -> x \+ y = Unit. +Proof. by rewrite joinCA=>/mxOx. Qed. + +Lemma mxyOx x y : valid (y \+ own \+ x) -> y \+ x = Unit. +Proof. by rewrite joinAC=>/mxOx. Qed. + +(* inversion principle for join *) +(* own and mx are prime elements, and unit does not factor *) +CoInductive mxjoin_spec (x y : mutex T) : + mutex T -> mutex T -> mutex T -> Type := +| bothnown of x = nown & y = nown : mxjoin_spec x y nown nown nown +| leftown of x = own & y = nown : mxjoin_spec x y own own nown +| rightown of x = nown & y = own : mxjoin_spec x y own nown own +| leftmx t of x = mx t & y = nown : mxjoin_spec x y (mx t) (mx t) nown +| rightmx t of x = nown & y = mx t : mxjoin_spec x y (mx t) nown (mx t) +| invalid of ~~ valid (x \+ y) : mxjoin_spec x y undef x y. + +Lemma mxPJ x y : mxjoin_spec x y (x \+ y) x y. +Proof. by case: x y=>[|||x][|||y]; constructor. Qed. + +End MutexLemmas. + +Prenex Implicits mxMx mxxM mxxyM mxMxy mxxMy mxyMx + mxOx mxxO mxxyO mxOxy mxxOy mxyOx. + +(* specific lemmas for binary mutexes *) + +Lemma mxON (x : mtx) : valid x -> x != own -> x = nown. +Proof. by case: x. Qed. + +Lemma mxNN (x : mtx) : valid x -> x != nown -> x = own. +Proof. by case: x. Qed. + +(* the next batch of lemmas is for automatic simplification *) + +Section MutexRewriting. +Variable T : eqType. +Implicit Types (t : T) (x : mutex T). + +Lemma mxE1 : ((@mx_undef T == nown) = false) * + ((@nown T == mx_undef) = false) * + ((@mx_undef T == own) = false) * + ((@own T == mx_undef) = false) * + ((@nown T == own) = false) * + ((@own T == nown) = false). +Proof. by []. Qed. + +Lemma mxE2 t : ((mx t == nown) = false) * + ((nown == mx t) = false) * + ((mx t == mx_undef) = false) * + ((mx_undef == mx t) = false) * + ((mx t == own) = false) * + ((own == mx t) = false). +Proof. by []. Qed. + +Lemma mxE3 t x : ((mx t \+ x == nown) = false) * + ((x \+ mx t == nown) = false) * + ((nown == mx t \+ x) = false) * + ((nown == x \+ mx t) = false) * + ((mx t \+ x == own) = false) * + ((x \+ mx t == own) = false) * + ((own \+ x == mx t) = false) * + ((x \+ own == mx t) = false). +Proof. by case: x. Qed. + +Lemma mxE4 x : + ((own \+ x == nown) = false) * + ((x \+ own == nown) = false) * + ((own \+ x == own) = (x == Unit)) * + ((x \+ own == own) = (x == Unit)). +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/idP/andP; [case/eqP/cancelMx=>->->|case=>/eqP->/eqP->]. +by do !split=>//; rewrite ?L // eq_sym L eq_sym. +Qed. + +Lemma mx_valid t : valid (mx t). +Proof. by []. Qed. + +Lemma mx_valid_own : valid (@own T). +Proof. by []. Qed. + +Lemma mx_injE t1 t2 : (mx t1 == mx t2) = (t1 == t2). +Proof. by []. Qed. + +Definition mxE := (mxE1, mxE2, mxE3, mxE4, mxE5, mx_injE, + mx_valid, mx_valid_own, + (* plus a bunch of safe simplifications *) + unitL, unitR, valid_unit, eq_refl, + valid_undef, undef_join, join_undef). + +End MutexRewriting. diff --git a/pcm/natmap.v b/pcm/natmap.v new file mode 100644 index 0000000..79cf14c --- /dev/null +++ b/pcm/natmap.v @@ -0,0 +1,607 @@ +(* +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 Coq Require Import ssreflect ssrbool ssrfun. +From mathcomp Require Import ssrnat eqtype seq path. +From fcsl Require Import prelude ordtype pcm finmap unionmap. +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +(************************) +(************************) +(* Maps over non-0 nats *) +(************************) +(************************) + +Module Type NMSig. +Parameter tp : Type -> Type. + +Section Params. +Variable A : Type. +Notation tp := (tp A). + +Parameter nm_undef : tp. +Parameter defined : tp -> bool. +Parameter empty : tp. +Parameter upd : nat -> A -> tp -> tp. +Parameter dom : tp -> seq nat. +Parameter dom_eq : tp -> tp -> bool. +Parameter free : nat -> tp -> tp. +Parameter find : nat -> tp -> option A. +Parameter union : tp -> tp -> tp. +Parameter nm_filter : pred nat -> tp -> tp. +Parameter empb : tp -> bool. +Parameter undefb : tp -> bool. +Parameter pts : nat -> A -> tp. + +Parameter from : tp -> @UM.base nat_ordType A (fun x => x != 0). +Parameter to : @UM.base nat_ordType A (fun x => x != 0) -> tp. + +Axiom ftE : forall b, from (to b) = b. +Axiom tfE : forall f, to (from f) = f. +Axiom undefE : nm_undef = to (@UM.Undef nat_ordType A (fun x => x != 0)). +Axiom defE : forall f, defined f = UM.valid (from f). +Axiom emptyE : empty = to (@UM.empty nat_ordType A (fun x => x != 0)). +Axiom updE : forall k v f, upd k v f = to (UM.upd k v (from f)). +Axiom domE : forall f, dom f = UM.dom (from f). +Axiom dom_eqE : forall f1 f2, dom_eq f1 f2 = UM.dom_eq (from f1) (from f2). +Axiom freeE : forall k f, free k f = to (UM.free k (from f)). +Axiom findE : forall k f, find k f = UM.find k (from f). +Axiom unionE : forall f1 f2, union f1 f2 = to (UM.union (from f1) (from f2)). +Axiom nmfiltE : forall q f, nm_filter q f = to (UM.um_filter q (from f)). +Axiom empbE : forall f, empb f = UM.empb (from f). +Axiom undefbE : forall f, undefb f = UM.undefb (from f). +Axiom ptsE : forall k v, pts k v = to (@UM.pts nat_ordType A (fun x => x != 0) k v). + +End Params. +End NMSig. + + +Module NMDef : NMSig. +Section NMDef. +Variable A : Type. + +Definition nonz x := x != 0. + +Definition tp : Type := @UM.base nat_ordType A nonz. + +Definition nm_undef := @UM.Undef nat_ordType A nonz. +Definition defined f := @UM.valid nat_ordType A nonz f. +Definition empty := @UM.empty nat_ordType A nonz. +Definition upd k v f := @UM.upd nat_ordType A nonz k v f. +Definition dom f := @UM.dom nat_ordType A nonz f. +Definition dom_eq f1 f2 := @UM.dom_eq nat_ordType A nonz f1 f2. +Definition free k f := @UM.free nat_ordType A nonz k f. +Definition find k f := @UM.find nat_ordType A nonz k f. +Definition union f1 f2 := @UM.union nat_ordType A nonz f1 f2. +Definition nm_filter q f := @UM.um_filter nat_ordType A nonz q f. +Definition empb f := @UM.empb nat_ordType A nonz f. +Definition undefb f := @UM.undefb nat_ordType A nonz f. +Definition pts k v := @UM.pts nat_ordType A nonz k v. + +Definition from (f : tp) : @UM.base nat_ordType A nonz := f. +Definition to (b : @UM.base nat_ordType A nonz) : tp := b. + +Lemma ftE b : from (to b) = b. Proof. by []. Qed. +Lemma tfE f : to (from f) = f. Proof. by []. Qed. +Lemma undefE : nm_undef = to (@UM.Undef nat_ordType A nonz). Proof. by []. Qed. +Lemma defE f : defined f = UM.valid (from f). Proof. by []. Qed. +Lemma emptyE : empty = to (@UM.empty nat_ordType A nonz). Proof. by []. Qed. +Lemma updE k v f : upd k v f = to (UM.upd k v (from f)). Proof. by []. Qed. +Lemma domE f : dom f = UM.dom (from f). Proof. by []. Qed. +Lemma dom_eqE f1 f2 : dom_eq f1 f2 = UM.dom_eq (from f1) (from f2). +Proof. by []. Qed. +Lemma freeE k f : free k f = to (UM.free k (from f)). Proof. by []. Qed. +Lemma findE k f : find k f = UM.find k (from f). Proof. by []. Qed. +Lemma unionE f1 f2 : union f1 f2 = to (UM.union (from f1) (from f2)). +Proof. by []. Qed. +Lemma nmfiltE q f : nm_filter q f = to (UM.um_filter q (from f)). +Proof. by []. Qed. +Lemma empbE f : empb f = UM.empb (from f). Proof. by []. Qed. +Lemma undefbE f : undefb f = UM.undefb (from f). Proof. by []. Qed. +Lemma ptsE k v : pts k v = to (@UM.pts nat_ordType A nonz k v). +Proof. by []. Qed. + +End NMDef. +End NMDef. + +Notation natmap := NMDef.tp. + +Definition natmapMixin A := + UMCMixin (@NMDef.ftE A) (@NMDef.tfE A) (@NMDef.defE A) + (@NMDef.undefE A) (@NMDef.emptyE A) (@NMDef.updE A) + (@NMDef.domE A) (@NMDef.dom_eqE A) (@NMDef.freeE A) + (@NMDef.findE A) (@NMDef.unionE A) (@NMDef.nmfiltE A) + (@NMDef.empbE A) (@NMDef.undefbE A) (@NMDef.ptsE A). + +Canonical nat_mapUMC A := + Eval hnf in UMC (natmap A) (@natmapMixin A). + +(* we add the canonical projections matching against naked type *) +(* thus, there are two ways to get a PCM for a union map: *) +(* generic one going through union_map_classPCM, and another by going *) +(* through union_mapPCM. Luckily, they produce equal results *) +(* and this is just a matter of convenience *) +(* Ditto for the equality type *) + +Definition nat_mapPCMMix A := union_map_classPCMMix (nat_mapUMC A). +Canonical nat_mapPCM A := Eval hnf in PCM (natmap A) (@nat_mapPCMMix A). + +Definition nat_mapCPCMMix A := union_map_classCPCMMix (nat_mapUMC A). +Canonical nat_mapCPCM A := Eval hnf in CPCM (natmap A) (@nat_mapCPCMMix A). + +Definition nat_mapTPCMMix A := union_map_classTPCMMix (nat_mapUMC A). +Canonical nat_mapTPCM A := Eval hnf in TPCM (natmap A) (@nat_mapTPCMMix A). + +Definition nat_map_eqMix (A : eqType) := + @union_map_class_eqMix nat_ordType A _ _ (@natmapMixin A). +Canonical nat_map_eqType (A : eqType) := + Eval hnf in EqType (natmap A) (@nat_map_eqMix A). + +Definition nm_pts A (k : nat) (v : A) := + @UMC.pts nat_ordType A (@nat_mapUMC A) k v. + +Notation "x \-> v" := (@nm_pts _ x v) (at level 30). + + +Lemma nm_dom0 A (h : natmap A) : (h = um_undef \/ h = Unit) <-> (dom h = [::]). +Proof. +split=>[|E]; first by case=>->; rewrite ?dom_undef ?dom0. +case V : (valid h); last by move/invalidE: (negbT V)=>->; left. +by rewrite (dom0E V) ?E //; right. +Qed. + +(***************************************) +(***************************************) +(* Constructions of last_key and fresh *) +(***************************************) +(***************************************) + +Section FreshLastKey. +Variable A : Type. +Implicit Type h : natmap A. + +Definition last_key h := last 0 (dom h). +Definition fresh h := (last_key h).+1. + +Lemma lastkey_undef : last_key um_undef = 0. +Proof. by rewrite /last_key !umEX. Qed. + +Lemma lastkey0 : last_key Unit = 0. +Proof. by rewrite /last_key /Unit /= !umEX. Qed. + +Lemma lastkey_dom h : valid h -> last_key h \notin dom h -> h = Unit. +Proof. +rewrite /valid /= /last_key /Unit /= !umEX /= -{4}[h]UMC.tfE. +case: (UMC.from h)=>//=; case=>s H /= H1 _ /seq_last_in. +rewrite /UM.empty UMC.eqE UM.umapE /supp fmapE /= {H H1}. +by elim: s. +Qed. + +Lemma dom_lastkey0P h : last_key h = 0 <-> dom h = [::]. +Proof. +rewrite -nm_dom0; split; last first. +- by case=>E; subst h; rewrite ?lastkey_undef ?lastkey0. +move=>E; case V : (valid h); last by left; move/invalidE: (negbT V). +right; apply: lastkey_dom=>//; rewrite E. +by apply: contraT; rewrite negbK; apply: dom_cond. +Qed. + +Lemma dom_lastkey h : valid h -> ~~ empb h -> last_key h \in dom h. +Proof. by move=>X; apply: contraR; move/(lastkey_dom X)=>->; apply/empbP. Qed. + +Lemma lastkeyxx h : valid h -> last_key h = 0 -> h = Unit. +Proof. +by move=>V H; apply: lastkey_dom V _; apply/negP=>/dom_cond; rewrite H. +Qed. + +Lemma dom_lastkeyE h a : a < last_key h -> last_key h \in dom h. +Proof. +move=>H; case V : (valid h); last first. +- by move/invalidE: (negbT V) H=>->; rewrite lastkey_undef. +by apply: dom_lastkey V (contraL _ H)=>/empbE ->; rewrite lastkey0. +Qed. + +Lemma lastkey_max h x : x \in dom h -> x <= last_key h. +Proof. +rewrite /last_key /= !umEX; case: (UMC.from h)=>//; case=>s H _ /=. +rewrite /supp /ord /= (leq_eqVlt x) orbC. +by apply: sorted_last_key_max (sorted_oleq H). +Qed. + +Lemma max_lastkey h x : + x \in dom h -> {in dom h, forall y, y <= x} -> last_key h = x. +Proof. +rewrite /last_key !umEX; case: (UMC.from h)=>//; case=>s H _ /=. +move=>H1 /= H2; apply: sorted_max_key_last (sorted_oleq H) H1 _. +by move=>z /(H2 z); rewrite leq_eqVlt orbC. +Qed. + +Lemma lastkeyPt (x : nat) v : x != 0 -> last_key (x \-> v) = x. +Proof. by rewrite /last_key domPtK /= => ->. Qed. + +Lemma hist_path h : path oleq 0 (dom h). +Proof. +rewrite !umEX; case: (UMC.from h)=>// {h} h /= _; case: h; case=>//= x s H. +rewrite {1}/oleq /ord /= orbC -leq_eqVlt /=. +by apply: sub_path H=>z y; rewrite /oleq=>->. +Qed. + +Lemma lastkey_mono h1 h2 : + {subset dom h1 <= dom h2} -> last_key h1 <= last_key h2. +Proof. by rewrite leq_eqVlt orbC; apply: seq_last_mono; apply: hist_path. Qed. + +Lemma lastkeyfUn h1 h2 : + valid (h1 \+ h2) -> last_key h1 <= last_key (h1 \+ h2). +Proof. by move=>X; apply: lastkey_mono=>x; rewrite domUn inE X => ->. Qed. + +Lemma lastkeyUnf h1 h2 : + valid (h1 \+ h2) -> last_key h2 <= last_key (h1 \+ h2). +Proof. by rewrite joinC; apply: lastkeyfUn. Qed. + +(* a convenient rephrasing of above lemmas *) + +Lemma lastkeyUn_mono h1 h2 t : + valid (h1 \+ h2) -> last_key (h1 \+ h2) < t -> last_key h1 < t. +Proof. +move=>V; apply/leq_trans/lastkey_mono. +by move=>x D; rewrite domUn inE V D. +Qed. + +Lemma lastkeyUn0 h1 h2 : + valid (h1 \+ h2) -> + last_key h1 = last_key h2 -> (h1 = Unit) * (h2 = Unit). +Proof. +move=>V E. +case E1: (empb h1). +- move/empbE: E1 E V=>->; rewrite unitL lastkey0. + by case/esym/dom_lastkey0P/nm_dom0=>-> //; rewrite valid_undef. +case E2: (empb h2). +- move/empbE: E2 E V=>->; rewrite unitR lastkey0. + by case/dom_lastkey0P/nm_dom0=>-> //; rewrite valid_undef. +case: validUn (V)=>// _ _ /(_ _ (dom_lastkey (validL V) (negbT E1))). +by rewrite E (dom_lastkey (validR V) (negbT E2)). +Qed. + +Lemma lastkeyUn h1 h2 : + last_key (h1 \+ h2) = + if valid (h1 \+ h2) then maxn (last_key h1) (last_key h2) else 0. +Proof. +have H k1 k2 : valid (k1 \+ k2) -> + last_key k1 < last_key k2 -> last_key (k1 \+ k2) = last_key k2. +- move=>V H; apply: max_lastkey=>[|x]. + - by rewrite domUn inE V (dom_lastkeyE H) orbT. + rewrite domUn inE V; case/orP; move/lastkey_max=>// T; + by apply: leq_trans T (ltnW H). +case V : (valid _); last first. +- by move/invalidE: (negbT V)=>->; rewrite lastkey_undef. +rewrite /maxn; case: ltngtP. +- by rewrite joinC in V *; apply: H. +- by apply: H. +by move/esym/(lastkeyUn0 V)=>E; rewrite !E unitL. +Qed. + +Lemma lastkeyPtUn h t u : + valid h -> last_key h < t -> valid (t \-> u \+ h). +Proof. +move=>V L; rewrite validPtUn; apply/and3P; split=>//=. +- by rewrite -lt0n; apply: leq_ltn_trans L. +by apply: contraL L; move/lastkey_max; case: leqP. +Qed. + +(* freshness *) + +Lemma dom_ordfresh h x : x \in dom h -> x < fresh h. +Proof. by move/lastkey_max. Qed. + +Lemma dom_freshn h n : fresh h + n \notin dom h. +Proof. by apply: contra (@dom_ordfresh _ _) _; rewrite -leqNgt leq_addr. Qed. + +Lemma dom_freshUn h1 h2 : valid h1 -> [pcm h2 <= h1] -> fresh h1 \notin dom h2. +Proof. +move=>V [x E]; rewrite {h1}E in V *; apply: contra (@dom_ordfresh _ _) _. +by rewrite joinC in V *; rewrite -leqNgt; apply: lastkeyUnf. +Qed. + +Lemma dom_fresh h : fresh h \notin dom h. +Proof. by move: (dom_freshn h 0); rewrite addn0. Qed. + +Lemma valid_freshUn h1 h2 v : + valid h1 -> [pcm h2 <= h1] -> valid (fresh h1 \-> v \+ h2) = valid h2. +Proof. +move=>V [x E]; rewrite {h1}E in V *. +by rewrite validPtUn dom_freshUn // andbT. +Qed. + +Lemma valid_fresh h v : valid (fresh h \-> v \+ h) = valid h. +Proof. by rewrite validPtUn dom_fresh andbT. Qed. + +Lemma lastkey_freshUn h1 h2 v : + valid h1 -> [pcm h2 <= h1] -> + last_key (fresh h1 \-> v \+ h2) = fresh h1. +Proof. +move=>V [x E]; rewrite {h1}E in V *. +apply: max_lastkey=>[|y] /=. +- by rewrite domUn inE valid_freshUn // (validL V) domPt inE eq_refl. +rewrite domUn inE valid_freshUn // (validL V) /= domPt inE /= eq_sym. +rewrite leq_eqVlt; case: eqP=>//= _ D. +by apply: lastkey_max; rewrite domUn inE V D. +Qed. + +Lemma lastkey_fresh h v : valid h -> last_key (fresh h \-> v \+ h) = fresh h. +Proof. +move=>Vf; apply: max_lastkey=>[|x] /=. +- by rewrite domUn inE valid_fresh Vf domPt inE eq_refl. +rewrite domUn inE /= valid_fresh Vf /= domPt inE /= eq_sym. +by rewrite leq_eqVlt; case: eqP=>//= _; apply: dom_ordfresh. +Qed. + +(* pcm induced ordering *) + +Lemma umpleq_lastkey h1 h2 : + valid h2 -> [pcm h1 <= h2] -> last_key h1 <= last_key h2. +Proof. +by move=>V H; case: H V=>z->V; apply: lastkey_mono=>k D; rewrite domUn inE V D. +Qed. + +(* backward induction on valid natmaps *) + +Lemma valid_indb (P : natmap A -> Prop) : + P Unit -> + (forall u, P (1 \-> u)) -> + (forall t u h, P h -> last_key h < t -> + valid (t \-> u \+ h) -> P (t \-> u \+ h)) -> + forall h, valid h -> P h. +Proof. +move=>H1 H2 H3; elim/um_indb=>//=. +- by rewrite -[valid _]negbK; move/negP/invalidE. +move=>k v f H V K _. +case E: (empb f); last first. +- apply: H3=>//; first by apply: H (validR V). + apply: K; apply: contraR (negbT E). + by rewrite -empbE; apply: lastkey_dom (validR V). +move/empbE: {K H} E V=>->; rewrite unitR => V. +move: (H3 k v Unit H1); rewrite unitR lastkey0 lt0n. +by apply=>//; rewrite validPt /= in V. +Qed. + +(* forward induction on valid natmaps *) + +Lemma valid_indf (P : natmap A -> Prop) : + P Unit -> + (forall t u h, P h -> + (forall x, x \in dom h -> t < x) -> + valid (t \-> u \+ h) -> P (t \-> u \+ h)) -> + forall h, valid h -> P h. +Proof. +move=>H1 H2; elim/um_indf=>//=. +- by rewrite -[valid _]negbK; move/negP/invalidE. +move=>k v f H V K _. +apply: H2=>//; first by apply: H (validR V). +move=>x; move/(order_path_min (@ordtype.trans _)): K. +by case: allP=>// X _; apply: X. +Qed. + +End FreshLastKey. + + +(*******************) +(*******************) +(* Gapless natmaps *) +(*******************) +(*******************) + +Section Gapless. +Variable A : Type. +Implicit Type h : natmap A. + +Definition gapless h := forall k, 0 < k <= last_key h -> k \in dom h. + +Lemma gp_undef : gapless um_undef. +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 (fresh h \-> u \+ h) <-> gapless h. +Proof. +case V : (valid h); last first. +- by move/invalidE: (negbT V)=>->; rewrite join_undefR. +split=>H k; move: (V); rewrite -(valid_fresh _ u)=>V'; last first. +- rewrite lastkey_fresh // domPtUn inE V' /= (leq_eqVlt k) eq_sym. + by move: (H k); rewrite -(ltnS k); case: ltngtP. +rewrite -(ltnS k) -/(fresh h); case/andP=>Z N. +move: (H k); rewrite lastkey_fresh // domPtUn inE V' Z /= leq_eqVlt eq_sym. +by case: ltngtP N=>//= _ _; apply. +Qed. + +Lemma gpPtUn h k u : + valid (k \-> u \+ h) -> + gapless (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 (k \-> u \+ h) = k. +- move: (C (fresh h)); rewrite Z (leq_eqVlt _ k) Y orbT andbT => /(_ erefl). + rewrite domPtUn inE (negbTE X) /=; case/andP=>_ /dom_ordfresh. + by rewrite ltnn. +apply: max_lastkey (find_some (findPtUn V)) _ => x. +rewrite domUn inE; case/andP=>_ /orP []. +- by rewrite domPt inE; case/andP=>_ /eqP ->. +by move/lastkey_max/leq_ltn_trans/(_ N)/ltnW. +Qed. +End Gapless. + +Arguments gp_fresh [A][h] u. + + +(*****************************) +(*****************************) +(* Natmaps with binary range *) +(*****************************) +(*****************************) + +Section AtVal. +Variable A : Type. +Implicit Type h : natmap (A * A). + +Definition atval v t h := oapp snd v (find t h). + +Lemma umpleq_atval v t h1 h2 : + gapless h1 -> valid h2 -> [pcm h1 <= h2] -> t <= last_key h1 -> + atval v t h2 = atval v t h1. +Proof. +move=>G V H K; rewrite /atval; case E1 : (find t h1)=>[a|]. +- by rewrite (umpleq_some V H E1). +case: t K E1 => [|t] K E1; last by move: (G t.+1 K) (find_none E1)=>->. +by case E2 : (find 0 h2)=>//; move/find_some/dom_cond: E2. +Qed. + +Definition last_val v h := atval v (last_key h) h. + +Lemma lastval0 v : last_val v Unit = v. +Proof. by rewrite /last_val /atval lastkey0 find0E. Qed. + +Lemma lastvalPt v p x : p != 0 -> last_val v (p \-> x) = x.2. +Proof. +by move=>V; rewrite /last_val /atval lastkeyPt // findPt /= V. +Qed. + +Lemma lastval_fresh v x h : + valid h -> last_val v (fresh h \-> x \+ h) = x.2. +Proof. +by move=>V; rewrite /last_val /atval lastkey_fresh // findPtUn // valid_fresh. +Qed. + +Lemma lastvalUn v h1 h2 : + last_val v (h1 \+ h2) = + if valid (h1 \+ h2) then + if last_key h1 < last_key h2 then last_val v h2 else last_val v h1 + else v. +Proof. +rewrite /last_val /atval lastkeyUn maxnC /maxn. +case: ifP; last by move/negbT/invalidE=>->; rewrite find_undef. +case: ltngtP=>N V. +- by rewrite findUnR // (dom_lastkeyE N). +- by rewrite findUnL // (dom_lastkeyE N). +by rewrite !(lastkeyUn0 V N) unitL lastkey0 find0E. +Qed. + +Lemma lastval_freshUn v x h1 h2 : + valid h1 -> [pcm h2 <= h1] -> + last_val v (fresh h1 \-> x \+ h2) = x.2. +Proof. +move=>V H; rewrite /last_val /atval. +rewrite lastkey_freshUn // findPtUn // valid_freshUn //. +by case: H V=>h -> /validL. +Qed. + +End AtVal. + +(* Continuous maps with binary range *) + +Section Continuous. +Variable A : Type. +Implicit Type h : natmap (A * A). + +Definition continuous v h := + forall k x, find k.+1 h = Some x -> oapp snd v (find k h) = x.1. + +Lemma cn_undef v : continuous v um_undef. +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 (fresh h \-> x \+ h) <-> + continuous v h /\ last_val v h = x.1. +Proof. +rewrite -(valid_fresh _ x)=>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_ordfresh /(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 // eq_refl ltn_eqF //; apply. +move=>k w; case: (ltnP k (last_key h))=>N; last first. +- by move/find_some /dom_ordfresh /(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 (k.+1 \-> (x, y) \+ h) -> continuous v (k.+1 \-> (x, y) \+ h) -> + oapp snd v (find k h) = x. +Proof. +by move=>V /(_ k (x, y)); rewrite !findPtUn2 // eq_refl ltn_eqF //; apply. +Qed. + +End Continuous. + +Arguments cn_fresh [A][v][h][x] _. + +(* Complete natmaps with binary range *) + +Section Complete. +Variable A : Type. +Implicit Type h : natmap (A * A). + +Definition complete v0 h vn := + [/\ valid h, gapless h, continuous v0 h & last_val 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 lastval0]. Qed. + +Lemma cm_fresh v0 vn h x : + complete v0 (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 lastval_fresh. +case=>-> [V] /(gp_fresh x) G C E; split=>//; +by [rewrite valid_fresh | apply/(cn_fresh V) | rewrite lastval_fresh]. +Qed. + +Lemma cmPtUn v0 vn h k x : + complete v0 (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 (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_val lastkeyPt // /atval findPt /= => <-; case: x. +Qed. + +Lemma cmCons v0 vn k x h : + complete v0 (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. + diff --git a/pcm/pcm.v b/pcm/pcm.v new file mode 100644 index 0000000..7f11c9c --- /dev/null +++ b/pcm/pcm.v @@ -0,0 +1,640 @@ +(* +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 Coq Require Import ssreflect ssrbool ssrfun. +From mathcomp Require Import ssrnat bigop. +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Delimit Scope pcm_scope with pcm. +Open Scope pcm_scope. + +(*******************************) +(* Partial Commutative Monoids *) +(*******************************) + +Module PCM. + +Record mixin_of (T : Type) := Mixin { + valid_op : T -> bool; + join_op : T -> T -> T; + unit_op : T; + _ : commutative join_op; + _ : associative join_op; + _ : left_id unit_op join_op; + (* monotonicity of valid *) + _ : forall x y, valid_op (join_op x y) -> valid_op x; + (* unit is valid *) + _ : valid_op unit_op}. + +Section ClassDef. + +Notation class_of := mixin_of (only parsing). + +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 pack c := @Pack T c. +Definition clone := fun c & cT -> T & phant_id (pack c) cT => pack c. + +Definition valid := valid_op class. +Definition join := join_op class. +Definition unit := unit_op class. + +End ClassDef. + +Arguments unit [cT]. + +Module Exports. +Coercion sort : type >-> Sortclass. +Notation pcm := type. +Notation PCMMixin := Mixin. +Notation PCM T m := (@pack T m). + +Notation "[ 'pcmMixin' 'of' T ]" := (class _ : mixin_of T) + (at level 0, format "[ 'pcmMixin' 'of' T ]") : pcm_scope. +Notation "[ 'pcm' 'of' T 'for' C ]" := (@clone T C _ idfun id) + (at level 0, format "[ 'pcm' 'of' T 'for' C ]") : pcm_scope. +Notation "[ 'pcm' 'of' T ]" := (@clone T _ _ id id) + (at level 0, format "[ 'pcm' 'of' T ]") : pcm_scope. + +Notation "x \+ y" := (join x y) + (at level 43, left associativity) : pcm_scope. +Notation valid := valid. +Notation Unit := unit. + +Arguments unit [cT]. +Prenex Implicits valid join unit. + +(* Restating the laws, with the notation. *) +(* Plus some additional derived lemmas. *) + +Section Laws. +Variable U V : pcm. + +Lemma joinC (x y : U) : x \+ y = y \+ x. +Proof. by case: U x y=>tp [v j z Cj *]; apply: Cj. Qed. + +Lemma joinA (x y z : U) : x \+ (y \+ z) = x \+ y \+ z. +Proof. by case: U x y z=>tp [v j z Cj Aj *]; apply: Aj. Qed. + +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 validL (x y : U) : valid (x \+ y) -> valid x. +Proof. case: U x y=>tp [v j z Cj Aj Uj /= Mj inv f ?]; apply: Mj. Qed. + +Lemma validR (x y : U) : valid (x \+ y) -> valid y. +Proof. by rewrite joinC; apply: validL. Qed. + +Lemma validE (x y : U) : valid (x \+ y) -> valid x * valid y. +Proof. by move=>X; rewrite (validL X) (validR X). Qed. + +Lemma unitL (x : U) : Unit \+ x = x. +Proof. by case: U x=>tp [v j z Cj Aj Uj *]; apply: Uj. Qed. + +Lemma unitR (x : U) : x \+ Unit = x. +Proof. by rewrite joinC unitL. Qed. + +Lemma valid_unit : valid (@Unit U). +Proof. by case: U=>tp [v j z Cj Aj Uj Vm Vu *]. Qed. + +(* 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. + +End Laws. + +Hint Resolve valid_unit. + +Section UnfoldingRules. +Variable U : pcm. + +Lemma pcm_joinE (x y : U) : x \+ y = join_op (class U) x y. +Proof. by []. Qed. + +Lemma pcm_validE (x : U) : valid x = valid_op (class U) x. +Proof. by []. Qed. + +Lemma pcm_unitE : unit = unit_op (class U). +Proof. by []. Qed. + +Definition pcmE := (pcm_joinE, pcm_validE, pcm_unitE). + +(* also a simple rearrangment equation *) +Definition pull (x y : U) := (joinC y x, joinCA y x). + +End UnfoldingRules. + +End Exports. + +End PCM. + +Export PCM.Exports. + +(*********************) +(* Cancellative PCMs *) +(*********************) + +(* definition of precision for an 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. + +Module CancellativePCM. + +Variant mixin_of (U : pcm) := Mixin of + forall x1 x2 x : U, valid (x1 \+ x) -> x1 \+ x = x2 \+ x -> x1 = x2. + +Section ClassDef. + +Record class_of (U : Type) := Class { + base : PCM.mixin_of U; + mixin : mixin_of (PCM.Pack base)}. + +Local Coercion base : class_of >-> PCM.mixin_of. + +Structure type : Type := Pack {sort : Type; _ : class_of sort}. +Local Coercion sort : type >-> Sortclass. + +Variables (T : Type) (cT : type). +Definition class := let: Pack _ c as cT' := cT return class_of cT' in c. +Definition clone c of phant_id class c := @Pack T c. + +(* produce a cancellative type out of the mixin *) +(* equalize m0 and m by means of a phantom *) +Definition pack b0 (m0 : mixin_of (PCM.Pack b0)) := + fun m & phant_id m0 m => Pack (@Class T b0 m). + +Definition pcm := PCM.Pack class. + +End ClassDef. + +Module Exports. +Coercion base : class_of >-> PCM.mixin_of. +Coercion sort : type >-> Sortclass. +Coercion pcm : type >-> PCM.type. +Canonical Structure pcm. + +Notation cpcm := type. +Notation CPCMMixin := Mixin. +Notation CPCM T m := (@pack T _ _ m id). + +Notation "[ 'cpcm' 'of' T 'for' cT ]" := (@clone T cT _ id) + (at level 0, format "[ 'cpcm' 'of' T 'for' cT ]") : pcm_scope. +Notation "[ 'cpcm' 'of' T ]" := (@clone T _ _ id) + (at level 0, format "[ 'cpcm' 'of' T ]") : pcm_scope. + +Section Lemmas. +Variable U : cpcm. + +Lemma joinKx (x1 x2 x : U) : valid (x1 \+ x) -> x1 \+ x = x2 \+ x -> x1 = x2. +Proof. by case: U x1 x2 x=>V [b][K] T; apply: K. Qed. + +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 Lemmas. +End Exports. + +End CancellativePCM. + +Export CancellativePCM.Exports. + +(***************) +(* Topped PCMs *) +(***************) + +(* PCM with an explicit undef value *) +(* we will want these constants to be decidable *) +(* so we also add unitb, to test if an element is unit *) +(* for undefb, that should be valid, so we don't add anything special *) +(* OTOH, unit should not be decomposable *) + +Module TPCM. + +Record mixin_of (U : pcm) := Mixin { + undef_op : U; + unitb_op : U -> bool; + _ : forall x, reflect (x = Unit) (unitb_op x); + (* next one doesn't hold of max PCM, so remove eventually *) + _ : forall x y : U, x \+ y = Unit -> x = Unit /\ y = Unit; +(* _ : forall x y z : U, valid (x \+ y \+ z) = + [&& valid (x \+ y), valid (y \+ z) & valid (x \+ z)]; *) + _ : ~~ valid undef_op; + _ : forall x, undef_op \+ x = undef_op}. + +Section ClassDef. + +Record class_of (U : Type) := Class { + base : PCM.mixin_of U; + mixin : mixin_of (PCM.Pack base)}. + +Local Coercion base : class_of >-> PCM.mixin_of. + +Structure type : Type := Pack {sort : Type; _ : class_of sort}. +Local Coercion sort : type >-> Sortclass. + +Variables (T : Type) (cT : type). +Definition class := let: Pack _ c as cT' := cT return class_of cT' in c. +Definition clone c of phant_id class c := @Pack T c. + +(* produce a topped pcm out of the mixin *) +(* equalize m0 and m by means of a phantom *) +Definition pack b0 (m0 : mixin_of (PCM.Pack b0)) := + fun m & phant_id m0 m => Pack (@Class T b0 m). + +Definition pcm := PCM.Pack class. +Definition unitb := unitb_op (mixin class). +Definition undef : pcm := undef_op (mixin class). + +End ClassDef. + +Module Exports. +Coercion base : class_of >-> PCM.mixin_of. +Coercion sort : type >-> Sortclass. +Coercion pcm : type >-> PCM.type. +Canonical Structure pcm. +Notation tpcm := type. +Notation TPCMMixin := Mixin. +Notation TPCM T m := (@pack T _ _ m id). + +Notation "[ 'tpcm' 'of' T 'for' cT ]" := (@clone T cT _ id) + (at level 0, format "[ 'tpcm' 'of' T 'for' cT ]") : pcm_scope. +Notation "[ 'tpcm' 'of' T ]" := (@clone T _ _ id) + (at level 0, format "[ 'tpcm' 'of' T ]") : pcm_scope. + +Notation undef := undef. +Notation unitb := unitb. +Arguments undef [cT]. +Prenex Implicits undef. + +Section Lemmas. +Variable U : tpcm. + +Lemma unitbP (x : U) : reflect (x = Unit) (unitb x). +Proof. by case: U x=>V [b][u]. Qed. + +Lemma unitbE : unitb (Unit : U). +Proof. by case: unitbP. Qed. + +Lemma joinE0 (x y : U) : x \+ y = Unit <-> (x = Unit) * (y = Unit). +Proof. +case: U x y=>V [b][u1 u2] H1 H2 H3 T x y; split; first by apply: H2. +by case=>->->; rewrite unitL. +Qed. + +Lemma valid_undefN : ~~ valid (@undef U). +Proof. by case: U=>V [b][u]. Qed. + +Lemma valid_undef : valid (@undef U) = false. +Proof. by rewrite (negbTE valid_undefN). Qed. + +Lemma undef_join (x : U) : undef \+ x = undef. +Proof. by case: U x=>V [b][u]. 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. + +Definition undefE := (undef_join, join_undef, valid_undef). + +End Lemmas. +End Exports. + +End TPCM. + +Export TPCM.Exports. + + +(***************************************) +(* Support for big operators over PCMs *) +(***************************************) + +Canonical pcm_monoid (U : pcm) := Monoid.Law (@joinA U) (@unitL U) (@unitR U). +Canonical pcm_comoid (U : pcm) := Monoid.ComLaw (@joinC U). + +Section BigPartialMorph. +Variables (R1 : Type) (R2 : pcm) (K : R1 -> R2 -> Type) (f : R2 -> R1). +Variables (id1 : R1) (op1 : R1 -> R1 -> R1). +Hypotheses + (f_op : forall x y : R2, valid (x \+ y) -> f (x \+ y) = op1 (f x) (f y)) + (f_id : f Unit = id1). + +Lemma big_pmorph I r (P : pred I) F : + valid (\big[PCM.join/Unit]_(i <- r | P i) F i) -> + f (\big[PCM.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 *) +Definition natPCMMix := + PCMMixin addnC addnA add0n (fun x y => @id true) (erefl _). +Canonical natPCM := Eval hnf in PCM nat natPCMMix. + +(* also with multiplication, but we don't make that one canonical *) +Definition multPCMMix := + PCMMixin mulnC mulnA mul1n (fun x y => @id true) (erefl _). +Definition multPCM := Eval hnf in PCM nat multPCMMix. + +(* with max too *) +Definition maxPCMMix := + PCMMixin maxnC maxnA max0n (fun x y => @id true) (erefl _). +Definition maxPCM := Eval hnf in PCM nat maxPCMMix. + +(* bools with disjunction are a pcm *) +Definition bool_orPCMMix := + PCMMixin orbC orbA orFb (fun x y => @id true) (erefl _). +Definition bool_orPCM := Eval hnf in PCM bool bool_orPCMMix. + +(* cartesian product of pcms is a pcm *) + +Module ProdPCM. +Section ProdPCM. +Variables (U V : pcm). +Local Notation tp := (U * V)%type. + +Definition pvalid := [fun x : tp => valid x.1 && valid x.2]. +Definition pjoin := [fun x1 x2 : tp => (x1.1 \+ x2.1, x1.2 \+ x2.2)]. +Definition punit : tp := (Unit, Unit). + +Lemma joinC x y : pjoin x y = pjoin y x. +Proof. +move: x y => [x1 x2][y1 y2] /=. +by rewrite (joinC x1) (joinC x2). +Qed. + +Lemma joinA x y z : pjoin x (pjoin y z) = pjoin (pjoin x y) z. +Proof. +move: x y z => [x1 x2][y1 y2][z1 z2] /=. +by rewrite (joinA x1) (joinA x2). +Qed. + +Lemma validL x y : pvalid (pjoin x y) -> pvalid x. +Proof. +move: x y => [x1 x2][y1 y2] /=. +by case/andP=>D1 D2; rewrite (validL D1) (validL D2). +Qed. + +Lemma unitL x : pjoin punit x = x. +Proof. by case: x=>x1 x2; rewrite /= !unitL. Qed. + +Lemma validU : pvalid punit. +Proof. by rewrite /pvalid /= !valid_unit. Qed. + +End ProdPCM. +End ProdPCM. + +Definition prodPCMMixin U V := + PCMMixin (@ProdPCM.joinC U V) (@ProdPCM.joinA U V) + (@ProdPCM.unitL U V) (@ProdPCM.validL U V) + (@ProdPCM.validU U V). +Canonical prodPCM U V := Eval hnf in PCM (_ * _) (@prodPCMMixin U V). + +(* 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 []. Qed. + +Lemma pcm_peta (x : prodPCM U V) : x = (x.1, x.2). +Proof. by case: x. Qed. + +Lemma pcmPV (x : prodPCM U V) : valid x = valid x.1 && valid x.2. +Proof. by rewrite pcmE. Qed. + +Definition pcmPE := (pcmPJ, pcmPV). + +End Simplification. + +(* product of TPCMs is a TPCM *) +(* With TPCMs we could really remove the pairs *) +(* where one element is valid and the other is not *) +(* But let's not bother now *) + +Section ProdTPCM. +Variables (U V : tpcm). + +Lemma prod_unitb (x : prodPCM U V) : + reflect (x = Unit) (unitb x.1 && unitb x.2). +Proof. +case: x=>x1 x2; case: andP=>/= H; constructor. +- by case: H=>/unitbP -> /unitbP ->. +by case=>X1 X2; elim: H; rewrite X1 X2 !unitbE. +Qed. + +Lemma prod_join0E (x y : prodPCM U V) : x \+ y = Unit -> x = Unit /\ y = Unit. +Proof. by case: x y=>x1 x2 [y1 y2][] /joinE0 [->->] /joinE0 [->->]. Qed. + +(* +Lemma prod_valid3 (x y z : prodPCM U V) : valid (x \+ y \+ z) = + [&& valid (x \+ y), valid (y \+ z) & valid (x \+ z)]. +Proof. +case: x y z=>x1 x2 [y1 y2][z1 z2]; rewrite pcmE /= !valid3 -!andbA. +by do !bool_congr. +Qed. +*) + +Lemma prod_valid_undef : ~~ valid (@undef U, @undef V). +Proof. by rewrite pcmPV negb_and !valid_undef. Qed. + +Lemma prod_undef_join x : (@undef U, @undef V) \+ x = (@undef U, @undef V). +Proof. by case: x=>x1 x2; rewrite pcmPE !undef_join. Qed. + +Definition prodTPCMMix := TPCMMixin prod_unitb prod_join0E + prod_valid_undef prod_undef_join. +Canonical prodTPCM := Eval hnf in TPCM (U * V) prodTPCMMix. + +End ProdTPCM. + +(* unit is a pcm; just for kicks *) + +Module UnitPCM. +Section UnitPCM. + +Definition uvalid (x : unit) := true. +Definition ujoin (x y : unit) := tt. +Definition uunit := tt. + +Lemma ujoinC x y : ujoin x y = ujoin y x. +Proof. by []. Qed. + +Lemma ujoinA x y z : ujoin x (ujoin y z) = ujoin (ujoin x y) z. +Proof. by []. Qed. + +Lemma uvalidL x y : uvalid (ujoin x y) -> uvalid x. +Proof. by []. Qed. + +Lemma uunitL x : ujoin uunit x = x. +Proof. by case: x. Qed. + +Lemma uvalidU : uvalid uunit. +Proof. by []. Qed. + +End UnitPCM. +End UnitPCM. + +Definition unitPCMMixin := + PCMMixin UnitPCM.ujoinC UnitPCM.ujoinA + UnitPCM.uunitL UnitPCM.uvalidL UnitPCM.uvalidU. +Canonical unitPCM := Eval hnf in PCM unit unitPCMMixin. + +(* but it's not a TPCM, as there is no undefined element *) +(* we have to lift for that *) + + +(* bools with conjunction are a pcm *) + +Module BoolConjPCM. +Lemma unitL x : true && x = x. Proof. by []. Qed. +End BoolConjPCM. + +Definition boolPCMMixin := PCMMixin andbC andbA BoolConjPCM.unitL + (fun x y => @id true) (erefl _). +Canonical boolConjPCM := Eval hnf in PCM bool boolPCMMixin. + + +(*************************) +(* 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 ->; rewrite joinA; apply: validL. 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. + +Hint Resolve pleq_unit pleq_refl pleq_joinr pleq_joinl. + +(*******************) +(* Local functions *) +(*******************) + +Definition local (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 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 U (fun x y => Some x). +Proof. by move=>p x y _ V [<-]; rewrite -joinA. Qed. + diff --git a/pcm/pred.v b/pcm/pred.v new file mode 100644 index 0000000..62d3dd0 --- /dev/null +++ b/pcm/pred.v @@ -0,0 +1,562 @@ +(* +Copyright 2010 IMDEA Software Institute +Licensed under the Apache License, Version 2.0 (the "License"); +you may not use this file except in compliance with the License. +You may obtain a copy of the License at + http://www.apache.org/licenses/LICENSE-2.0 +Unless required by applicable law or agreed to in writing, software +distributed under the License is distributed on an "AS IS" BASIS, +WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. +See the License for the specific language governing permissions and +limitations under the License. +*) + +From Coq Require Import ssreflect ssrbool ssrfun Setoid. +From mathcomp Require Import seq. +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. +Set Warnings "-projection-no-head-constant". + +(******************************************************************************) +(* This file re-implements some of ssrbool's entities in Prop universe *) +(******************************************************************************) + +Lemma andTp p : True /\ p <-> p. Proof. by intuition. Qed. +Lemma andpT p : p /\ True <-> p. Proof. by intuition. Qed. +Lemma andFp p : False /\ p <-> False. Proof. by intuition. Qed. +Lemma andpF p : p /\ False <-> False. Proof. by intuition. Qed. +Lemma orTp p : True \/ p <-> True. Proof. by intuition. Qed. +Lemma orpT p : p \/ True <-> True. Proof. by intuition. Qed. +Lemma orFp p : False \/ p <-> p. Proof. by intuition. Qed. +Lemma orpF p : p \/ False <-> p. Proof. by intuition. Qed. + +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 *) + +CoInductive Mem_Pred : Type := MemProp of Pred T. +Definition isMem pT toPred mem := mem = (fun p : pT => MemProp [eta toPred p]). + +(* the general structure for predicates *) + +Structure PredType : Type := PropPredType { + Pred_Sort :> Type; + toPred : Pred_Sort -> Pred T; + _ : {mem | isMem toPred mem}}. + +Definition mkPredType pT toP := + PropPredType (exist (@isMem pT toP) _ (erefl _)). + +(* Pred, SimplPred, Mem_Pred, pred and simpl_pred are PredType's *) +Canonical Structure PredPredType := Eval hnf in @mkPredType (Pred T) id. +Canonical Structure SimplPredPredType := Eval hnf in mkPredType Pred_of_Simpl. +Coercion Pred_of_Mem mp : Pred_Sort PredPredType := + let: MemProp p := mp in [eta p]. +Canonical Structure MemPredType := Eval hnf in mkPredType Pred_of_Mem. +Canonical Structure predPredType := Eval hnf in @mkPredType (pred T) id. +Canonical Structure simplpredPredType := + Eval hnf in @mkPredType (simpl_pred T) (fun p x => p x). + +End Predicates. + +Arguments Pred0 [T]. +Arguments PredT [T]. +Prenex Implicits Pred0 PredT PredI PredU PredC PredD Preim. + +Notation "r1 +p r2" := (PredU r1 r2 : Pred _) + (at level 55, right associativity) : rel_scope. +Notation "r1 *p r2" := (xPredI r1 r2 : Pred _) + (at level 45, right associativity) : rel_scope. + +Notation "[ 'Pred' : T | E ]" := (SimplPred (fun _ : T => E)) + (at level 0, format "[ 'Pred' : T | E ]") : fun_scope. +Notation "[ 'Pred' x | E ]" := (SimplPred (fun x => E)) + (at level 0, x ident, format "[ 'Pred' x | E ]") : fun_scope. +Notation "[ 'Pred' x : T | E ]" := (SimplPred (fun x : T => E)) + (at level 0, x ident, only parsing) : fun_scope. +Notation "[ 'Pred' x y | E ]" := (SimplPred (fun t => let: (x, y) := t in E)) + (at level 0, x ident, y ident, format "[ 'Pred' x y | E ]") : fun_scope. +Notation "[ 'Pred' x y : T | E ]" := + (SimplPred (fun t : (T*T) => let: (x, y) := t in E)) + (at level 0, x ident, y ident, only parsing) : fun_scope. + +Definition repack_Pred T pT := + let: PropPredType _ a mP := pT return {type of @PropPredType T for pT} -> _ in + fun k => k a mP. + +Notation "[ 'PredType' 'of' T ]" := (repack_Pred (fun a => @PropPredType _ T a)) + (at level 0, format "[ 'PredType' 'of' T ]") : form_scope. + +Notation Pred_Class := (Pred_Sort (PredPredType _)). +Coercion Sort_of_Simpl_Pred T (p : Simpl_Pred T) : Pred_Class := p : Pred T. + +Definition PredArgType := Type. +Coercion Pred_of_argType (T : PredArgType) : Simpl_Pred T := PredT. + +Notation "{ :: T }" := (T%type : PredArgType) + (at level 0, format "{ :: T }") : type_scope. + +(* These must be defined outside a Section because "cooking" kills the *) +(* nosimpl tag. *) +Definition Mem T (pT : PredType T) : pT -> Mem_Pred T := + nosimpl (let: PropPredType _ _ (exist mem _) := pT return pT -> _ in mem). +Definition InMem T x mp := nosimpl Pred_of_Mem T mp x. + +Prenex Implicits Mem. + +(* Membership Predicates can be used as simple ones *) +Coercion Pred_of_Mem_Pred T mp := [Pred x : T | InMem x mp]. + +(* equality and subset *) + +Definition EqPredType T (pT : PredType T) (p1 p2 : pT) := + forall x : T, toPred p1 x <-> toPred p2 x. + +Definition SubPredType T (pT : PredType T) (p1 p2 : pT) := + forall x : T, toPred p1 x -> toPred p2 x. + +Definition EqSimplPred T (p1 p2 : Simpl_Pred T) := EqPredType p1 p2. +Definition SubSimplPred T (p1 p2 : Simpl_Pred T) := SubPredType p1 p2. + +Definition EqPredFun T1 T2 (pT2 : PredType T2) p1 p2 := + forall x : T1, @EqPredType T2 pT2 (p1 x) (p2 x). +Definition SubPredFun T1 T2 (pT2 : PredType T2) p1 p2 := + forall x : T1, @SubPredType T2 pT2 (p1 x) (p2 x). + +Definition EqMem T p1 p2 := forall x : T, InMem x p1 <-> InMem x p2. +Definition SubMem T p1 p2 := forall x : T, InMem x p1 -> InMem x p2. + +Notation "A <~> B" := (@EqPredType _ _ A B) + (at level 70, no associativity) : rel_scope. +Notation "A ~> B" := (@SubPredType _ _ A B) + (at level 70, no associativity) : rel_scope. +Notation "A <~1> B" := (@EqPredFun _ _ _ A B) + (at level 70, no associativity) : rel_scope. +Notation "A ~1> B" := (@SubPredFun _ _ _ A B) + (at level 70, no associativity) : rel_scope. + +Notation "x \In A" := (InMem x (Mem A)) + (at level 70, no associativity) : rel_scope. +Notation "x \Notin A" := (~ (x \In A)) + (at level 70, no associativity) : rel_scope. +Notation "A =p B" := (EqMem (Mem A) (Mem B)) + (at level 70, no associativity) : type_scope. +Notation "A <=p B" := (SubMem (Mem A) (Mem B)) + (at level 70, no associativity) : type_scope. + +(* Some notation for turning PredTypes into Pred or Simple Pred *) +Notation "[ 'Mem' A ]" := (Pred_of_Simpl (Pred_of_Mem_Pred (Mem A))) + (at level 0, only parsing) : fun_scope. +Notation "[ 'PredI' A & B ]" := (PredI [Mem A] [Mem B]) + (at level 0, format "[ 'PredI' A & B ]") : fun_scope. +Notation "[ 'PredU' A & B ]" := (PredU [Mem A] [Mem B]) + (at level 0, format "[ 'PredU' A & B ]") : fun_scope. +Notation "[ 'PredD' A & B ]" := (PredD [Mem A] [Mem B]) + (at level 0, format "[ 'PredD' A & B ]") : fun_scope. +Notation "[ 'PredC' A ]" := (PredC [Mem A]) + (at level 0, format "[ 'PredC' A ]") : fun_scope. +Notation "[ 'Preim' f 'of' A ]" := (Preim f [Mem A]) + (at level 0, format "[ 'Preim' f 'of' A ]") : fun_scope. + +Notation "[ 'Pred' x \In A ]" := [Pred x | x \In A] + (at level 0, x ident, format "[ 'Pred' x \In A ]") : fun_scope. +Notation "[ 'Pred' x \In A | E ]" := [Pred x | (x \In A) /\ E] + (at level 0, x ident, format "[ 'Pred' x \In A | E ]") : fun_scope. +Notation "[ 'Pred' x y \In A & B | E ]" := + [Pred x y | (x \In A) /\ (y \In B) /\ E] + (at level 0, x ident, y ident, + format "[ 'Pred' x y \In A & B | E ]") : fun_scope. +Notation "[ 'Pred' x y \In A & B ]" := [Pred x y | (x \In A) /\ (y \In B)] + (at level 0, x ident, y ident, + format "[ 'Pred' x y \In A & B ]") : fun_scope. +Notation "[ 'Pred' x y \In A | E ]" := [Pred x y \In A & A | E] + (at level 0, x ident, y ident, + format "[ 'Pred' x y \In A | E ]") : fun_scope. +Notation "[ 'Pred' x y \In A ]" := [Pred x y \In A & A] + (at level 0, x ident, y ident, + format "[ 'Pred' x y \In A ]") : fun_scope. + +Section Simplifications. +Variables (T : Type) (pT : PredType T). + +Lemma Mem_toPred : forall (p : pT), Mem (toPred p) = Mem p. +Proof. by rewrite /Mem; case: pT => T1 app1 [mem1 /= ->]. Qed. + +Lemma toPredE x (p : pT) : toPred p x = (x \In p). +Proof. by rewrite -Mem_toPred. Qed. + +Lemma In_Simpl x (p : Simpl_Pred T) : (x \In p) = p x. +Proof. by []. Qed. + +Lemma Simpl_PredE (p : Pred T) : p <~> [Pred x | p x]. +Proof. by []. Qed. + +Lemma Mem_Simpl (p : Simpl_Pred T) : Mem p = p :> Pred T. +Proof. by []. Qed. + +Definition MemE := Mem_Simpl. (* could be extended *) + +Lemma Mem_Mem (p : pT) : (Mem (Mem p) = Mem p) * (Mem [Mem p] = Mem p). +Proof. by rewrite -Mem_toPred. Qed. + +End Simplifications. + +(**************************************) +(* Definitions and lemmas for setoids *) +(**************************************) + +Section RelProperties. +Variables (T : Type) (pT : PredType T). + +Lemma EqPredType_refl (r : pT) : EqPredType r r. Proof. by []. Qed. +Lemma SubPredType_refl (r : pT) : SubPredType r r. Proof. by []. Qed. + +Lemma EqPredType_sym (r1 r2 : pT) : EqPredType r1 r2 -> EqPredType r2 r1. +Proof. by move=>H1 x; split; move/H1. Qed. + +Lemma EqPredType_trans' (r1 r2 r3 : pT) : + EqPredType r1 r2 -> EqPredType r2 r3 -> EqPredType r1 r3. +Proof. by move=>H1 H2 x; split; [move/H1; move/H2 | move/H2; move/H1]. Qed. + +Lemma SubPredType_trans' (r1 r2 r3 : pT) : + SubPredType r1 r2 -> SubPredType r2 r3 -> SubPredType r1 r3. +Proof. by move=>H1 H2 x; move/H1; move/H2. Qed. + +Definition EqPredType_trans r2 r1 r3 := @EqPredType_trans' r1 r2 r3. +Definition SubPredType_trans r2 r1 r3 := @SubPredType_trans' r1 r2 r3. +End RelProperties. + +Hint Resolve EqPredType_refl SubPredType_refl. + +(* Declaration of relations *) + +Add Parametric Relation T (pT : PredType T) : pT (@EqPredType _ pT) + reflexivity proved by (@EqPredType_refl _ _) + symmetry proved by (@EqPredType_sym _ _) + transitivity proved by (@EqPredType_trans' _ _) as EqPredType_rel. + +Add Parametric Relation T : (Simpl_Pred T) (@EqSimplPred _) + reflexivity proved by (@EqPredType_refl _ _) + symmetry proved by (@EqPredType_sym _ _) + transitivity proved by (@EqPredType_trans' _ _) as EqSimplPred_rel. + +Add Parametric Relation T : (Simpl_Pred T) (@SubSimplPred _) + reflexivity proved by (@SubPredType_refl _ _) + transitivity proved by (@SubPredType_trans' _ _) as SubSimplPred_rel. + +Add Parametric Relation T : (Mem_Pred T) (@EqMem T) + reflexivity proved by (@EqPredType_refl _ _) + symmetry proved by (@EqPredType_sym _ _) + transitivity proved by (@EqPredType_trans' _ _) as EqMem_rel. + +Add Parametric Relation T : (Mem_Pred T) (@SubMem _) + reflexivity proved by (@SubPredType_refl _ _) + transitivity proved by (@SubPredType_trans' _ _) as SubMem_rel. + +(* Declaring morphisms. *) + +Add Parametric Morphism T : (@Pred_of_Simpl T) with signature + @EqSimplPred _ ==> @EqPredType T (PredPredType T) as Pred_of_Simpl_morph. +Proof. by []. Qed. + +Add Parametric Morphism T (pT : PredType T) : (@EqPredType T pT) with signature + @EqPredType _ _ ==> @EqPredType _ _ ==> iff as EqPredType_morph. +Proof. by move=>r1 s1 H1 r2 s2 H2; rewrite H1 H2. Qed. + +Add Parametric Morphism T (pT : PredType T) : (@SubPredType T pT) with signature + @EqPredType _ _ ==> @EqPredType _ _ ==> iff as SubPred_morph. +Proof. by move=>r1 s1 H1 r2 s2 H2; split=>H x; move/H1; move/H; move/H2. Qed. + +Add Parametric Morphism T : (@InMem T) with signature + @eq _ ==> @EqMem _ ==> iff as InMem_morph. +Proof. by move=>x r s H; split; move/H. Qed. + +Add Parametric Morphism T (pT : PredType T) : (@Mem T pT) with signature + @EqPredType _ _ ==> @EqMem _ as Mem_morhp. +Proof. by move=>x y H p; rewrite /EqPredType -!toPredE in H *; rewrite H. Qed. + +Add Parametric Morphism T : (@PredU T) with signature + @EqPredType _ _ ==> @EqPredType _ _ ==> @EqSimplPred _ as predU_morph. +Proof. +move=>r1 s1 H1 r2 h2 H2 x; split; +by case; [move/H1 | move/H2]=>/=; auto. +Qed. + +Add Parametric Morphism T : (@PredI T) with signature + @EqPredType _ _ ==> @EqPredType _ _ ==> @EqPredType _ _ as predI_morph. +Proof. +move=>r1 s1 H1 r2 s2 H2 x; split; +by case; move/H1=>T1; move/H2=>T2. +Qed. + +Add Parametric Morphism T : (@PredC T) with signature + @EqPredType _ _ ==> @EqPredType _ _ as predC_morph. +Proof. by move=>r s H x; split=>H1; apply/H. Qed. + +Section RelLaws. +Variable (T : Type). + +Lemma orrI (r : Pred nat) : r +p r <~> r. +Proof. by move=>x; split; [case | left]. Qed. + +Lemma orrC (r1 r2 : Pred T) : r1 +p r2 <~> r2 +p r1. +Proof. move=>x; split=>/=; tauto. Qed. + +Lemma orr0 (r : Pred T) : r +p Pred0 <~> r. +Proof. by move=>x; split; [case | left]. Qed. + +Lemma or0r (r : Pred T) : Pred0 +p r <~> r. +Proof. by rewrite orrC orr0. Qed. + +Lemma orrCA (r1 r2 r3 : Pred T) : r1 +p r2 +p r3 <~> r2 +p r1 +p r3. +Proof. by move=>x; split=>/=; intuition. Qed. + +Lemma orrAC (r1 r2 r3 : Pred T) : (r1 +p r2) +p r3 <~> (r1 +p r3) +p r2. +Proof. by move=>?; split=>/=; intuition. Qed. + +Lemma orrA (r1 r2 r3 : Pred T) : (r1 +p r2) +p r3 <~> r1 +p r2 +p r3. +Proof. by rewrite (orrC r2) orrCA orrC. Qed. + +(* absorption *) +Lemma orrAb (r1 a : Pred T) : r1 <~> r1 +p a <-> a ~> r1. +Proof. +split; first by move=>-> x /=; auto. +move=>H x /=; split; first by auto. +by case=>//; move/H. +Qed. + +Lemma sub_orl (r1 r2 : Pred T) : r1 ~> r1 +p r2. Proof. by left. Qed. +Lemma sub_orr (r1 r2 : Pred T) : r2 ~> r1 +p r2. Proof. by right. Qed. + +End RelLaws. + +Section SubMemLaws. +Variable T : Type. + +Lemma subp_refl (p : Pred T) : p <=p p. +Proof. by []. Qed. + +Lemma subp_asym (p1 p2 : Pred T) : p1 <=p p2 -> p2 <=p p1 -> p1 =p p2. +Proof. by move=>H1 H2 x; split; [move/H1 | move/H2]. Qed. + +Lemma subp_trans (p2 p1 p3 : Pred T) : p1 <=p p2 -> p2 <=p p3 -> p1 <=p p3. +Proof. by move=>H1 H2 x; move/H1; move/H2. Qed. + +Lemma subp_or (p1 p2 q : Pred T) : p1 <=p q /\ p2 <=p q <-> p1 +p p2 <=p q. +Proof. +split=>[[H1] H2 x|H1]; first by case; [move/H1 | move/H2]. +by split=>x H2; apply: H1; [left | right]. +Qed. + +Lemma subp_and (p1 p2 q : Pred T) : q <=p p1 /\ q <=p p2 <-> q <=p p1 *p p2. +Proof. +split=>[[H1] H2 x|] H; last by split=>x; case/H. +by split; [apply: H1 | apply: H2]. +Qed. + +Lemma subp_orl (p1 p2 q : Pred T) : p1 <=p p2 -> p1 +p q <=p p2 +p q. +Proof. by move=>H x; case; [move/H; left|right]. Qed. + +Lemma subp_orr (p1 p2 q : Pred T) : p1 <=p p2 -> q +p p1 <=p q +p p2. +Proof. by move=>H x; case; [left | move/H; right]. Qed. + +Lemma subp_andl (p1 p2 q : Pred T) : p1 <=p p2 -> p1 *p q <=p p2 *p q. +Proof. by by move=>H x [H1 H2]; split; [apply: H|]. Qed. + +Lemma subp_andr (p1 p2 q : Pred T) : p1 <=p p2 -> q *p p1 <=p q *p p2. +Proof. by move=>H x [H1 H2]; split; [|apply: H]. Qed. + +End SubMemLaws. + +Hint Resolve subp_refl. + +Section ListMembership. +Variable T : Type. + +Fixpoint Mem_Seq (s : seq T) := + if s is y::s' then (fun x => x = y \/ Mem_Seq s' x) else xPred0. + +Definition EqSeq_Class := seq T. +Identity Coercion seq_of_EqSeq : EqSeq_Class >-> seq. + +Coercion Pred_of_Eq_Seq (s : EqSeq_Class) : Pred_Class := [eta Mem_Seq s]. + +Canonical Structure seq_PredType := @mkPredType T (seq T) Pred_of_Eq_Seq. +(* The line below makes Mem_Seq a canonical instance of topred. *) +Canonical Structure Mem_Seq_PredType := mkPredType Mem_Seq. + +Lemma In_cons y s x : (x \In y :: s) <-> (x = y) \/ (x \In s). +Proof. by []. Qed. + +Lemma In_nil x : (x \In [::]) <-> False. +Proof. by []. Qed. + +Lemma Mem_Seq1 x y : (x \In [:: y]) <-> (x = y). +Proof. by rewrite In_cons orpF. Qed. + +Definition InE := (Mem_Seq1, In_cons, In_Simpl). + +Lemma Mem_cat x : forall s1 s2, (x \In s1 ++ s2) <-> x \In s1 \/ x \In s2. +Proof. +elim=>[|y s1 IH] s2 /=; first by split; [right | case]. +rewrite !InE /=. +split. +- case=>[->|/IH]; first by left; left. + by case; [left; right | right]. +case; first by case; [left | move=>H; right; apply/IH; left]. +by move=>H; right; apply/IH; right. +Qed. + +Lemma In_split x s : x \In s -> exists s1 s2, s = s1 ++ x :: s2. +Proof. +elim:s=>[|y s IH] //=; rewrite InE. +case=>[<-|]; first by exists [::], s. +by case/IH=>s1 [s2 ->]; exists (y :: s1), s2. +Qed. + +End ListMembership. + +Lemma Mem_map T T' (f : T -> T') x (s : seq T) : + x \In s -> f x \In (map f s). +Proof. +elim: s=>[|y s IH] //; rewrite InE /=. +by case=>[<-|/IH]; [left | right]. +Qed. + +Lemma Mem_map_inv T T' (f : T -> T') x (s : seq T) : + x \In (map f s) -> exists y, x = f y /\ y \In s. +Proof. +elim: s=>[|y s IH] //=; rewrite InE /=. +case; first by move=>->; exists y; split=>//; left. +by case/IH=>z [->]; exists z; split=>//; right. +Qed. + +Prenex Implicits Mem_map_inv. + +(* Setoids for extensional equality of functions *) + +Lemma eqfun_refl A B (f : A -> B) : f =1 f. Proof. by []. Qed. +Lemma eqfun_sym A B (f1 f2 : A -> B) : f1 =1 f2 -> f2 =1 f1. +Proof. by move=>H x; rewrite H. Qed. +Lemma eqfun_trans A B (f1 f2 f3 : A -> B) : f1 =1 f2 -> f2 =1 f3 -> f1 =1 f3. +Proof. by move=>H1 H2 x; rewrite H1 H2. Qed. + +Add Parametric Relation A B : (A -> B) (@eqfun _ _) + reflexivity proved by (@eqfun_refl A B) + symmetry proved by (@eqfun_sym A B) + transitivity proved by (@eqfun_trans A B) as eqfun_morph. + + +(***********************************) +(* Image of a collective predicate *) +(***********************************) + +Section Image. +Variables (A B : Type) (P : Pred A) (f : A -> B). + +Inductive image_spec b : Prop := Im_mem a of b = f a & a \In P. + +Definition Image' : Pred B := image_spec. + +End Image. + +(* swap to make the notation consider P before E; helps inference *) +Notation Image f P := (Image' P f). + +Notation "[ 'Image' E | i <- s ]" := (Image (fun i => E) s) + (at level 0, E at level 99, i ident, + format "[ '[hv' 'Image' E '/ ' | i <- s ] ']'") : rel_scope. + +Notation "[ 'Image' E | i <- s & C ]" := [Image E | i <- [PredI s & C]] + (at level 0, E at level 99, i ident, + format "[ '[hv' 'Image' E '/ ' | i <- s '/ ' & C ] ']'") : rel_scope. + +Notation "[ 'Image' E | i : T <- s ]" := (Image (fun i : T => E) s) + (at level 0, E at level 99, i ident, only parsing) : rel_scope. + +Notation "[ 'Image' E | i : T <- s & C ]" := + [Image E | i : T <- [PredI s & C]] + (at level 0, E at level 99, i ident, only parsing) : rel_scope. + +Lemma Image_mem A B (f : A -> B) (P : Pred A) x : x \In P -> f x \In Image f P. +Proof. by apply: Im_mem. Qed. + +Lemma Image_inj_sub A B (f : A -> B) (X1 X2 : Pred A) : + injective f -> Image f X1 <=p Image f X2 -> X1 <=p X2. +Proof. by move=>H E x /(Image_mem f) /E [y] /H ->. Qed. + +Lemma Image_inj_eqmem A B (f : A -> B) (X1 X2 : Pred A) : + injective f -> Image f X1 =p Image f X2 -> X1 =p X2. +Proof. by move=>H E; split; apply: Image_inj_sub H _ _; rewrite E. Qed. + +Lemma ImageU A B (f : A -> B) (X1 X2 : Pred A) : + Image f (PredU X1 X2) =p [PredU Image f X1 & Image f X2]. +Proof. +move=>x; split. +- by case=>y -> [H|H]; [left | right]; apply: Image_mem. +by case; case=>y -> H; apply: Image_mem; [left | right]. +Qed. + +Lemma ImageIm A B C (f1 : B -> C) (f2 : A -> B) (X : Pred A) : + Image f1 (Image f2 X) =p Image (f1 \o f2) X. +Proof. +move=>x; split; first by case=>_ -> [x' ->] H; exists x'. +by case=>a -> H; exists (f2 a)=>//; exists a. +Qed. + +Lemma ImageEq A B (f1 f2 : A -> B) (X : Pred A) : + f1 =1 f2 -> Image f1 X =p Image f2 X. +Proof. by move=>H x; split; case=>a ->; exists a. Qed. + diff --git a/pcm/prelude.v b/pcm/prelude.v new file mode 100644 index 0000000..fb486c4 --- /dev/null +++ b/pcm/prelude.v @@ -0,0 +1,247 @@ +(* +Copyright 2010 IMDEA Software Institute +Licensed under the Apache License, Version 2.0 (the "License"); +you may not use this file except in compliance with the License. +You may obtain a copy of the License at + http://www.apache.org/licenses/LICENSE-2.0 +Unless required by applicable law or agreed to in writing, software +distributed under the License is distributed on an "AS IS" BASIS, +WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. +See the License for the specific language governing permissions and +limitations under the License. +*) + +(******************************************************************************) +(* This file is Prelude -- often used notation definitions and lemmas that *) +(* are not included in the other libraries. *) +(******************************************************************************) + +From Coq Require Import ssreflect ssrbool ssrfun Eqdep. +From mathcomp Require Import ssrnat eqtype. +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +(***********) +(* Prelude *) +(***********) + +(* export inj_pair without exporting the whole Eqdep library *) +Definition inj_pair2 := @inj_pair2. +Arguments inj_pair2 [U P p x y] _. +Prenex Implicits inj_pair2. + +(* eta laws for pairs and units *) +Notation prod_eta := surjective_pairing. + +(* eta law often used with injection *) +Lemma prod_inj A B (x y : A * B) : x = y <-> (x.1, x.2) = (y.1, y.2). +Proof. by case: x y=>x1 x2 []. Qed. + +(* This declaration won't be needed after Coq 8.8.0 is out *) +Prenex Implicits Some_inj. + +(***************************) +(* operations on functions *) +(***************************) + +Lemma compA A B C D (h : A -> B) (g : B -> C) (f : C -> D) : + (f \o g) \o h = f \o (g \o h). +Proof. by []. Qed. + +Definition fprod A1 A2 B1 B2 (f1 : A1 -> B1) (f2 : A2 -> B2) := + fun (x : A1 * A2) => (f1 x.1, f2 x.2). + +Notation "f1 \* f2" := (fprod f1 f2) (at level 42, left associativity). + +Lemma ext A (B : A -> Type) (f1 f2 : forall x, B x) : + f1 = f2 -> forall x, f1 x = f2 x. +Proof. by move=>->. Qed. + +(* function splice *) +Definition fsplice A B1 B2 (f1 : A -> B1) (f2 : A -> B2) := + fun x : A => (f1 x, f2 x). + +Notation "[ 'fs' f1 , f2 ]" := (fsplice f1 f2) + (at level 0, format "[ 'fs' f1 , f2 ]"). + +(* composing relation and function *) +(* should i turn relations into functions on products? *) +(* then i won't need the next definition *) +(* it will be described using composition and splicing *) + +Definition relfuncomp A B (D : rel A) (f : B -> A) : rel B := + fun x y => D (f x) (f y). + +Notation "D \\o f" := (@relfuncomp _ _ D f) (at level 42, left associativity). + +(************************) +(* extension to ssrbool *) +(************************) + +Reserved Notation "[ /\ P1 , P2 , P3 , P4 , P5 & P6 ]" (at level 0, format + "'[hv' [ /\ '[' P1 , '/' P2 , '/' P3 , '/' P4 , '/' P5 ']' '/ ' & P6 ] ']'"). + +Reserved Notation "[ /\ P1 , P2 , P3 , P4 , P5 , P6 & P7 ]" (at level 0, format + "'[hv' [ /\ '[' P1 , '/' P2 , '/' P3 , '/' P4 , '/' P5 , '/' P6 ']' '/ ' & P7 ] ']'"). + +Reserved Notation "[ \/ P1 , P2 , P3 , P4 | P5 ]" (at level 0, format + "'[hv' [ \/ '[' P1 , '/' P2 , '/' P3 , '/' P4 ']' '/ ' | P5 ] ']'"). +Reserved Notation "[ \/ P1 , P2 , P3 , P4 , P5 | P6 ]" (at level 0, format + "'[hv' [ \/ '[' P1 , '/' P2 , '/' P3 , '/' P4 , '/' P5 ']' '/ ' | P6 ] ']'"). + +Reserved Notation "[ \/ P1 , P2 , P3 , P4 , P5 , P6 | P7 ]" (at level 0, format + "'[hv' [ \/ '[' P1 , '/' P2 , '/' P3 , '/' P4 , '/' P5 , '/' P6 ']' '/ ' | P7 ] ']'"). + +Inductive and6 (P1 P2 P3 P4 P5 P6 : Prop) : Prop := + And6 of P1 & P2 & P3 & P4 & P5 & P6. +Inductive and7 (P1 P2 P3 P4 P5 P6 P7 : Prop) : Prop := + And7 of P1 & P2 & P3 & P4 & P5 & P6 & P7. + +Inductive or5 (P1 P2 P3 P4 P5 : Prop) : Prop := + Or51 of P1 | Or52 of P2 | Or53 of P3 | Or54 of P4 | Or55 of P5. +Inductive or6 (P1 P2 P3 P4 P5 P6 : Prop) : Prop := + Or61 of P1 | Or62 of P2 | Or63 of P3 | Or64 of P4 | Or65 of P5 | Or66 of P6. +Inductive or7 (P1 P2 P3 P4 P5 P6 P7 : Prop) : Prop := + Or71 of P1 | Or72 of P2 | Or73 of P3 | Or74 of P4 | Or75 of P5 | Or76 of P6 | + Or77 of P7. + +Notation "[ /\ P1 , P2 , P3 , P4 , P5 & P6 ]" := + (and6 P1 P2 P3 P4 P5 P6) : type_scope. +Notation "[ /\ P1 , P2 , P3 , P4 , P5 , P6 & P7 ]" := + (and7 P1 P2 P3 P4 P5 P6 P7) : type_scope. +Notation "[ \/ P1 , P2 , P3 , P4 | P5 ]" := + (or5 P1 P2 P3 P4 P5) : type_scope. +Notation "[ \/ P1 , P2 , P3 , P4 , P5 | P6 ]" := + (or6 P1 P2 P3 P4 P5 P6) : type_scope. +Notation "[ \/ P1 , P2 , P3 , P4 , P5 , P6 | P7 ]" := + (or7 P1 P2 P3 P4 P5 P6 P7) : type_scope. + +Section ReflectConnectives. +Variable b1 b2 b3 b4 b5 b6 b7 : bool. + +Lemma and6P : reflect [/\ b1, b2, b3, b4, b5 & b6] [&& b1, b2, b3, b4, b5 & b6]. +Proof. +by case b1; case b2; case b3; case b4; case b5; case b6; + constructor; try by case. +Qed. + +Lemma and7P : reflect [/\ b1, b2, b3, b4, b5, b6 & b7] + [&& b1, b2, b3, b4, b5, b6 & b7]. +Proof. +by case b1; case b2; case b3; case b4; case b5; case b6; case: b7; + constructor; try by case. +Qed. + +Lemma or5P : reflect [\/ b1, b2, b3, b4 | b5] [|| b1, b2, b3, b4 | b5]. +Proof. +case b1; first by constructor; constructor 1. +case b2; first by constructor; constructor 2. +case b3; first by constructor; constructor 3. +case b4; first by constructor; constructor 4. +case b5; first by constructor; constructor 5. +by constructor; case. +Qed. + +Lemma or6P : reflect [\/ b1, b2, b3, b4, b5 | b6] [|| b1, b2, b3, b4, b5 | b6]. +Proof. +case b1; first by constructor; constructor 1. +case b2; first by constructor; constructor 2. +case b3; first by constructor; constructor 3. +case b4; first by constructor; constructor 4. +case b5; first by constructor; constructor 5. +case b6; first by constructor; constructor 6. +by constructor; case. +Qed. + +Lemma or7P : reflect [\/ b1, b2, b3, b4, b5, b6 | b7] + [|| b1, b2, b3, b4, b5, b6 | b7]. +Proof. +case b1; first by constructor; constructor 1. +case b2; first by constructor; constructor 2. +case b3; first by constructor; constructor 3. +case b4; first by constructor; constructor 4. +case b5; first by constructor; constructor 5. +case b6; first by constructor; constructor 6. +case b7; first by constructor; constructor 7. +by constructor; case. +Qed. + +End ReflectConnectives. + +Arguments and6P [b1 b2 b3 b4 b5 b6]. +Arguments and7P [b1 b2 b3 b4 b5 b6 b7]. +Arguments or5P [b1 b2 b3 b4 b5]. +Arguments or6P [b1 b2 b3 b4 b5 b6]. +Arguments or7P [b1 b2 b3 b4 b5 b6 b7]. +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]. + +(**********************************************) +(* Reflexive-transitive closure of a relation *) +(**********************************************) + +Fixpoint iter' T (g : T -> T -> Prop) n s1 s2 := + if n is n'.+1 then exists s, g s1 s /\ iter' g n' s s2 else s1 = s2. +Definition iter T (g : T -> T -> Prop) s1 s2 := exists n, iter' g n s1 s2. +(* curry the iteration *) +Definition iterc A T (g : A -> T -> A -> T -> Prop) a1 s1 a2 s2 := + iter (fun (a b : A * T) => g a.1 a.2 b.1 b.2) (a1, s1) (a2, s2). + +Section IteratedRels. +Variable T : Type. +Implicit Type g : T -> T -> Prop. + +Lemma iter_refl g s : iter g s s. +Proof. by exists 0. Qed. + +Lemma iter_trans g s1 s2 s3 : iter g s1 s2 -> iter g s2 s3 -> iter g s1 s3. +Proof. +case=>n; elim: n s1 =>[|n IH] s1 /=; first by move=>->. +by case=>s [H1 H2] /(IH _ H2) [m]; exists m.+1, s. +Qed. + +Lemma iterS g n s1 s2 : + iter' g n.+1 s1 s2 <-> exists s, iter' g n s1 s /\ g s s2. +Proof. +elim: n s1=>[|n IH] s1. +- by split; [case=>s [H <-]; exists s1 | case=>s [<- H]; exists s2]. +split; first by case=>s [H1] /IH [s'][H G]; exists s'; split=>//; exists s. +by case=>s [[s'][H1 H G]]; exists s'; split=>//; apply/IH; exists s. +Qed. + +Lemma iter'_sub g g' n s1 s2 : + (forall s1 s2, g s1 s2 -> g' s1 s2) -> + iter' g n s1 s2 -> iter' g' n s1 s2. +Proof. by move=>H; elim: n s1=>[|n IH] s1 //= [s][/H G] /IH; exists s. Qed. + +Lemma iter_sub g g' s1 s2 : + (forall s1 s2, g s1 s2 -> g' s1 s2) -> iter g s1 s2 -> iter g' s1 s2. +Proof. by move=>H [n]; exists n; apply: iter'_sub H _. Qed. + +Lemma iter1 g s1 s2 : g s1 s2 -> iter g s1 s2. +Proof. by exists 1, s2. Qed. + +End IteratedRels. + +Lemma iter2 A T (g : A -> T -> A -> T -> Prop) x1 s1 x2 s2 : + g x1 s1 x2 s2 -> iterc g x1 s1 x2 s2. +Proof. by apply: iter1. Qed. + +Prenex Implicits iter1 iter2. +Arguments iter_refl [T g s]. +Hint Resolve iter_refl. + +(**************) +(* empty type *) +(**************) + +Lemma emptyset_eqP : Equality.axiom (fun _ _ : Empty_set => true). +Proof. by case. Qed. + +Definition emptysetEqMix := EqMixin emptyset_eqP. +Canonical emptysetEqType := Eval hnf in EqType Empty_set emptysetEqMix. \ No newline at end of file diff --git a/pcm/unionmap.v b/pcm/unionmap.v new file mode 100644 index 0000000..642759b --- /dev/null +++ b/pcm/unionmap.v @@ -0,0 +1,2303 @@ +(* +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 eqtype seq path. +From fcsl Require Import ordtype finmap pcm. +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +(******************************************************************************) +(* This file contains the reference implementation of finite maps with keys *) +(* satisfying condition p and supporting disjoint union via a top element. *) +(******************************************************************************) + +(* I decided to have union_map_class be a class, rather than a +structure. The class packages a condition on keys. Ordinary union_maps +have a trivially true condition. Heaps have a condition that the +pointers are non-null. Then ordinary union maps, as well as heaps, +are declared as instances of this class, to automatically inherit all +the lemmas about the operations. + +An alternative implementation would have been to define +union_map_class as a structure parametrized by the condition, and then +obtain heaps by taking the parameter condition of non-null pointers, +and ordinary union_maps by taking parameter condition to true. + +I decided on the more complicated design to avoid the universe +jump. Heap values are dynamic, and live in Type(1), while most of the +times, ordinary union_maps store values from Type(0), and can be +nested. If the two structures (heaps and ordinary union maps) share +the same code, they will both lift to the common universe of Type(1), +preventing me from using maps at Type(0), and storing them in the heap. + +With the class design, no code is shared, only the lemmas (i.e., only +the class type, but that can freely live in an arbitrary +universe). The price to pay is that the code of the methods has to be +duplicated when declaring instances (such as in heaps.v, or below for +union_mapUMC), just to keep the universes from jumping. *) + +Module UM. +Section UM. +Variables (K : ordType) (V : Type) (p : pred K). + +Inductive base := + Undef | Def (f : {finMap K -> V}) of all p (supp f). + +Section FormationLemmas. +Variable (f g : {finMap K -> V}). + +Lemma all_supp_insP k v : p k -> all p (supp f) -> all p (supp (ins k v f)). +Proof. +move=>H1 H2; apply/allP=>x; rewrite supp_ins inE /=. +by case: eqP=>[->|_] //=; move/(allP H2). +Qed. + +Lemma all_supp_remP k : all p (supp f) -> all p (supp (rem k f)). +Proof. +move=>H; apply/allP=>x; rewrite supp_rem inE /=. +by case: eqP=>[->|_] //=; move/(allP H). +Qed. + +Lemma all_supp_fcatP : + all p (supp f) -> all p (supp g) -> all p (supp (fcat f g)). +Proof. +move=>H1 H2; apply/allP=>x; rewrite supp_fcat inE /=. +by case/orP; [move/(allP H1) | move/(allP H2)]. +Qed. + +Lemma all_supp_kfilterP q : + all p (supp f) -> all p (supp (kfilter q f)). +Proof. +move=>H; apply/allP=>x; rewrite supp_kfilt mem_filter. +by case/andP=>_ /(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: bool_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 (p 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 dom_eq f1 f2 := + match f1, f2 with + Def fs1 _, Def fs2 _ => supp fs1 == supp fs2 + | Undef, Undef => true + | _, _ => false + end. + +Definition free k f := + if f is Def fs pf then Def (all_supp_remP k pf) else Undef. + +Definition find k f := if f is Def fs _ then fnd k fs else None. + +Definition 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 um_filter q f := + if f is Def fs pf then Def (all_supp_kfilterP q pf) else Undef. + +Definition empb f := if f is Def fs _ then supp fs == [::] else false. + +Definition undefb f := if f is Undef then true else false. + +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] : p k /\ all p (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 (pts k v) f) -> + (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] : p k /\ all p (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); 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. + +(* a class of union_map types *) + +Module UMC. +Section MixinDef. +Variables (K : ordType) (V : Type) (p : pred K). + +Structure mixin_of (T : Type) := Mixin { + defined_op : T -> bool; + empty_op : T; + undef_op : T; + upd_op : K -> V -> T -> T; + dom_op : T -> seq K; + dom_eq_op : T -> T -> bool; + free_op : K -> T -> T; + find_op : K -> T -> option V; + union_op : T -> T -> T; + um_filter_op : pred K -> T -> T; + empb_op : T -> bool; + undefb_op : T -> bool; + pts_op : K -> V -> T; + + from_op : T -> UM.base V p; + to_op : UM.base V p -> T; + _ : forall b, from_op (to_op b) = b; + _ : forall f, to_op (from_op f) = f; + _ : forall f, defined_op f = UM.valid (from_op f); + _ : undef_op = to_op (UM.Undef V p); + _ : empty_op = to_op (UM.empty V p); + _ : forall k v f, upd_op k v f = to_op (UM.upd k v (from_op f)); + _ : forall f, dom_op f = UM.dom (from_op f); + _ : forall f1 f2, dom_eq_op f1 f2 = UM.dom_eq (from_op f1) (from_op f2); + _ : forall k f, free_op k f = to_op (UM.free k (from_op f)); + _ : forall k f, find_op k f = UM.find k (from_op f); + _ : forall f1 f2, + union_op f1 f2 = to_op (UM.union (from_op f1) (from_op f2)); + _ : forall q f, um_filter_op q f = to_op (UM.um_filter q (from_op f)); + _ : forall f, empb_op f = UM.empb (from_op f); + _ : forall f, undefb_op f = UM.undefb (from_op f); + _ : forall k v, pts_op k v = to_op (UM.pts p k v)}. +End MixinDef. + +Section ClassDef. +Variables (K : ordType) (V : Type). + +(* I used to package K and V into the class. I did that to get cond +function over keys to work without supplying the type parameter. I.e., +with K and V out, I have to write cond U k, where U : union_map_class +K V. But the complication wasn't worth it. One has to frequently +reorder arguments to various operations in a "less-ad-hoc" style, to +get the types to be inferred properly. While fun, it's just an +unnecessary hassle, since cond is not used that much. The condition p +can be kept in the structure, however, since no types depend on it. *) + +Structure class_of (T : Type) := Class { + p : pred K; + mixin : mixin_of V p T}. + +Structure type : Type := Pack {sort : Type; _ : class_of sort; _ : Type}. +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 T. +Definition pack p (m : @mixin_of K V p T) := + @Pack T (@Class T p m) T. + +(* Definition cond : pred K := @p _ class. *) +Definition cond := [pred x : K | @p _ class x]. +Definition from := from_op (mixin class). +Definition to := to_op (mixin class). +Definition defined := defined_op (mixin class). +Definition um_undef := undef_op (mixin class). +Definition empty := empty_op (mixin class). +Definition upd : K -> V -> cT -> cT := upd_op (mixin class). +Definition dom : cT -> seq K := dom_op (mixin class). +Definition dom_eq := dom_eq_op (mixin class). +Definition free : K -> cT -> cT := free_op (mixin class). +Definition find : K -> cT -> option V := find_op (mixin class). +Definition union := union_op (mixin class). +Definition um_filter : pred K -> cT -> cT := um_filter_op (mixin class). +Definition empb := empb_op (mixin class). +Definition undefb := undefb_op (mixin class). +Definition pts : K -> V -> cT := pts_op (mixin class). + +End ClassDef. + +Arguments um_undef [K V cT]. +Arguments empty [K V cT]. +Arguments pts [K V cT] _ _. +Prenex Implicits to um_undef empty pts. + +Section Lemmas. +Variables (K : ordType) (V : Type) (U : type K V). +Local Coercion sort : type >-> Sortclass. +Implicit Type f : U. + +Lemma ftE (b : UM.base V (cond U)) : from (to b) = b. +Proof. by case: U b=>x [p][*]. Qed. + +Lemma tfE f : to (from f) = f. +Proof. by case: U f=>x [p][*]. Qed. + +Lemma eqE (b1 b2 : UM.base V (cond U)) : + to b1 = to b2 <-> b1 = b2. +Proof. by split=>[E|-> //]; rewrite -[b1]ftE -[b2]ftE E. Qed. + +Lemma defE f : defined f = UM.valid (from f). +Proof. by case: U f=>x [p][*]. Qed. + +Lemma undefE : um_undef = to (UM.Undef V (cond U)). +Proof. by case: U=>x [p][*]. Qed. + +Lemma emptyE : empty = to (UM.empty V (cond U)). +Proof. by case: U=>x [p][*]. Qed. + +Lemma updE k v f : upd k v f = to (UM.upd k v (from f)). +Proof. by case: U k v f=>x [p][*]. Qed. + +Lemma domE f : dom f = UM.dom (from f). +Proof. by case: U f=>x [p][*]. Qed. + +Lemma dom_eqE f1 f2 : dom_eq f1 f2 = UM.dom_eq (from f1) (from f2). +Proof. by case: U f1 f2=>x [p][*]. Qed. + +Lemma freeE k f : free k f = to (UM.free k (from f)). +Proof. by case: U k f=>x [p][*]. Qed. + +Lemma findE k f : find k f = UM.find k (from f). +Proof. by case: U k f=>x [p][*]. Qed. + +Lemma unionE f1 f2 : union f1 f2 = to (UM.union (from f1) (from f2)). +Proof. by case: U f1 f2=>x [p][*]. Qed. + +Lemma um_filterE q f : um_filter q f = to (UM.um_filter q (from f)). +Proof. by case: U q f=>x [p][*]. Qed. + +Lemma empbE f : empb f = UM.empb (from f). +Proof. by case: U f=>x [p][*]. Qed. + +Lemma undefbE f : undefb f = UM.undefb (from f). +Proof. by case: U f=>x [p][*]. Qed. + +Lemma ptsE k v : pts k v = to (UM.pts (cond U) k v). +Proof. by case: U k v=>x [p][*]. Qed. + +End Lemmas. + +Module Exports. +Coercion sort : type >-> Sortclass. +Notation union_map_class := type. +Notation UMCMixin := Mixin. +Notation UMC T m := (@pack _ _ T _ m). + +Notation "[ 'umcMixin' 'of' T ]" := (mixin (class _ _ _ : class_of T)) + (at level 0, format "[ 'umcMixin' 'of' T ]") : form_scope. +Notation "[ 'um_class' 'of' T 'for' C ]" := (@clone _ _ T C _ id) + (at level 0, format "[ 'um_class' 'of' T 'for' C ]") : form_scope. +Notation "[ 'um_class' 'of' T ]" := (@clone _ _ T _ _ id) + (at level 0, format "[ 'um_class' 'of' T ]") : form_scope. + +Notation cond := cond. +Notation um_undef := um_undef. +Notation upd := upd. +Notation dom := dom. +Notation dom_eq := dom_eq. +Notation free := free. +Notation find := find. +Notation um_filter := um_filter. +Notation empb := empb. +Notation undefb := undefb. +Notation pts := pts. + +Definition umEX := + (ftE, tfE, eqE, undefE, defE, emptyE, updE, domE, dom_eqE, + freeE, findE, unionE, um_filterE, empbE, undefbE, ptsE, UM.umapE). + +End Exports. + +End UMC. + +Export UMC.Exports. + + +(***********************) +(* monoidal properties *) +(***********************) + +Module UnionMapClassPCM. +Section UnionMapClassPCM. +Variables (K : ordType) (V : Type) (U : union_map_class K V). +Implicit Type f : U. + +Local Notation "f1 \+ f2" := (@UMC.union _ _ _ f1 f2) + (at level 43, left associativity). +Local Notation valid := (@UMC.defined _ _ U). +Local Notation unit := (@UMC.empty _ _ U). + +Lemma joinC f1 f2 : f1 \+ f2 = f2 \+ f1. +Proof. +rewrite !umEX /UM.union. +case: (UMC.from f1)=>[|f1' H1]; case: (UMC.from f2)=>[|f2' H2] //. +by case: ifP=>E; rewrite disjC E // fcatC. +Qed. + +Lemma joinCA f1 f2 f3 : f1 \+ (f2 \+ f3) = f2 \+ (f1 \+ f3). +Proof. +rewrite !umEX /UM.union /=. +case: (UMC.from f1) (UMC.from f2) (UMC.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. +Qed. + +Lemma joinA f1 f2 f3 : f1 \+ (f2 \+ f3) = (f1 \+ f2) \+ f3. +Proof. by rewrite (joinC f2) joinCA joinC. Qed. + +Lemma validL f1 f2 : valid (f1 \+ f2) -> valid f1. +Proof. by rewrite !umEX; case: (UMC.from f1). Qed. + +Lemma unitL f : unit \+ f = f. +Proof. +rewrite -[f]UMC.tfE !umEX /UM.union /UM.empty. +by case: (UMC.from f)=>[//|f' H]; rewrite disjC disj_nil fcat0s. +Qed. + +Lemma validU : valid unit. +Proof. by rewrite !umEX. Qed. + +End UnionMapClassPCM. + +Module Exports. +Section Exports. +Variables (K : ordType) (V : Type) (U : union_map_class K V). + +(* generic structure for PCM for *all* union maps *) +(* I will add specific ones too for individual types *) +(* so that the projections can match against a concrete type *) +(* not against another projection, as is the case *) +(* with the generic class here *) +Definition union_map_classPCMMix := + PCMMixin (@joinC K V U) (@joinA K V U) (@unitL K V U) + (@validL K V U) (validU U). +Canonical union_map_classPCM := Eval hnf in PCM U union_map_classPCMMix. + +End Exports. +End Exports. + +End UnionMapClassPCM. + +Export UnionMapClassPCM.Exports. + + +(*****************) +(* Cancelativity *) +(*****************) + +Module UnionMapClassCPCM. +Section Cancelativity. +Variables (K : ordType) (V : Type) (U : union_map_class K V). +Implicit Type f : U. + +Lemma joinKf f1 f2 f : valid (f1 \+ f) -> f1 \+ f = f2 \+ f -> f1 = f2. +Proof. +rewrite -{3}[f1]UMC.tfE -{2}[f2]UMC.tfE !pcmE /= !umEX /UM.valid /UM.union. +case: (UMC.from f) (UMC.from f1) (UMC.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. + +End Cancelativity. + +Module Exports. +Section Exports. +Variables (K : ordType) (V : Type) (U : union_map_class K V). + +Definition union_map_classCPCMMix := CPCMMixin (@joinKf K V U). +Canonical union_map_classCPCM := Eval hnf in CPCM U union_map_classCPCMMix. + +End Exports. +End Exports. + +End UnionMapClassCPCM. + +Export UnionMapClassCPCM.Exports. + +(*********************************************************) +(* if V is an equality type, then union_map_class is too *) +(*********************************************************) + +Module UnionMapEq. +Section UnionMapEq. +Variables (K : ordType) (V : eqType) (p : pred K). +Variables (T : Type) (m : @UMC.mixin_of K V p T). + +Definition unionmap_eq (f1 f2 : UMC T m) := + match (UMC.from f1), (UMC.from f2) with + | UM.Undef, UM.Undef => true + | UM.Def f1' pf1, UM.Def f2' pf2 => f1' == f2' + | _, _ => false + end. + +Lemma unionmap_eqP : Equality.axiom unionmap_eq. +Proof. +move=>x y; rewrite -{1}[x]UMC.tfE -{1}[y]UMC.tfE /unionmap_eq. +case: (UMC.from x)=>[|f1 H1]; case: (UMC.from y)=>[|f2 H2] /=. +- by constructor. +- by constructor; move/(@UMC.eqE _ _ (UMC T m)). +- by constructor; move/(@UMC.eqE _ _ (UMC T m)). +case: eqP=>E; constructor; rewrite (@UMC.eqE _ _ (UMC T m)). +- by rewrite UM.umapE. +by case. +Qed. + +End UnionMapEq. + +Module Exports. +Section Exports. +Variables (K : ordType) (V : eqType) (p : pred K). +Variables (T : Type) (m : @UMC.mixin_of K V p T). + +Definition union_map_class_eqMix := EqMixin (@unionmap_eqP K V p T m). + +End Exports. +End Exports. + +End UnionMapEq. + +Export UnionMapEq.Exports. + + +(*******) +(* dom *) +(*******) + +Section DomLemmas. +Variables (K : ordType) (V : Type) (U : union_map_class K V). +Implicit Types (k : K) (v : V) (f g : U). + +Lemma dom_undef : dom (um_undef : U) = [::]. +Proof. by rewrite !umEX. Qed. + +Lemma dom0 : dom (Unit : U) = [::]. +Proof. by rewrite pcmE /= !umEX. Qed. + +Lemma dom0E f : valid f -> dom f =i pred0 -> f = Unit. +Proof. +rewrite !pcmE /= !umEX /UM.valid /UM.dom /UM.empty -{3}[f]UMC.tfE. +case: (UMC.from f)=>[|f' S] //= _; rewrite !umEX fmapE /= {S}. +by case: f'; case=>[|kv s] //= P /= /(_ kv.1); rewrite inE eq_refl. +Qed. + +Lemma domU k v f : + dom (upd k v f) =i + [pred x | cond U k & if x == k then valid f else x \in dom f]. +Proof. +rewrite pcmE /= !umEX /UM.dom /UM.upd /UM.valid /= /cond. +case: (UMC.from f)=>[|f' H] /= x. +- by rewrite !inE /= andbC; case: ifP. +case: decP; first by move=>->; rewrite supp_ins. +by move/(introF idP)=>->. +Qed. + +Lemma domF k f : + dom (free k f) =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 domUn f1 f2 : + dom (f1 \+ f2) =i + [pred x | valid (f1 \+ f2) & (x \in dom f1) || (x \in dom f2)]. +Proof. +rewrite !pcmE /= !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. + +Lemma dom_valid k f : k \in dom f -> valid f. +Proof. by rewrite pcmE /= !umEX; case: (UMC.from f). Qed. + +Lemma dom_cond k f : k \in dom f -> cond U k. +Proof. by rewrite !umEX; case: (UMC.from f)=>[|f' F] // /(allP F). 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 k f = f. +Proof. +rewrite -{3}[f]UMC.tfE !umEX /UM.dom /UM.free. +by case: (UMC.from f)=>[|f' H] //; apply: rem_supp. +Qed. + +CoInductive dom_find_spec k f : bool -> Type := +| dom_find_none' : find k f = None -> dom_find_spec k f false +| dom_find_some' v : find k f = Some v -> + f = upd k v (free k f) -> dom_find_spec k f true. + +Lemma dom_find k f : dom_find_spec k f (k \in dom f). +Proof. +rewrite !umEX /UM.dom -{1}[f]UMC.tfE. +case: (UMC.from f)=>[|f' H]. +- by apply: dom_find_none'; rewrite !umEX. +case: suppP (allP H k)=>[v|] 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 dom_umfilt p f : dom (um_filter p f) =i [predI p & dom f]. +Proof. +rewrite !umEX /UM.dom /UM.um_filter. +case: (UMC.from f)=>[|f' H] x; first by rewrite !inE /= andbF. +by rewrite supp_kfilt mem_filter. +Qed. + +Lemma dom_prec f1 f2 g1 g2 : + valid (f1 \+ g1) -> + f1 \+ g1 = f2 \+ g2 -> + dom f1 =i dom f2 -> f1 = f2. +Proof. +rewrite -{4}[f1]UMC.tfE -{3}[f2]UMC.tfE !pcmE /= !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 domK f1 f2 g1 g2 : + 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 !pcmE /= !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 umfilt_dom f1 f2 : + valid (f1 \+ f2) -> um_filter (mem (dom f1)) (f1 \+ f2) = f1. +Proof. +rewrite -{4}[f1]UMC.tfE !pcmE /= !umEX. +rewrite /UM.valid /UM.union /UM.um_filter /UM.dom. +case: (UMC.from f1) (UMC.from f2)=>[|f1' F1][|f2' F2] //. +case: ifP=>// D _; rewrite kfilt_fcat /=. +have E1: {in supp f1', supp f1' =i predT} by []. +have E2: {in supp f2', supp f1' =i pred0}. +- by move=>x; rewrite disjC in D; case: disjP D=>// D _ /D /negbTE ->. +rewrite (eq_in_kfilter E1) (eq_in_kfilter E2). +by rewrite kfilter_predT kfilter_pred0 fcats0. +Qed. + +Lemma sorted_dom f : sorted (@ord K) (dom f). +Proof. by rewrite !umEX; case: (UMC.from f)=>[|[]]. Qed. + +Lemma uniq_dom f : uniq (dom f). +Proof. +apply: sorted_uniq (sorted_dom f); +by [apply: ordtype.trans | apply: ordtype.irr]. +Qed. + +Lemma perm_domUn f1 f2 : + valid (f1 \+ f2) -> perm_eq (dom (f1 \+ f2)) (dom f1 ++ dom f2). +Proof. +move=>Vh; apply: uniq_perm_eq; 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 !pcmE /= !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/perm_eq_size; rewrite size_cat. Qed. + +End DomLemmas. + +Hint Resolve sorted_dom uniq_dom. +Prenex Implicits find_some find_none. + + +(**********) +(* filter *) +(**********) + +Section FilterLemmas. +Variables (K : ordType) (V : Type) (U : union_map_class K V). +Implicit Type f : U. + +Lemma eq_in_umfilt p1 p2 f : + {in dom f, p1 =1 p2} -> um_filter p1 f = um_filter p2 f. +Proof. +rewrite !umEX /UM.dom /UM.um_filter /= /dom. +by case: (UMC.from f)=>[|f' H] //=; apply: eq_in_kfilter. +Qed. + +Lemma umfilt0 q : um_filter q Unit = Unit :> U. +Proof. by rewrite !pcmE /= !umEX /UM.um_filter /UM.empty kfilt_nil. Qed. + +Lemma umfilt_undef q : um_filter q um_undef = um_undef :> U. +Proof. by rewrite !umEX. Qed. + +Lemma umfiltUn q f1 f2 : + valid (f1 \+ f2) -> + um_filter q (f1 \+ f2) = um_filter q f1 \+ um_filter q f2. +Proof. +rewrite !pcmE /= !umEX /UM.valid /UM.union. +case: (UMC.from f1)=>[|f1' H1]; case: (UMC.from f2)=>[|f2' H2] //=. +by case: ifP=>D //= _; rewrite kfilt_fcat disj_kfilt. +Qed. + +Lemma umfilt_pred0 f : valid f -> um_filter pred0 f = Unit. +Proof. +rewrite !pcmE /= !umEX /UM.valid /UM.empty. +case: (UMC.from f)=>[|f' H] //= _; case: f' H=>f' P H. +by rewrite fmapE /= /kfilter' filter_pred0. +Qed. + +Lemma umfilt_predT f : um_filter predT f = f. +Proof. +rewrite -[f]UMC.tfE !umEX /UM.um_filter. +by case: (UMC.from f)=>[|f' H] //; rewrite kfilter_predT. +Qed. + +Lemma umfilt_predI p1 p2 f : + um_filter (predI p1 p2) f = um_filter p1 (um_filter p2 f). +Proof. +rewrite -[f]UMC.tfE !umEX /UM.um_filter. +by case: (UMC.from f)=>[|f' H] //; rewrite kfilter_predI. +Qed. + +Lemma umfilt_predU p1 p2 f : + um_filter (predU p1 p2) f = + um_filter p1 f \+ um_filter (predD p2 p1) f. +Proof. +rewrite pcmE /= !umEX /UM.um_filter /UM.union /=. +case: (UMC.from f)=>[|f' H] //=. +rewrite in_disj_kfilt; last by move=>x _; rewrite /= negb_and orbA orbN. +rewrite -kfilter_predU. +by apply: eq_in_kfilter=>x _; rewrite /= orb_andr orbN. +Qed. + +Lemma umfilt_dpredU f p q : + {subset p <= predC q} -> + um_filter (predU p q) f = um_filter p f \+ um_filter q f. +Proof. +move=>D; rewrite umfilt_predU (eq_in_umfilt (p1:=predD q p) (p2:=q)) //. +by move=>k _ /=; case X : (p k)=>//=; move/D/negbTE: X. +Qed. + +Lemma umfiltUnK p f1 f2 : + valid (f1 \+ f2) -> + um_filter p (f1 \+ f2) = f1 -> + um_filter p f1 = f1 /\ um_filter p f2 = Unit. +Proof. +move=>V'; rewrite (umfiltUn _ V') => E. +have {V'} V' : valid (um_filter p f1 \+ um_filter p f2). +- by rewrite E; move/validL: V'. +have F : dom (um_filter p f1) =i dom f1. +- move=>x; rewrite dom_umfilt inE /=. + apply/andP/idP=>[[//]| H1]; split=>//; move: H1. + rewrite -{1}E domUn inE /= !dom_umfilt inE /= inE /=. + by case: (x \in p)=>//=; rewrite andbF. +rewrite -{2}[f1]unitR in E F; move/(dom_prec V' E): F=>X. +by rewrite X in E V' *; rewrite (joinxK V' E). +Qed. + +Lemma umfiltU p k v f : + um_filter p (upd k v f) = + if p k then upd k v (um_filter p f) else + if cond U k then um_filter p f else um_undef. +Proof. +rewrite !umEX /UM.um_filter /UM.upd /cond. +case: (UMC.from f)=>[|f' F]; first by case: ifP=>H1 //; case: ifP. +case: ifP=>H1; case: decP=>H2 //=. +- by rewrite !umEX kfilt_ins H1. +- by rewrite H2 !umEX kfilt_ins H1. +by case: ifP H2. +Qed. + +Lemma umfiltF p k f : + um_filter p (free k f) = + if p k then free k (um_filter p f) else um_filter p f. +Proof. +rewrite !umEX /UM.um_filter /UM.free. +by case: (UMC.from f)=>[|f' F]; case: ifP=>// E; rewrite !umEX kfilt_rem E. +Qed. + +End FilterLemmas. + + +(*********) +(* valid *) +(*********) + +Section ValidLemmas. +Variables (K : ordType) (V : Type) (U : union_map_class K V). +Implicit Types (k : K) (v : V) (f g : U). + +Lemma invalidE f : ~~ valid f <-> f = um_undef. +Proof. by rewrite !pcmE /= !umEX -2![f]UMC.tfE !umEX; case: (UMC.from f). Qed. + +Lemma valid_undef : valid (um_undef : U) = false. +Proof. by apply/negbTE; apply/invalidE. Qed. + +Lemma validU k v f : valid (upd k v f) = cond U k && valid f. +Proof. +rewrite !pcmE /= !umEX /UM.valid /UM.upd /cond. +case: (UMC.from f)=>[|f' F]; first by rewrite andbF. +by case: decP=>[|/(introF idP)] ->. +Qed. + +Lemma validF k f : valid (free k f) = valid f. +Proof. by rewrite !pcmE /= !umEX; case: (UMC.from f). Qed. + +CoInductive 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 !pcmE /= !umEX -{1}[f1]UMC.tfE -{1}[f2]UMC.tfE. +rewrite /UM.valid /UM.union /=. +case: (UMC.from f1) (UMC.from f2)=>[|f1' F1][|f2' F2] /=. +- by apply: valid_false1; rewrite pcmE /= !umEX. +- by apply: valid_false1; rewrite pcmE /= !umEX. +- by apply: valid_false2; rewrite pcmE /= !umEX. +case: ifP=>E. +- apply: valid_true; try by rewrite !pcmE /= !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 validFUn k f1 f2 : + valid (f1 \+ f2) -> valid (free k f1 \+ 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 k f2). +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 valid_umfilt p f : valid (um_filter p f) = valid f. +Proof. by rewrite !pcmE /= !umEX; case: (UMC.from f). Qed. + +Lemma dom_inNL k f1 f2 : + valid (f1 \+ f2) -> k \in dom f1 -> k \notin dom f2. +Proof. by case: validUn=>//_ _ H _; apply: H. Qed. + +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. + +End ValidLemmas. + + +(**********) +(* dom_eq *) +(**********) + +Section DomEqLemmas. +Variables (K : ordType) (V : Type) (U : union_map_class K V). +Implicit Type f : U. + +Lemma domeqP f1 f2 : + reflect (valid f1 = valid f2 /\ dom f1 =i dom f2) (dom_eq f1 f2). +Proof. +rewrite !pcmE /= !umEX /UM.valid /UM.dom /UM.dom_eq /in_mem. +case: (UMC.from f1) (UMC.from f2)=>[|f1' F1][|f2' F2] /=. +- by constructor. +- by constructor; case. +- by constructor; case. +by case: eqP=>H; constructor; [rewrite H | case=>_ /suppE]. +Qed. + +Lemma domeq0E f : dom_eq f Unit -> f = Unit. +Proof. by case/domeqP; rewrite valid_unit dom0; apply: dom0E. Qed. + +Lemma domeq_refl f : dom_eq f f. +Proof. by case: domeqP=>//; case. Qed. + +Hint Resolve domeq_refl. + +Lemma domeq_sym f1 f2 : dom_eq f1 f2 = dom_eq f2 f1. +Proof. +suff L f f' : dom_eq f f' -> dom_eq f' f by apply/idP/idP; apply: L. +by case/domeqP=>H1 H2; apply/domeqP; split. +Qed. + +Lemma domeq_trans f1 f2 f3 : + dom_eq f1 f2 -> dom_eq f2 f3 -> dom_eq f1 f3. +Proof. +case/domeqP=>E1 H1 /domeqP [E2 H2]; apply/domeqP=>//. +by split=>//; [rewrite E1 E2 | move=>x; rewrite H1 H2]. +Qed. + +Lemma domeqVUnE f1 f2 f1' f2' : + dom_eq f1 f2 -> dom_eq f1' f2' -> + valid (f1 \+ f1') = valid (f2 \+ f2'). +Proof. +have L f f' g : dom_eq f f' -> valid (f \+ g) -> valid (f' \+ g). +- case/domeqP=>E F; case: validUn=>// Vg1 Vf H _; case: validUn=>//. + - by rewrite -E Vg1. + - by rewrite Vf. + by move=>x; rewrite -F; move/H/negbTE=>->. +move=>H H'; case X: (valid (f1 \+ f1')); apply/esym. +- by move/(L _ _ _ H): X; rewrite !(joinC f2); move/(L _ _ _ H'). +rewrite domeq_sym in H; rewrite domeq_sym in H'. +apply/negbTE; apply: contra (negbT X); move/(L _ _ _ H). +by rewrite !(joinC f1); move/(L _ _ _ H'). +Qed. + +Lemma domeqVUn f1 f2 f1' f2' : + dom_eq f1 f2 -> dom_eq f1' f2' -> + valid (f1 \+ f1') -> valid (f2 \+ f2'). +Proof. by move=>D /(domeqVUnE D) ->. Qed. + +Lemma domeqUn f1 f2 f1' f2' : + dom_eq f1 f2 -> dom_eq f1' f2' -> + dom_eq (f1 \+ f1') (f2 \+ f2'). +Proof. +suff L f f' g : dom_eq f f' -> dom_eq (f \+ g) (f' \+ g). +- move=>H H'; apply: domeq_trans (L _ _ _ H). + by rewrite !(joinC f1); apply: domeq_trans (L _ _ _ H'). +move=>F; case/domeqP: (F)=>E H. +apply/domeqP; split; first by apply: domeqVUnE F _. +move=>x; rewrite !domUn /= inE. +by rewrite (domeqVUnE F (domeq_refl g)) H. +Qed. + +Lemma domeqfUn f f1 f2 f1' f2' : + 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' : + 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. + +Lemma domeqK f1 f2 f1' f2' : + valid (f1 \+ f1') -> + dom_eq (f1 \+ f1') (f2 \+ f2') -> + dom_eq f1 f2 -> dom_eq f1' f2'. +Proof. +move=>V1 /domeqP [E1 H1] /domeqP [E2 H2]; move: (V1); rewrite E1=>V2. +apply/domeqP; split; first by rewrite (validR V1) (validR V2). +by apply: domK V1 V2 H1 H2. +Qed. + +End DomEqLemmas. + +Hint Resolve domeq_refl. + + +(*************************) +(* some precision lemmas *) +(*************************) + +Section Precision. +Variables (K : ordType) (V : Type) (U : union_map_class K V). +Implicit Types (x y : U). + +Lemma prec_flip x1 x2 y1 y2 : + valid (x1 \+ y1) -> x1 \+ y1 = x2 \+ y2 -> + valid (y2 \+ x2) /\ y2 \+ x2 = y1 \+ x1. +Proof. by move=>X /esym E; rewrite joinC E X joinC. Qed. + +Lemma prec_domV x1 x2 y1 y2 : + valid (x1 \+ y1) -> x1 \+ y1 = x2 \+ y2 -> + reflect {subset dom x1 <= dom x2} (valid (x1 \+ y2)). +Proof. +move=>V1 E; case V12 : (valid (x1 \+ y2)); constructor. +- move=>n Hs; have : n \in dom (x1 \+ y1) by rewrite domUn inE V1 Hs. + rewrite E domUn inE -E V1 /= orbC. + by case: validUn V12 Hs=>// _ _ L _ /L /negbTE ->. +move=>Hs; case: validUn V12=>//. +- by rewrite (validL V1). +- by rewrite E in V1; rewrite (validR V1). +by move=>k /Hs; rewrite E in V1; case: validUn V1=>// _ _ L _ /L /negbTE ->. +Qed. + +Lemma prec_filtV x1 x2 y1 y2 : + valid (x1 \+ y1) -> x1 \+ y1 = x2 \+ y2 -> + reflect (x1 = um_filter (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_umfilt inE; case/andP. +move: (umfilt_dom V1); rewrite E umfiltUn -?E //. +rewrite (eq_in_umfilt (f:=y2) (p2:=pred0)). +- by rewrite umfilt_pred0 ?unitR //; rewrite E in V1; rewrite (validR V1). +by move=>n; case: validUn X=>// _ _ L _ /(contraL (L _)) /negbTE. +Qed. + +End Precision. + + +(**********) +(* update *) +(**********) + +Section UpdateLemmas. +Variable (K : ordType) (V : Type) (U : union_map_class K V). +Implicit Types (k : K) (v : V) (f : U). + +Lemma upd_invalid k v : upd k v um_undef = um_undef :> U. +Proof. by rewrite !umEX. Qed. + +Lemma upd_inj k v1 v2 f : + valid f -> cond U k -> upd k v1 f = upd k v2 f -> v1 = v2. +Proof. +rewrite !pcmE /= !umEX /UM.valid /UM.upd /cond. +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 k2 f) = + if k1 == k2 then upd k1 v f else free k2 (upd k1 v f). +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 !pcmE /= !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. + +End UpdateLemmas. + + +(********) +(* free *) +(********) + +Section FreeLemmas. +Variables (K : ordType) (V : Type) (U : union_map_class K V). +Implicit Types (k : K) (v : V) (f : U). + +Lemma free0 k : free k Unit = Unit :> U. +Proof. by rewrite !pcmE /= !umEX /UM.free /UM.empty rem_empty. Qed. + +Lemma freeU k1 k2 v f : + free k1 (upd k2 v f) = + if k1 == k2 then + if cond U k2 then free k1 f else um_undef + else upd k2 v (free k1 f). +Proof. +rewrite !umEX /UM.free /UM.upd /cond. +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 k1 (free k2 f) = + if k1 == k2 then free k1 f else free k2 (free k1 f). +Proof. +rewrite !umEX /UM.free. +by case: (UMC.from f)=>[|f' F]; case: ifP=>// E; rewrite !umEX rem_rem E. +Qed. + +Lemma freeUn k f1 f2 : + free k (f1 \+ f2) = + if k \in dom (f1 \+ f2) then free k f1 \+ free k f2 + else f1 \+ f2. +Proof. +rewrite !pcmE /= !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 k (f1 \+ f2) = free k f1 \+ free k f2. +Proof. +move=>V'; rewrite freeUn domUn V' /=; case: ifP=>//. +by move/negbT; rewrite negb_or; case/andP=>H1 H2; rewrite !dom_free. +Qed. + +Lemma freeUnL k f1 f2 : k \notin dom f1 -> free k (f1 \+ f2) = f1 \+ free k f2. +Proof. +move=>V1; rewrite freeUn domUn inE (negbTE V1) /=. +case: ifP; first by case/andP; rewrite dom_free. +move/negbT; rewrite negb_and; case/orP=>V2; last by rewrite dom_free. +suff: ~~ valid (f1 \+ free k f2) by move/invalidE: V2=>-> /invalidE ->. +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 freeUnR k f1 f2 : k \notin dom f2 -> free k (f1 \+ f2) = free k f1 \+ f2. +Proof. by move=>H; rewrite joinC freeUnL // joinC. Qed. + +Lemma free_umfilt p k f : + free k (um_filter p f) = + if p k then um_filter p (free k f) else um_filter p f. +Proof. +rewrite !umEX /UM.free /UM.um_filter. +case: (UMC.from f)=>[|f' F]; case: ifP=>// E; rewrite !umEX. +- by rewrite kfilt_rem E. +by rewrite rem_kfilt E. +Qed. + +End FreeLemmas. + + +(********) +(* find *) +(********) + +Section FindLemmas. +Variables (K : ordType) (V : Type) (U : union_map_class K V). +Implicit Types (k : K) (v : V) (f : U). + +Lemma find0E k : find k (Unit : U) = None. +Proof. by rewrite pcmE /= !umEX /UM.find /= fnd_empty. Qed. + +Lemma find_undef k : find k (um_undef : U) = None. +Proof. by rewrite !umEX /UM.find. Qed. + +Lemma find_cond k f : ~~ cond U k -> find k f = None. +Proof. +simpl. +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 cond U k2 && valid f then Some v else None + else if cond U k2 then find k1 f else None. +Proof. +rewrite pcmE /= !umEX /UM.valid /UM.find /UM.upd /cond. +case: (UMC.from f)=>[|f' F]; first by rewrite andbF !if_same. +case: decP=>H; first by rewrite H /= fnd_ins. +by rewrite (introF idP H) /= if_same. +Qed. + +Lemma findF k1 k2 f : + find k1 (free k2 f) = if k1 == k2 then None else find k1 f. +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 !pcmE /= !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 !pcmE /= !umEX /UM.valid /UM.find -{2 3}[f1]UMC.tfE -{2 3}[f2]UMC.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. + +Lemma find_umfilt p k f : + find k (um_filter p f) = if p k then find k f else None. +Proof. +rewrite !umEX /UM.find /UM.um_filter. +case: (UMC.from f)=>[|f' F]; first by rewrite if_same. +by rewrite fnd_kfilt. +Qed. + +End FindLemmas. + +(* if maps store units, i.e., we keep them just for sets *) +(* then we can simplify the find_eta lemma a bit *) + +Lemma domeq_eta (K : ordType) (U : union_map_class K unit) (f1 f2 : U) : + 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. + +(********) +(* empb *) +(********) + +Section EmpbLemmas. +Variables (K : ordType) (V : Type) (U : union_map_class K V). +Implicit Types (k : K) (v : V) (f : U). + +Lemma empb_undef : empb (um_undef : U) = false. +Proof. by rewrite !umEX. Qed. + +Lemma empbP f : reflect (f = Unit) (empb f). +Proof. +rewrite pcmE /= !umEX /UM.empty /UM.empb -{1}[f]UMC.tfE. +case: (UMC.from f)=>[|f' F]; first by apply: ReflectF; rewrite !umEX. +case: eqP=>E; constructor; rewrite !umEX. +- by move/supp_nilE: E=>->. +by move=>H; rewrite H in E. +Qed. + +Lemma empbU k v f : empb (upd k v f) = false. +Proof. +rewrite !umEX /UM.empb /UM.upd. +case: (UMC.from f)=>[|f' F] //; case: decP=>// H. +suff: k \in supp (ins k v f') by case: (supp _). +by rewrite supp_ins inE /= eq_refl. +Qed. + +Lemma empbUn f1 f2 : empb (f1 \+ f2) = empb f1 && empb f2. +Proof. +rewrite !pcmE /= !umEX /UM.empb /UM.union. +case: (UMC.from f1) (UMC.from f2)=>[|f1' F1][|f2' F2] //. +- by rewrite andbF. +case: ifP=>E; case: eqP=>E1; case: eqP=>E2 //; last 2 first. +- by move: E2 E1; move/supp_nilE=>->; rewrite fcat0s; case: eqP. +- by move: E1 E2 E; do 2![move/supp_nilE=>->]; case: disjP. +- by move/supp_nilE: E2 E1=>-> <-; rewrite fcat0s eq_refl. +have [k H3]: exists k, k \in supp f1'. +- case: (supp f1') {E1 E} E2=>[|x s _] //. + by exists x; rewrite inE eq_refl. +suff: k \in supp (fcat f1' f2') by rewrite E1. +by rewrite supp_fcat inE /= H3. +Qed. + +(* some transformation lemmas *) + +Lemma empbE f : f = Unit <-> empb f. +Proof. by case: empbP. Qed. + +Lemma empb0 : empb (Unit : U). +Proof. by apply/empbE. Qed. + +Lemma join0E f1 f2 : f1 \+ f2 = Unit <-> f1 = Unit /\ f2 = Unit. +Proof. by rewrite !empbE empbUn; case: andP. Qed. + +Lemma validEb f : reflect (valid f /\ forall k, k \notin dom f) (empb f). +Proof. +case: empbP=>E; constructor; first by rewrite E valid_unit dom0. +case=>V' H; apply: E; move: V' H. +rewrite !pcmE /= !umEX /UM.valid /UM.dom /UM.empty -{3}[f]UMC.tfE. +case: (UMC.from f)=>[|f' F] // _ H; rewrite !umEX. +apply/supp_nilE; case: (supp f') H=>// x s /(_ x). +by rewrite inE eq_refl. +Qed. + +Lemma validUnEb f : valid (f \+ f) = empb f. +Proof. +case E: (empb f); first by move/empbE: E=>->; rewrite unitL valid_unit. +case: validUn=>// V' _ L; case: validEb E=>//; case; split=>// k. +by case E: (k \in dom f)=>//; move: (L k E); rewrite E. +Qed. + +Lemma empb_umfilt p f : empb f -> empb (um_filter p f). +Proof. +rewrite !umEX /UM.empb /UM.um_filter. +case: (UMC.from f)=>[|f' F] //. +by move/eqP/supp_nilE=>->; rewrite kfilt_nil. +Qed. + +End EmpbLemmas. + + +(*********) +(* undef *) +(*********) + +Section UndefLemmas. +Variables (K : ordType) (V : Type) (U : union_map_class K V). +Implicit Types (k : K) (v : V) (f : U). + +Lemma undefb_undef : undefb (um_undef : U). +Proof. by rewrite !umEX. Qed. + +Lemma undefbP f : reflect (f = um_undef) (undefb f). +Proof. +rewrite !umEX /UM.undefb -{1}[f]UMC.tfE. +by case: (UMC.from f)=>[|f' F]; constructor; rewrite !umEX. +Qed. + +Lemma undefbE f : f = um_undef <-> undefb f. +Proof. by case: undefbP. Qed. + +Lemma join_undefL f : um_undef \+ f = um_undef. +Proof. by rewrite /PCM.join /= !umEX. Qed. + +Lemma join_undefR f : f \+ um_undef = um_undef. +Proof. by rewrite joinC join_undefL. Qed. + +End UndefLemmas. + + +(**********) +(* um_all *) +(**********) + +(* this is a defined notion; so it does not appear in the signature *) + +Section AllDefLemmas. +Variables (K : ordType) (V : Type) (U : union_map_class K V) (P : V -> Prop). +Implicit Types (k : K) (v : V) (f : U). + +Definition um_all f := forall k v, find k f = Some v -> P v. + +Lemma umall_undef : um_all um_undef. +Proof. by move=>k v; rewrite find_undef. Qed. + +Hint Resolve umall_undef. + +Lemma umall0 : um_all Unit. +Proof. by move=>k v; rewrite find0E. Qed. + +Lemma umallUn f1 f2 : um_all f1 -> um_all f2 -> um_all (f1 \+ f2). +Proof. +case W : (valid (f1 \+ f2)); last by move/invalidE: (negbT W)=>->. +by move=>X Y k v; rewrite findUnL //; case: ifP=>_; [apply: X|apply: Y]. +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 k v); rewrite findUnL // (find_some F). +Qed. + +Lemma umallUnR f1 f2 : valid (f1 \+ f2) -> um_all (f1 \+ f2) -> um_all f2. +Proof. by rewrite joinC; apply: umallUnL. Qed. + +End AllDefLemmas. + +Hint Resolve umall_undef umall0. + + +(********************************************) +(* Interaction of subset, dom and um_filter *) +(********************************************) + +Section Interaction. +Variables (K : ordType) (A : Type) (U : union_map_class K A). +Implicit Types (x y a b : U) (p q : pred K). + +(* filter p x is lower bound for p *) +Lemma subdom_umfilt x p : {subset dom (um_filter p x) <= p}. +Proof. by move=>a; rewrite dom_umfilt; case/andP. Qed. + + +(* Some equivalent forms for subset with dom *) +Lemma subdom_indomE x p : {subset dom x <= p} = {in dom x, p =1 predT}. +Proof. by []. Qed. + +Lemma subdom_umfiltE x p : {subset dom x <= p} <-> um_filter p x = x. +Proof. +split; last by move=><- a; rewrite dom_umfilt; case/andP. +by move/eq_in_umfilt=>->; rewrite umfilt_predT. +Qed. + +(* specific consequence of subdom_umfiltE *) +Lemma umfilt_memdomE x : um_filter (mem (dom x)) x = x. +Proof. by apply/subdom_umfiltE. Qed. + + +(* 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 subset_disjC p q : {subset p <= predC q} <-> {subset q <= predC p}. +Proof. by split=>H a; apply: contraL (H a). 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_umfilt0 x p : + valid x -> {subset dom x <= predC p} <-> um_filter p x = Unit. +Proof. +move=>V; split=>H. +- by rewrite (eq_in_umfilt (p2:=pred0)) ?umfilt_pred0 // => a /H /negbTE ->. +move=>k X; case: dom_find X H=>// a _ -> _; rewrite umfiltU !inE. +by case: ifP=>// _ /empbE; rewrite /= empbU. +Qed. + +End Interaction. + + +(************************) +(* PCM-induced ordering *) +(************************) + +(* Union maps and pcm ordering. *) + +Section UmpleqLemmas. +Variables (K : ordType) (A : Type) (U : union_map_class K A). +Implicit Types (x y a b : U) (p : pred K). + +Lemma umpleq_undef x : [pcm x <= um_undef]. +Proof. by exists um_undef; rewrite join_undefR. Qed. + +Hint Resolve umpleq_undef. + +(* pcm-induced preorder, is an order in the case of union maps *) + +Lemma umpleq_asym x y : [pcm x <= y] -> [pcm y <= x] -> x = y. +Proof. +case=>a -> [b]; case V : (valid x); last first. +- by move/invalidE: (negbT V)=>->; rewrite join_undefL. +rewrite -{1 2}[x]unitR -joinA in V *. +by case/(joinxK V)/esym/join0E=>->; rewrite unitR. +Qed. + +(* monotonicity lemmas *) + +Lemma umpleq_filt2 x y p : [pcm x <= y] -> [pcm um_filter p x <= um_filter p y]. +Proof. +move=>H; case V : (valid y). +- by case: H V=>a -> V; rewrite umfiltUn //; eexists _. +by move/invalidE: (negbT V)=>->; rewrite umfilt_undef; apply: umpleq_undef. +Qed. + +(* filter p x is lower bound for x *) +Lemma umpleq_filtI x p : [pcm um_filter p x <= x]. +Proof. +exists (um_filter (predD predT p) x); rewrite -umfilt_predU. +by rewrite -{1}[x]umfilt_predT; apply: eq_in_umfilt=>a; rewrite /= orbT. +Qed. + +Hint Resolve umpleq_filtI. + +(* in valid case, we can define the order by means of filter *) +Lemma umpleq_filtE a x : + valid x -> [pcm a <= x] <-> um_filter (mem (dom a)) x = a. +Proof. by move=>V; split=>[|<-] // H; case: H V=>b ->; apply: umfilt_dom. Qed. + +(* filter p x is largest lower bound for x and p *) +Lemma umpleq_filt_meet a x p : + {subset dom a <= p} -> [pcm a <= x] -> [pcm a <= um_filter p x]. +Proof. by move=>D /(umpleq_filt2 p); rewrite (eq_in_umfilt D) umfilt_predT. Qed. + +(* join is the least upper bound *) +Lemma umpleq_join x a b : + valid (a \+ b) -> [pcm a <= x] -> [pcm b <= x] -> [pcm a \+ b <= x]. +Proof. +case Vx : (valid x); last by move/invalidE: (negbT Vx)=>->. +move=>V /(umpleq_filtE _ Vx) <- /(umpleq_filtE _ Vx) <-. +by rewrite -umfilt_dpredU //; apply: valid_subdom V. +Qed. + +(* x <= y and subdom *) + +Lemma umpleq_subdom x y : valid y -> [pcm x <= y] -> {subset dom x <= dom y}. +Proof. by move=>V H; case: H V=>a -> V b D; rewrite domUn inE V D. Qed. + +Lemma subdom_umpleq a x y : + valid (x \+ y) -> [pcm a <= x \+ y] -> + {subset dom a <= dom x} -> [pcm a <= x]. +Proof. +move=>V H Sx; move: (umpleq_filt_meet Sx H); rewrite umfiltUn //. +rewrite umfilt_memdomE; move/subset_disjC: (valid_subdom V). +by move/(subdom_umfilt0 _ (validR V))=>->; rewrite unitR. +Qed. + +(* meet is the greatest lower bound *) +Lemma umpleq_meet a x y1 y2 : + valid (x \+ y1 \+ y2) -> + [pcm a <= x \+ y1] -> [pcm a <= x \+ y2] -> [pcm a <= x]. +Proof. +move=>V H1 H2. +have {V} [V1 V V2] : [/\ valid (x \+ y1), valid (y1 \+ y2) & valid (x \+ y2)]. +- rewrite (validL V); rewrite -joinA in V; rewrite (validR V). + by rewrite joinA joinAC in V; rewrite (validL V). +apply: subdom_umpleq (V1) (H1) _ => k D. +move: {D} (umpleq_subdom V1 H1 D) (umpleq_subdom V2 H2 D). +rewrite !domUn !inE V1 V2 /=; case : (k \in dom x)=>//=. +by case: validUn V=>// _ _ L _ /L /negbTE ->. +Qed. + +(* some/none lemmas *) + +Lemma umpleq_some x1 x2 t s : + valid x2 -> [pcm x1 <= x2] -> find t x1 = Some s -> find t x2 = Some s. +Proof. by move=>V H; case: H V=>a -> V H; rewrite findUnL // (find_some H). Qed. + +Lemma umpleq_none x1 x2 t : + valid x2 -> [pcm x1 <= x2] -> find t x2 = None -> find t x1 = None. +Proof. by case E: (find t x1)=>[a|] // V H <-; rewrite (umpleq_some V H E). Qed. + +End UmpleqLemmas. + + +(****************************) +(* Generic points-to lemmas *) +(****************************) + +(* Instances of union_maps have different definition of points-to, so +they need separate intances of each of following lemmas. I give the +lemmas complicated names prefixed by gen, because they are not +supposed to be used in applications *) + +Section PointsToLemmas. +Variables (K : ordType) (V : Type) (U : union_map_class K V). +Implicit Types (k : K) (v : V) (f : U). + +Lemma ptsU k v : pts k v = upd k v Unit :> U. +Proof. by rewrite !pcmE /= !umEX /UM.pts /UM.upd; case: decP. Qed. + +Lemma domPtK k v : dom (pts k v : U) = if cond U k then [:: k] else [::]. +Proof. +rewrite !umEX /= /UM.dom /supp /UM.pts /UM.upd /UM.empty /=. +by case D : (decP _)=>[a|a] /=; rewrite ?a ?(introF idP a). +Qed. + +(* a weaker version of gen_domPtK, but legacy *) +Lemma domPt k v : dom (pts k v : U) =i [pred x | cond U k & k == x]. +Proof. +by move=>x; rewrite ptsU domU !inE eq_sym valid_unit dom0; case: eqP. +Qed. + +Lemma validPt k v : valid (pts k v : U) = cond U k. +Proof. by rewrite ptsU validU valid_unit andbT. Qed. + +Lemma domeqPt k v1 v2 : dom_eq (pts k v1 : U) (pts k v2). +Proof. by apply/domeqP; rewrite !validPt; split=>// x; rewrite !domPt. Qed. + +Lemma findPt k v : find k (pts k v : U) = if cond U k then Some v else None. +Proof. by rewrite ptsU findU eq_refl /= valid_unit andbT. Qed. + +Lemma findPt2 k1 k2 v : + find k1 (pts k2 v : U) = + if (k1 == k2) && cond U k2 then Some v else None. +Proof. +by rewrite ptsU findU valid_unit andbT find0E; case: ifP=>//=; case: ifP. +Qed. + +Lemma findPt_inv k1 k2 v w : + find k1 (pts k2 v : U) = Some w -> [/\ k1 = k2, v = w & cond U k2]. +Proof. +rewrite ptsU findU; case: eqP; last by case: ifP=>//; rewrite find0E. +by move=>->; rewrite valid_unit andbT; case: ifP=>// _ [->]. +Qed. + +Lemma freePt2 k1 k2 v : + cond U k2 -> + free k1 (pts k2 v : U) = if k1 == k2 then Unit else pts k2 v. +Proof. by move=>N; rewrite ptsU freeU free0 N. Qed. + +Lemma freePt k v : + cond U k -> free k (pts k v : U) = 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 empbPt k v : empb (pts k v : U) = false. +Proof. by rewrite ptsU empbU. Qed. + +(* valid *) + +Lemma validPtUn k v f : + valid (pts k v \+ f) = [&& cond U 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) = [&& cond U k, valid f & k \notin dom f]. +Proof. by rewrite joinC; apply: validPtUn. Qed. + +(* the projections from validPtUn are often useful *) + +Lemma validPtUn_cond k v f : valid (pts k v \+ f) -> cond U k. +Proof. by rewrite validPtUn; case/and3P. Qed. + +Lemma validUnPt_cond k v f : valid (f \+ pts k v) -> cond U k. +Proof. by rewrite joinC; apply: validPtUn_cond. Qed. + +Lemma validPtUnV k v f : valid (pts k v \+ f) -> valid f. +Proof. by rewrite validPtUn; case/and3P. Qed. + +Lemma validUnPtV k v f : valid (f \+ pts k v) -> valid f. +Proof. by rewrite joinC; apply: validPtUnV. Qed. + +Lemma 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: (UMC.p _ 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 validxx f : valid (f \+ f) -> dom f =i pred0. +Proof. by case: validUn=>// _ _ L _ z; case: (_ \in _) (L z)=>//; elim. Qed. + +(* dom_eq *) + +Lemma domeqPtUn k v1 v2 f1 f2 : + dom_eq f1 f2 -> dom_eq (pts k v1 \+ f1) (pts k v2 \+ f2). +Proof. by apply: domeqUn=>//; apply: domeqPt. Qed. + +Lemma domeqUnPt k v1 v2 f1 f2 : + dom_eq f1 f2 -> dom_eq (f1 \+ pts k v1) (f2 \+ pts k v2). +Proof. by rewrite (joinC f1) (joinC f2); apply: domeqPtUn. 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 k1 (pts k2 v \+ f) = + if k1 == k2 then f else pts k2 v \+ free k1 f. +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 k1 (f \+ pts k2 v) = + if k1 == k2 then f else free k1 f \+ pts k2 v. +Proof. by rewrite !(joinC _ (pts k2 v)); apply: freePtUn2. Qed. + +Lemma freePtUn k v f : + valid (pts k v \+ f) -> free k (pts k v \+ f) = f. +Proof. by move=>V'; rewrite freePtUn2 // eq_refl. Qed. + +Lemma freeUnPt k v f : + valid (f \+ pts k v) -> free k (f \+ pts k v) = f. +Proof. by rewrite joinC; apply: freePtUn. 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_invalid. +Qed. + +Lemma updUnPt k v1 v2 f : upd k v1 (f \+ pts k v2) = f \+ pts k v1. +Proof. by rewrite !(joinC f); apply: updPtUn. Qed. + +(* empb *) + +Lemma empbPtUn k v f : empb (pts k v \+ f) = false. +Proof. by rewrite empbUn empbPt. Qed. + +Lemma empbUnPt k v f : empb (f \+ pts k v) = false. +Proof. by rewrite joinC; apply: empbPtUn. Qed. + +(* undef *) + +Lemma pts_undef k v1 v2 : pts k v1 \+ pts k v2 = um_undef :> U. +Proof. +apply/invalidE; rewrite validPtUn validPt domPt !negb_and negb_or eq_refl. +by case: (cond U k). +Qed. + +(* um_filter *) + +Lemma umfiltPt p k v : + um_filter p (pts k v : U) = + if p k then pts k v else if cond U k then Unit else um_undef. +Proof. by rewrite ptsU umfiltU umfilt0. Qed. + +Lemma umfiltPtUn p k v f : + um_filter p (pts k v \+ f) = + if valid (pts k v \+ f) then + if p k then pts k v \+ um_filter p f else um_filter p f + else um_undef. +Proof. +case: ifP=>X; last by move/invalidE: (negbT X)=>->; rewrite umfilt_undef. +rewrite umfiltUn // umfiltPt (validPtUn_cond X). +by case: ifP=>//; rewrite unitL. +Qed. + +(* um_all *) + +Lemma umallPt (P : V -> Prop) k v : P v -> um_all P (pts k v : U). +Proof. by move=>X u w /findPt_inv [_ <-]. Qed. + +Lemma umallPtUn (P : V -> Prop) k v f : + P v -> um_all P f -> um_all P (pts k v \+ f). +Proof. by move/(umallPt (k:=k)); apply: umallUn. Qed. + +Lemma umallPtE (P : V -> Prop) k v : cond U k -> um_all P (pts k v : U) -> P v. +Proof. by move=>C /(_ k v); rewrite findPt C; apply. Qed. + +Lemma umallPtUnE (P : V -> Prop) k v f : + valid (pts k v \+ f) -> um_all P (pts k v \+ f) -> P v /\ um_all P f. +Proof. +move=>W H; move: (umallUnL W H) (umallUnR W H)=>{H} H1 H2. +by split=>//; apply: umallPtE H1; apply: validPtUn_cond W. +Qed. + +(* others *) + +Lemma um_eta k f : + k \in dom f -> exists v, find k f = Some v /\ f = pts k v \+ free k f. +Proof. +case: dom_find=>// v E1 E2 _; exists v; split=>//. +by rewrite {1}E2 -{1}[free k f]unitL updUnR domF inE /= eq_refl ptsU. +Qed. + +Lemma um_eta2 k v f : + find k f = Some v -> f = pts k v \+ free k f. +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) (validR V'). +Qed. + +Lemma cancel2 k1 k2 f1 f2 v1 v2 : + valid (pts k1 v1 \+ f1) -> + pts k1 v1 \+ f1 = pts k2 v2 \+ f2 -> + if k1 == k2 then v1 = v2 /\ f1 = f2 + else [/\ free k1 f2 = free k2 f1, + f1 = pts k2 v2 \+ free k1 f2 & + f2 = pts k1 v1 \+ free k2 f1]. +Proof. +move=>V1 E; case: ifP=>X. +- by rewrite -(eqP X) in E; case/(cancel V1): E. +move: (V1); rewrite E => V2. +have E' : f2 = pts k1 v1 \+ free k2 f1. +- move: (erefl (free k2 (pts k1 v1 \+ f1))). + rewrite {2}E freeUn E domPtUn inE V2 eq_refl /=. + by rewrite ptsU freeU eq_sym X free0 -ptsU freePtUn. +suff E'' : free k2 f1 = free k1 f2. +- by rewrite -E''; rewrite E' joinCA in E; case/(cancel V1): E. +move: (erefl (free k1 f2)). +by rewrite {1}E' freePtUn // -E'; apply: (validR V2). +Qed. + +(* induction over union maps, expressed with pts and \+ *) + +(* forward progressing over keys *) +Lemma um_indf (P : U -> Prop) : + P um_undef -> P Unit -> + (forall k v f, P f -> valid (pts k v \+ f) -> + path ord k (dom f) -> + P (pts k v \+ f)) -> + forall f, P f. +Proof. +rewrite !pcmE /= !umEX => H1 H2 H3 f; rewrite -[f]UMC.tfE. +apply: (UM.base_indf (P := (fun b => P (UMC.to b))))=>//. +move=>k v g H V1; move: (H3 k v _ H); rewrite !pcmE /= !umEX. +by apply. +Qed. + +(* backward progressing over keys *) +Lemma um_indb (P : U -> Prop) : + P um_undef -> P Unit -> + (forall k v f, P f -> valid (pts k v \+ f) -> + (forall x, x \in dom f -> ord x k) -> + P (pts k v \+ f)) -> + forall f, P f. +Proof. +rewrite !pcmE /= !umEX => H1 H2 H3 f; rewrite -[f]UMC.tfE. +apply: (UM.base_indb (P := (fun b => P (UMC.to b))))=>//. +move=>k v g H V1; move: (H3 k v _ H); rewrite !pcmE /= !umEX. +by apply. +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/idP. +- move=>W; rewrite (validL W). + rewrite -joinA in W; rewrite (validR W). + by rewrite joinCA in W; rewrite (validR W). +case/and3P=>V1 V2 V3; case: validUn=>//. +- by rewrite V1. +- by rewrite (validR V2). +move=>k; rewrite domUn inE V1; case/orP. +- by case: validUn V3=>// _ _ X _ /X /negbTE ->. +by case: validUn V2=>// _ _ X _ /X /negbTE ->. +Qed. + +(* points-to is a prime element of the monoid *) +Lemma um_prime f1 f2 k v : + cond U k -> + f1 \+ f2 = pts k v -> + f1 = pts k v /\ f2 = Unit \/ + f1 = Unit /\ f2 = pts k v. +Proof. +move: f1 f2; apply: um_indf=>[f1|f2 _|k2 w2 f1 _ _ _ f X E]. +- rewrite join_undefL -(validPt _ v)=>W E. + 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=>-> [->] /join0E [->->]; rewrite unitR; left. +by move=>_ [_ _] /esym/empbP; rewrite empbPtUn. +Qed. + +End PointsToLemmas. + +Hint Resolve domeqPt domeqPtUn domeqUnPt. +Prenex Implicits validPtUn_cond findPt_inv um_eta2. + +(********************) +(* Topped structure *) +(********************) + +Module UnionClassTPCM. +Section UnionClassTPCM. +Variables (K : ordType) (V : Type) (U : union_map_class K V). +Implicit Type f : U. + +Lemma join0E f1 f2 : f1 \+ f2 = Unit -> f1 = Unit /\ f2 = Unit. +Proof. by rewrite join0E. Qed. + +Lemma valid_undefN : ~~ valid (um_undef: U). +Proof. by rewrite valid_undef. Qed. + +Lemma undef_join f : um_undef \+ f = um_undef. +Proof. by rewrite join_undefL. Qed. +End UnionClassTPCM. + +Module Exports. +Definition union_map_classTPCMMix K V (U : union_map_class K V) := + TPCMMixin (@empbP K V U) (@join0E K V U) + (@valid_undefN K V U) (@undef_join _ _ U). +Canonical union_map_classTPCM K V (U : union_map_class K V) := + Eval hnf in TPCM _ (@union_map_classTPCMMix K V U). +End Exports. +End UnionClassTPCM. + +Export UnionClassTPCM.Exports. + + +(********************************************************) +(********************************************************) +(* Instance of union maps with trivially true condition *) +(********************************************************) +(********************************************************) + +(* We build it over the base type with a trival condition on keys. We +seal the definition to avoid any slowdowns (just in case). *) + +Module Type UMSig. +Parameter tp : ordType -> Type -> Type. + +Section Params. +Variables (K : ordType) (V : Type). +Notation tp := (tp K V). + +Parameter um_undef : tp. +Parameter defined : tp -> bool. +Parameter empty : tp. +Parameter upd : K -> V -> tp -> tp. +Parameter dom : tp -> seq K. +Parameter dom_eq : tp -> tp -> bool. +Parameter free : K -> tp -> tp. +Parameter find : K -> tp -> option V. +Parameter union : tp -> tp -> tp. +Parameter um_filter : pred K -> tp -> tp. +Parameter empb : tp -> bool. +Parameter undefb : tp -> bool. +Parameter pts : K -> V -> tp. + +Parameter from : tp -> @UM.base K V predT. +Parameter to : @UM.base K V predT -> tp. + +Axiom ftE : forall b, from (to b) = b. +Axiom tfE : forall f, to (from f) = f. +Axiom undefE : um_undef = to (@UM.Undef K V predT). +Axiom defE : forall f, defined f = UM.valid (from f). +Axiom emptyE : empty = to (@UM.empty K V predT). +Axiom updE : forall k v f, upd k v f = to (UM.upd k v (from f)). +Axiom domE : forall f, dom f = UM.dom (from f). +Axiom dom_eqE : forall f1 f2, dom_eq f1 f2 = UM.dom_eq (from f1) (from f2). +Axiom freeE : forall k f, free k f = to (UM.free k (from f)). +Axiom findE : forall k f, find k f = UM.find k (from f). +Axiom unionE : forall f1 f2, union f1 f2 = to (UM.union (from f1) (from f2)). +Axiom umfiltE : forall q f, um_filter q f = to (UM.um_filter q (from f)). +Axiom empbE : forall f, empb f = UM.empb (from f). +Axiom undefbE : forall f, undefb f = UM.undefb (from f). +Axiom ptsE : forall k v, pts k v = to (@UM.pts K V predT k v). + +End Params. + +End UMSig. + + +Module UMDef : UMSig. +Section UMDef. +Variables (K : ordType) (V : Type). + +Definition tp : Type := @UM.base K V predT. + +Definition um_undef := @UM.Undef K V predT. +Definition defined f := @UM.valid K V predT f. +Definition empty := @UM.empty K V predT. +Definition upd k v f := @UM.upd K V predT k v f. +Definition dom f := @UM.dom K V predT f. +Definition dom_eq f1 f2 := @UM.dom_eq K V predT f1 f2. +Definition free k f := @UM.free K V predT k f. +Definition find k f := @UM.find K V predT k f. +Definition union f1 f2 := @UM.union K V predT f1 f2. +Definition um_filter q f := @UM.um_filter K V predT q f. +Definition empb f := @UM.empb K V predT f. +Definition undefb f := @UM.undefb K V predT f. +Definition pts k v := @UM.pts K V predT k v. + +Definition from (f : tp) : @UM.base K V predT := f. +Definition to (b : @UM.base K V predT) : tp := b. + +Lemma ftE b : from (to b) = b. Proof. by []. Qed. +Lemma tfE f : to (from f) = f. Proof. by []. Qed. +Lemma undefE : um_undef = to (@UM.Undef K V predT). Proof. by []. Qed. +Lemma defE f : defined f = UM.valid (from f). Proof. by []. Qed. +Lemma emptyE : empty = to (@UM.empty K V predT). Proof. by []. Qed. +Lemma updE k v f : upd k v f = to (UM.upd k v (from f)). Proof. by []. Qed. +Lemma domE f : dom f = UM.dom (from f). Proof. by []. Qed. +Lemma dom_eqE f1 f2 : dom_eq f1 f2 = UM.dom_eq (from f1) (from f2). +Proof. by []. Qed. +Lemma freeE k f : free k f = to (UM.free k (from f)). Proof. by []. Qed. +Lemma findE k f : find k f = UM.find k (from f). Proof. by []. Qed. +Lemma unionE f1 f2 : union f1 f2 = to (UM.union (from f1) (from f2)). +Proof. by []. Qed. +Lemma umfiltE q f : um_filter q f = to (UM.um_filter q (from f)). +Proof. by []. Qed. +Lemma empbE f : empb f = UM.empb (from f). Proof. by []. Qed. +Lemma undefbE f : undefb f = UM.undefb (from f). Proof. by []. Qed. +Lemma ptsE k v : pts k v = to (@UM.pts K V predT k v). Proof. by []. Qed. + +End UMDef. +End UMDef. + +Notation union_map := UMDef.tp. + +Definition unionmapMixin K V := + UMCMixin (@UMDef.ftE K V) (@UMDef.tfE K V) (@UMDef.defE K V) + (@UMDef.undefE K V) (@UMDef.emptyE K V) (@UMDef.updE K V) + (@UMDef.domE K V) (@UMDef.dom_eqE K V) (@UMDef.freeE K V) + (@UMDef.findE K V) (@UMDef.unionE K V) (@UMDef.umfiltE K V) + (@UMDef.empbE K V) (@UMDef.undefbE K V) (@UMDef.ptsE K V). + +Canonical union_mapUMC K V := + Eval hnf in UMC (union_map K V) (@unionmapMixin K V). + +(* we add the canonical projections matching against naked type *) +(* thus, there are two ways to get a PCM for a union map: *) +(* generic one going through union_map_classPCM, and another by going *) +(* through union_mapPCM. Luckily, they produce equal results *) +(* and this is just a matter of convenience *) +(* Ditto for the equality type *) + +Definition union_mapPCMMix K V := + union_map_classPCMMix (union_mapUMC K V). +Canonical union_mapPCM K V := + Eval hnf in PCM (union_map K V) (@union_mapPCMMix K V). + +Definition union_mapCPCMMix K V := + union_map_classCPCMMix (@union_mapUMC K V). +Canonical union_mapCPCM K V := + Eval hnf in CPCM (union_map K V) (@union_mapCPCMMix K V). + +Definition union_mapTPCMMix K V := + union_map_classTPCMMix (@union_mapUMC K V). +Canonical union_mapTPCM K V := + Eval hnf in TPCM (union_map K V) (@union_mapTPCMMix K V). + +Definition union_map_eqMix K (V : eqType) := + @union_map_class_eqMix K V _ _ (@unionmapMixin K V). +Canonical union_map_eqType K (V : eqType) := + Eval hnf in EqType (union_map K V) (@union_map_eqMix K V). + + +Definition um_pts (K : ordType) V (k : K) (v : V) := + @UMC.pts K V (@union_mapUMC K V) k v. + +Notation "x \\-> v" := (@um_pts _ _ x v) (at level 30). + +(* Finite sets are just union maps with unit range *) +Notation fset T := (@union_map T unit). +Notation "# x" := (x \\-> tt) (at level 20). + +(* Does the notation work? *) +Lemma xx : 1 \\-> true = 1 \\-> false. +Abort. + +(* does the pcm and the equality type work? *) +Lemma xx : ((1 \\-> true) \+ (2 \\-> false)) == (1 \\-> false). +simpl. +rewrite joinC. +Abort. + +(* can we use the base type? *) +Lemma xx (x : union_map nat_ordType nat) : x \+ x == x \+ x. +Abort. + +(* the default dom' lemmas don't work nicely *) +Lemma xx (f : union_map nat_ordType nat) : 3 \in dom (free 2 f). +try by rewrite domF' /=. +rewrite (@domF _ _ (union_mapUMC _ _)). +Abort. + +(* but the dom lemmas do work nicely *) +Lemma xx (f : union_map nat_ordType nat) : 3 \in dom (free 2 f). +rewrite domF /=. +Abort. + +Lemma xx : 1 \\-> (1 \\-> 3) = 2 \\-> (7 \\-> 3). +Abort. + + +(* Since the union map constructors are opaque, the decidable equality *) +(* does not simplify automatically; we need to export explicit equations *) +(* for simplification *) +(* so far, the ones I need are really just for points-to *) + +Section UMDecidableEquality. +Variables (K : ordType) (V : eqType) (U : union_map_class K V). + +Lemma umPtPtE (k1 k2 : K) (v1 v2 : V) : + (k1 \\-> v1 == k2 \\-> v2) = (k1 == k2) && (v1 == v2). +Proof. +rewrite {1}/eq_op /= /UnionMapEq.unionmap_eq /um_pts !umEX /=. +by rewrite {1}/eq_op /= /feq eqseq_cons andbT. +Qed. + +Lemma umPt0E (k : K) (v : V) : (k \\-> v == Unit) = false. +Proof. by apply: (introF idP)=>/eqP/empbP; rewrite empbPt. Qed. + +Lemma um0PtE (k : K) (v : V) : + (@Unit [pcm of union_map K V] == k \\-> v) = false. +Proof. by rewrite eq_sym umPt0E. Qed. + +Lemma umPtUndefE (k : K) (v : V) : (k \\-> v == um_undef) = false. +Proof. by rewrite /eq_op /= /UnionMapEq.unionmap_eq /um_pts !umEX. Qed. + +Lemma umUndefPtE (k : K) (v : V) : + ((um_undef : union_map_eqType K V) == k \\-> v) = false. +Proof. by rewrite eq_sym umPtUndefE. Qed. + +Lemma umUndef0E : ((um_undef : union_map_eqType K V) == Unit) = false. +Proof. by apply/(introF idP)=>/eqP/empbP; rewrite empb_undef. Qed. + +Lemma um0UndefE : ((Unit : union_mapPCM K V) == um_undef) = false. +Proof. by rewrite eq_sym umUndef0E. Qed. + +Lemma umPtUE (k : K) (v : V) f : (k \\-> v \+ f == Unit) = false. +Proof. by apply: (introF idP)=>/eqP/join0E/proj1/eqP; rewrite umPt0E. Qed. + +Lemma umUPtE (k : K) (v : V) f : (f \+ k \\-> v == Unit) = false. +Proof. by rewrite joinC umPtUE. Qed. + +Lemma umPtUPtE (k1 k2 : K) (v1 v2 : V) f : + (k1 \\-> v1 \+ f == k2 \\-> v2) = [&& k1 == k2, v1 == v2 & empb f]. +Proof. +apply/idP/idP; last first. +- by case/and3P=>/eqP -> /eqP -> /empbP ->; rewrite unitR. +move/eqP/(um_prime _); case=>//; case. +- move/(cancelPt2 _); rewrite validPt=>/(_ (erefl _)). + by case=>->->->; rewrite !eq_refl empb0. +by move/empbP; rewrite empbPt. +Qed. + +Lemma umPtPtUE (k1 k2 : K) (v1 v2 : V) f : + (k1 \\-> v1 == k2 \\-> v2 \+ f) = [&& k1 == k2, v1 == v2 & empb f]. +Proof. by rewrite eq_sym umPtUPtE (eq_sym k1) (eq_sym v1). Qed. + +Lemma umUPtPtE (k1 k2 : K) (v1 v2 : V) f : + (f \+ k1 \\-> v1 == k2 \\-> v2) = [&& k1 == k2, v1 == v2 & empb f]. +Proof. by rewrite joinC umPtUPtE. Qed. + +Lemma umPtUPt2E (k1 k2 : K) (v1 v2 : V) f : + (k1 \\-> v1 == f \+ k2 \\-> v2) = [&& k1 == k2, v1 == v2 & empb f]. +Proof. by rewrite joinC umPtPtUE. Qed. + +Definition umE := (umPtPtE, umPt0E, um0PtE, umPtUndefE, + umUndefPtE, um0UndefE, umUndef0E, + umPtUE, umUPtE, umPtUPtE, umPtPtUE, umUPtPtE, umPtUPt2E, + (* plus a bunch of safe simplifications *) + unitL, unitR, validPt, valid_unit, eq_refl, empb0, empbPt, + join_undefL, join_undefR, empb_undef). + +End UMDecidableEquality. + +