From 2d17970947e2c49b3cee9cbbd60a573df382f16e Mon Sep 17 00:00:00 2001 From: Karl Palmskog Date: Sat, 8 Sep 2018 23:11:42 -0500 Subject: [PATCH 1/7] port to fcsl-pcm --- .gitignore | 1 + .travis-ci.sh | 13 +- .travis.yml | 62 +- Core/Actions.v | 12 +- Core/Always.v | 46 +- Core/DepMaps.v | 31 +- Core/DiSeLExtraction.v | 2 +- Heaps/domain.v => Core/Domain.v | 62 +- Core/Freshness.v | 49 +- Core/HoareTriples.v | 37 +- Core/InductiveInv.v | 91 +- Core/InferenceRules.v | 21 +- Core/Injection.v | 93 +- Core/Make | 41 +- Core/NetworkSem.v | 12 +- Core/NewStatePredicates.v | 6 +- Core/Process.v | 22 +- Core/Protocols.v | 8 +- Core/Rely.v | 12 +- Core/State.v | 25 +- Core/StatePredicates.v | 46 +- Core/While.v | 10 +- Core/Worlds.v | 15 +- Core/opam | 9 +- Examples/Calculator/CalculatorClientLib.v | 41 +- Examples/Calculator/CalculatorExtraction.v | 4 +- Examples/Calculator/CalculatorInvariant.v | 34 +- Examples/Calculator/CalculatorProtocol.v | 23 +- Examples/Calculator/CalculatorServerLib.v | 42 +- .../Calculator/DelegatingCalculatorServer.v | 38 +- Examples/Calculator/SimpleCalculatorApp.v | 117 +- Examples/Calculator/SimpleCalculatorServers.v | 18 +- Examples/Greeter/Greeter.v | 62 +- Examples/LockResource/LockProtocol.v | 33 +- Examples/LockResource/ResourceProtocol.v | 22 +- Examples/Make | 2 +- Examples/Querying/QueryHooked.v | 178 +- Examples/Querying/QueryPlusTPC.v | 48 +- Examples/Querying/QueryProtocol.v | 18 +- Examples/TwoPhaseCommit/SimpleTPCApp.v | 51 +- Examples/TwoPhaseCommit/TwoPhaseCoordinator.v | 151 +- Examples/TwoPhaseCommit/TwoPhaseExtraction.v | 4 +- .../TwoPhaseCommit/TwoPhaseInductiveInv.v | 16 +- .../TwoPhaseCommit/TwoPhaseInductiveProof.v | 16 +- Examples/TwoPhaseCommit/TwoPhaseParticipant.v | 64 +- Examples/TwoPhaseCommit/TwoPhaseProtocol.v | 53 +- Heaps/Make | 15 - Heaps/Makefile | 19 - Heaps/coding.v | 844 ------- Heaps/descr | 1 - Heaps/finmap.v | 1335 ----------- Heaps/heap.v | 755 ------ Heaps/idynamic.v | 76 - Heaps/opam | 28 - Heaps/ordtype.v | 253 -- Heaps/pcm.v | 558 ----- Heaps/pperm.v | 298 --- Heaps/pred.v | 526 ----- Heaps/prelude.v | 167 -- Heaps/spcm.v | 255 --- Heaps/unionmap.v | 2026 ----------------- README.md | 19 +- _CoqProject | 18 +- 63 files changed, 944 insertions(+), 8010 deletions(-) rename Heaps/domain.v => Core/Domain.v (96%) delete mode 100644 Heaps/Make delete mode 100644 Heaps/Makefile delete mode 100644 Heaps/coding.v delete mode 100644 Heaps/descr delete mode 100644 Heaps/finmap.v delete mode 100644 Heaps/heap.v delete mode 100644 Heaps/idynamic.v delete mode 100644 Heaps/opam delete mode 100644 Heaps/ordtype.v delete mode 100644 Heaps/pcm.v delete mode 100644 Heaps/pperm.v delete mode 100644 Heaps/pred.v delete mode 100644 Heaps/prelude.v delete mode 100644 Heaps/spcm.v delete mode 100644 Heaps/unionmap.v diff --git a/.gitignore b/.gitignore index 1246aca..b1240d5 100644 --- a/.gitignore +++ b/.gitignore @@ -83,3 +83,4 @@ extraction/**/*.ml _build *.d.byte Makefile.coq.conf +.coqdeps.d diff --git a/.travis-ci.sh b/.travis-ci.sh index 53099f2..aeafd3f 100644 --- a/.travis-ci.sh +++ b/.travis-ci.sh @@ -1,15 +1,10 @@ +#!/usr/bin/env bash + set -ev -opam init --yes --no-setup eval $(opam config env) -opam repo add coq-released https://coq.inria.fr/opam/released -opam repo add coq-core-dev https://coq.inria.fr/opam/core-dev -opam repo add coq-extra-dev https://coq.inria.fr/opam/extra-dev - -opam pin add coq $COQ_VERSION --kind=version --yes -opam pin add coq-mathcomp-ssreflect $SSREFLECT_VERSION --kind=version --yes --verbose +opam update -opam pin add Heaps --yes --verbose -opam pin add Core --yes --verbose +opam pin add Core disel --yes --verbose make -j4 -C Examples diff --git a/.travis.yml b/.travis.yml index d21a7ba..46d51e3 100644 --- a/.travis.yml +++ b/.travis.yml @@ -1,16 +1,54 @@ -language: c -addons: - apt: - sources: - - avsm - packages: - - ocaml - - opam - - aspcud +language: generic +sudo: false + +services: + - docker + env: + global: + - THIS_REPO=disel matrix: - - COQ_VERSION=8.7.1 SSREFLECT_VERSION=1.6.4 -script: bash -ex .travis-ci.sh -sudo: false + - COQ_VERSION=coq8.7-32bit + - COQ_VERSION=coq8.8-32bit + +# The "docker run" command will pull if needed. +# Running this first gives two tries in case of network lossage. +before_script: + - timeout 5m docker pull palmskog/xenial-for-verdi-$COQ_VERSION || true + +# Using travis_wait here seems to cause the job to terminate after 1 minute +# with no error (!). +# The git commands are tried twice, in case of temporary network failure. +# The fcntl line works around a bug where Travis truncates logs and fails. +script: + - python -c "import fcntl; fcntl.fcntl(1, fcntl.F_SETFL, 0)" + - REMOTE_ORIGIN_URL=`git config --get remote.origin.url` + - echo "THIS_REPO=${THIS_REPO}" + - echo "COQ_VERSION=${COQ_VERSION}" + - echo "TRAVIS_BRANCH=${TRAVIS_BRANCH}" + - echo "REMOTE_ORIGIN_URL=${REMOTE_ORIGIN_URL}" + - echo "TRAVIS_EVENT_TYPE=${TRAVIS_EVENT_TYPE}" + - echo "TRAVIS_COMMIT=${TRAVIS_COMMIT}" + - echo "TRAVIS_PULL_REQUEST=${TRAVIS_PULL_REQUEST}" + - echo "TRAVIS_PULL_REQUEST_BRANCH=${TRAVIS_PULL_REQUEST_BRANCH}" + - echo "TRAVIS_PULL_REQUEST_SHA=${TRAVIS_PULL_REQUEST_SHA}" + - echo "TRAVIS_REPO_SLUG=${TRAVIS_REPO_SLUG}" + - >- + docker run palmskog/xenial-for-verdi-$COQ_VERSION /bin/bash -c "true && + if [ $TRAVIS_EVENT_TYPE = pull_request ] ; then + git clone --quiet --depth 9 $REMOTE_ORIGIN_URL $THIS_REPO || git clone --quiet --depth 9 $REMOTE_ORIGIN_URL $THIS_REPO + cd $THIS_REPO + git fetch origin +refs/pull/$TRAVIS_PULL_REQUEST/merge + git checkout -qf $TRAVIS_PULL_REQUEST_SHA + else + git clone --quiet --depth 9 -b $TRAVIS_BRANCH $REMOTE_ORIGIN_URL $THIS_REPO || git clone --quiet --depth 9 -b $TRAVIS_BRANCH $REMOTE_ORIGIN_URL $THIS_REPO + cd $THIS_REPO + git checkout -qf $TRAVIS_COMMIT + fi && + ./.travis-ci.sh" + +git: + depth: 9 + notifications: email: false diff --git a/Core/Actions.v b/Core/Actions.v index 19fad65..a2cda50 100644 --- a/Core/Actions.v +++ b/Core/Actions.v @@ -4,9 +4,9 @@ From mathcomp Require Import path. Require Import Eqdep. Require Import Relation_Operators. -From DiSeL.Heaps -Require Import pred prelude idynamic ordtype finmap pcm unionmap heap coding. -From DiSeL.Core +From fcsl +Require Import axioms pred prelude ordtype finmap pcm unionmap heap. +From DiSeL Require Import Freshness State EqTypeX Protocols Worlds NetworkSem. Require Classical_Prop. @@ -285,19 +285,19 @@ set s2 := let: d := getS s l in upd l (DStatelet f' s') s. exists s2, msg; split=>//; exists b; split=>//. move: (safe_safe (And4 C S J K))=> S''. -by rewrite -E (proof_irrelevance S'' S') . +by rewrite -E (pf_irr S'' S') . Qed. Lemma send_act_step_sem s1 (S : send_act_safe s1) s2 r: send_act_step S s2 r -> network_step W this s1 s2. Proof. case=>_[b][E Z]; case: (S)=>C S' /andP[D1] D2 K; subst s2=>/=. -rewrite (proof_irrelevance (safe_safe S) S') in E; clear S. +rewrite (pf_irr (safe_safe S) S') in E; clear S. rewrite /all_hooks_fire/filter_hooks in K. move: st S' E K pf'; clear pf' st; subst p=>st S' E K' pf'. apply: (@SendMsg W this s1 _ l st pf' to msg)=>////. move=>z lc hk E'; apply: (K' z); rewrite E'. -by rewrite find_um_filt/= eqxx. +by rewrite find_umfilt/= eqxx. Qed. Definition send_action_wrapper := diff --git a/Core/Always.v b/Core/Always.v index 1467cba..472df30 100644 --- a/Core/Always.v +++ b/Core/Always.v @@ -4,14 +4,12 @@ From mathcomp Require Import path. Require Import Eqdep. Require Import Relation_Operators. -From DiSeL.Heaps -Require Import pred prelude idynamic ordtype finmap pcm unionmap heap coding. -From DiSeL.Core +From fcsl +Require Import axioms pred prelude ordtype finmap pcm unionmap heap. +From DiSeL Require Import Freshness State EqTypeX DepMaps Protocols Worlds NetworkSem Rely. -From DiSeL.Core -Require Import Actions Injection Process. -From DiSeL.Core -Require InductiveInv. +From DiSeL +Require Import Actions Injection Process InductiveInv. Set Implicit Arguments. Unset Strict Implicit. @@ -27,7 +25,7 @@ Variable W : world. Notation coherent := (Coh W). -Implicit Arguments proc [W this]. +Arguments proc [this W]. Fixpoint always_sc A (s1 : state) p scs (P : state -> proc A -> Prop) : Prop := s1 \In coherent /\ @@ -152,7 +150,7 @@ Proof. move=>C H [|sc scs]; split=>// s2; case/H=>// H1[H2]H3//. split=>//; first by case: sc. move=>s3 q /stepAct [v][pf][_ -> St]. -rewrite (proof_irrelevance H1 pf) in H3. +rewrite (pf_irr H1 pf) in H3. apply: alw_ret'; last by move=>s4; apply: H3 St. by case: (step_coh (a_step_sem St)). Qed. @@ -207,7 +205,7 @@ move=>Ls; split. by move=>H ps; apply/(alwA' _ (Ls ps))=>x; apply: H. Qed. -Implicit Arguments alwA [A B s p P]. +Arguments alwA [A B s p P]. (* always commutes with implication, so we can weaken the postconditions *) @@ -240,7 +238,7 @@ move=>Ls; split; first by move=>H Hp scs; apply/alwI': Hp. by move=>H scs; apply/alwI'=>//; move/H; move/(_ scs). Qed. -Implicit Arguments alwI [A s p P Q]. +Arguments alwI [A s p P Q]. Lemma alw_bnd A B (p1 : proc A) (p12 : proc B) pp2 s1 @@ -336,8 +334,8 @@ split; apply: alw_imp=>t q _ I. by move=>v; move/I. Qed. -Implicit Arguments aftA [A B s p P]. -Implicit Arguments aftI [A s p P Q]. +Arguments aftA [A B s p P]. +Arguments aftI [A s p P Q]. End Always. @@ -364,8 +362,8 @@ case: (sem_split w C1 C2 N); case=>R E; [subst s2'|subst s2]; split=>//; apply: Idle; split=>//. case: (step_coh N)=>C _. case/(cohE w): (C)=>s3[s4][E]C' C''. -move: (coh_prec (cohS C) E C1 C')=>Z; subst s3. -by rewrite (joinfK (cohS C) E). +move: (coh_prec (cohS C) C1 C' E)=>Z; subst s3. +by rewrite (joinxK (cohS C) E). Qed. Lemma rely_split s1 s1' s2 s2' : @@ -382,12 +380,12 @@ elim: n s1 s1' E C1 C2=>[|n IH] /= s1 s1'; last first. + by case: G1=>m R; exists m.+1, z, s4. by case: G2=>m R; exists m.+1, z, s5. move=> [E1 E2] C1 C2. -move: (coh_prec (cohS E2) E1 C1 C2)=>Z; subst s2. -rewrite (joinfK (cohS E2) E1); split; exists 0=>//. -split=>//; rewrite -(joinfK (cohS E2) E1)=>{E1 s2' C2}. +move: (coh_prec (cohS E2) C1 C2 E1)=>Z; subst s2. +rewrite (joinxK (cohS E2) E1); split; exists 0=>//. +split=>//; rewrite -(joinxK (cohS E2) E1)=>{E1 s2' C2}. move/(cohE w): (E2)=>[t1][t2][E]C' C''. -move: ((coh_prec (cohS E2)) E C1 C')=>Z; subst t1. -by rewrite (joinfK (cohS E2) E). +move: ((coh_prec (cohS E2)) C1 C' E)=>Z; subst t1. +by rewrite (joinxK (cohS E2) E). Qed. @@ -423,13 +421,13 @@ case=>sc' [q'][x1][i2][y1][_ -> E -> {sc q s}] _ T Ls. have [E1 E2] : x1 = i1 /\ y1 = j1. - case: T=>Cx1 _. - move: (coh_prec (cohS C) E Ci1 Cx1) (E)=><-{E Cx1 x1}. - by move/(joinfK (cohS C)). + move: (coh_prec (cohS C) Ci1 Cx1 E) (E)=><-{E Cx1 x1}. + by move/(joinxK (cohS C)). rewrite {E x1}E1 {y1}E2 in T *. have C' : i2 \+ j1 \In Coh W. - move: (C)=>C'; rewrite (cohE w) in C *=>[[s1]][s2][E]D1 D2. - move: (coh_prec (cohS C') E Ci1 D1)=>Z; subst i1. - move: (joinfK (cohS C') E)=>Z; subst s2; clear E. + move: (coh_prec (cohS C') Ci1 D1 E)=>Z; subst i1. + move: (joinxK (cohS C') E)=>Z; subst s2; clear E. apply/(cohE w); exists i2, j1; split=>//. by case/step_coh: (pstep_network_sem T). move/(alw_step Ls): T=>{Ls} Ls. diff --git a/Core/DepMaps.v b/Core/DepMaps.v index 45b56fd..9bf9c70 100644 --- a/Core/DepMaps.v +++ b/Core/DepMaps.v @@ -4,15 +4,14 @@ From mathcomp Require Import path. Require Import Eqdep. Require Import Relation_Operators. -From DiSeL.Heaps -Require Import pred prelude idynamic ordtype finmap pcm unionmap heap coding. -From DiSeL.Core +From fcsl +Require Import pred prelude ordtype finmap pcm unionmap heap. +From DiSeL Require Import Freshness EqTypeX. Set Implicit Arguments. Unset Strict Implicit. Unset Printing Implicit Defensive. - (* An implementation of a dependent map structure *) Module DepMaps. @@ -25,7 +24,7 @@ Variable V : Type. Variable labF: V -> Label. Definition dmDom (u : union_map Label V) : bool := - all (fun l => if find l u is Some p then (labF p) == l else false) (keys_of u). + all (fun l => if find l u is Some p then (labF p) == l else false) (dom u). Record depmap := DepMap { dmap : union_map Label V; @@ -37,9 +36,7 @@ Section PCMOps. Variable dm : depmap. Lemma dmDom_unit : dmDom Unit. -Proof. -by apply/allP=>l; rewrite keys_dom dom0 inE. -Qed. +Proof. by apply/allP=>l; rewrite dom0. Qed. Definition unit := DepMap dmDom_unit. @@ -53,12 +50,12 @@ Lemma dmDom_join um1 um2: dmDom um1 -> dmDom um2 -> dmDom (um1 \+ um2). Proof. case; case W: (valid (um1 \+ um2)); last first. -- by move=> _ _; apply/allP=>l; rewrite keys_dom=>/dom_valid; rewrite W. -move/allP=>D1/allP D2; apply/allP=>l; rewrite keys_dom. +- by move=> _ _; apply/allP=>l; move/dom_valid; rewrite W. +move/allP=>D1/allP D2; apply/allP=>l. rewrite domUn inE=>/andP[_]/orP; rewrite findUnL//; case=>E; rewrite ?E. -- by apply: D1; rewrite keys_dom. +- by apply: D1. rewrite joinC in W; case: validUn W=>// _ _ /(_ l E)/negbTE->_. -by apply: D2; rewrite keys_dom. +by apply: D2. Qed. Definition join : depmap := DepMap (dmDom_join (@pf dm1) (@pf dm2)). @@ -84,21 +81,21 @@ Lemma joinC f1 f2 : f1 \+ f2 = f2 \+ f1. Proof. case: f1 f2=>d1 pf1[d2 pf2]; rewrite /join/=. move: (dmDom_join pf1 pf2) (dmDom_join pf2 pf1); rewrite joinC=>G1 G2. -by move: (proof_irrelevance G1 G2)=>->. +by move: (bool_irrelevance G1 G2)=>->. Qed. Lemma joinCA f1 f2 f3 : f1 \+ (f2 \+ f3) = f2 \+ (f1 \+ f3). Proof. case: f1 f2 f3=>d1 pf1[d2 pf2][d3 pf3]; rewrite /join/=. move: (dmDom_join pf1 (dmDom_join pf2 pf3)) (dmDom_join pf2 (dmDom_join pf1 pf3)). -by rewrite joinCA=>G1 G2; move: (proof_irrelevance G1 G2)=>->. +by rewrite joinCA=>G1 G2; move: (bool_irrelevance G1 G2)=>->. Qed. Lemma joinA f1 f2 f3 : f1 \+ (f2 \+ f3) = (f1 \+ f2) \+ f3. Proof. case: f1 f2 f3=>d1 pf1[d2 pf2][d3 pf3]; rewrite /join/=. move: (dmDom_join pf1 (dmDom_join pf2 pf3)) (dmDom_join (dmDom_join pf1 pf2) pf3). -by rewrite joinA=>G1 G2; move: (proof_irrelevance G1 G2)=>->. +by rewrite joinA=>G1 G2; move: (bool_irrelevance G1 G2)=>->. Qed. Lemma validL f1 f2 : valid (f1 \+ f2) -> valid f1. @@ -108,7 +105,7 @@ Lemma unitL f : unit \+ f = f. Proof. rewrite /join/unit/=; case: f=>//=u pf. move: pf (dmDom_join (dmDom_unit labF) pf); rewrite unitL=>g1 g2. -by move: (proof_irrelevance g1 g2)=>->. +by move: (bool_irrelevance g1 g2)=>->. Qed. Lemma validU : valid unit. @@ -127,7 +124,7 @@ Definition DepMap := DepMap. Lemma dep_unit (d : depmap labF) : dmap d = Unit -> d = unit labF. Proof. case: d=>u pf/=; rewrite /unit. move: (dmDom_unit labF)=>pf' Z; subst u. -by rewrite (proof_irrelevance pf). +by rewrite (bool_irrelevance pf). Qed. Coercion dmap := dmap. diff --git a/Core/DiSeLExtraction.v b/Core/DiSeLExtraction.v index d2dc957..a940f13 100644 --- a/Core/DiSeLExtraction.v +++ b/Core/DiSeLExtraction.v @@ -1,4 +1,4 @@ -From DiSeL.Core +From DiSeL Require Import While. Require Import ExtrOcamlBasic. diff --git a/Heaps/domain.v b/Core/Domain.v similarity index 96% rename from Heaps/domain.v rename to Core/Domain.v index 2be30ed..40a6682 100644 --- a/Heaps/domain.v +++ b/Core/Domain.v @@ -1,7 +1,7 @@ From mathcomp.ssreflect Require Import ssreflect ssrbool ssrfun ssrnat eqtype seq. -From DiSeL.Heaps -Require Import pred prelude. +From fcsl +Require Import axioms pred prelude. Set Implicit Arguments. Unset Strict Implicit. Unset Printing Implicit Defensive. @@ -122,7 +122,7 @@ Lemma sub_asym x y : sub_leq x y -> sub_leq y x -> x = y. Proof. move: x y=>[x Hx][y Hy]; rewrite /sub_leq /= => H1 H2. move: (poset_asym H1 H2) Hy=> <- Hy; congr exist. -by apply: proof_irrelevance. +by apply: pf_irr. Qed. Lemma sub_trans x y z : sub_leq x y -> sub_leq y z -> sub_leq x z. @@ -235,7 +235,7 @@ Lemma ideal_asym x y : ideal_leq x y -> ideal_leq y x -> x = y. Proof. move: x y=>[x1 H1][x2 H2]; rewrite /ideal_leq /= => H3 H4; move: H1 H2. rewrite (poset_asym H3 H4)=>H1 H2. -congr Ideal; apply: proof_irrelevance. +congr Ideal; apply: pf_irr. Qed. Lemma ideal_trans x y z : ideal_leq x y -> ideal_leq y z -> ideal_leq x z. @@ -344,7 +344,7 @@ Notation "[ 'lattice' 'of' T 'for' cT ]" := (@clone T cT _ id) Notation "[ 'lattice' 'of' T ]" := (@clone T _ _ id) (at level 0, format "[ 'lattice' 'of' T ]") : form_scope. -Implicit Arguments Lattice.sup [cT]. +Arguments Lattice.sup [cT]. Prenex Implicits Lattice.sup. Notation sup := Lattice.sup. @@ -399,9 +399,9 @@ Definition sup_closure (T : lattice) (s : Pred T) := End Lat. -Implicit Arguments lbot [T]. -Implicit Arguments sup_closed [T]. -Implicit Arguments sup_closure [T]. +Arguments lbot [T]. +Arguments sup_closed [T]. +Arguments sup_closure [T]. Prenex Implicits sup_closed sup_closure. @@ -628,20 +628,20 @@ End PredLattice. Lemma id_mono (T : poset) : monotone (@id T). Proof. by []. Qed. -Implicit Arguments id_mono [T]. +Arguments id_mono [T]. Prenex Implicits id_mono. Lemma const_mono (T1 T2 : poset) (y : T2) : monotone (fun x : T1 => y). Proof. by []. Qed. -Implicit Arguments const_mono [T1 T2 y]. +Arguments const_mono [T1 T2 y]. Prenex Implicits const_mono. Lemma comp_mono (T1 T2 T3 : poset) (f1 : T2 -> T1) (f2 : T3 -> T2) : monotone f1 -> monotone f2 -> monotone (f1 \o f2). Proof. by move=>M1 M2 x y H; apply: M1; apply: M2. Qed. -Implicit Arguments comp_mono [T1 T2 T3 f1 f2]. +Arguments comp_mono [T1 T2 T3 f1 f2]. Prenex Implicits comp_mono. Lemma proj1_mono (T1 T2 : poset) : monotone (@fst T1 T2). @@ -650,27 +650,27 @@ Proof. by case=>x1 x2 [y1 y2][]. Qed. Lemma proj2_mono (T1 T2 : poset) : monotone (@snd T1 T2). Proof. by case=>x1 x2 [y1 y2][]. Qed. -Implicit Arguments proj1_mono [T1 T2]. -Implicit Arguments proj2_mono [T1 T2]. +Arguments proj1_mono [T1 T2]. +Arguments proj2_mono [T1 T2]. Prenex Implicits proj1_mono proj2_mono. Lemma diag_mono (T : poset) : monotone (fun x : T => (x, x)). Proof. by []. Qed. -Implicit Arguments diag_mono [T]. +Arguments diag_mono [T]. Prenex Implicits diag_mono. Lemma app_mono A (T : poset) (x : A) : monotone (fun f : A -> T => f x). Proof. by move=>f1 f2; apply. Qed. -Implicit Arguments app_mono [A T]. +Arguments app_mono [A T]. Prenex Implicits app_mono. Lemma dapp_mono A (T : A -> poset) (x : A) : monotone (fun f : dfunPoset T => f x). Proof. by move=>f1 f2; apply. Qed. -Implicit Arguments dapp_mono [A T]. +Arguments dapp_mono [A T]. Prenex Implicits dapp_mono. Lemma prod_mono (S1 S2 T1 T2 : poset) (f1 : S1 -> T1) (f2 : S2 -> T2) : @@ -679,7 +679,7 @@ Proof. by move=>M1 M2 [x1] x2 [y1 y2][/= H1 H2]; split; [apply: M1 | apply: M2]. Qed. -Implicit Arguments prod_mono [S1 S2 T1 T2 f1 f2]. +Arguments prod_mono [S1 S2 T1 T2 f1 f2]. Prenex Implicits prod_mono. @@ -706,7 +706,7 @@ Lemma chainE (T : poset) (s1 s2 : chain T) : s1 = s2 <-> pred_of s1 =p pred_of s2. Proof. split=>[->//|]; move: s1 s2=>[s1 H1][s2 H2] /= E; move: H1 H2. -suff: s1 = s2 by move=>-> H1 H2; congr Chain; apply: proof_irrelevance. +suff: s1 = s2 by move=>-> H1 H2; congr Chain; apply: pf_irr. by apply: fext=>x; apply: pext; split; move/E. Qed. @@ -900,8 +900,8 @@ Notation "[ 'cpo' 'of' T 'for' cT ]" := (@clone T cT _ idfun) Notation "[ 'cpo' 'of' T ]" := (@clone T _ _ id) (at level 0, format "[ 'cpo' 'of' T ]") : form_scope. -Implicit Arguments CPO.bot [cT]. -Implicit Arguments CPO.lim [cT]. +Arguments CPO.bot [cT]. +Arguments CPO.lim [cT]. Prenex Implicits CPO.lim. Prenex Implicits CPO.bot. Notation lim := CPO.lim. @@ -1146,7 +1146,7 @@ Qed. End AdmissibleClosure. -Implicit Arguments chain_closed [T]. +Arguments chain_closed [T]. Prenex Implicits chain_closed. (* diagonal of an admissible set of pairs is admissible *) @@ -1233,7 +1233,7 @@ Lemma contE (s : chain D1) (C : continuous) : f (lim s) = lim [f ^^ s by cont_mono C]. Proof. case: C=>M E; rewrite E; congr (lim (image_chain _ _)). -apply: proof_irrelevance. +exact: pf_irr. Qed. End Continuity. @@ -1326,7 +1326,7 @@ Export Kleene.Exports. Lemma id_cont (D : cpo) : continuous (@id D). Proof. by exists id_mono; move=>d; rewrite id_chainE. Qed. -Implicit Arguments id_cont [D]. +Arguments id_cont [D]. Prenex Implicits id_cont. Lemma const_cont (D1 D2 : cpo) (y : D2) : continuous (fun x : D1 => y). @@ -1336,7 +1336,7 @@ exists const_mono; move=>s; apply: poset_asym. by apply: limM=>_ [x][->]. Qed. -Implicit Arguments const_cont [D1 D2 y]. +Arguments const_cont [D1 D2 y]. Prenex Implicits const_cont. Lemma comp_cont (D1 D2 D3 : cpo) (f1 : D2 -> D1) (f2 : D3 -> D2) : @@ -1346,7 +1346,7 @@ case=>M1 H1 [M2 H2]; exists (comp_mono M1 M2); move=>d. by rewrite /= H2 H1 comp_chainE. Qed. -Implicit Arguments comp_cont [D1 D2 D3 f1 f2]. +Arguments comp_cont [D1 D2 D3 f1 f2]. Prenex Implicits comp_cont. Lemma proj1_cont (D1 D2 : cpo) : continuous (@fst D1 D2). @@ -1355,8 +1355,8 @@ Proof. by exists proj1_mono. Qed. Lemma proj2_cont (D1 D2 : cpo) : continuous (@snd D1 D2). Proof. by exists proj2_mono. Qed. -Implicit Arguments proj1_cont [D1 D2]. -Implicit Arguments proj2_cont [D1 D2]. +Arguments proj1_cont [D1 D2]. +Arguments proj2_cont [D1 D2]. Prenex Implicits proj1_cont proj2_cont. Lemma diag_cont (D : cpo) : continuous (fun x : D => (x, x)). @@ -1365,7 +1365,7 @@ exists diag_mono=>d; apply: poset_asym; by split=>/=; [rewrite proj1_diagE | rewrite proj2_diagE]. Qed. -Implicit Arguments diag_cont [D]. +Arguments diag_cont [D]. Prenex Implicits diag_cont. Lemma app_cont A (D : cpo) x : continuous (fun f : A -> D => f x). @@ -1374,8 +1374,8 @@ Proof. by exists (app_mono x). Qed. Lemma dapp_cont A (D : A -> cpo) x : continuous (fun f : dfunCPO D => f x). Proof. by exists (dapp_mono x). Qed. -Implicit Arguments app_cont [A D]. -Implicit Arguments dapp_cont [A D]. +Arguments app_cont [A D]. +Arguments dapp_cont [A D]. Prenex Implicits app_cont dapp_cont. Lemma prod_cont (S1 S2 T1 T2 : cpo) (f1 : S1 -> T1) (f2 : S2 -> T2) : @@ -1385,5 +1385,5 @@ case=>M1 H1 [M2 H2]; exists (prod_mono M1 M2)=>d; apply: poset_asym; by (split=>/=; [rewrite proj1_prodE H1 | rewrite proj2_prodE H2]). Qed. -Implicit Arguments prod_cont [S1 S2 T1 T2 f1 f2]. +Arguments prod_cont [S1 S2 T1 T2 f1 f2]. Prenex Implicits prod_cont. diff --git a/Core/Freshness.v b/Core/Freshness.v index 6d5f9c4..89992b9 100644 --- a/Core/Freshness.v +++ b/Core/Freshness.v @@ -2,12 +2,30 @@ From mathcomp.ssreflect Require Import ssreflect ssrbool ssrnat eqtype ssrfun seq. From mathcomp Require Import path. -From DiSeL.Heaps -Require Import pred prelude idynamic ordtype finmap pcm coding unionmap heap. +From fcsl +Require Import pred prelude ordtype finmap pcm unionmap. Set Implicit Arguments. Unset Strict Implicit. Unset Printing Implicit Defensive. +Section Keys. +Variables (K : ordType) (V : Type) (U : union_map_class K V). +Implicit Types (k : K) (v : V) (f : U). + +Lemma keys_last_mono f1 f2 k : + path oleq k (dom f1) -> + path oleq k (dom f2) -> + (forall x, x \in dom f1 -> x \in dom f2) -> + oleq (last k (dom f1)) (last k (dom f2)). +Proof. +rewrite !umEX; case: (UMC.from f1); first by move=>_ H _; apply: path_last H. +move=>{f1} f1 /= _ H1; case: (UMC.from f2)=>/=. +- by move=>_ /allP; case: (supp f1)=>//; rewrite /oleq eq_refl orbT. +by move=>{f2} f2 /= _; apply: seq_last_mono H1. +Qed. + +End Keys. + (* Last_key and Fresh are constructs that work *) (* for any union map with integer keys *) (* Should be developed more generally for any union map class *) @@ -17,14 +35,14 @@ Section FreshLastKey. Variable V : Type. Implicit Type f : union_map [ordType of nat] V. -Definition last_key f := last 0 (keys_of f). +Definition last_key f := last 0 (dom f). Lemma last_key0 : last_key Unit = 0. -Proof. by rewrite /last_key /Unit /= !umE. Qed. +Proof. by rewrite /last_key /Unit /= !umEX. Qed. Lemma last_key_dom f : valid f -> last_key f \notin dom f -> f = Unit. Proof. -rewrite /valid /= /last_key /Unit /= !umE /= -{4}[f]UMC.tfE. +rewrite /valid /= /last_key /Unit /= !umEX /= -{4}[f]UMC.tfE. case: (UMC.from f)=>//=; case=>s H /= H1 _ /seq_last_in. rewrite /UM.empty UMC.eqE UM.umapE /supp fmapE /= {H H1}. by elim: s. @@ -35,7 +53,7 @@ Proof. by move=>X; apply: contraR; move/(last_key_dom X)=>->; apply/empbP. Qed. Lemma last_key_max f x : x \in dom f -> x <= last_key f. Proof. -rewrite /last_key /= !umE; case: (UMC.from f)=>//; case=>s H _ /=. +rewrite /last_key /= !umEX; case: (UMC.from f)=>//; case=>s H _ /=. rewrite /supp /ord /= (leq_eqVlt x) orbC. by apply: sorted_last_key_max (sorted_oleq H). Qed. @@ -43,17 +61,17 @@ Qed. Lemma max_key_last f x : x \in dom f -> {in dom f, forall y, y <= x} -> last_key f = x. Proof. -rewrite /last_key !umE; case: (UMC.from f)=>//; case=>s H _ /=. +rewrite /last_key !umEX; case: (UMC.from f)=>//; 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 last_keyPt (x : nat) v : last_key (x \\-> v) = x. -Proof. by rewrite /last_key /um_pts /= !umE. Qed. +Proof. by rewrite /last_key /um_pts /= !umEX. Qed. -Lemma hist_path f : path oleq 0 (keys_of f). +Lemma hist_path f : path oleq 0 (dom f). Proof. -rewrite !umE; case: (UMC.from f)=>// {f} f /= _; case: f; case=>//= x s H. +rewrite !umEX; case: (UMC.from f)=>// {f} f /= _; case: f; case=>//= x s H. rewrite {1}/oleq /ord /= orbC -leq_eqVlt /=. by apply: sub_path H=>z y; rewrite /oleq=>->. Qed. @@ -63,7 +81,7 @@ Lemma last_key_mono f1 f2 : Proof. rewrite leq_eqVlt orbC=>H; apply: (@keys_last_mono _ _ _ f1 f2); try by apply: hist_path. -by move=>x /=; move: (H x); rewrite -!keys_dom. +by move=>x /=; move: (H x). Qed. Lemma last_keyfUn f1 f2 : @@ -74,7 +92,6 @@ Lemma last_keyUnf f1 f2 : valid (f1 \+ f2) -> last_key f2 <= last_key (f1 \+ f2). Proof. by rewrite joinC; apply: last_keyfUn. Qed. - (* freshness *) Definition fresh f := (last_key f).+1. @@ -89,13 +106,13 @@ Lemma dom_fresh f : fresh f \notin dom f. Proof. by move: (dom_freshn f 0); rewrite addn0. Qed. Lemma valid_fresh f v : valid (f \+ fresh f \\-> v) = valid f. -Proof. by rewrite joinC um_validPtUn dom_fresh andbT. Qed. +Proof. by rewrite joinC validPtUn dom_fresh andbT. Qed. Lemma valid_fresh' f v i w : valid (f \+ i \\-> w) -> valid (f \+ fresh (f \+ i \\-> w) \\-> v). Proof. -move=> W; rewrite joinC um_validPtUn. +move=> W; rewrite joinC validPtUn. move: (dom_fresh (f \+ i \\-> w)); rewrite domUn inE; rewrite W/=. by rewrite negb_or=>/andP; case=>-> _;move/validL: W=>->. Qed. @@ -103,8 +120,8 @@ Qed. Lemma last_fresh f v : valid f -> last_key (f \+ fresh f \\-> v) = fresh f. Proof. move=>Vf; apply: max_key_last=>[|x] /=. -- by rewrite domUn inE valid_fresh Vf um_domPt inE eq_refl orbT. -rewrite domUn inE /= valid_fresh Vf /= um_domPt inE /= orbC eq_sym. +- by rewrite domUn inE valid_fresh Vf domPt inE eq_refl orbT. +rewrite domUn inE /= valid_fresh Vf /= domPt inE /= orbC eq_sym. by rewrite leq_eqVlt; case: eqP=>//= _; apply: dom_ordfresh. Qed. diff --git a/Core/HoareTriples.v b/Core/HoareTriples.v index ce4f78c..01cbd4b 100644 --- a/Core/HoareTriples.v +++ b/Core/HoareTriples.v @@ -4,15 +4,12 @@ From mathcomp Require Import path. Require Import Eqdep. Require Import Relation_Operators. -From DiSeL.Heaps -Require Import pred prelude idynamic ordtype finmap pcm unionmap heap coding. -From DiSeL.Core -Require Import Freshness State EqTypeX DepMaps Protocols Worlds NetworkSem Rely. -From DiSeL.Core +From fcsl +Require Import axioms pred prelude ordtype finmap pcm unionmap heap. +From DiSeL +Require Import Domain Freshness State EqTypeX DepMaps Protocols Worlds NetworkSem Rely. +From DiSeL Require Import Actions Injection Process Always. -From DiSeL.Heaps -Require Import domain. - Set Implicit Arguments. Unset Strict Implicit. Unset Printing Implicit Defensive. @@ -33,8 +30,8 @@ Variable A : Type. Notation coherent := (Coh W). -Implicit Arguments Prog [W A]. -Implicit Arguments Prog [W A this]. +Arguments Prog [W A]. +Arguments Prog [W A this]. Coercion set_of : prog >-> Funclass. @@ -44,7 +41,7 @@ Lemma progE (T1 T2 : prog W A this) : T1 = T2 <-> set_of T1 = set_of T2. Proof. split=>[->//|]; case: T1 T2=>m1 H1 [m2 H2] /= E. -by move: H2; rewrite -E => H2; rewrite (proof_irrelevance H1 H2). +by move: H2; rewrite -E => H2; rewrite (pf_irr H1 H2). Qed. (* Pre/poscondition and *) @@ -88,7 +85,7 @@ Lemma stsepE (W : world) A (s : spec A) (e1 e2 : DTbin this W s) : e1 = e2 <-> e1 = e2 :> prog W A this. Proof. split=>[->//|]; case: e1 e2=>T1 H1 [T2 H2] /= E. -by rewrite -{T2}E in H2 *; rewrite (proof_irrelevance H1 H2). +by rewrite -{T2}E in H2 *; rewrite (pf_irr H1 H2). Qed. (* Unfinished satisfies any specification *) @@ -108,9 +105,9 @@ Definition post_of W A := fun e : DT W A => (spec_of e).2. Definition code_of (W : world) A (e : DT W A) := let: with_spec _ c := e return DTbin this W (spec_of e) in c. -Implicit Arguments pre_of [W A]. -Implicit Arguments post_of [W A]. -Implicit Arguments with_spec [W A]. +Arguments pre_of [W A]. +Arguments post_of [W A]. +Arguments with_spec [W A]. Prenex Implicits pre_of post_of. Coercion with_spec : DTbin >-> DT. @@ -383,19 +380,19 @@ have : after i1 t' (s.2 i1) by case: (code_of e) H2=>p /= /(_ _ _ H1); apply. move/(aft_inject w C); apply: aft_imp=>{H1 t' H2} v m _. case=>i2 [j2][->{m} Ci2 S2 H2] i3 j3 E Ci3. suff [E1 E2]: i3 = i1 /\ j3 = j1. -- by rewrite {i3 E Ci3}E1 {j3}E2; exists i2, j2. -move: (coh_prec (cohS C) E Ci1 Ci3) (E)=>{Ci3 E} <-. -by move/(joinfK (cohS C)). +- by rewrite {i3 E Ci3}E1 {j3}E2; exists i2, j2. +move: (coh_prec (cohS C) Ci1 Ci3 E) (E)=>{Ci3 E} <-. +by move/(joinxK (cohS C)). Qed. Definition inject := with_spec (DTbin_make inject_has_spec). End Inject. +From DiSeL +Require Import InductiveInv. Section InductiveInv. -From DiSeL.Core -Require Import InductiveInv. Variable pr : protocol. (* Decompose the initial protocol *) diff --git a/Core/InductiveInv.v b/Core/InductiveInv.v index 1e1eee7..03676f5 100644 --- a/Core/InductiveInv.v +++ b/Core/InductiveInv.v @@ -4,12 +4,10 @@ From mathcomp Require Import path. Require Import Eqdep. Require Import Relation_Operators. -From DiSeL.Heaps -Require Import pred prelude idynamic ordtype finmap pcm unionmap heap coding. -From DiSeL.Core -Require Import Freshness State EqTypeX. -From DiSeL.Core -Require Import Protocols Worlds NetworkSem Rely. +From fcsl +Require Import axioms pred prelude ordtype finmap pcm unionmap heap. +From DiSeL +Require Import Freshness State EqTypeX Protocols Worlds NetworkSem Rely. Require FunctionalExtensionality. Set Implicit Arguments. Unset Strict Implicit. @@ -54,22 +52,22 @@ Definition CohI := CohPred (CohPredMixin cohIVd cohIVs cohIDom). (* Some helper lemmas *) Lemma st_helper d : (getStatelet (toSt d) l) = d. -Proof. by rewrite /getStatelet gen_findPt. Qed. +Proof. by rewrite /getStatelet findPt. Qed. Lemma cohSt d : coh d -> Coh W (toSt d). Proof. move=>C; split. -- by apply/andP; rewrite valid_unit; split=>//; apply: gen_validPt. -- by apply: gen_validPt. -- by move=>z lc ls st; rewrite dom0 inE. -- by move=>z; rewrite !gen_domPt inE. +- by apply/andP; rewrite valid_unit; split=>//; apply: validPt. +- by apply: validPt. +- by move=>z lc ls st; rewrite dom0. +- by move=>z; rewrite !domPt inE. move=>k; case B: (k \in dom (toSt d)); last first. - rewrite /getProtocol /getStatelet. case: (dom_find k (toSt d))=>[->|v /find_some]; last by rewrite B. - rewrite gen_domPt inE/= in B. - by case: (dom_find k (toSt p))=>[->|? /find_some]//; rewrite gen_domPt inE/= B. -rewrite gen_domPt inE/= in B; move/eqP: B=>B; subst k. -by rewrite /getProtocol /getStatelet !gen_findPt. + rewrite domPt inE/= in B. + by case: (dom_find k (toSt p))=>[->|? /find_some]//; rewrite domPt inE/= B. +rewrite domPt inE/= in B; move/eqP: B=>B; subst k. +by rewrite /getProtocol /getStatelet !findPt. Qed. Section SendInvWrapper. @@ -96,7 +94,7 @@ Proof. move: (s_safe_def st this to d msg)=>H; split. - move=>S; case: (S)=>/H[b][pf]E H1; exists b. exists (conj pf H1). - by rewrite /send_stepI (proof_irrelevance (proj1 (conj pf H1)) pf). + by rewrite /send_stepI (pf_irr (proj1 (conj pf H1)) pf). by case=>b[pf]E. Qed. @@ -113,10 +111,10 @@ Hypothesis HIstep : S_inv. Lemma s_step_cohI : s_step_coh_t CohI (t_snd st) send_stepI. Proof. -have E1: (getProtocol W l) = p by rewrite /getProtocol/W/=gen_findPt. +have E1: (getProtocol W l) = p by rewrite /getProtocol/W/=findPt. move/s_step_coh: (st)=>H. move=>this to d msg [S]H1 b. -rewrite /send_stepI (proof_irrelevance (proj1 (conj S H1)) S)=>E. +rewrite /send_stepI (pf_irr (proj1 (conj S H1)) S)=>E. split=>//= ; first by apply: (H this to d msg S b E). apply: (HIstep H1 E _)=>//. by move=>z N; rewrite /getLocal/= findU N. @@ -155,7 +153,7 @@ Notation msg_wfI := (fun d (C : CohI d) => msg_wf rt (proj1 C)). Lemma r_step_cohI : r_step_coh_t msg_wfI (t_rcv rt) receive_stepI. Proof. -have E1: (getProtocol W l) = p by rewrite /getProtocol/W/=gen_findPt. +have E1: (getProtocol W l) = p by rewrite /getProtocol/W/=findPt. move=>d from this i C F m D N Wf T. rewrite /receive_stepI/=. set d' := {| dstate := upd this (receive_step rt from m (proj1 C) F) (dstate d); @@ -248,7 +246,10 @@ case/Mem_map_inv; case=>st' stI/= [->]H1; case=>S Is E. rewrite /get_st/InMem!prEq; exists st'. split=>//. - by rewrite -HS/=; apply: (Mem_map ProtocolWithInvariant.st H1). rewrite/send_step/=/Transitions.send_step/=/send_stepI in E. -- by move=>???/sym/find_some; rewrite dom0 inE. +- move=>??? F. + apply sym_eq in F. + move: F. + by move/find_some; rewrite dom0. by exists (proj1 (conj S Is)). Qed. @@ -327,16 +328,16 @@ Lemma with_inv_coh pr I (ii : InductiveInv pr I) s: s \In Coh (mkWorld pr). Proof. case=>G1 G2 Gh G3 G4; split=>//. -- by apply/andP; rewrite valid_unit; split=>//; rewrite !gen_validPt in G1 *. -- by move=>???? ;rewrite dom0 inE. -- by move=>z; move: (G3 z); rewrite !um_domPt !inE. +- by apply/andP; rewrite valid_unit; split=>//; rewrite !validPt in G1 *. +- by move=>???? ;rewrite dom0. +- by move=>z; move: (G3 z); rewrite !domPt !inE. move=>l; move: (G4 l). case X: (l == plab pr); first by move/eqP:X=>X; subst l; rewrite !prEq; case. rewrite /getProtocol/mkWorld/=. suff [X1 X2]: l \notin dom (plab pr \\-> ProtocolWithIndInv ii) /\ l \notin dom (plab pr \\-> pr). - by case: dom_find X1 X2=>//=->_; case: dom_find=>//=->. -rewrite !gen_domPt !inE/=; suff: plab pr != l by []. +rewrite !domPt !inE/=; suff: plab pr != l by []. by apply/negbT; apply/negP=>/eqP Z; subst l; rewrite eqxx in X. Qed. @@ -346,12 +347,12 @@ Lemma with_inv_nodes pr I (ii : InductiveInv pr I) l : Proof. rewrite /getProtocol/mkWorld/=. case X: (l == plab pr); - first by move/eqP:X=>X; subst l; rewrite !gen_findPt. + first by move/eqP:X=>X; subst l; rewrite !findPt. rewrite /getProtocol/mkWorld/=. suff [X1 X2]: l \notin dom (plab pr \\-> ProtocolWithIndInv ii) /\ l \notin dom (plab pr \\-> pr). by case: dom_find X1 X2=>//=->_; case: dom_find=>//=->. -rewrite !gen_domPt !inE/=; suff: plab pr != l by []. +rewrite !domPt !inE/=; suff: plab pr != l by []. by apply/negbT; apply/negP=>/eqP Z; subst l; rewrite eqxx in X. Qed. @@ -370,7 +371,7 @@ case; first by case=>_<-; apply: Idle. (*** Send-transition of the host protocol ***) move=>l st Hs to msg h H1 H2 _ S A E ->/=. have Y : l = plab pr - by rewrite -(cohD C) gen_domPt inE/= in H2; move/eqP:H2. + by rewrite -(cohD C) domPt inE/= in H2; move/eqP:H2. subst l; move: st Hs H1 S E A; rewrite /get_st/InMem!prEq/==>st Hs H1 S E A. suff X: exists st', [/\ st' \In get_st (mkWorld (ProtocolWithIndInv ii)) (plab pr), @@ -390,11 +391,11 @@ have G: I (getStatelet s1 (plab pr)) (nodes pr (getStatelet s1 (plab pr))). + by case: C'=>_ _ _ _/(_ (plab pr)); rewrite prEq; case. move:(Hi _ Si)=>/={Hi Si}Hi; exists (@snd_transI pr I st stI); split=>//=. rewrite /send_safeI /send_stepI/=. -by exists (conj S G); rewrite (proof_irrelevance (proj1 (conj S G)) S). +by exists (conj S G); rewrite (pf_irr (proj1 (conj S G)) S). (*** Receive-transition of the host protocol ***) move=>l rt Hr i from H1 H2 C1 msg E[F]Hw/=. -have Y : l = plab pr by rewrite -(cohD C) gen_domPt inE/= in H2; move/eqP:H2. +have Y : l = plab pr by rewrite -(cohD C) domPt inE/= in H2; move/eqP:H2. subst l; move: rt Hr H1 E (coh_s _ C1) Hw. rewrite /get_rt/InMem/=!prEq=>rt Hr H1 E C1' Hw . have Hi: (coh (getProtocol (mkWorld (ProtocolWithIndInv ii)) (plab pr))) @@ -410,13 +411,13 @@ case:X=>rt'[Hr']E' Hw' Gr G. have pf: (z \in nodes (getProtocol (mkWorld (ProtocolWithIndInv ii)) (plab pr)) (getStatelet s1 (plab pr))) by rewrite prEq. move: (@ReceiveMsg _ z s1 s2 (plab pr) rt' Hr' i from pf)=>/=. -rewrite -(cohD C) /= gen_domPt inE eqxx/=; move/(_ is_true_true C' msg E')=>X. +rewrite -(cohD C) /= domPt inE eqxx/=; move/(_ is_true_true C' msg E')=>X. subst s2; apply X; split=>//; clear X. -- by rewrite (proof_irrelevance (coh_s (plab pr) C') Hi). +- by rewrite (pf_irr (coh_s (plab pr) C') Hi). congr (upd _ _); congr {| dstate := _ ; dsoup := _ |}; congr (upd _ _). rewrite Gr; clear E E' Hw Hw' Hr Hr' Gr rt F H2 H1 C1 C. -rewrite (proof_irrelevance pf Hz). -by rewrite (proof_irrelevance Hi (coh_s (plab pr) C')). +rewrite (pf_irr pf Hz). +by rewrite (pf_irr Hi (coh_s (plab pr) C')). (* Finding the corresponding receive-transition of rich protocol *) rewrite /get_rt/InMem/=; move: C1' H1 Hi Hz Hw; rewrite !prEq=>C1' H1 Hi Hz Hw. @@ -425,9 +426,9 @@ rewrite -HR in Hr; move/Mem_map_inv: Hr=>[[rt' rtI]]/=[Er]Ri; subst rt'. case=>C1'' Inv Hz. exists (@rcv_transI pr I rt rtI); split=>//. - by move: (@Hi _ Ri). rewrite /receive_step/=/receive_stepI/=/receive_step/=. -- by rewrite -(proof_irrelevance C1' (proj1 (conj C1'' Inv))). -by rewrite ?(proof_irrelevance C1' (proj1 (conj C1'' Inv))) - ?(proof_irrelevance H1 Hz). +- by rewrite -(pf_irr C1' (proj1 (conj C1'' Inv))). +by rewrite ?(pf_irr C1' (proj1 (conj C1'' Inv))) + ?(pf_irr H1 Hz). Qed. Lemma with_inv_step' pr I (ii : InductiveInv pr I) z s1 s2: @@ -439,7 +440,7 @@ case. (* Emulating send-transition *) move=>l st Hs to msg h H1 H2 C' S A E->{s2}. have Y : l = plab pr - by rewrite -(cohD C') gen_domPt inE/= in H2; move/eqP:H2. + by rewrite -(cohD C') domPt inE/= in H2; move/eqP:H2. subst l; move: st Hs H1 S E A; rewrite /get_st/InMem!prEq/==>st Hs H1 S E A. suff X: exists st', [/\ st' \In get_st (mkWorld pr) (plab pr), @@ -454,7 +455,7 @@ by apply: (getInvSendTrans (ii := ii)). (* Emulating a receive-transition *) move=>l rt Hr i from H1 H2 C1 msg E[F]Hw/=. -have Y : l = plab pr by rewrite -(cohD C1) gen_domPt inE/= in H2; move/eqP:H2. +have Y : l = plab pr by rewrite -(cohD C1) domPt inE/= in H2; move/eqP:H2. subst l; move: rt Hr H1 E (coh_s _ C1) Hw. rewrite /get_rt/InMem/=!prEq=>rt Hr H1 E C1' Hw. have Hi: (coh (getProtocol (mkWorld pr) (plab pr))) @@ -470,14 +471,14 @@ case:X=>rt'[Hr']E' Hw' Gr G. have pf: (z \in nodes (getProtocol (mkWorld pr) (plab pr)) (getStatelet s1 (plab pr))) by rewrite prEq. move: (@ReceiveMsg _ z s1 s2 (plab pr) rt' Hr' i from pf)=>/=. -rewrite -(cohD C1) gen_domPt inE eqxx/=. +rewrite -(cohD C1) domPt inE eqxx/=. move/(_ is_true_true (with_inv_coh C1) msg E')=>X. subst s2; apply X; split=>//; clear X. -- by rewrite (proof_irrelevance (coh_s (plab pr) (with_inv_coh C1)) Hi). +- by rewrite (pf_irr (coh_s (plab pr) (with_inv_coh C1)) Hi). congr (upd _ _); congr {| dstate := _ ; dsoup := _ |}; congr (upd _ _). rewrite Gr; clear E E' Hw Hw' Hr Hr' Gr rt F H2 H1. -rewrite (proof_irrelevance pf Hz). -by rewrite (proof_irrelevance Hi (coh_s (plab pr) (with_inv_coh C1))). +rewrite (pf_irr pf Hz). +by rewrite (pf_irr Hi (coh_s (plab pr) (with_inv_coh C1))). (* Identifying the appropriate receive transition *) case: ii C1 H2 rt Hr E H1 C1' Hw=>sts rts HS HR/=C1 H2 rt Hr E H1 C1' Hw. move: Hr; rewrite /rtsI. @@ -485,9 +486,9 @@ case/Mem_map_inv; case=>rt' rtI/= [Z] H1'; subst rt. rewrite /get_rt/InMem; move: C1' H1 Hi Hz Hw; rewrite !prEq=>C1' H1 Hi Hz Hw. exists rt'; split=>//; last first. - rewrite {1}/receive_step /rcv_transI /receive_stepI/=. - rewrite (proof_irrelevance H1 Hz). - by rewrite (proof_irrelevance (proj1 C1') Hi). -- by rewrite -(proof_irrelevance (proj1 C1') Hi). + rewrite (pf_irr H1 Hz). + by rewrite (pf_irr (proj1 C1') Hi). +- by rewrite -(pf_irr (proj1 C1') Hi). by rewrite -HR; apply: (Mem_map (@ProtocolWithInvariant.rt pr I) H1'). Qed. diff --git a/Core/InferenceRules.v b/Core/InferenceRules.v index 27288a2..70a0b13 100644 --- a/Core/InferenceRules.v +++ b/Core/InferenceRules.v @@ -4,17 +4,12 @@ From mathcomp Require Import path. Require Import Eqdep. Require Import Relation_Operators. -From DiSeL.Heaps -Require Import pred prelude idynamic ordtype finmap pcm unionmap. -From DiSeL.Heaps -Require Import heap coding domain. -From DiSeL.Core -Require Import Freshness State EqTypeX DepMaps Protocols Worlds NetworkSem Rely. -From DiSeL.Core -Require Import Actions Injection Process Always HoareTriples. -From DiSeL.Core -Require Import InductiveInv. - +From fcsl +Require Import pred prelude ordtype finmap pcm unionmap heap. +From DiSeL +Require Import Domain Freshness State EqTypeX DepMaps Protocols Worlds NetworkSem Rely. +From DiSeL +Require Import Actions Injection Process Always HoareTriples InductiveInv. Set Implicit Arguments. Unset Strict Implicit. Import Prenex Implicits. @@ -91,7 +86,7 @@ move=>H1 H2 Ci t [->|[t'][H3 H4]]. by apply: aft_bnd H3 _; move/(H1 Ci): H4; apply: aft_imp=>y j Cj H; apply: H2. Qed. -Implicit Arguments bind_rule [W A B e1 e2 i r]. +Arguments bind_rule [W A B e1 e2 i]. Lemma step W A B (e1 : DT this W A) (e2 : A -> DT this W B) i (r : cont B) : verify i e1 (fun y m => verify m (e2 y) r) -> @@ -181,7 +176,7 @@ Proof. by []. Qed. End GhostRules. -Implicit Arguments gh_ex [W A C s f i k]. +Arguments gh_ex [W A C s f]. Lemma act_rule W A (a: action W A this) i (r : cont A) : (forall j, network_rely W this i j -> a_safe a j /\ diff --git a/Core/Injection.v b/Core/Injection.v index 5808b58..01d32fe 100644 --- a/Core/Injection.v +++ b/Core/Injection.v @@ -4,12 +4,10 @@ From mathcomp Require Import path. Require Import Eqdep. Require Import Relation_Operators. -From DiSeL.Heaps -Require Import pred prelude idynamic ordtype finmap pcm unionmap heap coding. -From DiSeL.Core -Require Import Freshness State EqTypeX Protocols Worlds NetworkSem. -From DiSeL.Core -Require Import Actions. +From fcsl +Require Import axioms pred prelude ordtype finmap pcm unionmap heap. +From DiSeL +Require Import Freshness State EqTypeX Protocols Worlds NetworkSem Actions. Set Implicit Arguments. Unset Strict Implicit. Unset Printing Implicit Defensive. @@ -96,18 +94,18 @@ Lemma projectS_cohL W1 W2 s : s \In Coh (W1 \+ W2) -> hook_complete W1 -> projectS W1 s \In Coh W1. Proof. case=>V1 V2 G1 D H G2; split=>//; first by move/validL: V1. -- by rewrite valid_um_filter. +- by rewrite valid_umfilt. - move=>z; case B: (z \in dom (getc W1)). - + by rewrite dom_um_filt !inE B/= -D/=domUn !inE B/=; case/andP:V1=>->. - by rewrite dom_um_filt !inE B. + + by rewrite dom_umfilt !inE B/= -D/=domUn !inE B/=; case/andP:V1=>->. + by rewrite dom_umfilt !inE B. move=>l; move: (H l)=>{H}H. case B: (l \in dom (getc W1)); last first. - rewrite /getProtocol /getStatelet; move: (B). case: dom_find=>//-> _. suff X: ~~(l \in dom (projectS W1 s)) by case: dom_find X=>//-> _. - by rewrite /projectS dom_um_filt inE/= B. + by rewrite /projectS dom_umfilt inE/= B. have E1: find l s = find l (projectS W1 s). -- by rewrite /projectS/= find_um_filt B. +- by rewrite /projectS/= find_umfilt B. have E2: getProtocol (W1 \+ W2) l = getProtocol W1 l. - rewrite /getProtocol findUnL//?B//. by rewrite /valid/= in V1; case/andP: V1. @@ -132,9 +130,9 @@ have X: {in dom s, dom_filt (getc W2) =1 predD (dom_filt (getc W2)) (dom_filt (g move: I; rewrite domUn inE/==>/negbT; rewrite negb_and negb_or/=. have X: valid (W1 \+ W2) by []. by case/andP: X=>->/=_/andP[]->. -rewrite (eq_in_um_filt X) -um_filt_predU/=; clear X. +rewrite (eq_in_umfilt X) -umfilt_predU/=; clear X. suff X: {in dom s, predU (dom_filt (getc W1)) (dom_filt (getc W2)) =1 predT}. -- by rewrite (eq_in_um_filt X) um_filt_predT. +- by rewrite (eq_in_umfilt X) umfilt_predT. by move=>z; rewrite/= -D domUn inE=>/andP[]. Qed. @@ -157,7 +155,7 @@ rewrite -joinA in E. case/andP: H=>H1 H2. rewrite /PCM.join/= in H1 H2 E *. case: W2 H1 H2 E=>/=c2 h2 H1 H2 [E1 E2]. -by rewrite (joinfK H1 E1) (joinfK H2 E2). +by rewrite (joinxK H1 E1) (joinxK H2 E2). Qed. Lemma injExtR' W1 W2 K (pf : injects W2 (W1 \+ W2) K) : @@ -168,7 +166,7 @@ rewrite -(joinC W2) in E H. case/andP: H=>H1 H2; rewrite -joinA in E. rewrite /PCM.join/= in H1 H2 E *. case: W1 H1 H2 E=>/=c1 h1 H1 H2 [E1 E2]. -by rewrite (joinfK H1 E1) (joinfK H2 E2). +by rewrite (joinxK H1 E1) (joinxK H2 E2). Qed. Lemma injExtL W1 W2 (pf : injects W1 (W1 \+ W2) Unit) : @@ -197,11 +195,12 @@ suff X: s' = projectS W (s \+ s'). - by rewrite X; apply: (projectS_cohR H). suff X: s = projectS U (s \+ s'). - move: (cohS H)=>V; move/projectSE: (H)=>E. - rewrite E in V. rewrite {1}X in E; move/sym: E=>E. - by rewrite (joinfK V E). + rewrite E in V. + rewrite {1}X in E. + by rewrite (joinxK V (sym_eq E)). rewrite /projectS. -suff X: {in dom (s \+ s'), dom U.1 =1 dom s}. -- by rewrite (eq_in_um_filt X) um_filt_dom ?(cohS H)//. +suff X: {in dom (s \+ s'), dom U.1 =i dom s}. +- by rewrite (eq_in_umfilt X) umfilt_dom ?(cohS H)//. by move=>z _; move: (cohD C z); rewrite /in_mem. Qed. @@ -252,7 +251,7 @@ Proof. case=>V Vs Hk D L. split=>//; first by move/validL: V. - by apply: hook_completeL V Hk. -- move=>z; rewrite -D domUn !inE/= unitR dom0 inE orbC/=. +- move=>z; rewrite -D domUn !inE/= unitR dom0 orbC/=. by move/validL:V=>/andP[]->_. by move=>l; move: (L l); rewrite (get_protocol_hooks K l (validL V)). Qed. @@ -291,12 +290,12 @@ rewrite domUn inE=>/andP[Vs]/=/orP; case=>D C msg H2 [H3 H4]/=; [rewrite updUnL D|rewrite updUnR D]=>G;[left|right]. - have X: s1' = s2'. + move: (C2'); rewrite (joinC U) G -[upd _ _ _ \+ s1'](joinC s1'). - move/cohUnKR/(_ D1)/(_ Hu)=>C1''. - move: (coh_prec (cohS C2') G C2 C1'')=>Z; rewrite -Z in G; clear Z. - by rewrite (joinfK (cohS C2') G). + move/cohUnKR/(_ D1)/(_ Hu)=>C1''. + move: (coh_prec (cohS C2') C2 C1'' G)=>Z; rewrite -Z in G; clear Z. + by rewrite (joinxK (cohS C2') G). split=>//; subst s2'; rewrite -![_ \+ s1'](joinC s1') in G C2'. rewrite -[upd _ _ _ \+ s1'](joinC s1') in G; rewrite (joinC U) in C2'. - move: (joinfK (cohS C2') G)=>{G}G. + move: (joinxK (cohS C2') G)=>{G}G. have E: getProtocol U l = getProtocol (U \+ W) l. by rewrite (getPUn V)// (cohD C1). have E': getStatelet s1 l = getStatelet (s1 \+ s1') l. @@ -307,7 +306,7 @@ rewrite domUn inE=>/andP[Vs]/=/orP; case=>D C msg H2 [H3 H4]/=; (get_protocol_hooks K l V) E rt R H2=>_-><-rt R H2 coh_s H1 G H3 H4. apply: (ReceiveMsg R D H2 (i := i) (from := from) (s2 := s2) (C := C1)). split=>//=; move: (NetworkSem.coh_s l C1)=>coh_s'; - by rewrite -(proof_irrelevance coh_s coh_s'). + by rewrite -(pf_irr coh_s coh_s'). (* Second part or receive-step *) move: U W V {V1} Hw Hu s1 s2 s1' s2' C1 C2 C1' C2' D1 D2 rt H1 H2 Vs D C R H3 H4 G. @@ -317,10 +316,10 @@ move=>rt H1 H2 Vs D C R H3 H4 G. have X: s1' = s2'. - move: (C2'); rewrite (joinC U) G. move/cohUnKR/(_ D1)/(_ Hw)=>C1''; rewrite (joinC s1') in G. - move: (coh_prec (cohS C2') G C2 C1'')=>Z; rewrite -Z in G; clear Z. - by rewrite (joinfK (cohS C2') G). + move: (coh_prec (cohS C2') C2 C1'' G)=>Z; rewrite -Z in G; clear Z. + by rewrite (joinxK (cohS C2') G). split=>//; subst s2'; rewrite -!(joinC s1') in G C2'. -rewrite (joinC U) in C2'; move: (joinfK (cohS C2') G)=>{G}G. +rewrite (joinC U) in C2'; move: (joinxK (cohS C2') G)=>{G}G. rewrite joinC in V. have E: getProtocol U l = getProtocol (U \+ W) l. by rewrite (getPUn V)// (cohD C1). @@ -332,12 +331,12 @@ move: (getProtocol (U \+ W \+ (Unit, K)) l) (get_protocol_hooks K l V) E rt R H2=>_-><-rt R H2 coh_s H1 G H3 H4. apply: (ReceiveMsg R D H2 (i := i) (from := from) (s2 := s2) (C := C1)). split=>//=; move: (NetworkSem.coh_s l C1)=>coh_s'; -by rewrite -(proof_irrelevance coh_s coh_s'). +by rewrite -(pf_irr coh_s coh_s'). (* Idle Step *) -- case=>_ E; move: (coh_prec (cohS C1') E C1 C2)=>Z; subst s2. +- case=>_ E; move: (coh_prec (cohS C1') C1 C2 E)=>Z; subst s2. rewrite (joinC U) (joinC s1) in C1'; rewrite !(joinC s1) in E. - move: (coh_prec (cohS C1') E D1 D2)=>Z; subst s2'. + move: (coh_prec (cohS C1') D1 D2 E)=>Z; subst s2'. by left; split=>//; apply: Idle. (* Send Step *) @@ -351,10 +350,10 @@ by rewrite -(proof_irrelevance coh_s coh_s'). + have X: s1' = s2'. - move: (C2'); rewrite (joinC U) G -[upd _ _ _ \+ s1'](joinC s1'). move/cohUnKR/(_ D1)/(_ Hu)=>C1''. - move: (coh_prec (cohS C2') G C2 C1'')=>Z; rewrite -Z in G; clear Z. - by rewrite (joinfK (cohS C2') G). + move: (coh_prec (cohS C2') C2 C1'' G)=>Z; rewrite -Z in G; clear Z. + by rewrite (joinxK (cohS C2') G). split=>//; subst s2'; rewrite -!(joinC s1') in G C2'. - rewrite (joinC U) in C2'; move: (joinfK (cohS C2') G)=>{G}G. + rewrite (joinC U) in C2'; move: (joinxK (cohS C2') G)=>{G}G. rewrite (joinC s1') in G. have E: getProtocol U l = getProtocol (U \+ W) l. by rewrite (getPUn V)// (cohD C1). @@ -367,7 +366,9 @@ by rewrite -(proof_irrelevance coh_s coh_s'). apply: (SendMsg H1 H2 D C1 _ H3 G). (* Now proving the obligation about all_hooks_fire *) - move=>z lc hk/sym F A1 A2; move: (Hk z lc hk). + move=>z lc hk F A1 A2. + apply sym_eq in F. + move: (Hk z lc hk). rewrite -F -joinA !domUn !inE !Vs A1 A2 findUnL ?E' ?(find_some F)/=; last by case/andP:V1; rewrite-joinA =>_->. move/(_ erefl is_true_true is_true_true). @@ -376,10 +377,10 @@ by rewrite -(proof_irrelevance coh_s coh_s'). have X: s1' = s2'. - move: (C2'); rewrite (joinC U) G. move/cohUnKR/(_ D1)/(_ Hu)=>C1''; rewrite (joinC s1') in G. - move: (coh_prec (cohS C2') G C2 C1'')=>Z; rewrite -Z in G; clear Z. - by rewrite (joinfK (cohS C2') G). + move: (coh_prec (cohS C2') C2 C1'' G)=>Z; rewrite -Z in G; clear Z. + by rewrite (joinxK (cohS C2') G). split=>//; subst s2'; rewrite -!(joinC s1') in G C2'. - rewrite (joinC U) in C2'; move: (joinfK (cohS C2') G)=>{G}G. + rewrite (joinC U) in C2'; move: (joinxK (cohS C2') G)=>{G}G. rewrite (joinC s1') in G. have E: getProtocol U l = getProtocol (U \+ W) l. by rewrite (getPUn V)// (cohD C1). @@ -392,7 +393,9 @@ by rewrite -(proof_irrelevance coh_s coh_s'). apply: (SendMsg H1 H2 D C1 _ H3 G). (* Now proving the obligation about all_hooks_fire *) - move=>z lc hk/sym F A1 A2; move: (Hk z lc hk). + move=>z lc hk F A1 A2. + apply sym_eq in F. + move: (Hk z lc hk). rewrite -F -joinA !domUn !inE A1 A2 findUnL ?E' ?(find_some F)/=; last by rewrite joinA (joinC U.2); case/andP:V1=>_->. rewrite !(joinC s1') !Vs/= -!(orbC true). @@ -427,18 +430,20 @@ move=>G1 G2 G D' C1 C' N A z lc hk F D1 D2; move: F. case/andP: (cohW C')=>/=V1 V2. move: (cohUnKR (coh_hooks C') C1 G2) => C2. rewrite findUnL ?V2//=; case: ifP=>D3; last first. -- by move/sym/find_some/N/negP; rewrite eqxx. +- move => F; apply sym_eq in F; move: F. + by move/find_some/N/negP; rewrite eqxx. rewrite findUnR ?(validL V2)//; case: ifP=>[D|_]. + case/G2/andP: D=>_ D; rewrite (cohD C2) in D. by case: validUn (cohS C')=>//_ _/(_ _ D'); rewrite D. -move/sym=> F. +move => F. +apply sym_eq in F. have D'': lc \in dom s by case/andP:(G1 _ _ _ _ (find_some F)); rewrite (cohD C1). have E: getStatelet s l = getStatelet (s \+ s') l by rewrite (getSUn (cohS C'))// -?(cohD C1'). have E': getStatelet s lc = getStatelet (s \+ s') lc. by rewrite (getSUn (cohS C'))// -?(cohD C1'). move: (getStatelet (s \+ s') l) (getStatelet (s \+ s') lc) E E'. -by move=>y1 y2 Z1 Z2; subst y1 y2; move/sym: F=>F; apply: (A z). +by move=>y1 y2 Z1 Z2; subst y1 y2; apply sym_eq in F; apply: (A z). Qed. (********************************************************************) @@ -491,9 +496,9 @@ move: (getProtocol U l) (getStatelet s1 l) E E' C H2 (coh_s l C) rt G3 G G2 H1 G1=>z1 z2 Z1 Z2. subst z1 z2=>C pf C' G3 G G2 H1 H2 G1. apply: (ReceiveMsg H2 X G2 (i := msg) (from := from) (s2 := s2 \+ s)). -split=>//=; first by rewrite (proof_irrelevance (coh_s l C1) C'). +split=>//=; first by rewrite (pf_irr (coh_s l C1) C'). rewrite updUnL H3; congr (_ \+ _); move: (NetworkSem.coh_s l C1)=>pf'. -by rewrite (proof_irrelevance pf' C'). +by rewrite (pf_irr pf' C'). Qed. diff --git a/Core/Make b/Core/Make index 27b582c..3d4abd9 100644 --- a/Core/Make +++ b/Core/Make @@ -1,22 +1,23 @@ --Q . DiSeL.Core +-Q . DiSeL -arg "-w -notation-overridden,-local-declaration,-redundant-canonical-projection,-projection-no-head-constant" -./State.v -./Freshness.v -./Protocols.v -./EqTypeX.v -./DepMaps.v -./StatePredicates.v -./NewStatePredicates.v -./Worlds.v -./NetworkSem.v -./Rely.v -./Actions.v -./Injection.v -./InductiveInv.v -./Process.v -./Always.v -./HoareTriples.v -./InferenceRules.v -./While.v -./DiSeLExtraction.v +Domain.v +State.v +Freshness.v +Protocols.v +EqTypeX.v +DepMaps.v +StatePredicates.v +NewStatePredicates.v +Worlds.v +NetworkSem.v +Rely.v +Actions.v +Injection.v +InductiveInv.v +Process.v +Always.v +HoareTriples.v +InferenceRules.v +While.v +DiSeLExtraction.v diff --git a/Core/NetworkSem.v b/Core/NetworkSem.v index 9cdaccc..cb294e5 100644 --- a/Core/NetworkSem.v +++ b/Core/NetworkSem.v @@ -4,9 +4,9 @@ From mathcomp Require Import path. Require Import Eqdep. Require Import Relation_Operators. -From DiSeL.Heaps -Require Import pred prelude idynamic ordtype finmap pcm unionmap heap coding. -From DiSeL.Core +From fcsl +Require Import axioms pred prelude ordtype finmap pcm unionmap heap. +From DiSeL Require Import Freshness State EqTypeX Protocols Worlds. Set Implicit Arguments. Unset Strict Implicit. @@ -121,7 +121,7 @@ case=>[[H1 <-] | l st _ to a loc' pf D C S Ph Spf ->/= | have X: exists b pf, sstep this to (gets s1 l) a pf = Some b by exists loc', S. move/Y: X=>X; move: (G1 _ _ _ _ X) (G2 _ _ _ _ X)=>{G1 G2}G1 G2; apply: G3. rewrite /gets in Spf; rewrite Spf; move: (coh_s l C)=>G1'. - by rewrite -(proof_irrelevance S X). + by rewrite -(pf_irr X S). case: (C)=>W V K E H; split=>//; first by rewrite validU/= V. - move=>z; rewrite domU/= !inE V. by case b: (z == l)=>//; move/eqP: b=>?; subst; rewrite E D. @@ -131,9 +131,9 @@ rewrite V; move/eqP: b=>Z; subst k=>/=. have pf' : this \in dom (dstate (gets s1 l)) by move: (pf); rewrite -(Coh_dom D C). case: rt H1 H3 msg H4=>/= r_rcvwf mwf rstep G msg T F M. -rewrite -(proof_irrelevance (H l) (coh_s l C)) in M. +rewrite -(pf_irr (H l) (coh_s l C)) in M. move: (G (gets s1 l) from this i (H l) pf msg pf' T M F); rewrite /gets. -by move: (H l)=>G1'; rewrite -(proof_irrelevance G1 G1'). +by move: (H l)=>G1'; rewrite -(pf_irr G1 G1'). Qed. (* Stepping preserves the protocol structure *) diff --git a/Core/NewStatePredicates.v b/Core/NewStatePredicates.v index 0f88ac4..23b7bd6 100644 --- a/Core/NewStatePredicates.v +++ b/Core/NewStatePredicates.v @@ -4,9 +4,9 @@ From mathcomp Require Import path. Require Import Eqdep. Require Import Relation_Operators. -From DiSeL.Heaps -Require Import pred prelude idynamic ordtype finmap pcm unionmap heap coding. -From DiSeL.Core +From fcsl +Require Import pred prelude ordtype finmap pcm unionmap heap. +From DiSeL Require Import Freshness State EqTypeX DepMaps. Set Implicit Arguments. Unset Strict Implicit. diff --git a/Core/Process.v b/Core/Process.v index 8e6b863..ba7676e 100644 --- a/Core/Process.v +++ b/Core/Process.v @@ -4,11 +4,11 @@ From mathcomp Require Import path. Require Import Eqdep. Require Import Relation_Operators. -From DiSeL.Heaps -Require Import pred prelude idynamic ordtype finmap pcm unionmap heap coding. -From DiSeL.Core +From fcsl +Require Import pred prelude ordtype finmap pcm unionmap heap. +From DiSeL Require Import Freshness State EqTypeX DepMaps Protocols Worlds NetworkSem. -From DiSeL.Core +From DiSeL Require Import Actions Injection InductiveInv. Set Implicit Arguments. @@ -37,11 +37,11 @@ Inductive schedule := End ProcessSyntax. -Implicit Arguments Unfinished [this W A]. -Implicit Arguments Ret [this W A]. -Implicit Arguments Act [this W A]. -Implicit Arguments Seq [this W A B]. -Implicit Arguments WithInv [this W A]. +Arguments Unfinished [this W A]. +Arguments Ret [this W A]. +Arguments Act [this W A]. +Arguments Seq [this W A B]. +Arguments WithInv [this W A]. Section ProcessSemantics. @@ -211,7 +211,7 @@ case=>C; case: sc=>//=; last first. by left; exists s1'; exists v. move=>sc /= [s'][X] S [s1'][s2'][t'][t2'][??? C1'] T; subst q s1 s2. right; exists sc, t2', s1', s2', t'; do !split=>//. -by case: X=>t'' [E] Cs' _; rewrite (coh_prec (cohS C) E _ Cs'). +by case: X=>t'' [E] Cs' _; rewrite (coh_prec (cohS C) _ Cs' E). Qed. Lemma stepWithInv W A pr I (ii : InductiveInv pr I) s1 @@ -267,7 +267,7 @@ elim: sc W A s1 s2 t q=>/=. subst sc' q s1 s2=>C; move/HI=>S; apply: (sem_extend pf)=>//. apply/(cohE pf); exists s2', s; case: (step_coh S)=>C1 C2; split=>//. move/(cohE pf): (C)=>[s1][s2][E]C' H. - by move: (coh_prec (cohS C) E C1 C')=>Z; subst s1'; rewrite (joinfK (cohS C) E). + by move: (coh_prec (cohS C) C1 C' E)=>Z; subst s1'; rewrite (joinxK (cohS C) E). by move=>?????; case. - move=>W A s1 s2 p q; case: p; do?[by case|by move=>?; case]. + by move=>???; case. diff --git a/Core/Protocols.v b/Core/Protocols.v index 37a732e..111dab1 100644 --- a/Core/Protocols.v +++ b/Core/Protocols.v @@ -4,9 +4,9 @@ From mathcomp Require Import path. Require Import Eqdep. Require Import Relation_Operators. -From DiSeL.Heaps -Require Import pred prelude idynamic ordtype finmap pcm unionmap heap coding. -From DiSeL.Core +From fcsl +Require Import pred prelude ordtype finmap pcm unionmap heap. +From DiSeL Require Import Freshness State EqTypeX. Set Implicit Arguments. Unset Strict Implicit. @@ -26,7 +26,7 @@ Lemma getLocalU n m d s : valid (dstate d) -> m \in dom (dstate d) -> getLocal n d = (getLocal n {| dstate := upd m (getLocal m d) (dstate d); dsoup := s |}). Proof. -move=>V H2; move/gen_eta: (H2)=>[v2][F2 _]. +move=>V H2; move/um_eta: (H2)=>[v2][F2 _]. rewrite /getLocal F2/=; case X: (n == m); last by rewrite findU X/=. by move/eqP: X=>X; subst m; rewrite findU eqxx/=V F2. Qed. diff --git a/Core/Rely.v b/Core/Rely.v index c4b675b..6f2ce28 100644 --- a/Core/Rely.v +++ b/Core/Rely.v @@ -4,9 +4,9 @@ From mathcomp Require Import path. Require Import Eqdep. Require Import Relation_Operators. -From DiSeL.Heaps -Require Import pred prelude idynamic ordtype finmap pcm unionmap heap coding. -From DiSeL.Core +From fcsl +Require Import pred prelude ordtype finmap pcm unionmap heap. +From DiSeL Require Import Freshness State EqTypeX Protocols Worlds NetworkSem. Set Implicit Arguments. Unset Strict Implicit. @@ -74,7 +74,7 @@ move: (coh_coh l C); rewrite /gets findU; case B: (l == k)=>//=; move/eqP: B=>B; subst k; rewrite (stepV1 S). - case: (dom_find l s)=>[|d->_]; first by move/find_none; rewrite D. - move=> C' E; rewrite -E; rewrite joinC gen_findPtUn2; last first. + move=> C' E; rewrite -E; rewrite joinC findPtUn2; last first. + rewrite joinC valid_fresh; apply: (cohVs C'). case: ifP=>///eqP; move/find_some: E=>F Z. by move/negbTE: (dom_fresh (dsoup d)); rewrite -Z F. @@ -109,7 +109,7 @@ move: (coh_coh l C); rewrite /gets findU; case B: (l == k)=>//=; do?[by exists b]; move/eqP: B=>B; subst k; rewrite (stepV1 S). - case: (dom_find l s)=>[|d->_ C']; first by move/find_none; rewrite D. - rewrite joinC gen_findPtUn2; last first. + rewrite joinC findPtUn2; last first. + rewrite joinC valid_fresh; apply: (cohVs C'). case B: (m == fresh (dsoup d)); first by case=>_ Z _; subst; move/eqP: N. by move=>H; exists b; split. @@ -148,7 +148,7 @@ move: (coh_coh l C); rewrite /gets findU; case B: (l == k)=>//=; do?[by exists b]; move/eqP: B=>B; subst k; rewrite (stepV1 S). - case: (dom_find l s)=>[|d->_ C']; first by move/find_none; rewrite D. - rewrite joinC gen_findPtUn2; last first. + rewrite joinC findPtUn2; last first. + rewrite joinC valid_fresh; apply: (cohVs C'). case B: (m == fresh (dsoup d)); last by move=>->; exists b. move/eqP:B=>B; subst m; move: (dom_fresh (dsoup d))=>B. diff --git a/Core/State.v b/Core/State.v index cbf7220..f8575b1 100644 --- a/Core/State.v +++ b/Core/State.v @@ -4,9 +4,9 @@ From mathcomp Require Import path. Require Import Eqdep. Require Import Relation_Operators. -From DiSeL.Heaps -Require Import pred prelude idynamic ordtype finmap pcm unionmap heap coding. -From DiSeL.Core +From fcsl +Require Import pred prelude ordtype finmap pcm unionmap heap. +From DiSeL Require Import Freshness DepMaps EqTypeX. Set Implicit Arguments. Unset Strict Implicit. @@ -111,7 +111,7 @@ Section Shared. exists msg', find m s' = Some msg' /\ msg = mark_msg msg'. Proof. move=>V'; rewrite /consume_msg; case D: (m \in dom s'). - - move/gen_eta: D=>[msg'][->]_; rewrite findU eqxx/= V'. + - move/um_eta: D=>[msg'][->]_; rewrite findU eqxx/= V'. by case=><-; eexists _. by case: dom_find (D)=>//->_; move/find_some=>Z; rewrite Z in D. Qed. @@ -137,9 +137,9 @@ Section Shared. else (consume_msg s' j) \+ (i \\-> mm). Proof. rewrite ![_ \+ i \\-> _]joinC; rewrite eq_sym. - move=>V'; case B: (j==i); rewrite /consume_msg um_findPtUn2// B. - - by move/eqP: B=>?; subst j; rewrite um_updPtUn. - by case X: (find j s')=>//; rewrite updUnL um_domPt inE eq_sym B. + move=>V'; case B: (j==i); rewrite /consume_msg findPtUn2// B. + - by move/eqP: B=>?; subst j; rewrite updPtUn. + by case X: (find j s')=>//; rewrite updUnL domPt inE eq_sym B. Qed. Notation "'{{' m 'in' s 'at' id '}}'" := (find id s = Some m). @@ -164,19 +164,16 @@ Section Local. End Local. +(* Definition um_all {A:ordType} {B} (p : A -> B -> bool) (u : union_map A B) : bool := - gen_rect false true (fun k v f rec Hval Hpath => p k v && rec) u. + um_recf false true (fun k v f rec Hval Hpath => p k v && rec) u. Definition um_some {A:ordType} {B} (p : A -> B -> bool) (u : union_map A B) : bool := - gen_rect false false (fun k v f rec Hval Hpath => p k v || rec) u. - + um_recf false false (fun k v f rec Hval Hpath => p k v || rec) u. +*) Section Statelets. - Definition big_valid {U : encoded_pcm} (lst : lstate_type U) : bool := - valid lst && - valid (gen_rect (P := fun _ => U) Unit Unit (fun k v f rec Hval Hpath => v \+ rec) lst). - (* A particular statelet instance. The Label and the PCM are the parameters and are defined by the protocol. The lstate and dsop are subject of the evolution. diff --git a/Core/StatePredicates.v b/Core/StatePredicates.v index 842a840..0b72aef 100644 --- a/Core/StatePredicates.v +++ b/Core/StatePredicates.v @@ -4,9 +4,9 @@ From mathcomp Require Import path. Require Import Eqdep. Require Import Relation_Operators. -From DiSeL.Heaps -Require Import pred prelude idynamic ordtype finmap pcm unionmap heap coding. -From DiSeL.Core +From fcsl +Require Import pred prelude ordtype finmap pcm unionmap heap. +From DiSeL Require Import Freshness State EqTypeX DepMaps. Set Implicit Arguments. Unset Strict Implicit. @@ -53,8 +53,8 @@ Proof. move=>V H/negbTE N i to' tms' b/=. rewrite findUnR ?valid_fresh?V//. case: ifP; last by move=>_/H. -rewrite gen_domPt inE/==>/eqP Z; subst i. -rewrite gen_findPt/=; case=>Z1 Z2 Z3; subst to' tms' from'. +rewrite domPt inE/==>/eqP Z; subst i. +rewrite findPt/=; case=>Z1 Z2 Z3; subst to' tms' from'. by rewrite eqxx in N. Qed. @@ -176,8 +176,8 @@ Lemma no_msg_from_toE from to s tms to': Proof. move=>V H X/= i m b; rewrite findUnR ?valid_fresh?V//. case: ifP; last by move=>_/H. -rewrite gen_domPt inE/==>/eqP Z; subst i. -by rewrite gen_findPt/=; case=>_ Z; subst to'; rewrite eqxx in X. +rewrite domPt inE/==>/eqP Z; subst i. +by rewrite findPt/=; case=>_ Z; subst to'; rewrite eqxx in X. Qed. Lemma no_msg_from_toE' from to s tms from' to': @@ -187,8 +187,8 @@ Lemma no_msg_from_toE' from to s tms from' to': Proof. move=>V H X/= i m b; rewrite findUnR ?valid_fresh?V//. case: ifP; last by move=>_/H. -rewrite gen_domPt inE/==>/eqP Z; subst i. -by rewrite gen_findPt/=; case=>Z' Z; subst from'; rewrite eqxx in X. +rewrite domPt inE/==>/eqP Z; subst i. +by rewrite findPt/=; case=>Z' Z; subst from'; rewrite eqxx in X. Qed. Lemma msg_specE s from to tg cnt : @@ -199,12 +199,12 @@ Proof. move=>V H; split=>/=; last first. - move=>i t c; rewrite findUnR ?valid_fresh?V//. case: ifP; last by move=>_/H. - rewrite gen_domPt inE/==>/eqP Z; subst i. - by rewrite gen_findPt/=; case=>E Z; subst c t; rewrite !eqxx. + rewrite domPt inE/==>/eqP Z; subst i. + by rewrite findPt/=; case=>E Z; subst c t; rewrite !eqxx. exists (fresh s); split=>[|z[t][c]]. - exists tg, cnt; rewrite findUnR ?valid_fresh?V//. - by rewrite gen_domPt inE eqxx/=gen_findPt/=. -rewrite findUnR ?valid_fresh?V// gen_domPt !inE/=. + by rewrite domPt inE eqxx/=findPt/=. +rewrite findUnR ?valid_fresh?V// domPt !inE/=. by case: ifP=>[|_/H]//; move/eqP=>->. Qed. @@ -215,20 +215,20 @@ Lemma msg_specE' s from to to' tg cnt tms : Proof. move=>V N H; split=>//=; last first. - move=>i t c; rewrite findUnR ?valid_fresh?V//. - rewrite gen_domPt inE/=; case:ifP; last by move=>_; move/(proj2 H). - move/eqP=>Z; subst i; rewrite gen_findPt/=; case=>_ Z. + rewrite domPt inE/=; case:ifP; last by move=>_; move/(proj2 H). + move/eqP=>Z; subst i; rewrite findPt/=; case=>_ Z. by subst to'; rewrite eqxx in N. case: (H)=>H' _; case: H'=>i; case=>[[t]][c] U1 U2. exists i; split=>//. - exists t, c; rewrite findUnR ?valid_fresh?V//. - rewrite gen_domPt inE/=; case:ifP=>//. + rewrite domPt inE/=; case:ifP=>//. move/eqP=>Z; subst i. by move/find_some: U1=>E; move:(dom_fresh s); rewrite E. move=>z[t'][c']; rewrite findUnR ?valid_fresh?V//. -rewrite gen_domPt inE/=; case:ifP=>//; last first. +rewrite domPt inE/=; case:ifP=>//; last first. - by move=>_ G; apply: (U2 z); exists t', c'. move/eqP=>Z; subst z. -rewrite gen_findPt/=; case=>Z1 Z2. +rewrite findPt/=; case=>Z1 Z2. by subst to'; rewrite eqxx in N. Qed. @@ -239,20 +239,20 @@ Lemma msg_specE'' s from from' to to' tg cnt tms : Proof. move=>V N H; split=>//=; last first. - move=>i t c; rewrite findUnR ?valid_fresh?V//. - rewrite gen_domPt inE/=; case:ifP; last by move=>_; move/(proj2 H). - move/eqP=>Z; subst i; rewrite gen_findPt/=; case=>_ Z. + rewrite domPt inE/=; case:ifP; last by move=>_; move/(proj2 H). + move/eqP=>Z; subst i; rewrite findPt/=; case=>_ Z. by subst from'; rewrite eqxx in N. case: (H)=>H' _; case: H'=>i; case=>[[t]][c] U1 U2. exists i; split=>//. - exists t, c; rewrite findUnR ?valid_fresh?V//. - rewrite gen_domPt inE/=; case:ifP=>//. + rewrite domPt inE/=; case:ifP=>//. move/eqP=>Z; subst i. by move/find_some: U1=>E; move:(dom_fresh s); rewrite E. move=>z[t'][c']; rewrite findUnR ?valid_fresh?V//. -rewrite gen_domPt inE/=; case:ifP=>//; last first. +rewrite domPt inE/=; case:ifP=>//; last first. - by move=>_ G; apply: (U2 z); exists t', c'. move/eqP=>Z; subst z. -rewrite gen_findPt/=; case=>Z1 Z2. +rewrite findPt/=; case=>Z1 Z2. by subst from'; rewrite eqxx in N. Qed. diff --git a/Core/While.v b/Core/While.v index 95526ba..e9d86bb 100644 --- a/Core/While.v +++ b/Core/While.v @@ -4,18 +4,16 @@ From mathcomp Require Import path. Require Import Eqdep. Require Import Relation_Operators. -From DiSeL.Heaps -Require Import pred prelude idynamic ordtype finmap pcm unionmap heap coding domain. -From DiSeL.Core +From fcsl +Require Import pred prelude ordtype finmap pcm unionmap heap. +From DiSeL Require Import Freshness State EqTypeX DepMaps Protocols Worlds NetworkSem Rely. -From DiSeL.Core +From DiSeL Require Import Actions Injection Process Always HoareTriples InferenceRules. - Set Implicit Arguments. Unset Strict Implicit. Import Prenex Implicits. - Section While. Variable this : nid. Variable W : world. diff --git a/Core/Worlds.v b/Core/Worlds.v index ba6a9c9..5a3a705 100644 --- a/Core/Worlds.v +++ b/Core/Worlds.v @@ -4,9 +4,9 @@ From mathcomp Require Import path. Require Import Eqdep. Require Import Relation_Operators. -From DiSeL.Heaps -Require Import pred prelude idynamic ordtype finmap pcm unionmap heap coding. -From DiSeL.Core +From fcsl +Require Import pred prelude ordtype finmap pcm unionmap heap. +From DiSeL Require Import Freshness State EqTypeX Protocols. Set Implicit Arguments. Unset Strict Implicit. @@ -83,7 +83,7 @@ Definition hooks_consistent (c : context) (h : hooks) : Prop := Definition hook_complete w := hooks_consistent (getc w) (geth w). Lemma hook_complete0 c : hook_complete (c, Unit). -Proof. by move=>????; rewrite dom0 inE. Qed. +Proof. by move=>????; rewrite dom0. Qed. Definition Coh (w : world) : Pred state := fun s => let: c := fst w in @@ -161,8 +161,9 @@ Qed. (* Coherence is trivially precise wrt. statelets *) Lemma coh_prec w: precise (Coh w). Proof. -move=>s1 s2 t1 t2 V E C1 C2. -case: (C1)(C2)=>H1 G1 K1 D1 _ [H2 G2 K2 D2 _]. +move=>s1 s2 t1 t2 V C1 C2. +case: C1 => H1 G1 K1 D1 _. +case: C2 => H2 G2 K2 D2 _ H. by apply: (@dom_prec _ _ _ s1 s2 t1 t2)=>//z; rewrite -D1 -D2. Qed. @@ -201,7 +202,7 @@ Notation l := (plab p). Definition mkWorld : world := (l \\-> p, Unit). Lemma prEq : (getProtocol mkWorld l) = p. -Proof. by rewrite /getProtocol um_findPt. Qed. +Proof. by rewrite /getProtocol findPt. Qed. (* diff --git a/Core/opam b/Core/opam index c7c34e7..f3ed815 100644 --- a/Core/opam +++ b/Core/opam @@ -1,5 +1,4 @@ opam-version: "1.2" -name: "disel-core" version: "dev" maintainer: "palmskog@gmail.com" @@ -10,11 +9,11 @@ license: "BSD2" build: [ make "-j" "%{jobs}%" ] install: [ make "install" ] -remove: [ "sh" "-c" "rm -rf '%{lib}%/coq/user-contrib/DiSeL/Core'" ] +remove: [ "sh" "-c" "rm -rf '%{lib}%/coq/user-contrib/DiSeL'" ] depends: [ - "coq" { ((>= "8.7" & < "8.8~") | (= "dev")) } - "coq-mathcomp-ssreflect" { ((>= "1.6.2" & < "1.7~") | (= "dev")) } - "disel-heaps" {= "dev"} + "coq" { ((>= "8.7" & < "8.9~") | (= "dev")) } + "coq-mathcomp-ssreflect" { ((>= "1.6.2" & < "1.8~") | (= "dev")) } + "coq-fcsl-pcm" ] tags: [ diff --git a/Examples/Calculator/CalculatorClientLib.v b/Examples/Calculator/CalculatorClientLib.v index 6227884..89c40c2 100644 --- a/Examples/Calculator/CalculatorClientLib.v +++ b/Examples/Calculator/CalculatorClientLib.v @@ -4,19 +4,17 @@ From mathcomp Require Import path. Require Import Eqdep. Require Import Relation_Operators. -From DiSeL.Heaps -Require Import pred prelude idynamic ordtype finmap pcm unionmap. -From DiSeL.Heaps -Require Import heap coding domain. -From DiSeL.Core +From fcsl +Require Import axioms pred prelude ordtype finmap pcm unionmap heap. +From DiSeL Require Import Freshness State EqTypeX Protocols Worlds NetworkSem Rely. -From DiSeL.Core +From DiSeL Require Import Actions Injection Process Always HoareTriples InferenceRules. -From DiSeL.Core +From DiSeL Require Import InductiveInv. -From DiSeL.Examples +From DiSeL Require Import CalculatorProtocol CalculatorInvariant. -From DiSeL.Examples +From DiSeL Require Import SeqLib. Section CalculatorRecieve. @@ -51,7 +49,7 @@ Hypothesis Hc : cl \in cls. Program Definition tryrecv_resp_act := act (@tryrecv_action_wrapper W cl (fun k _ t b => (k == l) && (t == resp)) _). -Next Obligation. by case/andP:H=>/eqP->; rewrite gen_domPt inE/=. Qed. +Next Obligation. by case/andP:H=>/eqP->; rewrite domPt inE/=. Qed. Notation loc i := (getLocal cl (getStatelet i l)). Notation st := (ptr_nat 1). @@ -101,7 +99,7 @@ have Y : tms = head 0 tms :: behead tms. by case/andP: Hw=>_; case: (tms)=>//x xs _; exists x, xs. have Y' : from \in cs. - case: (proj1 C)=>Cs _ _ _. case: Cs=>Vs/(_ mid)Cs. - rewrite Es in Vs Cs; move: (um_findPtUn Vs)=>Ez. + rewrite Es in Vs Cs; move: (findPtUn Vs)=>Ez. by move: (Cs _ Ez)=>/=; rewrite/cohMsg/==>H; case: H. (* Using the invariant *) @@ -114,7 +112,7 @@ have X: (cl, from, (behead tms)) \in rs. - by case/andP: Hw; rewrite (getStK (proj1 cohs) L1). have P1: valid (dstate d) by apply: (cohVl C). have P2: valid i2 by apply: (cohS (proj2 (rely_coh R1))). -have P3: l \in dom i2 by rewrite -(cohD(proj2(rely_coh R1)))gen_domPt inE/=. +have P3: l \in dom i2 by rewrite -(cohD(proj2(rely_coh R1)))domPt inE/=. rewrite (rely_loc' _ R3)/= locE// /cr_step (getStK (proj1 cohs) L1)/=. clear R3 Hw P1 P2 P3; exists (remove_elem rs (cl, from, (behead tms))). move: (remove_elem_in rs (cl, from, (behead tms))); rewrite X. @@ -135,7 +133,7 @@ Definition receive_loop_inv (rs : reqs) := | None => loc i = st :-> rs end. -From DiSeL.Core +From DiSeL Require Import While. Program Definition receive_loop' : @@ -227,9 +225,13 @@ have C': coh cal (getStatelet i2 l) by case: C2=>_ _ _ _/(_ l);rewrite prEq. split=>//=. - split=>//=. + by split=>//; case: C'. - + rewrite/Actions.can_send -(cohD C2)/=gen_domPt inE/= eqxx. + + rewrite/Actions.can_send -(cohD C2)/=domPt inE/= eqxx. by rewrite mem_cat Hc orbC. - + by rewrite/Actions.filter_hooks um_filt0=>???/sym/find_some; rewrite dom0 inE. + + rewrite/Actions.filter_hooks umfilt0=>???. + move => F. + apply sym_eq in F. + move: F. + by move/find_some; rewrite dom0. move=>y i3 i4[S]/=;case=>Z[b]/=[F]E3 R3; subst y. case: F=>/=F; subst b i3=>/=. rewrite -(rely_loc' _ R1) in E1. @@ -237,7 +239,7 @@ rewrite (getStK _ E1) in R3. apply: (gh_ex (g:=[:: (cl, server, args)])). apply: call_rule=>//. - move=>C4; rewrite (rely_loc' _ R3) locE//; last by apply: (cohVl C'). - + by rewrite -(cohD C2) gen_domPt inE/=. + + by rewrite -(cohD C2) domPt inE/=. by apply: (cohS C2). clear R3=>v i5[rs'][from][args'][E5]P5 R C. suff X: args = args' /\ rs' = [::] by case: X=>Z X; subst args' rs'. @@ -246,7 +248,7 @@ suff X': rs' = [::]. move/P5: (cl, server, args). by rewrite inE eqxx inE/==>/esym/eqP; case=>_->. by case/perm_eq_size: P5=>/esym/size0nil. -Qed. +Qed. (**************************************************) (* @@ -309,8 +311,3 @@ by move=>i1/=[L1]??; apply: call_rule=>//; rewrite cats0. Qed. End CalculatorRecieve. - - - - - diff --git a/Examples/Calculator/CalculatorExtraction.v b/Examples/Calculator/CalculatorExtraction.v index 8749291..287b4fe 100644 --- a/Examples/Calculator/CalculatorExtraction.v +++ b/Examples/Calculator/CalculatorExtraction.v @@ -1,6 +1,6 @@ -From DiSeL.Core +From DiSeL Require Import DiSeLExtraction. -From DiSeL.Examples +From DiSeL Require Import SimpleCalculatorApp. Cd "extraction". diff --git a/Examples/Calculator/CalculatorInvariant.v b/Examples/Calculator/CalculatorInvariant.v index 79f31fc..3a129d3 100644 --- a/Examples/Calculator/CalculatorInvariant.v +++ b/Examples/Calculator/CalculatorInvariant.v @@ -4,17 +4,15 @@ From mathcomp Require Import path. Require Import Eqdep. Require Import Relation_Operators. -From DiSeL.Heaps -Require Import pred prelude idynamic ordtype finmap pcm unionmap. -From DiSeL.Heaps -Require Import heap coding domain. -From DiSeL.Core +From fcsl +Require Import axioms pred prelude ordtype finmap pcm unionmap heap. +From DiSeL Require Import Freshness State EqTypeX Protocols Worlds NetworkSem Rely. -From DiSeL.Core +From DiSeL Require Import Actions Injection Process Always HoareTriples InferenceRules. -From DiSeL.Core +From DiSeL Require Import InductiveInv StatePredicates. -From DiSeL.Examples +From DiSeL Require Import CalculatorProtocol. Section CalculatorInductiveInv. @@ -60,8 +58,8 @@ Notation cal' := (CalculatorProtocol f prec cs cls l). Notation coh' := (coh cal). Notation Sinv := (@S_inv cal (fun d _ => CalcInv d)). Notation Rinv := (@R_inv cal (fun d _ => CalcInv d)). -Notation PI := proof_irrelevance. -From DiSeL.Examples +Notation PI := pf_irr. +From DiSeL Require Import CalculatorProtocol. Program Definition s1: Sinv (server_send_trans f prec cs cls). @@ -71,7 +69,7 @@ move=>/= Hi E G C' n to' v' args' i s1 N1 N2/= Es. case: (S)=>_[_][C]/hasP[[[me cc]args]]_. case/andP=>/eqP Z1/andP[/eqP Y]/eqP Z2. move: (cohVs C')=>V; rewrite joinC/= in Es V. -move: (um_cancel2 V Es)=>/=; case: ifP. +move: (cancel2 V Es)=>/=; case: ifP. - move=>_; case. case=>Z3 Z4 Z5 Z6; subst s1 to to' msg. by simpl in Y; case: Z2=>->. move=>_[E1]/=E2 E3; subst to. clear Es V. @@ -84,7 +82,7 @@ Proof. move=>this to d msg S b. move=>/= Hi E G C' n to' v' args' i s1 N1 N2/= Es. move: (cohVs C')=>V; rewrite joinC/= in Es V. -move: (um_cancel2 V Es)=>/=; case: ifP; last first. +move: (cancel2 V Es)=>/=; case: ifP; last first. - move=>_[E1]/=E2 E3; clear Es V. by case: (S)=>_[_]C _; apply: (Hi C n to' v' args' i _ _ _ E2). move/eqP=>Z; subst i; by case; discriminate. @@ -106,13 +104,13 @@ move: (cohVs C)=>V; rewrite S1 in V. rewrite S1 joinC consumeUn ?eqxx joinC// in Es. suff V': valid (i \\-> mark_msg vm \+ free (cT:=union_mapUMC mid (msg TaggedMessage)) i (dsoup d)). -- move: (um_cancel2 V' Es); case: ifP=>B. +- move: (cancel2 V' Es); case: ifP=>B. - move/eqP:B=>B{S1 V V' Es}; subst i'. by case: vm=>????/=; rewrite /mark_msg/=; case; discriminate. by case=>_ X2 _; rewrite X2 joinCA in S1; rewrite S1; eexists _. move: (consume_valid i V). -rewrite /consume_msg/= findUnL// ?gen_domPt inE/= eqxx gen_findPt/=. -by rewrite updUnL/= gen_domPt/=!inE eqxx !gen_updPt/=. +rewrite /consume_msg/= findUnL// ?domPt inE/= eqxx findPt/=. +by rewrite updUnL/= domPt/=!inE eqxx !updPt/=. Qed. Program Definition r2: Rinv (client_recv_trans prec cs cls). @@ -131,13 +129,13 @@ move: (cohVs C)=>V; rewrite S1 in V. rewrite S1 joinC consumeUn ?eqxx joinC// in Es. suff V': valid (i \\-> mark_msg vm \+ free (cT:=union_mapUMC mid (msg TaggedMessage)) i (dsoup d)). -- move: (um_cancel2 V' Es); case: ifP=>B. +- move: (cancel2 V' Es); case: ifP=>B. - move/eqP:B=>B{S1 V V' Es}; subst i'. by case: vm=>????/=; rewrite /mark_msg/=; case; discriminate. by case=>_ X2 _; rewrite X2 joinCA in S1; rewrite S1; eexists _. move: (consume_valid i V). -rewrite /consume_msg/= findUnL// ?gen_domPt inE/= eqxx gen_findPt/=. -by rewrite updUnL/= gen_domPt/=!inE eqxx !gen_updPt/=. +rewrite /consume_msg/= findUnL// ?domPt inE/= eqxx findPt/=. +by rewrite updUnL/= domPt/=!inE eqxx !updPt/=. Qed. Definition sts' := [:: SI s1; SI s2]. diff --git a/Examples/Calculator/CalculatorProtocol.v b/Examples/Calculator/CalculatorProtocol.v index 113459e..aa0ebb9 100644 --- a/Examples/Calculator/CalculatorProtocol.v +++ b/Examples/Calculator/CalculatorProtocol.v @@ -4,15 +4,14 @@ From mathcomp Require Import path. Require Import Eqdep. Require Import Relation_Operators. -From DiSeL.Heaps -Require Import pred prelude idynamic ordtype finmap pcm unionmap heap coding domain. -From DiSeL.Core +From fcsl +Require Import axioms pred prelude ordtype finmap pcm unionmap heap. +From DiSeL Require Import Freshness State EqTypeX Protocols Worlds NetworkSem Rely. -From DiSeL.Core +From DiSeL Require Import Actions Injection Process Always HoareTriples InferenceRules. -From DiSeL.Examples +From DiSeL Require Import SeqLib. - Set Implicit Arguments. Unset Strict Implicit. Import Prenex Implicits. @@ -147,17 +146,17 @@ Proof. by rewrite mem_cat=>->. Qed. Lemma cohSt n d (C : CalCoh d) s: find st (getLocal n d) = Some s -> - idyn_tp s = cstate. + dyn_tp s = cstate. Proof. case: (C)=>_ _ _ G; case H: (n \in nodes). -- by move:(G _ H); case=>s'[]->_; rewrite hfindPt//=; case=><-. +- by move:(G _ H); case=>s'[]->_; rewrite findPt//=; case=><-. rewrite /getLocal; rewrite -(cohDom C) in H. by case: dom_find H=>//->; rewrite find0E. Qed. Definition getSt n d (C : CalCoh d) : cstate := match find st (getLocal n d) as f return _ = f -> _ with - Some v => fun epf => icoerce id (idyn_val v) (cohSt C epf) + Some v => fun epf => icast (sym_eq (cohSt C epf)) (dyn_val v) | _ => fun epf => [::] end (erefl _). @@ -166,7 +165,7 @@ Lemma getStK n d (C : CalCoh d) s : Proof. move=>E; rewrite /getSt/=. move: (cohSt C); rewrite !E/==>H. -by apply: ieqc. +by apply: eqc. Qed. Lemma getStE n i j C C' (pf : n \in nodes) : @@ -278,7 +277,7 @@ split=>/=. - split=>[|i ms/=]; first by rewrite valid_fresh (cohVs C). rewrite findUnL; last by rewrite valid_fresh (cohVs C). case: ifP=>E; first by case: C=>[[Vs]]H _ _ _; move/H. - move/um_findPt_inv=>[Z G]; subst i ms. + move/findPt_inv=>[Z G]; subst i ms. split; rewrite ?(proj2 (ss_safe_in pf))?(ss_safe_this pf)//. case: pf=>?[C'][tf]/hasP[]=>[[[n]]]h args _/andP[/eqP]Z1/andP[/eqP Z2]/eqP Z3//=. @@ -348,7 +347,7 @@ split=>/=. - split=>[|i ms/=]; first by rewrite valid_fresh (cohVs C). rewrite findUnL; last by rewrite valid_fresh (cohVs C). case: ifP=>E; first by case: C=>[[Vs]]H _ _ _; move/H. - move/um_findPt_inv=>[Z G]; subst i ms. + move/findPt_inv=>[Z G]; subst i ms. split=>//; rewrite ?(proj2 (cs_safe_in pf)); rewrite ?(proj2 (cs_safe_in pf))?(cs_safe_this pf)//=; by case: pf=>// _ _ _; exists msg. diff --git a/Examples/Calculator/CalculatorServerLib.v b/Examples/Calculator/CalculatorServerLib.v index 58d8dd3..bc78685 100644 --- a/Examples/Calculator/CalculatorServerLib.v +++ b/Examples/Calculator/CalculatorServerLib.v @@ -4,21 +4,19 @@ From mathcomp Require Import path. Require Import Eqdep. Require Import Relation_Operators. -From DiSeL.Heaps -Require Import pred prelude idynamic ordtype finmap pcm unionmap. -From DiSeL.Heaps -Require Import heap coding domain. -From DiSeL.Core +From fcsl +Require Import axioms pred prelude ordtype finmap pcm unionmap heap. +From DiSeL Require Import Freshness State EqTypeX Protocols Worlds NetworkSem Rely. -From DiSeL.Core +From DiSeL Require Import Actions Injection Process Always HoareTriples InferenceRules. -From DiSeL.Core +From DiSeL Require Import InductiveInv. -From DiSeL.Examples +From DiSeL Require Import CalculatorProtocol CalculatorInvariant. -From DiSeL.Examples +From DiSeL Require Import CalculatorClientLib. -From DiSeL.Examples +From DiSeL Require Import SeqLib. Section CalculatorServerLib. @@ -48,7 +46,7 @@ Export CalculatorProtocol. Program Definition tryrecv_req_act := act (@tryrecv_action_wrapper W sv (fun k _ t b => (k == l) && (t == req)) _). -Next Obligation. by case/andP:H=>/eqP->; rewrite gen_domPt inE/=. Qed. +Next Obligation. by case/andP:H=>/eqP->; rewrite domPt inE/=. Qed. (* Receive-transition for the calculator *) Program Definition tryrecv_req : @@ -76,14 +74,14 @@ move=>Z _; subst rt; move: H3; rewrite /msg_wf/=/sr_wf=>->; split=>//. set d := (getStatelet i2 l). have P1: valid (dstate d) by apply: (cohVl cohs). have P2: valid i2 by apply: (cohS (proj2 (rely_coh R1))). -have P3: l \in dom i2 by rewrite -(cohD(proj2(rely_coh R1)))gen_domPt inE/=. +have P3: l \in dom i2 by rewrite -(cohD(proj2(rely_coh R1)))domPt inE/=. rewrite -(rely_loc' _ R1) in E1. - by rewrite (rely_loc' _ R3)/= locE//=/sr_step Hs/= (getStK cohs E1). case: (cohs)=>Cs _ _ _. move/esym: F=> F. -by case: Cs=>_/(_ _ _ F); rewrite /cohMsg/= H2/=; case. +by case: Cs=>_/(_ _ _ F); rewrite /cohMsg/= H2/=; case. Qed. -Definition receive_req_loop_cond (res : option (nid * input)) := res == None . +Definition receive_req_loop_cond (res : option (nid * input)) := res == None. Definition receive_req_loop_inv (ps : reqs) := fun (r : option (nid * input)) i => @@ -94,7 +92,7 @@ Definition receive_req_loop_inv (ps : reqs) := | None => loc i = st :-> ps end. -From DiSeL.Core +From DiSeL Require Import While. Program Definition receive_req_loop : @@ -182,9 +180,13 @@ set d := (getStatelet i2 l). split=>//[|r i3 i4[Sf]St R3]. - split=>//; first 1 last. + by rewrite/Actions.can_send mem_cat Hs/= - -(cohD C2)/= gen_domPt/= inE eqxx. - + rewrite/Actions.filter_hooks um_filt0=>???/sym/find_some. - by rewrite dom0 inE. + -(cohD C2)/= domPt/= inE eqxx. + + rewrite/Actions.filter_hooks umfilt0=>???. + move => F. + apply sym_eq in F. + move: F. + move/find_some. + by rewrite dom0. split=>//; split=>//. exists C; rewrite -(rely_loc' _ R1) in L1; rewrite (getStK C L1). by apply/hasP; exists (to, sv, args)=>//=; rewrite H3 !eqxx. @@ -192,9 +194,9 @@ rewrite (rely_loc' _ R3)=>{R3}. case: St=>->[b]/=[][]->->/=; split=>//. have P1: valid (dstate (getStatelet i2 l)). by apply: (cohVl C). have P2: valid i2 by apply: (cohS (proj2 (rely_coh R1))). -have P3: l \in dom i2 by rewrite -(cohD(proj2(rely_coh R1))) gen_domPt inE/=. +have P3: l \in dom i2 by rewrite -(cohD(proj2(rely_coh R1))) domPt inE/=. rewrite -(rely_loc' _ R1) in L1. -by rewrite (proof_irrelevance (ss_safe_coh _ ) C) locE// (getStK C L1). +by rewrite (pf_irr (ss_safe_coh _ ) C) locE// (getStK C L1). Qed. (**************************************************) diff --git a/Examples/Calculator/DelegatingCalculatorServer.v b/Examples/Calculator/DelegatingCalculatorServer.v index 8fd4cf0..678ec08 100644 --- a/Examples/Calculator/DelegatingCalculatorServer.v +++ b/Examples/Calculator/DelegatingCalculatorServer.v @@ -4,19 +4,17 @@ From mathcomp Require Import path. Require Import Eqdep. Require Import Relation_Operators. -From DiSeL.Heaps -Require Import pred prelude idynamic ordtype finmap pcm unionmap. -From DiSeL.Heaps -Require Import heap coding domain. -From DiSeL.Core +From fcsl +Require Import pred prelude ordtype finmap pcm unionmap heap. +From DiSeL Require Import Freshness State EqTypeX Protocols Worlds NetworkSem Rely. -From DiSeL.Core +From DiSeL Require Import Actions Injection Process Always HoareTriples InferenceRules. -From DiSeL.Core +From DiSeL Require Import InductiveInv While. -From DiSeL.Examples +From DiSeL Require Import CalculatorProtocol CalculatorInvariant. -From DiSeL.Examples +From DiSeL Require Import CalculatorClientLib CalculatorServerLib. Export CalculatorProtocol. @@ -47,7 +45,7 @@ Definition V := W1 \+ W2. Lemma validV : valid V. Proof. rewrite /V; apply/andP=>/=. -split; first by rewrite gen_validPtUn/= gen_validPt/= gen_domPt inE/=. +split; first by rewrite validPtUn/= validPt/= domPt inE/=. by rewrite unitR valid_unit. Qed. @@ -96,47 +94,47 @@ Program Definition delegating_body : ret sv V tt). Lemma hook_complete_unit (c : context) : hook_complete (c, Unit). -Proof. by move=>????; rewrite dom0 inE. Qed. +Proof. by move=>????; rewrite dom0. Qed. Lemma hooks_consistent_unit (c : context) : hooks_consistent c Unit. -Proof. by move=>????; rewrite dom0 inE. Qed. +Proof. by move=>????; rewrite dom0. Qed. Next Obligation. rewrite -(unitR V)/V. have V: valid (W1 \+ W2 \+ Unit) by rewrite unitR validV. apply: (injectL V); do?[apply: hook_complete_unit | apply: hooks_consistent_unit]. -by move=>??????; rewrite dom0 inE. +by move=>??????; rewrite dom0. Defined. Next Obligation. rewrite -(unitR V)/V. have V: valid (W1 \+ W2 \+ Unit) by rewrite unitR validV. apply: (injectR V); do?[apply: hook_complete_unit | apply: hooks_consistent_unit]. -by move=>??????; rewrite dom0 inE. +by move=>??????; rewrite dom0. Qed. Next Obligation. rewrite -(unitR V)/V. have V: valid (W1 \+ W2 \+ Unit) by rewrite unitR validV. apply: (injectL V); do?[apply: hook_complete_unit | apply: hooks_consistent_unit]. -by move=>??????; rewrite dom0 inE. +by move=>??????; rewrite dom0. Defined. Lemma hcomp l : hook_complete (mkWorld l). -Proof. by move=>????; rewrite dom0 inE. Qed. +Proof. by move=>????; rewrite dom0. Qed. Next Obligation. move=>i/=[K1 L1]; apply: vrf_coh=>CD1; apply: step. move: (coh_split CD1 (hcomp cal1) (hcomp cal2))=>[i1[j1]][C1 D1 Z]; subst i. apply: inject_rule=>//. have E1 : loc (i1 \+ j1) l1 = loc i1 l1 - by rewrite (locProjL CD1 _ C1)// gen_domPt inE/=. + by rewrite (locProjL CD1 _ C1)// domPt inE/=. rewrite E1 in K1. apply: with_inv_rule=>//. apply: (gh_ex (g:=[::])). apply: call_rule=>//[[from args]] i2/=[K2]H1 H2 _ j2 CD2 R1. have E2 : loc (i1 \+ j1) l2 = loc j1 l2. - by rewrite (locProjR CD1 _ D1)// gen_domPt inE/=. + by rewrite (locProjR CD1 _ D1)// domPt inE/=. (* Adapt the second protocol's view *) rewrite E2 -(rely_loc' l2 R1) in L1. apply: step; clear C1 D1. @@ -154,9 +152,9 @@ rewrite -(rely_loc' _ R3) in L3; move: L3=>L4. apply: ret_rule=>m R _; rewrite (rely_loc' _ R). move/rely_coh: (R3)=>[]; rewrite injExtL ?(cohW CD2)//. move=>_ D4; clear R3; rewrite !(rely_loc' _ R); clear R. -have X: l2 \in dom W2.1 by rewrite gen_domPt inE eqxx. +have X: l2 \in dom W2.1 by rewrite domPt inE eqxx. rewrite (@locProjR _ _ _ _ _ CD4 X D4); split=>//. -have X': l1 \in dom W1.1 by rewrite gen_domPt inE eqxx. +have X': l1 \in dom W1.1 by rewrite domPt inE eqxx. rewrite /= eqxx in K4; rewrite (@locProjL _ _ _ _ _ CD4 X' _)//. by apply: (cohUnKL CD4 D4); apply: hook_complete_unit. Qed. diff --git a/Examples/Calculator/SimpleCalculatorApp.v b/Examples/Calculator/SimpleCalculatorApp.v index 65ffced..aa91230 100644 --- a/Examples/Calculator/SimpleCalculatorApp.v +++ b/Examples/Calculator/SimpleCalculatorApp.v @@ -4,21 +4,19 @@ From mathcomp Require Import path. Require Import Eqdep. Require Import Relation_Operators. -From DiSeL.Heaps -Require Import pred prelude idynamic ordtype finmap pcm unionmap. -From DiSeL.Heaps -Require Import heap coding domain. -From DiSeL.Core +From fcsl +Require Import pred prelude ordtype finmap pcm unionmap heap. +From DiSeL Require Import Freshness State EqTypeX Protocols Worlds NetworkSem Rely. -From DiSeL.Core +From DiSeL Require Import Actions Injection Process Always HoareTriples InferenceRules. -From DiSeL.Core +From DiSeL Require Import InductiveInv While. -From DiSeL.Examples +From DiSeL Require Import CalculatorProtocol CalculatorInvariant. -From DiSeL.Examples +From DiSeL Require Import CalculatorClientLib CalculatorServerLib. -From DiSeL.Examples +From DiSeL Require Import DelegatingCalculatorServer SimpleCalculatorServers. Export CalculatorProtocol. @@ -68,7 +66,7 @@ Definition V := W1 \+ W2. Lemma validV : valid V. Proof. rewrite /V; apply/andP=>/=. -split; first by rewrite gen_validPtUn/= gen_validPt/= gen_domPt inE/=. +split; first by rewrite validPtUn/= validPt/= domPt inE/=. by rewrite unitR valid_unit. Qed. @@ -100,15 +98,15 @@ Definition init_dstate2 := sv \\-> init_loc \+ sd \\-> init_loc. Lemma valid_init_dstate1 : valid init_dstate1. Proof. case: validUn=>//=; -do?[case: validUn=>//; do?[rewrite ?gen_validPt/=//]|by rewrite gen_validPt/=]. -by move=>k; rewrite !gen_domPt !inE/==>/eqP<-/eqP. +do?[case: validUn=>//; do?[rewrite ?validPt/=//]|by rewrite validPt/=]. +by move=>k; rewrite !domPt !inE/==>/eqP<-/eqP. Qed. Lemma valid_init_dstate2 : valid init_dstate2. Proof. case: validUn=>//=; -do?[case: validUn=>//; do?[rewrite ?gen_validPt/=//]|by rewrite gen_validPt/=]. -by move=>k; rewrite !gen_domPt !inE/==>/eqP<-/eqP. +do?[case: validUn=>//; do?[rewrite ?validPt/=//]|by rewrite validPt/=]. +by move=>k; rewrite !domPt !inE/==>/eqP<-/eqP. Qed. Notation init_dstatelet1 := (DStatelet init_dstate1 Unit). @@ -120,76 +118,76 @@ Definition init_state : state := Lemma validI : valid init_state. Proof. case: validUn=>//=; do?[case: validUn=>//; - do?[rewrite ?gen_validPt/=//]|by rewrite gen_validPt/=]; - by move=>k; rewrite !gen_domPt !inE/==>/eqP<-/eqP. + do?[rewrite ?gen_validPt/=//]|by rewrite validPt/=]; + by move=>k; rewrite !domPt !inE/==>/eqP<-/eqP. Qed. Lemma coh1': calcoh prec cs1 cls1 init_dstatelet1 /\ CalcInv l1 f prec cs1 cls1 init_dstatelet1. Proof. -split; last by move=>?????????/=/esym/empbP/=; rewrite um_empbPtUn. +split; last by move=>?????????/=/esym/empbP/=; rewrite empbPtUn. split=>//; rewrite ?valid_init_dstate1//. - split; first by rewrite valid_unit. by move=>m ms; rewrite find0E. - move=>z; rewrite /=/init_dstate1 domUn !inE/= valid_init_dstate1/=. - by rewrite !um_domPt !inE !(eq_sym z). + by rewrite !domPt !inE !(eq_sym z). move=>n/=; rewrite inE=>/orP; case=>//=. - move/eqP=>->/=; exists [::]=>/=. rewrite /getLocal/init_dstate1/= findUnL?valid_init_dstate1//. - by rewrite um_domPt/= gen_findPt/=. + by rewrite domPt/= findPt/=. rewrite inE=>/eqP=>->; exists [::]=>/=. rewrite /getLocal/init_dstate1/= findUnL?valid_init_dstate1//. -by rewrite um_domPt/= gen_findPt. +by rewrite domPt/= findPt. Qed. Lemma coh1 : l1 \\-> init_dstatelet1 \In Coh W1. Proof. split=>//. - apply/andP; split; last by rewrite valid_unit. - by rewrite ?gen_validPt. -- by rewrite gen_validPt/=. + by rewrite ?validPt. +- by rewrite validPt/=. - by apply: hook_complete_unit. -- by move=>z; rewrite !gen_domPt !inE/=. +- by move=>z; rewrite !domPt !inE/=. move=>k; case B: (l1==k); last first. - have X: (k \notin dom W1.1). - by rewrite /init_state/W1/=!gen_domPt !inE/=; move/negbT: B. - by rewrite /getProtocol /getStatelet/= ?gen_findPt2 eq_sym !B/=. -move/eqP:B=>B; subst k; rewrite prEq/getStatelet/init_state gen_findPt/=. -apply: coh1'. + by rewrite /init_state/W1/=!domPt !inE/=; move/negbT: B. + by rewrite /getProtocol /getStatelet/= ?findPt2 eq_sym !B/=. +move/eqP:B=>B; subst k; rewrite prEq/getStatelet/init_state findPt/=. +exact: coh1'. Qed. Lemma coh2' : calcoh prec cs2 cls2 init_dstatelet2 /\ CalcInv l2 f prec cs2 cls2 init_dstatelet2. Proof. -split; last by move=>?????????/=/esym/empbP/=; rewrite um_empbPtUn. +split; last by move=>?????????/=/esym/empbP/=; rewrite empbPtUn. split=>//; rewrite ?valid_init_dstate2//. - split; first by rewrite valid_unit. by move=>m ms; rewrite find0E//. - move=>z; rewrite /=/init_dstate2 domUn !inE/= valid_init_dstate2//=. - by rewrite !um_domPt !inE !(eq_sym z) orbC. + by rewrite !domPt !inE !(eq_sym z) orbC. move=>n/=; rewrite inE=>/orP; case=>//=. - move/eqP=>->/=; exists [::]=>/=. rewrite /getLocal/init_dstate2/= findUnL?valid_init_dstate2//. - by rewrite um_domPt/= gen_findPt/=. + by rewrite domPt/= findPt/=. rewrite inE=>/eqP=>->; exists [::]=>/=. rewrite /getLocal/init_dstate2/= findUnL?valid_init_dstate2//. -by rewrite um_domPt/= gen_findPt. +by rewrite domPt/= findPt. Qed. Lemma coh2 : l2 \\-> init_dstatelet2 \In Coh W2. Proof. split. - apply/andP; split; last by rewrite valid_unit. - by rewrite ?gen_validPt. -- by rewrite gen_validPt/=. + by rewrite ?validPt. +- by rewrite validPt/=. - by apply: hook_complete_unit. -- by move=>z; rewrite !gen_domPt !inE/=. +- by move=>z; rewrite !domPt !inE/=. move=>k; case B: (l2==k); last first. - have X: (k \notin dom W2.1). - by rewrite /init_state/W2/=!gen_domPt !inE/=; move/negbT: B. - by rewrite /getProtocol /getStatelet/= ?gen_findPt2 eq_sym !B/=. -move/eqP:B=>B; subst k; rewrite prEq/getStatelet/init_state gen_findPt/=. -apply: coh2'. + by rewrite /init_state/W2/=!domPt !inE/=; move/negbT: B. + by rewrite /getProtocol /getStatelet/= ?findPt2 eq_sym !B/=. +move/eqP:B=>B; subst k; rewrite prEq/getStatelet/init_state findPt/=. +exact: coh2'. Qed. Lemma init_coh : init_state \In Coh V. @@ -197,21 +195,22 @@ Proof. split=>//; first by apply: validV. - by apply: validI. - rewrite /V/=/init_state/==>z. -- by move=>???; rewrite domUn !inE/= dom0 inE andbC. +- by move=>???; rewrite domUn !inE/= dom0 andbC. - rewrite /V/init_state=>z; rewrite !domUn !inE; case/andP:validV=>->_/=. - by rewrite validI/= !gen_domPt. + by rewrite validI/= !domPt. move=>k; case B: ((l1 == k) || (l2 == k)); last first. - have X: (k \notin dom V.1). - + by rewrite /V domUn inE/= !gen_domPt!inE/= B andbC. + + by rewrite /V domUn inE/= !domPt!inE/= B andbC. rewrite /getProtocol /getStatelet/=. case: dom_find (X)=>//->_/=; rewrite /init_state. case/negbT/norP: B=>/negbTE N1/negbTE N2. - rewrite findUnL; rewrite ?validI// um_domPt inE N1. -by rewrite gen_findPt2 eq_sym N2/=. + rewrite findUnL; rewrite ?validI// domPt inE N1. + rewrite findPt2 eq_sym N1/=. + by rewrite findPt2 eq_sym N2/=. case/andP: validV=>V1 V2. case/orP:B=>/eqP Z; subst k; -rewrite /getProtocol/V findUnL/= ?V1 ?gen_domPt ?inE/= ?um_findPt; -rewrite /getStatelet ?findUnL/= ?validI// ?gen_domPt ?inE/= ?um_findPt; +rewrite /getProtocol/V findUnL/= ?V1 ?domPt ?inE/= ?findPt; +rewrite /getStatelet ?findUnL/= ?validI// ?domPt ?inE/= ?findPt; [by case: coh1'|by case coh2']. Qed. @@ -237,7 +236,7 @@ Next Obligation. rewrite -(unitR V)/V. have V: valid (W1 \+ W2 \+ Unit) by rewrite unitR validV. apply: (injectL V); do?[apply: hook_complete_unit | apply: hooks_consistent_unit]. -by move=>??????; rewrite dom0 inE. +by move=>??????; rewrite dom0. Qed. Next Obligation. @@ -250,10 +249,10 @@ case: (rely_ext X coh1 R)=>i1[j1][Z]C'; subst i. apply: inject_rule=>//. apply: call_rule=>C1{C'}/=; last by move=>m[H1]H2 H3. have E: (getStatelet i1 l1) = (getStatelet (i1 \+ j1) l1). -- by rewrite (locProjL (proj2 (rely_coh R)) _ C1)=>//; rewrite /W1 um_domPt. +- by rewrite (locProjL (proj2 (rely_coh R)) _ C1)=>//; rewrite /W1 domPt. rewrite E (rely_loc' _ R)/getLocal/=/getStatelet/=. -rewrite findUnL ?validI// um_domPt inE eqxx gen_findPt/=. -by rewrite /init_dstate1 findUnR?valid_init_dstate1// um_domPt/= gen_findPt/=. +rewrite findUnL ?validI// domPt inE eqxx findPt/=. +by rewrite /init_dstate1 findUnR?valid_init_dstate1// domPt/= findPt/=. Qed. (* [S1] Delegating server, serving the client's needs *) @@ -270,13 +269,13 @@ Next Obligation. move=>i/=R; apply: call_rule=>C1//=. rewrite (rely_loc' _ R)/getLocal/=/getStatelet/=. rewrite findUnL ?validI ?valid_init_dstate1//. -rewrite um_domPt inE eqxx gen_findPt/=. +rewrite domPt inE eqxx findPt/=. rewrite findUnR ?validI ?valid_init_dstate1//=. -rewrite gen_domPt inE/= gen_findPt/=; split=>//. +rewrite domPt inE/= findPt/=; split=>//. rewrite -(rely_loc _ R)/=/getStatelet findUnR ?validI ?valid_init_dstate1//=. -rewrite gen_domPt inE/= gen_findPt/= /init_dstate2/=. +rewrite domPt inE/= findPt/= /init_dstate2/=. rewrite findUnL ?validI ?valid_init_dstate2//. -by rewrite gen_domPt inE/= gen_findPt/= /init_dstate2/=. +by rewrite domPt inE/= findPt/= /init_dstate2/=. Qed. (* [S2] A memoizing server, serving as a delegate *) @@ -295,7 +294,7 @@ Next Obligation. rewrite -(unitR V)/V. have V: valid (W1 \+ W2 \+ Unit) by rewrite unitR validV. apply: (injectR V); do?[apply: hook_complete_unit | apply: hooks_consistent_unit]. -by move=>??????; rewrite dom0 inE. +by move=>??????; rewrite dom0. Qed. Next Obligation. @@ -312,11 +311,11 @@ case: (rely_ext X coh2 R)=>j1[i1][Z]C'; subst i. apply: inject_rule=>//=. apply: with_inv_rule; apply:call_rule=>//_. have E: (getStatelet j1 l2) = (getStatelet (j1 \+ i1) l2). -- by rewrite (locProjL (proj2 (rely_coh R)) _ C')=>//; rewrite /W1 um_domPt. +- by rewrite (locProjL (proj2 (rely_coh R)) _ C')=>//; rewrite /W1 domPt. rewrite E (rely_loc' _ R)/getLocal/=/getStatelet/=. rewrite findUnL ?validI//; last by rewrite joinC validI. -rewrite um_domPt/= gen_findPt/=. -by rewrite /init_dstate2 findUnL ?valid_init_dstate2 ?um_domPt/= ?gen_findPt. +rewrite domPt/= findPt/=. +by rewrite /init_dstate2 findUnL ?valid_init_dstate2 ?domPt/= ?findPt. Qed. End CalculatorApp. @@ -328,5 +327,3 @@ End CalculatorApp. Definition c_runner (u : unit) := client_run u. Definition s_runner1 (u : unit) := server1_run u. Definition s_runner2 (u : unit) := server2_run u. - - diff --git a/Examples/Calculator/SimpleCalculatorServers.v b/Examples/Calculator/SimpleCalculatorServers.v index cfcebfa..937768a 100644 --- a/Examples/Calculator/SimpleCalculatorServers.v +++ b/Examples/Calculator/SimpleCalculatorServers.v @@ -4,19 +4,17 @@ From mathcomp Require Import path. Require Import Eqdep. Require Import Relation_Operators. -From DiSeL.Heaps -Require Import pred prelude idynamic ordtype finmap pcm unionmap. -From DiSeL.Heaps -Require Import heap coding domain. -From DiSeL.Core +From fcsl +Require Import pred prelude ordtype finmap pcm unionmap heap. +From DiSeL Require Import Freshness State EqTypeX DepMaps Protocols Worlds NetworkSem Rely. -From DiSeL.Core +From DiSeL Require Import Actions Injection Process Always HoareTriples InferenceRules. -From DiSeL.Core +From DiSeL Require Import InductiveInv While. -From DiSeL.Examples +From DiSeL Require Import CalculatorProtocol CalculatorInvariant. -From DiSeL.Examples +From DiSeL Require Import CalculatorClientLib CalculatorServerLib. Export CalculatorProtocol. @@ -430,5 +428,3 @@ End MemoizingServer. Export OneShotServer. Export BatchingServer. Export MemoizingServer. - - diff --git a/Examples/Greeter/Greeter.v b/Examples/Greeter/Greeter.v index db4660a..8d9de1c 100644 --- a/Examples/Greeter/Greeter.v +++ b/Examples/Greeter/Greeter.v @@ -4,11 +4,11 @@ From mathcomp Require Import path. Require Import Eqdep. Require Import Relation_Operators. -From DiSeL.Heaps -Require Import pred prelude idynamic ordtype finmap pcm unionmap heap coding domain. -From DiSeL.Core +From fcsl +Require Import axioms pred prelude ordtype finmap pcm unionmap heap. +From DiSeL Require Import Freshness State EqTypeX Protocols Worlds NetworkSem Rely. -From DiSeL.Core +From DiSeL Require Import Actions Injection Process Always HoareTriples InferenceRules. Set Implicit Arguments. @@ -96,12 +96,12 @@ Proof. by move=> H; case: C => _ _ _/(_ n H). Qed. Lemma cohN n v : n \in fixed_nodes -> find counter (getLocal n d) = Some v -> - idyn_tp v = nat. -Proof. by move=>H; case: (lcoh H)=>m->; rewrite hfindPt /=; case=><-. Qed. + dyn_tp v = nat. +Proof. by move=>H; case: (lcoh H)=>m->; rewrite findPt /=; case=><-. Qed. Definition getN n (pf : n \in fixed_nodes) : nat := match find counter (getLocal n d) as f return _ = f -> _ with - Some v => fun epf => icoerce id (idyn_val v) (cohN pf epf) + Some v => fun epf => icast (sym_eq (cohN pf epf)) (dyn_val v) | None => fun epf => 0 end (erefl _). @@ -109,7 +109,7 @@ Lemma getNK n (Hn : n \in fixed_nodes) m: getLocal n d = counter :-> m -> getN Hn = m. Proof. move=>E; rewrite /getN; move: (cohN)=>H. -by move: (H n)=>{H}; rewrite E=>H/=; apply: ieqc. +by move: (H n)=>{H}; rewrite E=>H/=; apply: eqc. Qed. End GreetAux. @@ -129,7 +129,6 @@ Definition greet_step (this to : nid) (d : dstatelet) then Some (counter :-> (getN (greet_safe_coh pf) (this_in_pf pf)).+1) else None. - Lemma greet_safe_in this to d m : greet_safe this to d m -> this \in nodes d /\ to \in nodes d. Proof. by case. Qed. @@ -162,10 +161,10 @@ rewrite/greet_step; case:ifP=>///andP[/eqP Z1 H][]Z'; subst b=>/=. case:pf=>[F1]F2 F3 [Cs D]Cl Cn. split=>//=; do?[by rewrite validU/=]. - split; first by rewrite valid_fresh; case: Cs. - move=>m msg'; rewrite findUnR ?um_domPt ?inE ?valid_fresh; last by case: Cs. + move=>m msg'; rewrite findUnR ?domPt ?inE ?valid_fresh; last by case: Cs. case: ifP=>[/eqP Z|_]; last by case: Cs=>_ Cs; move/Cs. case:F3=>[m']/eqP E; subst msg. - by subst m; rewrite gen_findPt/=; case=><-. + by subst m; rewrite findPt/=; case=><-. - move=>n; rewrite domU inE/=; case: ifP=>[/eqP Z| _]//; subst n. by rewrite Cl F1. move: (And4 F1 F2 F3 (And4 Cs D Cl Cn)) => S. @@ -226,7 +225,7 @@ Definition W : world := (l \\-> gp, Unit). Variable this : nid. Lemma grEq : (getProtocol W l) = gp. -Proof. by rewrite /getProtocol um_findPt. Qed. +Proof. by rewrite /getProtocol findPt. Qed. (* Defining an action for greeting on top of the corresponding transition *) (* Sends and argument number and a suffix *) @@ -273,8 +272,12 @@ split=>[|r s5 s6 [Sf2] St R4]. - split; case/rely_coh: R3=>C3 C4=>//. + split=>//; first by eauto. by case: C4=>_ _ _ _/(_ l); rewrite grEq. - + by apply/andP; split=>//; move: (cohD C4)=><-; rewrite um_domPt inE. - by move=>???/=/sym/find_some/=; rewrite /Actions.filter_hooks um_filt0 dom0 inE. + + by apply/andP; split=>//; move: (cohD C4)=><-; rewrite domPt inE; apply/eqP. + move=>???/=. + move => F. + apply sym_eq in F. + move: F. + by move/find_some; rewrite /Actions.filter_hooks umfilt0 dom0. (*Bookkeeping for the postconidtion *) case: St=>Z1[h]/=[St]Z2; subst. @@ -289,7 +292,7 @@ case/rely_coh: (R3)=>C3 C4. erewrite getNK; last by rewrite (rely_loc' l R1); exact: Hloc. rewrite (rely_loc' l R4); split=>//. - rewrite /getLocal/getStatelet findU eqxx/= (cohS C4)/= findU eqxx/=. - rewrite getsE /=; last by rewrite -(cohD C4) um_domPt inE. + rewrite getsE /=; last by rewrite -(cohD C4) domPt inE; apply/eqP. by rewrite (cohVl (coh_coh l C4)). (* Reasoning about message soup *) @@ -298,8 +301,8 @@ case: (rely_send_other' R4 (m := fresh (msgs s4)) (l := l) (tm := {| tag := GreeterProtocol.greet_tag; tms_cont := n :: hello |}) (to := to) (b := true)); last by move=>b[->]/= _; exists b. rewrite /getStatelet findU eqxx/= (cohS C4)/=. -rewrite getsE /=; last by rewrite -(cohD C4) um_domPt inE. -rewrite joinC um_findPtUn // joinC valid_fresh. +rewrite getsE /=; last by rewrite -(cohD C4) domPt inE; apply/eqP. +rewrite joinC findPtUn // joinC valid_fresh. by apply: (cohVs (coh_coh l C4)). Qed. @@ -341,7 +344,7 @@ Definition V := (W l1 nodes) \+ (W l2 nodes). Lemma validV : valid V. Proof. rewrite /V; apply/andP=>/=. -split; first by rewrite gen_validPtUn/= gen_validPt/= gen_domPt inE/=. +split; first by rewrite validPtUn/= validPt/= domPt inE/=. by rewrite unitR valid_unit. Qed. @@ -366,10 +369,10 @@ Definition greeter_spec3 n1 n2 to := r = n1 + n2 + 2]). Lemma hook_complete_unit (c : context) : hook_complete (c, Unit). -Proof. by move=>????; rewrite dom0 inE. Qed. +Proof. by move=>????; rewrite dom0. Qed. Lemma hooks_consistent_unit (c : context) : hooks_consistent c Unit. -Proof. by move=>????; rewrite dom0 inE. Qed. +Proof. by move=>????; rewrite dom0. Qed. Program Definition greet3 n1 n2 to : greeter_spec3 n1 n2 to := Do (r1 <-- iinject (greet_prog1 n1 to); @@ -381,18 +384,18 @@ Next Obligation. rewrite -(unitR V)/V. have V: valid (W l1 nodes \+ W l2 nodes \+ Unit) by rewrite unitR validV. apply: (injectL V); do?[apply: hook_complete_unit | apply: hooks_consistent_unit]. -by move=>??????; rewrite dom0 inE. +by move=>??????; rewrite dom0. Qed. Next Obligation. exact Unit. Defined. Next Obligation. rewrite -(unitR V)/V. have V: valid (W l1 nodes \+ W l2 nodes \+ Unit) by rewrite unitR validV. apply: (injectR V); do?[apply: hook_complete_unit | apply: hooks_consistent_unit]. -by move=>??????; rewrite dom0 inE. +by move=>??????; rewrite dom0. Defined. Lemma hcomp l : hook_complete (W l nodes). -Proof. by move=>????; rewrite dom0 inE. Qed. +Proof. by move=>????; rewrite dom0. Qed. Next Obligation. move=>i/=[H1]H2 N. @@ -401,10 +404,10 @@ apply: step. move: (coh_split C (@hcomp l1) (@hcomp l2))=>[i1[j1]][C1 C2 Z]; subst i. apply: inject_rule=>//. have E1 : loc (i1 \+ j1) l1 = loc i1 l1 - by rewrite (locProjL C _ C1)// gen_domPt inE/=. + by rewrite (locProjL C _ C1)// domPt inE/=. rewrite E1 in H1; apply: call_rule=>// r1 m1 [Q1 Z1] C1' j' C' R1. have E2 : loc (i1 \+ j1) l2 = loc j1 l2. - by rewrite (locProjR C _ C2)// gen_domPt inE/=. + by rewrite (locProjR C _ C2)// domPt inE/=. rewrite E2 -(rely_loc' l2 R1) in H2. apply: step. rewrite joinC; apply: inject_rule. @@ -417,17 +420,12 @@ apply: ret_rule=>m3 R3[G1]G2 G3; split; last first. - by rewrite Z1 Z2 -[_.+1]addn1 -[n2.+1]addn1 addnAC !addnA !addn1 -addn2. - rewrite (rely_loc' l2 R3) joinC; rewrite joinC in C3'. - by rewrite (locProjR C3' _ C3)// gen_domPt inE/=. + by rewrite (locProjR C3' _ C3)// domPt inE/=. rewrite (rely_loc' l1 R3). move/rely_coh: (R2)=>[]; rewrite injExtR ?(cohW C)// =>_ C5. rewrite joinC in C3' *. -rewrite (locProjL C3' _ C5)//?/ddom ?gen_domPt ?inE//=. +rewrite (locProjL C3' _ C5)//?/ddom ?domPt ?inE//=. by rewrite (rely_loc' l1 R2). Qed. End CombineGreeters. - -(* Local Variables: *) -(* coq-prog-name: "coqtop" *) -(* coq-load-path: nil *) -(* End: *) diff --git a/Examples/LockResource/LockProtocol.v b/Examples/LockResource/LockProtocol.v index af915c1..cbbc072 100644 --- a/Examples/LockResource/LockProtocol.v +++ b/Examples/LockResource/LockProtocol.v @@ -4,11 +4,11 @@ From mathcomp Require Import path. Require Import Eqdep. Require Import Relation_Operators. -From DiSeL.Heaps -Require Import pred prelude idynamic ordtype finmap pcm unionmap heap coding domain. -From DiSeL.Core +From fcsl +Require Import axioms pred prelude ordtype finmap pcm unionmap heap. +From DiSeL Require Import Freshness State EqTypeX Protocols Worlds NetworkSem Rely. -From DiSeL.Core +From DiSeL Require Import Actions Injection Process Always HoareTriples InferenceRules. Set Implicit Arguments. @@ -97,7 +97,7 @@ Proof. move=>[H1 H2][y]Cm; split=>[|i ms/=]; first by rewrite valid_fresh. rewrite findUnL; last by rewrite valid_fresh. case: ifP=>E; first by move/H2. -by move/um_findPt_inv=>[Z G]; subst i m; exists y. +by move/findPt_inv=>[Z G]; subst i m; exists y. Qed. Definition state_coh d := @@ -184,28 +184,29 @@ Qed. Lemma getLocal_server_st_tp d (C : LockCoh d) s: find st (getLocal server d) = Some s -> - idyn_tp s = server_state. + dyn_tp s = server_state. Proof. have pf: server \in nodes by rewrite inE eqxx. move: (getLocal_coh C pf); rewrite eqxx; move =>[V][s']Z; rewrite Z in V *. -by rewrite hfindPt -(hvalidPt _ s') V; case=><-. +rewrite findPt /=. +by case =><-. Qed. Lemma getLocal_client_st_tp n d (C : LockCoh d) (H : n \in clients) s: find st (getLocal n d) = Some s -> - idyn_tp s = client_state. + dyn_tp s = client_state. Proof. have pf: n \in nodes by rewrite inE/= orbC H. move: (getLocal_coh C pf); rewrite H=>[[V]]. rewrite client_not_server//. move=>[_][cs] L. rewrite L in V *. -rewrite hfindPt -(hvalidPt _ cs) V. +rewrite findPt /=. by case=> <-. Qed. Definition getSt_server d (C : LockCoh d) : server_state := match find st (getLocal server d) as f return _ = f -> _ with - Some v => fun epf => icoerce id (idyn_val v) (getLocal_server_st_tp C epf) + Some v => fun epf => icast (sym_eq (getLocal_server_st_tp C epf)) (dyn_val v) | _ => fun epf => ServerState [::] 0 None end (erefl _). @@ -216,13 +217,13 @@ move=>E; rewrite /getSt_server/=. have pf: server \in nodes by rewrite inE eqxx. have V: valid (getLocal server d) by case: (getLocal_coh C pf). move: (getLocal_server_st_tp C); rewrite !E=>/= H. -by apply: ieqc. +by apply: eqc. Qed. Program Definition getSt_client c d (C : LockCoh d) (pf : c \in nodes) : client_state. case X: (c \in clients); last by exact: NotHeld. exact: (match find st (getLocal c d) as f return _ = f -> _ with - Some v => fun epf => icoerce id (idyn_val v) (getLocal_client_st_tp C X epf) + Some v => fun epf => icast (sym_eq (getLocal_client_st_tp C X epf)) (dyn_val v) | _ => fun epf => NotHeld end (erefl _)). Defined. @@ -233,7 +234,7 @@ Proof. move=>X E; rewrite /getSt_client/=. have V: valid (getLocal c d) by case: (getLocal_coh C pf). move: (getLocal_client_st_tp C); rewrite X !E=>/= H. -by apply: ieqc. +by apply: eqc. Qed. End GetterLemmas. @@ -298,7 +299,7 @@ move=>n Ni. rewrite /local_coh/=. rewrite /getLocal/=findU; case: ifP=>B; last by case: C=>_ _ _/(_ n Ni). move/eqP: B=>Z; subst n this; rewrite eqxx (cohVl C)/=. split. -by rewrite hvalidPt. +by rewrite validPt. by eexists. Qed. @@ -362,7 +363,7 @@ split=>/=; first by apply: consume_coh. move=>n Ni/=; rewrite /local_coh/=. rewrite /getLocal/=findU; case: ifP=>B/=; last by case: (C)=>_ _ _/(_ n Ni). move/eqP: B X=>Z/eqP X; subst n this; rewrite eqxx (cohVl C)/=. -split; first by rewrite hvalidPt. +split; first by rewrite validPt. by eexists. Qed. @@ -451,7 +452,7 @@ rewrite /getLocal/=findU; case: ifP=>B; last by case: C=>_ _ _/(_ n Ni). move/eqP: B=>Z; subst n. rewrite client_not_server// (cohVl C)/=. split. -- by rewrite hvalidPt. +- by rewrite validPt. split=>//. by eexists. Qed. diff --git a/Examples/LockResource/ResourceProtocol.v b/Examples/LockResource/ResourceProtocol.v index 5cf8b26..d7e5e2d 100644 --- a/Examples/LockResource/ResourceProtocol.v +++ b/Examples/LockResource/ResourceProtocol.v @@ -4,11 +4,11 @@ From mathcomp Require Import path. Require Import Eqdep. Require Import Relation_Operators. -From DiSeL.Heaps -Require Import pred prelude idynamic ordtype finmap pcm unionmap heap coding domain. -From DiSeL.Core +From fcsl +Require Import axioms pred prelude ordtype finmap pcm unionmap heap. +From DiSeL Require Import Freshness State EqTypeX Protocols Worlds NetworkSem Rely. -From DiSeL.Core +From DiSeL Require Import Actions Injection Process Always HoareTriples InferenceRules. Set Implicit Arguments. @@ -120,7 +120,7 @@ Proof. move=>[H1 H2]Cm; split=>[|i ms/=]; first by rewrite valid_fresh. rewrite findUnL; last by rewrite valid_fresh. case: ifP=>E; first by move/H2. -by move/um_findPt_inv=>[Z G]; subst i m. +by move/findPt_inv=>[Z G]; subst i m. Qed. Definition state_coh d := @@ -209,16 +209,16 @@ Qed. Lemma getLocal_server_st_tp d (C : ResourceCoh d) s: find st (getLocal server d) = Some s -> - idyn_tp s = server_state. + dyn_tp s = server_state. Proof. have pf: server \in nodes by rewrite inE eqxx. move: (getLocal_coh C pf); rewrite eqxx; move =>[V][s']Z; rewrite Z in V *. -by rewrite hfindPt -(hvalidPt _ s') V; case=><-. +by rewrite findPt /=; case=><-. Qed. Definition getSt_server d (C : ResourceCoh d) : server_state := match find st (getLocal server d) as f return _ = f -> _ with - Some v => fun epf => icoerce id (idyn_val v) (getLocal_server_st_tp C epf) + Some v => fun epf => icast (sym_eq (getLocal_server_st_tp C epf)) (dyn_val v) | _ => fun epf => ServerState 0 0 [::] end (erefl _). @@ -229,7 +229,7 @@ move=>E; rewrite /getSt_server/=. have pf: server \in nodes by rewrite inE eqxx. have V: valid (getLocal server d) by case: (getLocal_coh C pf). move: (getLocal_server_st_tp C); rewrite !E=>/= H. -by apply: ieqc. +by apply: eqc. Qed. End GetterLemmas. @@ -293,7 +293,7 @@ move=>n Ni. rewrite /local_coh/=. rewrite /getLocal/=findU; case: ifP=>B; last by case: C=>_ _ _/(_ n Ni). move/eqP: B=>Z; subst n this; rewrite eqxx (cohVl C)/=. split. -by rewrite hvalidPt. +by rewrite validPt. by eexists. Qed. @@ -377,7 +377,7 @@ split=>/=; first by apply: consume_coh. move=>n Ni/=; rewrite /local_coh/=. rewrite /getLocal/=findU; case: ifP=>B/=; last by case: (C)=>_ _ _/(_ n Ni). move/eqP: B X=>Z/eqP X; subst n this; rewrite eqxx (cohVl C)/=. -split; first by rewrite hvalidPt. +split; first by rewrite validPt. by eexists. Qed. diff --git a/Examples/Make b/Examples/Make index d53287b..5bc9af8 100644 --- a/Examples/Make +++ b/Examples/Make @@ -1,7 +1,7 @@ -Q . DiSeL.Examples -arg "-w -notation-overridden,-local-declaration,-redundant-canonical-projection,-projection-no-head-constant" -./SeqLib.v +SeqLib.v Greeter/Greeter.v Querying/QueryProtocol.v Querying/QueryHooked.v diff --git a/Examples/Querying/QueryHooked.v b/Examples/Querying/QueryHooked.v index b81c48a..b37d450 100644 --- a/Examples/Querying/QueryHooked.v +++ b/Examples/Querying/QueryHooked.v @@ -3,15 +3,15 @@ Require Import ssreflect ssrbool ssrnat eqtype ssrfun seq. From mathcomp Require Import path. Require Import Eqdep. -From DiSeL.Heaps -Require Import pred prelude idynamic ordtype finmap pcm unionmap heap coding domain. -From DiSeL.Core +From fcsl +Require Import axioms pred prelude ordtype finmap pcm unionmap heap. +From DiSeL Require Import Freshness State EqTypeX Protocols Worlds NetworkSem Rely Actions. -From DiSeL.Examples +From DiSeL Require Import SeqLib QueryProtocol. -From DiSeL.Core +From DiSeL Require Import NewStatePredicates. -From DiSeL.Core +From DiSeL Require Import Actions Injection Process Always HoareTriples InferenceRules. Section QueryHooked. @@ -46,23 +46,29 @@ Hypothesis Lab_neq: lq != (plab pc). Lemma W_valid : valid W. Proof. -apply/andP; split=>/=; last by rewrite um_validPt. -case: validUn=>//; rewrite ?um_validPt// =>l. -rewrite !um_domPt !inE=>/eqP=>Z; subst l=>/eqP=>Z. +apply/andP; split=>/=; last by rewrite validPt. +case: validUn=>//; rewrite ?validPt// =>l. +rewrite !domPt !inE=>/eqP=>Z; subst l=>/eqP=>Z. by subst lq; move/negbTE: Lab_neq; rewrite eqxx. Qed. Lemma W_complete : hook_complete W. Proof. -move=>z lc ls t/=; rewrite um_domPt inE=>/eqP[]_<-<-_. -rewrite !domUn !inE/= !um_domPt !inE !eqxx/= orbC. -by case/andP:W_valid=>->_. +move=>z lc ls t/=; rewrite domPt inE=>/eqP[]_<-<-_. +rewrite !domUn !inE/= !domPt !inE !eqxx/= orbC. +case/andP:W_valid=>->_. +rewrite 3!andTb. +rewrite inE. +by apply/orP; left. Qed. Lemma W_dom : dom W.1 =i [:: plab pc; lq]. Proof. move=>z/=; rewrite domUn !inE/=; case/andP: W_valid=>->/=_. -by rewrite !um_domPt !inE; rewrite -!(eq_sym z). +rewrite !domPt !inE; rewrite -!(eq_sym z). +rewrite /cond /= inE orbC. +apply sym_eq. +by rewrite orbC eq_sym. Qed. Lemma eqW : @@ -77,12 +83,17 @@ Lemma injW : injects (plab pc \\-> pc, Unit) W query_hookz. Proof. rewrite eqW; apply:injectL=>//; try by move=>????/=; rewrite dom0 inE. - by rewrite -eqW W_valid. -- move=>z lc ls t; rewrite um_domPt inE. +- move=>z lc ls t; rewrite domPt inE. + by rewrite /cond /= dom0. +- move=>z lc ls t; rewrite domPt inE. + by rewrite /cond /= dom0. +- move=>z lc ls t; rewrite domPt inE. case/eqP=>Z1 Z2 Z3 Z4; subst ls z lc t. - rewrite !domUn !inE !um_domPt !inE !eqxx/=. - by case/andP:W_valid=>/=->_/=; rewrite orbC. -move=>l; rewrite um_domPt inE=>/eqP=><-. -move=>z lc ls t; rewrite um_domPt inE=>/eqP[]_ _<-_. + rewrite !domUn !inE !domPt !inE !eqxx/=. + case/andP:W_valid=>/=->_/=; rewrite orbC inE. + by apply/orP; left. +move=>l; rewrite domPt inE=>/eqP=><-. +move=>z lc ls t; rewrite domPt inE=>/eqP[]_ _<-_. apply/negbT; apply/eqP=>Z; subst lq; move/negbTE: Lab_neq. by rewrite eqxx. Qed. @@ -90,13 +101,14 @@ Qed. Lemma labC : plab pc \in dom W.1. Proof. case/andP: W_valid=>V1 V2. -by rewrite domUn !inE V1/= um_domPt inE eqxx. +by rewrite domUn !inE V1/= domPt inE eqxx. Qed. Lemma labQ : lq \in dom W.1. Proof. case/andP: W_valid=>V1 V2. -by rewrite domUn !inE V1/= !um_domPt !inE eqxx orbC. +rewrite domUn !inE V1/= !domPt !inE /cond /= inE. +by apply/orP; right. Qed. Lemma injWQ : inj_ext injW = (lq \\-> pq, Unit). @@ -105,29 +117,29 @@ move: (W_valid)=>V; move: (cohK injW). rewrite {1}eqW/mkWorld/= -!joinA /PCM.join/= in V. case/andP: V=>/=V V'. rewrite {1}eqW/mkWorld/= -!joinA /PCM.join/=; case=>H K. -case: (um_cancel V H)=>_; rewrite !unitR=>_{H}H1. +case: (cancel V H)=>_; rewrite !unitR=>_{H}H1. rewrite [inj_ext _]surjective_pairing -H1{H1}; congr (_, _). rewrite !unitL joinC/=/query_hookz/= in V' K. rewrite -[_ \\-> _]unitR in V'. have Z: (1, plab pc, (lq, tresp)) \\-> query_hook \+ Unit = (1, plab pc, (lq, tresp)) \\-> query_hook \+ (inj_ext injW).2 by rewrite unitR. -by case: (um_cancel V' Z). +by case: (cancel V' Z). Qed. Lemma prEqC : getProtocol W (plab pc) = pc. rewrite /getProtocol/W/= findUnL; last by case/andP: W_valid. -by rewrite um_domPt inE eqxx um_findPt. +by rewrite domPt inE eqxx findPt. Qed. Lemma prEqQ : getProtocol W lq = pq. Proof. rewrite /getProtocol/W/= findUnR; last by case/andP: W_valid. -by rewrite um_domPt inE eqxx um_findPt. +by rewrite domPt inE eqxx findPt. Qed. Lemma prEqQ' : getProtocol (lq \\-> pq, Unit) lq = pq. -Proof. by rewrite /getProtocol um_findPt. Qed. +Proof. by rewrite /getProtocol findPt. Qed. (* Finished constructing the composite world *) @@ -174,28 +186,28 @@ move=> N H S; case: (step_coh S)=>C1 _. case: S; [by case=>_<- | move=>l st H1 to'/= msg n H2 H3 C H4 H5 H6->{s'}| move=>l rt H1 i from pf H3 C msg H2/=[H4]_->{s'}]; -rewrite -(cohD C) um_domPt inE in H3; move/eqP: H3=>H3; subst l; +rewrite -(cohD C) domPt inE in H3; move/eqP: H3=>H3; subst l; (* TODO: Get rid of irrelevant cases: by means of a tactic *) rewrite /query_init_state/getStatelet !findU !eqxx (cohS C1)/=; -have Cq: coh pq (getStatelet s lq) by move: (coh_coh lq C1); rewrite /getProtocol um_findPt. +have Cq: coh pq (getStatelet s lq) by move: (coh_coh lq C1); rewrite /getProtocol findPt. (* Send-transitions. *) - case:H=>G1 G2 G3 G4. move: st H1 H4 H5 H6; rewrite /get_st prEqQ'=>st H1 H4 H5 H6. case B: (to == z); rewrite /holds_res_perms/getLocal findU B/=; last first. + split=>//. move=>m t c; rewrite findUnR ?valid_fresh ?(cohVs Cq)//; case: ifP; last by move=>_/G3. - rewrite um_domPt inE=>/eqP<-{m}; rewrite um_findPt; case=>_ _ Z; subst z. + rewrite domPt inE=>/eqP<-{m}; rewrite findPt; case=>_ _ Z; subst z. by move/negbTE: N; rewrite eqxx. + move=>m t c; rewrite findUnR ?valid_fresh ?(cohVs Cq)//; case: ifP; last by move=>_/G4. - rewrite um_domPt inE=>/eqP<-{m}; rewrite um_findPt; case=>_ _ Z; subst z. + rewrite domPt inE=>/eqP<-{m}; rewrite findPt; case=>_ _ Z; subst z. by rewrite eqxx in B. move/eqP: B=>B; subst z; split=>//; first 1 last. + move=>m t c; rewrite findUnR ?valid_fresh ?(cohVs Cq)//; case: ifP; last by move=>_/G3. - rewrite um_domPt inE=>/eqP<-{m}; rewrite um_findPt; case=>_ _ Z; subst to. + rewrite domPt inE=>/eqP<-{m}; rewrite findPt; case=>_ _ Z; subst to. by move/negbTE: N; rewrite eqxx. (* Now two interesting assertions. *) + move=>m t c; rewrite findUnR ?valid_fresh ?(cohVs Cq)//; case: ifP; last by move=>_/G4. - rewrite um_domPt inE=>/eqP<-{m}; rewrite um_findPt; case=>Z' Z2 Z; subst to' t c. + rewrite domPt inE=>/eqP<-{m}; rewrite findPt; case=>Z' Z2 Z; subst to' t c. (* Consider two possible send-transitions. *) case: H1;[|case=>//]; move=>Z; subst st=>//. case: (H4)=>_[C']/=[rid][d][_]; case: G2=>rq[rs][E]H. @@ -214,7 +226,7 @@ have Cq: coh pq (getStatelet s lq) by move: (coh_coh lq C1); rewrite /getProtoco (* Receive-transitions *) case:H=>G1 G2 G3 G4. move: (coh_s _ _) rt pf H1 H2. -rewrite /get_rt prEqQ'/==>Cq'; rewrite !(proof_irrelevance Cq' Cq)=>{Cq'}. +rewrite /get_rt prEqQ'/==>Cq'; rewrite !(pf_irr Cq' Cq)=>{Cq'}. move=>rt pf H1 H2. case B: (to == z); rewrite /holds_res_perms/getLocal findU B/=; last first. - by split=>//; apply: no_msg_from_to_consume'=>//; rewrite ?(cohVs Cq). @@ -240,7 +252,7 @@ Proof. move=>H1 [n]H2; elim: n s H2 H1=>/=[s | n Hi s]; first by case=>Z _; subst s'. case=>z[s1][N]H1 H2 H3; apply: (Hi s1 H2). by apply: (query_init_step' _ _ _ _ N H3 H1). -Qed. +Qed. (* TODO: this is a good candidate for a general lemma in the framework. *) @@ -259,8 +271,8 @@ move=>i1[j1][C1 D1 Z]; subst s. case: (coh_split CD' _ _); try apply: hook_complete0. move=>i2[j2][C2 D2 Z]; subst s2. rewrite /query_init_state in Q *. -rewrite (locProjR CD _ D1) in Q; last by rewrite gen_domPt inE andbC eqxx. -rewrite (locProjR CD' _ D2); last by rewrite gen_domPt inE andbC eqxx. +rewrite (locProjR CD _ D1) in Q; last by rewrite domPt inE andbC eqxx. +rewrite (locProjR CD' _ D2); last by rewrite domPt inE andbC eqxx. case: (rely_split injW C1 C2 R)=>Rc Rq; rewrite injWQ in Rq. by apply: (query_init_rely' _ _ _ Q Rq). Qed. @@ -307,7 +319,7 @@ Hypothesis core_state_stable_step : forall z s data s' n, core_state_to_data n (getLc' s' n) data. Lemma prEqC' : (getProtocol (plab pc \\-> pc, Unit) (plab pc)) = pc. -Proof. by rewrite /getProtocol gen_findPt/=. Qed. +Proof. by rewrite /getProtocol findPt/=. Qed. (* Showing that the core assertion is stable *) Lemma core_state_stable_step_W s data s' z : @@ -329,14 +341,14 @@ move: (C1) (C2) =>C1' C2'. rewrite E1 E2 in H2 C1' C2'. have [X X'] : (forall z, getLc' s z = getLc' s1 z) /\ (forall z, getLc' s' z = getLc' s1' z). -- by split; move=>z'; rewrite /getStatelet find_um_filt um_domPt inE eqxx. +- by split; move=>z'; rewrite /getStatelet find_umfilt domPt inE eqxx. case: (sem_split injW _ _ H2). - apply: (@projectS_cohL (plab pc \\-> pc, Unit) ((plab pq \\-> pq, Unit) \+ (Unit, query_hookz)))=>//. - by move=>????; rewrite dom0 inE. + by move=>????; rewrite dom0. - suff H : hook_complete (plab pc \\-> pc, Unit). apply: (@projectS_cohL _ ((plab pq \\-> pq, Unit)\+(Unit, query_hookz)))=>//. - by move=>????; rewrite dom0 inE. + by move=>????; rewrite dom0. - case=>H _; rewrite X'; rewrite !X in L H1. by apply: (core_state_stable_step _ _ _ _ _ N H _ L H1). by case=>_ E; rewrite X'; rewrite !X in L H1; rewrite -E. @@ -387,7 +399,7 @@ move=>H5 H6; case: M=>G1 G2 G3 G4[rq][rs][E]Np; split=>//. - rewrite /getStatelet findU eqxx(cohS C)/==>z t c. rewrite findUnR ?(valid_fresh) ?(cohVs (cohQ s C))//. case: ifP; [|by move=>_; apply: G2]. - rewrite um_domPt inE=>/eqP<-; rewrite um_findPt; case=>???; subst t c to'. + rewrite domPt inE=>/eqP<-; rewrite findPt; case=>???; subst t c to'. case: H1;[|case=>//]; move=>Z; subst st=>//; simpl in H5, H6; rewrite/QueryProtocol.send_step (getStK _ E)/= in H6; case: H6=>H6/=; subst n. by case:H4=>_[C'][rid][_][_]; rewrite (getStK _ E); move/Np. @@ -395,12 +407,12 @@ move=>H5 H6; case: M=>G1 G2 G3 G4[rq][rs][E]Np; split=>//. case:G4; case=>i[[c']Q1 Q3] Q2; split; last first. + move=>j c; rewrite findUnR ?(valid_fresh) ?(cohVs (cohQ s C))//. case:ifP; last by move=>_; apply Q2. - rewrite um_domPt inE=>/eqP<-; rewrite um_findPt. + rewrite domPt inE=>/eqP<-; rewrite findPt. by case=>???; subst to; rewrite eqxx in N. exists i; split=>[|j[c1]];rewrite findUnL ?(valid_fresh) ?(cohVs (cohQ s C))//. + by exists c'; rewrite (find_some Q1). case: ifP=>_; first by move=> T; apply: Q3; exists c1. - by case/um_findPt_inv=>_[]_ _?; subst to; rewrite eqxx in N. + by case/findPt_inv=>_[]_ _?; subst to; rewrite eqxx in N. rewrite /getStatelet findU eqxx (cohS C)/=/holds_res_perms. rewrite /getLocal/=findU eqxx/= (cohVl (cohQ s C)). case: H1;[|case=>//]; move=>Z; subst st=>//; simpl in H5, H6; @@ -431,23 +443,23 @@ case: M=>G1 G2 G3 G4[rq][rs][E]Np; split=>//. - rewrite /getStatelet findU eqxx(cohS C)/==>z t c. rewrite findUnR ?(valid_fresh) ?(cohVs (cohQ s C))//. case: ifP; [|by move=>_; apply: G3]. - rewrite um_domPt inE=>/eqP<-; rewrite um_findPt; case=>????; subst t c to to'. + rewrite domPt inE=>/eqP<-; rewrite findPt; case=>????; subst t c to to'. by rewrite eqxx in N. - rewrite /getStatelet findU eqxx(cohS C)/=. case:G4; case=>i[[c']Q1 Q3] Q2; split; last first. + move=>j c; rewrite findUnR ?(valid_fresh) ?(cohVs (cohQ s C))//. case:ifP; last by move=>_; apply Q2. - rewrite um_domPt inE=>/eqP<-; rewrite um_findPt. + rewrite domPt inE=>/eqP<-; rewrite findPt. case=>???; subst to' c. case: H1;[|case=>//]; move=>Z; subst st=>//; simpl in H5, H6; - rewrite/QueryProtocol.send_step (getStK _ E)/= in H6; case: H6=>H6/=; subst n. + rewrite/QueryProtocol.send_step (getStK _ E)/= in H6; case: H6=>H6/=; subst n. by case:H4=>_[C']; rewrite (getStK _ E); case=>rid[d][->]/=/Np. exists i; split=>[|j[c1]];rewrite findUnL ?(valid_fresh) ?(cohVs (cohQ s C))//. + by exists c'; rewrite (find_some Q1). case: ifP=>_; first by move=> T; apply: Q3; exists c1. - case/um_findPt_inv=>_[]???; subst to'. + case/findPt_inv=>_[]???; subst to'. case: H1;[|case=>//]; move=>Z; subst st=>//; simpl in H5, H6; - rewrite/QueryProtocol.send_step (getStK _ E)/= in H6; case: H6=>H6/=; subst n. + rewrite/QueryProtocol.send_step (getStK _ E)/= in H6; case: H6=>H6/=; subst n. by case:H4=>_[C']; rewrite (getStK _ E); case=>rid[d][_]/=/Np. rewrite /getStatelet findU eqxx (cohS C)/=/holds_res_perms. rewrite /getLocal/=findU eqxx/= (cohVl (cohQ s C)). @@ -489,14 +501,14 @@ Ltac kill_g3 s C G3 to N := rewrite /getStatelet findU eqxx(cohS C)/==>z t c; rewrite findUnR ?(valid_fresh) ?(cohVs (cohQ s C))//; case: ifP; [|by move=>_; apply: G3]; - rewrite um_domPt inE=>/eqP<-; rewrite um_findPt; + rewrite domPt inE=>/eqP<-; rewrite findPt; by case=>_ _ Z _; subst to; rewrite eqxx in N. Ltac kill_g4 s C G4 to' t B:= rewrite /getStatelet findU eqxx(cohS C)/==>z t c; rewrite findUnR ?(valid_fresh) ?(cohVs (cohQ s C))//; case: ifP; [|by move=>_; apply: G4]; - rewrite um_domPt inE=>/eqP<-; rewrite um_findPt; + rewrite domPt inE=>/eqP<-; rewrite findPt; by case=>??; try move=>?; try subst t; try subst to'; try rewrite ?eqxx in B. (* now he can send us the message, this is the crucial step. *) @@ -529,7 +541,7 @@ case: (H4)=>_[C']/=[rid][d][Zm]Tr; subst msg. rewrite (getStK _ E)/= in Tr. rewrite Tr; constructor 3. (* Now exploit the hook! *) -move: (H5 1 (plab pc) query_hook); rewrite um_findPt -!(cohD C). +move: (H5 1 (plab pc) query_hook); rewrite findPt -!(cohD C). move/(_ (erefl _) labC labQ to rid d (erefl _))=>D. move: (core_state_to_data_inj _ _ _ _ D H)=>Z; subst d. split=>//; first 3 last. @@ -539,26 +551,26 @@ split=>//; first 3 last. move/Np/eqP: (mem_rem R)=>Zr; subst rn. move/Np:Tr=>/eqP=>Zr; subst rid. case: (C')=>_ _ _/(_ to); rewrite E (cohDom C')=>/(_ Qn). - case=>[[x1 x2]][]/(hcancelPtV _); rewrite hvalidPt/==>/(_ is_true_true). + case=>[[x1 x2]][]/(hcancelPtV _); rewrite validPt/==>/(_ is_true_true). case=>Z1 Z2; subst x1 x2=>/andP[_]G. by rewrite (rem_filter (this, req_num) G) mem_filter/= eqxx/= in R. + by rewrite /getStatelet findU eqxx(cohS C)/getLocal findU (negbTE N). + rewrite /getStatelet findU eqxx (cohS C)/==>z t c. rewrite findUnR ?(valid_fresh) ?(cohVs C')//. case: ifP=>[|_]; last by apply: G3. - by rewrite um_domPt inE=>/eqP<-{z}; rewrite um_findPt; case=><-. + by rewrite domPt inE=>/eqP<-{z}; rewrite findPt; case=><-. (* Here's the big hooking moment *) split=>[|i m]; rewrite /getStatelet findU eqxx/= (cohS C)/=; last first. + rewrite findUnR ?(valid_fresh) ?(cohVs C')//; case: ifP=>[|_]; last by move/G4. - rewrite um_domPt inE=>/eqP<-; rewrite um_findPt; case=><-. + rewrite domPt inE=>/eqP<-; rewrite findPt; case=><-. by move/Np/eqP: Tr=>->; rewrite eqxx. exists (fresh (dsoup (getStatelet s lq))); split. + exists (rid :: serialize data). - by rewrite findUnR ?(valid_fresh)?(cohVs C')// um_domPt inE eqxx um_findPt. + by rewrite findUnR ?(valid_fresh)?(cohVs C')// domPt inE eqxx findPt. move=>i[c]; rewrite findUnR ?(valid_fresh)?(cohVs C')//. case: ifP=>[|_]; last by move/G4. -by rewrite um_domPt inE=>/eqP<-; rewrite um_findPt. +by rewrite domPt inE=>/eqP<-; rewrite findPt. Qed. (* A lemma covering send-cases in the lc-part of the world *) @@ -747,7 +759,7 @@ move=> N H S; case: H=> Qn H L M; split=>//. case: S; [by case=>_<- | move=>l st H1 to'/= msg n H2 H3 C H4 H5 H6->{s'}| move=>l rt H1 i from pf H3 C msg H2/=[H4]_->{s'}]; - rewrite -(cohD C) domUn !inE !um_domPt !inE in H3; + rewrite -(cohD C) domUn !inE !domPt !inE in H3; case/andP:H3=>_ H3; case/orP: H3=>/eqP H3; subst l. (* Something sending in lc, should be irrelevant for us. *) @@ -788,7 +800,7 @@ move=> N H S; case: H=> Qn H L M; split=>//. case: S; [by case=>_<- | move=>l st H1 to'/= msg n H2 H3 C H4 H5 H6->{s'}| move=>l rt H1 i from pf H3 C msg H2/=[H4]_->{s'}]; - rewrite -(cohD C) domUn !inE !um_domPt !inE in H3; + rewrite -(cohD C) domUn !inE !domPt !inE in H3; case/andP:H3=>_ H3; case/orP: H3=>/eqP H3; subst l; do? [by rewrite /getStatelet findU (negbTE Lab_neq)]; rewrite /getStatelet findU eqxx(cohS C)/=. @@ -802,28 +814,28 @@ case:M; case=>G1 G2 G3 G4 G5; [constructor 1|constructor 2|constructor 3]. - split=>//=; first by rewrite /getLocal/= findU (negbTE N) in G1 *. + move=>i t c; rewrite findUnR ?(valid_fresh) ?(cohVs (cohQ s C))//. case:ifP; last by move=>_; apply G2. - rewrite um_domPt inE=>/eqP<-; rewrite um_findPt. + rewrite domPt inE=>/eqP<-; rewrite findPt. by case=>???; subst to; rewrite eqxx in A. + case: G4; case=>i[[c']Q1 Q3] Q2; split; last first. - move=>j c; rewrite findUnR ?(valid_fresh) ?(cohVs (cohQ s C))//. case:ifP; last by move=>_; apply Q2. - rewrite um_domPt inE=>/eqP<-; rewrite um_findPt. + rewrite domPt inE=>/eqP<-; rewrite findPt. by case=>????; subst z to' c; rewrite eqxx in N. exists i; split=>[|j[c1]];rewrite findUnL ?(valid_fresh) ?(cohVs (cohQ s C))//. - by exists c'; rewrite (find_some Q1). case: ifP=>_; first by move=> T; apply: Q3; exists c1. - by case/um_findPt_inv=>_[]????; subst to' z; rewrite eqxx in N. + by case/findPt_inv=>_[]????; subst to' z; rewrite eqxx in N. by case:G5=>rq[rs][Tr]Np; exists rq, rs; rewrite /getLocal/=findU A/=. (* Send-transition, case 2 *) - split=>//=; first by rewrite /getLocal/= findU (negbTE N) in G1 *. + move=>i t c; rewrite findUnR ?(valid_fresh) ?(cohVs (cohQ s C))//. case:ifP; last by move=>_; apply G3. - rewrite um_domPt inE=>/eqP<-; rewrite um_findPt. + rewrite domPt inE=>/eqP<-; rewrite findPt. by case=>????; subst z to'; rewrite eqxx in N. + move=>i t c; rewrite findUnR ?(valid_fresh) ?(cohVs (cohQ s C))//. case:ifP; last by move=>_; apply G4. - rewrite um_domPt inE=>/eqP<-; rewrite um_findPt. + rewrite domPt inE=>/eqP<-; rewrite findPt. by case=>????; subst z to'; rewrite eqxx in A. by case:G5=>rq[rs][Tr]Np; exists rq, rs; rewrite /getLocal/=findU A/=. @@ -831,17 +843,17 @@ case:M; case=>G1 G2 G3 G4 G5; [constructor 1|constructor 2|constructor 3]. - split=>//=; first by rewrite /getLocal/= findU (negbTE N) in G1 *. + move=>i t c; rewrite findUnR ?(valid_fresh) ?(cohVs (cohQ s C))//. case:ifP; last by move=>_; apply G3. - rewrite um_domPt inE=>/eqP<-; rewrite um_findPt. + rewrite domPt inE=>/eqP<-; rewrite findPt. by case=>????; subst to' z; rewrite eqxx in N. + case: G4; case=>i[[c']Q1 Q3] Q2; split; last first. - move=>j c; rewrite findUnR ?(valid_fresh) ?(cohVs (cohQ s C))//. case:ifP; last by move=>_; apply Q2. - rewrite um_domPt inE=>/eqP<-; rewrite um_findPt. + rewrite domPt inE=>/eqP<-; rewrite findPt. by case=>????; subst z to' c; rewrite eqxx in A. exists i; split=>[|j[c1]];rewrite findUnL ?(valid_fresh) ?(cohVs (cohQ s C))//. - by exists c'; rewrite (find_some Q1). case: ifP=>_; first by move=> T; apply: Q3; exists c1. - by case/um_findPt_inv=>_[]????; subst to' z; rewrite eqxx in A. + by case/findPt_inv=>_[]????; subst to' z; rewrite eqxx in A. by case:G5=>rq[rs][Tr]Np; exists rq, rs; rewrite /getLocal/=findU A/=. (** Receive-transitions **) @@ -869,7 +881,7 @@ case:M; case=>G1 G2 G3 G4 G5; [constructor 1|constructor 2|constructor 3]. by rewrite (negbTE N) orbC. by case:G5=>rq[rs][Tr]Np; exists rq, rs; rewrite /getLocal/=findU A/=. Qed. - + Lemma msg_story_rely req_num to data reqs resp s s2 : msg_story s req_num to data reqs resp -> @@ -953,8 +965,11 @@ apply: act_rule=>i1 R0; split=>//=[|r i2 i3[Hs]St R2]. split=>//; [split=>//; try by case:Q| | ]. + by exists Cq1; rewrite /QueryProtocol.send_req_prec (getStK _ P1)/= P2. + by apply/andP; split=>//; rewrite -(cohD C1) W_dom !inE eqxx orbC. - move=>z lc hk; rewrite find_um_filt eqxx /query_hookz/==>/sym. - by move/find_some; rewrite um_domPt !inE=>/eqP. + move=>z lc hk; rewrite find_umfilt eqxx /query_hookz. + move => F. + apply sym_eq in F. + move: F. + by move/find_some; rewrite domPt !inE=>/eqP. (* Postcondition *) have N: network_step W this i1 i2. - apply: (Actions.send_act_step_sem _ _ St)=>//; first by rewrite prEqQ. @@ -964,7 +979,7 @@ rewrite -(rely_loc' _ R0) in P1. move/rely_coh: (R0)=>[_]C1; move: (coh_coh lq C1);rewrite prEqQ=>Cq1. case: St=>->[h]/=[]. rewrite/QueryProtocol.send_step/QueryProtocol.send_step_fun/=. -rewrite (proof_irrelevance (QueryProtocol.send_safe_coh _) Cq1). +rewrite (pf_irr (QueryProtocol.send_safe_coh _) Cq1). rewrite (getStK _ P1); case=>Z Z'; subst h rid. rewrite Z' locE; last first; [by apply: cohVl Cq1| by apply: cohS C1| @@ -996,21 +1011,21 @@ split=>//; rewrite ?inE ?eqxx=>//=. - move=>m t c; rewrite /getStatelet findU eqxx (cohS C1)/=. set ds := (dsoup _); rewrite findUnR; last first. by rewrite valid_fresh; apply: cohVs Cq1. - rewrite gen_domPt !inE/=; case:ifP=>[/eqP<-|_]; last by apply: Q4. - by rewrite um_findPt; case=><-. + rewrite domPt !inE/=; case:ifP=>[/eqP<-|_]; last by apply: Q4. + by rewrite findPt; case=><-. (* TODO: refactor this into a separate lemma about those state predicates *) - rewrite /getStatelet findU eqxx (cohS C1)/=. set ds := (dsoup _); split=>//. exists (fresh ds); split=>//. + exists [:: fresh_id reqs]. rewrite findUnR; last by rewrite valid_fresh; apply: cohVs Cq1. - by rewrite um_domPt inE eqxx um_findPt. + by rewrite domPt inE eqxx findPt. + move=>x[c]; rewrite findUnR; last by rewrite valid_fresh; apply: cohVs Cq1. - case:ifP=>[|_[]]; first by rewrite um_domPt inE=>/eqP->. + case:ifP=>[|_[]]; first by rewrite domPt inE=>/eqP->. by move/(Q3 x treq c). move=>x c ; rewrite findUnR; last by rewrite valid_fresh; apply: cohVs Cq1. case:ifP=>[|_[]]; last by move/(Q3 x treq c). - by rewrite um_domPt inE=>/eqP->; rewrite um_findPt;case=>->. + by rewrite domPt inE=>/eqP->; rewrite findPt;case=>->. set ds := (dsoup _). case: Q2=>reqs'[resp'][G1 G2]. case X: (to == this). @@ -1039,9 +1054,9 @@ Program Definition tryrecv_resp (rid : nat) (to : nid) := (fun k n t (b : seq nat) => [&& k == lq, n == to, t == tresp, head 0 b == rid & to \in qnodes]) _). Next Obligation. -case/andP:H=>/eqP=>->_; rewrite joinC domUn inE um_domPt inE eqxx andbC/=. -case: validUn=>//=; rewrite ?um_validPt//. -move=>k; rewrite !um_domPt !inE=>/eqP<-/eqP Z; subst lq. +case/andP:H=>/eqP=>->_; rewrite joinC domUn inE domPt inE eqxx andbC/=. +case: validUn=>//=; rewrite ?validPt//. +move=>k; rewrite !domPt !inE=>/eqP<-/eqP Z; subst lq. by move/negbTE: Lab_neq; rewrite eqxx. Qed. @@ -1065,7 +1080,7 @@ Definition recv_resp_inv (rid : nat) to local_indicator data (getLc i) & msg_story i rid to data ((to, rid) :: reqs) resp]. -From DiSeL.Core +From DiSeL Require Import While. Program Definition receive_resp_loop (rid : nat) to : @@ -1140,7 +1155,7 @@ move: (msg_story_rely _ _ _ _ _ _ _ Q3 R1)=>{Q3}Q3. have P1: valid (dstate (getSq i1)) by case: (C1')=>P1 P2 P3 P4. have P2: valid i1 by apply: (cohS C1). have P3: lq \in dom i1. -- rewrite -(cohD C1) domUn inE !um_domPt !inE/= eqxx orbC andbC/=. + rewrite -(cohD C1) domUn inE !domPt !inE/= inE eqxx orbC andbC/=. by case/andP: W_valid. clear Hin R1 C0 i0 i3 Hw. (* Consider different cases for msg_story. *) @@ -1155,7 +1170,8 @@ case=>X1 _ X2 X3 X4; rewrite /query_init_state. subst i2. rewrite /receive_step/=/QueryProtocol.receive_step !(getStK C1' X1)/=!inE!eqxx/=. rewrite !locE ?Qn/=;[|by case: C1'|by apply: cohS C1|by case: C1']. -move/sym:E=>E; split=>//; last 1 first. +apply sym_eq in E. +split=>//; last 1 first. - case:X3; case=>m'[[c]]E'/(_ mid) H. have Z: m' = mid by apply: H; exists (tms_cont tms); rewrite E; case: (tms) E2=>/=t c'->. @@ -1171,7 +1187,7 @@ split=>//. by rewrite /getLocal/= findU; case:ifP=>//=/eqP Z; subst to; rewrite eqxx in B. move/eqP:B=>B; subst to; case: X4=> rq[rs][G1 G2]. rewrite G1 in X1. - have V: valid (qst :-> (rq, rs)) by apply: hvalidPt. + have V: valid (qst :-> (rq, rs)) by apply: validPt. case: (hcancelPtV V X1)=>Z1 Z2; subst rq rs; exists reqs, resp; split=>//. rewrite /getStatelet/= findU eqxx (cohS C1)/=/getLocal/=findU eqxx/=. by case: C1'=>_->. diff --git a/Examples/Querying/QueryPlusTPC.v b/Examples/Querying/QueryPlusTPC.v index 1ff6510..038a43c 100644 --- a/Examples/Querying/QueryPlusTPC.v +++ b/Examples/Querying/QueryPlusTPC.v @@ -3,21 +3,21 @@ Require Import ssreflect ssrbool ssrnat eqtype ssrfun seq. From mathcomp Require Import path. Require Import Eqdep. -From DiSeL.Heaps -Require Import pred prelude idynamic ordtype finmap pcm unionmap heap coding domain. -From DiSeL.Core +From fcsl +Require Import pred prelude ordtype finmap pcm unionmap heap. +From DiSeL Require Import Freshness State EqTypeX Protocols Worlds NetworkSem Rely. -From DiSeL.Core +From DiSeL Require Import NewStatePredicates. -From DiSeL.Examples +From DiSeL Require Import SeqLib. -From DiSeL.Core +From DiSeL Require Import Actions Injection Process Always HoareTriples InferenceRules. -From DiSeL.Examples +From DiSeL Require Import TwoPhaseProtocol TwoPhaseCoordinator TwoPhaseParticipant. -From DiSeL.Examples +From DiSeL Require TwoPhaseInductiveProof. -From DiSeL.Examples +From DiSeL Require Import QueryProtocol QueryHooked. Section QueryPlusTPC. @@ -63,12 +63,12 @@ Proof. rewrite/core_state_to_data. case:ifP=>_ E; rewrite E ![_ \+ log :-> _]joinC=>{E}E. - have V: valid (log :-> d.2 \+ st :-> (d.1, CInit)). - - by case: validUn=>//k; rewrite !hdomPt !inE/==>/eqP<-. + - by case: validUn=>//k; rewrite !domPt !inE/==>/eqP<-. case: (hcancelV V E)=>E2=>{V E}V E. case: (hcancelPtV V E)=>E1. by rewrite [d]surjective_pairing [d']surjective_pairing E1 E2. have V: valid (log :-> d.2 \+ st :-> (d.1, PInit)). -- by case: validUn=>//k; rewrite !hdomPt !inE/==>/eqP<-. +- by case: validUn=>//k; rewrite !domPt !inE/==>/eqP<-. case: (hcancelV V E)=>E2=>{V E}V E. case: (hcancelPtV V E)=>E1. by rewrite [d]surjective_pairing [d']surjective_pairing E1 E2. @@ -183,46 +183,46 @@ move: (C0)=>CD0; rewrite /W eqW in CD0; move: (coh_hooks CD0)=>{CD0}CD0. case: (coh_split CD0); try apply: hook_complete0. move=>i1[j1][C1 D1 Z]. subst i0; apply: inject_rule=>//. -have E : loc_tpc (i1 \+ j1) = loc_tpc i1 by rewrite (locProjL CD0 _ C1)// gen_domPt inE andbC eqxx. +have E : loc_tpc (i1 \+ j1) = loc_tpc i1 by rewrite (locProjL CD0 _ C1)// domPt inE andbC eqxx. rewrite E{E} in P1. apply: with_inv_rule'. apply: call_rule=>//_ i2 [chs]L2 C2 Inv j2 CD2/= R. (* Massaging the complementary state *) -have E : loc_qry (i1 \+ j1) = loc_qry j1 by rewrite (locProjR CD0 _ D1)// gen_domPt inE andbC eqxx. +have E : loc_qry (i1 \+ j1) = loc_qry j1 by rewrite (locProjR CD0 _ D1)// domPt inE andbC eqxx. rewrite E {E} -(rely_loc' _ R) in P3. case: (rely_coh R)=>_ D2. rewrite /W eqW in CD2; move: (coh_hooks CD2)=>{CD2}CD2. rewrite /mkWorld/= in C2. have C2': i2 \In Coh (plab pc \\-> pc, Unit). - split=>//=. - + by rewrite /valid/= valid_unit um_validPt. + + by rewrite /valid/= valid_unit validPt. + by apply: (cohS C2). + by apply: hook_complete0. - + by move=>z; rewrite -(cohD C2) !um_domPt. + + by move=>z; rewrite -(cohD C2) !domPt. move=>l; case B: (lc == l). - + move/eqP:B=>B; subst l; rewrite /getProtocol um_findPt; split=>//. - by move: (coh_coh lc C2); rewrite /getProtocol um_findPt. - have X: l \notin dom i2 by rewrite -(cohD C2) um_domPt inE; move/negbT: B. + + move/eqP:B=>B; subst l; rewrite /getProtocol findPt; split=>//. + by move: (coh_coh lc C2); rewrite /getProtocol findPt. + have X: l \notin dom i2 by rewrite -(cohD C2) domPt inE; move/negbT: B. rewrite /getProtocol/= (find_empty _ _ X). - have Y: l \notin dom (lc \\-> pc) by rewrite um_domPt inE; move/negbT: B. + have Y: l \notin dom (lc \\-> pc) by rewrite domPt inE; move/negbT: B. by case: dom_find Y=>//->_. have D2': j2 \In Coh (lq \\-> pq lq Data qnodes serialize, Unit) by apply: (cohUnKR CD2 _); try apply: hook_complete0. -rewrite -(locProjL CD2 _ C2') in L2; last by rewrite um_domPt inE eqxx. -rewrite -(locProjR CD2 _ D2') in P3; last by rewrite um_domPt inE eqxx. +rewrite -(locProjL CD2 _ C2') in L2; last by rewrite domPt inE eqxx. +rewrite -(locProjR CD2 _ D2') in P3; last by rewrite domPt inE eqxx. clear C2 D2. (* So what's important is for the precondition ofattachment to be *) (* independent of the core protocol. *) rewrite injWQ in R. rewrite /query_init_state/= in P4. -rewrite (locProjR CD0 _ D1) in P4; last by rewrite um_domPt inE eqxx. +rewrite (locProjR CD0 _ D1) in P4; last by rewrite domPt inE eqxx. have Q4: qry_init to j2. - by apply: (query_init_rely' lq Data qnodes serialize cn to _ _ P4 R). clear P4. rewrite /query_init_state/= -(locProjR CD2 _ D2') in Q4; - last by rewrite um_domPt inE eqxx. + last by rewrite domPt inE eqxx. (* Now ready to use the spec for querying. *) apply (gh_ex (g:=(rq, rs, (size ds, seq.zip chs ds)))). @@ -230,7 +230,7 @@ apply: call_rule=>//=; last by move=>d m[->->T1 T2->]_; eexists _. move=>CD2'; split=>//. case/orP: P2=>[|P]; first by move/eqP=>Z; subst to; rewrite /core_state_to_data eqxx. rewrite !(locProjL CD2 _ C2') in L2 *; - last by rewrite um_domPt inE eqxx. + last by rewrite domPt inE eqxx. move: (coh_coh lc C2'); rewrite prEq; case=>C3 _. rewrite /core_state_to_data; case:ifP=>//[|_]; first by move=>/eqP Z; subst to. by apply: (@cn_agree lc cn pts [::] Hnin _ _ _ to C3 _ Inv). diff --git a/Examples/Querying/QueryProtocol.v b/Examples/Querying/QueryProtocol.v index 0a1410d..0877e78 100644 --- a/Examples/Querying/QueryProtocol.v +++ b/Examples/Querying/QueryProtocol.v @@ -3,11 +3,11 @@ Require Import ssreflect ssrbool ssrnat eqtype ssrfun seq. From mathcomp Require Import path. Require Import Eqdep. -From DiSeL.Heaps -Require Import pred prelude idynamic ordtype finmap pcm unionmap heap coding domain. -From DiSeL.Core +From fcsl +Require Import axioms pred prelude ordtype finmap pcm unionmap heap. +From DiSeL Require Import Freshness State EqTypeX Protocols Worlds NetworkSem Rely Actions. -From DiSeL.Examples +From DiSeL Require Import SeqLib. Set Implicit Arguments. @@ -113,7 +113,7 @@ Lemma send_soupCoh d m : Proof. move=>[H1 H2]Cm; split=>[|i ms/=]; first by rewrite valid_fresh. rewrite findUnL; last by rewrite valid_fresh. -case: ifP=>E; [by move/H2|by move/um_findPt_inv=>[Z G]; subst i m]. +case: ifP=>E; [by move/H2|by move/findPt_inv=>[Z G]; subst i m]. Qed. Lemma consume_coh d m : QCoh d -> soupCoh (consume_msg (dsoup d) m). @@ -141,23 +141,23 @@ Qed. (************** TODO: Refactor this!!! **************) Lemma cohSt n d (C : QCoh d) s: find st (getLocal n d) = Some s -> - idyn_tp s = qstate. + dyn_tp s = qstate. Proof. case: (C)=>_ _ D G; case H: (n \in nodes); rewrite -D in H. -- by move:(G _ H); case=>s'[]->_; rewrite hfindPt//=; case=><-. +- by move:(G _ H); case=>s'[]->_; rewrite findPt//=; case=><-. by rewrite /getLocal; case: dom_find H=>//->; rewrite find0E. Qed. Definition getSt n d (C : QCoh d) : qstate := match find st (getLocal n d) as f return _ = f -> _ with - Some v => fun epf => icoerce id (idyn_val v) (cohSt C epf) + Some v => fun epf => icast (sym_eq (cohSt C epf)) (dyn_val v) | _ => fun epf => ([::], [::]) end (erefl _). Lemma getStK n d (C : QCoh d) s : getLocal n d = st :-> s -> getSt n C = s. Proof. -by move=>E; rewrite /getSt/=; move: (cohSt C); rewrite !E/==>H; apply: ieqc. +by move=>E; rewrite /getSt/=; move: (cohSt C); rewrite !E/==>H; apply: eqc. Qed. Lemma getStE n i j C C' (pf : n \in nodes) : diff --git a/Examples/TwoPhaseCommit/SimpleTPCApp.v b/Examples/TwoPhaseCommit/SimpleTPCApp.v index 13476c0..64993c7 100644 --- a/Examples/TwoPhaseCommit/SimpleTPCApp.v +++ b/Examples/TwoPhaseCommit/SimpleTPCApp.v @@ -4,15 +4,15 @@ From mathcomp Require Import path. Require Import Eqdep. Require Import Relation_Operators. -From DiSeL.Heaps -Require Import pred prelude ordtype finmap pcm unionmap heap coding domain. -From DiSeL.Core +From fcsl +Require Import pred prelude ordtype finmap pcm unionmap heap. +From DiSeL Require Import State Protocols Worlds NetworkSem Rely. -From DiSeL.Core +From DiSeL Require Import HoareTriples InferenceRules While. -From DiSeL.Examples +From DiSeL Require Import TwoPhaseProtocol TwoPhaseCoordinator TwoPhaseParticipant. -From DiSeL.Examples +From DiSeL Require TwoPhaseInductiveProof. Section SimpleTpcApp. @@ -69,11 +69,11 @@ Definition init_dstate := Lemma valid_init_dstate : valid init_dstate. Proof. case: validUn=>//=; -do?[case: validUn=>//; do?[rewrite ?gen_validPt/=//]|by rewrite gen_validPt/=]. -- by move=>k; rewrite !gen_domPt !inE/==>/eqP<-/eqP. -- by move=>k; rewrite domUn!inE/==>/andP[_]/orP[]; rewrite !gen_domPt!inE/==>/eqP=><-. -move=>k; rewrite domUn!inE/==>/andP[_]/orP[]; last by rewrite !gen_domPt!inE/==>/eqP=><-. -by rewrite domUn!inE/==>/andP[_]/orP[]; rewrite !gen_domPt!inE/==>/eqP=><-. +do?[case: validUn=>//; do?[rewrite ?validPt/=//]|by rewrite validPt/=]. +- by move=>k; rewrite !domPt !inE/==>/eqP<-/eqP. +- by move=>k; rewrite domUn!inE/==>/andP[_]/orP[]; rewrite !domPt!inE/==>/eqP=><-. +move=>k; rewrite domUn!inE/==>/andP[_]/orP[]; last by rewrite !domPt!inE/==>/eqP=><-. +by rewrite domUn!inE/==>/andP[_]/orP[]; rewrite !domPt!inE/==>/eqP=><-. Qed. Notation init_dstatelet := (DStatelet init_dstate Unit). @@ -83,7 +83,7 @@ Definition init_state : state := l \\-> init_dstatelet. Lemma getCnLoc : getLocal cn init_dstatelet = init_heap_c. Proof. -rewrite /getLocal/init_dstate -!joinA um_findPtUn//. +rewrite /getLocal/init_dstate -!joinA findPtUn//. by rewrite !joinA valid_init_dstate. Qed. @@ -91,9 +91,9 @@ Lemma getPLoc p : p \in pts -> getLocal p init_dstatelet = init_heap_p. Proof. rewrite /pts !inE=>/orP[];[|move=>/orP[]]=>/eqP->{p}; move: valid_init_dstate; rewrite /getLocal/init_dstate. -- by rewrite -!joinA joinCA=>V; rewrite um_findPtUn//. -- by rewrite -joinA joinCA=>V; rewrite um_findPtUn//. -by rewrite joinC=>V; rewrite um_findPtUn//. +- by rewrite -!joinA joinCA=>V; rewrite findPtUn//. +- by rewrite -joinA joinCA=>V; rewrite findPtUn//. +by rewrite joinC=>V; rewrite findPtUn//. Qed. @@ -101,32 +101,33 @@ Qed. Notation W := (mkWorld (TwoPhaseCoordinator.tpc l cn pts others Hnin)). Lemma hook_complete_unit (c : context) : hook_complete (c, Unit). -Proof. by move=>????; rewrite dom0 inE. Qed. +Proof. by move=>????; rewrite dom0. Qed. Lemma hooks_consistent_unit (c : context) : hooks_consistent c Unit. -Proof. by move=>????; rewrite dom0 inE. Qed. +Proof. by move=>????; rewrite dom0. Qed. Lemma init_coh : init_state \In Coh W. Proof. split. - apply/andP; split; last by rewrite valid_unit. - by rewrite gen_validPt. -- by rewrite/init_state gen_validPt. + by rewrite validPt. +- by rewrite/init_state validPt. - by apply: hook_complete_unit. -- by move=>z; rewrite /init_state !um_domPt inE/=. +- by move=>z; rewrite /init_state !domPt inE/=. move=>k; case B: (l==k); last first. - have X: (k \notin dom init_state) /\ (k \notin dom W.1). - by rewrite /init_state/W/=!gen_domPt !inE/=; move/negbT: B. + by rewrite /init_state/W/=!domPt !inE/=; move/negbT: B. rewrite /getProtocol /getStatelet/=. case: dom_find X=>//; last by move=>? _ _[]. by move=>->[_]; case: dom_find=>//->. -move/eqP:B=>B; subst k; rewrite prEq/getStatelet/init_state gen_findPt/=. +move/eqP:B=>B; subst k; rewrite prEq/getStatelet/init_state findPt/=. split=>//=; do?[by apply: valid_init_dstate]; first by split=>//m ms; rewrite find0E. - move=>z; rewrite /init_dstate/TPCProtocol.nodes/=/others. rewrite !domUn !inE valid_init_dstate/=. rewrite !domUn !inE (validL valid_init_dstate)/=. rewrite !domUn !inE (validL (validL valid_init_dstate))/=. - rewrite !gen_domPt!inE/= !(eq_sym z). + rewrite !domPt!inE/= !(eq_sym z). + rewrite 3!inE. by case:(cn==z)(p1==z)(p2==z)(p3==z);case;case;case. move=>z/=; rewrite !inE=>/orP []. - move/eqP=>Z; subst z; rewrite getCnLoc; split=>//=. @@ -154,7 +155,7 @@ Next Obligation. move=>i/=R. apply: with_inv_rule'. apply:call_rule=>//. -- by rewrite (rely_loc' _ R) /getStatelet gen_findPt/=getCnLoc. +- by rewrite (rely_loc' _ R) /getStatelet findPt/=getCnLoc. move=>_ m [chs] CS C I _. exists chs. split=>//. @@ -178,7 +179,7 @@ Next Obligation. move=>i/=R. apply: with_inv_rule'. apply:call_rule=>//. -- by rewrite (rely_loc' _ R)/getStatelet gen_findPt/= (getPLoc _ pf). +- by rewrite (rely_loc' _ R)/getStatelet findPt/= (getPLoc _ pf). move=>_ m [bs][ds] PS C I _. exists (seq.zip bs ds), (size choices). move /(coh_coh l) in C. diff --git a/Examples/TwoPhaseCommit/TwoPhaseCoordinator.v b/Examples/TwoPhaseCommit/TwoPhaseCoordinator.v index 7a46640..6025bfe 100644 --- a/Examples/TwoPhaseCommit/TwoPhaseCoordinator.v +++ b/Examples/TwoPhaseCommit/TwoPhaseCoordinator.v @@ -4,15 +4,15 @@ From mathcomp Require Import path. Require Import Eqdep. Require Import Relation_Operators. -From DiSeL.Heaps -Require Import pred prelude idynamic ordtype finmap pcm unionmap heap coding domain. -From DiSeL.Core +From fcsl +Require Import axioms pred prelude ordtype finmap pcm unionmap heap. +From DiSeL Require Import Freshness State EqTypeX DepMaps Protocols Worlds NetworkSem Rely. -From DiSeL.Core +From DiSeL Require Import Actions Injection Process Always HoareTriples InferenceRules. -From DiSeL.Core +From DiSeL Require Import InductiveInv While. -From DiSeL.Examples +From DiSeL Require Import TwoPhaseProtocol. Module TwoPhaseCoordinator. @@ -53,15 +53,15 @@ Program Definition tryrecv_prep_resp := act (@tryrecv_action_wrapper W cn (* filter *) (fun k _ t b => (k == l) && ((t == prep_yes) || (t == prep_no))) _). (* TODO: automate these kinds of proofs *) -Next Obligation. by case/andP: H=>/eqP->_; rewrite /ddom gen_domPt inE/=. Qed. +Next Obligation. by case/andP: H=>/eqP->_; rewrite /ddom domPt inE/=. Qed. Program Definition tryrecv_commit_ack := act (@tryrecv_action_wrapper W cn (fun k _ t b => (k == l) && (t == commit_ack)) _). -Next Obligation. by case/andP: H=>/eqP->_; rewrite /ddom gen_domPt inE/=. Qed. +Next Obligation. by case/andP: H=>/eqP->_; rewrite /ddom domPt inE/=. Qed. Program Definition tryrecv_abort_ack := act (@tryrecv_action_wrapper W cn (fun k _ t b => (k == l) && (t == abort_ack)) _). -Next Obligation. by case/andP: H=>/eqP->_; rewrite /ddom gen_domPt inE/=. Qed. +Next Obligation. by case/andP: H=>/eqP->_; rewrite /ddom domPt inE/=. Qed. (************** Coordinator code **************) @@ -91,9 +91,11 @@ apply: act_rule=>j R; split=>[|r k m]; first by case: (rely_coh R). case=>/=H1[Cj]Z; subst j=>->R'. split; first by rewrite (rely_loc' l R') (rely_loc' _ R). case: (rely_coh R')=>_; case=>_ _ _ _/(_ l)=>/= pf; rewrite prEq in pf. -exists pf; move: (rely_loc' l R') =>/sym E'. +exists pf; move: (rely_loc' l R'). +move => E'. +apply sym_eq in E'. suff X: getStC (Actions.safe_local (prEq tpc) H1) = getStC pf by rewrite X. -by apply: (getStCE pf _ E'). +by apply: (getStCE pf _ E'). Qed. (*******************************************) @@ -143,8 +145,12 @@ case=>[[E1 P1 C1]|]. left; exists e; split; last by exists d. by rewrite -(rely_loc' _ R1) in E1; rewrite (getStC_K _ E1). + rewrite /Actions.can_send /nodes inE eqxx andbC/=. - by rewrite -(cohD (proj2 (rely_coh R1)))/ddom gen_domPt inE/=. - + by rewrite /Actions.filter_hooks um_filt0=>???/sym/find_some; rewrite dom0 inE. + by rewrite -(cohD (proj2 (rely_coh R1)))/ddom domPt inE/=. + + rewrite /Actions.filter_hooks umfilt0=>???. + move => F. + apply sym_eq in F. + move: F. + by move/find_some; rewrite dom0. case: {-1}(Sf)=>_/=[]Hc[C][]; last first. - move=>[n][d'][ps][E1'][]Z1 Z2 _; subst n d'. rewrite -(rely_loc' _ R1) in E1. @@ -163,26 +169,30 @@ have Pre: have Y: pts == [:: to] by rewrite (perm_eq_small _ Hp). rewrite /cstep_send/= Y in G. move: (proj2 Hc)=>Y'; rewrite Y' in G=>{Y'}. - rewrite [cn_safe_coh _ ](proof_irrelevance _ C) E1' in G. (* TADA! *) + rewrite [cn_safe_coh _ ](pf_irr _ C) E1' in G. (* TADA! *) rewrite (rely_loc' l R2); subst k; rewrite -(rely_loc' _ R1) in E1. rewrite locE; last apply: (cohVl C). - + by rewrite -(proof_irrelevance (cn_in cn pts others) (cn_this_in _ _)) + + by rewrite -(pf_irr (cn_in cn pts others) (cn_this_in _ _)) (getStL_Kc _ (cn_in cn pts others) E1). - + by rewrite -(cohD (proj2 (rely_coh R1)))/ddom gen_domPt inE/=. + + by rewrite -(cohD (proj2 (rely_coh R1)))/ddom domPt inE/=. by apply: (cohS (proj2 (rely_coh R1))). - exists [:: to]; split; last by rewrite cat_cons/=. have Y: pts == [:: to] = false. + apply/negP=>/eqP Z; rewrite Z in Hp. - by move/perm_eq_size: (Hp); case=>/sym/size0nil=>Z'; rewrite Z' eqxx in X. + move/perm_eq_size: (Hp). + move => F. + apply sym_eq in F. + move: F. + by case=>/size0nil=>Z'; rewrite Z' eqxx in X. (* This part is similar to the previous step *) rewrite /cstep_send/= Y in G. move: (proj2 Hc)=>Y'; rewrite Y' in G=>{Y'}. - rewrite [cn_safe_coh _ ](proof_irrelevance _ C) E1' in G. (* TADA! *) + rewrite [cn_safe_coh _ ](pf_irr _ C) E1' in G. (* TADA! *) rewrite (rely_loc' l R2); subst k; rewrite -(rely_loc' _ R1) in E1. rewrite locE; last apply: (cohVl C). - + by rewrite -(proof_irrelevance (cn_in cn pts others) (cn_this_in _ _)) + + by rewrite -(pf_irr (cn_in cn pts others) (cn_this_in _ _)) (getStL_Kc _ (cn_in cn pts others) E1). - + by rewrite -(cohD (proj2 (rely_coh R1)))/ddom gen_domPt inE/=. + + by rewrite -(cohD (proj2 (rely_coh R1)))/ddom domPt inE/=. by apply: (cohS (proj2 (rely_coh R1))). apply: call_rule'=>/=[Cm|r2 m2]; first by right. @@ -210,12 +220,19 @@ case: Y=>to[tos] Z; subst to_send=>{X}. case: (proj2 (rely_coh R1))=>_ _ _ _/(_ l); rewrite prEq=>C; exists C. right; exists e, d, ps; split=>//. by rewrite -(rely_loc' _ R1) in E1; rewrite (getStC_K _ E1). - + move/perm_eq_uniq: Hp; rewrite Puniq sym. + + move/perm_eq_uniq: Hp; rewrite Puniq. + move => F. + apply sym_eq in F. + move: F. rewrite -cat_rcons cat_uniq -cats1 cat_uniq=>/andP[]/andP[_]/andP[]. by rewrite /= orbC/=. + rewrite /Actions.can_send /nodes inE eqxx andbC. - by rewrite -(cohD (proj2 (rely_coh R1)))/ddom gen_domPt inE/=. - by rewrite /Actions.filter_hooks um_filt0=>???/sym/find_some; rewrite dom0 inE. + by rewrite -(cohD (proj2 (rely_coh R1)))/ddom domPt inE/=. + rewrite /Actions.filter_hooks umfilt0=>???. + move => F. + apply sym_eq in F. + move: F. + by move/find_some; rewrite dom0. (* dismiss the bogus branch *) case: {-1}(Sf)=>_/=[]Hc[C][]. - case=>b[E1'][d'][Z1 Z2]_; subst b d'. @@ -238,32 +255,36 @@ suff Pre: - case X: ([::] == tos); [move/eqP: X=>X; subst tos; rewrite eqxx| rewrite eq_sym X]. rewrite /cstep_send/= (proj2 Hc)/= in G. - rewrite [cn_safe_coh _ ](proof_irrelevance _ C) E1' in G. + rewrite [cn_safe_coh _ ](pf_irr _ C) E1' in G. have Y: perm_eq (to :: ps) pts. rewrite (perm_eq_sym pts) in Hp. by apply/perm_eqlE; rewrite -cat1s perm_catC; apply/perm_eqlP. rewrite Y/= in G. rewrite (rely_loc' l R2); subst k; rewrite locE; last apply: (cohVl C). - + by rewrite -(proof_irrelevance (cn_in cn pts others) (cn_this_in _ _)) + + by rewrite -(pf_irr (cn_in cn pts others) (cn_this_in _ _)) (getStL_Kc _ (cn_in cn pts others) E1). - + by rewrite -(cohD (proj2 (rely_coh R1)))/ddom gen_domPt inE/=. + + by rewrite -(cohD (proj2 (rely_coh R1)))/ddom domPt inE/=. by apply: (cohS (proj2 (rely_coh R1))). rewrite /cstep_send/= (proj2 Hc)/= in G. -rewrite [cn_safe_coh _ ](proof_irrelevance _ C) E1' in G. +rewrite [cn_safe_coh _ ](pf_irr _ C) E1' in G. have Y : perm_eq (to :: ps) pts = false. - apply/negP=>Hp'; move: (perm_eq_trans Hp' Hp). rewrite -[to::ps]cat1s -[to::tos]cat1s. move/perm_eqlE: (perm_catC ps [::to])=>Hs. move/(perm_eq_trans Hs); rewrite -[_++[::_]]cats0 catA perm_cat2l. - by move/perm_eq_size/sym/size0nil=>Z; subst tos. + move/perm_eq_size. + move => F. + apply sym_eq in F. + move: F. + by move/size0nil=>Z; subst tos. rewrite Y/= in G. rewrite (rely_loc' l R2); subst k; rewrite locE; last apply: (cohVl C). - + rewrite -(proof_irrelevance (cn_in cn pts others) (cn_this_in _ _)) + + rewrite -(pf_irr (cn_in cn pts others) (cn_this_in _ _)) (getStL_Kc _ (cn_in cn pts others) E1); exists (to::ps). split=>//; move: Hp. by rewrite -cat_rcons -cat1s -!catA !(perm_eq_sym pts) -perm_catCA catA cats1. - + by rewrite -(cohD (proj2 (rely_coh R1)))/ddom gen_domPt inE/=. + + by rewrite -(cohD (proj2 (rely_coh R1)))/ddom domPt inE/=. by apply: (cohS (proj2 (rely_coh R1))). Qed. @@ -328,7 +349,7 @@ have D: rt = cn_receive_prep_yes_trans _ _ _ \/ rt = cn_receive_prep_no_trans _ have P1: valid (dstate (getS j)) by apply: (@cohVl _ TPCCoh); case: (Cj')=>P1 P2 P3 P4; split=>//=; done. have P2: valid j by apply: (cohS (proj2 (rely_coh R1))). -have P3: l \in dom j by rewrite -(cohD (proj2 (rely_coh R1)))/ddom gen_domPt inE/=. +have P3: l \in dom j by rewrite -(cohD (proj2 (rely_coh R1)))/ddom domPt inE/=. (* Two cases: received yes or no from a participant *) case: D=>{Hin G Hw}Z; subst rt; rewrite/= in E1 R; case:ifP=>G1; apply: ret_rule=>m' R'=>{d lg I1}[[d lg][I1]]; rewrite /rc_prep_inv ?E1 ?eqxx/=; @@ -400,9 +421,9 @@ apply: step; apply: (gh_ex (g:=e)); apply: (gh_ex (g:=d)); apply: (gh_ex (g:=lg)); apply: (gh_ex (g:=res)). apply: call_rule=>// b s4 [E4]->{b}C4/=. apply: ret_rule=>i5 R5 lg'[e'] E0'; exists e. -rewrite E0 in E0'; case: (hcancelV _ E0'); first by rewrite hvalidPtUn. +rewrite E0 in E0'; case: (hcancelV _ E0'); first by rewrite validPtUn. case=>Z1 _; subst e'; move/(hcancelPtV _)=>/=. -by rewrite hvalidPt (rely_loc' _ R5)=>/(_ is_true_true)=><-. +by rewrite validPt (rely_loc' _ R5)=>/(_ is_true_true)=><-. Qed. (*******************************************) @@ -442,9 +463,13 @@ apply: step; apply: act_rule=>s2 R2/=. have Pre: Actions.send_act_safe W (p:=tpc) cn l (cn_send_commit_trans cn pts others) [:: e] to s2. - split; [by case: (rely_coh R2) | | |]; last first. - + by rewrite /Actions.filter_hooks um_filt0=>???/sym/find_some; rewrite dom0 inE. + + rewrite /Actions.filter_hooks umfilt0=>???. + move => F. + apply sym_eq in F. + move: F. + by move/find_some; rewrite dom0. + rewrite /Actions.can_send /nodes inE eqxx andbC/=. - by rewrite -(cohD (proj2 (rely_coh R2)))/ddom gen_domPt inE/=. + by rewrite -(cohD (proj2 (rely_coh R2)))/ddom domPt inE/=. case: (proj2 (rely_coh R2))=>_ _ _ _/(_ l); rewrite prEq=>C; split. + split=>//; case: H; first by case=>?[_]<-; rewrite inE eqxx. by case=>ps[_]/perm_eq_mem->; rewrite mem_cat orbC inE eqxx. @@ -452,7 +477,10 @@ have Pre: Actions.send_act_safe W (p:=tpc) cn l rewrite -(rely_loc' _ R2) in P1; rewrite (getStC_K _ P1); first by exists e, d, res=>//. exists e, d, ps; split=>//. - move/perm_eq_uniq: P2; rewrite Puniq sym. + move/perm_eq_uniq: P2; rewrite Puniq. + move => F. + apply sym_eq in F. + move: F. rewrite -cat_rcons cat_uniq -cats1 cat_uniq=>/andP[]/andP[_]/andP[]. by rewrite /= orbC. @@ -488,7 +516,11 @@ have X: perm_eq (to :: ps) pts = (tos == [::]). rewrite -[to::ps]cat1s -[to::tos]cat1s. move/perm_eqlE: (perm_catC ps [::to])=>Hs. move/(perm_eq_trans Hs); rewrite -[_++[::_]]cats0 catA perm_cat2l. - by move/perm_eq_size/sym/size0nil=>Z; subst tos. + move/perm_eq_size. + move => F. + apply sym_eq in F. + move: F. + by move/size0nil=>Z; subst tos. move/eqP=>Z; subst tos; rewrite perm_eq_sym; apply: (perm_eq_trans P2). by apply/perm_eqlE; move: (perm_catC ps [:: to]). rewrite X in G=>{X}; subst i3. @@ -544,9 +576,13 @@ apply: step; apply: act_rule=>s2 R2/=. have Pre: Actions.send_act_safe W (p:=tpc) cn l (cn_send_abort_trans cn pts others) [:: e] to s2. - split; [by case: (rely_coh R2) | | |]; last first. - + by rewrite /Actions.filter_hooks um_filt0=>???/sym/find_some; rewrite dom0 inE. + + rewrite /Actions.filter_hooks umfilt0=>???. + move => F. + apply sym_eq in F. + move: F. + by move/find_some; rewrite dom0. + rewrite /Actions.can_send /nodes inE eqxx andbC/=. - by rewrite -(cohD (proj2 (rely_coh R2)))/ddom gen_domPt inE/=. + by rewrite -(cohD (proj2 (rely_coh R2)))/ddom domPt inE/=. case: (proj2 (rely_coh R2))=>_ _ _ _/(_ l); rewrite prEq=>C; split. + split=>//; case: H; first by case=>?[_]<-; rewrite inE eqxx. by case=>ps[_]/perm_eq_mem->; rewrite mem_cat orbC inE eqxx. @@ -554,7 +590,10 @@ have Pre: Actions.send_act_safe W (p:=tpc) cn l rewrite -(rely_loc' _ R2) in P1; rewrite (getStC_K _ P1); first by exists e, d, res=>//. exists e, d, ps; split=>//. - move/perm_eq_uniq: P2; rewrite Puniq sym. + move/perm_eq_uniq: P2; rewrite Puniq. + move => F. + apply sym_eq in F. + move: F. rewrite -cat_rcons cat_uniq -cats1 cat_uniq=>/andP[]/andP[_]/andP[]. by rewrite /= orbC. @@ -592,7 +631,11 @@ have X: perm_eq (to :: ps) pts = (tos == [::]). rewrite -[to::ps]cat1s -[to::tos]cat1s. move/perm_eqlE: (perm_catC ps [::to])=>Hs. move/(perm_eq_trans Hs); rewrite -[_++[::_]]cats0 catA perm_cat2l. - by move/perm_eq_size/sym/size0nil=>Z; subst tos. + move/perm_eq_size. + move => F. + apply sym_eq in F. + move: F. + by move/size0nil=>Z; subst tos. move/eqP=>Z; subst tos; rewrite perm_eq_sym; apply: (perm_eq_trans P2). by apply/perm_eqlE; move: (perm_catC ps [:: to]). rewrite X in G=>{X}; subst i3. @@ -672,7 +715,7 @@ have D: rt = cn_receive_commit_ack_trans _ _ _. have P1: valid (dstate (getS j)) by apply: (@cohVl _ TPCCoh); case: (Cj')=>P1 P2 P3 P4; split=>//=; done. have P2: valid j by apply: (cohS (proj2 (rely_coh R1))). -have P3: l \in dom j by rewrite -(cohD (proj2 (rely_coh R1)))/ddom gen_domPt inE/=. +have P3: l \in dom j by rewrite -(cohD (proj2 (rely_coh R1)))/ddom domPt inE/=. (* The final blow *) by subst rt; rewrite/= in E1 R; case:ifP=>G1; apply: ret_rule=>m' R'; rewrite /rc_commit_cond=>{d lg I1}[[d lg][I1]]; @@ -690,7 +733,12 @@ apply: ghC=>i[d lg]E1 C1. have Pre: rc_commit_inv e (d, lg) [::] i. - rewrite /rc_commit_inv/= E1/=. have X: perm_eq [::] pts = false. - - apply/negP. move/perm_eq_size/sym/size0nil=>Z. + - apply/negP. + move/perm_eq_size. + move => F. + apply sym_eq in F. + move: F. + move/size0nil=>Z. by move: (PtsNonEmpty); rewrite Z. by rewrite X. apply: call_rule'=>[|acc m]; first by exists (d, lg). @@ -750,7 +798,7 @@ have D: rt = cn_receive_abort_ack_trans _ _ _. have P1: valid (dstate (getS j)) by apply: (@cohVl _ TPCCoh); case: (Cj')=>P1 P2 P3 P4; split=>//=; done. have P2: valid j by apply: (cohS (proj2 (rely_coh R1))). -have P3: l \in dom j by rewrite -(cohD (proj2 (rely_coh R1)))/ddom gen_domPt inE/=. +have P3: l \in dom j by rewrite -(cohD (proj2 (rely_coh R1)))/ddom domPt inE/=. (* The final blow *) by subst rt; rewrite/= in E1 R; case:ifP=>G1; apply: ret_rule=>m' R'; rewrite /rc_abort_cond=>{d lg I1}[[d lg][I1]]; @@ -768,7 +816,12 @@ apply: ghC=>i[d lg]E1 C1. have Pre: rc_abort_inv e (d, lg) [::] i. - rewrite /rc_abort_inv/= E1/=. have X: perm_eq [::] pts = false. - - apply/negP. move/perm_eq_size/sym/size0nil=>Z. + - apply/negP. + move/perm_eq_size. + move => F. + apply sym_eq in F. + move: F. + move/size0nil=>Z. by move: (PtsNonEmpty); rewrite Z. by rewrite X. apply: call_rule'=>[|acc m]; first by exists (d, lg). @@ -813,17 +866,17 @@ case:ifP=>A. move=>s5 E5 C5; apply: (gh_ex (g:=(d, lg))). apply: call_rule=>//_ s6 E6 C6. apply: ret_rule=>i6 R6 e' lg' E0'. - rewrite E0 in E0'; case: (hcancelV _ E0'); first by rewrite hvalidPtUn. + rewrite E0 in E0'; case: (hcancelV _ E0'); first by rewrite validPtUn. + case=>Z1 _; subst e'; move/(hcancelPtV _)=>/=. - by rewrite hvalidPt (rely_loc' _ R6)=>/(_ is_true_true)=><-. + by rewrite validPt (rely_loc' _ R6)=>/(_ is_true_true)=><-. do![apply: step]; apply: (gh_ex (g:=lg)). apply: call_rule=>_; first by exists res; rewrite has_predC A. move=>s5 E5 C5; apply: (gh_ex (g:=(d, lg))). apply: call_rule=>//_ s6 E6 C6. apply: ret_rule=>i6 R6 e' lg' E0'. -rewrite E0 in E0'; case: (hcancelV _ E0'); first by rewrite hvalidPtUn. +rewrite E0 in E0'; case: (hcancelV _ E0'); first by rewrite validPtUn. - case=>Z1 _; subst e'; move/(hcancelPtV _)=>/=. -by rewrite hvalidPt (rely_loc' _ R6)=>/(_ is_true_true)=><-. +by rewrite validPt (rely_loc' _ R6)=>/(_ is_true_true)=><-. Qed. (**************************************************) diff --git a/Examples/TwoPhaseCommit/TwoPhaseExtraction.v b/Examples/TwoPhaseCommit/TwoPhaseExtraction.v index 2867484..97a572c 100644 --- a/Examples/TwoPhaseCommit/TwoPhaseExtraction.v +++ b/Examples/TwoPhaseCommit/TwoPhaseExtraction.v @@ -1,6 +1,6 @@ -From DiSeL.Core +From DiSeL Require Import DiSeLExtraction. -From DiSeL.Examples +From DiSeL Require Import SimpleTPCApp. Cd "extraction". diff --git a/Examples/TwoPhaseCommit/TwoPhaseInductiveInv.v b/Examples/TwoPhaseCommit/TwoPhaseInductiveInv.v index d934a7c..98ed2ec 100644 --- a/Examples/TwoPhaseCommit/TwoPhaseInductiveInv.v +++ b/Examples/TwoPhaseCommit/TwoPhaseInductiveInv.v @@ -4,17 +4,15 @@ From mathcomp Require Import path. Require Import Eqdep. Require Import Relation_Operators. -From DiSeL.Heaps -Require Import pred prelude idynamic ordtype finmap pcm unionmap. -From DiSeL.Heaps -Require Import heap coding domain. -From DiSeL.Core +From fcsl +Require Import axioms pred prelude ordtype finmap pcm unionmap heap. +From DiSeL Require Import Freshness State EqTypeX DepMaps Protocols Worlds NetworkSem Rely. -From DiSeL.Core +From DiSeL Require Import Actions Injection Process Always HoareTriples InferenceRules. -From DiSeL.Core +From DiSeL Require Import InductiveInv While StatePredicates. -From DiSeL.Examples +From DiSeL Require Import TwoPhaseProtocol. Require Import Omega. @@ -206,7 +204,7 @@ Definition Inv (d : dstatelet) := (***********************************************************) Notation coh d := (coh tpc d). -Notation PI := proof_irrelevance. +Notation PI := pf_irr. Export TPCProtocol. Lemma inv_init d e (C : coh d): diff --git a/Examples/TwoPhaseCommit/TwoPhaseInductiveProof.v b/Examples/TwoPhaseCommit/TwoPhaseInductiveProof.v index 341152e..9a3fad5 100644 --- a/Examples/TwoPhaseCommit/TwoPhaseInductiveProof.v +++ b/Examples/TwoPhaseCommit/TwoPhaseInductiveProof.v @@ -4,17 +4,15 @@ From mathcomp Require Import path. Require Import Eqdep. Require Import Relation_Operators. -From DiSeL.Heaps -Require Import pred prelude idynamic ordtype finmap pcm unionmap. -From DiSeL.Heaps -Require Import heap coding domain. -From DiSeL.Core +From fcsl +Require Import axioms pred prelude ordtype finmap pcm unionmap heap. +From DiSeL Require Import Freshness State EqTypeX DepMaps Protocols Worlds NetworkSem Rely. -From DiSeL.Core +From DiSeL Require Import Actions Injection Process Always HoareTriples InferenceRules. -From DiSeL.Core +From DiSeL Require Import InductiveInv While StatePredicates. -From DiSeL.Examples +From DiSeL Require Import TwoPhaseProtocol TwoPhaseInductiveInv. Set Implicit Arguments. @@ -39,7 +37,7 @@ Notation loc z d := (getLocal z d). Notation Sinv := (@S_inv tpc (fun d _ => Inv cn pts d)). Notation Rinv := (@R_inv tpc (fun d _ => Inv cn pts d)). Notation coh d := (coh tpc d). -Notation PI := proof_irrelevance. +Notation PI := pf_irr. Export TPCProtocol. diff --git a/Examples/TwoPhaseCommit/TwoPhaseParticipant.v b/Examples/TwoPhaseCommit/TwoPhaseParticipant.v index c585cdd..1a3d595 100644 --- a/Examples/TwoPhaseCommit/TwoPhaseParticipant.v +++ b/Examples/TwoPhaseCommit/TwoPhaseParticipant.v @@ -4,15 +4,15 @@ From mathcomp Require Import path. Require Import Eqdep. Require Import Relation_Operators. -From DiSeL.Heaps -Require Import pred prelude idynamic ordtype finmap pcm unionmap heap coding domain. -From DiSeL.Core +From fcsl +Require Import axioms pred prelude ordtype finmap pcm unionmap heap. +From DiSeL Require Import Freshness State EqTypeX DepMaps Protocols Worlds NetworkSem Rely. -From DiSeL.Core +From DiSeL Require Import Actions Injection Process Always HoareTriples InferenceRules. -From DiSeL.Core +From DiSeL Require Import InductiveInv While. -From DiSeL.Examples +From DiSeL Require Import TwoPhaseProtocol. Module TwoPhaseParticipant. @@ -39,7 +39,7 @@ Section ParticipantImplementation. Program Definition tryrecv_prep_req := act (@tryrecv_action_wrapper W p (fun k _ t b => (k == l) && (t == prep_req)) _). (* TODO: automate these kinds of proofs *) -Next Obligation. by case/andP: H=>/eqP->_; rewrite /ddom gen_domPt inE/=. Qed. +Next Obligation. by case/andP: H=>/eqP->_; rewrite /ddom domPt inE/=. Qed. (* @@ -62,7 +62,7 @@ Program Definition tryrecv_commabrt_req c := act (@tryrecv_action_wrapper W p (fun k _ t b => (k == l) && (if c then (t == commit_req) || (t == abort_req) else (t == abort_req))) _). -Next Obligation. by case/andP: H=>/eqP->_; rewrite /ddom gen_domPt inE/=. Qed. +Next Obligation. by case/andP: H=>/eqP->_; rewrite /ddom domPt inE/=. Qed. (* Four send-actions, e -- id of the current era *) @@ -107,7 +107,9 @@ apply: act_rule=>j R; split=>[|r k m]; first by case: (rely_coh R). case=>/=H1[Cj]Z; subst j=>->R'. split; first by rewrite (rely_loc' l R') (rely_loc' _ R). case: (rely_coh R')=>_; case=>_ _ _ _/(_ l)=>/= pf; rewrite prEq in pf. -exists pf; move: (rely_loc' l R') =>/sym E'. +exists pf; move: (rely_loc' l R'). +move => E'. +apply sym_eq in E'. suff X: getStP Hnin (Actions.safe_local (prEq tpc) H1) Pin = getStP Hnin pf Pin by rewrite X. by apply: (TPCProtocol.getStPE Hnin pf _ Pin Hpin E'). Qed. @@ -163,7 +165,7 @@ subst rt=>{G}. have P1: valid (dstate (getS i2)) by apply: (@cohVl _ coh); case: (Cj')=>P1 P2 P3 P4; split=>//=; done. have P2: valid i2 by apply: (cohS (proj2 (rely_coh R1))). -have P3: l \in dom i2 by rewrite -(cohD (proj2 (rely_coh R1)))/ddom gen_domPt inE/=. +have P3: l \in dom i2 by rewrite -(cohD (proj2 (rely_coh R1)))/ddom domPt inE/=. case: ifP=>G1; apply: ret_rule=>//i5 R4. - rewrite /rp_prep_req_inv; rewrite (rely_loc' _ R4) (rely_loc' _ R) locE//=. rewrite /TPCProtocol.rp_step Hpin/=. case/andP:G1=>/eqP=>Z1/eqP H2; subst from. @@ -208,8 +210,12 @@ have Pre: forall i2, network_rely W p i i2 -> case: doCommit; split=>//; exists X, C, e, d; by rewrite -(rely_loc' _ R1) in E1; rewrite (getStP_K _ _ _ _ E1). + rewrite /Actions.can_send /nodes inE/= mem_cat Hpin orbC. - by rewrite -(cohD (proj2 (rely_coh R1)))/ddom gen_domPt inE/= eqxx. - by rewrite /Actions.filter_hooks um_filt0=>???/sym/find_some; rewrite dom0 inE. + by rewrite -(cohD (proj2 (rely_coh R1)))/ddom domPt inE/= eqxx. + rewrite /Actions.filter_hooks umfilt0=>???. + move => F. + apply sym_eq in F. + move: F. + by move/find_some; rewrite dom0. case: doCommit Pre=>Pre; apply: act_rule=>i2 R1/=; move:(Pre i2 R1)=>{Pre}Pre. @@ -219,11 +225,11 @@ case: {-1}(Sf)=>_/=[]_[H][]C'//; rewrite -(rely_loc' _ R1) in E1. rewrite (getStP_K Hnin C' (pn_this_in _ H) Hpin E1)/=. case=>e'[d'][][]Z1 Z2 _ G; subst d' e'. move: St; rewrite/Actions.send_act_step/==>[][_][h][]. -rewrite /pn_step -!(proof_irrelevance C' (pn_safe_coh _)) - -!(proof_irrelevance (pn_this_in _ H) (pn_this_in _ _)). +rewrite /pn_step -!(pf_irr C' (pn_safe_coh _)) + -!(pf_irr (pn_this_in _ H) (pn_this_in _ _)). rewrite (getStP_K Hnin C' (pn_this_in _ H) Hpin E1)(getStL_Kp _ (pn_this_in _ H) E1)/=. case=>Z Z';subst h i3; rewrite (rely_loc' _ R3) locE//; last by apply: (cohVl C'). -+ by rewrite -(cohD (proj2 (rely_coh R1)))/ddom gen_domPt inE/=. ++ by rewrite -(cohD (proj2 (rely_coh R1)))/ddom domPt inE/=. by apply: (cohS (proj2 (rely_coh R1))). (* Send no-answer *) @@ -233,11 +239,11 @@ case: {-1}(Sf)=>_/=[]_[H][]C'//; rewrite -(rely_loc' _ R1) in E1. rewrite (getStP_K Hnin C' (pn_this_in _ H) Hpin E1)/=. case=>e'[d'][][]Z1 Z2 _ G; subst d' e'. move: St; rewrite/Actions.send_act_step/==>[][_][h][]. -rewrite /pn_step -!(proof_irrelevance C' (pn_safe_coh _)) - -!(proof_irrelevance (pn_this_in _ H) (pn_this_in _ _)). +rewrite /pn_step -!(pf_irr C' (pn_safe_coh _)) + -!(pf_irr (pn_this_in _ H) (pn_this_in _ _)). rewrite (getStP_K Hnin C' (pn_this_in _ H) Hpin E1)(getStL_Kp _ (pn_this_in _ H) E1)/=. case=>Z Z';subst h i3; rewrite (rely_loc' _ R3) locE//; last by apply: (cohVl C'). -+ by rewrite -(cohD (proj2 (rely_coh R1)))/ddom gen_domPt inE/=. ++ by rewrite -(cohD (proj2 (rely_coh R1)))/ddom domPt inE/=. by apply: (cohS (proj2 (rely_coh R1))). Qed. @@ -294,7 +300,7 @@ move=>rt pf Cj' Hin R E2 Hw G E. have P1: valid (dstate (getS i2)) by apply: (@cohVl _ coh); case: (Cj')=>P1 P2 P3 P4; split=>//=; done. have P2: valid i2 by apply: (cohS (proj2 (rely_coh R1))). -have P3: l \in dom i2 by rewrite-(cohD (proj2 (rely_coh R1)))/ddom gen_domPt inE/=. +have P3: l \in dom i2 by rewrite-(cohD (proj2 (rely_coh R1)))/ddom domPt inE/=. (* See the remark [Overly cautious filtering] *) have D: if c then rt = pn_receive_commit_ack_trans _ Hnin \/ @@ -353,8 +359,12 @@ have Pre: forall i2, network_rely W p i i2 -> case: hasCommitted E1; split=>//; exists X, C, e, d; by rewrite -(rely_loc' _ R1) in E1; rewrite (getStP_K _ _ _ _ E1). + rewrite /Actions.can_send /nodes inE/= mem_cat Hpin orbC. - by rewrite -(cohD (proj2 (rely_coh R1)))/ddom gen_domPt inE/= eqxx. - by rewrite /Actions.filter_hooks um_filt0=>???/sym/find_some; rewrite dom0 inE. + by rewrite -(cohD (proj2 (rely_coh R1)))/ddom domPt inE/= eqxx. + rewrite /Actions.filter_hooks umfilt0=>???. + move => F. + apply sym_eq in F. + move: F. + by move/find_some; rewrite dom0. case: hasCommitted E1 Pre=>E1 Pre; apply: act_rule=>i2 R1/=; move:(Pre i2 R1)=>{Pre}Pre. (* Send commit-ack *) @@ -363,11 +373,11 @@ case: {-1}(Sf)=>_/=[]_[H][]C'//; rewrite -(rely_loc' _ R1) in E1. rewrite (getStP_K Hnin C' (pn_this_in _ H) Hpin E1)/=. case=>e'[d'][][]Z1 Z2 _ G; subst d' e'. move: St; rewrite/Actions.send_act_step/==>[][_][h][]. -rewrite /pn_step -!(proof_irrelevance C' (pn_safe_coh _)) - -!(proof_irrelevance (pn_this_in _ H) (pn_this_in _ _)). +rewrite /pn_step -!(pf_irr C' (pn_safe_coh _)) + -!(pf_irr (pn_this_in _ H) (pn_this_in _ _)). rewrite (getStP_K Hnin C' (pn_this_in _ H) Hpin E1)(getStL_Kp _ (pn_this_in _ H) E1)/=. case=>Z Z';subst h i3; rewrite (rely_loc' _ R3) locE//; last by apply: (cohVl C'). -+ by rewrite -(cohD (proj2 (rely_coh R1)))/ddom gen_domPt inE/=. ++ by rewrite -(cohD (proj2 (rely_coh R1)))/ddom domPt inE/=. by apply: (cohS (proj2 (rely_coh R1))). (* Send abort-ack *) (* TODO: Remove proof duplication!!! *) @@ -376,11 +386,11 @@ case: {-1}(Sf)=>_/=[]_[H][]C'//; rewrite -(rely_loc' _ R1) in E1. rewrite (getStP_K Hnin C' (pn_this_in _ H) Hpin E1)/=. case=>e'[d'][][]Z1 Z2 _ G; subst d' e'. move: St; rewrite/Actions.send_act_step/==>[][_][h][]. -rewrite /pn_step -!(proof_irrelevance C' (pn_safe_coh _)) - -!(proof_irrelevance (pn_this_in _ H) (pn_this_in _ _)). +rewrite /pn_step -!(pf_irr C' (pn_safe_coh _)) + -!(pf_irr (pn_this_in _ H) (pn_this_in _ _)). rewrite (getStP_K Hnin C' (pn_this_in _ H) Hpin E1)(getStL_Kp _ (pn_this_in _ H) E1)/=. case=>Z Z';subst h i3; rewrite (rely_loc' _ R3) locE//; last by apply: (cohVl C'). -+ by rewrite -(cohD (proj2 (rely_coh R1)))/ddom gen_domPt inE/=. ++ by rewrite -(cohD (proj2 (rely_coh R1)))/ddom domPt inE/=. by apply: (cohS (proj2 (rely_coh R1))). Qed. diff --git a/Examples/TwoPhaseCommit/TwoPhaseProtocol.v b/Examples/TwoPhaseCommit/TwoPhaseProtocol.v index e497477..4ecb858 100644 --- a/Examples/TwoPhaseCommit/TwoPhaseProtocol.v +++ b/Examples/TwoPhaseCommit/TwoPhaseProtocol.v @@ -4,11 +4,11 @@ From mathcomp Require Import path. Require Import Eqdep. Require Import Relation_Operators. -From DiSeL.Heaps -Require Import pred prelude idynamic ordtype finmap pcm unionmap heap coding domain. -From DiSeL.Core +From fcsl +Require Import axioms pred prelude ordtype finmap pcm unionmap heap. +From DiSeL Require Import Freshness State EqTypeX DepMaps Protocols Worlds NetworkSem Rely. -From DiSeL.Core +From DiSeL Require Import Actions Injection Process Always HoareTriples InferenceRules. Set Implicit Arguments. @@ -176,7 +176,7 @@ Proof. move=>[H1 H2][y]Cm; split=>[|i ms/=]; first by rewrite valid_fresh. rewrite findUnL; last by rewrite valid_fresh. case: ifP=>E; first by move/H2. -by move/um_findPt_inv=>[Z G]; subst i m; exists y. +by move/findPt_inv=>[Z G]; subst i m; exists y. Qed. Lemma trans_updDom this d s : @@ -228,28 +228,29 @@ Qed. Lemma cohStC d (C : TPCCoh d) s: find st (getLocal cn d) = Some s -> - idyn_tp s = CStateT. + dyn_tp s = CStateT. Proof. have pf: cn \in nodes by rewrite inE eqxx. -move: (locCn C pf); rewrite eqxx; move =>[V][s'][l']Z; rewrite Z in V *; -by rewrite (@hfindPtUn _ st s' _ V); case=><-/=. +move: (locCn C pf); rewrite eqxx; move =>[V][s'][l']Z; rewrite Z in V *. +rewrite findPtUn //. +by case=><-/=. Qed. Lemma cohStP n d (C : TPCCoh d) (H : n \in pts) s: find st (getLocal n d) = Some s -> - idyn_tp s = PStateT. + dyn_tp s = PStateT. Proof. have pf: n \in nodes by rewrite inE/=orbC mem_cat H. move: (locCn C pf); rewrite H=>[[V]]. case E: (n == cn); last first. - move=>[s'][l']Z; rewrite Z in V *. - by rewrite findUnL//; rewrite hdomPt inE/= hfindPt/=; case=><-. + by rewrite findUnL//; rewrite domPt inE/= findPt/=; case=><-. by move/eqP: E=>E; subst n; move: Hnin; rewrite H. Qed. Definition getStC d (C : TPCCoh d) : CStateT := match find st (getLocal cn d) as f return _ = f -> _ with - Some v => fun epf => icoerce id (idyn_val v) (cohStC C epf) + Some v => fun epf => icast (sym_eq (cohStC C epf)) (dyn_val v) | _ => fun epf => (0, CInit) end (erefl _). @@ -260,14 +261,14 @@ move=>E; rewrite /getStC/=. have pf : cn \in nodes by rewrite inE eqxx. have V: valid (getLocal cn d) by case: (locCn C pf). move: (cohStC C); rewrite !E=>/= H. -by apply: ieqc. +by apply: eqc. Qed. Program Definition getStP n d (C : TPCCoh d) (pf : n \in nodes) : PStateT. Proof. case X: (n \in pts); last by exact: (0, PInit). exact: (match find st (getLocal n d) as f return _ = f -> _ with - Some v => fun epf => icoerce id (idyn_val v) (cohStP C X epf) + Some v => fun epf => icast (sym_eq (cohStP C X epf)) (dyn_val v) | _ => fun epf => (0, PInit) end (erefl _)). Defined. @@ -280,28 +281,28 @@ have V: valid (getLocal n d) by case: (locCn C pf). rewrite E in V. move: (cohStP C); case B: (n \in pts)=>//=; last by rewrite X in B. move=>H; move: (H (erefl true))=>{H}; rewrite E=>/=H. -by apply: ieqc. +by apply: eqc. Qed. (* Log getter *) Lemma cohStL d (C : TPCCoh d) n (H : n \in nodes) l: - find log (getLocal n d) = Some l -> idyn_tp l = Log. + find log (getLocal n d) = Some l -> dyn_tp l = Log. Proof. move: (locCn C H)=>[V]. case B: (n == cn)=>/=. - move=>[s'][l']Z; rewrite Z in V *; - by rewrite joinC in V *; rewrite (@hfindPtUn _ log l' _ V); case=><-/=. + by rewrite joinC in V *; rewrite findPtUn //; case=><-/=. rewrite inE in H; case/orP: H; first by rewrite B. rewrite/= mem_cat; case X: (n \in pts)=>/=_. - move=>[s'][l']Z; rewrite Z in V *; - by rewrite joinC in V *; rewrite (@hfindPtUn _ log l' _ V); case=><-/=. + by rewrite joinC in V *; rewrite findPtUn //; case=><-/=. by move=>H; move/find_some=>Y; rewrite Y in H. Qed. Definition getStL n d (C : TPCCoh d) (pf : n \in nodes) : Log := match find log (getLocal n d) as f return _ = f -> _ with - Some v => fun epf => icoerce id (idyn_val v) (cohStL C pf epf) + Some v => fun epf => icast (sym_eq (cohStL C pf epf)) (dyn_val v) | _ => fun epf => [::] end (erefl _). @@ -311,7 +312,7 @@ Lemma getStL_Kc n d (C : TPCCoh d) (pf : n \in nodes) (m : CStateT) (l : Log): Proof. move=>E; rewrite /getStL/=. have V: valid (getLocal n d) by case: (locCn C pf). -by rewrite E in V; move: (cohStL C pf); rewrite !E/==>H; apply: ieqc. +by rewrite E in V; move: (cohStL C pf); rewrite !E/==>H; apply: eqc. Qed. Lemma getStL_Kp n d (C : TPCCoh d) (pf : n \in nodes) (m : PStateT) (l : Log): @@ -319,7 +320,7 @@ Lemma getStL_Kp n d (C : TPCCoh d) (pf : n \in nodes) (m : PStateT) (l : Log): Proof. move=>E; rewrite /getStL/=. have V: valid (getLocal n d) by case: (locCn C pf). -by rewrite E in V; move: (cohStL C pf); rewrite !E/==>H; apply: ieqc. +by rewrite E in V; move: (cohStL C pf); rewrite !E/==>H; apply: eqc. Qed. Lemma cn_in : cn \in nodes. @@ -532,13 +533,13 @@ split=>/=. - apply: send_soupCoh; first by case:(cn_safe_coh pf). exists (getStC (cn_safe_coh pf)).1. case: (pf)=>H[C']P/=; move: (conj H _)=>pf'. - by move: (cn_prec_safe H P); rewrite -(proof_irrelevance C' (cn_safe_coh pf')). + by move: (cn_prec_safe H P); rewrite -(pf_irr C' (cn_safe_coh pf')). - by apply: trans_updDom=>//; case: (cn_safe_in pf). - by rewrite validU; apply: cohVl C. move=>n Ni. rewrite /localCoh/=. rewrite /getLocal/=findU; case: ifP=>B; last by case: C=>_ _ _/(_ n Ni). move/eqP: B=>Z; subst n this; rewrite eqxx (cohVl C)/=. -by split; rewrite ?hvalidPtUn//; last by eexists _, _. +by split; rewrite ?validPtUn//; last by eexists _, _. Qed. Lemma cn_safe_def this to d msg : @@ -750,8 +751,8 @@ split=>/=. - apply: send_soupCoh; first by case:(pn_safe_coh pf). exists (getStP C (pn_this_in (proj1 pf))).1. case: (pf)=>H[H'][C']P/=; move: (conj H _)=>pf'. - move: (pn_prec_safe H P); rewrite (proof_irrelevance C' C)/=. - by rewrite (proof_irrelevance (pn_this_in H') _)//; apply: pn_this_in. + move: (pn_prec_safe H P); rewrite (pf_irr C' C)/=. + by rewrite (pf_irr (pn_this_in H') _)//; apply: pn_this_in. - by apply: trans_updDom=>//; case: (pn_safe_in pf). - by rewrite validU; apply: cohVl C. move=>n Ni. rewrite /localCoh/=. @@ -760,7 +761,7 @@ move/eqP: B=>Z; subst n=>/=. have X : this == cn = false by apply/negP=>/eqP Z; subst this; move: E (Hnin)=>->. rewrite X (cohVl C)/=; split=>//. move: (pstep_send _ _ _)=>ps; rewrite E. -rewrite ?hvalidPtUn//; last by eexists _, _. +rewrite ?validPtUn//; last by eexists _, _. Qed. Lemma pn_safe_def this to d msg : @@ -968,5 +969,3 @@ End TPCProtocol. Export TPCProtocol.States. Export TPCProtocol.Exports. - - diff --git a/Heaps/Make b/Heaps/Make deleted file mode 100644 index e00b734..0000000 --- a/Heaps/Make +++ /dev/null @@ -1,15 +0,0 @@ --Q . DiSeL.Heaps --arg "-w -notation-overridden,-local-declaration,-redundant-canonical-projection,-projection-no-head-constant" - -./coding.v -./domain.v -./finmap.v -./heap.v -./idynamic.v -./ordtype.v -./pcm.v -./pperm.v -./pred.v -./prelude.v -./spcm.v -./unionmap.v diff --git a/Heaps/Makefile b/Heaps/Makefile deleted file mode 100644 index 0898c8e..0000000 --- a/Heaps/Makefile +++ /dev/null @@ -1,19 +0,0 @@ -ifeq "$(COQBIN)" "" -COQBIN=$(dir $(shell which coqtop))/ -endif - -default: Makefile.coq - $(MAKE) -f Makefile.coq - -install: Makefile.coq - $(MAKE) -f Makefile.coq install - -clean: - if [ -f Makefile.coq ]; then \ - $(MAKE) -f Makefile.coq clean; fi - rm -f Makefile.coq Makefile.coq.conf - -Makefile.coq: Make - coq_makefile -f Make -o Makefile.coq - -.PHONY: default clean install diff --git a/Heaps/coding.v b/Heaps/coding.v deleted file mode 100644 index 746bfa3..0000000 --- a/Heaps/coding.v +++ /dev/null @@ -1,844 +0,0 @@ -From mathcomp.ssreflect -Require Import ssreflect ssrbool ssrnat eqtype fintype ssrfun seq. -From mathcomp -Require Import finfun. -Require Import Eqdep. -From DiSeL.Heaps -Require Import pred prelude idynamic ordtype domain pcm spcm unionmap heap. -Set Implicit Arguments. -Unset Strict Implicit. -Unset Printing Implicit Defensive. - -(*****************************************) -(* Codes for types, small types and pcms *) -(*****************************************) - -(* first a data structure of small types *) -(* this is similar to pcms, except it has no operations on it *) -(* i will register a number of types as sets, but, importantly, *) -(* not heaps, which are a larger entity *) - -Module SetType. - -Section ClassDef. -Structure type : Type := Pack {sort : Type}. -Local Coercion sort : type >-> Sortclass. - -Variables (T : Type) (cT : type). -Definition pack := @Pack T. -Definition clone := fun (_ : cT -> T) & phant_id pack cT => pack. -End ClassDef. - -Module Exports. -Coercion sort : type >-> Sortclass. -Notation set := type. -Notation SET T := (@pack T). - -Notation "[ 'set' 'of' T 'for' C ]" := (@clone T C idfun id) - (at level 0, format "[ 'set' 'of' T 'for' C ]") : form_scope. -Notation "[ 'set' 'of' T ]" := (@clone T _ id id) - (at level 0, format "[ 'set' 'of' T ]") : form_scope. - -End Exports. -End SetType. - -Export SetType.Exports. - -Canonical unitSet := SET unit. -Canonical natSet := SET nat. -Canonical boolSet := SET bool. -Canonical ptrSet := SET ptr. -Canonical mutexSet := SET mutex. -Canonical optionSet (s : set) := SET (option s). -Canonical seqSet (s : set) := SET (seq s). -Canonical prodSet (s1 s2 : set) := SET (prod s1 s2). -Canonical arrSet (s1 s2 : set) := SET (s1 -> s2). -(* Canonical natfinsetSet := SET natfinset. *) - -(*************************************************) -(* Now the actual codes for types, sets and pcms *) -(*************************************************) - -(* so far no need for mutual induction *) -(* It's possible to define this mutually inductively *) -(* but Coq's default induction principles won't account for it *) -(* One has to use Scheme command for that (check CPDT) *) - -(* small types *) -Inductive st_code := - st_unit | st_nat | st_bool | st_ptr | st_mtx | (* st_natfinset | *) (* nullary *) - st_option of st_code | st_seq of st_code | (* unary *) - st_prod of st_code & st_code | st_arr of st_code & st_code. (* binary *) - -(* small pcms, unused so far *) -Inductive spcm_code := - spcm_unit | spcm_nat | spcm_natmap of st_code | spcm_ptrmap of st_code | - spcm_prod of spcm_code & spcm_code. - -(* general pcm's *) -Inductive pcm_code := - pcm_unit | pcm_nat | pcm_mtx | pcm_heap | pcm_natmap of st_code | - pcm_ptrmap of st_code | pcm_prod of pcm_code & pcm_code. - -(* general types *) -Inductive tp_code := - tp_unit | tp_nat | tp_bool | tp_ptr | tp_mtx | tp_heap | (* nullary *) - tp_option of tp_code | tp_seq of tp_code | (* unary *) - tp_prod of tp_code & tp_code | tp_arr of tp_code & tp_code | (* binary *) - tp_array of nat & pcm_code | tp_ptrmap of st_code. - -Fixpoint st_sort (x : st_code) : set := - match x with - | st_unit => [set of unit] - | st_nat => [set of nat] - | st_bool => [set of bool] - | st_ptr => [set of ptr] - | st_mtx => [set of mutex] - (* | st_natfinset => [set of natfinset] *) - | st_option y => [set of option (st_sort y)] - | st_seq y => [set of seq (st_sort y)] - | st_prod y z => [set of (st_sort y * st_sort z)] - | st_arr y z => [set of st_sort y -> st_sort z] - end. - -Fixpoint spcm_sort (x : spcm_code) : spcm := - match x with - | spcm_unit => unitSPCM - | spcm_nat => natSPCM - | spcm_natmap t2 => natmapSPCM (st_sort t2) - | spcm_ptrmap t2 => ptrmapSPCM (st_sort t2) - | spcm_prod y z => prodSPCM (spcm_sort y) (spcm_sort z) - end. - -Fixpoint pcm_sort (x : pcm_code) : pcm := - match x with - | pcm_unit => [pcm of unit] - | pcm_nat => [pcm of nat] - | pcm_mtx => [pcm of lift [unlifted of mutex]] - | pcm_heap => [pcm of heap] - | pcm_natmap t2 => [pcm of union_map nat_ordType (st_sort t2)] - | pcm_ptrmap t2 => [pcm of union_map ptr_ordType (st_sort t2)] - | pcm_prod y z => [pcm of pcm_sort y * pcm_sort z] - end. - -Fixpoint tp_sort (x : tp_code) : Type := - match x with - | tp_unit => unit - | tp_nat => nat - | tp_bool => bool - | tp_ptr => ptr - | tp_mtx => mutex - | tp_heap => heap - | tp_option y => option (tp_sort y) - | tp_seq y => seq (tp_sort y) - | tp_prod y z => (tp_sort y * tp_sort z)%type - | tp_arr y z => tp_sort y -> tp_sort z - | tp_array n z => {ffun 'I_n -> pcm_sort z} - | tp_ptrmap z => union_map ptr_ordType (st_sort z) - end. - -(* deciding equality on each of these types *) - -Fixpoint st_eq x y := - match x, y with - | st_unit, st_unit => true - | st_nat, st_nat => true - | st_bool, st_bool => true - | st_ptr, st_ptr => true - | st_mtx, st_mtx => true - (* | st_natfinset, st_natfinset => true *) - | st_option x1, st_option y1 => st_eq x1 y1 - | st_seq x1, st_seq y1 => st_eq x1 y1 - | st_prod x1 y1, st_prod x2 y2 => - st_eq x1 x2 && st_eq y1 y2 - | st_arr x1 y1, st_arr x2 y2 => - st_eq x1 x2 && st_eq y1 y2 - | _, _ => false - end. - -Fixpoint spcm_eq x y := - match x, y with - | spcm_unit, spcm_unit => true - | spcm_nat, spcm_nat => true - | spcm_natmap t1, spcm_natmap t2 => st_eq t1 t2 - | spcm_ptrmap t1, spcm_ptrmap t2 => st_eq t1 t2 - | spcm_prod x1 y1, spcm_prod x2 y2 => - spcm_eq x1 x2 && spcm_eq y1 y2 - | _, _ => false - end. - -Fixpoint pcm_eq x y := - match x, y with - | pcm_unit, pcm_unit => true - | pcm_nat, pcm_nat => true - | pcm_mtx, pcm_mtx => true - | pcm_heap, pcm_heap => true - | pcm_natmap t1, pcm_natmap t2 => st_eq t1 t2 - | pcm_ptrmap t1, pcm_ptrmap t2 => st_eq t1 t2 - | pcm_prod x1 y1, pcm_prod x2 y2 => - pcm_eq x1 x2 && pcm_eq y1 y2 - | _, _ => false - end. - -Fixpoint tp_eq x y := - match x, y with - | tp_unit, tp_unit => true - | tp_nat, tp_nat => true - | tp_bool, tp_bool => true - | tp_ptr, tp_ptr => true - | tp_mtx, tp_mtx => true - | tp_heap, tp_heap => true - | tp_option x1, tp_option y1 => tp_eq x1 y1 - | tp_seq x1, tp_seq y1 => tp_eq x1 y1 - | tp_prod x1 y1, tp_prod x2 y2 => - tp_eq x1 x2 && tp_eq y1 y2 - | tp_arr x1 y1, tp_arr x2 y2 => - tp_eq x1 x2 && tp_eq y1 y2 - | tp_array n1 y1, tp_array n2 y2 => - (n1 == n2) && (pcm_eq y1 y2) - | tp_ptrmap y1, tp_ptrmap y2 => st_eq y1 y2 - | _, _ => false - end. - - -(* all are equality types *) - -Lemma st_eqP : Equality.axiom st_eq. -Proof. -elim; -(* nullaries *) -try by [case; constructor]; -(* unaries *) -try by - [move=>x IH; case; try by [constructor]; - move=>y; apply: iffP (IH y) _ _; [move=>-> | case]]; -(* binaries *) -by [move=>x1 IHx y1 IHy; case; try by [constructor]; - move=>x2 y2 /=; apply: iffP andP _ _; - case; move/IHx=>->; move/IHy=>->]. -Qed. - -Lemma spcm_eqP : Equality.axiom spcm_eq. -Proof. -elim; -try by [case; constructor]. -- move=>t; case=>/=; try constructor=>//; move=>y. - by case: st_eqP=>[->|H]; constructor=>//; case. -- move=>t; case=>/=; try constructor=>//; move=>y. - by case: st_eqP=>[->|H]; constructor=>//; case. -by [move=>x1 IHx y1 IHy; case; try by [constructor]; - move=>x2 y2 /=; apply: iffP andP _ _; - case; move/IHx=>->; move/IHy=>->]. -Qed. - -Lemma pcm_eqP : Equality.axiom pcm_eq. -Proof. -elim; -try by [case; constructor]. -- move=>t; case=>/=; try constructor=>//; move=>y. - by case: st_eqP=>[->|H]; constructor=>//; case. -- move=>t; case=>/=; try constructor=>//; move=>y. - by case: st_eqP=>[->|H]; constructor=>//; case. -by [move=>x1 IHx y1 IHy; case; try by [constructor]; - move=>x2 y2 /=; apply: iffP andP _ _; - case; move/IHx=>->; move/IHy=>->]. -Qed. - -Lemma tp_eqP : Equality.axiom tp_eq. -Proof. -elim; -(* nullaries *) -try by [case; constructor]; -(* unaries *) -try by - [move=>x IH; case; try by [constructor]; - move=>y; apply: iffP (IH y) _ _; [move=>-> | case]]; -(* binaries *) -try by [move=>x1 IHx y1 IHy; case; try by [constructor]; - move=>x2 y2 /=; apply: iffP andP _ _; - case; move/IHx=>->; move/IHy=>->]. -(* array case *) -move=>n1 p1 []; try by [constructor]. -move=>n2 p2 /=; case: eqP=>[->|H]; last by constructor; case. -by case: pcm_eqP=>[->|H]; constructor=>//; case. -(* map case *) -move=>s1 []; try by [constructor]. -by move=>s2 /=; case: st_eqP=>[->|H]; constructor=>//; case. -Qed. - -(* declaration of equality structures *) -Definition st_code_eqMixin := EqMixin st_eqP. -Canonical st_code_eqType := EqType st_code st_code_eqMixin. - -Definition spcm_code_eqMixin := EqMixin spcm_eqP. -Canonical spcm_code_eqType := EqType spcm_code spcm_code_eqMixin. - -Definition pcm_code_eqMixin := EqMixin pcm_eqP. -Canonical pcm_code_eqType := EqType pcm_code pcm_code_eqMixin. - -Definition tp_code_eqMixin := EqMixin tp_eqP. -Canonical tp_code_eqType := EqType tp_code tp_code_eqMixin. - -(*********************************) -(* Now the structures for coding *) -(*********************************) - -Module Code. - -Structure mixin_of (T : Type) := Mixin { - code_dom_op : eqType; - encode_op : code_dom_op -> T; - to_type_op : T -> Type}. - -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 code_dom := code_dom_op class. -Definition encode := @encode_op _ class. -Definition to_type := to_type_op class. - -End ClassDef. - -Module Exports. -Coercion sort : type >-> Sortclass. -Notation code := type. -Notation codeMixin c i := (@Mixin _ c i). -Notation Code T m := (@pack T m). - -Notation code_dom := code_dom. -Notation encode := encode. -Notation to_type := to_type. - -Implicit Arguments to_type [cT]. -Implicit Arguments encode [cT]. - -Prenex Implicits to_type encode. - -Notation "[ 'codeMixin' 'of' T ]" := (class _ : mixin_of T) - (at level 0, format "[ 'codeMixin' 'of' T ]") : form_scope. -Notation "[ 'code' 'of' T 'for' C ]" := (@clone T C _ idfun id) - (at level 0, format "[ 'code' 'of' T 'for' C ]") : form_scope. -Notation "[ 'code' 'of' T ]" := (@clone T _ _ id id) - (at level 0, format "[ 'code' 'of' T ]") : form_scope. - -End Exports. - -End Code. - -Export Code.Exports. - - -(* each of the above datatypes is a code according to above definition *) - -Definition set_codeMixin := codeMixin st_code_eqType st_sort SetType.sort. -Canonical setCode := Code set set_codeMixin. - -Definition spcm_codeMixin := codeMixin spcm_code_eqType spcm_sort SPCM.sort. -Canonical spcmCode := Code spcm spcm_codeMixin. - -Definition pcm_codeMixin := codeMixin pcm_code_eqType pcm_sort PCM.sort. -Canonical pcmCode := Code pcm pcm_codeMixin. - -Definition type_codeMixin := codeMixin tp_code_eqType tp_sort id. -Canonical typeCode := Code Type type_codeMixin. - -(*********************************************) -(* Computing code instance for a PCM or type *) -(*********************************************) - -(* this is a canonical structure setup that inverts the interp function. *) -(* important so that we avoid excesive pcm_code or tp_code annotations. *) - -Module Encoded. -Section Encoded. -Variable C : code. - -Structure mixin_of (v : C) := Mixin { - decode_op : code_dom C; - _ : encode decode_op = v}. - -Section ClassDef. - -Notation class_of := mixin_of (only parsing). - -Structure type : Type := Pack {sort : C; _ : class_of sort; _ : C}. -Local Coercion sort : type >-> Code.sort. - -Variables (v : C) (cT : type). -Definition class := let: Pack _ c _ as cT' := cT return class_of cT' in c. - -Definition pack c := @Pack v c v. - -Definition clone := - fun c & to_type cT -> to_type v & phant_id (pack c) cT => pack c. - -Definition decode := decode_op class. - -End ClassDef. -End Encoded. - -Module Exports. -Coercion sort : type >-> Code.sort. -Notation encoded := type. -Notation encodeMixin := Mixin. -Notation Encode C v m := (@pack C v m). - -Notation decode := decode. - -Implicit Arguments encodeMixin []. - -Section Lemmas. -Variable C : code. -Implicit Type U : encoded C. - -(* the next lemma is problematic to use in practice *) -(* as the rewriting gets in trouble with dependencies *) -(* the solution is to break up U, just like in proof below *) - -Lemma encdecE U : encode (decode U) = U. -Proof. by case: U=>tp []. Qed. - -Lemma decode_inj U1 U2 : decode U1 = decode U2 -> U1 = U2 :> C. -Proof. -case: U1=>s1 [dec1 pf1 t1]; case: U2=>s2 [dec2 pf2 t2] /=. -by rewrite /decode /= -pf1 -pf2=>->. -Qed. - -Definition code_dyn U (v : to_type U) := - @idyn (code_dom C) (to_type \o encode) (decode U) - (icoerce to_type v (esym (encdecE U))). - -(* every idynamic can be recast into a code_dyn form *) - -Section CodeEta. -Variable u : idynamic (to_type \o @encode C). - -Definition code_enc : encoded C := - let: v := encode (idyn_tp u) in - let: m := @encodeMixin C v (idyn_tp u) (erefl _) in - Encode C v m. - -Definition code_val : to_type code_enc := idyn_val u. - -Lemma code_eta : u = code_dyn code_val. -Proof. by rewrite -{1}(idyn_eta u); apply: idynE=>// pf; rewrite !ieqc. Qed. - -End CodeEta. - -End Lemmas. - -Notation "[ 'encodedMixin' C 'for' v ]" := (@class C _ : @mixin_of C v) - (at level 0, format "[ 'encodedMixin' C 'for' v ]") : form_scope. -Notation "[ 'encoded' C 'of' v 'for' R ]" := (@clone C v R _ idfun id) - (at level 0, format "[ 'encoded' C 'of' v 'for' R ]") : form_scope. -Notation "[ 'encoded' C 'of' v ]" := (@clone C v _ _ idfun id) - (at level 0, format "[ 'encoded' C 'of' v ]") : form_scope. - -End Exports. - -End Encoded. - -Export Encoded.Exports. - -Definition encoded_set := encoded setCode. -Definition encodeMixinST := encodeMixin setCode. -Definition EncodeST v m : encoded_set := @Encoded.pack setCode v m. -Implicit Arguments encodeMixinST []. -Implicit Arguments EncodeST []. - -Definition encoded_spcm := encoded spcmCode. -Definition encodeMixinSPCM := encodeMixin spcmCode. -Definition EncodeSPCM v m : encoded_spcm := @Encoded.pack spcmCode v m. -Implicit Arguments encodeMixinSPCM []. -Implicit Arguments EncodeSPCM []. - -Definition encoded_pcm := encoded pcmCode. -Definition encodeMixinPCM := encodeMixin pcmCode. -Definition EncodePCM v m : encoded_pcm := @Encoded.pack pcmCode v m. -Implicit Arguments encodeMixinPCM []. -Implicit Arguments EncodePCM []. - -Definition encoded_type := encoded typeCode. -Definition encodeMixinTP := encodeMixin typeCode. -Definition EncodeTP v m : encoded_type := @Encoded.pack typeCode v m. -Implicit Arguments encodeMixinTP []. -Implicit Arguments EncodeTP []. - -(* add the explicit coercions *) - -Definition encoded_sort_set : encoded_set -> set := - fun U => Encoded.sort U. -Coercion encoded_sort_set : encoded_set >-> set. - -Definition encoded_sort_spcm : encoded_spcm -> spcm := - fun U => Encoded.sort U. -Coercion encoded_sort_spcm : encoded_spcm >-> spcm. - -Definition encoded_sort_pcm : encoded_pcm -> pcm := - fun U => Encoded.sort U. -Coercion encoded_sort_pcm : encoded_pcm >-> pcm. - -Definition encoded_sort_type : encoded_type -> Type := - fun U => Encoded.sort U. -Coercion encoded_sort_type : encoded_type >-> Sortclass. - -(* Can I coerce spcms to pcms? *) - -Fixpoint spcm_to_pcm (x : spcm_code) : pcm_code := - match x with - | spcm_unit => pcm_unit - | spcm_nat => pcm_nat - | spcm_natmap t2 => pcm_natmap t2 - | spcm_ptrmap t2 => pcm_ptrmap t2 - | spcm_prod y z => pcm_prod (spcm_to_pcm y) (spcm_to_pcm z) - end. - -Lemma spcm_pcm_pf (V : encoded_spcm) : - @encode pcmCode (spcm_to_pcm (decode V)) = V. -Proof. -case: V=>V [d E /=] s; rewrite /decode /= -{}E /encode. -elim: d=>[||d|d|/= s1 -> s2 ->]; -by congr PCM.pack; congr PCM.Mixin; try apply: proof_irrelevance. -Qed. - -Definition spcm_pcm_encodedMixin (V : encoded_spcm) := - encodeMixinPCM V (spcm_to_pcm (decode V)) (spcm_pcm_pf V). -Canonical spcm_pcm_Encoded (V : encoded_spcm) := - EncodePCM V (spcm_pcm_encodedMixin V). - -Coercion spcm_pcm_Encoded : encoded_spcm >-> encoded_pcm. - -(* Moving on; special notation *) - -Definition pcm_dyn (U : encoded_pcm) (v : U) := @code_dyn pcmCode U v. - -Section PCMEta. -Variable u : idynamic (to_type \o pcm_sort). - -Definition pcm_enc : encoded_pcm := @code_enc pcmCode u. -Definition pcm_val : pcm_enc := @code_val pcmCode u. -Definition pcm_eta : u = pcm_dyn pcm_val := @code_eta pcmCode u. - -End PCMEta. - -(****************************) -(* inversion table for sets *) -(****************************) - -Definition unit_st_encodedMixin := encodeMixinST unitSet st_unit (erefl _). -Canonical unit_st_Encoded := EncodeST unitSet unit_st_encodedMixin. - -Definition nat_st_encodedMixin := encodeMixinST natSet st_nat (erefl _). -Canonical nat_st_Encoded := EncodeST natSet nat_st_encodedMixin. - -Definition bool_st_encodedMixin := encodeMixinST boolSet st_bool (erefl _). -Canonical bool_st_Encoded := EncodeST boolSet bool_st_encodedMixin. - -Definition ptr_st_encodedMixin := encodeMixinST ptrSet st_ptr (erefl _). -Canonical ptr_st_Encoded := EncodeST ptrSet ptr_st_encodedMixin. - -Definition mtx_st_encodedMixin := encodeMixinST mutexSet st_mtx (erefl _). -Canonical mtx_st_Encoded := EncodeST mutexSet mtx_st_encodedMixin. - -(* -Definition natfinset_st_encodedMixin := - encodeMixinST natfinsetSet st_natfinset (erefl _). -Canonical natfinset_st_Encoded := EncodeST natfinsetSet natfinset_st_encodedMixin. -*) - -Lemma option_st_pf (U : encoded_set) : - @encode setCode (st_option (decode U)) = optionSet U. -Proof. by rewrite /decode; case: U=>st [/= dp <-]. Qed. - -Definition option_st_encodedMixin (U : encoded_set) := - encodeMixin setCode (optionSet U) (st_option (decode U)) (option_st_pf U). -Canonical option_st_Encoded (U : encoded_set) := - EncodeST (optionSet U) (option_st_encodedMixin U). - -Lemma seq_st_pf (U : encoded_set) : - @encode setCode (st_seq (decode U)) = seqSet U. -Proof. by rewrite /decode; case: U=>st [/= dp <-]. Qed. - -Definition seq_st_encodedMixin (U : encoded_set) := - encodeMixin setCode (seqSet U) (st_seq (decode U)) (seq_st_pf U). -Canonical seq_st_Encoded (U : encoded_set) := - EncodeST (seqSet U) (seq_st_encodedMixin U). - -Lemma prod_st_pf (U1 U2 : encoded_set) : - @encode setCode (st_prod (decode U1) (decode U2)) = prodSet U1 U2. -Proof. by rewrite /decode; case: U1 U2=>t1 [/= ? <-] _ [t2 [/= ? <-]]. Qed. - -Definition prod_st_encodedMixin (U1 U2 : encoded_set) := - encodeMixinST (prodSet U1 U2) - (st_prod (decode U1) (decode U2)) (prod_st_pf U1 U2). -Canonical prod_st_Encoded (U1 U2 : encoded_set) := - EncodeST (prodSet U1 U2) (prod_st_encodedMixin U1 U2). - -(* this one is kinda pointless since we can't trigger inference of lambda's *) -(* so functions will have to be decorated with their types by hand *) - -Lemma arr_st_pf (U1 U2 : encoded_set) : - @encode setCode (st_arr (decode U1) (decode U2)) = arrSet U1 U2. -Proof. by rewrite /decode; case: U1 U2=>t1 [/= ? <-] _ [t2 [/= ? <-]]. Qed. - -Definition arr_st_encodedMixin (U1 U2 : encoded_set) := - encodeMixinST (arrSet U1 U2) (st_arr (decode U1) (decode U2)) (arr_st_pf U1 U2). -Canonical arr_st_Encoded (U1 U2 : encoded_set) := - EncodeST (arrSet U1 U2) (arr_st_encodedMixin U1 U2). - -(* -Notation "[ 'encoded_set' 'of' v ]" := - ([encoded setCode of [set of v]]) - (at level 0, format "[ 'encoded_set' 'of' v ]") : form_scope. -*) - -Definition set_clone v cT c b1 b2 : encoded_set := - @Encoded.clone setCode v cT c b1 b2. - -Notation "[ 'encoded_set' 'of' v ]" := (@set_clone [set of v] _ _ idfun id) - (at level 0, format "[ 'encoded_set' 'of' v ]") : form_scope. - -(*****************************) -(* inversion table for types *) -(*****************************) - -Definition unit_tp_encodedMixin := encodeMixinTP (unit : Type) tp_unit (erefl _). -Canonical unit_tp_Encoded := EncodeTP unit unit_tp_encodedMixin. - -Definition nat_tp_encodedMixin := encodeMixinTP nat tp_nat (erefl _). -Canonical nat_tp_Encoded := EncodeTP nat nat_tp_encodedMixin. - -Definition bool_tp_encodedMixin := encodeMixinTP bool tp_bool (erefl _). -Canonical bool_tp_Encoded := EncodeTP bool bool_tp_encodedMixin. - -Definition ptr_tp_encodedMixin := encodeMixinTP ptr tp_ptr (erefl _). -Canonical ptr_tp_Encoded := EncodeTP ptr ptr_tp_encodedMixin. - -Definition mtx_tp_encodedMixin := encodeMixinTP mutex tp_mtx (erefl _). -Canonical mtx_tp_Encoded := EncodeTP mutex mtx_tp_encodedMixin. - -Definition heap_tp_encodedMixin := encodeMixinTP heap tp_heap (erefl _). -Canonical heap_tp_Encoded := EncodeTP heap heap_tp_encodedMixin. - -Lemma option_tp_pf (U : encoded_type) : - @encode typeCode (tp_option (decode U)) = option U. -Proof. by rewrite /decode; case: U=>tp [/= dp <-]. Qed. - -Definition option_tp_encodedMixin (U : encoded_type) := - encodeMixin typeCode (option U) (tp_option (decode U)) (option_tp_pf U). -Canonical option_tp_Encoded (U : encoded_type) := - EncodeTP (option U) (option_tp_encodedMixin U). - -Lemma seq_tp_pf (U : encoded_type) : - @encode typeCode (tp_seq (decode U)) = seq U. -Proof. by rewrite /decode; case: U=>tp [/= dp <-]. Qed. - -Definition seq_tp_encodedMixin (U : encoded_type) := - encodeMixin typeCode (seq U) (tp_seq (decode U)) (seq_tp_pf U). -Canonical seq_tp_Encoded (U : encoded_type) := - EncodeTP (seq U) (seq_tp_encodedMixin U). - -Lemma prod_tp_pf (U1 U2 : encoded_type) : - @encode typeCode (tp_prod (decode U1) (decode U2)) = prod U1 U2. -Proof. by rewrite /decode; case: U1 U2=>t1 [/= ? <-] _ [t2 [/= ? <-]]. Qed. - -Definition prod_tp_encodedMixin (U1 U2 : encoded_type) := - encodeMixinTP (prod U1 U2) - (tp_prod (decode U1) (decode U2)) (prod_tp_pf U1 U2). -Canonical prod_tp_Encoded (U1 U2 : encoded_type) := - EncodeTP (prod U1 U2) (prod_tp_encodedMixin U1 U2). - -(* this one is kinda pointless since we can't trigger inference of lambda's *) -(* so functions will have to be decorated with their types by hand *) - -Lemma arr_tp_pf (U1 U2 : encoded_type) : - @encode typeCode (tp_arr (decode U1) (decode U2)) = (U1 -> U2). -Proof. by rewrite /decode; case: U1 U2=>t1 [/= ? <-] _ [t2 [/= ? <-]]. Qed. - -Definition arr_tp_encodedMixin (U1 U2 : encoded_type) := - encodeMixinTP (U1 -> U2) (tp_arr (decode U1) (decode U2)) (arr_tp_pf U1 U2). -Canonical arr_tp_Encoded (U1 U2 : encoded_type) := - EncodeTP (U1 -> U2) (arr_tp_encodedMixin U1 U2). - -Lemma array_tp_pf (n : nat) (U : encoded_pcm) : - @encode typeCode (tp_array n (decode U)) = {ffun 'I_n -> U}. -Proof. by rewrite /decode; case: U=>t [/= d <- _]. Qed. - -Definition array_tp_encodedMixin n (U : encoded_pcm) := - encodeMixinTP {ffun 'I_n -> U} - (tp_array n (decode U)) (array_tp_pf n U). -Canonical array_tp_Encoded n (U : encoded_pcm) := - EncodeTP {ffun 'I_n -> U} (array_tp_encodedMixin n U). - -Lemma ptrmap_tp_pf (V : encoded_set) : - @encode typeCode (tp_ptrmap (decode V)) = union_map ptr_ordType V. -Proof. by rewrite /decode; case: V=>t [/= d <- _]. Qed. - -Definition ptrmap_tp_encodedMixin (V : encoded_set) := - encodeMixinTP (union_map ptr_ordType V) (tp_ptrmap (decode V)) - (ptrmap_tp_pf V). -Canonical ptrmap_tp_Encoded (V : encoded_set) := - EncodeTP (union_map ptr_ordType V) (ptrmap_tp_encodedMixin V). - -(* -Notation "[ 'encoded_type' 'of' v ]" := - ([encoded typeCode of v]) - (at level 0, format "[ 'encoded_type' 'of' v ]") : form_scope. -*) - -Definition type_clone v cT c b1 b2 : encoded_type := - @Encoded.clone typeCode v cT c b1 b2. - -Notation "[ 'encoded_type' 'of' v ]" := (@type_clone v _ _ idfun id) - (at level 0, format "[ 'encoded_type' 'of' v ]") : form_scope. - -(****************************) -(* inversion table for pcms *) -(****************************) - -Definition unit_pcm_encodedMixin := encodeMixinPCM unitPCM pcm_unit (erefl _). -Canonical unit_pcm_Encoded := EncodePCM unitPCM unit_pcm_encodedMixin. - -Definition nat_pcm_encodedMixin := encodeMixinPCM natPCM pcm_nat (erefl _). -Canonical nat_pcm_Encoded := EncodePCM natPCM nat_pcm_encodedMixin. - -Definition mtx_pcm_encodedMixin := - encodeMixinPCM (liftPCM mutexUnlifted) pcm_mtx (erefl _). -Canonical mtx_pcm_Encoded := - EncodePCM (liftPCM mutexUnlifted) mtx_pcm_encodedMixin. - -Definition heap_pcm_encodedMixin := encodeMixinPCM heapPCM pcm_heap (erefl _). -Canonical heap_pcm_Encoded := EncodePCM heapPCM heap_pcm_encodedMixin. - -Lemma natmap_pcm_pf (V : encoded_set) : - @encode pcmCode (pcm_natmap (decode V)) = union_mapPCM nat_ordType V. -Proof. by rewrite /decode; case: V=>_ [/= ? <-]. Qed. - -Definition natmap_pcm_encodedMixin (V : encoded_set) := - encodeMixinPCM (union_mapPCM nat_ordType V) (pcm_natmap (decode V)) - (natmap_pcm_pf V). -Canonical natmap_pcm_Encoded (V : encoded_set) := - EncodePCM (union_mapPCM nat_ordType V) (natmap_pcm_encodedMixin V). - -Lemma ptrmap_pcm_pf (V : encoded_set) : - @encode pcmCode (pcm_ptrmap (decode V)) = union_mapPCM ptr_ordType V. -Proof. by rewrite /decode; case: V=>_ [/= ? <-]. Qed. - -Definition ptrmap_pcm_encodedMixin (V : encoded_set) := - encodeMixinPCM (union_mapPCM ptr_ordType V) (pcm_ptrmap (decode V)) - (ptrmap_pcm_pf V). -(* this one overlaps with natmap, so i have to assign it by hand *) -(* it can't be canonical *) -(* this should be fixed by assigning dedicated names to these *) -(* unionmap types; e.g. specifically like natmap and ptrmap *) -Definition ptrmap_pcm_Encoded (V : encoded_set) := - EncodePCM (union_mapPCM ptr_ordType V) (ptrmap_pcm_encodedMixin V). - -Lemma prod_pcm_pf (U1 U2 : encoded_pcm) : - @encode pcmCode (pcm_prod (decode U1) (decode U2)) = prodPCM U1 U2. -Proof. by rewrite /decode; case: U1 U2=>t1 [/= ? <-] _ [t2 [/= ? <-]]. Qed. - -Definition prod_pcm_encodedMixin (U1 U2 : encoded_pcm) := - encodeMixinPCM (prodPCM U1 U2) - (pcm_prod (decode U1) (decode U2)) (prod_pcm_pf U1 U2). -Canonical prod_pcm_Encoded (U1 U2 : encoded_pcm) := - EncodePCM (prodPCM U1 U2) (prod_pcm_encodedMixin U1 U2). - -(* -Notation "[ 'encoded_pcm' 'of' v ]" := - ([encoded pcmCode of [pcm of v]]) - (at level 0, format "[ 'encoded_pcm' 'of' v ]") : form_scope. -*) - -Definition pcm_clone v cT c b1 b2 : encoded_pcm := - @Encoded.clone pcmCode v cT c b1 b2. - -Notation "[ 'encoded_pcm' 'of' v ]" := (@pcm_clone [pcm of v] _ _ idfun id) - (at level 0, format "[ 'encoded_pcm' 'of' v ]") : form_scope. - - -(*****************************) -(* inversion table for spcms *) -(*****************************) - -Definition unit_spcm_encodedMixin := - encodeMixinSPCM unitSPCM spcm_unit (erefl _). -Canonical unit_spcm_Encoded := EncodeSPCM unitSPCM unit_spcm_encodedMixin. - -Definition nat_spcm_encodedMixin := encodeMixinSPCM natSPCM spcm_nat (erefl _). -Canonical nat_spcm_Encoded := EncodeSPCM natSPCM nat_spcm_encodedMixin. - -Lemma natmap_spcm_pf (V : encoded_set) : - @encode spcmCode (spcm_natmap (decode V)) = natmapSPCM V. -Proof. by rewrite /decode; case: V=>_ [/= ? <-]. Qed. - -Definition natmap_spcm_encodedMixin (V : encoded_set) := - encodeMixinSPCM (natmapSPCM V) (spcm_natmap (decode V)) - (natmap_spcm_pf V). -Canonical natmap_spcm_Encoded (V : encoded_set) := - EncodeSPCM (natmapSPCM V) (natmap_spcm_encodedMixin V). - -Lemma ptrmap_spcm_pf (V : encoded_set) : - @encode spcmCode (spcm_ptrmap (decode V)) = ptrmapSPCM V. -Proof. by rewrite /decode; case: V=>_ [/= ? <-]. Qed. - -Definition ptrmap_spcm_encodedMixin (V : encoded_set) := - encodeMixinSPCM (ptrmapSPCM V) (spcm_ptrmap (decode V)) - (ptrmap_spcm_pf V). -Canonical ptrmap_spcm_Encoded (V : encoded_set) := - EncodeSPCM (ptrmapSPCM V) (ptrmap_spcm_encodedMixin V). - - -Lemma prod_spcm_pf (U1 U2 : encoded_spcm) : - @encode spcmCode (spcm_prod (decode U1) (decode U2)) = prodSPCM U1 U2. -Proof. by rewrite /decode; case: U1 U2=>t1 [/= ? <-] _ [t2 [/= ? <-]]. Qed. - -Definition prod_spcm_encodedMixin (U1 U2 : encoded_spcm) := - encodeMixinSPCM (prodSPCM U1 U2) - (spcm_prod (decode U1) (decode U2)) (prod_spcm_pf U1 U2). -Canonical prod_spcm_Encoded (U1 U2 : encoded_spcm) := - EncodeSPCM (prodSPCM U1 U2) (prod_spcm_encodedMixin U1 U2). - -(* -Notation "[ 'encoded_spcm' 'of' v ]" := - ([encoded spcmCode of [spcm of v]]) - (at level 0, format "[ 'encoded_spcm' 'of' v ]") : form_scope. -*) - -Definition spcm_clone v cT c b1 b2 : encoded_spcm := - @Encoded.clone spcmCode v cT c b1 b2. - -Notation "[ 'encoded_spcm' 'of' v ]" := (@spcm_clone [spcm of v] _ _ idfun id) - (at level 0, format "[ 'encoded_spcm' 'of' v ]") : form_scope. - -(* -Print Canonical Projections. -Check [encoded pcmCode of natPCM]. -Check [encoded_spcm of nat]. -Check [encoded_set of fset _]. -Check [encoded_pcm of heap]. -Check [encoded_set of nat]. -Check [encoded typeCode of unit]. -Check (@code_dyn pcmCode _ 3). -Check (@code_dyn typeCode _ [:: 3]). -Check (@code_dyn pcmCode _ (3, 4)). -Check [encoded_pcm of nat]. -Check [encoded_type of nat -> nat]. -Check [encoded_type of (union_map ptr_ordType [encoded_set of nat])]. - -Print Canonical Projections. - -Check [encoded_pcm of union_map nat_ordType [encoded_set of bool]]. -*) - - diff --git a/Heaps/descr b/Heaps/descr deleted file mode 100644 index 5ecfca0..0000000 --- a/Heaps/descr +++ /dev/null @@ -1 +0,0 @@ -A theory of partial finite heaps diff --git a/Heaps/finmap.v b/Heaps/finmap.v deleted file mode 100644 index 31435c9..0000000 --- a/Heaps/finmap.v +++ /dev/null @@ -1,1335 +0,0 @@ -From mathcomp.ssreflect -Require Import ssreflect ssrbool ssrnat eqtype ssrfun seq fintype. -From mathcomp -Require Import path. -From DiSeL.Heaps -Require Import pred prelude 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). - -Structure finMap : Type := 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. - -(* Forward recursion principle *) -Definition fmap_rect' (P : fmap -> Type) : - 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 _ _). -- by apply: disj_true; case: allP E. -move: E; rewrite all_predC; move/negbFE. -by case: hasPx=>// x H1 H2 _; apply: disj_false H1 H2. -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. - -(* heh, a theory of submaps would be good here *) -(* but i don't have time for a decent development *) -(* so let's do a quick lemma that's needed for feaps *) -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. -by 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 zipC f1 f2 : zip f1 f2 = zip f2 f1. -Proof. -case: f1 f2=>s1 H1 [s2 H2] /=. -move: {-}(erefl (zip' s1 s2)) (zip_sorted' H1) (zip_sorted' H2). -move: {-1}(zip' s1 s2)=>s; rewrite zipC' => <-{s}. -case: (zip' s2 s1)=>// a qf1 qf2. -by congr (Some (FinMap _)); apply: proof_irrelevance. -Qed. - -Lemma zipA f1 f2 f3 : - obind (zip f1) (zip f2 f3) = obind (zip^~ f3) (zip f1 f2). -Proof. -case: f1 f2 f3=>s1 H1 [s2 H2] [s3 H3] /=. -move: (zip_sorted' _) (zip_sorted' _). -case E1: (zip' s2 s3)=>[w1|]; -case E2: (zip' s1 s2)=>[w3|] //= qf1 qf2; last first. -- by move: (zipA' s1 s2 s3) (zip_sorted' _); rewrite E1 E2 /= => <-. -- by move: (zipA' s1 s2 s3) (zip_sorted' _); rewrite E1 E2 /= => ->. -move: (zipA' s1 s2 s3) (zip_sorted' _) (zip_sorted' _). -rewrite E1 E2 /= => ->; case: (zip' w3 s3)=>// s rf1 rf2. -by congr (Some (FinMap _)); apply: proof_irrelevance. -Qed. - -Lemma zip_unitL f : zip (zip_unit f) f = Some f. -Proof. -case: f=>s H /=; move: (zip_sorted' _); rewrite zip_unitL' => pf. -by congr (Some (FinMap _)); apply: proof_irrelevance. -Qed. - -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. - - -Section Experiments. - -Variable mtype : Type. -Definition soup := {finMap nat -> mtype}. - -Structure cohPred := Coh { - (* Validity predicate *) - (* isValid : soup -> Prop; *) - }. - - -Structure protocol := Protocol { - label : nat; - Lstate : Type; - (* tagType : eqType; *) - coh : cohPred; -}. - - -Definition blah := {finMap nat -> protocol}. - - -End Experiments. - - diff --git a/Heaps/heap.v b/Heaps/heap.v deleted file mode 100644 index 84feb72..0000000 --- a/Heaps/heap.v +++ /dev/null @@ -1,755 +0,0 @@ -From mathcomp.ssreflect -Require Import ssreflect ssrbool ssrnat eqtype ssrfun seq. -From mathcomp -Require Import path div. -Require Import Eqdep. -From DiSeL.Heaps -Require Import pred idynamic 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. - -Local Notation idyn v := (@idyn _ id _ v). - -Inductive heap := - Undef | Def (finmap : {finMap ptr -> idynamic id}) of - null \notin supp finmap. - -Section NullLemmas. -Variables (f g : {finMap ptr -> idynamic id}) (x : ptr) (v : idynamic 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. -move: pf2; rewrite {f2}E in pf1 *; move=>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 : pred ptr := - if h is Def f _ then [fun x => x \in supp f] else pred0. - -Definition dom_eq h1 h2 := - (def h1 == def 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 keys_of h : seq ptr := - if h is Def f _ then supp f else [::]. - -Local Notation base := - (@UM.base ptr_ordType (idynamic 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 ptsE k (v : idynamic id) : pts k v = to (@UM.pts _ _ _ k v). -Proof. -by rewrite /pts /UM.pts /UM.upd /=; case: decP=>// H; rewrite heapE. -Qed. - -Lemma keys_ofE f : keys_of f = UM.keys_of (from f). -Proof. by case: f. Qed. - -Module Exports. - -(* the inheritance from PCMs *) - -Definition heapUMCMixin := - UMCMixin ftE tfE defE undefE emptyE updE domE dom_eqE freeE - findE unionE umfiltE empbE ptsE keys_ofE. -Canonical heapUMC := Eval hnf in UMC heap heapUMCMixin. - -(* Can I get heap's own PCM? *) - -Definition heapPCMMixin := union_map_classPCMMixin heapUMC. -Canonical heapPCM := Eval hnf in PCM heap heapPCMMixin. - -End Exports. - -End Heap. - -Export Heap.Exports. - -Notation heap := Heap.heap. - -Definition heap_pts A (x : ptr) (v : A) := - @UMC.pts _ _ heapUMC x (@idyn _ id _ v). - -(* -Notation "x :-> v" := (@heap_pts _ x v) (at level 30). -*) - -(* -Notation "x :-> v" := (@UMC.pts heapUMC x (idyn v)) (at level 30). -*) - -(* The above notation is not getting printed *) -(* I think something is overriding it; possibly *) -(* that it's head is the same as in other notations *) -(* namely, UMC.pts, and it's only types that differentiate *) -(* which notation should be printed *) - -(* if I introduce an explicit definition to override any other *) -(* conflicting notation, it is printed *) -Notation "x :-> v" := (@heap_pts _ x v) (at level 30). - -(*****************************) -(* Specific points-to lemmas *) -(*****************************) - -Section HeapPointsToLemmas. -Implicit Types (x : ptr) (h : heap). -(* A: upd requires explicit annotation; change that *) -(* why is notation not being displayed? *) -Local Notation idyn v := (@idyn _ id _ v). - -Lemma hptsU A x (v : A) : x :-> v = upd x (idyn v) (Unit : heap). -Proof. exact: gen_ptsU. Qed. - -Lemma hdomPt A x (v : A) : dom (x :-> v) =i [pred y | x != null & x == y]. -Proof. exact: gen_domPt. Qed. - -Lemma hvalidPt A x (v : A) : valid (x :-> v) = (x != null). -Proof. exact: gen_validPt. Qed. - -Lemma hfindPt A x (v : A) : - find x (x :-> v) = if x != null then Some (idyn v) else None. -Proof. exact: gen_findPt. Qed. - -Lemma hfindPt2 A x1 x2 (v : A) : - find x1 (x2 :-> v) = - if (x1 == x2) && (x2 != null) then Some (idyn v) else None. -Proof. exact: gen_findPt2. Qed. - -Lemma hfreePt A x (v : A) : x != null -> free x (x :-> v) = Unit. -Proof. by move=>N; apply: gen_freePt. Qed. - -Lemma hfreePt2 A x1 x2 (v : A) : - x2 != null -> - free x1 (x2 :-> v) = if x1 == x2 then Unit else x2 :-> v. -Proof. by move=>N; apply: gen_freePt2. Qed. - -Lemma hdomeqPt A1 A2 x (v1 : A1) (v2 : A2) : - dom_eq (x :-> v1) (x :-> v2). -Proof. exact: gen_domeqPt. Qed. - -Lemma hcancelPt A1 A2 x (v1 : A1) (v2 : A2) : - valid (x :-> v1) -> x :-> v1 = x :-> v2 -> idyn v1 = idyn v2. -Proof. exact: gen_cancelPt. Qed. - -Lemma hcancelPtT A1 A2 x (v1 : A1) (v2 : A2) : - valid (x :-> v1) -> x :-> v1 = x :-> v2 -> A1 = A2. -Proof. by move=>V; move/(hcancelPt V); move/idyn_injT. Qed. - -Lemma hcancelPtV A x (v1 v2 : A) : - valid (x :-> v1) -> x :-> v1 = x :-> v2 -> v1 = v2. -Proof. by move=>V; move/(hcancelPt V)/idyn_inj. Qed. - -Lemma hupdPt A1 A2 x (v1 : A1) (v2 : A2) : - upd x (idyn v1) (x :-> v2) = x :-> v1. -Proof. exact: gen_updPt. Qed. - -Lemma hempbPt A x (v : A) : empb (x :-> v) = false. -Proof. exact: gen_empbPt. Qed. - -(* valid *) - -Lemma hvalidPtUn A x (v : A) h : - valid (x :-> v \+ h) = [&& x != null, valid h & x \notin dom h]. -Proof. exact: gen_validPtUn. Qed. - -Lemma hvalidPt_cond A x (v : A) h : valid (x :-> v \+ h) -> x != null. -Proof. exact: gen_validPt_cond. Qed. - -Lemma hvalidPtV A x (v : A) h : valid (x :-> v \+ h) -> valid h. -Proof. exact: gen_validPtV. Qed. - -Lemma hvalidPtD A x (v : A) h : valid (x :-> v \+ h) -> x \notin dom h. -Proof. exact: gen_validPtD. Qed. - -Prenex Implicits hvalidPtD. - -(* dom *) - -Lemma hdomPtUn A x (v : A) h : - dom (x :-> v \+ h) =i - [pred y | valid (x :-> v \+ h) & (x == y) || (y \in dom h)]. -Proof. exact: gen_domPtUn. Qed. - -(* find *) - -Lemma hfindPtUn A x (v : A) h : - valid (x :-> v \+ h) -> find x (x :-> v \+ h) = Some (idyn v). -Proof. exact: gen_findPtUn. Qed. - -Lemma hfindPtUn2 A x1 x2 (v : A) h : - valid (x2 :-> v \+ h) -> - find x1 (x2 :-> v \+ h) = if x1 == x2 then Some (idyn v) else find x1 h. -Proof. exact: gen_findPtUn2. Qed. - -(* free *) - -Lemma hfreePtUn A x (v : A) h : - valid (x :-> v \+ h) -> free x (x :-> v \+ h) = h. -Proof. exact: gen_freePtUn. Qed. - -Lemma hfreePtUn2 A x1 x2 (v : A) h : - valid (x2 :-> v \+ h) -> - free x1 (x2 :-> v \+ h) = if x1 == x2 then h else x2 :-> v \+ free x1 h. -Proof. exact: gen_freePtUn2. Qed. - -(* upd *) - -Lemma hupdPtUn A1 A2 x (v1 : A1) (v2 : A2) h : - upd x (idyn v1) (x :-> v2 \+ h) = x :-> v1 \+ h. -Proof. exact: gen_updPtUn. 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/gen_eta; case=>A v H; exists A, v. Qed. - -(* to get rid of pesky idyn *) -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 hcancel A1 A2 x (v1 : A1) (v2 : A2) h1 h2 : - valid (x :-> v1 \+ h1) -> - x :-> v1 \+ h1 = x :-> v2 \+ h2 -> - [/\ idyn v1 = idyn v2, valid h1 & h1 = h2]. -Proof. exact: gen_cancel. 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/(hcancel V); move/idyn_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/(hcancel V); move/idyn_inj. Qed. - -Lemma hcancel2 A1 A2 x1 x2 (v1 : A1) (v2 : A2) h1 h2 : - valid (x1 :-> v1 \+ h1) -> - x1 :-> v1 \+ h1 = x2 :-> v2 \+ h2 -> - if x1 == x2 then idyn v1 = idyn v2 /\ h1 = h2 - else [/\ free x1 h2 = free x2 h1, - h1 = x2 :-> v2 \+ free x1 h2 & - h2 = x1 :-> v1 \+ free x2 h1]. -Proof. exact: gen_cancel2. 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 /(gen_cancel2 V); case: ifP=>// _ [/idyn_inj]. Qed. - -Lemma humfiltPt A p x (v : A) : - um_filter p (x :-> v) = - if p x then x :-> v else if x != null then Unit else um_undef. -Proof. exact: gen_umfiltPt. Qed. - -End HeapPointsToLemmas. - -Prenex Implicits hvalidPt_cond 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 -topredE /=. -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 hdomPtUn 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. - -Lemma dom_notin_notin (h1 h2 : heap) x : - valid (h1 \+ h2) -> x \notin dom (h1 \+ h2) -> x \notin dom h1. -Proof. by move=>D; rewrite domUn inE /= negb_and negb_or /= D; case/andP. Qed. - -Lemma dom_in_notin (h1 h2 : heap) x : - valid (h1 \+ h2) -> x \in dom h1 -> x \notin dom h2. -Proof. by case: validUn=>// D1 D2 H _; apply: H. 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 hdomPtUn 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 hvalidPtUn 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; apply: updiVm'. -case: xs=>[|v xs]; rewrite ptr0 ?andbF ?andbT //=. -by rewrite hdomPtUn 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 /= hdomPtUn inE /= eq_refl andbT -updiS updiD H. -rewrite -addn1 addnC -ptrA updiS hdomPtUn 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 hempbPt. -by case/(hcancelV D)=><- {D} D; move/(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. - -(*************************************) -(* 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 hvalidPtUn. Qed. - -Implicit Arguments defPtUnO [A h 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: hvalidPt_cond. Qed. - -Implicit 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: hvalidPtV. Qed. - -Implicit Arguments defPt_defO [A h 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: hvalidPtD. Qed. - -Implicit Arguments defPt_domO [A h 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: hdomPtUn. Qed. - -Implicit Arguments domPtUnO [A h v f]. - -Lemma lookPtUnO A h x (v : A) (f : partition (x :-> v) h) : - valid (untag f) -> find x (untag f) = Some (@idyn _ id _ v). -Proof. by rewrite partitionE; apply: hfindPtUn. Qed. - -Implicit 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: hfreePtUn. Qed. - -Implicit 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 _ id _ v2) (untag (f (x :-> v1))) = f (x :-> v2). -Proof. by rewrite !partitionE; apply: hupdPtUn. Qed. - -Implicit Arguments updPtUnO [A1 A2 x i v1 v2]. - -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. - -Implicit Arguments cancelTO [A1 A2 h1 h2 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. - -Implicit Arguments cancelO [A h1 h2 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. - - diff --git a/Heaps/idynamic.v b/Heaps/idynamic.v deleted file mode 100644 index 28bc424..0000000 --- a/Heaps/idynamic.v +++ /dev/null @@ -1,76 +0,0 @@ -From mathcomp.ssreflect -Require Import ssreflect ssrfun. -Require Import Eqdep. -From DiSeL.Heaps -Require Import prelude. -Set Implicit Arguments. -Unset Strict Implicit. -Unset Printing Implicit Defensive. - -(*************************************************************************) -(* A new variant of dynamic type, which doesn't directly quantify over a *) -(* type. Rather, we quantify over an index I and a function sort : I -> *) -(* Type, which determines the type. By choosing I = Type and sort = id, *) -(* we recover the old dynamic. *) -(* *) -(* The new dynamic makes it possible to produce small dynamic types by *) -(* choosing the index I to be the type of codes. Hence, this approach is *) -(* more flexible when compared to the old type dynamic. *) -(* Also, it seems we don't need jmeq on pairs *) -(*************************************************************************) - -(* Eventually get rid of the old dynamic, and work with this new type *) -(* But for now, to avoid conflict, I'll call this one idynamic, for *) -(* indexed dynamic *) - -Section IndexedDynamic. -Variable (I : Type) (sort : I -> Type). - -Inductive idynamic := idyn (A : I) of sort A. - -Definition idyn_tp (x : idynamic) : I := let: idyn A _ := x in A. - -Definition idyn_val (x : idynamic) : sort (idyn_tp x) := - let: idyn _ v := x in v. - -Lemma idyn_eta (x : idynamic) : idyn (idyn_val x) = x. -Proof. by case: x. Qed. - -Lemma idyn_injT (A1 A2 : I) (v1 : sort A1) (v2 : sort A2) : - idyn v1 = idyn v2 -> A1 = A2. -Proof. by case. Qed. - -Lemma idyn_inj (A : I) (v1 v2 : sort A) : idyn v1 = idyn v2 -> v1 = v2. -Proof. by case=>H; apply: inj_pairT2 H. Qed. - -Definition icoerce A B (v : sort A) : A = B -> sort B := - [eta eq_rect A [eta sort] v B]. - -Lemma ieqc A (v : sort A) (pf : A = A) : icoerce v pf = v. -Proof. by move: pf; apply: Streicher_K. Qed. - -Definition ijmeq A B (v1 : sort A) (v2 : sort B) := - forall pf, icoerce v1 pf = v2. - -Lemma ijmeq_refl A (v : sort A) : ijmeq v v. -Proof. by move=>pf; rewrite ieqc. Qed. - -Hint Resolve ijmeq_refl. - -Lemma ijmE A (v1 v2 : sort A) : ijmeq v1 v2 <-> v1 = v2. -Proof. by split=>[|->]; first by move/(_ (erefl _)). Qed. - -Lemma idynE (A1 A2 : I) (v1 : sort A1) (v2 : sort A2) : - A1 = A2 -> ijmeq v1 v2 -> idyn v1 = idyn v2. -Proof. by move=>E; rewrite -{A2}E in v2 *; move/ijmE=>->. Qed. - -End IndexedDynamic. - -Prenex Implicits idyn_tp idyn_val idyn_injT idyn_inj. - -Implicit Arguments icoerce [I A B]. - -Hint Resolve ijmeq_refl. - - - diff --git a/Heaps/opam b/Heaps/opam deleted file mode 100644 index b8a6574..0000000 --- a/Heaps/opam +++ /dev/null @@ -1,28 +0,0 @@ -opam-version: "1.2" -name: "disel-heaps" -version: "dev" -maintainer: "palmskog@gmail.com" - -homepage: "https://github.com/DistributedComponents/disel" -dev-repo: "https://github.com/DistributedComponents/disel.git" -bug-reports: "https://github.com/DistributedComponents/disel/issues" -license: "BSD2" - -build: [ make "-j" "%{jobs}%" ] -install: [ make "install" ] -remove: [ "sh" "-c" "rm -rf '%{lib}%/coq/user-contrib/DiSeL'" ] -depends: [ - "coq" { ((>= "8.7" & < "8.8~") | (= "dev")) } - "coq-mathcomp-ssreflect" { ((>= "1.6.2" & < "1.7~") | (= "dev")) } -] - -tags: [ - "category:Computer Science/Concurrent Systems and Protocols/Theory of concurrent systems" - "keyword:program verification" - "keyword:separation logic" - "keyword:distributed algorithms" -] -authors: [ - "Ilya Sergey <>" - "James R. Wilcox <>" -] diff --git a/Heaps/ordtype.v b/Heaps/ordtype.v deleted file mode 100644 index 819b64c..0000000 --- a/Heaps/ordtype.v +++ /dev/null @@ -1,253 +0,0 @@ -From mathcomp.ssreflect -Require Import ssreflect ssrbool ssrnat eqtype ssrfun seq fintype. -From mathcomp -Require Import path. -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 : forall x y, [|| x < y, x == y | y < x]. -Proof. by move=>*; 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 : forall x y, [|| ordf x y, x == y | ordf y x]. -Proof. -move=>x y; 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 : forall 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. diff --git a/Heaps/pcm.v b/Heaps/pcm.v deleted file mode 100644 index 6eb30a2..0000000 --- a/Heaps/pcm.v +++ /dev/null @@ -1,558 +0,0 @@ -From mathcomp.ssreflect -Require Import ssreflect ssrbool ssrnat eqtype ssrfun seq. -From DiSeL.Heaps -Require Import pred prelude ordtype finmap. -Set Implicit Arguments. -Unset Strict Implicit. -Unset Printing Implicit Defensive. - -(*******************************) -(* 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; _ : 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 valid := valid_op class. -Definition join := join_op class. -Definition unit := unit_op class. - -End ClassDef. - -Implicit Arguments unit [cT]. - -Definition morph_axiom (A B : type) (f : sort A -> sort B) := - f unit = unit /\ forall x y, f (join x y) = join (f x) (f y). - -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 ]") : form_scope. -Notation "[ 'pcm' 'of' T 'for' C ]" := (@clone T C _ idfun id) - (at level 0, format "[ 'pcm' 'of' T 'for' C ]") : form_scope. -Notation "[ 'pcm' 'of' T ]" := (@clone T _ _ id id) - (at level 0, format "[ 'pcm' 'of' T ]") : form_scope. - -Notation "x \+ y" := (join x y) (at level 43, left associativity). -Notation valid := valid. -Notation Unit := unit. - -Implicit Arguments unit [cT]. -Prenex Implicits join unit. - -Section Morphism. -Variables A B : pcm. - -Structure pcm_morph_type := PCMMorph { - pcm_func :> A -> B; - _ : morph_axiom pcm_func}. - -Definition pcm_morph_for of phant (A -> B) := pcm_morph_type. -Identity Coercion type_of_pcm_morph : pcm_morph_for >-> pcm_morph_type. - -End Morphism. - -Notation "{ 'pcm_morph' fT }" := (pcm_morph_for (Phant fT)) - (at level 0, format "{ 'pcm_morph' '[hv' fT ']' }") : type_scope. - -(* 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 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. - -Variable f : {pcm_morph U -> V}. - -Lemma fjoin x y : f (x \+ y) = f x \+ f y. -Proof. by case: f=>? []. Qed. - -Lemma funit : f Unit = Unit. -Proof. by case: f=>? []. 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). - -End UnfoldingRules. - -End Exports. - -End PCM. - -Export PCM.Exports. - -(* definition of precision for an arbitrary PCM U *) - -Definition precise (U : pcm) (P : Pred U) := - forall s1 s2 t1 t2, - valid (s1 \+ t1) -> - s1 \+ t1 = s2 \+ t2 -> - s1 \In P -> s2 \In P -> s1 = s2. - - -(****************************************************) -(* Sometimes we want to construct PCM's by lifting. *) -(* We need a structure which by lifting becomes PCM *) -(* This will give us uniform notation for lifting *) -(****************************************************) - -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. - -Implicit 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 *) -(**************************************************) - -Module Lift. -Section Lift. -Variable A : unlifted. - -Inductive lift : Type := Undef | Up of A. - -Let unit := Up ounit. - -Let valid := - [fun x : lift => if x is Up _ then true else false]. - -Let join := - [fun x y : lift => - if (x, y) is (Up v, Up w) then - if ojoin v w is Some u then Up u - else Undef - else Undef]. - -Lemma joinC (x y : lift) : join x y = join y x. -Proof. by case: x y=>[|x][|y] //=; rewrite ojoinC. Qed. - -Lemma joinA (x y z : lift) : 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. -End Lift. - -Notation lift A := (@Lift.lift A). -Notation undef := (@Lift.Undef _). -Notation up a := (@Lift.Up _ a). - -(* A view for pattern-matching lifted pcm's *) - -CoInductive lift_spec A : lift A -> Type := -| undef_spec : lift_spec undef -| up_spec : forall a : A, lift_spec (up a). - -Lemma liftP A (x : lift A) : lift_spec x. -Proof. by case: x=>[|a]; [apply: undef_spec | apply: up_spec]. Qed. - -Definition liftPCMMixin A := - PCMMixin (@Lift.joinC A) (@Lift.joinA A) - (@Lift.unitL A) (@Lift.validL A) (@Lift.validU A). -Canonical liftPCM A := Eval hnf in PCM (lift A) (liftPCMMixin A). - -(* simplifying up-up expressions *) - -Lemma upupE (A : unlifted) (a1 a2 : A) : - up a1 \+ up a2 = - if ojoin a1 a2 is Some a then up a else undef. -Proof. by []. Qed. - -(* We can prove that lifting preserves equality types *) - -Module LiftEqType. -Section LiftEqType. -Variable (A : eqType) (c : Unlifted.mixin_of A). - -Let U := (Unlifted A c). - -Definition lift_eq (u v : lift U) := - match u, v with - Lift.Up a, Lift.Up b => a == b - | Lift.Undef, Lift.Undef => true - | _, _ => false - end. - -Lemma lift_eqP : Equality.axiom lift_eq. -Proof. -case=>[|x][|y] /=; try by constructor. -by apply: (iffP eqP)=>[->|[]]. -Qed. - -End LiftEqType. -End LiftEqType. - -Definition liftEqMixin A c := EqMixin (@LiftEqType.lift_eqP A c). -Canonical liftEqType A c := Eval hnf in EqType _ (@liftEqMixin A c). - -(******************************) -(* Now some PCM constructions *) -(******************************) - -(* heaps are pcms, albeit large *) -(* remove this line eventually; this is here for legacy *) -(* -Definition old_heapPCMMixin := PCMMixin unC unA un0h defUnl def0. -Canonical old_heapPCM := Eval hnf in PCM heap old_heapPCMMixin. -*) - -(* nats with addition are a pcm *) - -Definition natPCMMixin := - PCMMixin addnC addnA add0n (fun x y => @id true) (erefl _). -Canonical natPCM := Eval hnf in PCM nat natPCMMixin. - -(* also with multiplication, but we don't make that one canonical *) - -Definition multPCMMixin := - PCMMixin mulnC mulnA mul1n (fun x y => @id true) (erefl _). -Definition multPCM := Eval hnf in PCM nat multPCMMixin. - -(* with max too *) - -Definition maxPCMMixin := - PCMMixin maxnC maxnA max0n (fun x y => @id true) (erefl _). -Definition maxPCM := Eval hnf in PCM nat maxPCMMixin. - -(* mutexes are an unlifted pcm and an equality type *) - -Inductive mutex := own | nown. - -Module MutexUnlift. - -Definition mutex_eq x y := - match x, y with - own, own => true - | nown, nown => true - | _, _ => false - end. - -Lemma mutex_eqP : Equality.axiom mutex_eq. -Proof. by case; case; constructor. Qed. - -Definition ojoin x y := - match x, y with - own, nown => Some own - | nown, own => Some own - | nown, nown => Some nown - | _, _ => None - end. - -Let ounit := nown. - -Lemma ojoinC x y : ojoin x y = ojoin y x. -Proof. by case: x; case: y. Qed. - -Lemma ojoinA x y z : - obind (ojoin x) (ojoin y z) = obind (ojoin^~ z) (ojoin x y). -Proof. by case: x; case: y; case: z. Qed. - -Lemma ounitL x : ojoin ounit x = Some x. -Proof. by case: x. Qed. - -End MutexUnlift. - -Definition mutexEqMixin := EqMixin MutexUnlift.mutex_eqP. -Canonical mutexEqType := Eval hnf in EqType mutex mutexEqMixin. - -Definition mutexUnliftedMixin := - UnliftedMixin MutexUnlift.ojoinC MutexUnlift.ojoinA MutexUnlift.ounitL. -Canonical mutexUnlifted := Eval hnf in Unlifted mutex mutexUnliftedMixin. - -(* some lemmas for mutexes *) - -Lemma mtx00E m1 m2 : m1 \+ m2 = up nown -> m1 = up nown /\ m2 = up nown. -Proof. by case: (liftP m1)=>//; case: (liftP m2)=>//; case; case. Qed. - -(* pairs of pcms are 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). - -(* 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. - -(* 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. - -(* finite maps with disjoint union are a pcm *) - -Module FinMapPCM. -Section FinMapPCM. -Variables (K : ordType) (V : Type). - -Inductive finmap := Undef | Def of {finMap K -> V}. - -Definition valid (f : finmap) := - if f is Def _ then true else false. - -Definition join (f1 f2 : finmap) := - if (f1, f2) is (Def m1, Def m2) then - if disj m1 m2 then Def (fcat m1 m2) - else Undef - else Undef. - -Definition unit := Def (@nil K V). - -Lemma joinC f1 f2 : join f1 f2 = join f2 f1. -Proof. -case: f1 f2=>[|m1][|m2] //; rewrite /join. -by case: ifP=>E; rewrite disjC E // fcatC. -Qed. - -Lemma joinCA f1 f2 f3 : join f1 (join f2 f3) = join f2 (join f1 f3). -Proof. -case: f1 f2 f3=>[|m1][|m2][|m3] //. -rewrite /join; case E1: (disj m2 m3); last first. -- by case E2: (disj m1 m3)=>//; rewrite disj_fcat E1 andbF. -rewrite disj_fcat andbC. -case E2: (disj m1 m3)=>//; rewrite disj_fcat (disjC m2) E1 /= andbT. -by case E3: (disj m1 m2)=>//; rewrite fcatAC // E1 E2 E3. -Qed. - -Lemma joinA f1 f2 f3 : join f1 (join f2 f3) = join (join f1 f2) f3. -Proof. by rewrite (joinC f2) joinCA joinC. Qed. - -Lemma validL f1 f2 : valid (join f1 f2) -> valid f1. -Proof. by case: f1. Qed. - -Lemma unitL f : join unit f = f. -Proof. by case: f=>[//|m]; rewrite /join/unit disjC disj_nil fcat0s. Qed. - -Lemma validU : valid unit. -Proof. by []. Qed. - -End FinMapPCM. - -Module Exports. -Notation finmap := finmap. - -Section Exports. -Variables (K : ordType) (V : Type). - -Definition finmapPCMMixin := - PCMMixin (@joinC K V) (@joinA K V) (@unitL K V) (@validL K V) (@validU K V). -Canonical finmapPCM := Eval hnf in PCM (finmap K V) finmapPCMMixin. - -End Exports. -End Exports. -End FinMapPCM. - -Export FinMapPCM.Exports. - - - diff --git a/Heaps/pperm.v b/Heaps/pperm.v deleted file mode 100644 index 25e89d9..0000000 --- a/Heaps/pperm.v +++ /dev/null @@ -1,298 +0,0 @@ -From mathcomp.ssreflect -Require Import ssreflect ssrfun ssrbool ssrnat seq eqtype. -From mathcomp -Require Import path. -From DiSeL.Heaps -Require Import pred. -Set Implicit Arguments. -Unset Strict Implicit. -Import Prenex Implicits. - -(****************************************************) -(* A theory of permutations over non-equality types *) -(****************************************************) - -Section Permutations. -Variable A : Type. - -Lemma in_split (x : A) (s : seq A) : - x \In s -> exists s1, exists s2, s = s1 ++ x :: s2. -Proof. -elim:s=>[|y s IH] //=; rewrite InE. -case=>[<-|]; first by exists [::]; exists s. -by case/IH=>s1 [s2] ->; exists (y :: s1); exists s2. -Qed. - -Inductive perm (s1 s2 : seq A) : Prop := -| permutation_nil of s1 = [::] & s2 = [::] -| permutation_skip x t1 t2 of s1 = x :: t1 & s2 = x :: t2 & perm t1 t2 -| permutation_swap x y t of s1 = x :: y :: t & s2 = y :: x :: t -| permutation_trans t of perm s1 t & perm t s2. - -Lemma perm_nil (s : seq A) : perm [::] s <-> s = [::]. -Proof. -split=>[H|]; last by move=>->; apply: permutation_nil. -move: {1 2}[::] s H (erefl (Nil A)). -apply: perm_ind=>[|s1 s2 x t1 t2 ->|s1 s2 x y t ->|s1 s2 t _ IH1 _ IH2] //. -by move/IH1; move/IH2. -Qed. - -Lemma perm_refl (s : seq A) : perm s s. -Proof. -elim:s=>[|e s IH]; first by apply: permutation_nil. -by apply: (permutation_skip (x:=e)) IH. -Qed. - -Hint Resolve perm_refl. - -Lemma perm_sym s1 s2 : perm s1 s2 <-> perm s2 s1. -Proof. -suff L: forall s1 s2, perm s1 s2 -> perm s2 s1 by split; apply: L. -apply: perm_ind=>s1' s2'. -- by move=>->->; apply: permutation_nil. -- by move=>x t1 t2 -> -> H1; apply: permutation_skip. -- by move=>x y t -> ->; apply: permutation_swap. -by move=>t _ H1 _ H2; apply: permutation_trans H2 H1. -Qed. - -Lemma perm_trans s2 s1 s3 : perm s1 s2 -> perm s2 s3 -> perm s1 s3. -Proof. by apply: permutation_trans. Qed. - -Lemma perm_in s1 s2 x : perm s1 s2 -> x \In s1 -> x \In s2. -Proof. -move: s1 s2; apply: perm_ind=>s1 s2. -- by move=>->->. -- move=>y t1 t2 -> -> H; rewrite !InE; tauto. -- by move=>y z t -> ->; rewrite !InE; tauto. -by move=>t _ IH1 _ IH2; move/IH1; move/IH2. -Qed. - -Lemma perm_cat2lL s s1 s2 : perm s1 s2 -> perm (s ++ s1) (s ++ s2). -Proof. by elim:s=>[|e s IH] //=; move/IH; apply: permutation_skip. Qed. - -Lemma perm_cat2rL s s1 s2 : perm s1 s2 -> perm (s1 ++ s) (s2 ++ s). -Proof. -move=>H; move: s1 s2 H s; apply: perm_ind=>s1 s2. -- by move=>->->. -- by move=>x t1 t2 -> -> H IH s /=; apply: permutation_skip. -- by move=>x y t -> -> s /=; apply: permutation_swap. -by move=>t H1 IH1 H2 IH2 s; apply: permutation_trans (IH2 s). -Qed. - -Lemma perm_catL s1 t1 s2 t2 : - perm s1 s2 -> perm t1 t2 -> perm (s1 ++ t1) (s2 ++ t2). -Proof. -move=>H; move: s1 s2 H t1 t2; apply: perm_ind=>s1 s2. -- by move=>->->. -- move=>x t1 t2 -> -> H IH r1 r2. - by move/IH; apply: permutation_skip. -- move=>x y t -> -> t1 t2 H. - by apply: (permutation_trans (t:=[:: x, y & t] ++ t2)); - [apply: perm_cat2lL | simpl; apply: permutation_swap]. -move=>t H1 IH1 H2 IH2 t1 t2 H. -by apply: permutation_trans (IH2 _ _ H); apply: IH1. -Qed. - -Lemma perm_cat_consL s1 t1 s2 t2 x : - perm s1 s2 -> perm t1 t2 -> perm (s1 ++ x :: t1) (s2 ++ x :: t2). -Proof. -by move=>H1 H2; apply: perm_catL H1 _; apply: permutation_skip H2. -Qed. - -Lemma perm_catC s1 s2 : perm (s1 ++ s2) (s2 ++ s1). -Proof. -elim:s1 s2=>[|x s1 IH1] s2 /=; first by rewrite cats0. -apply: (@perm_trans (x::s2++s1)); first by apply: permutation_skip (IH1 s2). -elim: s2=>[|y s2 IH2] //=. -apply: (@perm_trans (y::x::s2++s1)); first by apply: permutation_swap. -by apply: permutation_skip IH2. -Qed. - -Hint Resolve perm_catC. - -Lemma perm_cons_catCA s1 s2 x : perm (x :: s1 ++ s2) (s1 ++ x :: s2). -Proof. -rewrite -cat_rcons -cats1 -cat_cons -cat1s. -by apply: perm_cat2rL; apply: perm_catC. -Qed. - -Lemma perm_cons_catAC s1 s2 x : perm (s1 ++ x :: s2) (x :: s1 ++ s2). -Proof. by apply/perm_sym; apply: perm_cons_catCA. Qed. - -Hint Resolve perm_cons_catCA perm_cons_catAC. - -Lemma perm_cons_cat_consL s1 s2 s x : - perm s (s1 ++ s2) -> perm (x :: s) (s1 ++ x :: s2). -Proof. -case: s1=>[|a s1] /= H; first by apply: permutation_skip. -apply: (@perm_trans (x::a::s1++s2)); first by apply: permutation_skip. -apply: (@perm_trans (a::x::s1++s2)); first by apply: permutation_swap. -by apply: permutation_skip=>//; apply: perm_catCA. -Qed. - -(* a somewhat generalized induction principle *) -Lemma perm_ind2 (P : seq A -> seq A -> Prop) : - P [::] [::] -> - (forall x s1 s2, perm s1 s2 -> P s1 s2 -> - P (x :: s1) (x :: s2)) -> - (forall x y s1 s2, perm s1 s2 -> P s1 s2 -> - P (y :: x :: s1) (x :: y :: s2)) -> - (forall s2 s1 s3, perm s1 s2 -> P s1 s2 -> - perm s2 s3 -> P s2 s3 -> P s1 s3) -> - forall s1 s2, perm s1 s2 -> P s1 s2. -Proof. -move=>H1 H2 H3 H4; apply: perm_ind=>s1 s2; last 1 first. -- by move=>t; apply: H4. -- by move=>->->. -- by move=>x t1 t2 -> ->; apply: H2. -move=>x y t -> ->. -have R : forall t, P t t by elim=>[|e t1 IH] //; apply: H2. -apply: (H4 (y :: x :: t))=>//; last by apply: H3. -by apply: permutation_swap. -Qed. - -Lemma perm_size l1 l2: perm l1 l2 -> size l1 = size l2. -Proof. -move: l1 l2; apply: perm_ind=>s1 s2=>[|x t1 t2|x y t|t]; first by move=>->->. -- by move=>->->{s1 s2} _ /=->. -- by move=>->->{s1 s2}/=. -- by move=>_->_->. -Qed. - -(* Now the hard part; the opposite implications *) -Lemma perm_cat_consR s1 t1 s2 t2 x : - perm (s1 ++ x :: t1) (s2 ++ x :: t2) -> perm (s1 ++ t1) (s2 ++ t2). -Proof. -move: s1 t1 s2 t2 x. -suff H: - forall r1 r2, perm r1 r2 -> forall x s1 t1 s2 t2, - r1 = s1 ++ x :: t1 -> r2 = s2 ++ x :: t2 -> perm (s1 ++ t1) (s2 ++ t2). -- by move=>s1 t1 s2 t2 x; move/H; apply. -apply: perm_ind2; last 1 first. -- move=>s2 s1 s3 H1 IH1 H2 IH2 x r1 t1 r2 t2 E1 E2. - case: (@in_split x s2). - - apply: perm_in H1 _; rewrite E1; apply: (@perm_in (x::r1++t1))=>//. - by rewrite InE; left. - move=>s4 [s5] E; apply: (@perm_trans (s4++s5)); first by apply: IH1 E1 E. - by apply: IH2 E E2. -- by move=>x []. -- move=>x t1 t2 H IH y s1 s2 p1 p2 E1 E2. - case: s1 E1=>[|b s1] /=; case: p1 E2=>[|c p1] /= E1 E2. - - by case: E1 E2=><- <- [<-]. - - apply: (@perm_trans (p1 ++ c :: p2))=>//. - by case: E1 H=><- ->; case: E2=><- ->. - - case: E1 E2 H=><- <- [<-] ->; apply: (@perm_trans (s1 ++ x:: s2)). - by apply: perm_cons_cat_consL. - case: E1 E2 H IH=><- -> [<-] -> H IH. - by apply: permutation_skip=>//; apply: IH. -move=>x y p1 p2 H IH z s1 t1 s2 t2 E1 E2. -case: s1 E1 H IH=>[|b s1]; case: s2 E2=>[|c s2]=>/=. -- by case=><- <- [<-] <- H IH; apply: permutation_skip. -- case=><-; case: s2=>[|b s2] /=. - - by case=><- <-; case=><- H IH; apply: permutation_skip H. - case=><- -> [<-] <- H IH. - by apply: permutation_skip=>//; apply: perm_trans H _. -- case=><- <- [<-]; case: s1=>[|a s1] /=. - - by case=><- H IH; apply: permutation_skip. - by case=><- -> H IH; apply: permutation_skip=>//; apply: perm_trans H. -case=><-; case: s2=>[|a s2] /=; case: s1=>[|d s1] /=. -- by case=><- <- [<-] <- <- H IH; apply: permutation_skip. -- case=><- <- [<-] <- -> H IH. - apply: (@perm_trans (x::y::s1 ++ t1)); first by apply: permutation_swap. - by apply: permutation_skip=>//; apply: perm_trans H. -- case=><- -> [<-] <- <- H IH. - apply: (@perm_trans (y::x::s2++t2)); last by apply: permutation_swap. - by apply: permutation_skip=>//; apply: perm_trans H _. -case=><- -> [<-] <- -> H IH. -apply: (@perm_trans (x::y::s1++t1)); first by apply: permutation_swap. -by apply: permutation_skip=>//; apply: permutation_skip=>//; apply: IH. -Qed. - -Lemma perm_cons x s1 s2 : perm (x :: s1) (x :: s2) <-> perm s1 s2. -Proof. -split; last by apply: permutation_skip. -by move/(@perm_cat_consR [::] s1 [::] s2 x). -Qed. - -Lemma perm_cons_cat_cons x s1 s2 s : - perm (x :: s) (s1 ++ x :: s2) <-> perm s (s1 ++ s2). -Proof. -split=>[|H]; first by by move/(@perm_cat_consR [::] s s1 s2 x). -by apply: (@perm_trans (x :: s1++s2))=>//; apply: permutation_skip. -Qed. - -Lemma perm_cat_cons x s1 s2 t1 t2 : - perm (s1 ++ x :: t1) (s2 ++ x :: t2) <-> perm (s1 ++ t1) (s2 ++ t2). -Proof. -split=>[|H]; first by apply: perm_cat_consR. -apply: (@perm_trans (x::s1++t1))=>//; apply: (@perm_trans (x::s2++t2))=>//. -by apply/perm_cons. -Qed. - -Lemma perm_cat2l s1 s2 s3: perm (s1 ++ s2) (s1 ++ s3) <-> perm s2 s3. -Proof. -split; last by apply: perm_cat2lL. -elim: s1 s2 s3=>[|x s1 IH] s2 s3 //= H. -by apply: IH; move/perm_cons: H. -Qed. - -Lemma perm_cat2r s1 s2 s3 : perm (s2 ++ s1) (s3 ++ s1) <-> perm s2 s3. -Proof. -split; last by apply: perm_cat2rL. -elim: s1 s2 s3=>[|x s1 IH] s2 s3 /=; first by rewrite !cats0. -by move=>H; apply: IH; apply: perm_cat_consR H. -Qed. - -Lemma perm_catAC s1 s2 s3 : perm ((s1 ++ s2) ++ s3) ((s1 ++ s3) ++ s2). -Proof. by move=>*; rewrite -!catA perm_cat2l. Qed. - -Lemma perm_catCA s1 s2 s3 : perm (s1 ++ s2 ++ s3) (s2 ++ s1 ++ s3). -Proof. by move=>*; rewrite !catA perm_cat2r. Qed. - -End Permutations. - -Hint Resolve perm_refl perm_catC perm_cons_catCA - perm_cons_catAC perm_catAC perm_catCA. - - -(* perm and map *) - -Lemma perm_map A B (f : A -> B) (s1 s2 : seq A) : - perm s1 s2 -> perm (map f s1) (map f s2). -Proof. -set P := fun s1 s2 => perm (map f s1) (map f s2). -suff {s1 s2} L: forall s1 s2, perm s1 s2 -> P s1 s2 by apply: L. -apply: perm_ind2=>//; last 1 first. -- move=>s2 s1 s3 _ IH1 _ IH2. - by apply: perm_trans IH1 IH2. -- by move=>x s1 s2 H IH; rewrite /P perm_cons. -move=>x1 x2 s1 s2 H IH. -have L: perm [:: f x2, f x1 & map f s1] ([:: f x1] ++ [:: f x2 & map f s1]). -- by apply/perm_cons_cat_cons. -by apply: perm_trans L _; rewrite !perm_cons. -Qed. - -(* mapping to ssreflect decidable perm *) - -Lemma perm_eq_perm (A : eqType) (s1 s2 : seq A) : - reflect (perm s1 s2) (perm_eq s1 s2). -Proof. -apply: (iffP idP); last first. -- elim: s1 s2 /; last 1 first. - - by move=>s1 s2 t _ H2 _; apply: perm_eq_trans H2. - - by move=>s1 s2 ->->. - - by move=>s1 s2 x t1 t2 ->-> H1 H2; rewrite seq.perm_cons. - move=>s1 s2 x y t ->->. - by rewrite -(perm_rot 1) /rot /= seq.perm_cons seq.perm_catC. -elim: s1 s2=>[|x s1 IH] /= s2. -- move/perm_eq_mem=>H; apply/perm_nil. - by case: s2 H=>// a s2 /(_ a); rewrite inE eq_refl. -move=>H; move: (perm_eq_mem H x); rewrite inE eq_refl; move/esym=>/= H'. -move/(perm_eq_trans H): {H} (perm_to_rem H'); rewrite seq.perm_cons. -move/IH/(perm_cons x)=>H; apply: perm_trans H _. -elim: s2 x H'=>[|y s2 IH2] //= x. -rewrite inE; case/orP; first by move/eqP=><-; rewrite eq_refl. -case: eqP=>[->//|_ H]; move/(perm_cons y): (IH2 _ H). -by apply: perm_trans; apply: permutation_swap. -Qed. - diff --git a/Heaps/pred.v b/Heaps/pred.v deleted file mode 100644 index ebc7bb6..0000000 --- a/Heaps/pred.v +++ /dev/null @@ -1,526 +0,0 @@ -From mathcomp.ssreflect -Require Import ssreflect ssrfun ssrbool eqtype seq. -Require Import Setoid Morphisms. -Set Implicit Arguments. -Unset Strict Implicit. -Unset Printing Implicit Defensive. - -(* First some basic propositional equalities Basically, we need to repeat *) -(* most of ssrbool.v here but we'll do it as we go. *) - -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. - -Implicit Arguments Pred0 [T]. -Implicit 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 EqPred T (p1 p2 : Pred T) := EqPredType p1 p2. *) -(* Definition SubPred T (p1 p2 : Pred T) := SubPredType p1 p2. *) -Definition EqSimplPred T (p1 p2 : Simpl_Pred T) := EqPredType p1 p2. -Definition SubSimplPred T (p1 p2 : Simpl_Pred T) := SubPredType p1 p2. -(* -Definition EqMem T (p1 p2 : Mem_Pred T) := EqPredType p1 p2. -Definition SubMem T (p1 p2 : Mem_Pred T) := SubPredType p1 p2. -*) - -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 : forall x (p : pT), toPred p x = (x \In p). -Proof. by move=> *; rewrite -Mem_toPred. Qed. - -Lemma In_Simpl : forall x (p : Simpl_Pred T), (x \In p) = p x. -Proof. by []. Qed. - -Lemma Simpl_PredE : forall (p : Pred T), p <~> [Pred x | p x]. -Proof. by []. Qed. - -(* Definition InE := (In_Simpl, Simpl_PredE). (* to be extended *) *) - -Lemma Mem_Simpl : forall (p : Simpl_Pred T), Mem p = p :> Pred T. -Proof. by []. Qed. - -Definition MemE := Mem_Simpl. (* could be extended *) - -Lemma Mem_Mem : forall p : pT, (Mem (Mem p) = Mem p) * (Mem [Mem p] = Mem p). -Proof. by move=> p; 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 *) - -(* Unfortunately, Coq setoids don't seem to understand implicit coercions *) -(* and canonical structures so we have to repeat relation declarations *) -(* for all instances. This is really annoying, but at least I don't have *) -(* to reprove the lemmas on refl, sym and trans *) -(* *) -Add Parametric Relation T (pT : PredType T) : pT (@EqPredType _ pT) - reflexivity proved by (@EqPredType_refl _ _) - symmetry proved by (@EqPredType_sym _ _) - transitivity proved by (@EqPredType_trans' _ _) as EqPredType_rel. - -(* *) -(* Add Parametric Relation T (pT : PredType T) : pT (@SubPredType _ pT) *) -(* reflexivity proved by (@SubPredType_refl _ _) *) -(* transitivity proved by (@SubPredType_trans' _ _) as SubPredType_rel. *) - -(* Add Parametric Relation T : (Pred T) (@EqPred _) - reflexivity proved by (@EqPredType_refl _ _) - symmetry proved by (@EqPredType_sym _ _) - transitivity proved by (@EqPredType_trans' _ _) as EqPred_rel. - -Add Parametric Relation T : (Pred T) (@SubPred _) - reflexivity proved by (@SubPredType_refl _ _) - transitivity proved by (@SubPredType_trans' _ _) as SubPred_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. *) -(* Annoyingly, even the coercions must be declared *) - -Add Parametric Morphism T : (@Pred_of_Simpl T) with signature - @EqSimplPred _ ==> @EqPredType T (PredPredType T) as Pred_of_Simpl_morph. -Proof. by []. Qed. - -(* Do we need other coercions? We'll discover as we go *) - -(* Now the other morphisms. Again, not clear which ones are needed. *) -(* However, for all this to work, it seems that morphisms must be *) -(* declared with most specific signatures, or else the system *) -(* complains. For example, we use EqPred _ instead of EqPredType _ _, *) -(* even though the former is an instance of the later. *) - -(* -Add Parametric Morphism T : (@EqPred T) with signature - @EqPred _ ==> @EqPred _ ==> iff as EqPred_morph. -Proof. by move=>r1 s1 H1 r2 s2 H2; rewrite H1 H2. Qed. -*) - -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 : forall y s x, (x \In y :: s) <-> (x = y) \/ (x \In s). -Proof. by []. Qed. - -Lemma In_nil : forall x, (x \In [::]) <-> False. -Proof. by []. Qed. - -Lemma Mem_Seq1 : forall x y, (x \In [:: y]) <-> (x = y). -Proof. by move=> x y; rewrite In_cons orpF. Qed. - -Definition InE := (Mem_Seq1, In_cons, In_Simpl). -(* I also wanted to add Simpl_PredE, but setoid rewrite returns an error *) -(* and instead of trying the other rules in the tuple, it just stops *) -(* This is ridiculuous *) - -Lemma Mem_cat : forall x s1 s2, (x \In s1 ++ s2) <-> x \In s1 \/ x \In s2. -Proof. -move=>x; 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. - -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. - - - diff --git a/Heaps/prelude.v b/Heaps/prelude.v deleted file mode 100644 index 9941278..0000000 --- a/Heaps/prelude.v +++ /dev/null @@ -1,167 +0,0 @@ -From mathcomp.ssreflect -Require Import ssreflect ssrbool ssrnat eqtype ssrfun seq. -From DiSeL.Heaps -Require Import pred. -Require Import 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 proof_irrelevance (P : Prop) (p1 p2 : P) : p1 = p2. -Proof. by apply: ext_prop_dep_proof_irrel_cic; apply: pext. Qed. - -Lemma eta A (B : A -> Type) (f : forall x, B x) : f = [eta f]. -Proof. by apply: fext. Qed. - -Lemma ext A (B : A -> Type) (f1 f2 : forall x, B x) : - f1 = f2 -> forall x, f1 x = f2 x. -Proof. by move=>->. Qed. - -(*******************) -(* Setoid renaming *) -(*******************) - -(* Setoid library takes up some important arrow notations *) -(* used by ssreflect and elsewhere; so we must rename *) -Ltac add_morphism_tactic := SetoidTactics.add_morphism_tactic. -Notation " R ===> R' " := (@Morphisms.respectful _ _ R R') - (right associativity, at level 55) : signature_scope. - -(***********) -(* Prelude *) -(***********) - -(* often used notation definitions and lemmas that are *) -(* not included in the other libraries *) - -Definition inj_pair2 := @inj_pair2. -Implicit Arguments inj_pair2 [U P p x y]. -Prenex Implicits inj_pair2. - -Lemma inj_sval A P : injective (@sval A P). -Proof. -move=>[x Hx][y Hy] /= H; move: Hx Hy; rewrite H=>*. -congr exist; apply: proof_irrelevance. -Qed. - -Lemma svalE A (P : A -> Prop) x H : sval (exist P x H) = x. -Proof. by []. Qed. - -(* rewrite rule for propositional symmetry *) -Lemma sym A (x y : A) : x = y <-> y = x. -Proof. by []. Qed. - -(* inferrable reflexivity axiom *) -Definition erefli A (x : A) := erefl x. -Arguments erefli [A x]. - -(* selecting a list element *) -(* should really be in seq.v *) - -Section HasSelect. -Variables (A : eqType) (p : pred A). - -CoInductive has_spec (s : seq A) : bool -> Type := -| has_true x of x \in s & p x : has_spec s true -| has_false of (all (predC p) s) : has_spec s false. - -Lemma hasPx : forall s, has_spec s (has p s). -Proof. -elim=>[|x s IH] /=; first by apply: has_false. -rewrite orbC; case: IH=>/=. -- by move=>k H1; apply: has_true; rewrite inE H1 orbT. -case E: (p x)=>H; last by apply: has_false; rewrite /= E H. -by apply: has_true E; rewrite inE eq_refl. -Qed. - -End HasSelect. - -(***************************) -(* 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. - -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. - -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). - -(************************) -(* 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 ]" (at level 0, format - "'[hv' [ \/ '[' P1 , '/' P2 , '/' P3 , '/' P4 ']' '/ ' & P5 ] ']'"). -Reserved Notation "[ \/ P1 , P2 , P3 , P4 , P5 & P6 ]" (at level 0, format - "'[hv' [ \/ '[' P1 , '/' P2 , '/' P3 , '/' P4 , '/' P5 ']' '/ ' & P6 ] ']'"). - -Inductive and6 (P1 P2 P3 P4 P5 P6 : Prop) : Prop := - And6 of P1 & P2 & P3 & P4 & P5 & P6. - -Inductive 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. - -Notation "[ /\ P1 , P2 , P3 , P4 , P5 & P6 ]" := (and6 P1 P2 P3 P4 P5 P6) : 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. - -Section ReflectConnectives. - -Variable b1 b2 b3 b4 b5 b6 : 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 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. - -End ReflectConnectives. - -Implicit Arguments and6P [b1 b2 b3 b4 b5 b6]. -Implicit Arguments or5P [b1 b2 b3 b4 b5]. -Implicit Arguments or6P [b1 b2 b3 b4 b5 b6]. -Prenex Implicits and6P or5P or6P. - - diff --git a/Heaps/spcm.v b/Heaps/spcm.v deleted file mode 100644 index a41c1e7..0000000 --- a/Heaps/spcm.v +++ /dev/null @@ -1,255 +0,0 @@ -From mathcomp.ssreflect -Require Import ssreflect ssrbool ssrnat eqtype ssrfun seq. -From DiSeL.Heaps -Require Import pred prelude ordtype domain unionmap pcm heap. -Set Implicit Arguments. -Unset Strict Implicit. -Unset Printing Implicit Defensive. - -(*************************************) -(* Small Partial Commutative Monoids *) -(* can be stored into heaps without *) -(* universe inconsistencies *) -(*************************************) - -Module SPCM. - -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; _ : 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 valid := valid_op class. -Definition join := join_op class. -Definition unit := unit_op class. - -End ClassDef. - -Implicit Arguments unit [cT]. - -Definition morph_axiom (A B : type) (f : sort A -> sort B) := - f unit = unit /\ forall x y, f (join x y) = join (f x) (f y). - -Module Exports. -Coercion sort : type >-> Sortclass. -Notation spcm := type. -Notation SPCMMixin := Mixin. -Notation SPCM T m := (@pack T m). - -Notation "[ 'spcmMixin' 'of' T ]" := (class _ : mixin_of T) - (at level 0, format "[ 'spcmMixin' 'of' T ]") : form_scope. -Notation "[ 'spcm' 'of' T 'for' C ]" := (@clone T C _ idfun id) - (at level 0, format "[ 'spcm' 'of' T 'for' C ]") : form_scope. -Notation "[ 'spcm' 'of' T ]" := (@clone T _ _ id id) - (at level 0, format "[ 'spcm' 'of' T ]") : form_scope. - -Section TOPCM. -Variable U : spcm. -Local Notation tp := (sort U). - -Definition pvalid (x : tp) := valid x. -Definition pjoin (x1 x2 : tp) := (join x1 x2) : tp. -Definition punit : tp := unit. - -Lemma pjoinC (x y : tp) : pjoin x y = pjoin y x. -Proof. by rewrite /pjoin; case: U x y=>T [v j u c *]; apply: c. Qed. - -Lemma pjoinA (x y z : tp) : pjoin x (pjoin y z) = pjoin (pjoin x y) z. -Proof. by rewrite /pjoin; case: U x y z=>T [v j y c a *]; apply: a. Qed. - -Lemma punitL x : pjoin punit x = x. -Proof. by rewrite /pjoin /punit; case: U x=>T [v j y c a l *]; apply: l. Qed. - -Lemma pmonoV x y : pvalid (pjoin x y) -> pvalid x. -Proof. -rewrite /pvalid /pjoin /valid. -by case: U x y=>T [v j y c a l w ????]; apply: w. -Qed. - -Lemma pvalidU : pvalid punit. -Proof. by rewrite /pvalid /punit; case: U=>T [v j y c a l w z *]; apply: z. Qed. - -Definition spcmPCMMixin := PCMMixin pjoinC pjoinA punitL pmonoV pvalidU. -Canonical spcmPCM := Eval hnf in PCM tp spcmPCMMixin. - -End TOPCM. - -Coercion spcmPCM : spcm >-> pcm. - -End Exports. -End SPCM. - -Export SPCM.Exports. - -(******************************) -(* Now some PCM constructions *) -(******************************) - -(* nats with addition are a pcm *) - -Definition natSPCMMixin := - SPCMMixin addnC addnA add0n (fun x y => @id true) (erefl _). -Canonical natSPCM := Eval hnf in SPCM nat natSPCMMixin. - -(* also with multiplication, but we don't make that one canonical *) - -Definition multSPCMMixin := - SPCMMixin mulnC mulnA mul1n (fun x y => @id true) (erefl _). -Definition multSPCM := Eval hnf in SPCM nat multSPCMMixin. - -(* with max too *) - -Definition maxSPCMMixin := - SPCMMixin maxnC maxnA max0n (fun x y => @id true) (erefl _). -Definition maxSPCM := Eval hnf in SPCM nat maxSPCMMixin. - -Module ProdSPCM. -Section ProdSPCM. -Variables (U V : spcm). -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 ProdSPCM. -End ProdSPCM. - -Definition prodSPCMMixin U V := - SPCMMixin (@ProdSPCM.joinC U V) (@ProdSPCM.joinA U V) - (@ProdSPCM.unitL U V) (@ProdSPCM.validL U V) (@ProdSPCM.validU U V). -Canonical prodSPCM U V := Eval hnf in SPCM (_ * _) (@prodSPCMMixin U V). - -(* unit is a pcm; just for kicks *) - -Module UnitSPCM. -Section UnitSPCM. - -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 UnitSPCM. -End UnitSPCM. - -Definition unitSPCMMixin := - SPCMMixin UnitSPCM.ujoinC UnitSPCM.ujoinA - UnitSPCM.uunitL UnitSPCM.uvalidL UnitSPCM.uvalidU. -Canonical unitSPCM := Eval hnf in SPCM unit unitSPCMMixin. - -(* bools with conjunction are a pcm *) - -Module BoolConjSPCM. -Lemma unitL x : true && x = x. Proof. by []. Qed. -End BoolConjSPCM. - -Definition boolSPCMMixin := SPCMMixin andbC andbA BoolConjSPCM.unitL - (fun x y => @id true) (erefl _). -Canonical boolConjSPCM := Eval hnf in SPCM bool boolSPCMMixin. - - -(* unionmaps with integer domains are small pcms *) - -Section NatMapSPCM. -Variable A : Type. - -Let U := union_mapPCM nat_ordType A. - -Definition natmapSPCMMixin := - SPCMMixin (@joinC U) (@joinA U) (@unitL U) (@validL U) (@valid_unit U). -Canonical natmapSPCM := Eval hnf in SPCM (union_map nat_ordType A) natmapSPCMMixin. - -End NatMapSPCM. - -(* so are unionmaps with pointer domain *) - -Section PtrMapSPCM. -Variable A : Type. -Let U := union_mapPCM ptr_ordType A. - -Definition ptrmapSPCMMixin := - SPCMMixin (@joinC U) (@joinA U) (@unitL U) (@validL U) (@valid_unit U). -Canonical ptrmapSPCM := Eval hnf in SPCM (union_map ptr_ordType A) ptrmapSPCMMixin. - -End PtrMapSPCM. - -(* nat maps with unit range are just finite nat sets *) -(* introduce a special name for that *) - -Module NatFinSet. -Definition natfinset := union_map [ordType of nat] unit. -Definition natfinsetSPCMMixin := natmapSPCMMixin unit. -Canonical natfinsetSPCM := Eval hnf in SPCM natfinset natfinsetSPCMMixin. - -Module Exports. -Notation natfinset := natfinset. -Notation natfinsetSPCMMixin := natfinsetSPCMMixin. -Canonical natfinsetSPCM. -End Exports. - -End NatFinSet. - -Export NatFinSet.Exports. - - - - diff --git a/Heaps/unionmap.v b/Heaps/unionmap.v deleted file mode 100644 index e76581b..0000000 --- a/Heaps/unionmap.v +++ /dev/null @@ -1,2026 +0,0 @@ -From mathcomp.ssreflect -Require Import ssreflect ssrbool ssrnat eqtype ssrfun seq. -From mathcomp -Require Import path. -From DiSeL.Heaps -Require Import idynamic ordtype finmap pcm. -Set Implicit Arguments. -Unset Strict Implicit. -Unset Printing Implicit Defensive. - -(* 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 : pred K := - if f is Def fs _ then fun x => x \in supp fs else pred0. - -Definition dom_eq f1 f2 := - (valid f1 == valid 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 pts k v := upd k v empty. - -Definition keys_of f : seq K := - if f is Def fs _ then supp fs else [::]. - -(* 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 (keys_of 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. - -(* forward recursion principle *) -Definition base_rectf (P : base -> Type) : - P Undef -> P empty -> - (forall k v f, P f -> valid (union (pts k v) f) -> - path ord k (keys_of f) -> - P (union (pts k v) f)) -> - forall f, P f. -Proof. -rewrite /empty => H1 H2 H3; apply: base_rect=>//. -apply: fmap_rect'=>[|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 keys_of 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 -> pred 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; - pts_op : K -> V -> T; - keys_of_op : T -> seq K; - - 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 k v, pts_op k v = to_op (UM.pts p k v); - _ : forall f, keys_of_op f = (UM.keys_of (from_op f))}. -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 "leas-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 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 -> pred 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 pts : K -> V -> cT := pts_op (mixin class). -Definition keys_of : cT -> seq K := keys_of_op (mixin class). - -End ClassDef. - -Implicit Arguments um_undef [K V cT]. -Implicit Arguments empty [K V cT]. -Implicit 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 ptsE k v : pts k v = to (UM.pts (cond U) k v). -Proof. by case: U k v=>x [p][*]. Qed. - -Lemma keys_ofE f : keys_of f = UM.keys_of (from f). -Proof. by case: U f=>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 pts := pts. -Notation keys_of := keys_of. - -Definition umE := - (ftE, tfE, eqE, undefE, defE, emptyE, updE, domE, dom_eqE, - freeE, findE, unionE, um_filterE, empbE, ptsE, keys_ofE, 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 !umE /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 !umE /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 !umE; case: (UMC.from f1). Qed. - -Lemma unitL f : unit \+ f = f. -Proof. -rewrite -[f]UMC.tfE !umE /UM.union /UM.empty. -by case: (UMC.from f)=>[//|f' H]; rewrite disjC disj_nil fcat0s. -Qed. - -Lemma validU : valid unit. -Proof. by rewrite !umE. 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_classPCMMixin := - 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_classPCMMixin. - -End Exports. -End Exports. - -End UnionMapClassPCM. - -Export UnionMapClassPCM.Exports. - - -(*****************) -(* Cancelativity *) -(*****************) - -Section Cancelativity. -Variables (K : ordType) (V : Type) (U : union_map_class K V). -Implicit Type f : U. - -Lemma joinKf f f1 f2 : valid (f1 \+ f) -> f1 \+ f = f2 \+ f -> f1 = f2. -Proof. -rewrite -{3}[f1]UMC.tfE -{2}[f2]UMC.tfE !pcmE /= !umE /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. - -Lemma joinfK f f1 f2 : valid (f \+ f1) -> f \+ f1 = f \+ f2 -> f1 = f2. -Proof. by rewrite !(joinC f); apply: joinKf. Qed. - -End Cancelativity. - - -(*********************************************************) -(* 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_eqMixin := EqMixin (@unionmap_eqP K V p T m). -(* don't declare canonical here, but do so for every type individually *) -(* we don't have a generic theory of decidable equality for union_maps *) -(* so this is not needed here *) -(* unlike pcm notation, which we use in all the lemmas *) -(* so we need a generic projection from union_map to pcm *) -(* -Canonical union_map_class_eqType := - Eval hnf in EqType (UMC T m) union_map_class_eqMixin. -*) - -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) = pred0. -Proof. by rewrite !umE. Qed. - -Lemma dom0 : dom (Unit : U) = pred0. -Proof. by rewrite pcmE /= !umE. Qed. - -Lemma dom0E f : valid f -> dom f =i pred0 -> f = Unit. -Proof. -rewrite !pcmE /= !umE /UM.valid /UM.dom /UM.empty -{3}[f]UMC.tfE. -case: (UMC.from f)=>[|f' S] //= _; rewrite !umE 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 /= !umE /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 !umE; case: (UMC.from f)=>[|f' H] x; rewrite -!topredE /=; -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 /= !umE /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 /= !umE; case: (UMC.from f). Qed. - -Lemma dom_cond k f : k \in dom f -> cond U k. -Proof. by rewrite !umE; case: (UMC.from f)=>[|f' F] // /(allP F). Qed. - -Lemma dom_free k f : k \notin dom f -> free k f = f. -Proof. -rewrite -{3}[f]UMC.tfE !umE /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 !umE /UM.dom -{1}[f]UMC.tfE. -case: (UMC.from f)=>[|f' H]. -- by apply: dom_find_none'; rewrite !umE. -case: suppP (allP H k)=>[v|] H1 I; last first. -- by apply: dom_find_none'; rewrite !umE. -apply: (dom_find_some' (v:=v)); rewrite !umE // /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_um_filt p f : dom (um_filter p f) =i [predI p & dom f]. -Proof. -rewrite !umE /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 /= !umE. -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; rewrite -!topredE /= => 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 /= !umE /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 -!topredE /= !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 um_filt_dom f1 f2 : - valid (f1 \+ f2) -> um_filter (dom f1) (f1 \+ f2) = f1. -Proof. -rewrite -{4}[f1]UMC.tfE !pcmE /= !umE. -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. - -End DomLemmas. - -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 um_filtUn q f1 f2 : - valid (f1 \+ f2) -> - um_filter q (f1 \+ f2) = um_filter q f1 \+ um_filter q f2. -Proof. -rewrite !pcmE /= !umE /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 um_filt0 q : um_filter q Unit = Unit :> U. -Proof. by rewrite !pcmE /= !umE /UM.um_filter /UM.empty kfilt_nil. Qed. - -Lemma um_filt_pred0 f : valid f -> um_filter pred0 f = Unit. -Proof. -rewrite !pcmE /= !umE /UM.valid /UM.empty. -case: (UMC.from f)=>[|f' H] //= _; case: f' H=>f' P H. -by rewrite fmapE /= /kfilter' filter_pred0. -Qed. - -Lemma um_filt_predT f : um_filter predT f = f. -Proof. -rewrite -[f]UMC.tfE !umE /UM.um_filter. -by case: (UMC.from f)=>[|f' H] //; rewrite kfilter_predT. -Qed. - -Lemma um_filt_predI p1 p2 f : - um_filter (predI p1 p2) f = um_filter p1 (um_filter p2 f). -Proof. -rewrite -[f]UMC.tfE !umE /UM.um_filter. -by case: (UMC.from f)=>[|f' H] //; rewrite kfilter_predI. -Qed. - -Lemma um_filt_predU p1 p2 f : - um_filter (predU p1 p2) f = - um_filter p1 f \+ um_filter (predD p2 p1) f. -Proof. -rewrite pcmE /= !umE /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 eq_in_um_filt p1 p2 f : - {in dom f, p1 =1 p2} -> um_filter p1 f = um_filter p2 f. -Proof. -rewrite !umE /UM.dom /UM.um_filter /= /dom. -by case: (UMC.from f)=>[|f' H] //=; apply: eq_in_kfilter. -Qed. - -Lemma um_filtUnK 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 (um_filtUn _ 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_um_filt inE /=. - apply/andP/idP=>[[//]| H1]; split=>//; move: H1. - rewrite -{1}E domUn inE /= !dom_um_filt 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'; move/(joinfK V'): E. -Qed. - -Lemma um_filtU 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 !umE /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 !umE kfilt_ins H1. -- by rewrite H2 !umE kfilt_ins H1. -by case: ifP H2. -Qed. - -Lemma um_filtF 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 !umE /UM.um_filter /UM.free. -by case: (UMC.from f)=>[|f' F]; case: ifP=>// E; rewrite !umE 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 /= !umE -2![f]UMC.tfE !umE; 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 /= !umE /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 /= !umE; 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 /= !umE -{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 /= !umE. -- by apply: valid_false1; rewrite pcmE /= !umE. -- by apply: valid_false2; rewrite pcmE /= !umE. -case: ifP=>E. -- apply: valid_true; try by rewrite !pcmE /= !umE. - by move=>k; rewrite !umE => H; case: disjP E=>//; move/(_ _ H). -case: disjP E=>// k T1 T2 _. -by apply: (valid_false3 (k:=k)); rewrite !umE. -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 !inE /= (dom_cond H2) 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_um_filter p f : valid (um_filter p f) = valid f. -Proof. by rewrite !pcmE /= !umE; case: (UMC.from f). 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 /= !umE /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 domeq_validUn 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 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: domeq_validUn F _. -move=>x; rewrite !domUn /= inE. -by rewrite (domeq_validUn F (domeq_refl g)) H. -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 *) -(*************************) - -(* should really be part of dom section, but proofs use lemmas *) -(* which haven't been proved before the dom section *) - -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 (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_um_filt inE; case/andP. -move: (um_filt_dom V1); rewrite E um_filtUn -?E //. -rewrite (eq_in_um_filt (f:=y2) (p2:=pred0)). -- by rewrite um_filt_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 !umE. 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 /= !umE /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 !umE /UM.upd. -case: (UMC.from f)=>[|f' H']; case: ifP=>// E; -case: decP=>H1; case: decP=>H2 //; rewrite !umE. -- 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 !umE /UM.dom /UM.upd /UM.free. -case: (UMC.from f)=>[|f' F] //; case: ifP=>// H1; -by case: decP=>H2 //; rewrite !umE 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 /= !umE /UM.upd /UM.union /UM.dom -!topredE. -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 !umE; case: disjP D=>// H _. - by rewrite (H _ H1) /= fcat_inss // (H _ H1). -- by rewrite disj_ins H1 D /= !umE 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 /= !umE /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 !umE /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 !umE rem_ins H1. -- by case: ifP H2. -by rewrite !umE 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 !umE /UM.free. -by case: (UMC.from f)=>[|f' F]; case: ifP=>// E; rewrite !umE 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 /= !umE /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 !umE rem_supp // supp_fcat inE E2. -rewrite disj_rem; last by rewrite disjC disj_rem // disjC. -rewrite !umE; 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_um_filt 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 !umE /UM.free /UM.um_filter. -case: (UMC.from f)=>[|f' F]; case: ifP=>// E; rewrite !umE. -- 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 /= !umE /UM.find /= fnd_empty. Qed. - -Lemma find_undef k : find k (um_undef : U) = None. -Proof. by rewrite !umE /UM.find. 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 /= !umE /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 !umE /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 /= !umE /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 /= !umE /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 !umE; apply/fmapP=>k; move: (H k); rewrite !umE. -Qed. - -Lemma find_um_filt p k f : - find k (um_filter p f) = if p k then find k f else None. -Proof. -rewrite !umE /UM.find /UM.um_filter. -case: (UMC.from f)=>[|f' F]; first by rewrite if_same. -by rewrite fnd_kfilt. -Qed. - -End FindLemmas. - - -(********) -(* 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 !umE. Qed. - -Lemma empbP f : reflect (f = Unit) (empb f). -Proof. -rewrite pcmE /= !umE /UM.empty /UM.empb -{1}[f]UMC.tfE. -case: (UMC.from f)=>[|f' F]; first by apply: ReflectF; rewrite !umE. -case: eqP=>E; constructor; rewrite !umE. -- 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 !umE /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 /= !umE /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 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 /= !umE /UM.valid /UM.dom /UM.empty -{3}[f]UMC.tfE. -case: (UMC.from f)=>[|f' F] // _ H; rewrite !umE. -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_um_filt p f : empb f -> empb (um_filter p f). -Proof. -rewrite !umE /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 join_undefL f : um_undef \+ f = um_undef. -Proof. by rewrite /PCM.join /= !umE. Qed. - -Lemma join_undefR f : f \+ um_undef = um_undef. -Proof. by rewrite joinC join_undefL. Qed. - -End UndefLemmas. - - -(***********) -(* keys_of *) -(***********) - -(* basically, expose that dom is a sequence *) -(* makes you wonder if we should just declare dom to be *) -(* a seq, and remove key_of. But I can't bother now *) - -Section KeysOfLemmas. -Variables (K : ordType) (V : Type) (U : union_map_class K V). -Implicit Types (k : K) (v : V) (f : U). - -Lemma keys_dom f : keys_of f =i dom f. -Proof. -rewrite !umE /UM.keys_of /=. -by case: (UMC.from f). -Qed. - -Lemma keys0 : keys_of (Unit : U) = [::]. -Proof. -have : forall s : seq K, s =i [::] -> s = [::]. -- by elim=>[|x s IH] // /(_ x); rewrite inE eq_refl. -by apply=>x; rewrite keys_dom dom0. -Qed. - -Lemma keys_undef : keys_of (um_undef : U) = [::]. -Proof. -have : forall s : seq K, s =i [::] -> s = [::]. -- by elim=>[|x s IH] // /(_ x); rewrite inE eq_refl. -by apply=>x; rewrite keys_dom dom_undef. -Qed. - -Lemma keys_sorted f : sorted (@ord K) (keys_of f). -Proof. by rewrite !umE; case: (UMC.from f)=>[|[]]. Qed. - -Lemma keys_uniq f : uniq (keys_of f). -Proof. -apply: sorted_uniq (keys_sorted f); -by [apply: ordtype.trans | apply: ordtype.irr]. -Qed. - -Lemma keysUn_perm f1 f2 : - valid (f1 \+ f2) -> - perm_eq (keys_of (f1 \+ f2)) (keys_of f1 ++ keys_of f2). -Proof. -move=>Vh; apply: uniq_perm_eq; last 1 first. -- by move=>x; rewrite mem_cat !keys_dom domUn inE Vh. -- by apply: keys_uniq. -rewrite cat_uniq !keys_uniq /= andbT; apply/hasPn=>x. -by apply: contraL; rewrite /= !keys_dom; case: validUn Vh=>// _ _ /(_ x). -Qed. - -Lemma keysUn_size f1 f2 : - valid (f1 \+ f2) -> - size (keys_of (f1 \+ f2)) = size (keys_of f1) + size (keys_of f2). -Proof. by move/keysUn_perm/perm_eq_size; rewrite size_cat. Qed. - -(* helper lemmas about last_key *) - -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 A) 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 (ordtype.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 A) 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 A) 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 A) x s1 -> path (@oleq A) x s2 -> - (forall x, x \in s1 -> x \in 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. - -(* last key is monotone *) - -Prenex Implicits oleq. - -Lemma keys_last_mono f1 f2 k : - path oleq k (keys_of f1) -> - path oleq k (keys_of f2) -> - (forall x, x \in keys_of f1 -> x \in keys_of f2) -> - oleq (last k (keys_of f1)) (last k (keys_of f2)). -Proof. -rewrite !umE; case: (UMC.from f1); first by move=>_ H _; apply: path_last H. -move=>{f1} f1 /= _ H1; case: (UMC.from f2)=>/=. -- by move=>_ /allP; case: (supp f1)=>//; rewrite /oleq eq_refl orbT. -by move=>{f2} f2 /= _; apply: seq_last_mono H1. -Qed. - -End KeysOfLemmas. - - -(****************************) -(* 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 gen_ptsU k v : pts k v = upd k v Unit :> U. -Proof. by rewrite !pcmE /= !umE /UM.pts /UM.upd; case: decP. Qed. - -Lemma gen_domPt k v : dom (pts k v : U) =i [pred x | cond U k & k == x]. -Proof. -by move=>x; rewrite gen_ptsU domU !inE eq_sym valid_unit dom0; case: eqP. -Qed. - -Lemma gen_validPt k v : valid (pts k v : U) = cond U k. -Proof. by rewrite gen_ptsU validU valid_unit andbT. Qed. - -Lemma gen_domeqPt k v1 v2 : dom_eq (pts k v1 : U) (pts k v2). -Proof. -by apply/domeqP; rewrite !gen_validPt; split=>// x; rewrite !gen_domPt. -Qed. - -Lemma gen_findPt k v : find k (pts k v : U) = if cond U k then Some v else None. -Proof. by rewrite gen_ptsU findU eq_refl /= valid_unit andbT. Qed. - -Lemma gen_findPt2 k1 k2 v : - find k1 (pts k2 v : U) = - if (k1 == k2) && cond U k2 then Some v else None. -Proof. -by rewrite gen_ptsU findU valid_unit andbT find0E; case: ifP=>//=; case: ifP. -Qed. - -Lemma gen_findPt_inv k1 k2 v w : - find k1 (pts k2 v : U) = Some w -> [/\ k1 = k2, v = w & cond U k2]. -Proof. -rewrite gen_ptsU findU; case: eqP; last by case: ifP=>//; rewrite find0E. -by move=>->; rewrite valid_unit andbT; case: ifP=>// _ [->]. -Qed. - -Lemma gen_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 gen_ptsU freeU free0 N. Qed. - -Lemma gen_freePt k v : - cond U k -> free k (pts k v : U) = Unit. -Proof. by move=>N; rewrite gen_freePt2 // eq_refl. Qed. - -Lemma gen_cancelPt k v1 v2 : - valid (pts k v1 : U) -> pts k v1 = pts k v2 :> U -> v1 = v2. -Proof. by rewrite gen_validPt !gen_ptsU; apply: upd_inj. Qed. - -Lemma gen_updPt k v1 v2 : upd k v1 (pts k v2) = pts k v1 :> U. -Proof. by rewrite !gen_ptsU updU eq_refl. Qed. - -Lemma gen_empbPt k v : empb (pts k v : U) = false. -Proof. by rewrite gen_ptsU empbU. Qed. - -(* valid *) - -Lemma gen_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 gen_validPt=>H1 -> H2; rewrite H1 (H2 k) //. - by rewrite gen_domPt inE /= eq_refl H1. -- by rewrite gen_validPt; move/negbTE=>->. -- by move/negbTE=>->; rewrite andbF. -move=>x; rewrite gen_domPt inE /=. -by case/andP=>-> /eqP <- ->; rewrite !andbF. -Qed. - -(* the projections from validPtUn are often useful *) - -Lemma gen_validPt_cond k v f : valid (pts k v \+ f) -> cond U k. -Proof. by rewrite gen_validPtUn; case/and3P. Qed. - -Lemma gen_validPtV k v f : valid (pts k v \+ f) -> valid f. -Proof. by rewrite gen_validPtUn; case/and3P. Qed. - -Lemma gen_validPtD k v f : valid (pts k v \+ f) -> k \notin dom f. -Proof. by rewrite gen_validPtUn; case/and3P. Qed. - -(* dom *) - -Lemma gen_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 !gen_validPtUn gen_domPt !inE. -by rewrite -!andbA; case: (cond U k). -Qed. - -Lemma gen_domPtUnE k v f : k \in dom (pts k v \+ f) = valid (pts k v \+ f). -Proof. by rewrite gen_domPtUn inE eq_refl andbT. Qed. - -Lemma gen_validxx f : valid (f \+ f) -> dom f =i pred0. -Proof. by case: validUn=>// _ _ L _ z; case: (_ \in _) (L z)=>//; elim. Qed. - -(* find *) - -Lemma gen_findPtUn k v f : - valid (pts k v \+ f) -> find k (pts k v \+ f) = Some v. -Proof. -move=>V'; rewrite findUnL // gen_domPt inE /=. -by rewrite gen_ptsU findU eq_refl valid_unit (gen_validPt_cond V'). -Qed. - -Lemma gen_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 // gen_domPt inE eq_sym. -by rewrite gen_findPt2 (gen_validPt_cond V') andbT /=; case: ifP=>// ->. -Qed. - -(* free *) - -Lemma gen_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 gen_domPtUn inE V' /= eq_sym. -rewrite gen_ptsU freeU (gen_validPt_cond V') free0. -case: eqP=>/= E; first by rewrite E unitL (dom_free (gen_validPtD V')). -by case: ifP=>// N; rewrite dom_free // N. -Qed. - -Lemma gen_freePtUn k v f : - valid (pts k v \+ f) -> free k (pts k v \+ f) = f. -Proof. by move=>V'; rewrite gen_freePtUn2 // eq_refl. Qed. - -(* upd *) - -Lemma gen_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 gen_domPt inE eq_refl gen_updPt (gen_validPt_cond V1). -have V2 : valid (pts k v1 \+ f) = false. -- by rewrite !gen_validPtUn in V1 *. -move/negbT/invalidE: V1=>->; move/negbT/invalidE: V2=>->. -by apply: upd_invalid. -Qed. - -(* empb *) - -Lemma gen_empbPtUn k v f : empb (pts k v \+ f) = false. -Proof. by rewrite empbUn gen_empbPt. Qed. - -(* others *) - -Lemma gen_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 gen_ptsU. -Qed. - -Lemma gen_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 !gen_findPtUn -?E //; case=>X. -by rewrite -{}X in E *; rewrite -(joinfK V' E) (validR V'). -Qed. - -Lemma gen_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/(gen_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 gen_domPtUn inE V2 eq_refl /=. - by rewrite gen_ptsU freeU eq_sym X free0 -gen_ptsU gen_freePtUn. -suff E'' : free k2 f1 = free k1 f2. -- by rewrite -E''; rewrite E' joinCA in E; case/(gen_cancel V1): E. -move: (erefl (free k1 f2)). -by rewrite {1}E' gen_freePtUn // -E'; apply: (validR V2). -Qed. - -Lemma gen_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 gen_ptsU um_filtU um_filt0. Qed. - -(* keys *) - -Lemma gen_keysPt k v : - keys_of (pts k v : U) = if cond U k then [:: k] else [::]. -Proof. -rewrite !umE /= /UM.keys_of /supp /UM.pts /UM.upd /UM.empty /=. -by case D : (decP _)=>[a|a] /=; rewrite ?a ?(introF idP a). -Qed. - -(* induction and recursion over union maps, expressed with pts and \+ *) - -(* forward progressing over keys *) -Lemma gen_indf (P : U -> Prop) : - P um_undef -> P Unit -> - (forall k v f, P f -> valid (pts k v \+ f) -> - path ord k (keys_of f) -> - P (pts k v \+ f)) -> - forall f, P f. -Proof. -rewrite !pcmE /= !umE => 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 /= !umE. -by apply. -Qed. - -Lemma gen_rect (P : U -> Type) : - P um_undef -> P Unit -> - (forall k v f, P f -> valid (pts k v \+ f) -> - path ord k (keys_of f) -> - P (pts k v \+ f)) -> - forall f, P f. -Proof. -rewrite !pcmE /= !umE => H1 H2 H3 f; rewrite -[f]UMC.tfE. -apply: (UM.base_rectf (P := (fun b => P (UMC.to b))))=>//. -move=>k v g H V1; move: (H3 k v _ H); rewrite !pcmE /= !umE. -by apply. -Qed. - -(* backward progressing over keys *) -Lemma gen_indb (P : U -> Prop) : - P um_undef -> P Unit -> - (forall k v f, P f -> valid (pts k v \+ f) -> - (forall x, x \in keys_of f -> ord x k) -> - P (pts k v \+ f)) -> - forall f, P f. -Proof. -rewrite !pcmE /= !umE => 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 /= !umE. -by apply. -Qed. - -End PointsToLemmas. - -Prenex Implicits gen_findPt_inv. - -(*******************************) -(*******************************) -(* Main instance of union maps *) -(*******************************) -(*******************************) - -(* 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 -> pred 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 pts : K -> V -> tp. -Parameter keys_of : tp -> seq K. - -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 ptsE : forall k v, pts k v = to (@UM.pts K V predT k v). -Axiom keys_ofE : forall f, keys_of f = UM.keys_of (from f). - -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 pts k v := @UM.pts K V predT k v. -Definition keys_of f := @UM.keys_of K V predT f. - -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 ptsE k v : pts k v = to (@UM.pts K V predT k v). Proof. by []. Qed. -Lemma keys_ofE f : keys_of f = UM.keys_of (from f). 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.ptsE K V) (@UMDef.keys_ofE 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_mapPCMMixin K V := union_map_classPCMMixin (union_mapUMC K V). -Canonical union_mapPCM K V := - Eval hnf in PCM (union_map K V) (@union_mapPCMMixin K V). - -Definition union_map_eqMixin K (V : eqType) := - @union_map_class_eqMixin K V _ _ (unionmapMixin K V). -Canonical union_map_eqType K (V : eqType) := - Eval hnf in EqType (union_map K V) (@union_map_eqMixin 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. - -(* can i store maps into maps without universe inconsistencies? *) -(* yes, the idea of the class works *) -Lemma xx : 1 \\-> (1 \\-> 3) = 2 \\-> (7 \\-> 3). -Abort. - -(*************************************************) -(* Points-to lemmas specific to plain union maps *) -(*************************************************) - -Section UMPointsToLemmas. -Variables (K : ordType) (V : Type). -Notation U := (union_mapUMC K V). -Implicit Types (k : K) (v : V) (f : U). - -Lemma um_ptsU k v : k \\-> v = upd k v (Unit : U). -Proof. exact: gen_ptsU. Qed. - -Lemma um_domPt k v : dom (k \\-> v) =i [pred x | k == x]. -Proof. exact: gen_domPt. Qed. - -Lemma um_validPt k v : valid (k \\-> v) = cond U k. -Proof. exact: gen_validPt. Qed. - -Lemma um_domeqPt k v1 v2 : dom_eq (k \\-> v1) (k \\-> v2). -Proof. exact: gen_domeqPt. Qed. - -Lemma um_findPt k v : find k (k \\-> v) = Some v. -Proof. exact: gen_findPt. Qed. - -Lemma um_freePt2 k1 k2 v : - cond U k2 -> - free k1 (k2 \\-> v) = if k1 == k2 then Unit else pts k2 v. -Proof. exact: gen_freePt2. Qed. - -Lemma um_freePt k v : - cond U k -> free k (k \\-> v) = Unit. -Proof. exact: gen_freePt. Qed. - -Lemma um_cancelPt k v1 v2 : k \\-> v1 = k \\-> v2 -> v1 = v2. -Proof. by move/(@gen_cancelPt K V); rewrite gen_validPt; apply. Qed. - -Lemma um_updPt k v1 v2 : upd k v1 (k \\-> v2) = k \\-> v1. -Proof. exact: gen_updPt. Qed. - -Lemma um_empbPt k v : empb (k \\-> v) = false. -Proof. exact: gen_empbPt. Qed. - -Lemma um_validPtUn k v f : - valid (k \\-> v \+ f) = valid f && (k \notin dom f). -Proof. exact: gen_validPtUn. Qed. - -Lemma um_validPtV k v f : valid (k \\-> v \+ f) -> valid f. -Proof. exact: gen_validPtV. Qed. - -Lemma um_validPtD k v f : valid (k \\-> v \+ f) -> k \notin dom f. -Proof. exact: gen_validPtD. Qed. - -Lemma um_domPtUn k v f : - dom (k \\-> v \+ f) =i - [pred x | valid (k \\-> v \+ f) & (k == x) || (x \in dom f)]. -Proof. exact: gen_domPtUn. Qed. - -Lemma um_domPtUnE k v f : k \in dom (k \\-> v \+ f) = valid (k \\-> v \+ f). -Proof. exact: gen_domPtUnE. Qed. - -Lemma um_validxx f : valid (f \+ f) -> dom f =i pred0. -Proof. exact: gen_validxx. Qed. - -Lemma um_findPtUn k v f : - valid (k \\-> v \+ f) -> find k (k \\-> v \+ f) = Some v. -Proof. exact: gen_findPtUn. Qed. - -Lemma um_findPtUn2 k1 k2 v f : - valid (k2 \\-> v \+ f) -> - find k1 (k2 \\-> v \+ f) = if k1 == k2 then Some v else find k1 f. -Proof. exact: gen_findPtUn2. Qed. - -Lemma um_freePtUn k v f : - valid (k \\-> v \+ f) -> free k (k \\-> v \+ f) = f. -Proof. exact: gen_freePtUn. Qed. - -Lemma um_freePtUn2 k1 k2 v f : - valid (k2 \\-> v \+ f) -> - free k1 (k2 \\-> v \+ f) = - if k1 == k2 then f else k2 \\-> v \+ free k1 f. -Proof. exact: gen_freePtUn2. Qed. - -Lemma um_updPtUn k v1 v2 f : upd k v1 (k \\-> v2 \+ f) = k \\-> v1 \+ f. -Proof. exact: gen_updPtUn. Qed. - -Lemma um_findPt_inv k1 k2 v w : - find k1 (k2 \\-> v) = Some w -> k1 = k2 /\ v = w. -Proof. by case/gen_findPt_inv. Qed. - -Lemma um_empbPtUn k v f : empb (k \\-> v \+ f) = false. -Proof. exact: gen_empbPtUn. Qed. - -Lemma um_eta k f : - k \in dom f -> - exists v, find k f = Some v /\ f = k \\-> v \+ free k f. -Proof. exact: gen_eta. Qed. - -Lemma um_eta2 k v f : - find k f = Some v -> f = k \\-> v \+ free k f. -Proof. by move=>E; case/um_eta: (find_some E)=>v' []; rewrite E; case=><-. Qed. - -Lemma um_cancel k v1 v2 f1 f2 : - valid (k \\-> v1 \+ f1) -> - k \\-> v1 \+ f1 = k \\-> v2 \+ f2 -> [/\ v1 = v2, valid f1 & f1 = f2]. -Proof. exact: gen_cancel. Qed. - -Lemma um_cancel2 k1 k2 f1 f2 v1 v2 : - valid (k1 \\-> v1 \+ f1) -> - k1 \\-> v1 \+ f1 = k2 \\-> v2 \+ f2 -> - if k1 == k2 then v1 = v2 /\ f1 = f2 - else [/\ free k1 f2 = free k2 f1, - f1 = k2 \\-> v2 \+ free k1 f2 & - f2 = k1 \\-> v1 \+ free k2 f1]. -Proof. exact: gen_cancel2. Qed. - -Lemma um_umfiltPt p k v : - um_filter p (k \\-> v) = if p k then k \\-> v else Unit. -Proof. exact: gen_umfiltPt. Qed. - -(* keys *) - -Lemma um_keysPt k v : keys_of (k \\-> v) = [:: k]. -Proof. by rewrite gen_keysPt. Qed. - -Lemma um_indf (P : U -> Prop) : - P um_undef -> P Unit -> - (forall k v f, P f -> valid (k \\-> v \+ f) -> - path ord k (keys_of f) -> P (k \\-> v \+ f)) -> - forall f, P f. -Proof. exact: gen_indf. Qed. - -Lemma um_indb (P : U -> Prop) : - P um_undef -> P Unit -> - (forall k v f, P f -> valid (k \\-> v \+ f) -> - (forall x, x \in keys_of f -> ord x k) -> - P (k \\-> v \+ f)) -> - forall f, P f. -Proof. by move=>H1 H2 H3; apply: gen_indb. Qed. - - -End UMPointsToLemmas. - -Prenex Implicits um_findPt_inv dom um_eta2. diff --git a/README.md b/README.md index bb18cfa..3ece2d7 100644 --- a/README.md +++ b/README.md @@ -5,19 +5,16 @@ compositional verification of distributed systems. This code accompanies the paper entitled **Programming and Proving with Distributed Protocols** by Ilya Sergey, James R. Wilcox and -Zachary Tatlock, accepted for publication at POPL 2018. +Zachary Tatlock, in the POPL 2018 proceedings. ## Building the Project -A VM has been provided for your convenience and is described below. If -you would like to use your own machine instead, you should clone this branch of -the GitHub repository; the following dependencies are necessary. - ### Requirements -* Coq 8.7 (available from https://coq.inria.fr/download) -* Mathematical Components 1.6.2 or later (http://math-comp.github.io/math-comp/) -* OCaml 4.05.0 or later (to compile and run the extracted applications) +* [Coq 8.7 or 8.8](https://coq.inria.fr) +* [Mathematical Components 1.6.2 or later](http://math-comp.github.io/math-comp/) (`ssreflect`) +* [FCSL PCM library 1.0.0 or later](https://github.com/imdea-software/fcsl-pcm) +* [OCaml 4.05.0 or later](https://ocaml.org) (to compile and run the extracted applications) ### Building Manually @@ -45,14 +42,14 @@ Make sure OPAM is installed and use the following commands: ``` opam repo add coq-released https://coq.inria.fr/opam/released opam repo add distributedcomponents-dev http://opam-dev.distributedcomponents.net -opam install disel-core +opam install disel ``` +As an alternative, a VM for a previous version has been provided for +your convenience and is described below. ## Project Structure -* `Heaps` -- A theory of partial finite heaps; - * `Core` -- Disel implementation, metatheory and inference rules; * `Examples` -- Case studies implemented in Disel diff --git a/_CoqProject b/_CoqProject index 953c8a4..6891512 100644 --- a/_CoqProject +++ b/_CoqProject @@ -1,18 +1,8 @@ --Q . DiSeL +-Q Core DiSeL +-Q Examples DiSeL -arg "-w -notation-overridden,-local-declaration,-redundant-canonical-projection,-projection-no-head-constant" -Heaps/coding.v -Heaps/domain.v -Heaps/finmap.v -Heaps/heap.v -Heaps/idynamic.v -Heaps/ordtype.v -Heaps/pcm.v -Heaps/pperm.v -Heaps/pred.v -Heaps/prelude.v -Heaps/spcm.v -Heaps/unionmap.v +Core/Domain.v Core/State.v Core/Freshness.v Core/Protocols.v @@ -52,3 +42,5 @@ Examples/TwoPhaseCommit/TwoPhaseExtraction.v Examples/TwoPhaseCommit/TwoPhaseInductiveInv.v Examples/TwoPhaseCommit/TwoPhaseInductiveProof.v Examples/Querying/QueryPlusTPC.v +LockResource/LockProtocol.v +LockResource/ResourceProtocol.v From 5c8b43463691c541c37bbab425b90d9c2417f294 Mon Sep 17 00:00:00 2001 From: Karl Palmskog Date: Sat, 8 Sep 2018 23:18:11 -0500 Subject: [PATCH 2/7] fix permission and file list --- .travis-ci.sh | 0 _CoqProject | 4 ++-- 2 files changed, 2 insertions(+), 2 deletions(-) mode change 100644 => 100755 .travis-ci.sh diff --git a/.travis-ci.sh b/.travis-ci.sh old mode 100644 new mode 100755 diff --git a/_CoqProject b/_CoqProject index 6891512..df542ff 100644 --- a/_CoqProject +++ b/_CoqProject @@ -42,5 +42,5 @@ Examples/TwoPhaseCommit/TwoPhaseExtraction.v Examples/TwoPhaseCommit/TwoPhaseInductiveInv.v Examples/TwoPhaseCommit/TwoPhaseInductiveProof.v Examples/Querying/QueryPlusTPC.v -LockResource/LockProtocol.v -LockResource/ResourceProtocol.v +Examples/LockResource/LockProtocol.v +Examples/LockResource/ResourceProtocol.v From 3cc255e0c23986b356cf2d930332a6e54dd53290 Mon Sep 17 00:00:00 2001 From: Karl Palmskog Date: Sat, 8 Sep 2018 23:25:07 -0500 Subject: [PATCH 3/7] fix pinning --- .travis-ci.sh | 2 +- Core/opam | 1 + 2 files changed, 2 insertions(+), 1 deletion(-) diff --git a/.travis-ci.sh b/.travis-ci.sh index aeafd3f..f2dcc81 100755 --- a/.travis-ci.sh +++ b/.travis-ci.sh @@ -6,5 +6,5 @@ eval $(opam config env) opam update -opam pin add Core disel --yes --verbose +opam pin add Core --yes --verbose make -j4 -C Examples diff --git a/Core/opam b/Core/opam index f3ed815..3111ea5 100644 --- a/Core/opam +++ b/Core/opam @@ -1,4 +1,5 @@ opam-version: "1.2" +name: "disel" version: "dev" maintainer: "palmskog@gmail.com" From 420247ee3d0bdfa372fe23a3a8e8078404f20dd3 Mon Sep 17 00:00:00 2001 From: Karl Palmskog Date: Sat, 8 Sep 2018 23:45:28 -0500 Subject: [PATCH 4/7] minor adjustments --- Examples/Make | 2 +- Examples/SeqLib.v | 4 +--- 2 files changed, 2 insertions(+), 4 deletions(-) diff --git a/Examples/Make b/Examples/Make index 5bc9af8..57cbcd0 100644 --- a/Examples/Make +++ b/Examples/Make @@ -1,4 +1,4 @@ --Q . DiSeL.Examples +-Q . DiSeL -arg "-w -notation-overridden,-local-declaration,-redundant-canonical-projection,-projection-no-head-constant" SeqLib.v diff --git a/Examples/SeqLib.v b/Examples/SeqLib.v index 3c42ac2..d3f8d07 100644 --- a/Examples/SeqLib.v +++ b/Examples/SeqLib.v @@ -1,7 +1,5 @@ -From mathcomp.ssreflect -Require Import ssreflect ssrbool ssrnat eqtype ssrfun seq. From mathcomp -Require Import path. +Require Import ssreflect ssrbool ssrnat eqtype ssrfun seq path. Require Import Eqdep. (*************************************************************) From 4b3e0c33c067716f703c8e021f2de7f1267817a9 Mon Sep 17 00:00:00 2001 From: Karl Palmskog Date: Sun, 9 Sep 2018 12:38:13 -0500 Subject: [PATCH 5/7] makefile adjustments --- Core/Makefile | 5 ++-- Examples/Make | 31 +++++++++++++----------- Examples/Makefile | 5 ++-- Examples/TwoPhaseCommit/TwoPhaseClient.v | 11 +++++---- Makefile | 5 ++-- _CoqProject | 3 ++- 6 files changed, 31 insertions(+), 29 deletions(-) diff --git a/Core/Makefile b/Core/Makefile index 0898c8e..954cd22 100644 --- a/Core/Makefile +++ b/Core/Makefile @@ -8,9 +8,8 @@ default: Makefile.coq install: Makefile.coq $(MAKE) -f Makefile.coq install -clean: - if [ -f Makefile.coq ]; then \ - $(MAKE) -f Makefile.coq clean; fi +clean: Makefile.coq + $(MAKE) -f Makefile.coq cleanall rm -f Makefile.coq Makefile.coq.conf Makefile.coq: Make diff --git a/Examples/Make b/Examples/Make index 57cbcd0..58d3d7f 100644 --- a/Examples/Make +++ b/Examples/Make @@ -1,23 +1,26 @@ -Q . DiSeL -arg "-w -notation-overridden,-local-declaration,-redundant-canonical-projection,-projection-no-head-constant" -SeqLib.v -Greeter/Greeter.v -Querying/QueryProtocol.v -Querying/QueryHooked.v -Calculator/CalculatorProtocol.v +Calculator/CalculatorExtraction.v Calculator/CalculatorInvariant.v -Calculator/CalculatorClientLib.v -Calculator/CalculatorServerLib.v -Calculator/SimpleCalculatorServers.v Calculator/DelegatingCalculatorServer.v +Calculator/SimpleCalculatorServers.v +Calculator/CalculatorClientLib.v +Calculator/CalculatorProtocol.v Calculator/SimpleCalculatorApp.v -TwoPhaseCommit/TwoPhaseProtocol.v -TwoPhaseCommit/TwoPhaseCoordinator.v -TwoPhaseCommit/TwoPhaseParticipant.v +Calculator/CalculatorServerLib.v +SeqLib.v +LockResource/ResourceProtocol.v +LockResource/LockProtocol.v TwoPhaseCommit/SimpleTPCApp.v TwoPhaseCommit/TwoPhaseInductiveInv.v -TwoPhaseCommit/TwoPhaseInductiveProof.v +TwoPhaseCommit/TwoPhaseClient.v +TwoPhaseCommit/TwoPhaseInductiveProof.v +TwoPhaseCommit/TwoPhaseParticipant.v +TwoPhaseCommit/TwoPhaseExtraction.v +TwoPhaseCommit/TwoPhaseProtocol.v +TwoPhaseCommit/TwoPhaseCoordinator.v +Greeter/Greeter.v Querying/QueryPlusTPC.v -LockResource/LockProtocol.v -LockResource/ResourceProtocol.v +Querying/QueryProtocol.v +Querying/QueryHooked.v diff --git a/Examples/Makefile b/Examples/Makefile index 0898c8e..954cd22 100644 --- a/Examples/Makefile +++ b/Examples/Makefile @@ -8,9 +8,8 @@ default: Makefile.coq install: Makefile.coq $(MAKE) -f Makefile.coq install -clean: - if [ -f Makefile.coq ]; then \ - $(MAKE) -f Makefile.coq clean; fi +clean: Makefile.coq + $(MAKE) -f Makefile.coq cleanall rm -f Makefile.coq Makefile.coq.conf Makefile.coq: Make diff --git a/Examples/TwoPhaseCommit/TwoPhaseClient.v b/Examples/TwoPhaseCommit/TwoPhaseClient.v index 5a27f97..26df7d6 100644 --- a/Examples/TwoPhaseCommit/TwoPhaseClient.v +++ b/Examples/TwoPhaseCommit/TwoPhaseClient.v @@ -4,11 +4,12 @@ From mathcomp Require Import path. Require Import Eqdep. Require Import Relation_Operators. -Require Import pred prelude idynamic ordtype finmap pcm unionmap heap coding domain. -Require Import State DepMaps Protocols Worlds NetworkSem Rely. -Require Import HoareTriples InferenceRules. -Require Import While. -Require Import TwoPhaseProtocol TwoPhaseInductiveInv TwoPhaseCoordinator TwoPhaseParticipant. +From fcsl +Require Import pred prelude ordtype finmap pcm unionmap heap. +From DiSeL +Require Import State DepMaps Protocols Worlds NetworkSem Rely HoareTriples InferenceRules. +From DiSeL +Require Import While TwoPhaseProtocol TwoPhaseInductiveInv TwoPhaseCoordinator TwoPhaseParticipant. Section TwoPhaseClient. diff --git a/Makefile b/Makefile index 7fd65bf..9977771 100644 --- a/Makefile +++ b/Makefile @@ -8,9 +8,8 @@ default: Makefile.coq install: Makefile.coq $(MAKE) -f Makefile.coq install -clean: - if [ -f Makefile.coq ]; then \ - $(MAKE) -f Makefile.coq clean; fi +clean: Makefile.coq + $(MAKE) -f Makefile.coq cleanall rm -f Makefile.coq Makefile.coq.conf Makefile.coq: _CoqProject diff --git a/_CoqProject b/_CoqProject index df542ff..0841cad 100644 --- a/_CoqProject +++ b/_CoqProject @@ -40,7 +40,8 @@ Examples/TwoPhaseCommit/TwoPhaseParticipant.v Examples/TwoPhaseCommit/SimpleTPCApp.v Examples/TwoPhaseCommit/TwoPhaseExtraction.v Examples/TwoPhaseCommit/TwoPhaseInductiveInv.v -Examples/TwoPhaseCommit/TwoPhaseInductiveProof.v +Examples/TwoPhaseCommit/TwoPhaseInductiveProof.v +Examples/TwoPhaseCommit/TwoPhaseClient.v Examples/Querying/QueryPlusTPC.v Examples/LockResource/LockProtocol.v Examples/LockResource/ResourceProtocol.v From c0ee1963ef7c7a2b80751e8925888a1d63b99dac Mon Sep 17 00:00:00 2001 From: Karl Palmskog Date: Sun, 9 Sep 2018 12:50:14 -0500 Subject: [PATCH 6/7] extraction must be done from the toplevel --- Examples/Make | 2 -- 1 file changed, 2 deletions(-) diff --git a/Examples/Make b/Examples/Make index 58d3d7f..e09cf67 100644 --- a/Examples/Make +++ b/Examples/Make @@ -1,7 +1,6 @@ -Q . DiSeL -arg "-w -notation-overridden,-local-declaration,-redundant-canonical-projection,-projection-no-head-constant" -Calculator/CalculatorExtraction.v Calculator/CalculatorInvariant.v Calculator/DelegatingCalculatorServer.v Calculator/SimpleCalculatorServers.v @@ -17,7 +16,6 @@ TwoPhaseCommit/TwoPhaseInductiveInv.v TwoPhaseCommit/TwoPhaseClient.v TwoPhaseCommit/TwoPhaseInductiveProof.v TwoPhaseCommit/TwoPhaseParticipant.v -TwoPhaseCommit/TwoPhaseExtraction.v TwoPhaseCommit/TwoPhaseProtocol.v TwoPhaseCommit/TwoPhaseCoordinator.v Greeter/Greeter.v From def3f891b171f6a07ced465c2ce6824d28c38d72 Mon Sep 17 00:00:00 2001 From: Karl Palmskog Date: Mon, 17 Sep 2018 21:43:18 -0500 Subject: [PATCH 7/7] fix final deprecation warnings --- Examples/Calculator/CalculatorClientLib.v | 5 +---- Examples/Calculator/CalculatorInvariant.v | 2 -- Examples/Calculator/CalculatorServerLib.v | 5 +---- Examples/Querying/QueryHooked.v | 14 +++----------- Examples/TwoPhaseCommit/TwoPhaseParticipant.v | 2 +- 5 files changed, 6 insertions(+), 22 deletions(-) diff --git a/Examples/Calculator/CalculatorClientLib.v b/Examples/Calculator/CalculatorClientLib.v index 89c40c2..2dd6a89 100644 --- a/Examples/Calculator/CalculatorClientLib.v +++ b/Examples/Calculator/CalculatorClientLib.v @@ -11,7 +11,7 @@ Require Import Freshness State EqTypeX Protocols Worlds NetworkSem Rely. From DiSeL Require Import Actions Injection Process Always HoareTriples InferenceRules. From DiSeL -Require Import InductiveInv. +Require Import InductiveInv While. From DiSeL Require Import CalculatorProtocol CalculatorInvariant. From DiSeL @@ -133,9 +133,6 @@ Definition receive_loop_inv (rs : reqs) := | None => loc i = st :-> rs end. -From DiSeL -Require Import While. - Program Definition receive_loop' : {(rs : reqs)}, DHT [cl, W] (fun i => loc i = st :-> rs, diff --git a/Examples/Calculator/CalculatorInvariant.v b/Examples/Calculator/CalculatorInvariant.v index 3a129d3..733a169 100644 --- a/Examples/Calculator/CalculatorInvariant.v +++ b/Examples/Calculator/CalculatorInvariant.v @@ -59,8 +59,6 @@ Notation coh' := (coh cal). Notation Sinv := (@S_inv cal (fun d _ => CalcInv d)). Notation Rinv := (@R_inv cal (fun d _ => CalcInv d)). Notation PI := pf_irr. -From DiSeL -Require Import CalculatorProtocol. Program Definition s1: Sinv (server_send_trans f prec cs cls). Proof. diff --git a/Examples/Calculator/CalculatorServerLib.v b/Examples/Calculator/CalculatorServerLib.v index bc78685..db62483 100644 --- a/Examples/Calculator/CalculatorServerLib.v +++ b/Examples/Calculator/CalculatorServerLib.v @@ -11,7 +11,7 @@ Require Import Freshness State EqTypeX Protocols Worlds NetworkSem Rely. From DiSeL Require Import Actions Injection Process Always HoareTriples InferenceRules. From DiSeL -Require Import InductiveInv. +Require Import InductiveInv While. From DiSeL Require Import CalculatorProtocol CalculatorInvariant. From DiSeL @@ -92,9 +92,6 @@ Definition receive_req_loop_inv (ps : reqs) := | None => loc i = st :-> ps end. -From DiSeL -Require Import While. - Program Definition receive_req_loop : {ps : reqs}, DHT [sv, W] (fun i => loc i = st :-> ps, diff --git a/Examples/Querying/QueryHooked.v b/Examples/Querying/QueryHooked.v index b37d450..c7b15a3 100644 --- a/Examples/Querying/QueryHooked.v +++ b/Examples/Querying/QueryHooked.v @@ -1,18 +1,14 @@ From mathcomp.ssreflect -Require Import ssreflect ssrbool ssrnat eqtype ssrfun seq. -From mathcomp -Require Import path. +Require Import ssreflect ssrbool ssrnat eqtype ssrfun seq path. Require Import Eqdep. From fcsl Require Import axioms pred prelude ordtype finmap pcm unionmap heap. From DiSeL Require Import Freshness State EqTypeX Protocols Worlds NetworkSem Rely Actions. From DiSeL -Require Import SeqLib QueryProtocol. +Require Import SeqLib QueryProtocol NewStatePredicates Actions. From DiSeL -Require Import NewStatePredicates. -From DiSeL -Require Import Actions Injection Process Always HoareTriples InferenceRules. +Require Import Injection Process Always HoareTriples InferenceRules While. Section QueryHooked. @@ -1080,9 +1076,6 @@ Definition recv_resp_inv (rid : nat) to local_indicator data (getLc i) & msg_story i rid to data ((to, rid) :: reqs) resp]. -From DiSeL -Require Import While. - Program Definition receive_resp_loop (rid : nat) to : {(rrd : (seq (nid * nat) * seq (nid * nat) * Data))}, DHT [this, W] (fun i => let: (reqs, resp, data) := rrd in @@ -1245,4 +1238,3 @@ by apply: (core_state_stable _ _ _ _ R _ T2 T4); case: T3. Qed. End QueryHooked. - diff --git a/Examples/TwoPhaseCommit/TwoPhaseParticipant.v b/Examples/TwoPhaseCommit/TwoPhaseParticipant.v index 1a3d595..d437a5a 100644 --- a/Examples/TwoPhaseCommit/TwoPhaseParticipant.v +++ b/Examples/TwoPhaseCommit/TwoPhaseParticipant.v @@ -85,7 +85,7 @@ Next Obligation. by rewrite !InE; do![right]. Qed. (************** Participant code **************) -Implicit Arguments TPCProtocol.TPCCoh [cn pts others]. +Arguments TPCProtocol.TPCCoh [cn pts others]. Notation coh := (@TPCProtocol.TPCCoh cn pts others). Notation getS s := (getStatelet s l). Notation loc i := (getLocal p (getStatelet i l)).