From 1d1fefd5e19b77bf7b857625ef2263f047cd3b5f Mon Sep 17 00:00:00 2001 From: Aleksandar Nanevski Date: Mon, 30 Sep 2024 15:33:49 +0200 Subject: [PATCH] v2.0 (#27) * blah * adding files * renaming * blah * removed interlude.v * started refactoring model.v * minor * started refactoring examples * continuing refactoring * refactored a bunch of examples. going on with graph.v * started refactoring graph.v * minor * wip * continuing with refactoring graph.v. * blah * blah * blah * comments about do we want to follow dangling edges to their dangling target node or not. i think not. * comments * minor * minor * blah * wip * wip * refactored schorr.v * committed schorr * minor * minor * minor * minor * moved some helper lemmas to seqext.v * blah * blah * blah * blah * blah * blah * blah * blah * wip * some cleaning up in congprog.v * cleanup in congprog.v * cleanup * cleanup * blah * ported proofs of congr_prog * modified congprog to use postconditions with vfun. * abstracting the code in congprog.v * defining shape predicate * made abstract type of root pointers of arrays and kvmaps be small (ie.. declared as Set) * developed a number of lemmas about star. to be moved to other files later * defining shape predicate * defining shape predicate * blah * reorgainzation of pred.v prelude.v and consequences. also, largely cleaned up kvmaps.v and hashtable.v to get the inheritance and hiding to behave properly * removed some files that shouldn't be here * moved additional lemmas from individual files to fcsl-pcm. also, cleaned up congprog, hashtab, kvmaps, etc. * minor * sectioning of congmath.v * minor * modifying comments * changed natmap to add view for last_val * hm * removed extraneous lemmas from unionmap.v * minor * changed In_dom_umfilt lemma to use exists2 instead of exists. * found a way to view "weird" lemmas in natmap.v as non-weird. removed the corresponding comment, and repositioned the lemmas. * propagating changed to natmap.v from mathador * added alternative lemma for seq_lt irreflexivity, one that isn't given as an equation, but as implication into False. * renaming slt_irrN into sltnn * added validPt2 and domPt2 * propagating changes from mathador * blah * ibalh * blah * blah * blah * many changes introduced to deal with graphs * preparing for release * removed files inheritted from fcsl-pcm * added Marcos to the list of authors * regenerated .opam files from meta.yml * removed devcomments * removed some lemmas that have recently been included into mathcomp * regenerated htt/dune file * playing with dune * playing with dune * forgot some files * blah * blah * blah * blah * changed mustache files --------- Co-authored-by: kevinlopez --- .github/workflows/docker-action.yml | 7 +- README.md | 8 +- _CoqProject | 35 +- coq-htt-examples.opam => coq-htt-core.opam | 13 +- coq-htt.opam | 14 +- dune-project | 4 +- examples/array.v | 149 +- examples/bintree.v | 128 +- examples/bst.v | 168 +- examples/bubblesort.v | 362 ++--- examples/congmath.v | 914 ++++++----- examples/congprog.v | 637 ++++++++ examples/counter.v | 35 +- examples/cyclic.v | 185 ++- examples/dlist.v | 130 +- examples/exploit.v | 13 + examples/gcd.v | 29 +- examples/hashtab.v | 240 +-- examples/kvmaps.v | 619 +++++--- examples/llist.v | 164 +- examples/queue.v | 83 +- examples/quicksort.v | 265 ++-- examples/stack.v | 64 +- examples/tree.v | 60 +- examples/union_find.v | 167 +- htt/domain.v | 1031 +++++++++++++ {theories => htt}/dune | 2 +- {theories => htt}/heapauto.v | 164 +- {theories => htt}/model.v | 989 ++++++------ htt/options.v | 12 + meta.yml | 28 +- regenerate.sh | 4 +- ...es.opam.mustache => coq-htt.opam.mustache} | 2 +- templates-extra/docker-action.yml.mustache | 5 +- templates-extra/dune.mustache | 2 +- theories/domain.v | 1367 ----------------- theories/interlude.v | 192 --- theories/options.v | 3 - 38 files changed, 4475 insertions(+), 3819 deletions(-) rename coq-htt-examples.opam => coq-htt-core.opam (86%) mode change 100755 => 100644 examples/congmath.v create mode 100644 examples/congprog.v mode change 100755 => 100644 examples/counter.v mode change 100755 => 100644 examples/exploit.v mode change 100755 => 100644 examples/hashtab.v mode change 100755 => 100644 examples/kvmaps.v mode change 100755 => 100644 examples/queue.v mode change 100755 => 100644 examples/stack.v create mode 100644 htt/domain.v rename {theories => htt}/dune (91%) rename {theories => htt}/heapauto.v (82%) rename {theories => htt}/model.v (56%) create mode 100644 htt/options.v rename templates-extra/{coq-htt-examples.opam.mustache => coq-htt.opam.mustache} (95%) delete mode 100644 theories/domain.v delete mode 100644 theories/interlude.v delete mode 100644 theories/options.v diff --git a/.github/workflows/docker-action.yml b/.github/workflows/docker-action.yml index a774d92..761fe01 100644 --- a/.github/workflows/docker-action.yml +++ b/.github/workflows/docker-action.yml @@ -16,14 +16,15 @@ jobs: strategy: matrix: image: - - 'mathcomp/mathcomp:1.17.0-coq-8.15' - - 'mathcomp/mathcomp:1.17.0-coq-8.18' + - 'mathcomp/mathcomp:2.2.0-coq-8.19' + - 'mathcomp/mathcomp:2.2.0-coq-8.20' + - 'mathcomp/mathcomp:latest-coq-dev' fail-fast: false steps: - uses: actions/checkout@v2 - uses: coq-community/docker-coq-action@v1 with: - opam_file: 'coq-htt-examples.opam' + opam_file: 'coq-htt.opam' custom_image: ${{ matrix.image }} diff --git a/README.md b/README.md index 0fe5f74..e9a1336 100644 --- a/README.md +++ b/README.md @@ -37,13 +37,13 @@ that HTT implements Separation logic as a shallow embedding in Coq. - Aleksandar Nanevski (initial) - Germán Andrés Delbianco - Alexander Gryzlov + - Marcos Grandury - License: [Apache-2.0](LICENSE) -- Compatible Coq versions: Coq 8.15 to 8.17 +- Compatible Coq versions: Coq 8.19 to 8.20 - Additional dependencies: - - [MathComp ssreflect 1.17](https://math-comp.github.io) + - [MathComp ssreflect 2.2](https://math-comp.github.io) - [MathComp algebra](https://math-comp.github.io) - - [MathComp fingroup](https://math-comp.github.io) - - [FCSL-PCM 1.8](https://github.com/imdea-software/fcsl-pcm) + - [FCSL-PCM 2.0](https://github.com/imdea-software/fcsl-pcm) - [Dune](https://dune.build) 2.5 or later - Coq namespace: `htt` - Related publication(s): diff --git a/_CoqProject b/_CoqProject index b2abb00..f4eab2f 100644 --- a/_CoqProject +++ b/_CoqProject @@ -1,26 +1,35 @@ --Q . htt --arg "-w -notation-overridden,-local-declaration,-redundant-canonical-projection,-projection-no-head-constant" +-Q examples htt +-Q htt htt +-docroot docs # where the documentation should go -theories/options.v -theories/interlude.v -theories/domain.v -theories/model.v -theories/heapauto.v +-arg -w -arg -notation-overridden +-arg -w -arg -redundant-canonical-projection + +# release-specific arguments +-arg -w -arg -notation-incompatible-prefix # specific to coq8.20.0 +-arg -w -arg -deprecated-from-Coq # specific to coq8.21 +-arg -w -arg -deprecated-dirpath-Coq # specific to coq8.21 + +htt/options.v +htt/domain.v +htt/model.v +htt/heapauto.v examples/exploit.v examples/gcd.v +examples/counter.v examples/llist.v -examples/array.v -examples/bubblesort.v -examples/quicksort.v -examples/stack.v examples/dlist.v +examples/array.v examples/queue.v examples/cyclic.v +examples/stack.v examples/bintree.v examples/bst.v examples/kvmaps.v examples/hashtab.v -examples/counter.v +examples/bubblesort.v +examples/quicksort.v examples/congmath.v -examples/tree.v +examples/congprog.v +examples/tree.v examples/union_find.v diff --git a/coq-htt-examples.opam b/coq-htt-core.opam similarity index 86% rename from coq-htt-examples.opam rename to coq-htt-core.opam index caeded0..95ee94a 100644 --- a/coq-htt-examples.opam +++ b/coq-htt-core.opam @@ -1,4 +1,5 @@ # This file was generated from `meta.yml`, please do not edit manually. +# Follow the instructions on https://github.com/coq-community/templates to regenerate. opam-version: "2.0" maintainer: "fcsl@software.imdea.org" @@ -30,14 +31,13 @@ variables). The connection reconciles dependent types with effects of state and establishes Separation logic as a type theory for such effects. In implementation terms, it means that HTT implements Separation logic as a shallow embedding in Coq.""" -build: [make "-j%{jobs}%"] -install: [make "install"] +build: ["dune" "build" "-p" name "-j" jobs] depends: [ - "coq" { (>= "8.15" & < "8.19~") | (= "dev") } - "coq-mathcomp-ssreflect" { (>= "1.17.0" & < "1.18~") | (= "dev") } + "dune" {>= "3.6"} + "coq" { (>= "8.19" & < "8.21~") | (= "dev") } + "coq-mathcomp-ssreflect" { (>= "2.2.0" & < "2.3~") | (= "dev") } "coq-mathcomp-algebra" - "coq-mathcomp-fingroup" - "coq-fcsl-pcm" { (>= "1.8.0" & < "1.9~") | (= "dev") } + "coq-fcsl-pcm" { (>= "2.0.0" & < "2.1~") | (= "dev") } ] tags: [ @@ -50,4 +50,5 @@ authors: [ "Aleksandar Nanevski" "Germán Andrés Delbianco" "Alexander Gryzlov" + "Marcos Grandury" ] diff --git a/coq-htt.opam b/coq-htt.opam index 1517652..7086e27 100644 --- a/coq-htt.opam +++ b/coq-htt.opam @@ -1,6 +1,4 @@ # This file was generated from `meta.yml`, please do not edit manually. -# Follow the instructions on https://github.com/coq-community/templates to regenerate. - opam-version: "2.0" maintainer: "fcsl@software.imdea.org" version: "dev" @@ -31,14 +29,13 @@ variables). The connection reconciles dependent types with effects of state and establishes Separation logic as a type theory for such effects. In implementation terms, it means that HTT implements Separation logic as a shallow embedding in Coq.""" -build: ["dune" "build" "-p" name "-j" jobs] +build: [make "-j%{jobs}%"] +install: [make "install"] depends: [ - "dune" {>= "2.5"} - "coq" { (>= "8.15" & < "8.19~") | (= "dev") } - "coq-mathcomp-ssreflect" { (>= "1.17.0" & < "1.18~") | (= "dev") } + "coq" { (>= "8.19" & < "8.21~") | (= "dev") } + "coq-mathcomp-ssreflect" { (>= "2.2.0" & < "2.3~") | (= "dev") } "coq-mathcomp-algebra" - "coq-mathcomp-fingroup" - "coq-fcsl-pcm" { (>= "1.8.0" & < "1.9~") | (= "dev") } + "coq-fcsl-pcm" { (>= "2.0.0" & < "2.1~") | (= "dev") } ] tags: [ @@ -51,4 +48,5 @@ authors: [ "Aleksandar Nanevski" "Germán Andrés Delbianco" "Alexander Gryzlov" + "Marcos Grandury" ] diff --git a/dune-project b/dune-project index bf8c8f3..4d97b3b 100644 --- a/dune-project +++ b/dune-project @@ -1,3 +1,3 @@ -(lang dune 2.5) -(using coq 0.2) +(lang dune 3.6) +(using coq 0.6) (name htt) diff --git a/examples/array.v b/examples/array.v index 42cbe10..d79fb40 100644 --- a/examples/array.v +++ b/examples/array.v @@ -1,42 +1,78 @@ +(* +Copyright 2010 IMDEA Software Institute +Licensed under the Apache License, Version 2.0 (the "License"); +you may not use this file except in compliance with the License. +You may obtain a copy of the License at + http://www.apache.org/licenses/LICENSE-2.0 +Unless required by applicable law or agreed to in writing, software +distributed under the License is distributed on an "AS IS" BASIS, +WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. +See the License for the specific language governing permissions and +limitations under the License. +*) + From Coq Require Import ssreflect ssrbool ssrfun. From mathcomp Require Import ssrnat eqtype seq path fintype tuple finfun finset. From pcm Require Import options axioms prelude pred. From pcm Require Import pcm unionmap heap. From htt Require Import options model heapauto. -Definition indx {I : finType} (x : I) := index x (enum I). - -Prenex Implicits indx. - -(***********************************) -(* Arrays indexed by a finite type *) -(***********************************) +(*********************************) +(* Arrays indexed by finite type *) +(*********************************) -(* an array is a pointer to a contiguous memory region holding its values *) +(* array is (pointer to) a contiguous memory region *) +(* holding the array values *) -Record array (I : finType) (T : Type) : Type := Array {orig :> ptr}. -Arguments Array [I T]. +Record array (I : finType) (T : Type) : Set := Array {orig :> ptr}. +Arguments Array {I T}. Definition array_for (I : finType) (T : Type) of phant (I -> T) := array I T. Identity Coercion array_for_array : array_for >-> array. Notation "{ 'array' aT }" := (array_for (Phant aT)) (at level 0, format "{ 'array' '[hv' aT ']' }") : type_scope. -Module Array. +Module Type ArraySig. +Parameter shape : forall {I : finType} {T : Type}, + {array I -> T} -> {ffun I -> T} -> Pred heap. + +(* build new array with all cells initialized to x *) +Parameter new : forall {I : finType} {T : Type} (x : T), + STsep (emp, [vfun (a : {array I -> T}) h => h \In shape a [ffun => x]]). + +(* build new array with cells initialized by finite function f *) +Parameter newf : forall {I : finType} {T : Type} (f : {ffun I -> T}), + STsep (emp, [vfun a h => h \In shape a f]). + +(* free the array *) +Parameter free : forall {I : finType} {T : Type} (a : {array I -> T}), + STsep (fun i => exists f, i \In shape a f, [vfun _ : unit => emp]). + +(* read k-th cell *) +Parameter read : forall {I : finType} {T : Type} (a : {array I -> T}) (k : I), + STsep {f h} (fun i => i = h /\ i \In shape a f, + [vfun (y : T) m => m = h /\ y = f k]). + +(* write k-th cell *) +Parameter write : forall {I : finType} {T : Type} (a : {array I -> T}) k x, + STsep {f} (shape a f, + [vfun (_ : unit) h => h \In shape a (finfun [eta f with k |-> x])]). +End ArraySig. + + +Module Array : ArraySig. Section Array. -Variable (I : finType) (T : Type). +Context {I : finType} {T : Type}. Notation array := {array I -> T}. -(* an array is specified by a finite function *) - -Definition shape (a : array) (f : {ffun I -> T}) := - [Pred h | h = updi a (fgraph f) /\ valid h]. +(* array is specified by finite function *) +Definition shape (a : array) (f : {ffun I -> T}) : Pred heap := + [Pred h | h = updi a (fgraph f)]. (* main methods *) -(* a new empty array preallocates all cells for all possible index values *) +(* new empty array preallocates all cells for all possible index values *) (* initializing all of them to `x` *) - Program Definition new (x : T) : STsep (emp, [vfun a => shape a [ffun => x]]) := Do (x <-- allocb x #|I|; @@ -45,22 +81,21 @@ Next Obligation. (* pull out ghost vars, run the program *) move=>x [] _ /= ->; step=>y; step. (* simplify *) -rewrite unitR=>H; split=>//; congr updi; rewrite codomE cardE. +rewrite unitR=>_; congr updi; rewrite /= codomE cardE. by elim: (enum I)=>[|t ts] //= ->; rewrite (ffunE _ _). Qed. -Opaque new. +(* new array corresponding to a domain of a finite function f *) -(* a new array corresponding to a domain of a finite function f *) - -(* the loop invariant: *) -(* a partially filled array corresponds to some finite function g acting as a prefix of f *) +(* loop invariant: *) +(* partially filled array corresponds to finite function g *) +(* acting as prefix of f *) Definition fill_loop a (f : {ffun I -> T}) : Type := forall s : seq I, STsep (fun i => exists g s', [/\ i \In shape a g, s' ++ s = enum I & forall x, x \in s' -> g x = f x], - [vfun a => shape a f]). + [vfun a => shape a f]). Program Definition newf (f : {ffun I -> T}) : STsep (emp, [vfun a => shape a f]) := @@ -70,21 +105,21 @@ Program Definition newf (f : {ffun I -> T}) : (* preallocate a new array *) x <-- new (f v); (* fill it with values of f on I *) - let fill := Fix (fun (loop : fill_loop x f) s => - Do (if s is k::t return _ then - x .+ (indx k) ::= f k;; - loop t - else ret (Array x))) + let fill := ffix (fun (loop : fill_loop x f) s => + Do (if s is k::t return _ then + x .+ (indx k) ::= f k;; + loop t + else ret x)) in fill (enum I) else ret (Array null)). (* first the loop *) Next Obligation. (* pull out preconditions (note that there are no ghost vars), split *) -move=>f v x loop s [] /= _ [g][s'][[-> _]]; case: s=>[|k t] /= H1 H2. +move=>f v x loop s [] /= _ [g][s'][->]; case: s=>[|k t] /= H1 H2. - (* we've reached the end, return the array *) - rewrite cats0 in H1; step=>H; split=>//. (* g spans the whole f *) - by rewrite (_ : g = f) // -ffunP=>y; apply: H2; rewrite H1 mem_enum. + rewrite cats0 in H1; step=>_; rewrite (_ : g = f) // -ffunP=>y. + by apply: H2; rewrite H1 mem_enum. (* run the loop iteration, split the heap and save its validity *) rewrite (updi_split x k); step; apply: vrfV=>V; apply: [gE]=>//=. (* add the new index+value to g *) @@ -100,11 +135,11 @@ Next Obligation. (* pull out params, check if I is empty *) move=>f [] _ ->; case: fintype.pickP=>[v|] H /=. - (* run the `new` subroutine, simplify *) - apply: [stepE]=>//= a _ [-> V]. + apply: [stepE]=>//= a _ ->. (* invoke the loop, construct g from the first value of f *) by apply: [gE]=>//=; exists [ffun => f v], nil. (* I is empty, so should be the resulting heap *) -step=>_; split=>//; rewrite codom_ffun. +step; rewrite /shape /= codom_ffun. suff L: #|I| = 0 by case: (fgraph f)=>/=; rewrite L; case. by rewrite cardE; case: (enum I)=>[|x s] //; move: (H x). Qed. @@ -114,20 +149,20 @@ Qed. (* the loop invariant: *) (* a partially freed array still contains valid #|I| - k cells *) (* corresponding to some suffix xs of the original array's spec *) -Definition free_loop (a : array) : Type := +Definition free_loop (a : ptr) : Type := forall k, STsep (fun i => exists xs: seq T, [/\ i = updi (a .+ k) xs, valid i & size xs + k = #|I|], - [vfun _ : unit => emp]). + [vfun _ : unit => emp]). Program Definition free (a : array) : STsep (fun i => exists f, i \In shape a f, - [vfun _ : unit => emp]) := - Do (let go := Fix (fun (loop : free_loop a) k => - Do (if k == #|I| then ret tt - else dealloc a.+k;; - loop k.+1)) + [vfun _ : unit => emp]) := + Do (let go := ffix (fun (loop : free_loop a) k => + Do (if k == #|I| then ret tt + else dealloc a.+k;; + loop k.+1)) in go 0). (* first the loop *) Next Obligation. @@ -136,40 +171,38 @@ move=>a loop k [] i /= [[|v xs]][->] /= _; first by rewrite add0n=>/eqP ->; step (* the suffix is non-empty so k < #|I| *) case: eqP=>[->|_ H]; first by move/eqP; rewrite -{2}(add0n #|I|) eqn_add2r. (* run the program, simplify *) -step; apply: vrfV=>V; apply: [gE]=>//=. -by exists xs; rewrite V ptrA addn1 -addSnnS unitL. +step; apply: vrfV=>V; apply: [gE]=>//=; exists xs. +by rewrite V unitL -addSnnS H -addnS. Qed. (* now the outer program *) Next Obligation. (* pull out params, invoke the loop *) -move=>a [] /= _ [f][-> V]; apply: [gE]=>//=. +move=>a [] /= _ [f ->]; apply: vrfV=>V; apply: [gE]=>//=. (* the suffix xs is the whole codomain of f *) exists (tval (fgraph f))=>/=. by rewrite ptr0 V {3}codomE size_map -cardE addn0. Qed. (* reading from an array, doesn't modify the heap *) - Program Definition read (a : array) (k : I) : - {f h}, STsep (fun i => i = h /\ i \In shape a f, - [vfun (y : T) m => m = h /\ y = f k]) := + STsep {f h} (fun i => i = h /\ i \In shape a f, + [vfun (y : T) m => m = h /\ y = f k]) := Do (!a .+ (indx k)). Next Obligation. (* pull out ghost vars *) -move=>a k [f][_][] _ [->][/= -> _]. +move=>a k [f][_][] _ [->->]. (* split the heap and run the program *) by rewrite (updi_split a k); step. Qed. (* writing to an array, updates the spec function with a new value *) - Program Definition write (a : array) (k : I) (x : T) : - {f}, STsep (shape a f, - [vfun _ : unit => shape a (finfun [eta f with k |-> x])]) := + STsep {f} (shape a f, + [vfun _ : unit => shape a (finfun [eta f with k |-> x])]) := Do (a .+ (indx k) ::= x). Next Obligation. (* pull out ghost vars, split the heap *) -move=>a k x [f][] _ [-> _] /=; rewrite /shape !(updi_split a k). +move=>a k x [f][] _ -> /=; rewrite /shape !(updi_split a k). (* run the program, simplify *) by step; rewrite takeord dropord ffunE /= eq_refl. Qed. @@ -177,13 +210,13 @@ Qed. End Array. End Array. + (* Tabulating a finite function over another one *) (* Useful in building linked data structures that are pointed to by *) (* array elements, such as hashtables *) Section Table. - -Variables (I : finType) (T S : Type) (x : {array I -> T}) +Variables (I : finType) (T S : Type) (x : ptr) (Ps : T -> S -> Pred heap). Definition table (t : I -> T) (b : I -> S) (i : I) : Pred heap := @@ -192,7 +225,8 @@ Definition table (t : I -> T) (b : I -> S) (i : I) : Pred heap := Lemma tableP (s : {set I}) t1 t2 b1 b2 h : (forall x, x \in s -> t1 x = t2 x) -> (forall x, x \in s -> b1 x = b2 x) -> - h \In sepit s (table t1 b1) -> h \In sepit s (table t2 b2). + h \In sepit s (table t1 b1) -> + h \In sepit s (table t2 b2). Proof. move=>H1 H2 H. apply/(sepitF (f1 := table t1 b1))=>//. @@ -203,7 +237,8 @@ Lemma tableP2 (s1 s2 : {set I}) t1 t2 b1 b2 h : s1 =i s2 -> (forall x, s1 =i s2 -> x \in s1 -> t1 x = t2 x) -> (forall x, s1 =i s2 -> x \in s1 -> b1 x = b2 x) -> - h \In sepit s1 (table t1 b1) -> h \In sepit s2 (table t2 b2). + h \In sepit s1 (table t1 b1) -> + h \In sepit s2 (table t2 b2). Proof. move=>H1 H2 H3; rewrite (eq_sepit _ H1). by apply: tableP=>y; rewrite -H1=>R; [apply: H2 | apply: H3]. diff --git a/examples/bintree.v b/examples/bintree.v index 0b41d66..81552af 100644 --- a/examples/bintree.v +++ b/examples/bintree.v @@ -1,3 +1,15 @@ +(* +Copyright 2021 IMDEA Software Institute +Licensed under the Apache License, Version 2.0 (the "License"); +you may not use this file except in compliance with the License. +You may obtain a copy of the License at + http://www.apache.org/licenses/LICENSE-2.0 +Unless required by applicable law or agreed to in writing, software +distributed under the License is distributed on an "AS IS" BASIS, +WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. +See the License for the specific language governing permissions and +limitations under the License. +*) From Coq Require Import ssreflect ssrbool ssrfun. From mathcomp Require Import eqtype seq ssrnat. @@ -23,25 +35,24 @@ Fixpoint inorder {A} (t : tree A) : seq A := else [::]. (* Add an element to the rightmost branch *) - Fixpoint addr {A} (y : A) (t : tree A) : tree A := if t is Node l x r then Node l x (addr y r) else Node leaf y leaf. -(* Tree heap predicate: [left branch pointer, value, right branch pointer] *) - +(* Tree heap predicate: *) +(* [left branch pointer, value, right branch pointer] *) Fixpoint shape {A} (p : ptr) (t : tree A) := if t is Node l a r then [Pred h | exists l' r' h', - h = p :-> l' \+ (p .+ 1 :-> a \+ (p .+ 2 :-> r' \+ h')) + h = p :-> l' \+ (p.+1 :-> a \+ (p.+2 :-> r' \+ h')) /\ h' \In shape l' l # shape r' r] else [Pred h | p = null /\ h = Unit]. (* Null pointer represents a leaf *) - Lemma shape_null {A} (t : tree A) h : - valid h -> h \In shape null t -> + valid h -> + h \In shape null t -> t = leaf /\ h = Unit. Proof. move=>V; case: t=>/= [|l a r]; first by case=>_->. @@ -49,56 +60,55 @@ by case=>?[?][?][E]; rewrite E validPtUn in V. Qed. (* Non-null pointer represents a node *) - Lemma shape_cont {A} (t : tree A) p h : - p != null -> h \In shape p t -> + p != null -> + h \In shape p t -> exists l a r l' r' h', [/\ t = Node l a r, - h = p :-> l' \+ (p .+ 1 :-> a \+ (p .+ 2 :-> r' \+ h')) + h = p :-> l' \+ (p.+1 :-> a \+ (p.+2 :-> r' \+ h')) & h' \In shape l' l # shape r' r]. Proof. move=>E; case: t=>/= [|l a r]. - by case=>E0; rewrite E0 in E. -case=>l'[r'][h'][E1 E2]. -by exists l,a,r,l',r',h'. +case=>l' [r'][h'][E1 E2]. +by exists l, a, r, l', r', h'. Qed. Section Tree. Variable A : Type. (* Node creation *) - Program Definition mknode (x : A) : STsep (emp, [vfun p => shape p (Node leaf x leaf)]) := Do (n <-- allocb null 3; - n.+ 1 ::= x;; + n.+1 ::= x;; ret n). Next Obligation. (* the starting heap is empty *) move=>x [] _ /= ->. (* run all the steps *) -step=>n; rewrite !unitR ptrA; step; step=>_. +step=>n; rewrite !unitR; step; step=>_. (* the postcondition is satisfied *) by exists null, null, Unit; vauto; rewrite unitR. Qed. (* Recursive tree disposal *) -(* We start from a well-formed tree and arrive at an empty heap *) +(* Start from a well-formed tree, and arrive at empty heap *) Definition disposetreeT : Type := - forall p, {t : tree A}, STsep (shape p t, [vfun _ : unit => emp]). + forall p, STsep {t : tree A} (shape p t, [vfun _ : unit => emp]). Program Definition disposetree : disposetreeT := - Fix (fun (loop : disposetreeT) p => + ffix (fun (loop : disposetreeT) p => Do (if p == null then ret tt else l <-- !p; - r <-- !(p.+ 2); + r <-- !p.+2; loop l;; loop r;; dealloc p;; - dealloc p.+ 1;; - dealloc p.+ 2 + dealloc p.+1;; + dealloc p.+2 )). Next Obligation. (* pull out ghost var + precondition, branch on null check *) @@ -115,31 +125,31 @@ apply: [stepX r]@hr=>//= _ _ ->; rewrite unitR. by do 3!step; rewrite !unitL. Qed. -(* Calculation of tree size *) +(* Calculate tree size *) (* loop invariant: *) (* the subtree size is added to the accumulator *) Definition treesizeT : Type := forall (ps : ptr * nat), - {t : tree A}, STsep (shape ps.1 t, - [vfun s h => s == ps.2 + size_tree t /\ shape ps.1 t h]). + STsep {t : tree A} (shape ps.1 t, + [vfun s h => s == ps.2 + size_tree t /\ shape ps.1 t h]). Program Definition treesize p : - {t : tree A}, STsep (shape p t, - [vfun s h => s == size_tree t /\ shape p t h]) := - Do (let len := Fix (fun (go : treesizeT) '(p, s) => + STsep {t : tree A} (shape p t, + [vfun s h => s == size_tree t /\ shape p t h]) := + Do (let len := ffix (fun (go : treesizeT) '(p, s) => Do (if p == null then ret s else l <-- !p; - r <-- !(p.+ 2); + r <-- !p.+2; ls <-- go (l, s); go (r, ls.+1))) in len (p, 0)). Next Obligation. (* pull out ghost var + precondition, branch on null check *) move=>_ go _ p s /= [t][] i /= H; case: eqP H=>[{p}->|/eqP Ep] H. -- (* empty tree has size 0 *) - by step=>V; case: (shape_null V H)=>->->/=; rewrite addn0. +(* empty tree has size 0 *) +- by step=>V; case: (shape_null V H)=>->->/=; rewrite addn0. (* non-null pointer is a node, deconstruct it, read branch pointers *) - case: (shape_cont Ep H)=>l[a][r][l'][r'][_][{t H}-> {i}-> [hl][hr][-> Hl Hr]] /=. +case: (shape_cont Ep H)=>l[a][r][l'][r'][_][{t H}-> {i}-> [hl][hr][-> Hl Hr]]. do 2!step. (* calculate left branch size, update the accumulator *) apply: [stepX l]@hl=>//= _ hl' [/eqP -> Hl']. @@ -149,57 +159,57 @@ apply: [gX r]@hr=>//= _ hr' [/eqP -> Hr'] _. by split; [rewrite addnS addSn addnA | vauto]. Qed. Next Obligation. -(* pull out ghost var + precondition, start the loop with empty accumulator *) +(* pull out ghost var + precondition, start loop with empty accumulator *) by move=>/= p [t][] i /= H; apply: [gE t]. Qed. (* Tree in-order traversal, storing the values in a linked list *) -Opaque insert new. - (* loop invariant: *) (* the subtree is unchanged, its values are prepended to the accumulator list *) Definition inordertravT : Type := forall (ps : ptr * ptr), - {(t : tree A) (l : seq A)}, - STsep (shape ps.1 t # lseq ps.2 l, - [vfun s h => h \In shape ps.1 t # lseq s (inorder t ++ l)]). + STsep {(t : tree A) (l : seq A)} + (shape ps.1 t # lseq ps.2 l, + [vfun s h => h \In shape ps.1 t # lseq s (inorder t ++ l)]). Program Definition inordertrav p : - {t : tree A}, STsep (shape p t, - [vfun s h => h \In shape p t # lseq s (inorder t)]) := - Do (let loop := Fix (fun (go : inordertravT) '(p, s) => + STsep {t : tree A} (shape p t, + [vfun s h => h \In shape p t # lseq s (inorder t)]) := + Do (let loop := ffix (fun (go : inordertravT) '(p, s) => Do (if p == null then ret s else l <-- !p; - a <-- !(p.+ 1); - r <-- !(p.+ 2); + a <-- !p.+1; + r <-- !p.+2; s1 <-- go (r, s); s2 <-- insert s1 (a : A); go (l, s2))) in n <-- new A; loop (p, n)). Next Obligation. -(* pull out ghosts + precondition, deconstruct the heap, branch on null check *) +(* pull out ghosts + precondition, destruct heap, branch on null check *) move=>_ go _ p s /= [t][xs][] _ /= [h1][h2][-> H1 H2]. case: eqP H1=>[{p}->|/eqP Ep] H1. -- (* return the accumulated list - empty tree has no values *) - by step=>V; case: (shape_null (validL V) H1)=>->->/=; vauto. -(* non-empty tree is a node, deconstruct the node, read the pointers and the value *) -case: (shape_cont Ep H1)=>l[a][r][l'][r'][_][{t H1}-> {h1}-> [hl][hr][-> Hl Hr]] /=. -do 3!step. -(* run traversal on the right branch first (it's cheaper to grow a linked list to the left) *) +(* return the accumulated list - empty tree has no values *) +- by step=>V; case: (shape_null (validL V) H1)=>->->/=; vauto. +(* non-empty tree is a node *) +(* deconstruct the node, read the pointers and the value *) +case: (shape_cont Ep H1)=>l[a][r][l'][r'][_][{t H1}-> {h1}-> +[hl][hr][-> Hl Hr]] /=; do 3!step. +(* run traversal on the right branch first *) +(* (it's cheaper to grow a linked list to the left) *) apply: [stepX r, xs]@(hr \+ h2)=>//=; first by vauto. -(* we have subheaps corresponding to the left+right branches and the updated list *) +(* subheaps exist corresponding to left/right branches and updated list *) move=>s1 _ [hr'][hs][-> Hr' Hs]. (* prepend the node value to the list *) apply: [stepX (inorder r ++ xs)]@hs=>//= pa _ [s2][h'][-> H']. (* run the traversal on the left branch with updated list *) -apply: [gX l, (a::inorder r ++ xs)]@(hl \+ pa :-> a \+ pa.+ 1 :-> s2 \+ h')=>//=. -- (* the precondition is satisfied by simple heap manipulation *) - move=>_; exists hl, (pa :-> a \+ (pa.+ 1 :-> s2 \+ h')). +apply: [gX l, (a::inorder r ++ xs)]@(hl \+ pa :-> a \+ pa.+1 :-> s2 \+ h')=>//=. +(* the precondition is satisfied by simple heap manipulation *) +- exists hl, (pa :-> a \+ (pa.+1 :-> s2 \+ h')). by split=>//=; [rewrite !joinA | vauto]. (* the postcondition is also satisfied by simple massaging *) move=>s3 _ [hl''][hs'][-> Hl'' Hs'] _. -exists (p :-> l' \+ (p.+ 1 :-> a \+ (p.+ 2 :-> r' \+ (hl'' \+ hr')))), hs'. +exists (p :-> l' \+ (p.+1 :-> a \+ (p.+2 :-> r' \+ (hl'' \+ hr')))), hs'. split; try by hhauto. by rewrite -catA. Qed. @@ -217,21 +227,19 @@ Qed. (* Expanding the tree to the right *) -Opaque mknode. - (* loop invariant: the value is added to the rightmost branch *) Definition expandrightT x : Type := forall (p : ptr), - {t : tree A}, STsep (shape p t, - [vfun p' => shape p' (addr x t)]). + STsep {t : tree A} (shape p t, + [vfun p' => shape p' (addr x t)]). Program Definition expandright x : expandrightT x := - Fix (fun (go : expandrightT x) p => + ffix (fun (go : expandrightT x) p => Do (if p == null then n <-- mknode x; ret n - else pr <-- !(p.+ 2); + else pr <-- !p.+2; p' <-- go pr; - p.+ 2 ::= p';; + p.+2 ::= p';; ret p)). Next Obligation. (* pull out ghost + precondition, branch on null check *) diff --git a/examples/bst.v b/examples/bst.v index e61fabf..3deedab 100644 --- a/examples/bst.v +++ b/examples/bst.v @@ -1,58 +1,69 @@ +(* +Copyright 2021 IMDEA Software Institute +Licensed under the Apache License, Version 2.0 (the "License"); +you may not use this file except in compliance with the License. +You may obtain a copy of the License at + http://www.apache.org/licenses/LICENSE-2.0 +Unless required by applicable law or agreed to in writing, software +distributed under the License is distributed on an "AS IS" BASIS, +WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. +See the License for the specific language governing permissions and +limitations under the License. +*) + From Coq Require Import ssreflect ssrbool ssrfun. From mathcomp Require Import eqtype ssrnat seq path. -From pcm Require Import options axioms pred ordtype. +From pcm Require Import options axioms pred ordtype seqext. From pcm Require Import pcm unionmap heap autopcm automap. -From htt Require Import interlude model heapauto. -From htt Require Import bintree. +From htt Require Import options model heapauto bintree. Section BST. Context {T : ordType}. -(* A binary _search_ tree remains a binary tree structurally, plus: *) -(* 1. its elements must have an ordering defined on them *) -(* 2. a recursive invariant should be satisfied: *) -(* all left branch elements are smaller than the node value & *) -(* all right branch elements are larger than the node value *) +(* Binary _search_ tree is a binary tree structurally, plus: *) +(* 1. its elements have an ordering defined on them *) +(* 2. the following recursive invariant is satisfied: *) +(* - all left branch elements are smaller than the node value & *) +(* - all right branch elements are larger than the node value *) (* Search tree operations *) Fixpoint insert x (t : tree T) : tree T := - if t is Node l a r - then - if x == a then Node l a r - else if ord x a then Node (insert x l) a r - else Node l a (insert x r) + if t is Node l a r then + if x == a then Node l a r + else if ord x a then Node (insert x l) a r + else Node l a (insert x r) else Node leaf x leaf. Fixpoint search (t : tree T) x : bool := - if t is Node l a r - then - if x == a then true - else if ord x a then search l x - else search r x + if t is Node l a r then + if x == a then true + else if ord x a then search l x + else search r x else false. (* Search tree invariant & its interaction with operations *) Fixpoint bst (t : tree T) : bool := - if t is Node l a r - then [&& all (ord^~ a) (inorder l), all (ord a) (inorder r), bst l & bst r] - else true. + if t is Node l a r then + [&& all (ord^~ a) (inorder l), all (ord a) (inorder r), bst l & bst r] + else true. -(* Both BSTs and sorted lists can be used to implement lookup structures, but trees *) -(* are more efficient computationally, while lists are simpler to reason about, *) -(* since the operations on them are associative and commute in many cases. We can *) -(* switch between two specification styles via the in-order traversal. *) +(* BSTs and sorted lists can both be used to implement lookup structures *) +(* but trees are more efficient computationally, while lists are simpler *) +(* to reason about, since the operations on them are associative and *) +(* commute in many cases. One can switch between two specification styles *) +(* via in-order traversal. *) Lemma bst_to_sorted (t : tree T) : bst t = sorted ord (inorder t). Proof. elim: t=>//=l -> a r ->. -by rewrite sorted_cat_cons_cat /= cats1 sorted_rconsE // +by rewrite sorted_cat_cons_cat /= cats1 sorted_rconsE //= (path_sortedE (@trans T)) andbACA -andbA. Qed. -(* An in-order specification for insertion *) +(* In-order specification for insertion *) Lemma inorder_insert x (t : tree T) : bst t -> perm_eq (inorder (insert x t)) @@ -61,25 +72,22 @@ Proof. elim: t=>//=l IHl a r IHr /and4P [Hal Har /IHl Hl /IHr Hr] {IHl IHr}. rewrite mem_cat inE; case: (ifP [|| _, _ | _]). - case/or3P=>H. - - rewrite H in Hl. - move/allP: Hal=>/(_ x H) /[dup] Hxa /ord_neq/negbTE ->; rewrite Hxa /=. - by apply/permEl/perm_catr. + - rewrite H in Hl; move/allP: Hal=>/(_ x H) /[dup] Hxa /ord_neq/negbTE ->. + by rewrite Hxa; apply/permEl/perm_catr. - by rewrite H. - rewrite H in Hr. - move/allP: Har=>/(_ x H) /[dup] /nsym/negP/negbTE ->. + rewrite H in Hr; move/allP: Har=>/(_ x H) /[dup] /nsym/negP/negbTE ->. move/ord_neq; rewrite eq_sym =>/negbTE -> /=. by apply/permEl/perm_catl; rewrite perm_cons. -move/negbT; rewrite !negb_or; case/and3P=>/negbTE Hxl /negbTE -> /negbTE Hxr; +move/negbT; rewrite !negb_or; case/and3P=>/negbTE Hxl /negbTE -> /negbTE Hxr. rewrite {}Hxl in Hl; rewrite {}Hxr in Hr. -case: ifP=>/= H. -- by rewrite -cat_cons; apply/permEl/perm_catr. +case: ifP=>/= H; first by rewrite -cat_cons; apply/permEl/perm_catr. rewrite -(cat1s x) -(cat1s a) -(cat1s a (inorder r)). rewrite perm_sym perm_catC -!catA catA perm_sym catA. apply/permEl/perm_catl; apply: (perm_trans Hr). by rewrite cats1 -perm_rcons. Qed. -(* As a corollary, we can show that insertion preserves the tree invariant *) +(* Corollary: insertion preserves the tree invariant *) Lemma bst_insert x (t : tree T) : bst t -> bst (insert x t). Proof. elim: t=>//=l IHl a r IHr /and4P [Hal Har Hl Hr]. @@ -124,9 +132,10 @@ by apply: (perm_trans Hi2); rewrite Hx2. Qed. (* Moreover, such representations are equal *) -Corollary insert_insert x1 x2 (t : tree T) : - bst t -> - inorder (insert x1 (insert x2 t)) = inorder (insert x2 (insert x1 t)). +Lemma insert_insert x1 x2 (t : tree T) : + bst t -> + inorder (insert x1 (insert x2 t)) = + inorder (insert x2 (insert x1 t)). Proof. move=>H; apply: ord_sorted_eq. - by rewrite -bst_to_sorted; do 2!apply: bst_insert. @@ -152,36 +161,36 @@ Qed. (* Pointer-based procedures *) -Opaque mknode. - (* Inserting into the BST *) Definition inserttreeT x : Type := forall p, - {t : tree T}, STsep (shape p t, - [vfun p' => shape p' (insert x t)]). + STsep {t : tree T} (shape p t, + [vfun p' => shape p' (insert x t)]). Program Definition inserttree x : inserttreeT x := - Fix (fun (go : inserttreeT x) p => + ffix (fun (go : inserttreeT x) p => Do (if p == null then n <-- mknode x; ret n - else a <-- !(p.+ 1); + else a <-- !p.+1; if x == a then ret p - else if ord x a - then l <-- !p; - l' <-- go l; - p ::= l';; - ret p - else r <-- !(p.+ 2); - r' <-- go r; - p.+ 2 ::= r';; - ret p)). + else + if ord x a then + l <-- !p; + l' <-- go l; + p ::= l';; + ret p + else + r <-- !p.+2; + r' <-- go r; + p.+2 ::= r';; + ret p)). Next Obligation. (* pull out ghost + precondition, branch on null check *) move=>x go p [t][] i /= H; case: eqP H=>[{p}->|/eqP E] H. -- (* the tree is empty, make a new node *) - apply: vrfV=>V; case: (shape_null V H)=>{t H}->{i V}->. +(* the tree is empty, make a new node *) +- apply: vrfV=>V; case: (shape_null V H)=>{t H}->{i V}->. by apply: [stepE]=>// n m H; step. (* the tree is a node, deconstruct it *) case: (shape_cont E H)=>l[z][r][pl][pr][_][{t H}->{i}->][hl][hr][-> Hl Hr]. @@ -190,8 +199,8 @@ case: (shape_cont E H)=>l[z][r][pl][pr][_][{t H}->{i}->][hl][hr][-> Hl Hr]. step; case: eqP=>_; first by step; vauto. (* branch on comparison, read corresponding pointer *) case: ifP=>Ho; step. -- (* insert in the left branch, update the left pointer *) - apply: [stepX l]@hl=>//= p' h' H'. +(* insert in the left branch, update the left pointer *) +- apply: [stepX l]@hl=>//= p' h' H'. by do 2!step; vauto. (* insert in the right branch, update the right pointer *) apply: [stepX r]@hr=>//= p' h' H'. @@ -203,33 +212,33 @@ Qed. (* lopp invariant: the tree is unchanged, return true if the element is found *) Definition searchtreeT x : Type := forall p, - {t : tree T}, STsep (shape p t, - [vfun b h => shape p t h /\ b == search t x]). + STsep {t : tree T} (shape p t, + [vfun b h => shape p t h /\ b == search t x]). Program Definition searchtree x : searchtreeT x := - Fix (fun (go : searchtreeT x) p => - Do (if p == null - then ret false - else a <-- !(p.+ 1); + ffix (fun (go : searchtreeT x) p => + Do (if p == null then ret false + else a <-- !p.+1; if x == a then ret true - else if ord x a - then l <-- !p; - go l - else r <-- !(p.+ 2); - go r)). + else if ord x a then + l <-- !p; + go l + else + r <-- !p.+2; + go r)). Next Obligation. (* pull out ghost + precondition, branch on null check *) move=>x go p [t][] i /= H; case: eqP H=>[{p}->|/eqP E] H. -- (* tree is empty, it can't contain anything *) - by apply: vrfV=>V; case: (shape_null V H)=>{t H}->{i V}->; step. +(* tree is empty, it can't contain anything *) +- by apply: vrfV=>V; case: (shape_null V H)=>{t H}->{i V}->; step. (* the tree is a node, deconstruct it *) case: (shape_cont E H)=>l[z][r][pl][pr][_][{t H}->{i}->][hl][hr][-> Hl Hr]. (* read the value, compare it to the element, return true if it matches *) step; case: eqP=>_; first by step; vauto. (* branch on comparison, read corresponding pointer *) case: ifP=>Ho; step. -- (* loop on the left branch *) - by apply: [gX l]@hl=>//= b h' [H' E'] _; vauto. +(* loop on the left branch *) +- by apply: [gX l]@hl=>//= b h' [H' E'] _; vauto. (* loop on the right branch *) by apply: [gX r]@hr=>//= b h' [H' E'] _; vauto. Qed. @@ -237,10 +246,10 @@ Qed. (* test that the API is consistent, i.e. BST invariant is preserved *) (* and lookup finds previously inserted elements *) Program Definition test p x1 x2 : - {t : tree T}, STsep (fun h => shape p t h /\ bst t, - [vfun (pb : ptr * bool) h => - let t' := insert x2 (insert x1 t) in - [/\ shape pb.1 t' h, bst t' & pb.2]]) := + STsep {t : tree T} (fun h => shape p t h /\ bst t, + [vfun (pb : ptr * bool) h => + let t' := insert x2 (insert x1 t) in + [/\ shape pb.1 t' h, bst t' & pb.2]]) := Do (p1 <-- inserttree x1 p; p2 <-- inserttree x2 p1; b <-- searchtree x1 p2; @@ -256,8 +265,9 @@ step=>_. (* insertions preserve the invariant *) have Hi2: bst (insert x2 t) by apply: bst_insert. have Hi21: bst (insert x2 (insert x1 t)) by do 2!apply: bst_insert. -(* at this point we're done with the separation logic part, i.e. the mutable state *) -(* the only non-trivial goal remaining is search being consistent with insert *) +(* separation logic part (i.e. part about mutable state) is done *) +(* the only remaining non-trivial goal is showing that *) +(* search is consistent with insert *) split=>//{p2 h3 H3}; rewrite {b}Hb. (* switch to in-order lookup *) move: (inorder_search Hi21 x1); rewrite -topredE /= =>->. diff --git a/examples/bubblesort.v b/examples/bubblesort.v index fe2731f..8fb6810 100644 --- a/examples/bubblesort.v +++ b/examples/bubblesort.v @@ -1,25 +1,37 @@ +(* +Copyright 2022 IMDEA Software Institute +Licensed under the Apache License, Version 2.0 (the "License"); +you may not use this file except in compliance with the License. +You may obtain a copy of the License at + http://www.apache.org/licenses/LICENSE-2.0 +Unless required by applicable law or agreed to in writing, software +distributed under the License is distributed on an "AS IS" BASIS, +WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. +See the License for the specific language governing permissions and +limitations under the License. +*) + From Coq Require Import ssreflect ssrbool ssrfun. -From mathcomp Require Import ssrnat eqtype seq path fintype tuple finfun fingroup perm interval. +From mathcomp Require Import ssrnat eqtype seq path fintype tuple. +From mathcomp Require Import finfun fingroup perm interval. From pcm Require Import options axioms prelude pred ordtype slice. -From pcm Require Import pcm unionmap heap. -From htt Require Import options interlude model heapauto. +From pcm Require Import seqext pcm unionmap heap. +From htt Require Import options model heapauto. From htt Require Import array. - -(* hack to avoid "_ *p _" notation clash *) From mathcomp Require order. Import order.Order.NatOrder order.Order.TTheory. +Local Open Scope order_scope. -(* Mathematics of (bubble) array sorting: *) -(* A theory of permutations built out of (adjacent-element) swaps acting on *) -(* finite functions from bounded nats to ordered values. *) +(* Brief mathematics of (bubble) array sorting: *) +(* Theory of permutations built out of (adjacent-element) swaps acting on *) +(* finite functions from bounded nats to ordered values. *) (* Shorthands for previous/next indices *) (* TODO use fintype.lift instead ? *) Section OrdArith. (* widen by 1 *) -Definition Wo n : 'I_n -> 'I_n.+1 := - widen_ord (leqnSn _). +Definition Wo n : 'I_n -> 'I_n.+1 := widen_ord (leqnSn _). (* increase by 1 *) Definition So n : 'I_n -> 'I_n.+1 := @@ -33,8 +45,8 @@ Proof. by case: i. Qed. Lemma So_WoN n (i : 'I_n) : So i != Wo i. Proof. -case: i=>/= m prf; rewrite /Wo /widen_ord /=; apply/negP=>/eqP; case=>/eqP. -by rewrite eqn_leq leqnSn andbT ltnn. +case: i=>/= m prf; rewrite /Wo /widen_ord /=. +by apply/negP=>/eqP [/eqP]; rewrite eqn_leq leqnSn andbT ltnn. Qed. (* decrease by 1 *) @@ -57,34 +69,31 @@ Proof. by case: i prf=>/= m prf0 prf1; apply/ord_inj. Qed. (* subtract 2 and decrease bound by 1 *) Definition M2lo n (i : 'I_n) (pi : 1%N < i) : 'I_n.-1. Proof. -case: i pi=>m prf /= Hm. -apply: (Ordinal (n:=n.-1) (m:=m.-2)). +case: i pi=>m prf /= Hm; apply: (Ordinal (n:=n.-1) (m:=m.-2)). rewrite ltn_predRL; case: m prf Hm=>//=; case=>//= m Hm _. by apply: ltnW. Defined. Lemma M2lo_eq n (i : 'I_n) (prf : 1%N < i) : nat_of_ord (M2lo prf) = i.-2. Proof. by case: i prf. Qed. - End OrdArith. (* Taking current/next element *) Section OnthCodom. -Variable (A : Type). +Variable A : Type. -Corollary onth_codom1 {n} (i : 'I_n) (f: {ffun 'I_n.+1 -> A}) : - onth (fgraph f) i = Some (f (Wo i)). +Lemma onth_codom1 {n} (i : 'I_n) (f: {ffun 'I_n.+1 -> A}) : + onth (fgraph f) i = Some (f (Wo i)). Proof. by rewrite (onth_codom (Wo i)). Qed. -Corollary onth_codom1S {n} (i : 'I_n) (f: {ffun 'I_n.+1 -> A}) : - onth (fgraph f) i.+1 = Some (f (So i)). +Lemma onth_codom1S {n} (i : 'I_n) (f: {ffun 'I_n.+1 -> A}) : + onth (fgraph f) i.+1 = Some (f (So i)). Proof. by rewrite -(onth_codom (So i)) /= So_eq. Qed. - End OnthCodom. (* Slicing on finfun codomains *) Section CodomSlice. -Variable (n : nat) (A : Type). +Variables (n : nat) (A : Type). Lemma codom1_split (f : {ffun 'I_n.+1 -> A}) (i : 'I_n) : codom f = &:(fgraph f) `]-oo, i : nat[ ++ @@ -92,10 +101,10 @@ Lemma codom1_split (f : {ffun 'I_n.+1 -> A}) (i : 'I_n) : &:(fgraph f) `]i.+1, +oo[. Proof. set i0 := Wo i; set i1 := So i. -rewrite {1}codomE (enum_split i0) /= {2}(enum_split i1) /= /heap.indx - (index_enum_ord i0) (index_enum_ord i1) drop_cat size_take - size_enum_ord ltn_ord So_eq ltnn subnn /= map_cat /= map_take map_drop -codomE. -by rewrite /slice /= addn0 addn1 /= drop0 take_size. +rewrite {1}codomE (enum_split i0) /= {2}(enum_split i1) /= /indx + (index_enum_ord i0) (index_enum_ord i1) drop_cat size_take size_enum_ord + ltn_ord So_eq ltnn subnn /= map_cat /= map_take map_drop -codomE. +by rewrite /slice.slice/= addn0 addn1 /= drop0 take_size. Qed. (* TODO generalize? *) @@ -106,14 +115,15 @@ Proof. by rewrite slice_xL // onth_codom1S. Qed. (* TODO notation fails *) Lemma codom1_ax_rcons (f : {ffun 'I_n.+1 -> A}) (a : itv_bound nat) (i : 'I_n) : order.Order.le a (BLeft (i : nat)) -> - &:(fgraph f) (Interval a (BRight (i : nat))) = + &:(fgraph f) (Interval a (BRight (i : nat))) = rcons (&:(fgraph f) (Interval a (BLeft (i : nat)))) (f (Wo i)). Proof. by move=>H; rewrite slice_xR //= onth_codom1 /=. Qed. Lemma codom1_ax_rcons2 (f : {ffun 'I_n.+1 -> A}) (a : itv_bound nat) (i : 'I_n) : order.Order.le a (BLeft (i : nat)) -> &:(fgraph f) (Interval a (BRight i.+1)) = - rcons (rcons (&:(fgraph f) (Interval a (BLeft (i : nat)))) (f (Wo i))) (f (So i)). + rcons (rcons (&:(fgraph f) (Interval a (BLeft (i : nat)))) + (f (Wo i))) (f (So i)). Proof. move=>H; rewrite slice_xR; last first. - by apply: ltW; case: a H=>[[] ax|ab] //; rewrite !bnd_simp ltEnat /= =>/ltnW. @@ -132,17 +142,15 @@ Qed. Lemma codom1_Skk (f : {ffun 'I_n.+1 -> A}) (i : 'I_n) : &:(fgraph f) `[i.+1, i.+1] = [:: f (So i)]. Proof. -rewrite slice_kk /=. -pose i' := cast_ord (esym (card_ord n.+1)) (So i). +rewrite slice_kk /=; pose i' := cast_ord (esym (card_ord n.+1)) (So i). move: (onth_tnth (codom_tuple f) i')=>/=; rewrite So_eq =>->. by move: (@tnth_fgraph _ _ f i'); rewrite enum_val_ord {2}/i' cast_ordKV=>->. Qed. - End CodomSlice. (* Sortedness on codomain slices *) Section CodomSliceOrd. -Variable (n : nat) (A : ordType). +Variables (n : nat) (A : ordType). Lemma codom1_sortedS (f: {ffun 'I_n.+1 -> A}) (i : 'I_n) : oleq (f (Wo i)) (f (So i)) -> @@ -159,14 +167,15 @@ End CodomSliceOrd. (* Swapping adjacent elements and its interaction with slicing *) Section SwapNext. -Variable (n : nat) (A : Type). +Variables (n : nat) (A : Type). Definition swnx (i : 'I_n) : 'S_n.+1 := tperm (Wo i) (So i). Lemma swnxE1 (f : {ffun 'I_n.+1 -> A}) (i : 'I_n) : pffun (swnx i) f (Wo i) = f (So i). Proof. -by rewrite /swnx ffunE; case: tpermP=>// /eqP; rewrite eq_sym (negbTE (So_WoN i)). +rewrite /swnx ffunE; case: tpermP=>// /eqP. +by rewrite eq_sym (negbTE (So_WoN i)). Qed. Lemma swnxE2 (f : {ffun 'I_n.+1 -> A}) (i : 'I_n) : @@ -204,24 +213,22 @@ move=>_ Hx; case: j Hx Hij=>[[] jx|[]] //=. by rewrite addn1 ltEnat /= ltnS leEnat =>->. Qed. -Corollary swnx_ao (f : {ffun 'I_n.+1 -> A}) (i : 'I_n) x k : - (k <= i)%N -> - &:(fgraph (pffun (swnx i) f)) (Interval x (BLeft k)) = - &:(fgraph f) (Interval x (BLeft k)). +Lemma swnx_ao (f : {ffun 'I_n.+1 -> A}) (i : 'I_n) x k : + (k <= i)%N -> + &:(fgraph (pffun (swnx i) f)) (Interval x (BLeft k)) = + &:(fgraph f) (Interval x (BLeft k)). Proof. -move=>Hk. -apply: swnx_notin; rewrite in_itv /= negb_and ltEnat /= -leqNgt; apply/orP; right=>//. -by apply: ltnW. +move=>Hk; apply: swnx_notin; rewrite in_itv /= negb_and ltEnat /= -leqNgt; +by apply/orP; right=>//; apply: ltnW. Qed. -Corollary swnx_oa (f : {ffun 'I_n.+1 -> A}) (i : 'I_n) k y : - (i.+1 <= k)%N -> - &:(fgraph (pffun (swnx i) f)) (Interval (BRight k) y) = - &:(fgraph f) (Interval (BRight k) y). +Lemma swnx_oa (f : {ffun 'I_n.+1 -> A}) (i : 'I_n) k y : + (i.+1 <= k)%N -> + &:(fgraph (pffun (swnx i) f)) (Interval (BRight k) y) = + &:(fgraph f) (Interval (BRight k) y). Proof. -move=>Hk. -apply: swnx_notin; rewrite in_itv /= negb_and ltEnat /= -leqNgt; apply/orP; left=>//. -by apply: ltnW. +move=>Hk; apply: swnx_notin; rewrite in_itv /= negb_and ltEnat /= -leqNgt; +by apply/orP; left=>//; apply: ltnW. Qed. Lemma swnx_split (f : {ffun 'I_n.+1 -> A}) (i : 'I_n) : @@ -229,8 +236,8 @@ Lemma swnx_split (f : {ffun 'I_n.+1 -> A}) (i : 'I_n) : [:: f (So i); f (Wo i)] ++ &:(fgraph f) `]i.+1, +oo[. Proof. -rewrite /= (codom1_split _ i) /=; set i0 := Wo i; set i1 := So i. -by rewrite swnx_ao // swnx_oa //= swnxE1 swnxE2. +rewrite [LHS](codom1_split _ i) /=; set i0 := Wo i; set i1 := So i. +by rewrite swnx_ao // swnx_oa // swnxE1 swnxE2. Qed. Lemma swnx_ux_rcons (f : {ffun 'I_n.+1 -> A}) (i : 'I_n) : @@ -251,12 +258,11 @@ Proof. by rewrite codom1_xu_cons swnx_oa // swnxE2. Qed. Lemma swnx_Skk (f : {ffun 'I_n.+1 -> A}) (i : 'I_n) : &:(fgraph (pffun (swnx i) f)) `[i.+1, i.+1] = [:: f (Wo i)]. Proof. by rewrite codom1_Skk swnxE2. Qed. - End SwapNext. (* Swaps/slices with decidable equality *) Section SwapNextEq. -Variable (n : nat) (A : eqType). +Variables (n : nat) (A : eqType). Lemma perm_swnx (f : {ffun 'I_n.+1 -> A}) (i : 'I_n) : perm_eq (fgraph f) (fgraph (pffun (swnx i) f)). @@ -267,22 +273,26 @@ Qed. Lemma perm_swnx_ux (f : {ffun 'I_n.+1 -> A}) (i : 'I_n) k : i <= k -> - perm_eq (&:(fgraph f) `]-oo, k.+1]) (&:(fgraph (pffun (swnx i) f)) `]-oo, k.+1]). + perm_eq &:(fgraph f) `]-oo, k.+1] + &:(fgraph (pffun (swnx i) f)) `]-oo, k.+1]. Proof. move=>H; move: (perm_swnx f i). set f' : {ffun 'I_n.+1 -> A} := pffun (swnx i) f. rewrite {1}(slice_uxou (fgraph f) k.+1) {1}(slice_uxou (fgraph f') k.+1). by rewrite swnx_oa /=; [rewrite perm_cat2r | rewrite ltnS]. Qed. - End SwapNextEq. Section Bubble. -Variable (n : nat) (A : ordType). +Variables (n : nat) (A : ordType). -Opaque Array.write Array.read. +(*****************) +(*****************) +(* Verifications *) +(*****************) +(*****************) -(* We verify 3 versions of the algorithm. *) +(* Verified 3 versions of the algorithm. *) (* *) (* The flag/swapped index is threaded through as a functional argument *) (* instead of storing and reading it from the heap. This would be *) @@ -319,8 +329,8 @@ Opaque Array.write Array.read. (*****************************************************) Program Definition cas_next (a : {array 'I_n.+2 -> A}) (i : 'I_n.+1) : - {f : {ffun 'I_n.+2 -> A}}, - STsep (Array.shape a f, + STsep {f : {ffun 'I_n.+2 -> A}} + (Array.shape a f, [vfun (b : bool) h => exists (p : 'S_n.+2), let i0 := Wo i in let i1 := So i in h \In Array.shape a (pffun p f) /\ @@ -336,25 +346,24 @@ Program Definition cas_next (a : {array 'I_n.+2 -> A}) (i : 'I_n.+1) : ret true else ret false). Next Obligation. -move=>a i /= [f][] h /= [E V]. -set i0 := Wo i; set i1 := So i. +move=>a i /= [f][] h /= E; set i0 := Wo i; set i1 := So i. do 2!apply: [stepE f, h]=>//= _ _ [->->]. case: oleqP=>H; first by step=>_; exists 1%g; split=>//; rewrite pffunE1. -apply: [stepE f ]=>//= _ _ [-> Vs ]; set fs := finfun [eta f with i0 |-> f i1]. -apply: [stepE fs]=>//= _ _ [-> Vsw]; set fsw := finfun [eta fs with i1 |-> f i0]. +apply: [stepE f]=>//= _ {E}h E. +set fs := finfun [eta f with i0 |-> f i1]. +apply: [stepE fs]=>//= _ {E}h E. +set fsw := (finfun _) in E. step=>_; exists (swnx i); do!split=>//=. -suff: fsw = pffun (swnx i) f by move=>->. +suff: fsw = pffun (swnx i) f by move=><-. rewrite /fsw /fs /swnx; apply/ffunP=>/= x; rewrite !ffunE /= ffunE /=. case: tpermP=>[->|->|/eqP/negbTE->/eqP/negbTE->] //; rewrite eqxx //. by rewrite /i1 eq_sym (negbTE (So_WoN i)). Qed. -Opaque cas_next. - Definition bubble_pass_loop (a : {array 'I_n.+2 -> A}) := forall isw : 'I_n.+1 * bool, - {f : {ffun 'I_n.+2 -> A}}, - STsep (fun h => h \In Array.shape a f /\ + STsep {f : {ffun 'I_n.+2 -> A}} + (fun h => h \In Array.shape a f /\ (* if no swaps have happened yet, the prefix is sorted *) ~~ isw.2 ==> sorted oleq (&:(fgraph f) `]-oo, isw.1 : nat]), [vfun (b : bool) h => @@ -364,13 +373,13 @@ Definition bubble_pass_loop (a : {array 'I_n.+2 -> A}) := ~~ b ==> (p == 1%g) && sorted oleq (fgraph f)]]). Program Definition bubble_pass (a : {array 'I_n.+2 -> A}) : - {f : {ffun 'I_n.+2 -> A}}, - STsep (Array.shape a f, + STsep {f : {ffun 'I_n.+2 -> A}} + (Array.shape a f, [vfun (b : bool) h => exists (p : 'S_n.+2), h \In Array.shape a (pffun p f) /\ ~~ b ==> (p == 1%g) && sorted oleq (fgraph f)]) := - Do (let go := Fix (fun (loop : bubble_pass_loop a) '(i, sw) => + Do (let go := ffix (fun (loop : bubble_pass_loop a) '(i, sw) => Do (sw0 <-- cas_next a i; let sw1 := sw || sw0 in if decP (b := i < n) idP is left pf @@ -381,10 +390,10 @@ Next Obligation. move=>a loop _ i sw [f][] h /= [H Hs]. set i0 := Wo i; set i1 := So i. apply: [stepE f]=>//= sw0 m [p][{}H Hsw]; case: decP=>Hin. -- (* i < n, recursive call *) - apply: [gE (pffun p f)]=>//=. - - (* precondition *) - split=>//; rewrite negb_or andbC Sbo_eq. +(* i < n, recursive call *) +- apply: [gE (pffun p f)]=>//=. + (* precondition *) + - split=>//; rewrite negb_or andbC Sbo_eq. case: sw0 Hsw=>//=; case=>{p H}-> Hf; rewrite pffunE1. (* swap didn't happen in this iteration *) apply/(implyb_trans Hs)/implyP=>{Hs}. @@ -402,8 +411,10 @@ step=>Vm; exists p; split=>//; first by apply/implyP=>->. (* swap didn't happen in last iteration, check initial flag *) case: sw0 Hsw=>/=; first by rewrite orbT. (* flag wasn't set, so the pass didn't swap anything - the array is sorted *) -case=>-> Hf; rewrite orbF eqxx /=; apply/(implyb_trans Hs)/implyP=>{}Hs. -rewrite (slice_usize (codom f)) size_codom /= card_ord slice_oSR -[X in `]-oo, X.+1]]Hin. +case=>-> Hf; rewrite orbF eqxx /=. +apply/(implyb_trans Hs)/implyP=>{}Hs. +rewrite (slice_usize (codom f)) size_codom /= card_ord +slice_oSR -[X in `]-oo, X.+1]]Hin. by apply: codom1_sortedS Hs. Qed. Next Obligation. @@ -413,25 +424,24 @@ split=>//; apply: sorted1. by rewrite slice_size /= add0n subn0 size_codom /= card_ord. Qed. -Opaque bubble_pass. - Definition bubble_sort_loop (a : {array 'I_n.+2 -> A}) : Type := unit -> - {f : {ffun 'I_n.+2 -> A}}, - STsep (Array.shape a f, + STsep {f : {ffun 'I_n.+2 -> A}} + (Array.shape a f, [vfun (_ : unit) h => exists (p : 'S_n.+2), h \In Array.shape a (pffun p f) /\ sorted oleq (fgraph (pffun p f))]). -Program Definition bubble_sort (a : {array 'I_n.+2 -> A}) : bubble_sort_loop a := - Fix (fun (go : bubble_sort_loop a) _ => +Program Definition bubble_sort (a : {array 'I_n.+2 -> A}) : + bubble_sort_loop a := + ffix (fun (go : bubble_sort_loop a) _ => Do (sw <-- bubble_pass a; if sw then go tt : ST _ else skip)). Next Obligation. -move=>a go _ [f][] h /= [E V]. +move=>a go _ [f][] h /= E. apply: [stepE f]=>//; case=>m /=. - (* the pass did swap something, loop *) case=>p [Hm _]. @@ -447,8 +457,6 @@ End Bubble. Section BubbleOpt. Variable (n : nat) (A : ordType). -Opaque cas_next. - (*******************) (* optimization #1 *) (*******************) @@ -473,31 +481,31 @@ Opaque cas_next. Definition bubble_pass_opt_loop (a : {array 'I_n.+2 -> A}) (k : 'I_n.+1) := forall isw : 'I_n.+1 * bool, - {f : {ffun 'I_n.+2 -> A}}, - STsep (fun h => [/\ h \In Array.shape a f, - (* all elements before f(k+2) are smaller than it *) - allrel oleq (&:(fgraph f) `]-oo, k.+1]) - (&:(fgraph f) `]k.+1, +oo[), - (* and the suffix is sorted *) - sorted oleq (&:(fgraph f) `]k.+1, +oo[), - (* we don't touch it *) - isw.1 <= k, - (* all elements before f(i) are smaller than it *) - all (oleq^~ (f (Wo isw.1))) (&:(fgraph f) `]-oo, isw.1 : nat[) & - ~~ isw.2 ==> sorted oleq (&:(fgraph f) `]-oo, isw.1 : nat[)], - [vfun (b : bool) h => - exists (p : 'S_n.+2), let f' := pffun p f in - [/\ h \In Array.shape a f', - isw.2 ==> b & - if b then - allrel oleq (&:(fgraph f') `]-oo, k : nat]) - (&:(fgraph f') `]k : nat, +oo[) && - sorted oleq (&:(fgraph f') `]k : nat, +oo[) - else (p == 1%g) && sorted oleq (fgraph f)]]). + STsep {f : {ffun 'I_n.+2 -> A}} + (fun h => [/\ h \In Array.shape a f, + (* all elements before f(k+2) are smaller than it *) + allrel oleq (&:(fgraph f) `]-oo, k.+1]) + (&:(fgraph f) `]k.+1, +oo[), + (* and the suffix is sorted *) + sorted oleq (&:(fgraph f) `]k.+1, +oo[), + (* we don't touch it *) + isw.1 <= k, + (* all elements before f(i) are smaller than it *) + all (oleq^~ (f (Wo isw.1))) (&:(fgraph f) `]-oo, isw.1 : nat[) & + ~~ isw.2 ==> sorted oleq (&:(fgraph f) `]-oo, isw.1 : nat[)], + [vfun (b : bool) h => + exists (p : 'S_n.+2), let f' := pffun p f in + [/\ h \In Array.shape a f', + isw.2 ==> b & + if b then + allrel oleq (&:(fgraph f') `]-oo, k : nat]) + (&:(fgraph f') `]k : nat, +oo[) && + sorted oleq (&:(fgraph f') `]k : nat, +oo[) + else (p == 1%g) && sorted oleq (fgraph f)]]). Program Definition bubble_pass_opt (a : {array 'I_n.+2 -> A}) (k : 'I_n.+1) : - {f : {ffun 'I_n.+2 -> A}}, - STsep (fun h => [/\ h \In Array.shape a f, + STsep {f : {ffun 'I_n.+2 -> A}} + (fun h => [/\ h \In Array.shape a f, allrel oleq (&:(fgraph f) `]-oo, k.+1]) (&:(fgraph f) `]k.+1, +oo[) & sorted oleq (&:(fgraph f) `]k.+1, +oo[)], @@ -509,29 +517,28 @@ Program Definition bubble_pass_opt (a : {array 'I_n.+2 -> A}) (k : 'I_n.+1) : (&:(fgraph f') `]k : nat, +oo[) && sorted oleq (&:(fgraph f') `]k : nat, +oo[) else (p == 1%g) && sorted oleq (fgraph f)]) := - Do (let go := Fix (fun (loop : bubble_pass_opt_loop a k) '(i, sw) => + Do (let go := ffix (fun (loop : bubble_pass_opt_loop a k) '(i, sw) => Do (sw0 <-- cas_next a i; let sw1 := sw || sw0 in if decP (b := i < k) idP is left pf - then loop (Sbo (i:=i) pf, sw1) - else ret sw1)) + then loop (Sbo (i:=i) pf, sw1) + else ret sw1)) in go (ord0, false)). Next Obligation. (* Sbo is always safe *) -move=>_ k _ _ i _ _ /= E; rewrite E ltnS; symmetry. -by apply: (leq_trans E); rewrite -ltnS; apply: ltn_ord. +by move=>_ k _ _ i _ _ /= E; rewrite E ltnS (leq_trans E) // -ltnS ltn_ord. Qed. Next Obligation. move=>a k loop _ i sw [f][] h /= [H Hak Hsk Hik Hai Hsi]. set i0 := Wo i; set i1 := So i. apply: [stepE f]=>//= sw0 m [p][Hm Hsw]; case: decP=>{}H. -- (* i < k, recursive call *) - apply: [gE (pffun p f)]=>//=. - - (* precondition *) - rewrite Sbo_eq Wo_Sbo_eq negb_or slice_oSR. +(* i < k, recursive call *) +- apply: [gE (pffun p f)]=>//=. + (* precondition *) + - rewrite Sbo_eq Wo_Sbo_eq negb_or slice_oSR. case: sw0 Hsw=>/=; case=>Ep Hf; rewrite {p}Ep ?pffunE1 in Hm *. - - (* swap happened before the call *) - rewrite andbF /= swnx_oa //; split=>//. + (* swap happened before the call *) + - rewrite andbF /= swnx_oa //; split=>//. - rewrite (@allrel_in_l _ _ _ _ &:(codom f) `]-oo, k.+1] ) //. by apply/perm_mem; rewrite perm_sym; apply: perm_swnx_ux. by rewrite swnx_ux_rcons all_rcons swnxE2 Hai andbT; apply: ordW. @@ -552,8 +559,8 @@ have {H}Hik : nat_of_ord i = k. step=>Vm; exists p; split=>//; first by apply/implyP=>->. rewrite -Hik in Hak Hsk *; rewrite -slice_oSL. case: sw0 Hsw=>/=; case=>-> Hf. -- (* swap happened on last iteration *) - rewrite orbT swnx_xu_cons /=; apply/andP; split. +(* swap happened on last iteration *) +- rewrite orbT swnx_xu_cons /=; apply/andP; split. - rewrite swnx_ux_rcons allrel_rconsl allrel_consr /=. move: Hak; rewrite codom1_ax_rcons2 // !allrel_rconsl -/i0 -/i1 -andbA. by case/and3P=>-> _ ->; rewrite (ordW Hf) /= !andbT. @@ -562,8 +569,8 @@ case: sw0 Hsw=>/=; case=>-> Hf. by rewrite codom1_ax_rcons2 // 4!(mem_rcons, inE) eqxx /= orbT. (* swap didn't happen on last iteration *) rewrite orbF pffunE1 eqxx /=; case: sw Hsi=>/=[_|]. -- (* flag was set *) - apply/andP; split. +(* flag was set *) +- apply/andP; split. - rewrite codom1_ax_rcons // codom1_xu_cons allrel_rconsl allrel_consr /=. move: Hak; rewrite codom1_ax_rcons2 // !allrel_rconsl -andbA -/i0 -/i1. case/and3P=>->-> _; rewrite Hf /= !andbT. @@ -571,10 +578,13 @@ rewrite orbF pffunE1 eqxx /=; case: sw Hsi=>/=[_|]. rewrite codom1_xu_cons /= (path_sortedE (@otrans A)) Hsk andbT. by move: Hak; rewrite codom1_ax_rcons2 // !allrel_rconsl -andbA; case/and3P. (* flag wasn't set, so the pass didn't swap anything - the array is sorted *) -move=>Hs2; rewrite (codom1_split f i) -/i0 -/i1 /=; rewrite Hik in Hai Hs2 Hsk *. -rewrite (sorted_pairwise (@otrans A)) pairwise_cat /= -!(sorted_pairwise (@otrans A)). +move=>Hs2; rewrite (codom1_split f i) -/i0 -/i1 /=. +rewrite Hik in Hai Hs2 Hsk *. +rewrite (sorted_pairwise (@otrans A)) pairwise_cat /=. +rewrite -!(sorted_pairwise (@otrans A)). rewrite Hf Hsk Hs2 /= andbT !allrel_consr -!andbA Hai /=. -move: Hak; rewrite -Hik codom1_ax_rcons2 // !allrel_rconsl -andbA; case/and3P=>->->->/=. +move: Hak; rewrite -Hik codom1_ax_rcons2 // !allrel_rconsl -andbA. +case/and3P=>->->->/=. by rewrite andbT Hik; apply/sub_all/Hai=>z /otrans; apply. Qed. Next Obligation. @@ -583,12 +593,10 @@ apply: [gE f]=>//= [|v m [p][Hm _ Hv] _]; last by exists p. by rewrite itv_under0R. Qed. -Opaque bubble_pass_opt. - Definition bubble_sort_opt_loop (a : {array 'I_n.+2 -> A}) : Type := forall (k : 'I_n.+1), - {f : {ffun 'I_n.+2 -> A}}, - STsep (fun h => [/\ h \In Array.shape a f, + STsep {f : {ffun 'I_n.+2 -> A}} + (fun h => [/\ h \In Array.shape a f, allrel oleq (&:(fgraph f) `]-oo, k.+1]) (&:(fgraph f) `]k.+1, +oo[) & sorted oleq (&:(fgraph f) `]k.+1, +oo[)], @@ -598,13 +606,13 @@ Definition bubble_sort_opt_loop (a : {array 'I_n.+2 -> A}) : Type := sorted oleq (fgraph f')]). Program Definition bubble_sort_opt (a : {array 'I_n.+2 -> A}) : - {f : {ffun 'I_n.+2 -> A}}, - STsep (Array.shape a f, + STsep {f : {ffun 'I_n.+2 -> A}} + (Array.shape a f, [vfun (_ : unit) h => exists (p : 'S_n.+2), h \In Array.shape a (pffun p f) /\ sorted oleq (fgraph (pffun p f))]) := - Do (let go := Fix (fun (loop : bubble_sort_opt_loop a) k => + Do (let go := ffix (fun (loop : bubble_sort_opt_loop a) k => Do (sw <-- bubble_pass_opt a k; if sw then loop (Po k) : ST _ @@ -612,9 +620,10 @@ Program Definition bubble_sort_opt (a : {array 'I_n.+2 -> A}) : in go ord_max). Next Obligation. move=>a go k [f][] h /= [H Ha Hs]. -apply: [stepE f]=>{H Ha Hs}//=; case=>m; case=>p [Hm /andP [H Hs]]; last first. -- (* id permutation *) - rewrite {p H}(eqP H) in Hm. +apply: [stepE f]=>{H Ha Hs}//=; case=>m; case=>p [Hm /andP [H Hs]]; +last first. +(* id permutation *) +- rewrite {p H}(eqP H) in Hm. by step=>_; exists 1%g; split=>//; rewrite pffunE1. apply: [gE (pffun p f)]=>//=; last first. - by move=> _ m2 [p'][H2 Hs'] _; exists (p' * p)%g; rewrite pffunEM. @@ -629,7 +638,7 @@ rewrite codom_ux_rcons /= allrel_rconsl Ha1 /=. by have -> //: Ordinal P = @Ordinal n.+2 1 (ltn0Sn n); apply/ord_inj. Qed. Next Obligation. -move=>a [f][] h /= [E V]. +move=>a [f][] h /= E. have Hs: size (codom f) <= n.+2 by rewrite size_codom /= card_ord. apply: [gE f]=>//=; split=>//; rewrite (@itv_overL _ _ (BRight n.+1)) //=; try by rewrite addn1. @@ -641,8 +650,6 @@ End BubbleOpt. Section BubbleOpt2. Variable (n : nat) (A : ordType). -Opaque cas_next. - (*******************) (* optimization #2 *) (*******************) @@ -668,8 +675,8 @@ Opaque cas_next. Definition bubble_pass_opt2_loop (a : {array 'I_n.+2 -> A}) (k : 'I_n.+1) := forall ils : 'I_n.+1 * 'I_n.+2, - {f : {ffun 'I_n.+2 -> A}}, - STsep (fun h => [/\ h \In Array.shape a f, + STsep {f : {ffun 'I_n.+2 -> A}} + (fun h => [/\ h \In Array.shape a f, allrel oleq (&:(fgraph f) `]-oo, k.+1]) (&:(fgraph f) `]k.+1, +oo[), sorted oleq (&:(fgraph f) `]k.+1, +oo[), @@ -686,8 +693,8 @@ Definition bubble_pass_opt2_loop (a : {array 'I_n.+2 -> A}) (k : 'I_n.+1) := sorted oleq (&:(fgraph f') `[j : nat, +oo[)]]). Program Definition bubble_pass_opt2 (a : {array 'I_n.+2 -> A}) (k : 'I_n.+1) : - {f : {ffun 'I_n.+2 -> A}}, - STsep (fun h => [ /\ h \In Array.shape a f, + STsep {f : {ffun 'I_n.+2 -> A}} + (fun h => [ /\ h \In Array.shape a f, allrel oleq (&:(fgraph f) `]-oo, k.+1]) (&:(fgraph f) `]k.+1, +oo[) & sorted oleq (&:(fgraph f) `]k.+1, +oo[)], @@ -697,7 +704,7 @@ Program Definition bubble_pass_opt2 (a : {array 'I_n.+2 -> A}) (k : 'I_n.+1) : allrel oleq (&:(fgraph f') `]-oo, j : nat[) (&:(fgraph f') `[j : nat, +oo[) & sorted oleq (&:(fgraph f') `[j : nat, +oo[)]]) := - Do (let go := Fix (fun (loop : bubble_pass_opt2_loop a k) '(i, ls) => + Do (let go := ffix (fun (loop : bubble_pass_opt2_loop a k) '(i, ls) => Do (sw <-- cas_next a i; let ls1 := if sw then So i else ls in if decP (b := i < k) idP is left pf @@ -706,25 +713,25 @@ Program Definition bubble_pass_opt2 (a : {array 'I_n.+2 -> A}) (k : 'I_n.+1) : in go (ord0, ord0)). Next Obligation. (* show that Sbo is always safe *) -move=>_ k _ _ i _ _ /= E; rewrite E ltnS; symmetry. -by apply: (leq_trans E); rewrite -ltnS; apply: ltn_ord. +by move=>_ k _ _ i _ _ /= E; rewrite E ltnS (leq_trans E) // -ltnS ltn_ord. Qed. Next Obligation. move=>a k loop _ i ls [f][] h /= [Hs Hak Hsk /andP [Hils Hik] Hai Hsi]. set i0 := Wo i; set i1 := So i. apply: [stepE f]=>//= sw m [p][Hm Hsw]; case: decP=>H. -- (* i < k, recursive call *) - apply: [gE (pffun p f)]=>//=. - - (* precondition *) - rewrite Sbo_eq. +(* i < k, recursive call *) +- apply: [gE (pffun p f)]=>//=. + (* precondition *) + - rewrite Sbo_eq. case: sw Hsw=>/=; case=>Ep Hf; rewrite {p}Ep ?pffunE1 in Hm *. - - (* swap happened before the call *) - rewrite So_eq swnx_oa; last by rewrite ltnS; apply: ltnW. + (* swap happened before the call *) + - rewrite So_eq swnx_oa; last by rewrite ltnS; apply: ltnW. rewrite swnx_Skk /=; split=>//. - rewrite (@allrel_in_l _ _ _ _ &:(codom f) `]-oo, k.+1]) //. by apply/perm_mem; rewrite perm_sym; apply: perm_swnx_ux. - by rewrite ltnS leqnn. - rewrite allrel1r slice_oSR (slice_split (x:=ls) _ true); last by rewrite in_itv. + rewrite allrel1r slice_oSR (slice_split (x:=ls) _ true); + last by rewrite in_itv. rewrite /= all_cat; apply/andP; split. - rewrite swnx_ao //; move: Hai. by rewrite codom1_ax_rcons // allrel_rconsr; case/andP. @@ -741,7 +748,8 @@ apply: [stepE f]=>//= sw m [p][Hm Hsw]; case: decP=>H. suff: all (oleq^~ (f (Wo i))) (&:(codom f) `]-oo, ls:nat[). - by apply/sub_all=>z /otrans; apply. by move: Hai; rewrite codom1_ax_rcons // allrel_rconsr; case/andP. - rewrite codom1_ax_rcons2 // (sorted_rconsE _ _ (@otrans A)) -{2}codom1_ax_rcons // Hsi andbT. + rewrite codom1_ax_rcons2 //. + rewrite (sorted_rconsE _ _ (@otrans A)) -{2}codom1_ax_rcons // Hsi andbT. rewrite all_rcons Hf /=. move: Hsi; rewrite codom1_ax_rcons // (sorted_rconsE _ _ (@otrans A)). by case/andP=>+ _; apply/sub_all=>z Hz; apply/otrans/Hf. @@ -756,13 +764,15 @@ have {H}Hik : nat_of_ord i = k. by move/negP: H; rewrite -leqNgt. rewrite -{loop k}Hik in Hak Hsk *. step=>Vm; exists p; case: sw Hsw=>/=; case=>Ep Hf. -- (* swap happened on last iteration *) - rewrite So_eq; split=>//=; first by apply: ltnW. - - rewrite Ep slice_oSR swnx_ux_rcons swnx_xu_cons allrel_rconsl allrel_consr /=. +(* swap happened on last iteration *) +- rewrite So_eq; split=>//=; first by apply: ltnW. + - rewrite Ep slice_oSR swnx_ux_rcons swnx_xu_cons + allrel_rconsl allrel_consr /=. move: Hak; rewrite codom1_ax_rcons2 // !allrel_rconsl -/i0 -/i1 -andbA. case/and3P=>-> _ ->; rewrite (ordW Hf) /= !andbT. - move: Hai; rewrite codom1_ax_rcons //= allrel_rconsr; case/andP=>_ Hals. - move: Hsi; rewrite codom1_ax_rcons // sorted_rconsE //; case/andP =>Halsi _. + move: Hai; rewrite codom1_ax_rcons //= allrel_rconsr=>/andP [_ Hals]. + move: Hsi; rewrite codom1_ax_rcons // sorted_rconsE //=. + case/andP => Halsi _. move: Hils; rewrite leq_eqVlt; case/orP=>[/eqP <-|Hlsi] //. rewrite (slice_split (x:=ls) _ true); last by rewrite in_itv. by rewrite /= all_cat Hals. @@ -770,18 +780,22 @@ step=>Vm; exists p; case: sw Hsw=>/=; case=>Ep Hf. by move: Hak; rewrite codom1_ax_rcons2 // !allrel_rconsl -!andbA; case/and3P. (* swap didn't happen on last iteration *) rewrite {p}Ep pffunE1 /= in Hm *. -rewrite (slice_split (x:=i) (i:=`[ls:nat, +oo[) _ false); last by rewrite in_itv /= andbT. +rewrite (slice_split (x:=i) (i:=`[ls:nat, +oo[) _ false); last first. +- by rewrite in_itv /= andbT. split=>//=. -- rewrite allrel_catr Hai /= -slice_oSL codom1_xu_cons allrel_consr; apply/andP; split. +- rewrite allrel_catr Hai /= -slice_oSL codom1_xu_cons allrel_consr. + apply/andP; split. - suff: all (oleq^~ (f (Wo i))) (&:(codom f) `]-oo, ls:nat[). - by apply/sub_all=>z /otrans; apply. apply/sub_all/Hai=>z /allP; apply. by rewrite codom1_ax_rcons // mem_rcons inE eqxx. apply/allrel_sub_l/Hak/slice_subset. by rewrite subitvE /= bnd_simp leEnat; apply/ltnW. -rewrite (sorted_pairwise (@otrans A)) pairwise_cat -!(sorted_pairwise (@otrans A)) Hsi /=. +rewrite (sorted_pairwise (@otrans A)) pairwise_cat. +rewrite -!(sorted_pairwise (@otrans A)) Hsi /=. rewrite -slice_oSL codom1_xu_cons /= allrel_consr -andbA; apply/and3P; split. -- move: Hsi; rewrite (sorted_pairwise (@otrans A)) codom1_ax_rcons // pairwise_rcons. +- move: Hsi. + rewrite (sorted_pairwise (@otrans A)) codom1_ax_rcons // pairwise_rcons. case/andP=>H' _; rewrite all_rcons Hf /=. by apply/sub_all/H'=>z Hz; apply/otrans/Hf. - apply/allrel_sub_l/Hak/slice_subset. @@ -796,12 +810,10 @@ split=>//; first by rewrite itv_underR. by apply: sorted1; rewrite slice_size size_codom card_ord. Qed. -Opaque bubble_pass_opt2. - Definition bubble_sort_opt2_loop (a : {array 'I_n.+2 -> A}) : Type := forall (k : 'I_n.+1), - {f : {ffun 'I_n.+2 -> A}}, - STsep (fun h => [/\ h \In Array.shape a f, + STsep {f : {ffun 'I_n.+2 -> A}} + (fun h => [/\ h \In Array.shape a f, allrel oleq (&:(fgraph f) `]-oo, k.+1]) (&:(fgraph f) `]k.+1, +oo[) & sorted oleq (&:(fgraph f) `]k.+1, +oo[)], @@ -811,13 +823,13 @@ Definition bubble_sort_opt2_loop (a : {array 'I_n.+2 -> A}) : Type := sorted oleq (fgraph f')]). Program Definition bubble_sort_opt2 (a : {array 'I_n.+2 -> A}) : - {f : {ffun 'I_n.+2 -> A}}, - STsep (Array.shape a f, + STsep {f : {ffun 'I_n.+2 -> A}} + (Array.shape a f, [vfun (_ : unit) h => exists (p : 'S_n.+2), let f' := pffun p f in h \In Array.shape a f' /\ sorted oleq (fgraph f')]) := - Do (let go := Fix (fun (loop : bubble_sort_opt2_loop a) k => + Do (let go := ffix (fun (loop : bubble_sort_opt2_loop a) k => Do (k1 <-- bubble_pass_opt2 a k; if decP (b := 1 < k1) idP is left pf then loop (M2lo (i:=k1) pf) : ST _ @@ -832,8 +844,10 @@ case: decP=>Hx; last first. - by move=>Hx _; rewrite Hx -slice_0L slice_uu in Hs. case: (nat_of_ord x) Ha Hs=>//; case=>//= Ha1 Hs1 _ _. rewrite (slice_uoxu (codom (pffun p f)) 1). - rewrite (sorted_pairwise (@otrans A)) pairwise_cat -!(sorted_pairwise (@otrans A)). - by apply/and3P; split=>//; apply: sorted1; rewrite slice_size size_codom card_ord. + rewrite (sorted_pairwise (@otrans A)) pairwise_cat. + rewrite -!(sorted_pairwise (@otrans A)). + apply/and3P; split=>//; apply: sorted1. + by rewrite slice_size size_codom card_ord. apply: [gE (pffun p f)]=>//=; last first. - move=> _ m2 [p'][H2 Hs'] _; exists (p' * p)%g. by rewrite pffunEM. @@ -841,7 +855,7 @@ rewrite -slice_oSL -slice_oSR (M2lo_eq Hx); suff: x.-2.+2 = x by move=>->. by rewrite -subn2 -addn2 addnBAC // addnK. Qed. Next Obligation. -move=>a [f][] h /= [E V]. +move=>a [f][] h /= E. apply: [gE f]=>//=. rewrite (@itv_overL _ _ _ +oo) /=; first by split=>//; exact: allrel0r. by rewrite leEnat addn1 size_codom /= card_ord. diff --git a/examples/congmath.v b/examples/congmath.v old mode 100755 new mode 100644 index fde0566..fa4c2d3 --- a/examples/congmath.v +++ b/examples/congmath.v @@ -1,25 +1,47 @@ +(* +Copyright 2009 IMDEA Software Institute +Licensed under the Apache License, Version 2.0 (the "License"); +you may not use this file except in compliance with the License. +You may obtain a copy of the License at + http://www.apache.org/licenses/LICENSE-2.0 +Unless required by applicable law or agreed to in writing, software +distributed under the License is distributed on an "AS IS" BASIS, +WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. +See the License for the specific language governing permissions and +limitations under the License. +*) + +From HB Require Import structures. +From Coq Require Import Recdef Setoid ssreflect ssrbool ssrfun. +From mathcomp Require Import eqtype choice ssrnat seq bigop fintype finfun. +From pcm Require Import options prelude ordtype finmap pred seqext. + (**********************) (* Congruence closure *) (**********************) -From Coq Require Import Recdef ssreflect ssrbool ssrfun. -From mathcomp Require Import eqtype choice ssrnat seq bigop fintype finfun. -From pcm Require Import options prelude ordtype finmap pred seqext. -From htt Require Import interlude. +(* Background mathematics for the verification of the *) +(* Barcelogic Congruence Closure algorithm. *) +(* Described in the POPL10 paper: *) +(* Structure the verification of heap-manipulating programs *) +(* by Nanevski, Vafeiadis and Berdine *) -Ltac add_morphism_tactic := SetoidTactics.add_morphism_tactic. +(* This file contains the proofs of the mathematical properties *) +(* needed for the verification. The implementation and the *) +(* proofs of the actual stateful programs is in the *) +(* accompanying file congprog.v *) (**************************) (* Constants with arities *) (**************************) -Inductive constant : Type := const_with_arity of nat & nat. +Inductive constant : Set := const_with_arity of nat & nat. (* constants are an equality type *) Definition eqcnst (c1 c2 : constant) : bool := match c1, c2 with - | const_with_arity s1 n1, const_with_arity s2 n2 => - (s1 == s2) && (n1 == n2) + | const_with_arity s1 n1, const_with_arity s2 n2 => + (s1 == s2) && (n1 == n2) end. Lemma eqconstantP : Equality.axiom eqcnst. @@ -29,8 +51,7 @@ do ![case: eqP=>/=[->|H]; last by apply: ReflectF; case]. by apply: ReflectT. Qed. -Canonical Structure constant_eqMixin := EqMixin eqconstantP. -Canonical Structure constant_eqType := EqType _ constant_eqMixin. +HB.instance Definition _ := hasDecEq.Build constant eqconstantP. (* constants are a countable type *) @@ -46,38 +67,73 @@ Proof. by case. Qed. Lemma consttagK : cancel consttag tagconst. Proof. by case. Qed. -Definition const_countMixin := CanCountMixin tagconstK. -Definition const_choiceMixin := CountChoiceMixin const_countMixin. -Canonical Structure const_choiceType := Eval hnf in ChoiceType _ const_choiceMixin. -Canonical Structure const_countType := Eval hnf in CountType _ const_countMixin. +HB.instance Definition _ := Countable.copy constant (can_type tagconstK). (***********************************************************) (* The symbols are a predetermined finite set of constants *) (***********************************************************) -Variable symbs : seq constant. -Definition symb := @seq_sub const_choiceType symbs. -Canonical Structure symb_eqType := [eqType of symb]. -Canonical Structure symb_finType := [finType of symb for @seq_sub_finType _ _]. -(* symb is an ordered type *) -Definition symb_ordMixin := [fin_ordMixin of symb]. -Canonical Structure symb_ordType := OrdType _ symb_ordMixin. +(* Given sequences s, symb s is the type of elements of s. *) +(* One could use seq_sub from fintype.v here, and just set *) +(* Definition symb (s : seq constant) := seq_sub s. *) +(* However, this is expensive in terms of universes, *) +(* as it forces symb to be Type, not Set. *) +(* Thus, it's best here to roll out a bespoke fintype structure *) +(* copying relevant code from seq_sub from fintype.v *) +(* but setting universe to Set from the beginning. *) + +Record symb (s : seq constant) : Set := + SeqSymb {symb_val : constant; + symb_valP : in_mem symb_val (@mem constant _ s)}. + +#[warning="-projection-no-head-constant"] +HB.instance Definition _ s := [isSub for (@symb_val s)]. +HB.instance Definition _ s := [Equality of symb s by <:]. +Definition symb_enum s : seq (symb s) := undup (pmap insub s). +Lemma mem_symb_enum s x : x \in symb_enum s. +Proof. by rewrite mem_undup mem_pmap -valK map_f ?symb_valP. Qed. +Lemma val_symb_enum s : uniq s -> map val (symb_enum s) = s. +Proof. +move=> Us; rewrite /symb_enum undup_id ?pmap_sub_uniq //. +rewrite (pmap_filter (insubK _)); apply/all_filterP. +by apply/allP => x; rewrite isSome_insub. +Qed. +Definition symb_pickle s x := index x (symb_enum s). +Definition symb_unpickle s n := nth None (map some (symb_enum s)) n. +Lemma symb_pickleK s : pcancel (@symb_pickle s) (@symb_unpickle s). +Proof. +rewrite /symb_unpickle => x. +by rewrite (nth_map x) ?nth_index ?index_mem ?mem_symb_enum. +Qed. +Definition symb_isCountable s := isCountable.Build (symb s) (@symb_pickleK s). +Fact symb_axiom s : Finite.axiom (symb_enum s). +Proof. exact: Finite.uniq_enumP (undup_uniq _) (@mem_symb_enum s). Qed. +Definition symb_isFinite s := isFinite.Build (symb s) (@symb_axiom s). +HB.instance Definition _ s := [Choice of symb s by <:]. +HB.instance Definition _ s : isCountable (symb s) := symb_isCountable s. +HB.instance Definition _ s : isFinite (symb s) := symb_isFinite s. +HB.instance Definition symb_ord_mix s := + isOrdered.Build (symb s) (fintype_is_ordtype (symb s)). +(* manual canonical declaration, as HB fails to declare it *) +Canonical symb_ordType s : ordType := + Ordered.Pack (sort:=symb s) (Ordered.Class (symb_ord_mix s)). (****************************) (* Expressions over symbols *) (****************************) -Inductive exp : Type := const of symb | app of exp & exp. +(* Consequently, expressions must be Type, not Set *) +Inductive exp s : Set := const of symb s | app of exp s & exp s. (* expressions are an equality type *) -Fixpoint eqexp (e1 e2 : exp) {struct e1} : bool := +Fixpoint eqexp s (e1 e2 : exp s) {struct e1} : bool := match e1, e2 with | const c1, const c2 => c1 == c2 | app e1 e2, app e3 e4 => (eqexp e1 e3) && (eqexp e2 e4) | _, _ => false end. -Lemma eqexpP : Equality.axiom eqexp. +Lemma eqexpP s : Equality.axiom (@eqexp s). Proof. elim=>[c1|e1 IH1 e2 IH2][c2|e3 e4] /=; try by apply: ReflectF. - case: eqP=>[->|H]; last by apply: ReflectF; case. @@ -87,24 +143,34 @@ case: IH2=>[->|H]; last by apply: ReflectF; case. by apply: ReflectT. Qed. -Canonical Structure exp_eqMixin := EqMixin eqexpP. -Canonical Structure exp_eqType := EqType _ exp_eqMixin. +HB.instance Definition _ s := hasDecEq.Build (@exp s) (@eqexpP s). + +Notation rel_exp s := (Pred (exp s * exp s)). (****************************************) (* Congruence relations over expression *) (****************************************) (* R is monotone if it is preserved by application *) -Notation rel_exp := (Pred (exp*exp)). - -Definition Reflexive (r : rel_exp) := forall x, (x, x) \In r. -Definition Symmetric (r : rel_exp) := forall x y, (x, y) \In r -> (y, x) \In r. -Definition Transitive (r : rel_exp) := forall y x z, (x, y) \In r -> (y, z) \In r -> (x, z) \In r. -Definition Equivalence (r : rel_exp) := Reflexive r /\ Symmetric r /\ Transitive r. -Definition Antisymmetric (r : rel_exp) := forall x y, (x, y) \In r -> (y, x) \In r -> x = y. - +Section Congruence. +Variable s : seq constant. +Notation symb := (symb s). +Notation exp := (exp s). +Notation rel_exp := (rel_exp s). + +Definition Reflexive (r : rel_exp) := + forall x, (x, x) \In r. +Definition Symmetric (r : rel_exp) := + forall x y, (x, y) \In r -> (y, x) \In r. +Definition Transitive (r : rel_exp) := + forall y x z, (x, y) \In r -> (y, z) \In r -> (x, z) \In r. +Definition Equivalence (r : rel_exp) := + Reflexive r /\ Symmetric r /\ Transitive r. +Definition Antisymmetric (r : rel_exp) := + forall x y, (x, y) \In r -> (y, x) \In r -> x = y. Definition monotone (R : rel_exp) : Prop := - forall f1 f2 e1 e2, (f1, f2) \In R -> (e1, e2) \In R -> (app f1 e1, app f2 e2) \In R. + forall f1 f2 e1 e2, (f1, f2) \In R -> + (e1, e2) \In R -> (app f1 e1, app f2 e2) \In R. Definition congruence (R : rel_exp) := Equivalence R /\ monotone R. @@ -121,7 +187,7 @@ Definition mono_cong (R : rel_exp) : congruence R -> monotone R. Proof. by move=>[[]]. Qed. Definition closure (R : rel_exp) : rel_exp := - [Pred e1 e2 | forall C, R ~> C -> congruence C -> (e1, e2) \In C]. + [Pred e | forall C, R ~> C -> congruence C -> e \In C]. Lemma cong_clos (R : rel_exp) : congruence (closure R). Proof. @@ -134,7 +200,6 @@ move=>f1 f2 e1 e2 H1 H2 C H3 H4. by apply: mono_cong (H1 C H3 H4) (H2 C H3 H4). Qed. -#[export] Hint Resolve cong_clos : core. Lemma reflC (R : rel_exp) : Reflexive (closure R). @@ -149,12 +214,13 @@ Proof. by apply: trans_cong. Qed. Lemma monoC (R : rel_exp) : monotone (closure R). Proof. by apply: mono_cong. Qed. -#[export] Hint Resolve reflC symC : core. (* lemmas about closure *) -Lemma closE (R1 R2 : rel_exp) : R1 <~> R2 -> closure R1 <~> closure R2. +Lemma closE (R1 R2 : rel_exp) : + R1 <~> R2 -> + closure R1 <~> closure R2. Proof. suff H: forall R1 R2, R1 <~> R2 -> closure R1 ~> closure R2. - by split; apply: H; [| symmetry]. @@ -162,11 +228,14 @@ move=>{}R1 {}R2 H [x y] H1 C H2 H3; apply: H1 H3. by rewrite H. Qed. +(* to be repeated outside of the section *) Add Morphism closure with - signature Morphisms.respectful (fun e1 e2 => e1 <~> e2) (fun e1 e2 => e1 <~> e2) as closure_morph. + signature Morphisms.respectful (fun e1 e2 => e1 <~> e2) + (fun e1 e2 => e1 <~> e2) as closure_morph. Proof. by apply: closE. Qed. -Lemma clos_clos (R Q : rel_exp) : closure (closure R +p Q) <~> closure (R +p Q). +Lemma clos_clos (R Q : rel_exp) : + closure (closure R \+p Q) <~> closure (R \+p Q). Proof. case=>e1 e2; split=>H1 C H2 H3; apply: H1 =>// [[x y]] H4. - by case: H4=>H4; [apply: H4=>// [[x1 y1]] H4|]; apply: H2; [left | right]. @@ -174,26 +243,30 @@ case: H4=>H4; apply: H2; [left | right]=>//. by move=>C0 H5 _; apply: H5. Qed. -Lemma clos_idemp (R : rel_exp) : closure (closure R) <~> closure R. +Lemma clos_idemp (R : rel_exp) : + closure (closure R) <~> closure R. Proof. by rewrite -(orr0 (closure R)) clos_clos !orr0. Qed. Lemma closKR (R1 R2 K : rel_exp) : - closure R1 <~> closure R2 -> closure (K +p R1) <~> closure (K +p R2). + closure R1 <~> closure R2 -> + closure (K \+p R1) <~> closure (K \+p R2). Proof. move=>H. by rewrite 2!(orrC K) -(clos_clos R1) -(clos_clos R2) H. Qed. Lemma closRK (R1 R2 K : rel_exp) : - closure R1 <~> closure R2 -> closure (R1 +p K) <~> closure (R2 +p K). + closure R1 <~> closure R2 -> + closure (R1 \+p K) <~> closure (R2 \+p K). Proof. by rewrite 2!(orrC _ K); apply: closKR. Qed. -Lemma closI (R : rel_exp) : R ~> closure R. +Lemma closI (R : rel_exp) : R <=p closure R. Proof. by move=>[x y] H1 T H2 H3; apply: H2. Qed. (* absorption and closure *) Lemma closKA (K R : rel_exp) : - closure R <~> closure (K +p R) <-> K ~> closure R. + closure R <~> closure (K \+p R) <-> + K <=p closure R. Proof. rewrite -orrAb; split=>[->|]. - by rewrite orrAb; move=>[x y] H; apply: closI; left. @@ -201,19 +274,22 @@ by rewrite (orrC _ R) -clos_clos orrC=><-; rewrite clos_idemp. Qed. Lemma closAK (K R : rel_exp) : - closure R <~> closure (R +p K) <-> K ~> closure R. + closure R <~> closure (R \+p K) <-> + K <=p closure R. Proof. by rewrite orrC closKA. Qed. Lemma clos_or (R1 R2 R : rel_exp) : - closure R1 ~> closure R -> closure R2 ~> closure R -> - closure (R1 +p R2) ~> closure R. + closure R1 <=p closure R -> + closure R2 <=p closure R -> + closure (R1 \+p R2) <=p closure R. Proof. rewrite -!closAK !(orrC R) !clos_clos -!(orrC R) => H1 H2. by rewrite H1 -clos_clos H2 clos_clos orrAC orrA. Qed. Lemma sub_closKR (R1 R2 K : rel_exp) : - (closure R1 ~> closure R2) -> closure (K +p R1) ~> closure (K +p R2). + (closure R1 <=p closure R2) -> + closure (K \+p R1) <=p closure (K \+p R2). Proof. rewrite -!closAK=>H. rewrite -(orrC (closure _)) clos_clos orrAC -orrA orrI. @@ -223,32 +299,59 @@ by rewrite orrC. Qed. Lemma sub_closRK (R1 R2 K : rel_exp) : - (closure R1 ~> closure R2) -> closure (R1 +p K) ~> closure (R2 +p K). + (closure R1 <=p closure R2) -> + closure (R1 \+p K) <=p closure (R2 \+p K). Proof. by move=>H; rewrite -!(orrC K); apply: sub_closKR. Qed. Lemma sub_closI (R1 R2 : rel_exp) : - R1 ~> R2 -> closure R1 ~> closure R2. + R1 <=p R2 -> + closure R1 <=p closure R2. Proof. rewrite -orrAb -closAK -(orrC (closure _)) clos_clos => ->. by rewrite (orrC R2) -orrA orrI. Qed. -(* relation that only relates the constant symbols with their image under a function *) +(* relation that only relates the constant symbols *) +(* with their image under a function *) Definition graph (f : symb -> exp) : rel_exp := - [Pred e1 e2 | (e1, e2) \in image (fun s => (const s, f s)) predT]. + [Pred e | e \in image (fun s => (const s, f s)) predT]. + +End Congruence. + +Hint Resolve reflC symC reflC symC : core. + +(* repeat the morphism declaration outside the section *) +Add Parametric Morphism s : (@closure s) with + signature Morphisms.respectful (fun e1 e2 => e1 <~> e2) + (fun e1 e2 => e1 <~> e2) as closure_morph'. +Proof. by apply: closE. Qed. + +(********************************************************************) +(* The definition of the data structures involved in the algorithm. *) +(********************************************************************) + +Section Structures. +Variable s : seq constant. +Notation symb := (symb s). +Notation exp := (exp s). +Notation rel_exp := (rel_exp s). (* Equations in canonical form, as required by the congruence closure *) (* algorithm. An equation is in canonical form if it is an equation *) (* between two constants, or between a constant and an application of two *) (* constants. *) -Inductive Eq : Type := simp_eq of symb & symb | comp_eq of symb & symb & symb. +Inductive Eq : Type := + simp_eq of symb & symb | + comp_eq of symb & symb & symb. (* equations are an equality type *) Definition eqEq (eq1 eq2 : Eq) : bool := match eq1, eq2 with - simp_eq c1 c2, simp_eq d1 d2 => (c1 == d1) && (c2 == d2) - | comp_eq c c1 c2, comp_eq d d1 d2 => (c == d) && (c1 == d1) && (c2 == d2) + simp_eq c1 c2, simp_eq d1 d2 => + (c1 == d1) && (c2 == d2) + | comp_eq c c1 c2, comp_eq d d1 d2 => + (c == d) && (c1 == d1) && (c2 == d2) | _, _ => false end. @@ -259,42 +362,43 @@ do ![case: eqP=>[->|H]; last by apply: ReflectF; case]; by apply: ReflectT. Qed. -Canonical Structure Eq_eqMixin := EqMixin eqEqP. -Canonical Structure Eq_eqType := EqType _ Eq_eqMixin. +HB.instance Definition _ := hasDecEq.Build Eq eqEqP. (* interpreting equations as relations *) Definition eq2rel (eq : Eq) : rel_exp := match eq with - | simp_eq c1 c2 => [Pred x y | x = const c1 /\ y = const c2] - | comp_eq c c1 c2 => [Pred x y | x = const c /\ y = app (const c1) (const c2)] + | simp_eq c1 c2 => Pred1 (const c1, const c2) + | comp_eq c c1 c2 => + Pred1 (const c, app (const c1) (const c2)) end. Fixpoint eqs2rel (eqs : seq Eq) : rel_exp := - if eqs is hd::tl then eq2rel hd +p eqs2rel tl else Pred0. + if eqs is hd::tl then eq2rel hd \+p eqs2rel tl else Pred0. (* Coercing triples of symbols into equations *) Definition symb2eq (t : symb*symb*symb) := let: (c, c1, c2) := t in comp_eq c c1 c2. (* inverting relations equations represented as triples *) -Definition invert (x y : exp) (s : seq (symb*symb*symb)) : - (x, y) \In eqs2rel (map symb2eq s) <-> +Definition invert (x y : exp) (ss : seq (symb*symb*symb)) : + (x, y) \In eqs2rel (map symb2eq ss) <-> exists c c1 c2, - [/\ x = const c, y = app (const c1) (const c2) & (c, c1, c2) \in s]. + [/\ x = const c, y = app (const c1) (const c2) & + (c, c1, c2) \in ss]. Proof. split. -- elim: s=>[|[[c c1] c2] s IH] /=; first by case. +- elim: ss=>[|[[c c1] c2] ss IH] /=; first by case. case; first by case=>->->; exists c, c1, c2; rewrite inE eq_refl. move/IH=>[d][d1][d2][->->H]. by exists d, d1, d2; rewrite inE H orbT. -elim: s=>[|[[c c1] c2] s IH] [d][d1][d2][H1 H2] //. +elim: ss=>[|[[c c1] c2] ss IH] [d][d1][d2][H1 H2] //. rewrite inE; case/orP; first by rewrite H1 H2; case/eqP=>->->->; left. move=>T; apply: sub_orr; apply: IH. by exists d, d1, d2; rewrite H1 H2 T. Qed. -(* We also keep track of pending equations. These are equations that have *) -(* been placed into the input list, but have not yet been processed and *) +(* One must also keep track of pending equations. These are equations *) +(* placed into the input list, but have not yet been processed and *) (* inserted into the data structure for computing the relation. The *) (* pending equations are either simple equations of the form c1 = c2 or *) (* are composite ones of the form (c = c1 c2, d = d1 d2), where c1, d1 *) @@ -321,8 +425,7 @@ do ![case: eqP=>[->|H]; last by apply: ReflectF; case]; by apply: ReflectT. Qed. -Canonical Structure pend_eqMixin := EqMixin eq_pendP. -Canonical Structure pend_eqType := EqType _ pend_eqMixin. +HB.instance Definition _ := hasDecEq.Build pend eq_pendP. (* pending equations are interpreted as follows for *) (* purposes of computing the relations they encode *) @@ -332,11 +435,6 @@ Definition pend2eq (p : pend) := | comp_pend (c, c1, c2) (d, d1, d2) => simp_eq c d end. -(********************************************************************) -(* The definition of the data structures involved in the algorithm. *) -(********************************************************************) - - Record data : Type := Data {(* An array, indexed by symbs, containing for each symb its *) (* current representative (for the relation obtained by the *) @@ -355,41 +453,37 @@ Record data : Type := (* The list of equations pending to be processed *) pending : seq pend}. -(* We can collect the representatives of all the symbols into a finite list. *) -(* We remove the duplicates. *) +(* Collect the representatives of all the symbols *) +(* into a finite list. Remove the duplicates. *) Definition reps D : seq symb := undup (map (rep D) (enum predT)). Lemma uniq_reps D : uniq (reps D). Proof. by rewrite undup_uniq. Qed. -#[export] Hint Resolve uniq_reps : core. Lemma rep_in_reps D a : rep D a \in reps D. Proof. by move=>*; rewrite mem_undup; apply: map_f; rewrite mem_enum. Qed. -#[export] Hint Resolve rep_in_reps : core. (* The symbols and their representatives are a base case in defining the closure *) Definition rep2rel D := graph (fun x => const (rep D x)). -Lemma clos_rep D R a : (const a, const (rep D a)) \In closure (rep2rel D +p R). +Lemma clos_rep D R a : (const a, const (rep D a)) \In closure (rep2rel D \+p R). Proof. -apply: (@closI _ (const a, const (rep D a))); apply: sub_orl; rewrite /= /rep2rel /graph /=. -by rewrite mem_map /= ?mem_enum // => x1 x2 [->]. +by apply: closI; apply: sub_orl; rewrite InE mem_map ?mem_enum // =>x1 x2 []. Qed. -#[export] Hint Resolve clos_rep : core. (* relation defined by the lookup table *) Definition lookup_rel D : rel_exp := - [Pred e1 e2 | + [Pred e | exists a b c c1 c2, - [/\ e1 = app (const a) (const b), + [/\ e.1 = app (const a) (const b), a \in reps D, b \in reps D, - fnd (a, b) (lookup D) = Some (c, c1, c2) & e2 = const (rep D c)]]. + fnd (a, b) (lookup D) = Some (c, c1, c2) & e.2 = const (rep D c)]]. (* inverting relations equations represented as triples *) Lemma invert_look D (x y : exp) : @@ -405,22 +499,34 @@ by exists a, b, c, c1, c2. Qed. (* The relation defined by the data structure is the following *) -Definition CRel D := closure (rep2rel D +p - lookup_rel D +p +Definition CRel D := closure (rep2rel D \+p + lookup_rel D \+p eqs2rel (map pend2eq (pending D))). Lemma cong_rel D : congruence (CRel D). Proof. by move=>*; apply: cong_clos. Qed. -#[export] Hint Resolve cong_rel : core. -Lemma clos_rel D R a : (const a, const (rep D a)) \In closure (CRel D +p R). +Lemma clos_rel D R a : (const a, const (rep D a)) \In closure (CRel D \+p R). Proof. by rewrite /CRel clos_clos orrA; apply: clos_rep. Qed. -(*******************************************) -(* Congruence Closure Code and Termination *) -(*******************************************) +End Structures. + +Hint Resolve uniq_reps rep_in_reps clos_rep cong_rel : core. + +(************************************) +(* Mathematical (purely-functional) *) +(* model of the algorithm. *) +(* Also, proof of termination. *) +(************************************) + +Section MathModel. +Variable s : seq constant. +Notation symb := (symb s). +Notation exp := (exp s). +Notation rel_exp := (rel_exp s). +Notation data := (data s). (* termination metric is the number of equations in the use and pending lists *) Definition metric (D : data) : nat := @@ -447,10 +553,11 @@ Lemma join_class_perm (D : data) (a' b' : symb) : (filter (predC1 a') (reps D)). Proof. have ffun_comp: forall (A B C : finType) (f1 : B -> A) (f2 : C -> B), - finfun (f1 \o f2) =1 f1 \o f2 :> (C -> A) by move=>A B C f1 f2 x; rewrite ffunE. + finfun (f1 \o f2) =1 f1 \o f2 :> (C -> A). +- by move=>A B C f1 f2 x; rewrite ffunE. have maps_ffun_comp : forall (A B C : finType) (f1 : B -> C) (f2 : A -> B) s, map (finfun (comp f1 f2)) s = map f1 (map f2 s). -- by move=>A B C f1 f2 s; elim: s=>[|x s IH] //=; rewrite IH /= ffunE. +- by move=>A B C f1 f2 ss; elim: ss=>[|x ss IH] //=; rewrite IH /= ffunE. move=>H H2. apply: uniq_perm; [|apply: filter_uniq|]; try by rewrite undup_uniq. move: H2. @@ -462,17 +569,17 @@ have L1 : forall x, a' != f x. have L2 : forall x, x != a' -> x = f x. - by rewrite /f=>x; case: eqP. rewrite eq_sym in H. -move: (map _ _)=>s H1 x. +move: (map _ _)=>ss H1 x. rewrite !mem_undup. case E1 : (x == a'). - rewrite (eqP E1) mem_filter /= eq_refl /=. - elim: s {H1} =>// t s H1. + elim: ss {H1} =>// t ss H1. by rewrite inE /= (negbTE (L1 t)) H1. case E2 : (x == b'). - rewrite mem_filter /= E1 (eqP E2) H1 (L2 b') //=. by apply: map_f. rewrite mem_filter /= E1 /=. -elim: s {H1}=>// t s; rewrite /= 2!inE=>->. +elim: ss {H1}=>// t ss; rewrite /= 2!inE=>->. case: (t =P a')=>[->|H2]. - by rewrite E1 /f eq_refl E2. by rewrite -(L2 t) //; case: eqP H2. @@ -492,31 +599,33 @@ Lemma join_class_metric (D : data) (a' b' : symb) : Proof. move=>H1 H2 H3; rewrite /metric /= -addnA; congr plus. by rewrite (perm_big _ (join_class_perm H3 H2)) - -(perm_big _ (permEl (perm_filterC (pred1 a') (reps D)))) - big_cat filter_pred1_uniq // !big_filter big_cons /= big_nil addn0 addnC. + -(perm_big _ (permEl (perm_filterC (pred1 a') (reps D)))) + big_cat filter_pred1_uniq // !big_filter big_cons /= big_nil addn0 addnC. Qed. (* joining the use lists *) Fixpoint join_use' (D : data) (a' b' : symb) - (old_use : seq (symb * symb * symb)) {struct old_use} : data := + (old_use : seq (symb * symb * symb)) {struct old_use} : data := match old_use with | [::] => D | (c, c1, c2)::p' => if fnd (rep D c1, rep D c2) (lookup D) is Some (d, d1, d2) then - let: newD := Data (rep D) (class D) - (finfun (fun r => if r == a' then behead (use D r) else - use D r)) - (lookup D) - (comp_pend (c, c1, c2) (d, d1, d2) :: pending D) - in join_use' newD a' b' p' + let: newD := + Data (rep D) (class D) + (finfun (fun r => if r == a' then behead (use D r) + else use D r)) + (lookup D) + (comp_pend (c, c1, c2) (d, d1, d2) :: pending D) + in join_use' newD a' b' p' else - let: newD := Data (rep D) (class D) - (finfun (fun r => if r == a' then behead (use D r) else - if r == b' then (c, c1, c2) :: use D r else - use D r)) - (ins (rep D c1, rep D c2) (c, c1, c2) (lookup D)) - (pending D) - in join_use' newD a' b' p' + let: newD := + Data (rep D) (class D) + (finfun (fun r => if r == a' then behead (use D r) else + if r == b' then (c, c1, c2) :: use D r + else use D r)) + (ins (rep D c1, rep D c2) (c, c1, c2) (lookup D)) + (pending D) + in join_use' newD a' b' p' end. Definition join_use D a' b' := join_use' D a' b' (use D a'). @@ -527,7 +636,8 @@ Lemma join_use_metric (D : data) (a' b' : symb) : metric (join_use D a' b') = (metric D) + size (use D a'). Proof. rewrite /join_use; move E: (use D a')=>old_use. -elim: old_use D E=>[|[[c c1] c2] old_use IH] D E H1 H2 H3 /=; first by rewrite addn0. +elim: old_use D E=>[|[[c c1] c2] old_use IH] D E H1 H2 H3 /=. +- by rewrite addn0. move: (mem_split_uniq H2 (uniq_reps D))=>[s1][s2][L1 L2 L3]. have L4 : undup (map (rep D) (enum predT)) = reps D by []. case: (fnd _ _)=>[[[d d1] d2]|]; last first. @@ -556,20 +666,19 @@ case S2: (x == a')=>//=. by rewrite ffunE S2. Qed. -Let pend0 (e : pend) := - match e with simp_pend a b => a | comp_pend (a,_,_) (b,_,_) => a end. -Let pend1 (e : pend) := - match e with simp_pend a b => b | comp_pend (a,_,_) (b,_,_) => b end. -Notation "e ..0" := (pend0 e) (at level 2). -Notation "e ..1" := (pend1 e) (at level 2). +(* left/right symbol in pending equation *) +Definition pendL (e : pend s) := + let: (simp_pend a _ | comp_pend (a,_,_) _) := e in a. +Definition pendR (e : pend s) := + let: (simp_pend _ b | comp_pend _ (b,_,_)) := e in b. (* loop through the equations in the pending list *) Function propagate (D : data) {measure metric D} : data := match (pending D) with | [::] => D | e :: p' => - let: a' := rep D (e..0) in - let: b' := rep D (e..1) in + let: a' := rep D (pendL e) in + let: b' := rep D (pendR e) in let: D' := Data (rep D) (class D) (use D) (lookup D) p' in if a' == b' then propagate D' else @@ -583,7 +692,7 @@ Proof. by rewrite /metric eq addSn; apply: ltP; apply: ltnSn. move=>D e p' eq H. rewrite join_use_metric ?H //; last first. -- by rewrite join_class_eq ?mem_filter ?rep_in_reps //= ?(eq_sym (rep D e..1)) H. +- by rewrite join_class_eq ?mem_filter ?rep_in_reps //= ?(eq_sym (rep D (pendR e))) H. - by rewrite join_class_eq ?mem_filter ?rep_in_reps ?H //= eq_refl. rewrite -join_class_metric ?H ?rep_in_reps //. by apply: ltP; rewrite /metric /reps /= eq /= addSn. @@ -603,9 +712,20 @@ Fixpoint norm (D : data) (t : exp) {struct t} : exp := end end. -(************************************) -(* Some invariants of the algorithm *) -(************************************) +End MathModel. + +(*******************************) +(* Invariants of the algorithm *) +(*******************************) + +Section Invariants. +Variable s : seq constant. +Notation symb := (symb s). +Notation exp := (exp s). +Notation rel_exp := (rel_exp s). +Notation data := (data s). +Notation pend2eq := (@pend2eq s). +Implicit Type D : data. (* the rep function is idempotent *) Definition rep_idemp D := forall a, rep D (rep D a) = rep D a. @@ -615,27 +735,34 @@ Definition class_inv D := forall x a, (rep D x == a) = (x \in class D a). (* use lists only store equations with appropriate representatives *) Definition use_inv D := - forall a c c1 c2, a \in reps D -> (c, c1, c2) \in use D a -> rep D c1 = a \/ rep D c2 = a. + forall a c c1 c2, a \in reps D -> (c, c1, c2) \in use D a -> + rep D c1 = a \/ rep D c2 = a. -(* intermediary invariant which holds after the class of a' has been joined to b' *) +(* intermediary invariant which holds after *) +(* the class of a' has been joined to b' *) Definition use_inv1 D a' b' := forall c c1 c2, - (forall a, a \in reps D -> (c, c1, c2) \in use D a -> rep D c1 = a \/ rep D c2 = a) /\ + (forall a, a \in reps D -> (c, c1, c2) \in use D a -> + rep D c1 = a \/ rep D c2 = a) /\ ((c, c1, c2) \in use D a' -> rep D c1 = b' \/ rep D c2 = b'). -(* lookup table only stores equations with appropriate representatives *) +(* lookup table only stores equations with *) +(* appropriate representatives *) Definition lookup_inv D := forall a b c c1 c2, a \in reps D -> b \in reps D -> - fnd (a, b) (lookup D) = Some (c, c1, c2) -> rep D c1 = a /\ rep D c2 = b. + fnd (a, b) (lookup D) = Some (c, c1, c2) -> + rep D c1 = a /\ rep D c2 = b. (* two symbols are similar if they are either related by representatives *) (* or are to be related after the pending list is processed *) Definition similar D a b := - (const a, const b) \In closure (rep2rel D +p eqs2rel (map pend2eq (pending D))). + (const a, const b) \In closure (rep2rel D \+p eqs2rel + (map pend2eq (pending D))). Definition similar1 D a' b' a b := - (const a, const b) \In closure (rep2rel D +p [Pred x y | x = const a' /\ y = const b'] +p - eqs2rel (map pend2eq (pending D))). + (const a, const b) \In + closure (rep2rel D \+p Pred1 (const a', const b') \+p + eqs2rel (map pend2eq (pending D))). (* invariants tying use lists with the lookup table *) Definition use_lookup_inv D := @@ -646,11 +773,14 @@ Definition use_lookup_inv D := Definition lookup_use_inv D := forall a b d d1 d2, - a \in reps D -> b \in reps D -> fnd (a, b) (lookup D) = Some (d, d1, d2) -> + a \in reps D -> b \in reps D -> + fnd (a, b) (lookup D) = Some (d, d1, d2) -> [/\ exists c c1 c2, - (c, c1, c2) \in use D a /\ rep D c1 = a /\ rep D c2 = b /\ similar D c d & + (c, c1, c2) \in use D a /\ rep D c1 = a /\ + rep D c2 = b /\ similar D c d & exists c c1 c2, - (c, c1, c2) \in use D b /\ rep D c1 = a /\ rep D c2 = b /\ similar D c d]. + (c, c1, c2) \in use D b /\ rep D c1 = a /\ + rep D c2 = b /\ similar D c d]. (* intermediary invariants after removing an equation from the pending list *) Definition use_lookup_inv1 D a' b' := @@ -663,9 +793,11 @@ Definition lookup_use_inv1 D a' b' := forall a b d d1 d2, a \in reps D -> b \in reps D -> fnd (a, b) (lookup D) = Some (d, d1, d2) -> [/\ exists c c1 c2, - (c, c1, c2) \in use D a /\ rep D c1 = a /\ rep D c2 = b /\ similar1 D a' b' c d & + (c, c1, c2) \in use D a /\ rep D c1 = a /\ + rep D c2 = b /\ similar1 D a' b' c d & exists c c1 c2, - (c, c1, c2) \in use D b /\ rep D c1 = a /\ rep D c2 = b /\ similar1 D a' b' c d]. + (c, c1, c2) \in use D b /\ rep D c1 = a /\ + rep D c2 = b /\ similar1 D a' b' c d]. (* intermediary invariants after join_class and during join_use *) Definition use_lookup_inv2 D a' b' := @@ -708,40 +840,64 @@ Definition use_lookup_inv2 D a' b' := Definition lookup_use_inv2 D a' b' := forall d d1 d2, - [/\ forall b, b \in reps D -> fnd (a', b) (lookup D) = Some (d, d1, d2) -> + [/\ forall b, b \in reps D -> + fnd (a', b) (lookup D) = Some (d, d1, d2) -> rep D d1 = b' /\ exists c c1 c2, - (c, c1, c2) \in use D b /\ rep D c1 = b' /\ rep D c2 = b /\ similar D c d, - forall a, a \in reps D -> fnd (a, a') (lookup D) = Some (d, d1, d2) -> + (c, c1, c2) \in use D b /\ rep D c1 = b' /\ + rep D c2 = b /\ similar D c d, + forall a, a \in reps D -> + fnd (a, a') (lookup D) = Some (d, d1, d2) -> rep D d2 = b' /\ exists c c1 c2, - (c, c1, c2) \in use D a /\ rep D c1 = a /\ rep D c2 = b' /\ similar D c d & + (c, c1, c2) \in use D a /\ rep D c1 = a /\ + rep D c2 = b' /\ similar D c d & forall a b, a \in reps D -> b \in reps D -> fnd (a, b) (lookup D) = Some (d, d1, d2) -> [/\ exists c c1 c2, - (c, c1, c2) \in use D a /\ rep D c1 = a /\ rep D c2 = b /\ similar D c d & + (c, c1, c2) \in use D a /\ rep D c1 = a /\ + rep D c2 = b /\ similar D c d & exists c c1 c2, - (c, c1, c2) \in use D b /\ rep D c1 = a /\ rep D c2 = b /\ similar D c d]]. + (c, c1, c2) \in use D b /\ rep D c1 = a /\ + rep D c2 = b /\ similar D c d]]. (* the invariant of the propagate routine *) Definition propagate_inv D := - rep_idemp D /\ use_inv D /\ lookup_inv D /\ use_lookup_inv D /\ lookup_use_inv D. + rep_idemp D /\ use_inv D /\ lookup_inv D /\ + use_lookup_inv D /\ lookup_use_inv D. + +End Invariants. + +(**********************************) +(* Verification of the math model *) +(**********************************) -(****************) -(* Verification *) -(****************) +(* proving properties of the mathematical model *) + +Section MathModelProofs. +Variable s : seq constant. +Notation symb := (symb s). +Notation exp := (exp s). +Notation rel_exp := (rel_exp s). +Notation data := (data s). +Notation pend2eq := (@pend2eq s). +Notation symb2eq := (@symb2eq s). +Notation closure := (@closure s). +Implicit Type D : data. (* first some basic rewrite rules *) -Lemma reps_rep (D : data) (a : symb) : rep_idemp D -> a \in reps D -> rep D a = a. +Lemma reps_rep D (a : symb) : + rep_idemp D -> a \in reps D -> rep D a = a. Proof. by move=>H; rewrite mem_undup; case/mapP=>x _ ->; apply: H. Qed. Lemma symR R x y : (x, y) \In closure R <-> (y, x) \In closure R. Proof. by split=>T; apply: symC. Qed. Lemma app_rep D R a b x : - (x, app (const a) (const b)) \In closure (rep2rel D +p R) <-> - (x, app (const (rep D a)) (const (rep D b))) \In closure (rep2rel D +p R). + (x, app (const a) (const b)) \In closure (rep2rel D \+p R) <-> + (x, app (const (rep D a)) (const (rep D b))) + \In closure (rep2rel D \+p R). Proof. split=>T. - apply: (transC (y:= app (const a) (const (rep D b)))); last first. @@ -755,18 +911,18 @@ by apply: monoC; [apply: symC; apply: clos_rep | apply: reflC]. Qed. Lemma app_rel D R a b x : - (x, app (const a) (const b)) \In closure (CRel D +p R) <-> - (x, app (const (rep D a)) (const (rep D b))) \In closure (CRel D +p R). + (x, app (const a) (const b)) \In closure (CRel D \+p R) <-> + (x, app (const (rep D a)) (const (rep D b))) \In closure (CRel D \+p R). Proof. by rewrite /CRel !clos_clos !orrA; apply: app_rep. Qed. Lemma const_rep D R a x : - (const a, x) \In closure (rep2rel D +p R) <-> - (const (rep D a), x) \In closure (rep2rel D +p R). + (const a, x) \In closure (rep2rel D \+p R) <-> + (const (rep D a), x) \In closure (rep2rel D \+p R). Proof. by split=>T; apply: transC T; [apply: symC|]; apply: clos_rep. Qed. Lemma const_rel D R a x : - (const a, x) \In closure (CRel D +p R) <-> - (const (rep D a), x) \In closure (CRel D +p R). + (const a, x) \In closure (CRel D \+p R) <-> + (const (rep D a), x) \In closure (CRel D \+p R). Proof. by rewrite /CRel !clos_clos !orrA; apply: const_rep. Qed. (* Now the lemmas *) @@ -776,13 +932,13 @@ Proof. by rewrite /CRel !clos_clos !orrA; apply: const_rep. Qed. (* table and the pending list all start empty. *) Definition init : data := - Data [ffun x => x] [ffun c => [:: c]] [ffun c => [::]] (nil _ _) [::]. + Data [ffun x => x] [ffun c => [:: c]] [ffun c => [::]] nil [::]. -Lemma initP : propagate_inv init /\ CRel init <~> closure (graph const). +Lemma initP : propagate_inv init /\ CRel init <~> closure (graph (@const s)). Proof. have S: lookup_rel init <~> Pred0. - by split=>//; case: x=> ??[?][?][?][?][?][]. -have E: graph const <~> graph (fun x => const ([ffun x => x] x)). +have E: graph (@const s) <~> graph (fun x => const ([ffun x => x] x)). - by split; case: x =>x1 x2 /=; case/imageP=>[x] _ [->->]; apply/imageP; exists x; rewrite ?ffunE. rewrite /CRel /= S 2!orr0 /=; split; last by rewrite E. @@ -795,18 +951,20 @@ Qed. (* Lemmas about join_class *) Lemma join_class_repE (D : data) (a' b' : symb) : - a' \in reps D -> b' \in reps D -> rep_idemp D -> + a' \in reps D -> + b' \in reps D -> + rep_idemp D -> closure (rep2rel (join_class D a' b')) <~> - closure (rep2rel D +p [Pred x y | x = const a' /\ y = const b']). + closure (rep2rel D \+p Pred1 (const a', const b')). Proof. move=>H2 H3 H1 [x y]; split=>/= T. -- pose ty z := PredU (rep2rel D) [Pred x0 y0 | x0 = const z /\ y0 = const b']. +- pose ty z := PredU (rep2rel D) (Pred1 (const z, const b')). move: (clos_idemp (ty a') (x,y))=>/=<-. apply: sub_closI T=>{x y} [[x y]]. case/imageP=>z _ /= [->->] {x y}; rewrite ffunE /=. case: eqP=>[<-|_]; last by apply: clos_rep. apply: (transC (y:=const (rep D z))); first by apply: clos_rep. - by apply: (@closI _ (const (rep D z), const b')); right. + by apply: closI; right. move: (clos_idemp (rep2rel (join_class D a' b')) (x,y))=>/=<-. apply: sub_closI T=>{x y} [[x y]]. case; last first. @@ -814,16 +972,21 @@ case; last first. by rewrite /= !ffunE /= reps_rep // eq_refl. case/imageP=>z _ /= [->->] {x y}. case E: (rep D z == a'); last first. -- by apply: (@closI _ (const z, const (rep D z))); apply/imageP; exists z; rewrite // ffunE /= E. +- apply: closI; apply/imageP. + by exists z; rewrite // ffunE /= E. apply: (transC (y:=const b')). -- by apply: (@closI _ (const z, const b')); apply/imageP; exists z; rewrite // ffunE /= E. -apply: symC; apply: (@closI _ (const (rep D z), const b')); apply/imageP. +- apply: closI; apply/imageP. + by exists z; rewrite // ffunE /= E. +apply: symC; apply: closI; apply/imageP. by exists (rep D z); rewrite // ffunE /= H1 // E. Qed. Lemma join_class_useP (D : data) (a' b' : symb) : - a' \in reps D -> b' \in reps D -> a' != b' -> - use_inv D -> use_inv1 (join_class D a' b') a' b'. + a' \in reps D -> + b' \in reps D -> + a' != b' -> + use_inv D -> + use_inv1 (join_class D a' b') a' b'. Proof. move=>H1 H2 H3 H4 c c1 c2; split=>/= [a|H5]; rewrite !ffunE /=; last first. - by case: (H4 _ _ _ _ H1 H5)=>->; [left | right]; rewrite /= eq_refl. @@ -833,8 +996,11 @@ by case: (H4 _ _ _ _ H6 H7)=>->; [left | right]; rewrite (negbTE H5). Qed. Lemma join_class_lookupP (D : data) (a' b' : symb) : - a' \in reps D -> b' \in reps D -> a' != b' -> - lookup_inv D -> lookup_inv (join_class D a' b'). + a' \in reps D -> + b' \in reps D -> + a' != b' -> + lookup_inv D -> + lookup_inv (join_class D a' b'). Proof. move=>H1 H2 H3 H a b c c1 c2. rewrite !join_class_eq // !mem_filter /= !ffunE /=. @@ -844,17 +1010,20 @@ by rewrite (negbTE T1) (negbTE Q1). Qed. Lemma join_class_simE (D : data) (a' b' : symb) : - a' \in reps D -> b' \in reps D -> rep_idemp D -> - forall a b, similar1 D a' b' a b <-> similar (join_class D a' b') a b. + a' \in reps D -> + b' \in reps D -> + rep_idemp D -> + forall a b, similar1 D a' b' a b <-> + similar (join_class D a' b') a b. Proof. move=>H1 H2 H a b. by rewrite /similar /similar1 -orrA -clos_clos -join_class_repE // clos_clos. Qed. -Module Dummy321. End Dummy321. - Lemma join_class_classP (D : data) (a' b' : symb) : - a' != b' -> class_inv D -> class_inv (join_class D a' b'). + a' != b' -> + class_inv D -> + class_inv (join_class D a' b'). Proof. rewrite /join_class /class_inv=>E /= => H x a; rewrite !ffunE /=. case: ifP=>H1. @@ -867,36 +1036,40 @@ by rewrite H. Qed. Lemma join_classP (D : data) (a' b' : symb) : - a' \in reps D -> b' \in reps D -> a' != b' -> - rep_idemp D -> use_inv D -> lookup_inv D -> - use_lookup_inv1 D a' b' -> lookup_use_inv1 D a' b' -> + a' \in reps D -> + b' \in reps D -> + a' != b' -> + rep_idemp D -> + use_inv D -> + lookup_inv D -> + use_lookup_inv1 D a' b' -> + lookup_use_inv1 D a' b' -> use_lookup_inv2 (join_class D a' b') a' b' /\ lookup_use_inv2 (join_class D a' b') a' b'. Proof. move=>H1 H2 H3 L0 L1 L2 T1 T2. split; last first. -- split=>[b|a|a b]; rewrite !join_class_eq // !mem_filter /=; case/andP=>H4 H5. - - move=>F. - rewrite ffunE /=. +- split=>[b|a|a b]; rewrite !join_class_eq // !mem_filter /=; + case/andP=>H4 H5. + - move=>F; rewrite ffunE /=. case: (L2 a' b d d1 d2 H1 H5 F)=>R1 R2; rewrite R1 eq_refl; do !split=>//. case: (T2 a' b d d1 d2 H1 H5 F)=>_ [c][c1][c2][Q1][Q2][Q3]Q4. - exists c, c1, c2; rewrite !ffunE /= Q2 Q3 eq_refl (negbTE H4); do !split=>//. - by rewrite -join_class_simE. - - move=>F. - rewrite ffunE /=. + exists c, c1, c2; rewrite !ffunE /= Q2 Q3 eq_refl (negbTE H4). + by do !split=>//; rewrite -join_class_simE. + - move=>F; rewrite ffunE /=. case: (L2 a a' d d1 d2 H5 H1 F)=>R1 R2; rewrite R2 eq_refl; do !split=>//. case: (T2 a a' d d1 d2 H5 H1 F)=>[[c][c1][c2][Q1][Q2][Q3]Q4 _]. - exists c, c1, c2; rewrite !ffunE /= Q2 Q3 eq_refl (negbTE H4); do !split=>//. - by rewrite -join_class_simE. + exists c, c1, c2; rewrite !ffunE /= Q2 Q3 eq_refl (negbTE H4). + by do !split=>//; rewrite -join_class_simE. case/andP=>H6 H7 F. case: (T2 a b d d1 d2 H5 H7 F). move=>[c][c1][c2][Q1][Q2][Q3] Q4. move=>[e][e1][e2][F1][F2][F3] F4. split. - - exists c, c1, c2; rewrite !ffunE /= Q2 Q3 (negbTE H4) (negbTE H6); do !split=>//. - by rewrite -join_class_simE. - exists e, e1, e2; rewrite !ffunE /= F2 F3 (negbTE H4) (negbTE H6); do !split=>//. - by rewrite -join_class_simE. + - exists c, c1, c2; rewrite !ffunE /= Q2 Q3 (negbTE H4) (negbTE H6). + by do !split=>//; rewrite -join_class_simE. + exists e, e1, e2; rewrite !ffunE /= F2 F3 (negbTE H4) (negbTE H6). + by do !split=>//; rewrite -join_class_simE. split=>[c c1 c2 /= H4 | a c c1 c2]; last first. - rewrite join_class_eq // mem_filter /=. case/andP=>H4 H5 H6. @@ -930,97 +1103,100 @@ rewrite -join_class_simE=>//. by rewrite (eqP H6). Qed. -Module Dummy123. End Dummy123. - -Definition pull {T : Type} (r : Pred T) := (orrC r, orrCA r). - +Definition rpull {T : Type} (r : Pred T) := (orrC r, orrCA r). Lemma join_classE (D : data) (a' b' : symb) : - a' \in reps D -> b' \in reps D -> a' != b' -> - rep_idemp D -> use_lookup_inv1 D a' b' -> lookup_use_inv1 D a' b' -> - closure (CRel (join_class D a' b') +p eqs2rel (map symb2eq (use D a'))) <~> - closure (CRel D +p [Pred x y | x = const a' /\ y = const b']). + a' \in reps D -> + b' \in reps D -> + a' != b' -> + rep_idemp D -> + use_lookup_inv1 D a' b' -> + lookup_use_inv1 D a' b' -> + closure (CRel (join_class D a' b') \+p + eqs2rel (map symb2eq (use D a'))) <~> + closure (CRel D \+p Pred1 (const a', const b')). Proof. -pose ty := [Pred x0 y0 | x0 = const a' /\ y0 = const b']. -move=>L1 L2 L3 H1 H4 H5 [x y]; split. +pose ty := Pred1 (const a', const b'). +move=>L1 L2 L3 H1 H4 H5 [x y]; rewrite !toPredE; split. - apply: clos_or; move=>{x y}[x y]; last first. - - move: (clos_idemp (CRel D +p ty) (x,y))=><-. + - rewrite -(clos_idemp (CRel D \+p ty)). apply: sub_closI=>{x y}[[x y]] /=. move/invert=>[c][c1][c2][->->H6]. move: (H4 a' c c1 c2 L1 H6)=>[d][d1][d2][Q1][Q2][Q3] Q4. apply/app_rel. apply: (transC (y := const d)). - rewrite /CRel clos_clos !orrA. - rewrite -!(pull (eqs2rel _)) -3!(pull ty) -!(pull (rep2rel _)) -!orrA. - move: Q4; apply: sub_closI=>{x y} [[x y]] Q4; apply: sub_orl. - by rewrite !toPredE in Q4 *; rewrite orrA. + rewrite -!(rpull (eqs2rel _)) -3!(rpull ty) -!(rpull (rep2rel _)) -!orrA. + move: Q4; apply: sub_closI=>{x y} [[x y]] Q4; apply: sub_orl. + by rewrite orrA. rewrite const_rel /CRel clos_clos !orrA symR. - apply: (@closI _ (app (const (rep D c1)) (const (rep D c2)), const (rep D d))). - apply: sub_orr; apply: sub_orl. - rewrite toPredE invert_look. + apply: closI; apply: sub_orr; apply: sub_orl. + rewrite invert_look. exists (rep D d1), (rep D d2), d, d1, d2. by rewrite -Q2 -Q3 !rep_in_reps. - rewrite /CRel !toPredE clos_idemp clos_clos !orrA /=. + rewrite /CRel clos_idemp clos_clos !orrA /=. rewrite -clos_clos join_class_repE // clos_clos orrA. - rewrite -!(pull (eqs2rel (map pend2eq _))). + rewrite -!(rpull (eqs2rel (map pend2eq _))). apply: sub_closKR=>{x y}; apply: sub_closKR=>[[x y]] T1. - rewrite toPredE -clos_idemp. + rewrite -clos_idemp. apply: sub_closI T1=>{x y} [[x y]]. - rewrite toPredE. case=>[T|]; first by apply: closI; right. - case=>a[b][c][c1][c2][-> +++ ->]. + case=>a [b][c][c1][c2][/= -> +++ ->]. rewrite !join_class_eq //= !mem_filter /=. case/andP=>T1 T1'; case/andP=>T2 T2' T3; rewrite ffunE /=. case: eqP=>[H6|_]; [apply: (transC (y:= const a')); last by [apply closI; right] |]; - apply closI; apply: sub_orl; rewrite toPredE invert_look; + apply closI; apply: sub_orl; rewrite invert_look; exists a, b, c, c1, c2=>//. by rewrite H6. -rewrite /CRel !toPredE !clos_clos !orrA /= => {H4} T. -rewrite -clos_clos join_class_repE // clos_clos !orrA -(pull ty). -rewrite -3!(pull ty) -!(pull (rep2rel _)) -!(pull (lookup_rel _)) in T *. +rewrite /CRel !clos_clos !orrA /= => {H4} T. +rewrite -clos_clos join_class_repE // clos_clos !orrA -(rpull ty). +rewrite -3!(rpull ty) -!(rpull (rep2rel _)) -!(rpull (lookup_rel _)) in T *. rewrite -clos_idemp; apply: sub_closI T=>{x y} [[x y]]. case=>[|T]; last by apply closI; apply: sub_orr; - rewrite toPredE -!orrA; apply: sub_orl; rewrite toPredE !orrA. -case=>a[b][c][c1][c2][-> T1 T2 T3 ->]. + rewrite -!orrA; apply: sub_orl; rewrite !orrA. +case=>a [b][c][c1][c2][/= -> T1 T2 T3 ->]. case: (H5 a b c c1 c2 T1 T2 T3). move=>[d][d1][d2][Q1][Q2][Q3] Q4 [f][f1][f2][F1][F2][F3] F4. -case E1: (a == a'). -- rewrite toPredE !(pull (lookup_rel _)) !orrA -Q2 -Q3 symR -app_rep -const_rep symR. +case E1: (a == a'). +- rewrite !(rpull (lookup_rel _)) !orrA -Q2 -Q3 symR + -app_rep -const_rep symR. apply: (transC (y:= const d)); last first. - move: Q4; apply: sub_closI=>{x y} [[x y]]. - by rewrite !toPredE -!orrA=>Q4; do 2!apply: sub_orl. - apply: symC; apply closI; do 3!apply: sub_orr; apply: sub_orl; rewrite toPredE invert -(eqP E1). + by rewrite -!orrA=>Q4; do 2!apply: sub_orl. + apply: symC; apply closI; do 3!apply: sub_orr; apply: sub_orl. + rewrite invert -(eqP E1). by exists d, d1, d2. case E2: (b == a'). -- rewrite toPredE !(pull (lookup_rel _)) !orrA -F2 -F3 symR -app_rep -const_rep symR. +- rewrite !(rpull (lookup_rel _)) !orrA. + rewrite -F2 -F3 symR -app_rep -const_rep symR. apply: (transC (y:= const f)); last first. - move: F4; apply: sub_closI=>{x y} [[x y]]. - by rewrite !toPredE -!orrA=>F4; do 2!apply: sub_orl. + by rewrite -!orrA=>F4; do 2!apply: sub_orl. apply: symC; apply closI; do 3!apply: sub_orr; apply: sub_orl. - rewrite toPredE invert -(eqP E2). + rewrite invert -(eqP E2). by exists f, f1, f2. case E3: (rep D c == a'); last first. -- apply: closI; apply: sub_orl; rewrite toPredE invert_look. +- apply: closI; apply: sub_orl; rewrite invert_look. exists a, b, c, c1, c2. by rewrite !join_class_eq // !mem_filter /= E1 E2 !ffunE /= E3. rewrite (eqP E3). apply: (transC (y := const b')); last first. - by apply: symC; apply closI; do 2!apply: sub_orr; apply: sub_orl. -apply closI; apply: sub_orl; rewrite toPredE invert_look. +apply closI; apply: sub_orl; rewrite invert_look. exists a, b, c, c1, c2. by rewrite !join_class_eq // !mem_filter /= E1 E2 !ffunE /= E3. Qed. -Module Dummy1. End Dummy1. - (* Lemmas about join_use *) Lemma join_use_classP (D : data) (a' b' : symb) : - a' != b' -> class_inv D -> class_inv (join_use D a' b'). + a' != b' -> + class_inv D -> + class_inv (join_use D a' b'). Proof. rewrite /class_inv /join_use; move E: (use _ _)=>x. -elim: x D E=>[|[[c c1] c2] s IH] D E H1 H2 x a //=. +elim: x D E=>[|[[c c1] c2] ss IH] D E H1 H2 x a //=. by case F : (fnd _ _)=>[[[d d1] d2]|]; rewrite IH //= ffunE eq_refl E. Qed. @@ -1030,22 +1206,26 @@ Lemma join_useT t (D : data) (a' b' : symb) (l1 l2 : seq (symb*symb*symb)) : join_use' (join_use' D a' b' l1) a' b' l2. Proof. elim: l1 D a' b'=>[|[[c1 c2] c] l1 IH] D a' b' //= H. -by case F: (fnd _ _)=>[[[d d1] d2]|] /=; rewrite IH //= ffunE eq_refl H /=; eauto. +case F: (fnd _ _)=>[[[d d1] d2]|] /=; +by rewrite IH //= ffunE eq_refl H /=; eauto. Qed. Lemma join_use_useE (D : data) (a' b' : symb) (l1 l2 : seq (symb*symb*symb)) : - use D a' = l1 ++ l2 -> use (join_use' D a' b' l1) a' = l2. + use D a' = l1 ++ l2 -> + use (join_use' D a' b' l1) a' = l2. Proof. elim: l1 D a' b'=>[|[[c1 c2] c] l1 IH] D a' b' //= H. by case F: (fnd _ _)=>[[[d d1] d2]|] //=; apply: IH; rewrite /= ffunE eq_refl H. Qed. Lemma join_use_useP (D : data) (a' b' : symb) : - a' \notin reps D -> b' \in reps D -> - use_inv1 D a' b' -> use_inv1 (join_use D a' b') a' b'. + a' \notin reps D -> + b' \in reps D -> + use_inv1 D a' b' -> + use_inv1 (join_use D a' b') a' b'. Proof. rewrite /join_use; move E: (use _ _)=>x. -elim: x D E=>[|[[c c1] c2] s IH] D E H1 H2 H3 //=. +elim: x D E=>[|[[c c1] c2] ss IH] D E H1 H2 H3 //=. case F: (fnd _ _)=>[[[d d1] d2]|]; (apply: IH=>//=; first by [rewrite ffunE eq_refl E]; move=>e e1 e2; rewrite /reps /=; split=>[a|]; rewrite ffunE; @@ -1061,11 +1241,13 @@ by rewrite E inE eq_refl. Qed. Lemma join_use_lookupP (D : data) (a' b' : symb) : - a' \notin reps D -> b' \in reps D -> - lookup_inv D -> lookup_inv (join_use D a' b'). + a' \notin reps D -> + b' \in reps D -> + lookup_inv D -> + lookup_inv (join_use D a' b'). Proof. rewrite /join_use; move E: (use _ _)=>x. -elim: x D E=>[|[[c c1] c2] s IH] D E H1 H2 H3 //=. +elim: x D E=>[|[[c c1] c2] ss IH] D E H1 H2 H3 //=. case F: (fnd _ _)=>[[[d d1] d2]|]. - by apply: IH=>//=; first by rewrite ffunE eq_refl E. apply: IH=>//=; first by rewrite ffunE eq_refl E. @@ -1074,47 +1256,55 @@ by case: eqP=>[[-> -> _ _] [_ -> ->]|_]; last by apply: H3. Qed. Lemma join_usePE (D : data) (a' b' : symb) : - a' \notin reps D -> b' \in reps D -> - rep_idemp D -> use_inv1 D a' b' -> lookup_inv D -> - use_lookup_inv2 D a' b' -> lookup_use_inv2 D a' b' -> + a' \notin reps D -> + b' \in reps D -> + rep_idemp D -> + use_inv1 D a' b' -> + lookup_inv D -> + use_lookup_inv2 D a' b' -> + lookup_use_inv2 D a' b' -> rep_idemp (join_use D a' b') /\ use_inv1 (join_use D a' b') a' b' /\ lookup_inv (join_use D a' b') /\ use_lookup_inv2 (join_use D a' b') a' b' /\ lookup_use_inv2 (join_use D a' b') a' b' /\ - closure (CRel D +p eqs2rel (map symb2eq (use D a'))) <~> + closure (CRel D \+p eqs2rel (map symb2eq (use D a'))) <~> CRel (join_use D a' b'). Proof. rewrite /join_use; move E: (use _ _)=>x. -elim: x D E=>[|[[c c1] c2] s IH] D E L1 L2 H1 H2 H3 H4 H5 /=. +elim: x D E=>[|[[c c1] c2] ss IH] D E L1 L2 H1 H2 H3 H4 H5 /=. - by rewrite orr0 /CRel clos_idemp. case F: (fnd _ _)=>[[[d d1] d2]|]; set D' := Data _ _ _ _ _. -- have S1: closure (rep2rel D' +p eqs2rel (map pend2eq (pending D'))) <~> - closure (rep2rel D +p [Pred x y | x = const c /\ y = const d] +p +- have S1: closure (rep2rel D' \+p eqs2rel (map pend2eq (pending D'))) <~> + closure (rep2rel D \+p Pred1 (const c, const d) \+p eqs2rel (map pend2eq (pending D))) by []. have S2: forall e1 e2, similar D e1 e2 -> similar D' e1 e2. - rewrite /similar=>e1 e2; move: (const e1) (const e2)=>{e1 e2} x y. - by rewrite S1; apply: sub_closI=>{x y} [[x y]]; case; [left | right; right]. - have T2 : closure (CRel D' +p eqs2rel (map symb2eq s)) <~> - closure (CRel D +p [Pred x y | x = const c /\ y = app (const c1) (const c2)] +p - eqs2rel (map symb2eq s)). + by rewrite S1; apply: sub_closI=>{x y} [[x y]]; + case; [left | right; right]. + have T2 : closure (CRel D' \+p eqs2rel (map symb2eq ss)) <~> + closure (CRel D \+p Pred1 (const c, app (const c1) (const c2)) + \+p eqs2rel (map symb2eq ss)). - rewrite /CRel {2 3}/D' /= !clos_clos !orrA. - rewrite -!(pull (eqs2rel (map pend2eq _))) -!(pull (eqs2rel (map symb2eq s))). + rewrite -!(rpull (eqs2rel (map pend2eq _))). + rewrite -!(rpull (eqs2rel (map symb2eq ss))). do 2!apply: closKR. rewrite -!orrA=>[[x y]]; split=>T; rewrite toPredE -clos_idemp; apply: sub_closI T=>{x y} [[x y]]. - case; first by move=>H7; apply: closI; apply: sub_orl. - case=>->->; rewrite toPredE !orrA symR const_rep symR. - apply: (transC (y:= app (const (rep D c1)) (const (rep D c2)))); last first. - - apply closI; apply: sub_orr; apply: sub_orl; rewrite toPredE invert_look. + case=>->->; rewrite !orrA symR const_rep symR. + apply: (transC (y:= app (const (rep D c1)) (const (rep D c2)))); + last first. + - apply closI; apply: sub_orr; apply: sub_orl. + rewrite invert_look. by exists (rep D c1), (rep D c2), d, d1, d2. by rewrite -app_rep; apply closI; do 2!right. case; first by move=>H7; apply closI; apply: sub_orl. - case=>->->; rewrite toPredE !orrA app_rep. + case=>->->; rewrite !orrA app_rep. apply: (transC (y := const d)); first by apply closI; right; right. apply: (transC (y := const (rep D d))); first by apply: clos_rep. apply: symC; apply closI; apply: sub_orr; apply: sub_orl. - rewrite toPredE invert_look. + rewrite invert_look. exists (rep D c1), (rep D c2), d, d1, d2. by rewrite !rep_in_reps. rewrite -T2. @@ -1132,8 +1322,10 @@ case F: (fnd _ _)=>[[[d d1] d2]|]; set D' := Data _ _ _ _ _. by exists f, f1, f2; rewrite Q2 Q3; do ![split=>//]; apply: S2. case: eqP L3 L1=>[->-> //|_] L3 L1 H6. case: (H4.2 a e e1 e2 L3 H6)=>[[R1][R2]|[R]|[R1][R2]|[R]]; - [move=>[[h][h1][h2][F1][F2][F3] F4]| idtac | move=>[[h][h1][h2][F1][F2][F3] F4]| idtac]; - move=>[f][f1][f2][Q1][Q2][Q3] Q4; [idtac | apply: Or42 | idtac | apply: Or44]; + [move=>[[h][h1][h2][F1][F2][F3] F4]| idtac | + move=>[[h][h1][h2][F1][F2][F3] F4]| idtac]; + move=>[f][f1][f2][Q1][Q2][Q3] Q4; + [idtac | apply: Or42 | idtac | apply: Or44]; try by [do !split=>//; exists f, f1, f2; do !split=>//; apply: S2]; rewrite E inE in F1; case/orP: F1=>F1; [apply: Or42 | apply: Or41 | apply: Or44 | apply: Or43]; @@ -1162,11 +1354,10 @@ case F: (fnd _ _)=>[[[d d1] d2]|]; set D' := Data _ _ _ _ _. by case: eqP H6 L1=>[->-> //| _] H6 L1; do !split=>//; apply: S2. exists h, h1, h2; rewrite {1}/D' /= !ffunE. by case: eqP H7 L1=>[->-> //| _] H7 L1; do !split=>//; apply: S2. -have T1: lookup_rel D' <~> lookup_rel D +p - [Pred x y | x = app (const (rep D c1)) (const (rep D c2)) /\ - y = const (rep D c)]. +have T1: lookup_rel D' <~> lookup_rel D \+p + Pred1 (app (const (rep D c1)) (const (rep D c2)), const (rep D c)). - rewrite /lookup_rel /D' /= /reps /= =>[[x y]]; split. - - move=>[a][b][d][d1][d2][-> T2 T3]. + - move=>[a][b][d][d1][d2][/= -> T2 T3]. rewrite fnd_ins. case: eqP=>[[->->] [<- _ _] ->|_ T4 T5]; [by right | left]. by exists a, b, d, d1, d2. @@ -1180,19 +1371,20 @@ have T1: lookup_rel D' <~> lookup_rel D +p exists (rep D c1), (rep D c2), c, c1, c2. rewrite fnd_ins !rep_in_reps. by case: eqP. -have T2: closure (CRel D' +p eqs2rel (map symb2eq s)) <~> - closure (CRel D +p [Pred x y | x = const c /\ y = app (const c1) (const c2)] +p - eqs2rel (map symb2eq s)). -- pose ty1 := [Pred x y | x = app (const (rep D c1)) (const (rep D c2)) /\ y = const (rep D c)]. - pose ty2 := [Pred x y | x = const c /\ y = app (const c1) (const c2)]. - rewrite /CRel {3}/D' /= T1 !clos_clos !orrA -2!(pull ty1) -3!(pull ty2) -!orrA -!(pull (rep2rel _)). +have T2: closure (CRel D' \+p eqs2rel (map symb2eq ss)) <~> + closure (CRel D \+p Pred1 (const c, app (const c1) (const c2)) + \+p eqs2rel (map symb2eq ss)). +- pose ty1 := Pred1 (app (const (rep D c1)) (const (rep D c2)), const (rep D c)). + pose ty2 := Pred1 (const c, app (const c1) (const c2)). + rewrite /CRel {3}/D' /= T1 !clos_clos !orrA -2!(rpull ty1) + -3!(rpull ty2) -!orrA -!(rpull (rep2rel _)). do 3!apply: closRK; move=>[x y]; split=>T; rewrite toPredE -clos_idemp; apply: sub_closI T=>{x y} [[x y]]. - case; first by move=>T; apply: closI; left. - case=>->->; rewrite toPredE symR -const_rep -app_rep. + case=>->->; rewrite symR -const_rep -app_rep. by apply closI; right. case=>[T|[->->]]; first by apply: closI; left. - by rewrite toPredE const_rep app_rep symR; apply closI; right. + by rewrite const_rep app_rep symR; apply closI; right. rewrite -T2. apply: IH=>//=; first by rewrite ffunE eq_refl E. - move=>e e1 e2; split=>/= [a H6|]; rewrite !ffunE; last first. @@ -1218,12 +1410,14 @@ apply: IH=>//=; first by rewrite ffunE eq_refl E. case: eqP=>[->|H7 H8]; [rewrite inE; case/orP=>[|H8] | idtac]. - case/eqP=>->->->. have H8: (c, c1, c2) \in use D a' by [rewrite E inE eq_refl]. - case: ((H2 c c1 c2).2 H8)=>H9; [apply: Or42 | apply: Or44]; do ![split=>//]; - exists c, c1, c2; rewrite fnd_ins -H9 eq_refl; + case: ((H2 c c1 c2).2 H8)=>H9; [apply: Or42 | apply: Or44]; + do ![split=>//]; exists c, c1, c2; rewrite fnd_ins -H9 eq_refl; by do !split=>//; apply: reflC. - case: (H4.2 b' d d1 d2 L2 H8)=>[[R1][R2]|[R]|[R1][R2]|[R]]; - [move=>[[f][f1][f2][F1][F2][F3] F4]|idtac|move=>[[f][f1][f2][F1][F2][F3] F4]| idtac]; - move=>[e][e1][e2][Q1][Q2][Q3] Q4; [idtac | apply: Or42 | idtac | apply: Or44]; + [move=>[[f][f1][f2][F1][F2][F3] F4]|idtac| + move=>[[f][f1][f2][F1][F2][F3] F4]| idtac]; + move=>[e][e1][e2][Q1][Q2][Q3] Q4; + [idtac | apply: Or42 | idtac | apply: Or44]; try by [do !split=>//; exists e, e1, e2; rewrite fnd_ins; case: eqP Q1=>[[->->]|]; first by rewrite F]; rewrite E inE in F1; case/orP: F1=>F1; @@ -1236,8 +1430,10 @@ apply: IH=>//=; first by rewrite ffunE eq_refl E. case/eqP: F1 F4=><-<-<- F4; by rewrite F3 F2 ?R1 ?R2 eq_refl. case: (H4.2 a d d1 d2 H6 H8)=>[[R1][R2]|[R]|[R1][R2]|[R]]; - [move=>[[f][f1][f2][F1][F2][F3] F4]|idtac|move=>[[f][f1][f2][F1][F2][F3] F4]| idtac]; - move=>[e][e1][e2][Q1][Q2][Q3] Q4; [idtac | apply: Or42 | idtac | apply: Or44]; + [move=>[[f][f1][f2][F1][F2][F3] F4]|idtac| + move=>[[f][f1][f2][F1][F2][F3] F4]| idtac]; + move=>[e][e1][e2][Q1][Q2][Q3] Q4; + [idtac | apply: Or42 | idtac | apply: Or44]; try by [do !split=>//; exists e, e1, e2; rewrite fnd_ins; case: eqP Q1=>[[->->]|]; first by rewrite F]; rewrite E inE in F1; case/orP: F1=>F1; @@ -1312,39 +1508,41 @@ exists f, f1, f2; do !split=>//. by apply: (transC (y:= const e))=>//; apply: symC. Qed. -Module Dummy3. End Dummy3. - (* Lemmas about propagate *) Lemma propagatePE (D : data) : - propagate_inv D -> propagate_inv (propagate D) /\ - pending (propagate D) = [::] /\ - CRel D <~> CRel (propagate D). + propagate_inv D -> + propagate_inv (propagate D) /\ + pending (propagate D) = [::] /\ + CRel D <~> CRel (propagate D). Proof. move: D. -pose ty x0 y0 := [Pred x y | x = const x0 /\ y = const y0]. +pose ty x0 y0 := Pred1 (@const s x0, @const s y0). have L: forall D a b, - closure (rep2rel D +p ty a b) <~> - closure (rep2rel D +p ty (rep D a) (rep D b)). -- move=>D a b [x y]; split=>T; rewrite toPredE -clos_idemp; apply: sub_closI T=>{x y} [[x y]]; + closure (rep2rel D \+p ty a b) <~> + closure (rep2rel D \+p ty (rep D a) (rep D b)). +- move=>D a b [x y]; split=>T; rewrite toPredE -clos_idemp; + apply: sub_closI T=>{x y} [[x y]]; (case; first by move=>H; apply: closI; left); case=>->->; - [rewrite toPredE const_rep symR const_rep symR | rewrite toPredE -const_rep symR -const_rep symR]; + [rewrite const_rep symR const_rep symR | + rewrite -const_rep symR -const_rep symR]; by apply closI; right. pose inv D := propagate_inv D. have L1: forall D pend_eq p' a b a' b' D', - pending D = pend_eq :: p' -> pend2eq pend_eq = simp_eq a b -> - rep D a = a' -> rep D b = b' -> Data (rep D) (class D) (use D) (lookup D) p' = D' -> - (a' == b') -> - (inv D' -> inv (propagate D') /\ pending (propagate D') = [::] /\ - CRel D' <~> CRel (propagate D')) -> - inv D -> inv (propagate D') /\ pending (propagate D') = [::] /\ - CRel D <~> CRel (propagate D'). + pending D = pend_eq :: p' -> pend2eq pend_eq = simp_eq a b -> + rep D a = a' -> rep D b = b' -> + Data (rep D) (class D) (use D) (lookup D) p' = D' -> + (a' == b') -> + (inv D' -> inv (propagate D') /\ pending (propagate D') = [::] /\ + CRel D' <~> CRel (propagate D')) -> + inv D -> inv (propagate D') /\ pending (propagate D') = [::] /\ + CRel D <~> CRel (propagate D'). - move=>D pend_eq p' a b a' b' D' H1 H H2 H3 H4 H5 IH [H6][H7][H8][H9] H10. have L2: forall a1 b1, similar D a1 b1 -> similar D' a1 b1. - - move=>a1 b1; rewrite /similar -{2}H4 H1 /= H /= -(pull (ty a b)) => T. - rewrite -clos_idemp; move: T; apply: sub_closI=>{a1 b1} [[x y]]; rewrite -H4. - case=>[[->->]|T]; last by apply: closI. - rewrite toPredE const_rep symR const_rep symR /=. + - move=>a1 b1; rewrite /similar -{2}H4 H1 /= H /= -(rpull (ty a b)) => T. + rewrite -clos_idemp; move: T; apply: sub_closI=>{a1 b1} [[x y]]; + rewrite -H4; case=>[[->->]|T]; last by apply: closI. + rewrite const_rep symR const_rep symR /=. by rewrite H2 H3 (eqP H5); apply: reflC. suff T: CRel D <~> CRel D'. - rewrite T; apply: IH; rewrite -H4; do 4!split=>//. @@ -1357,33 +1555,36 @@ have L1: forall D pend_eq p' a b a' b' D', split; [exists d, d1, d2 | exists e, e1, e2]; by do !split=>//; rewrite H4; apply: L2. rewrite /CRel H1 /= -{-1}H4 /=. - rewrite -!(pull (lookup_rel _)) -!(pull (eqs2rel (map pend2eq _))). + rewrite -!(rpull (lookup_rel _)) -!(rpull (eqs2rel (map pend2eq _))). do 2![apply: closKR]; rewrite H /= L -H4. symmetry; rewrite closAK=>[[x y]]; case=>->->. by rewrite H2 H3 (eqP H5); apply: reflC. have L2: forall D pend_eq p' a b a' b' D' D'', - pending D = pend_eq :: p' -> pend2eq pend_eq = simp_eq a b -> - rep D a = a' -> rep D b = b' -> - Data (rep D) (class D) (use D) (lookup D) p' = D' -> a' != b' -> - join_class D' a' b' = D'' -> - (inv (join_use D'' a' b') -> inv (propagate (join_use D'' a' b')) /\ - pending (propagate (join_use D'' a' b')) = [::] /\ - CRel (join_use D'' a' b') <~> CRel (propagate (join_use D'' a' b'))) -> - inv D -> inv (propagate (join_use D'' a' b')) /\ - pending (propagate (join_use D'' a' b')) = [::] /\ - CRel D <~> CRel (propagate (join_use D'' a' b')). -- move=>D pend_eq p' a b a' b' D' D'' H1 H H2 H3 H4 H5 H6 IH [H7][H8][H9][H10] H11. - have T: CRel D <~> closure (CRel D' +p [Pred x y | x = const a' /\ y = const b']). + pending D = pend_eq :: p' -> pend2eq pend_eq = simp_eq a b -> + rep D a = a' -> rep D b = b' -> + Data (rep D) (class D) (use D) (lookup D) p' = D' -> a' != b' -> + join_class D' a' b' = D'' -> + (inv (join_use D'' a' b') -> inv (propagate (join_use D'' a' b')) /\ + pending (propagate (join_use D'' a' b')) = [::] /\ + CRel (join_use D'' a' b') <~> CRel (propagate (join_use D'' a' b'))) -> + inv D -> inv (propagate (join_use D'' a' b')) /\ + pending (propagate (join_use D'' a' b')) = [::] /\ + CRel D <~> CRel (propagate (join_use D'' a' b')). +- move=>D pend_eq p' a b a' b' D' D'' H1 H H2 H3 H4 H5 H6 IH + [H7][H8][H9][H10] H11. + have T: CRel D <~> closure (CRel D' \+p + Pred1 (const a', const b')). - rewrite /CRel H1 -{2 3}H4 /= clos_clos !orrA H /=. - rewrite -!(pull (lookup_rel _)) -!(pull (eqs2rel (map pend2eq _))). + rewrite -!(rpull (lookup_rel _)) -!(rpull (eqs2rel (map pend2eq _))). by do 2![apply: closKR]; rewrite L H2 H3 -H4 /= /rep2rel. have S2: forall a1 b1, similar D a1 b1 -> similar1 D' a' b' a1 b1. - move=>a1 b1. - rewrite /similar /similar1 H1 -H4 /= H /= (pull (ty a b)) (pull (ty a' b')) -!orrA => Q4. + rewrite /similar /similar1 H1 -H4 /= H /=. + rewrite (rpull (ty a b)) (rpull (ty a' b')) -!orrA => Q4. rewrite -clos_idemp. move: Q4; apply: sub_closI=>[[x y]]. case; first by move=>Q4; apply: closI; left. - case=>->->; rewrite toPredE !orrA const_rep symR const_rep symR /= H2 H3. + case=>->->; rewrite !orrA const_rep symR const_rep symR /= H2 H3. by apply closI; do 2!right. have [L2' L3']: a' \in reps D' /\ b' \in reps D' by rewrite -H2 -H3 -H4 !rep_in_reps. @@ -1427,27 +1628,29 @@ have L2: forall D pend_eq p' a b a' b' D' D'', have L13'' : use_lookup_inv (join_use D'' a' b'). - have U: forall D, use (join_use D a' b') a' = [::]. - rewrite /join_use=>D1; move E: (use D1 _)=>x. - elim: x D1 E=>[|[[f f1] f2] s IH1] D1 E //=. - by case F: (fnd _ _)=>[[[d d1] d2]|]; apply: IH1; rewrite /= !ffunE eq_refl E. + elim: x D1 E=>[|[[f f1] f2] ss IH1] D1 E //=. + case F: (fnd _ _)=>[[[d d1] d2]|]; apply: IH1; + by rewrite /= !ffunE eq_refl E. move=>a1 c c1 c2 H12 H13. case: (T9.2 a1 c c1 c2 H12 H13); first 1 last; try by [case=><- [d][d1][d2][Q1][Q2][Q3] Q4; exists d, d1, d2]; by case=>[_][_][[d][d1][d2]][]; rewrite U. rewrite -H6 /= join_classE // T H6=>->. by apply: IH. -apply/(@propagate_ind (fun d d' => inv d -> inv d' /\ pending d' = [::] /\ CRel d <~> CRel d')); first by []. +apply/(@propagate_ind s (fun d d' => inv d -> inv d' /\ + pending d' = [::] /\ CRel d <~> CRel d'))=>//. - by move=>D e p' H1 a' H2 b' H3 D'; apply: L1 H1 _ H2 H3; case: e=>// [[[c c1]] c2] [[d d1] d2]. -move=>D e p' H1 a' H2 b' H3 D' E [] // H _ D''; apply: L2 H1 _ H2 H3 E (negbT H). +move=>D e p' H1 a' H2 b' H3 D' E [] // H _ D''; +apply: L2 H1 _ H2 H3 E (negbT H). by case: e=>// [[[c c1]] c2] [[d d1] d2]. Qed. -Module Dummy4. End Dummy4. - (* Lemmas about interaction of propagate with pending and closure *) Lemma propagate_pendP d eq : propagate_inv d -> - propagate_inv (Data (rep d) (class d) (use d) (lookup d) (eq :: pending d)). + @propagate_inv s (Data (rep d) (class d) (use d) + (lookup d) (eq :: pending d)). Proof. move=>H; set d' := Data _ _ _ _ _. have L: forall a b, similar d a b -> similar d' a b. @@ -1466,19 +1669,19 @@ Lemma propagate_clos_pendP d c c1 c2 e e1 e2 : fnd (rep d c1, rep d c2) (lookup d) = Some (e, e1, e2) -> CRel (Data (rep d) (class d) (use d) (lookup d) (comp_pend (c, c1, c2) (e, e1, e2) :: pending d)) - <~> closure (CRel d +p - [Pred x y | x = const c /\ y = app (const c1) (const c2)]). + <~> closure (CRel d \+p + Pred1 (const c, app (const c1) (const c2))). Proof. move=>PI H. have [R1 R2]: rep d e1 = rep d c1 /\ rep d e2 = rep d c2. -- by move: PI=>[_][_][L1] _; apply: (L1 _ _ e). +- by move: PI=>[_][_][L1] _; apply: L1 H; apply: rep_in_reps. rewrite /CRel /= /eq2rel /=. -rewrite clos_clos -!(pull (eqs2rel _)) !orrA; apply: closKR. +rewrite clos_clos -!(rpull (eqs2rel _)) !orrA; apply: closKR. move=>[z y]; split=>O; rewrite toPredE -clos_idemp; move: O; apply: sub_closI; move=>{z y} [z y]. - case=>[O|]; first by apply: closI; left. case=>[O|]; first by apply: closI; right; left. - case=>->->; rewrite toPredE symR const_rep symR. + case=>->->; rewrite symR const_rep symR. apply: (transC (y := app (const (rep d e1)) (const (rep d e2)))). - by rewrite R1 R2 -app_rep; apply closI; right; right. apply closI; right; left; rewrite R1 R2. @@ -1491,8 +1694,6 @@ exists (rep d c1), (rep d c2), e, e1, e2. by rewrite !rep_in_reps. Qed. -Module DummyT. End DummyT. - Section NoPend. Variables (d : data) (c c1 c2 : symb). Hypotheses (PI : propagate_inv d) @@ -1504,23 +1705,25 @@ Notation u2' := [ffun z => if z == rep d c2 then (c, c1, c2) :: u1' z Lemma propagate_nopendP : propagate_inv (Data (rep d) (class d) u2' - (ins (rep d c1, rep d c2) (c, c1, c2) (lookup d)) (pending d)). + (ins (rep d c1, rep d c2) (c, c1, c2) + (lookup d)) (pending d)). Proof. have E : forall (a c c1 c2 : symb) l, - (c, c1, c2) :: (if a == rep d c1 then (c, c1, c2) :: l else l) =i (c, c1, c2) :: l. + (c, c1, c2) :: (if a == rep d c1 then (c, c1, c2) :: l else l) =i + (c, c1, c2) :: l. - move=>a f f1 f2 l z; rewrite !inE; case: ifP=>H //. by rewrite inE orbA orbb. move: PI=>[R1][U1][L1][UL1] LU1; do 2!split=>//. - move=>a e e1 e2; rewrite /reps /= !ffunE /= => T1. case: ifP=>E1. - rewrite E (eqP E1) inE. - by case/orP=>[|O]; [case/eqP=>_->->; right | apply: U1 O]. + by case/orP=>[/eqP [_ ->->]|O]; [right|apply/U1/O/rep_in_reps]. case: ifP=>E2; last by apply: U1 T1. rewrite (eqP E2) inE. - by case/orP=>[|O]; [case/eqP=>_->->; left | apply: U1 O]. + by case/orP=>[/eqP [_ ->->]|O]; [left|apply/U1/O/rep_in_reps]. split=>[a b e e1 e2|]. - rewrite /reps /= fnd_ins. - by case: ifP=>E1; [case/eqP: E1=>->-> T1 T2 [_->->] | apply: L1]. + by case: ifP=>E1; [case/eqP: E1=>->-> T1 T2 [_->->]|apply: L1]. split=>[a e e1 e2|]. - rewrite /reps /= !ffunE /= => T1. case: ifP=>E1. @@ -1554,36 +1757,37 @@ Qed. Lemma propagate_clos_nopendP : CRel (Data (rep d) (class d) u2' - (ins (rep d c1, rep d c2) (c, c1, c2) (lookup d)) (pending d)) <~> - closure (CRel d +p [Pred x y | x = const c /\ y = app (const c1) (const c2)]). + (ins (rep d c1, rep d c2) (c, c1, c2) (lookup d)) + (pending d)) <~> + closure (CRel d \+p + Pred1 (const c, app (const c1) (const c2))). Proof. rewrite /CRel clos_clos orrA orrAC -!orrA; apply: closRK; rewrite !orrA. move=>[xx yy]; split=>O; rewrite toPredE -clos_idemp; move: O; apply: sub_closI; move=>{xx yy} [xx yy]. - case=>[O|]; first by apply: closI; left. - rewrite /lookup_rel =>[[a]][b][e][e1][e2][-> R1 R2 T1 ->]. + rewrite /lookup_rel =>[[a]][b][e][e1][e2][/= -> R1 R2 T1 ->]. rewrite fnd_ins in T1; case: ifP T1=>[|O T1]. - - case/eqP=>->->[->->->]; rewrite toPredE symR -app_rep -const_rep. + - case/eqP=>->->[->->->]; rewrite symR -app_rep -const_rep. by apply closI; right; right. by apply: closI; right; left; exists a, b, e, e1, e2. case=>[O|]; first by apply: closI; left. case=>[O|[->->]]. -- apply: closI; right; move: O=>[a][b][e][e1][e2][-> R1 R2 O ->]. +- apply: closI; right; move: O=>[a][b][e][e1][e2][/= -> R1 R2 O ->]. exists a, b, e, e1, e2; do !split=>//. by rewrite fnd_ins; case: eqP O=>//; case=>->->; rewrite Ev. -rewrite toPredE app_rep const_rep symR; apply closI; right. +rewrite app_rep const_rep symR; apply closI; right. exists (rep d c1), (rep d c2), c, c1, c2. by rewrite !rep_in_reps fnd_ins eq_refl. Qed. End NoPend. -Module DummyQ. End DummyQ. - (* Lemma about normalization *) -Lemma norm_const (D : data) (t : exp) (s : symb) : - norm D t = const s -> s \in reps D. +Lemma norm_const (D : data) (t : exp) (ss : symb) : + norm D t = const ss -> + ss \in reps D. Proof. elim: t=>[t|t1 IH1 t2 IH2] /=; first by case=><-. do 2![case: (norm D _)=>//] => q2 q1. @@ -1601,7 +1805,8 @@ by rewrite (norm_const E1) (norm_const E2). Qed. Lemma norm_rel (D : data) (t : exp) : - rep_idemp D -> (t, norm D t) \In CRel D. + rep_idemp D -> + (t, norm D t) \In CRel D. Proof. move=>H; elim: t=>[t|t1 IH1 t2 IH2] /=; first by apply: clos_rep. case E1: (norm D t1) IH1=>[q1|] IH1; last by apply: monoC. @@ -1614,7 +1819,8 @@ by rewrite (norm_const E1) (norm_const E2). Qed. Lemma normP (D : data) (x y : exp) : - rep_idemp D -> pending D = [::] -> + rep_idemp D -> + pending D = [::] -> reflect ((x, y) \In CRel D) (norm D x == norm D y). Proof. move=> H1 H2; case: eqP=>H; constructor. @@ -1622,12 +1828,18 @@ move=> H1 H2; case: eqP=>H; constructor. apply: (transC (y := norm D y)); last by apply: symC; apply: norm_rel. by rewrite H; apply: reflC. rewrite /CRel H2 /= orr0 => H3; case: H. -suff: (x, y) \In [Pred e1 e2 | norm D e1 = norm D e2] by []. +suff: (x, y) \In [Pred e | norm D e.1 = norm D e.2] by []. apply: {x y} H3; last first. - split; last by move=>????; rewrite !InE /= => -> ->. by do !split=>//; move=> x y z; rewrite !InE /= => ->. move=>[x y]; case. - by case/imageP=>a _ [->] ->; rewrite /= H1. -move=>[a][b][c][c1][c2][-> Q1 Q2 Q3 ->] /=; do 2!rewrite reps_rep //. +move=>[a][b][c][c1][c2][/= -> Q1 Q2 Q3 ->] /=; do 2!rewrite reps_rep //. by rewrite Q3 H1. Qed. + +End MathModelProofs. + +(* empty congruence only relates constant symbols to themselves *) +Definition empty_cong s := closure (graph (@const s)). +Arguments empty_cong : clear implicits. diff --git a/examples/congprog.v b/examples/congprog.v new file mode 100644 index 0000000..31f0328 --- /dev/null +++ b/examples/congprog.v @@ -0,0 +1,637 @@ +(* +Copyright 2009 IMDEA Software Institute +Licensed under the Apache License, Version 2.0 (the "License"); +you may not use this file except in compliance with the License. +You may obtain a copy of the License at + http://www.apache.org/licenses/LICENSE-2.0 +Unless required by applicable law or agreed to in writing, software +distributed under the License is distributed on an "AS IS" BASIS, +WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. +See the License for the specific language governing permissions and +limitations under the License. +*) + +From HB Require Import structures. +From mathcomp Require Import ssreflect ssrbool ssrnat eqtype ssrfun. +From mathcomp Require Import div finset seq fintype finfun choice. +From pcm Require Import axioms prelude ordtype finmap pred pcm. +From pcm Require Import unionmap heap autopcm automap. +From htt Require Import options model heapauto llist array. +From htt Require Import kvmaps hashtab congmath. + +(***********************************) +(* Congruence closure verification *) +(***********************************) + +(* Spec, code and proofs for the verification of the *) +(* Barcelogic Congruence Closure algorithm. *) +(* Described in the POPL10 paper: *) +(* Structure the verification of heap-manipulating programs *) +(* by Nanevski, Vafeiadis and Berdine *) + +(* This file contains the verification of the stateful programs. *) +(* The associated math properties used in the verification are *) +(* proved in the accompanying file congmath.v *) + +Notation finE := finset.inE. + +(*************) +(* Signature *) +(*************) + +Module Type CongrSig. +(* abstract type for collection of root pointers *) +(* seq const is the list of constants over which *) +(* structure computes *) +(* this is a global argument, hence exposed by tp *) +Parameter root : seq constant -> Set. + +(* abstract predicate tying roots, (congruence) relation, heap *) +Parameter shape : forall {s}, root s -> rel_exp s -> Pred heap. + +(* initialize empty congruence structure; return roots *) +Parameter init : forall {s}, + STsep (emp, [vfun rt m => m \In shape rt (empty_cong s)]). + +(* merge new equation into the structure rooted by rt *) +Parameter merge : forall {s} rt (e : Eq s), + STsep {R} (fun i => i \In shape rt R, + [vfun (_ : unit) m => + m \In shape rt (closure (R \+p eq2rel e))]). + +(* check if two expressions are congruent in the structure rooted by rt *) +Parameter check : forall {s} rt (t1 t2 : exp s), + STsep {R} (fun i => i \In shape rt R, + [vfun (b : bool) m => m \In shape rt R /\ + (b <-> (t1, t2) \In R)]). +End CongrSig. + +(******************) +(* Implementation *) +(******************) + +(* faithful to Barcelogic's implementation *) + +Module Congr : CongrSig. +Section Congr. +Variable s : seq constant. +Notation symb := (symb s). + +(* the lookup table is represented as a hash table with 10 buckets *) +Local Definition K : Set := symb * symb. +Local Definition V : Set := symb * symb * symb. + +Definition hash (n : nat) (k : K) := index k (enum (@predT K)) %% n. + +Lemma hash_ord n k : + 0 < n -> + hash n k < n. +Proof. exact: ltn_pmod. Qed. + +Definition hash10 k : 'I_10 := Ordinal (@hash_ord 10 k erefl). + +(* the tables relating arrays with the content of the lists *) +(* ctab is for class lists, utab is for use lists *) + +Definition llist_root (T : Set) : Set := ptr. + +Definition ctab := @table symb ptr (seq symb) (@lseq symb). +Definition utab := + @table symb ptr (seq (symb * symb * symb)) (@lseq (symb * symb * symb)). + +Definition n := #|{: symb}|. + +(* roots for the structure *) +(* all components are isomorphic to pointers *) +Inductive ptrs : Set := + Ptrs of + (* array of representatives; for each c *) + (* returns symbol that represents c's class *) + {array symb -> symb} & + (* class list; for each c *) + (* singly-linked list storing whole class of c *) + {array symb -> llist_root symb} & + (* use list; internal structure *) + (* see the paper for description *) + {array symb -> llist_root (symb*symb*symb)} & + (* hash table; for each pair of representatives *) + (* stores equations; see paper for description *) + htab V hash10 & + (* list of pending equations *) + ptr. + +(* renaming type to satisfy the signature check *) +Definition root := ptrs. + +Section ShapePredicates. +Variable rt : ptrs. +Notation r := (let: Ptrs r _ _ _ _ := rt in r). +Notation clist := (let: Ptrs _ clist _ _ _ := rt in clist). +Notation ulist := (let: Ptrs _ _ ulist _ _ := rt in ulist). +Notation htb := (let: Ptrs _ _ _ htb _ := rt in htb). +Notation p := (let: Ptrs _ _ _ _ p := rt in p). + +(* Data structure's layout in the heap *) + +Definition ashape D := + [Pred h | let: (d, ct, ut) := D in + h \In + Array.shape r (rep d : {ffun symb -> symb}) # + (Array.shape clist ct # sepit setT (ctab ct (class d))) # + (Array.shape ulist ut # sepit setT (utab ut (use d))) # + kvm_shape htb (lookup d) # + [Pred h' | exists l, h' \In Pred1 (p :-> l) # lseq l (pending d)]]. + +Definition bshape d := + [Pred h | class_inv d /\ exists ct ut, h \In ashape (d, ct, ut)]. + +Definition shape (R : rel_exp s) : Pred heap := + [Pred h | exists d, h \In bshape d /\ propagate_inv d /\ + pending d = [::] /\ CRel d <~> R]. + +End ShapePredicates. + +(* Initialization procedure to generate *) +(* the empty structure and its root pointers *) + +Definition iT (clist : {array symb -> llist_root symb}): Type := forall k, + STsep (fun i => k <= n /\ exists f, i \In Array.shape clist f # + sepit [set c | indx c < k] (ctab f [ffun c => [:: c]]), + [vfun (_ : unit) m => exists f, m \In Array.shape clist f # + sepit setT (ctab f [ffun c => [:: c]])]). + +Program Definition init : + STsep (emp, [vfun rt m => m \In shape rt (empty_cong s)]) := + Do (rx <-- Array.newf [ffun x : symb => x]; + cl <-- Array.new null; + ffix (fun (loop : iT cl) k => + Do (if decP (b:= k < n) idP is left pf then + x <-- allocb (ith k pf) 2; + x.+1 ::= null;; + Array.write cl (ith k pf) x;; + loop k.+1 + else ret tt)) 0;; + ul <-- Array.new null; + htb <-- kvm_new (htab V hash10); + p <-- alloc null; + ret (Ptrs rx cl ul htb p)). +Next Obligation. +move=>_ cl loop k [i][pf][/= f][hc][hct][->{i} Hc Hct]. +case: decP=>[{}pf|] /=; last first. +- case: (ltngtP k n) pf=>// Ekn _ _; step=>_. + exists f, hc, hct; split=>//. + apply: tableP2 Hct=>// x; rewrite !finE Ekn. + by rewrite /n cardE index_mem /index mem_enum. +step=>x; step; apply: [stepX f]@hc=>//= [[m]] Em. +apply: [gE]=>//=; split=>//. +eexists [ffun z => if z == ith k pf then x else f z]. +rewrite (_ : _ \+ _ = m \+ (x :-> ith k pf \+ + x.+1 :-> null \+ hct)); last by heap_congr. +hhauto; rewrite (sepitS (ith k pf)) finE /= indx_ith ltnSn. +rewrite /ctab/table !ffunE eqxx; hhauto. +apply: tableP2 Hct=>// a. +- by rewrite !finE ltnS indx_injE; case: ltngtP. +by rewrite !finE !ffunE indx_injE; case: eqP=>// ->; rewrite ltnn. +Qed. +Next Obligation. +case=>_ ->; apply: [stepE]=>//= rx hr Er; apply: [stepU]=>//= cl hc Ec. +apply: [stepX]@hc=>//=. +- split=>//; exists [ffun x => null], hc, Unit; rewrite unitR. + by split=>//; rewrite (_ : [set c | indx c < 0] = set0) // sepit0. +case=>_ [f][hc'][hrest][-> Hc' Hrest]. +apply: [stepU]=>//= ul hu Ehu; apply: [stepU]=>//= htb ht Ht. +set d := Data [ffun x => x] [ffun c => [:: c]] [ffun c => [::]] (@nil K V) [::]. +step=>px; step; exists d; split; last by case: (initP s). +split=>[a b|/=]; first by rewrite !ffunE !inE. +exists f, [ffun s => null]. +rewrite (_ : px :-> null \+ _ = hr \+ ((hc' \+ hrest) \+ (hu \+ Unit \+ + (ht \+ (px :-> null \+ Unit))))); last by rewrite unitR; heap_congr. +hhauto; rewrite /sepit sepitseq_emp // => k. +by rewrite /utab/table !ffunE; split=>//; case=>_ ->. +Qed. + +(* Method implementations use various input root pointers. *) +(* These are given explicit names as projections from rt. *) + +Section Internal. +Variable rt : ptrs. +Notation ashape' := (ashape rt). +Notation bshape' := (bshape rt). +Notation shape' := (shape rt). + +Notation r := (let: Ptrs r _ _ _ _ := rt in r). +Notation clist := (let: Ptrs _ clist _ _ _ := rt in clist). +Notation ulist := (let: Ptrs _ _ ulist _ _ := rt in ulist). +Notation htb := (let: Ptrs _ _ _ htb _ := rt in htb). +Notation p := (let: Ptrs _ _ _ _ p := rt in p). + +Definition cT (a' b' : symb) : Type := + forall x : unit, STsep {D} + (fun i => i \In ashape' D /\ a' != b', + [vfun (_ : unit) m => exists ct, exists ut, + let: (d, _, _) := D in + m \In ashape' + (Data [ffun s => if s \in class d a' then b' + else rep d s] + [ffun s => if s == a' then [::] else + if s == b' then rev (class d a') ++ class d b' + else class d s] + (use d) (lookup d) (pending d), ct, ut)]). + +Program Definition join_hclass (a' b' : symb) : + STsep {d} (fun i => i \In bshape' d /\ a' != b', + [vfun (_ : unit) m => m \In bshape' (join_class d a' b')]) := + Do (ffix (fun (loop : cT a' b') (xx : unit) => + Do (a <-- Array.read clist a'; + b <-- Array.read clist b'; + if a == null then ret tt + else + s <-- !a; + next <-- !a.+1; + a.+1 ::= b;; + Array.write clist b' a;; + Array.write clist a' next;; + Array.write r s b';; + loop tt)) tt). +Next Obligation. +move=>a' b' loop [][[[/= d ct] ut]][i][H N] /=. +case: H=>/= rh [_][->{i} Rh][_][h][->][th][ctx][->] Th Ctx H. +rewrite (sepitT1 a') in Ctx; case: Ctx=>cta [w][->{ctx}Cta]. +rewrite (sepitS b') !finE eq_sym {1}N /=. +case=>ctb [ctx][->{w}Ctb Ctx]; rewrite /ctab/table in Cta Ctb. +apply: [stepX ct, th]@th=>//= _ _ [->->]. +apply: [stepX ct, th]@th=>//= _ _ [->->]. +apply: vrfV=>V; case: (ct a' =P null) Cta=>[/[dup] Ea' ->|/eqP Na']. +- case/(lseq_null (validX V))=>/= ->->{cta V}; step=>{}V; exists ct, ut. + rewrite (_ : rh \+ _ = rh \+ (th \+ (Unit \+ + (ctb \+ ctx)) \+ h)); last by heap_congr. + rewrite -fin_eta; hhauto; rewrite (sepitT1 a'); hhauto. + - by rewrite /ctab/table ffunE eqxx /= /= Ea'. + rewrite (sepitS b') !finE eq_sym N; hhauto. + - by rewrite /ctab/table /= ffunE eqxx eq_sym (negbTE N). + apply: tableP Ctx=>// x; rewrite !finE andbT ffunE. + by case/andP=>/negbTE -> /negbTE ->. +case/(lseq_pos Na')=>a {V} [next][cta'][Eca ->{cta}Cta']. +do 3!step; apply: [stepX ct]@th=>//= [_] {th Th} th1 Th1. +set ct1 := (finfun _) in Th1. +apply: [stepX ct1]@th1=>//= [_] {th1 Th1} th2 Th2. +set ct2 := (finfun _) in Th2. +apply: [stepX rep d]@rh=>//= _ {rh Rh} rh1 Rh1. +set r1 := (finfun _) in Rh1. +set cv2 := [ffun z => if z == a' then (behead (class d a')) + else if z == b' then a :: class d b' else class d z]. +apply: [gE (Data r1 cv2 (use d) (lookup d) (pending d), ct2, ut)]=>/=; +last by move=>?? []. +- rewrite (_ : rh1 \+ _ = rh1 \+ (th2 \+ (cta' \+ (ct a' :-> a \+ + ((ct a').+1 :-> ct b' \+ ctb) \+ ctx)) \+ h)); last by heap_congr. + hhauto; rewrite (sepitT1 a'); hhauto. + - by rewrite /ctab/table/ct2/cv2 !ffunE /= eqxx. + rewrite (sepitS b') !finE eq_sym N; hhauto. + - rewrite /ctab/table/ct2/cv2/ct1 !ffunE /= ffunE /= eqxx; + by rewrite eq_sym (negbTE N); hhauto. + apply: tableP Ctx=>x; rewrite /ct2/cv2/ct1 !ffunE /= ?ffunE /=; + by rewrite !finE andbT; case/andP=>/negbTE -> /negbTE ->. +case=>m [{Th2}ct2][ut2] /= Hm _; exists ct2, ut2. +congr (m \In ashape' (Data _ _ _ _ _, ct2, ut2)): Hm; apply/ffunP=>x. +- by rewrite !ffunE eqxx {2}Eca inE /=; case: (x =P a)=>//= _; rewrite if_same. +rewrite !ffunE !eqxx /= (eq_sym b') (negbTE N). +case: (x =P a')=>// _; case: (x =P b')=>// _. +by rewrite Eca rev_cons cat_rcons. +Qed. +Next Obligation. +move=>a' b' [d][i][[C]][/= ct][ut H] N. +apply: [gE (d, ct, ut)]=>//= [[m]][ct1][ut1] W _. +set d' := (Data _ _ _ _ _) in W; suff E : join_class d a' b' = d'. +- by split; [apply: join_class_classP|rewrite E; eauto]. +by congr Data; apply/ffunP=>x; rewrite !ffunE /= C. +Qed. + +Definition uT (a' b' : symb) := forall x : unit, + STsep {d} (fun i => exists don, a' != b' /\ + i \In bshape' (join_use' d a' b' don) /\ + use d a' = don ++ use (join_use' d a' b' don) a', + [vfun (_ : unit) m => m \In bshape' (join_use d a' b')]). + +Program Definition join_huse (a' b' : symb) : + STsep {d} (fun i => a' != b' /\ i \In bshape' d, + [vfun (_ : unit) m => m \In bshape' (join_use d a' b')]) := + Do (ffix (fun (loop : uT a' b') xx => + Do (a <-- Array.read ulist a'; + if a == null then ret tt + else + eq <-- !a; + next <-- !a.+1; + Array.write ulist a' next;; + c1 <-- Array.read r eq.1.2; + c2 <-- Array.read r eq.2; + v <-- kvm_lookup htb (c1, c2); + if v is Some d then + dealloc a;; + dealloc a.+1;; + p' <-- !p; + q <-- insert p' (comp_pend eq d); + p ::= q;; + loop tt + else + kvm_insert htb (c1, c2) eq;; + b <-- Array.read ulist b'; + a.+1 ::= b;; + Array.write ulist b' a;; + loop tt)) tt). +Next Obligation. +move=>a' b' loop [][/= d][i][don][N][H Eu]. +set d1 := join_use' d a' b' don in H Eu. +case: H=>C [/= ct][ut][rh][_][->{i} Rh][cth][_][-> Htc][_][_][->] +[ru][hu][-> Ut Hu][ht][_][-> Ht][p'][_][hp][->-> Hp]. +move: Hu; rewrite (sepitT1 a'); case=>ha' [w][->{hu} Ua]. +rewrite (sepitS b') !finE eq_sym {1}N; case=>hb' [h][->{w} Ub R]. +rewrite /utab/table/= in Ua Ub. +apply: [stepX ut, ru]@ru=>//= _ _ [->->]; apply: vrfV=>V. +case: (ut a' =P null) Ua=>[/[dup] Ea' ->|/eqP Na' {V}]. +- case/(lseq_null (validX V))=>/= {V} Eu1 ->{ha'}; step. + rewrite (_ : rh \+ _ = rh \+ (cth \+ (ru \+ (Unit \+ (hb' \+ h)) \+ + (ht \+ (p:-> p' \+ hp))))); last by heap_congr. + rewrite /join_use Eu Eu1 cats0 -/d1; split=>//=; exists ct, ut. + hhauto; rewrite (sepitT1 a'); hhauto; first by rewrite /utab/table Ea' Eu1. + by rewrite (sepitS b') !finE eq_sym N; hhauto. +case/(lseq_pos Na')=>[[[c1 c2 c3]]][nxt][hh][Eu1 ->{ha'} Hh]; step; step. +set c := (c1, c2, c3) in Eu1 *; apply: [stepX ut]@ru=>//= [[]] {Ut}ru Ut2. +set ut2 := (finfun _) in Ut2. +apply: [stepX rep d1, rh]@rh=>//= _ _ [->->]. +apply: [stepX rep d1, rh]@rh=>//= _ _ [->->]. +apply: [stepX lookup d1]@ht=>//= v {Ht}ht [Ht Eqv]. +rewrite Eu1 in Eu; set d2 := join_use' d a' b' (don ++ [:: c]). +set a1' := behead (use d1 a') in Eu. +have Eu2 : use d2 a' = a1'. +- by rewrite /d2 (join_useT (t:=behead (use d1 a'))) //; apply: join_use_useE. +have Ej2: join_use d2 a' b' = join_use d1 a' b'. +- rewrite /join_use/d2 Eu2. + by rewrite -!(join_useT (t:=[::])) ?cats0 -?catA //= ?Eu -?Eu1. +case: v Eqv=>[[[e1 e2 e3]]|] /= /esym Eqv. +- set e := (e1, e2, e3) in Eqv. + do 3!step; apply: [stepX pending d1]@hp=>//= q _ [r0][{Hp}hp [-> Hp]]. + step; apply: [gE d]=>[||??[]] //=. + exists (don ++ [:: c]); rewrite -/d2 Eu2 -catA; do 2!split=>//. + rewrite /bshape'/class_inv/ashape/d2 (join_useT (t:=a1')) //= Eqv -/d1 /=. + rewrite (_ : _ \+ _ = rh \+ (cth \+ ((ru \+ (hh \+ (hb' \+ h))) \+ (ht \+ + (p :-> q \+ (q :-> comp_pend c e \+ + (q.+1 :-> r0 \+ hp))))))); last by heap_congr. + case: Htc=>x [y][->] X1 X2; hhauto; [eauto|eauto|exact: Ut2|]. + rewrite (sepitT1 a'); hhauto. + - by rewrite /utab/table/ut2 !ffunE /= eqxx. + rewrite (sepitS b') !finE eq_sym N; hhauto. + - by rewrite /utab/table/ut2 !ffunE /= eq_sym (negbTE N). + apply: tableP R=>a /=; rewrite !finE andbT /ut2 ?ffunE /=; + by case/andP=>_ /negbTE ->. +apply: [stepX lookup d1]@ht=>//= [[{Ht}ht Ht]]. +apply: [stepX ut2, ru]@ru=>//= _ _ [->->]; step. +apply: [stepX ut2]@ru=>//= [[]] {Ut2}ru Ut3; set ut3 := (finfun _) in Ut3. +apply: [gE d]=>[||??[]] //=; exists (don ++ [:: c]). +rewrite -/d2 Eu2 -catA; do 2!split=>//. +rewrite /bshape'/class_inv/ashape/d2 (join_useT (t:=a1')) //= Eqv -/d1 /=. +rewrite (_ : _ \+ _ = rh \+ (cth \+ ((ru \+ (hh \+ ((ut a' :-> c \+ + ((ut a').+1 :-> ut2 b' \+ hb')) \+ h))) \+ (ht \+ (p :-> p' \+ hp))))); + last by heap_congr. +case: Htc=>x [y][->] X1 X2; hhauto; [eauto|eauto|exact: Ut3|]. +rewrite (sepitT1 a'); hhauto. +- by rewrite /utab/table/ut3 ffunE /= (negbTE N) /ut2 !ffunE /= eqxx. +rewrite (sepitS b') !finE eq_sym N /=; hhauto. +- by rewrite /utab/table/=/ut3 !ffunE /= eqxx eq_sym (negbTE N); hhauto. +apply: tableP R=>a /=; rewrite !finE andbT /ut3/ut2 !ffunE /= ?ffunE /=; +by case/andP=>/negbTE -> /negbTE ->. +Qed. +Next Obligation. +by move=>a' b' [d][i][N H]; apply: [gE d]=>[||??[]] //; exists [::]. +Qed. + +Definition pT : Type := forall x : unit, + STsep {d} (fun i => i \In bshape' d, + [vfun (_ : unit) m => m \In bshape' (propagate d)]). + +(* left/right symbol in pending equation *) +Definition pendL (e : pend s) := + let: (simp_pend a _ | comp_pend (a,_,_) _) := e in a. +Definition pendR (e : pend s) := + let: (simp_pend _ b | comp_pend _ (b,_,_)) := e in b. + +Program Definition hpropagate := + ffix (fun (loop : pT) x => + Do (q <-- @read ptr p; + if q == null then ret tt (* pending list is empty *) + else + eq <-- !q; + next <-- !q.+1; + p ::= (next : ptr);; + dealloc q;; + dealloc q.+1;; + a' <-- Array.read r (pendL eq); + b' <-- Array.read r (pendR eq); + if a' == b' then loop tt : ST _ + else + join_hclass a' b';; + join_huse a' b';; + loop tt)) tt. +Next Obligation. +move=>loop [[d]][i][C][/= ct][ut][hr][_][-> Hr][hc][_][-> Hc] +[hu][_][-> Hu][ht][_][-> Ht][q][_][hp][->] -> Hp. +step; case: (q =P null) Hp=>[->{q}|/eqP N] Hp. +- apply: vrfV=>V; case/(lseq_null (validX V)): Hp=>{V} Ep /= ->{hp}. + step; rewrite propagate_equation Ep; split=>//=. + by exists ct, ut; hhauto; rewrite Ep. +case/(lseq_pos N): Hp=>eq [next][hnext][Ep] ->{hp} Hp; do !step. +apply: [stepX rep d, hr]@hr=>//= a' _ [-> Ea']. +apply: [stepX rep d, hr]@hr=>//= b' _ [-> Eb']. +set d1 := Data (rep d) (class d) (use d) (lookup d) (behead (pending d)). +case: ifPn=>[E|E]. +- have T1 : propagate d = propagate d1. + - rewrite propagate_equation Ep. + by case: eq Ep Ea' Eb' E=>[a2 b2 _|[[a2 ??]][[b2 ??]] _] ->->->. + apply: [gE d1]=>[||??[] //] /=; last by rewrite T1. + by split=>//=; exists ct, ut; hhauto. +apply: [stepE d1]=>//=; first by do 2split=>//; exists ct, ut; hhauto. +set d2 := (join_class _ _ _); case=>m H; apply: [stepE d2]=>//= [[n]] Hn. +set d3 := (join_use _ _ _) in Hn. +suff -> : propagate d = propagate d3 by apply: [gE d3]. simpl in Hu. +rewrite /d3/d2 propagate_equation Ep; +by case: eq Ep Ea' Eb' E=>[a2 b2 _|[[a2 ??]][[b2 ??]] _] ->-> /negbTE ->. +Qed. + +Program Definition merge (e : Eq s) : + STsep {R} (fun i => i \In shape' R, + [vfun (_ : unit) m => + m \In shape' (closure (R \+p eq2rel e))]) := + match e with + | simp_eq a b => + Do (q <-- !p; + x <-- insert q (simp_pend a b); + p ::= x;; + hpropagate) + | comp_eq c c1 c2 => + Do (c1' <-- Array.read r c1; + c2' <-- Array.read r c2; + v <-- kvm_lookup htb (c1', c2'); + if v is Some b then + q <-- !p; + x <-- insert q (comp_pend (c, c1, c2) b); + p ::= x;; + hpropagate + else + kvm_insert htb (c1', c2') (c, c1, c2);; + u1 <-- Array.read ulist c1'; + x <-- insert u1 (c, c1, c2); + (* if c1' == c2' the equation will be added twice *) + (* but this is okay, so we need not check for equality *) + (* this will rarely occur in practice, because an equation *) + (* c = c1' c1' means that a function c1' is applied to itself *) + (* so by avoiding the check, we optimize for the common case *) + (* this complicates the proof somewhat, however *) + Array.write ulist c1' x;; + u2 <-- Array.read ulist c2'; + x <-- insert u2 (c, c1, c2); + Array.write ulist c2' x) + end. +Next Obligation. +move=>e a b [R][_][d][[C]][/= ct][ut][hr][_][-> Hr][hc][_][-> Hc][_][_] +[->][hu][hu'][-> Hu Hu'][ht][_][-> Ht][q][_][hp][->] -> Hp [PI][Ep Erel]. +step; apply: [stepX (pending d)]@hp=>//= x _ [r0][{Hp}hp][-> Hp]. +set d1:=Data (rep d) (class d) (use d) (lookup d) (simp_pend a b :: pending d). +step; apply: [gE d1]=>//=. +- rewrite Ep /= in Hp; case: Hp=>->{r0} ->{hp}. + rewrite (_ : _ \+ _ = hr \+ (hc \+ (hu \+ hu' \+ (ht \+ + (p :-> x \+ (x :-> simp_pend a b \+ (x.+1 :-> null \+ Unit))))))); + last by heap_congr. + by split=>//=; exists ct, ut; hhauto; rewrite Ep. +case=>m Hm _. +have L: propagate_inv d1 by apply: propagate_pendP PI. +case: (propagatePE L)=>L1 [L2 L3]; exists (propagate d1); do 3!split=>//. +by rewrite -L3 -Erel clos_clos /CRel -!orrA orrAC. +Qed. +Next Obligation. +move=>e c c1 c2 [R][_][d][[C]][/= ct][ut][hr][_][-> Hr][hc][_][-> H] +[_][_][->][hu][hu'][-> Hu Hu'][ht][_][-> Ht][q][_] +[hp][->] -> Hp [PI][Ep Erel]; set cx := (c, c1, c2). +apply: [stepX rep d, hr]@hr=>//= _ _ [->->]. +apply: [stepX rep d, hr]@hr=>//= _ _ [->->]. +apply: [stepX lookup d]@ht=>//= v {Ht}ht [Ht Ev]. +case: v Ev=>[[[e0 e1 e2]]|] Ev. +- set ex := (e0, e1, e2). + set d1 := Data (rep d) (class d) (use d) + (lookup d) (comp_pend cx ex :: pending d). + step; apply: [stepX pending d]@hp=>//= x _ [x1][{}hp][-> {}Hp]. + step; apply: [gE d1]=>//=. + - rewrite (_ : _ \+ _ = hr \+ (hc \+ (hu \+ hu' \+ (ht \+ (p :-> x \+ + (x :-> comp_pend cx ex \+ (x.+1 :-> x1 \+ hp))))))); + last by heap_congr. + by split=>//; exists ct, ut; hhauto. + case=>m Hm _. + have L : propagate_inv d1 by apply: propagate_pendP PI. + case: (propagatePE L)=>L1 [L2] L3; exists (propagate d1). + by rewrite -L3 -Erel propagate_clos_pendP. +apply: [stepX lookup d]@ht=>//= _ {Ht}ht Ht. +apply: [stepX ut, hu]@hu=>//= _ _ [->->]. +move: Hu'; rewrite (sepitT1 (rep d c1)). +case=>hu'' [hu2][->{hu'} Hu'' Hu']. +apply: [stepX use d (rep d c1)]@hu''=>//= x _ [r0][{Hu''}hu''][-> Hu'']. +apply: [stepX ut]@hu=>//= _ {Hu}hu Hu; set ut1 := (finfun _) in Hu. +apply: [stepX ut1, hu]@hu=>//= _ _ [->->]. +set u1' := [ffun z => if z == rep d c1 then cx :: use d z else use d z]. +set u' := [ffun z => if z == rep d c2 then cx :: u1' z else u1' z]. +set l' := (ins _ _ _) in Ht. +set d1 := Data (rep d) (class d) u' l' (pending d). +pose ut2 x' := [ffun z => if z == rep d c2 then x' else ut1 z]. +case E : (rep d c2 == rep d c1). +- apply: [stepX cx::use d (rep d c1)]@(x :-> _ \+ x.+1 :-> _ \+ hu'')=>//=. + - by exists r0, hu''; rewrite joinA (eqP E) /ut1 !ffunE /= eqxx. + move=>v {Hu''}_ [x'][_][->][r1][{}hu''][-> Hu'']. + apply: [gX ut1]@hu=>//= [[]] {Hu}hu Hu _. + rewrite (_ : _ \+ _ = hr \+ (hc \+ (hu \+ (v :-> cx \+ + (v.+1 :-> x' \+ (x' :-> cx \+ (x'.+1 :-> r1 \+ hu''))) \+ + hu2) \+ (ht \+ (p :-> q \+ hp))))); last by heap_congr. + exists d1; split=>//; last first. + - split; first by apply: propagate_nopendP. + by rewrite -Erel propagate_clos_nopendP. + split=>//=; exists ct, (ut2 v); hhauto. + rewrite (sepitT1 (rep d c1)) /utab/table/ut2 /= !ffunE eq_sym eqxx E /=. + hhauto; apply: tableP Hu'=>// a; rewrite !finE /u'/u1' !ffunE /= andbT; + by rewrite (eqP E)=>/negbTE ->. +move: Hu'; rewrite (sepitS (rep d c2)) !finE E /=. +case=>hu1 [hu3][->{hu2} Hu1' Hu2']. +apply: [stepX use d (rep d c2)]@hu1=>//=; first by rewrite /ut1 ffunE /= E. +move=>v _ [x'][{Hu1'}hu1][-> Hu1']. +apply: [gX ut1]@hu=>//= [[]] {Hu}hu Hu _. +rewrite (_ : _ \+ _ = hr \+ (hc \+ (hu \+ + (x :-> cx \+ (x.+1 :-> r0 \+ hu'') \+ + ((v :-> cx \+ (v.+1 :-> x' \+ hu1) \+ hu3))) \+ + ((ht \+ (p :-> q \+ hp)))))); last by heap_congr. +exists d1; split; last first. +- split; first by apply: propagate_nopendP. + split; first by rewrite /d1. + by rewrite -Erel; apply: propagate_clos_nopendP. +split=>//; exists ct, (ut2 v); hhauto. +rewrite (sepitT1 (rep d c1))/utab/table/ut2 !ffunE /= eqxx eq_sym E. +hhauto; rewrite (sepitS (rep d c2)) !finE !ffunE /= !eqxx E /=. +hhauto; apply: tableP Hu2'=>/= a; rewrite !finE !ffunE /= andbT; +by case/andP=>/negbTE -> /negbTE ->. +Qed. + +Definition nT : Type := + forall t, STsep {d} (fun i => i \In bshape' d, + [vfun y m => y = norm d t /\ m \In bshape' d]). + +Program Definition hnorm := + ffix (fun (hnorm : nT) (t : exp s) => + Do (match t with + | const a => + a' <-- Array.read r a; + ret (const a') + | app t1 t2 => + u1 <-- hnorm t1; + u2 <-- hnorm t2; + if (u1, u2) is (const w1, const w2) then + v <-- kvm_lookup htb (w1, w2); + if v is Some a then + a' <-- Array.read r a.1.1; + ret (const a') + else ret (app u1 u2) + else ret (app u1 u2) + end)). +Next Obligation. +move=>hnorm t [d][_][Ci][/= ct][ut][hr][hrest][-> Hr Hrest]. +case: t=>[a|t1 t2]. +- apply: [stepX rep d, hr]@hr=>//= _ _ [->->]. + by step; do 2!split=>//; exists ct, ut; hhauto. +apply: [stepE d]=>//=; first by split=>//; exists ct, ut; hhauto. +move=>u1 m [E1 H]; apply: [stepE d]=>//=. +move=>u2 {}m [E2 {}H] {Ci ct ut hr hrest Hr Hrest}. +case: H=>Ci [/= ct][ut][hr][w][->{m} Hr][hc][w'] +[->{w} Hc][hu][w][->{w'} Hu][ht][hrest][->{w} Ht Hrest]. +case: u1 E1=>[w1|??] /= E1; last first. +- by rewrite -E1 E2; step; do 2!split=>//; exists ct, ut; hhauto. +case: u2 E2=>[w2|??] /= E2; last first. +- by rewrite -E1 -E2; step; do 2!split=>//; exists ct, ut; hhauto. +rewrite -E1 -E2; apply: [stepX lookup d]@ht=>//= v {Ht}ht [Ht Ev]. +case: v Ev=>[[[a1 a2 a3]]|] <-; last first. +- by step; do 2!split=>//; exists ct, ut; hhauto. +apply: [stepX rep d, hr]@hr=>//= _ _ [->->]. +by step; do 2!split=>//; exists ct, ut; hhauto. +Qed. + +Program Definition check t1 t2 : + STsep {R} (fun i => i \In shape' R, + [vfun (b : bool) m => m \In shape' R /\ + (b <-> (t1, t2) \In R)]) := + Do (u1 <-- hnorm t1; + u2 <-- hnorm t2; + ret (u1 == u2)). +Next Obligation. +move=>t1 t2 [R][h][d][H][[RI X]][Ep PI]. +apply: [stepX d]@h=>//= _ {H}h [-> H]. +apply: [stepX d]@h=>//= _ {H}h [-> H]. +step; split; first by exists d. +by case: normP=>//; rewrite PI. +Qed. + +End Internal. +End Congr. +End Congr. + diff --git a/examples/counter.v b/examples/counter.v old mode 100755 new mode 100644 index fdb4081..aa80a10 --- a/examples/counter.v +++ b/examples/counter.v @@ -1,8 +1,21 @@ +(* +Copyright 2010 IMDEA Software Institute +Licensed under the Apache License, Version 2.0 (the "License"); +you may not use this file except in compliance with the License. +You may obtain a copy of the License at + http://www.apache.org/licenses/LICENSE-2.0 +Unless required by applicable law or agreed to in writing, software +distributed under the License is distributed on an "AS IS" BASIS, +WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. +See the License for the specific language governing permissions and +limitations under the License. +*) + From mathcomp Require Import ssreflect ssrbool ssrfun. From mathcomp Require Import ssrnat eqtype. From pcm Require Import options axioms prelude pred. From pcm Require Import pcm unionmap heap. -From htt Require Import options interlude model heapauto. +From htt Require Import options model heapauto. (* counter that hides local state with an abstract predicate *) @@ -10,8 +23,8 @@ Prenex Implicits exist. Record cnt : Type := Counter {inv : nat -> Pred heap; - method : {v : nat}, STsep (inv v, - [vfun y m => y = v /\ m \In inv v.+1])}. + method : STsep {v : nat} (inv v, + [vfun y m => y = v /\ m \In inv v.+1])}. Program Definition new : STsep (emp, [vfun y m => m \In inv y 0]) := Do (x <-- alloc 0; @@ -28,7 +41,7 @@ move=>[] _ /= ->. (* but it doesn't; we suspect because universe levels are off *) (* but error messages aren't helpful *) (* one can still finish the step manually *) -apply/vrf_bind/vrf_alloc=>x; rewrite unitR=>_. +apply/vrf_bnd/vrf_alloc=>x; rewrite unitR=>_. (* after which, automation proceeds to work *) by step. Qed. @@ -53,8 +66,9 @@ Definition interp (r : rep) : Pred heap := [Pred h | h = rptr r :-> rval r]. Record scnt : Type := SCount {sinv : nat -> rep; - smethod : {v : nat}, STsep (interp (sinv v), - [vfun y m => y = v /\ m \In interp (sinv v.+1)])}. + smethod : STsep {v : nat} + (interp (sinv v), + [vfun y m => y = v /\ m \In interp (sinv v.+1)])}. Program Definition snew : STsep (emp, [vfun y m => m \In interp (sinv y 0)]) := @@ -69,7 +83,7 @@ by do 3!step. Qed. Next Obligation. move=>[] _ /= ->. -apply/vrf_bind/vrf_alloc=>x; rewrite unitR=>_. +apply/vrf_bnd/vrf_alloc=>x; rewrite unitR=>_. by step. Qed. @@ -93,8 +107,9 @@ Definition pinterp (R : rep -> Prop) : Pred heap := Record pcnt : Type := PCount {pinv : nat -> rep -> Prop; - pmethod : {v : nat}, STsep (pinterp (pinv v), - [vfun y m => y = v /\ m \In pinterp (pinv v.+1)])}. + pmethod : STsep {v : nat} + (pinterp (pinv v), + [vfun y m => y = v /\ m \In pinterp (pinv v.+1)])}. Program Definition pnew : STsep (emp, [vfun y m => m \In pinterp (pinv y 0)]) := @@ -109,7 +124,7 @@ by do 3!step; move=>?; split=>// _ ->. Qed. Next Obligation. move=>[] _ /= ->. -apply/vrf_bind/vrf_alloc=>x; rewrite unitR=>_. +apply/vrf_bnd/vrf_alloc=>x; rewrite unitR=>_. by step=>_ /= _ ->. Qed. diff --git a/examples/cyclic.v b/examples/cyclic.v index 7b8501a..968ba47 100644 --- a/examples/cyclic.v +++ b/examples/cyclic.v @@ -1,13 +1,28 @@ +(* +Copyright 2021 IMDEA Software Institute +Licensed under the Apache License, Version 2.0 (the "License"); +you may not use this file except in compliance with the License. +You may obtain a copy of the License at + http://www.apache.org/licenses/LICENSE-2.0 +Unless required by applicable law or agreed to in writing, software +distributed under the License is distributed on an "AS IS" BASIS, +WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. +See the License for the specific language governing permissions and +limitations under the License. +*) + From Coq Require Import ssreflect ssrbool ssrfun. From mathcomp Require Import eqtype seq ssrnat. -From pcm Require Import options axioms pred. -From pcm Require Import pcm unionmap heap automap autopcm. -From htt Require Import options interlude model heapauto. +From pcm Require Import options axioms pred seqext. +From pcm Require Import pcm unionmap heap auto automap autopcm. +From htt Require Import options model heapauto. From htt Require Import llist. -(* a queue variant that has a fixed capacity and can overwrite data in a circular way *) +(* queue variant with fixed capacity *) +(* that overwrites data in a circular way *) -(* the structure is a pair of pointers, a mutable length and immutable capacity *) +(* the structure is a pair of pointers: *) +(* a mutable length and immutable capacity *) Record buffer (T : Type) : Type := Buf {active: ptr; inactive: ptr; len: ptr; capacity: nat}. @@ -29,15 +44,14 @@ Definition is_buffer (a i : ptr) (m n : nat) (xs : seq T) := n = size xs + size ys, m = size xs, ha \In lseg a i xs & - hi \In lseg i a ys] - ]. + hi \In lseg i a ys]]. (* the structure itself requires three extra memory cells *) Definition shape (b : buffer) (xs : seq T) := [Pred h | exists a i m h', - [/\ valid (active b :-> a \+ (inactive b :-> i \+ (len b :-> m \+ h'))), - h = active b :-> a \+ (inactive b :-> i \+ (len b :-> m \+ h')) & - h' \In is_buffer a i m (capacity b) xs]]. + [/\ valid (active b :-> a \+ (inactive b :-> i \+ (len b :-> m \+ h'))), + h = active b :-> a \+ (inactive b :-> i \+ (len b :-> m \+ h')) & + h' \In is_buffer a i m (capacity b) xs]]. (* main methods *) @@ -46,14 +60,14 @@ Definition shape (b : buffer) (xs : seq T) := (* loop invariant for allocating the inner structure *) Definition new_loopT (n : nat) (init : T) : Type := forall (pk : ptr * nat), - {q}, STsep (fun h => pk.2 < n /\ h \In lseg pk.1 q (nseq pk.2 init), + STsep {q} (fun h => pk.2 < n /\ h \In lseg pk.1 q (nseq pk.2 init), [vfun p' => lseg p' q (nseq n.-1 init)]). (* allocate the buffer with capacity n prefilled by init *) Program Definition new (n : nat) (init : T) : STsep (emp, [vfun v => shape v [::]]) := (* allocate the buffer in a loop *) - Do (let run := Fix (fun (go : new_loopT n init) '(r,k) => + Do (let run := ffix (fun (go : new_loopT n init) '(r,k) => Do (if k < n.-1 then p' <-- allocb r 2; p' ::= init;; @@ -67,7 +81,7 @@ Program Definition new (n : nat) (init : T) : (* allocate the remaining n-1 nodes by prepending to it *) q <-- run (p, 0); (* form the cycle *) - p.+ 1 ::= q;; + p.+1 ::= q;; (* allocate the structure *) m <-- alloc 0; pi <-- alloc q; @@ -93,23 +107,23 @@ Qed. (* now the outer program *) Next Obligation. (* simplify, match on n *) -move=>n init [] _ /= ->; case: ifP. +move=>n init [] _ /= ->; case: ifP=>[N0|_]. - (* 0 < n, allocate the initial node *) - move=>H0; step=>p; step; rewrite !unitR. + step=>p; step; rewrite !unitR. (* run the loop by framing on an empty heap *) apply: [stepU p]=>//= q h H. (* run the rest of the program *) step; step=>m; step=>pi; step=>pa; step=>V. (* the structure is well-formed if the buffer is *) - exists q, q, 0, (h \+ (p :-> init \+ p.+ 1 :-> q)); split=>//. + exists q, q, 0, (h \+ (p :-> init \+ p.+1 :-> q)); split=>//. (* most of the conditions hold trivially or by simple arithmetics *) - exists (nseq n init), Unit, (h \+ (p :-> init \+ p.+ 1 :-> q)); split=>//=. + exists (nseq n init), Unit, (h \+ (p :-> init \+ p.+1 :-> q)); split=>//=. - by rewrite unitL. - by rewrite add0n size_nseq. (* the cycle is well-formed *) - by rewrite -(ltn_predK H0) -rcons_nseq; apply/lseg_rcons; exists p, h. + by rewrite -(ltn_predK N0) -rcons_nseq; apply/lseg_rcons; exists p, h. (* n = 0, allocate a trivial structure with a null buffer *) -move=>_; step=>m; step=>pi; step=>pa; step=>V. +step=>m; step=>pi; step=>pa; step=>V. (* it is well-formed *) exists null, null, 0, Unit=>/=; split=>//. by exists [::], Unit, Unit; split=>//; rewrite unitR. @@ -117,16 +131,16 @@ Qed. (* "polite" write, checks the buffer length, throws an exception on overflow *) Program Definition write (x : T) (b : buffer) : - {xs}, STsep (shape b xs, - fun y h => match y with - | Val _ => h \In shape b (rcons xs x) - | Exn e => e = BufferFull /\ h \In shape b xs - end) := + STsep {xs} (shape b xs, + fun y h => match y with + | Val _ => h \In shape b (rcons xs x) + | Exn e => e = BufferFull /\ h \In shape b xs + end) := Do (m <-- !len b; if m < capacity b then i <-- !inactive b; i ::= x;; - r <-- !i.+ 1; + r <-- !i.+1; inactive b ::= (r : ptr);; len b ::= m.+1 else throw BufferFull). @@ -138,11 +152,11 @@ rewrite Ec; step; case: ltnP. - (* buffer is not full, so the inactive part is non-empty *) rewrite {h}E -{1}(addn0 (size xs)) ltn_add2l => Hys. (* read the inactive pointer, unroll the inactive heap so that we can proceed *) - step; case/(lseg_lt0n Hys): Hi=>y[r][h][_ {hi}-> H]; do 4![step]=>{y}V. + step; case/(lseg_lt0n Hys): Hi=>y [r][h][_ {hi}-> H]; do 4![step]=>{y}V. (* the new structure is well-formed if the buffer is *) - exists a, r, (size xs).+1, (ha \+ (i :-> x \+ (i.+ 1 :-> r \+ h))); split=>//. + exists a, r, (size xs).+1, (ha \+ (i :-> x \+ (i.+1 :-> r \+ h))); split=>//. (* most of the conditions hold trivially or by simple arithmetics *) - exists (behead ys), (ha \+ (i :-> x \+ i.+ 1 :-> r)), h; split=>//. + exists (behead ys), (ha \+ (i :-> x \+ i.+1 :-> r)), h; split=>//. - by rewrite !joinA. - by rewrite Ec size_rcons size_behead -subn1 addnABC // subn1. - by rewrite size_rcons. @@ -157,14 +171,16 @@ exists a, i, (size xs), h; split=>//. by exists [::], ha, hi; rewrite Eys in Ec Hi. Qed. -(* a version that overwrites data in a cyclic fashion *) -(* we make checking for capacity != 0 the client's problem so it can be dealt with globally *) +(* version that overwrites data in a cyclic fashion *) +(* checking that capacity != 0 is the client's problem *) +(* so it can be dealt with globally *) Program Definition overwrite (x : T) (b : buffer) : - {xs}, STsep (fun h => 0 < capacity b /\ h \In shape b xs, - [vfun _ => shape b (drop ((size xs).+1 - capacity b) (rcons xs x))]) := + STsep {xs} (fun h => 0 < capacity b /\ h \In shape b xs, + [vfun _ => shape b (drop ((size xs).+1 - capacity b) + (rcons xs x))]) := Do (i <-- !inactive b; i ::= x;; - r <-- !i.+ 1; + r <-- !i.+1; inactive b ::= (r : ptr);; m <-- !len b; if m < capacity b then @@ -177,69 +193,72 @@ move=>x b [xs []] _ /= [Hc][a][i][_][_][_ -> [ys][ha][hi][-> Ec -> Ha Hi]]. (* read the inactive pointer, case split on inactive part *) (* we do this early on because we need to unroll the heap to proceed *) rewrite Ec in Hc *; step; case: ys Ec Hi Hc=>/=. -- (* inactive part is empty, so the buffer is full & the active part is the whole cycle *) - rewrite addn0=>Ec [Ei {hi}->] Hxs; rewrite unitR; rewrite {i}Ei in Ha *. - (* unroll the (active) heap *) +(* inactive part is empty, so the buffer is full *) +(* and the active part is the whole cycle *) +- rewrite addn0=>Ec [Ei {hi}->] Hxs; rewrite unitR; rewrite {i}Ei in Ha *. +(* unroll the (active) heap *) case/(lseg_lt0n Hxs): Ha=>z[r][h'][Exs {ha}-> H']. - (* run the rest of the program, the postcondition simplifies to beheading the extended xs *) +(* run the rest of the program *) +(* the postcondition simplifies to beheading the extended xs *) do 4!step; rewrite ltnn subSnn drop1; step=>V. - (* the new structure is well-formed if the buffer is *) - exists r, r, (size xs), (a :-> x \+ (a.+ 1 :-> r \+ h')); split=>//. - (* most of the conditions hold trivially or by simple arithmetics *) - exists [::], (a :-> x \+ (a.+ 1 :-> r \+ h')), Unit; split=>//=. +(* the new structure is well-formed if the buffer is *) + exists r, r, (size xs), (a :-> x \+ (a.+1 :-> r \+ h')); split=>//. +(* most of the conditions hold trivially or by simple arithmetics *) + exists [::], (a :-> x \+ (a.+1 :-> r \+ h')), Unit; split=>//=. - by rewrite unitR. - by rewrite addn0 size_behead size_rcons. - by rewrite size_behead size_rcons. - (* xs is non-empty so behead commutes with rcons *) +(* xs is non-empty so behead commutes with rcons *) rewrite behead_rcons //; apply/lseg_rcons. - (* the segment is well-formed *) +(* the segment is well-formed *) by exists a, h'; split=>//; rewrite [in RHS]joinC joinA. (* inactive part is non-empty, unroll the inactive heap *) move=>y ys Ec [r][h'][{hi}-> H'] _. -(* run the rest of the program, the postcondition simplifies to just the extended xs *) +(* run the rest of the program *) +(* the postcondition simplifies to just the extended xs *) do 4![step]=>{y}; rewrite -{1}(addn0 (size xs)) ltn_add2l /=; step=>V. rewrite addnS subSS subnDA subnn sub0n drop0. (* the new structure is well-formed if the buffer is *) -exists a, r, (size xs).+1, (ha \+ (i :-> x \+ (i.+ 1 :-> r \+ h'))); split=>//. +exists a, r, (size xs).+1, (ha \+ (i :-> x \+ (i.+1 :-> r \+ h'))); split=>//. (* the buffer is well-formed by simple arithmetics *) -exists ys, (ha \+ (i :-> x \+ i.+ 1 :-> r)), h'; split=>//. +exists ys, (ha \+ (i :-> x \+ i.+1 :-> r)), h'; split=>//. - by rewrite !joinA. - by rewrite Ec size_rcons addnS addSn. - by rewrite size_rcons. by apply/lseg_rcons; exists i, ha. Qed. -(* reading (or rather popping) a value from the buffer *) +(* reading (popping) a value from the buffer *) Program Definition read (b : buffer) : - {xs}, STsep (shape b xs, - fun y h => match y with - | Val x => h \In shape b (behead xs) /\ ohead xs = Some x - | Exn e => e = BufferEmpty /\ xs = [::] - end) := + STsep {xs} (shape b xs, + fun y h => match y with + | Val x => h \In shape b (behead xs) /\ ohead xs = Some x + | Exn e => e = BufferEmpty /\ xs = [::] + end) := Do (m <-- !len b; if 0 < m then a <-- !active b; x <-- !a; - r <-- !a.+ 1; + r <-- !a.+1; active b ::= (r : ptr);; len b ::= m.-1;; ret x else throw BufferEmpty). Next Obligation. -(* pull out ghost, destructure the precondition *) +(* pull out ghost, destructure precondition *) move=>b [xs []] _ /= [a][i][_][_][_ -> [ys][ha][hi][-> Hc -> Ha Hi]]. (* read the length, match on it *) step; case: ltnP. -- (* buffer is non-empty, read the active pointer *) - move=>Hxs; step. - (* unroll the active heap to proceed *) +(* buffer is non-empty, read the active pointer *) +- move=>Hxs; step. +(* unroll the active heap to proceed *) case/(lseg_lt0n Hxs): Ha=>x[r][h'][Exs {ha}-> H']. - (* run the rest of the program *) +(* run the rest of the program *) do 5![step]=>V; split; last by rewrite Exs. - (* the new structure is well-formed if the buffer is *) - exists r, i, (size xs).-1, (a :-> x \+ (a.+ 1 :-> r \+ h') \+ hi); split=>//. - (* the buffer is well-formed by simple arithmetics *) - exists (rcons ys x), h', (hi \+ (a :-> x \+ a.+ 1 :-> r)); split=>//. +(* the new structure is well-formed if the buffer is *) + exists r, i, (size xs).-1, (a :-> x \+ (a.+1 :-> r \+ h') \+ hi); split=>//. +(* the buffer is well-formed by simple arithmetics *) + exists (rcons ys x), h', (hi \+ (a :-> x \+ a.+1 :-> r)); split=>//. - by rewrite [LHS](pullX (h' \+ hi)) !joinA. - by rewrite size_rcons Hc {1}Exs /= addnS addSn. - by rewrite size_behead. @@ -249,23 +268,23 @@ by rewrite leqn0=>/nilP->; step. Qed. (* deallocating the buffer *) -(* we do check for capacity = 0 here since this should be a rare operation *) +(* check that capacity = 0 here, since this should be a rare operation *) (* of course, it could also be moved into the precondition as for overwrite *) (* loop invariant for deallocating the inner structure *) Definition free_loopT (n : nat) : Type := forall (pk : ptr * nat), - {q (xs : seq T)}, STsep (fun h => h \In lseg pk.1 q xs /\ size xs = n - pk.2, - [vfun _ : unit => emp]). + STsep {q (xs : seq T)} (fun h => h \In lseg pk.1 q xs /\ size xs = n - pk.2, + [vfun _ : unit => emp]). Program Definition free (b : buffer) : - {xs}, STsep (fun h => h \In shape b xs, - [vfun _ => emp]) := - Do (let run := Fix (fun (go : free_loopT (capacity b)) '(r,k) => + STsep {xs} (fun h => h \In shape b xs, + [vfun _ => emp]) := + Do (let run := ffix (fun (go : free_loopT (capacity b)) '(r,k) => Do (if k < capacity b then - p' <-- !r.+ 1; + p' <-- !r.+1; dealloc r;; - dealloc r.+ 1;; + dealloc r.+1;; go (p', k.+1) else ret tt)) in (if 0 < capacity b then @@ -280,13 +299,13 @@ Program Definition free (b : buffer) : Next Obligation. (* pull out ghosts, destructure the preconditions, match on k *) move=>b go _ r k [q][xs][] h /= [H Hs]; case: ltnP. -- (* k < capacity, so the spec is still non-empty *) - move=>Hk; have {Hk}Hxs: 0 < size xs by rewrite Hs subn_gt0. - (* unroll the heap *) +(* k < capacity, so the spec is still non-empty *) +- move=>Hk; have {Hk}Hxs: 0 < size xs by rewrite Hs subn_gt0. +(* unroll the heap *) case/(lseg_lt0n Hxs): H=>x[p][h'][E {h}-> H']. - (* save the next node pointer and deallocate the node *) +(* save the next node pointer and deallocate the node *) do 3!step; rewrite !unitL. - (* run the recursive call, simplify *) +(* run the recursive call, simplify *) apply: [gE q, behead xs]=>//=; split=>//. by rewrite subnS; move: Hs; rewrite {1}E /= =><-. (* k = capacity, so the spec is empty *) @@ -298,21 +317,21 @@ Qed. Next Obligation. (* pull out ghost, destructure the precondition, match on capacity *) move=>b [xs][] _ /= [a][i][m][_][_ -> [ys][ha][hi][-> Hc -> Ha Hi]]; case: ltnP. -- (* 0 < capacity, first apply the vrf_bind primitive to "step into" the `if` block *) - (* run the loop starting from the active part, frame on the corresponding heaps *) - move=>_; apply/vrf_bind; step; apply: [stepX a, xs++ys]@(ha \+ hi)=>//=. - - (* concatenate the active and inactive parts to satisfy the loop precondition *) - move=>_; split; last by rewrite size_cat subn0. +(* 0 < capacity, first apply vrf_bind to "step into" the `if` block *) +(* run the loop starting from the active part, frame the corresponding heaps *) +- move=>_; apply/vrf_bnd; step; apply: [stepX a, xs++ys]@(ha \+ hi)=>//=. +(* concatenate the active and inactive parts to satisfy the loop precondition *) + - move=>_; split; last by rewrite size_cat subn0. by apply/lseg_cat; exists i, ha, hi. - (* by deallocating the structure we empty the heap fully *) +(* by deallocating the structure, the heap is emptied *) by move=>_ _ ->; rewrite unitR; step=>V; do 3![step]=>_; rewrite !unitR. (* capacity = 0, so just deallocate the structure *) rewrite leqn0=>/eqP E; do 4!step; rewrite !unitL=>_. (* both parts of the spec are empty and so are the corresponding heaps *) -move: Hc; rewrite E =>/eqP; rewrite eq_sym addn_eq0; case/andP=>/nilP Ex /nilP Ey. +move: Hc; rewrite E =>/eqP; rewrite eq_sym addn_eq0=>/andP [/nilP Ex /nilP Ey]. move: Ha; rewrite Ex /=; case=>_->; move: Hi; rewrite Ey /=; case=>_->. by rewrite unitR. Qed. End Buffer. -End Buffer. \ No newline at end of file +End Buffer. diff --git a/examples/dlist.v b/examples/dlist.v index b54cefd..72e551d 100644 --- a/examples/dlist.v +++ b/examples/dlist.v @@ -1,3 +1,15 @@ +(* +Copyright 2021 IMDEA Software Institute +Licensed under the Apache License, Version 2.0 (the "License"); +you may not use this file except in compliance with the License. +You may obtain a copy of the License at + http://www.apache.org/licenses/LICENSE-2.0 +Unless required by applicable law or agreed to in writing, software +distributed under the License is distributed on an "AS IS" BASIS, +WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. +See the License for the specific language governing permissions and +limitations under the License. +*) From Coq Require Import ssreflect ssrbool ssrfun. From mathcomp Require Import ssrnat eqtype seq. @@ -5,15 +17,15 @@ From pcm Require Import options axioms pred. From pcm Require Import pcm unionmap heap autopcm. From htt Require Import options model heapauto. -(* Doubly linked lists, follows the same structure as singly-linked *) -(* ones, adding a second pointer pointing backwards. *) +(* Doubly-linked lists, follows the same structure as singly-linked *) +(* lists, adding a second pointer pointing backwards. *) (* The arguments are: preceding node, first node, last node, succeeding node *) (* Internally we use three cells for a node: value, next node, last node *) Fixpoint dseg {A} (pp p q qn : ptr) (xs : seq A) := if xs is hd::tl then [Pred h | exists r h', - h = p :-> hd \+ (p.+ 1 :-> r \+ (p.+ 2:-> pp \+ h')) + h = p :-> hd \+ (p.+1 :-> r \+ (p.+2:-> pp \+ h')) /\ h' \In dseg p r q qn tl] else [Pred h | [/\ p = qn, pp = q & h = Unit]]. @@ -23,11 +35,11 @@ Variable A : Type. (* structure of a list segment with appended node *) (* adding the predecessor pointer breaks the symmetry with inductive lists *) -(* so we rely on this lemma a lot to manually restructure the spec *) +(* so we rely on this lemma to manually restructure the spec *) Lemma dseg_rcons (xs : seq A) x pp p q qn h : h \In dseg pp p q qn (rcons xs x) <-> exists r h', - h = h' \+ (q :-> x \+ (q.+ 1 :-> qn \+ q.+ 2 :-> r)) + h = h' \+ (q :-> x \+ (q.+1 :-> qn \+ q.+2 :-> r)) /\ h' \In dseg pp p r q xs. Proof. elim: xs pp p h => [|y xs IH] pp p h/=. @@ -36,27 +48,29 @@ elim: xs pp p h => [|y xs IH] pp p h/=. by exists qn, Unit; rewrite unitR. split. - case=>r[_][-> /IH][s][h'][-> H']. - exists s, (p :-> y \+ (p.+ 1 :-> r \+ (p.+ 2 :-> pp \+ h'))). + exists s, (p :-> y \+ (p.+1 :-> r \+ (p.+2 :-> pp \+ h'))). rewrite !joinA; split=>//. by exists r, h'; rewrite !joinA. case=>r[_][->][s][h'][-> H']. -exists s, (h' \+ (q :-> x \+ (q.+ 1 :-> qn \+ q.+ 2 :-> r))). +exists s, (h' \+ (q :-> x \+ (q.+1 :-> qn \+ q.+2 :-> r))). rewrite !joinA; split=>//; apply/IH. by exists r, h'; rewrite !joinA. Qed. -(* first node being null means the list is empty *) +(* if first node is null, then list is empty *) Lemma dseg_nullL (xs : seq A) pp q qn h : - valid h -> h \In dseg pp null q qn xs -> + valid h -> + h \In dseg pp null q qn xs -> [/\ qn = null, pp = q, xs = [::] & h = Unit]. Proof. case: xs=>[|x xs] /= D H; first by case: H. by case: H D=>r[h'][-> _]; rewrite validPtUn eq_refl. Qed. -(* last node being null also means the list is empty *) +(* if last node is null, then list is empty *) Lemma dseg_nullR (xs : seq A) pp p qn h : - valid h -> h \In dseg pp p null qn xs -> + valid h -> + h \In dseg pp p null qn xs -> [/\ p = qn, pp = null, xs = [::] & h = Unit]. Proof. case/lastP: xs=>[|xs x] D /=; first by case. @@ -64,12 +78,13 @@ case/dseg_rcons=>r[h'][]; move: D=>/[swap]->/validR. by rewrite validPtUn. Qed. -(* deconstruct a non-trivial segment from the left *) +(* deconstruct non-trivial segment from the left *) Lemma dseg_neqL (xs : seq A) (pp p q qn : ptr) h : - p != qn -> h \In dseg pp p q qn xs -> + p != qn -> + h \In dseg pp p q qn xs -> exists x r h', [/\ xs = x :: behead xs, - h = p :-> x \+ (p.+ 1 :-> r \+ (p.+ 2 :-> pp \+ h')) & + h = p :-> x \+ (p.+1 :-> r \+ (p.+2 :-> pp \+ h')) & h' \In dseg p r q qn (behead xs)]. Proof. case: xs=>[|x xs] /= H; last first. @@ -77,12 +92,13 @@ case: xs=>[|x xs] /= H; last first. by case=>E; rewrite E eq_refl in H. Qed. -(* deconstruct a non-trivial segment from the right *) +(* deconstruct non-trivial segment from the right *) Lemma dseg_neqR (xs : seq A) (pp p q qn : ptr) h : - pp != q -> h \In dseg pp p q qn xs -> + pp != q -> + h \In dseg pp p q qn xs -> exists s x r h', [/\ xs = rcons s x, - h = h' \+ (q :-> x \+ (q.+ 1 :-> qn \+ q.+ 2 :-> r)) & + h = h' \+ (q :-> x \+ (q.+1 :-> qn \+ q.+2 :-> r)) & h' \In dseg pp p r q s]. Proof. case/lastP: xs=>[|xs x] /= H. @@ -91,17 +107,17 @@ case/dseg_rcons=>r[h'][{h}-> H']. by exists xs, x, r, h'. Qed. -(* splitting the list corresponds to splitting the heap *) +(* concatenating/splitting lists = concatenating/splitting heaps *) Lemma dseg_cat (xs ys : seq A) pp p q qn h : - h \In dseg pp p q qn (xs++ys) <-> - exists jn j, h \In dseg pp p j jn xs # dseg j jn q qn ys. + h \In dseg pp p q qn (xs ++ ys) <-> + exists jn j, h \In dseg pp p j jn xs # dseg j jn q qn ys. Proof. elim: xs pp p h=>/=. - move=>pp p h; split; first by move=>H; exists p, pp, Unit, h; rewrite unitL. by case=>jn [j][h1][h2][{h}-> [->->->]]; rewrite unitL. move=>x xs IH pp p h; split. - case=>r [_][{h}-> /IH][jn][j][h1][h2][-> H1 H2]. - exists jn, j, (p :-> x \+ p.+ 1 :-> r \+ p.+ 2 :-> pp \+ h1), h2. + exists jn, j, (p :-> x \+ p.+1 :-> r \+ p.+2 :-> pp \+ h1), h2. by rewrite !joinA; split=>//; exists r, h1; rewrite !joinA. case=>jn[j][_][h2][{h}->][r][h1][-> H1 H2]. exists r, (h1 \+ h2); rewrite !joinA; split=>//. @@ -120,106 +136,108 @@ Variable A : Type. (* specializing the segment lemmas *) Lemma dseq_nullL (xs : seq A) q h : - valid h -> h \In dseq null q xs -> + valid h -> + h \In dseq null q xs -> [/\ q = null, xs = [::] & h = Unit]. Proof. by move=>D; case/(dseg_nullL D). Qed. Lemma dseq_nullR (xs : seq A) p h : - valid h -> h \In dseq p null xs -> + valid h -> + h \In dseq p null xs -> [/\ p = null, xs = [::] & h = Unit]. Proof. by move=>D; case/(dseg_nullR D). Qed. Lemma dseq_posL (xs : seq A) p q h : - p != null -> h \In dseq p q xs -> + p != null -> + h \In dseq p q xs -> exists x r h', [/\ xs = x :: behead xs, - h = p :-> x \+ (p.+ 1 :-> r \+ (p.+ 2 :-> null \+ h')) & + h = p :-> x \+ (p.+1 :-> r \+ (p.+2 :-> null \+ h')) & h' \In dseg p r q null (behead xs)]. Proof. by apply: dseg_neqL. Qed. Lemma dseq_posR (xs : seq A) p q h : - q != null -> h \In dseq p q xs -> + q != null -> + h \In dseq p q xs -> exists s x r h', [/\ xs = rcons s x, - h = h' \+ (q :-> x \+ (q.+ 1 :-> null \+ q.+ 2 :-> r)) & + h = h' \+ (q :-> x \+ (q.+1 :-> null \+ q.+2 :-> r)) & h' \In dseg null p r q s]. Proof. by rewrite eq_sym=>Hq /(dseg_neqR Hq). Qed. (* main methods *) -(* prepend a value, return pointers to new first and last nodes *) - +(* prepend value x, return pointers to new first and last nodes *) Program Definition insertL p q (x : A) : - {l}, STsep (dseq p q l, [vfun pq => dseq pq.1 pq.2 (x::l)]) := + STsep {l} (dseq p q l, [vfun pq => dseq pq.1 pq.2 (x :: l)]) := Do (r <-- allocb x 3; - r.+ 1 ::= p;; - r.+ 2 ::= null;; + r.+1 ::= p;; + r.+2 ::= null;; if p == null then ret (r,r) - else p.+ 2 ::= r;; + else p.+2 ::= r;; ret (r,q)). Next Obligation. (* pull out ghost + precondition *) move=>p q x [l][] i /= H. (* create a new node in first 3 steps (+ rearrange pointer arith), branch *) -step=>r; rewrite unitR ptrA; do 2!step; case: ifP H=>[/eqP->|/negbT N]. +step=>r; rewrite unitR; do 2!step; case: ifP H=>[/eqP ->|/negbT N]. - (* the list is empty, so new first node = last node *) move/dseq_nullL=>H; step; rewrite joinA=>/validR/H [_->->] /=. by exists null, Unit; rewrite !joinA. (* deconstruct non-empty list, run the rest *) case/(dseq_posL N)=>y[z][h'][E {i}-> H']; do 2![step]=>V. (* massage the heap to fit the goal *) -exists p, (p :-> y \+ (p.+ 1 :-> z \+ (p.+ 2 :-> r \+ h'))). +exists p, (p :-> y \+ (p.+1 :-> z \+ (p.+2 :-> r \+ h'))). by rewrite !joinA; split=>//; rewrite E /=; exists z, h'; rewrite !joinA. Qed. -(* append a value (in constant), return pointers to new first and last nodes *) - +(* append value x, return pointers to new first and last nodes *) Program Definition insertR p q (x : A) : - {l}, STsep (dseq p q l, [vfun pq => dseq pq.1 pq.2 (rcons l x)]) := + STsep {l} (dseq p q l, [vfun pq => dseq pq.1 pq.2 (rcons l x)]) := Do (r <-- allocb x 3; - r.+ 1 ::= null;; - r.+ 2 ::= q;; + r.+1 ::= null;; + r.+2 ::= q;; if q == null then ret (r,r) - else q.+ 1 ::= r;; - ret (p,r)). + else q.+1 ::= r;; + ret (p,r)). Next Obligation. (* pull out ghost + precondition *) move=>p q x [l []] i /= H. (* create a new node in first 3 steps (+ rearrange pointer arith), branch *) -step=>r; rewrite unitR ptrA; do 2!step; case: ifP H=>[/eqP->|/negbT N]. +step=>r; rewrite unitR; do 2!step; case: ifP H=>[/eqP->|/negbT N]. - (* the list is empty, so new first node = last node *) move/dseq_nullR=>H; step; rewrite joinA=>/validR/H [_->->] /=. by exists null, Unit; rewrite !joinA. (* deconstruct non-empty list, run the rest, restructure the goal *) case/(dseq_posR N)=>s[y][z][h'][{l}-> {i}-> H']; do 2![step]=>_; apply/dseg_rcons. (* massage the heap and simplify *) -exists q, (h' \+ (q :-> y \+ (q.+ 1 :-> r \+ q.+ 2 :-> z))). +exists q, (h' \+ (q :-> y \+ (q.+1 :-> r \+ q.+2 :-> z))). split; first by rewrite joinC. (* restructure the goal once more *) by apply/dseg_rcons; vauto. Qed. -(* traversing the dlist backwards and consing all elements *) +(* travers the dlist backwards and cons all elements *) (* reifies the specification *) -(* the loop invariant: *) -(* 1. the heap is unchanged *) -(* 2. the result is remainder + accumulator *) -(* we carry the pointer to the accumulator as a ghost var *) +(* loop invariant: *) +(* 1. heap is unchanged *) +(* 2. result = remainder + accumulator *) +(* carry the pointer to the accumulator as a logical var *) Definition traverse_backT (p : ptr) : Type := forall (qs : ptr * seq A), - {l nx}, STsep (dseg null p qs.1 nx l, - [vfun r h => h \In dseg null p qs.1 nx l - /\ r = l ++ qs.2]). + STsep {l nx} (dseg null p qs.1 nx l, + [vfun r h => h \In dseg null p qs.1 nx l /\ + r = l ++ qs.2]). Program Definition traverse_back p q : - {l}, STsep (dseq p q l, - [vfun r h => h \In dseq p q l /\ r = l]) := + STsep {l} (dseq p q l, + [vfun r h => h \In dseq p q l /\ r = l]) := Do (let tb := - Fix (fun (go : traverse_backT p) '(r, acc) => + ffix (fun (go : traverse_backT p) '(r, acc) => Do (if r == null then ret acc else x <-- !r; - rnxt <-- !(r.+ 2); + rnxt <-- !r.+2; go (rnxt, x :: acc))) in tb (q, [::])). (* first, the loop *) diff --git a/examples/exploit.v b/examples/exploit.v old mode 100755 new mode 100644 index 0a6c959..f42ffa3 --- a/examples/exploit.v +++ b/examples/exploit.v @@ -1,3 +1,16 @@ +(* +Copyright 2009 IMDEA Software Institute +Licensed under the Apache License, Version 2.0 (the "License"); +you may not use this file except in compliance with the License. +You may obtain a copy of the License at + http://www.apache.org/licenses/LICENSE-2.0 +Unless required by applicable law or agreed to in writing, software +distributed under the License is distributed on an "AS IS" BASIS, +WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. +See the License for the specific language governing permissions and +limitations under the License. +*) + From Coq Require Import ssreflect ssrbool Logic.Hurkens. (* This file shows the unsoundness of the axiom pack_injective assumed in *) diff --git a/examples/gcd.v b/examples/gcd.v index bd85688..40c39b0 100644 --- a/examples/gcd.v +++ b/examples/gcd.v @@ -1,6 +1,19 @@ +(* +Copyright 2021 IMDEA Software Institute +Licensed under the Apache License, Version 2.0 (the "License"); +you may not use this file except in compliance with the License. +You may obtain a copy of the License at + http://www.apache.org/licenses/LICENSE-2.0 +Unless required by applicable law or agreed to in writing, software +distributed under the License is distributed on an "AS IS" BASIS, +WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. +See the License for the specific language governing permissions and +limitations under the License. +*) + From Coq Require Import ssreflect ssrbool ssrfun. From mathcomp Require Import ssrnat eqtype div. -From pcm Require Import axioms pred pcm heap. +From pcm Require Import axioms pred ordtype pcm heap. From htt Require Import options model heapauto. (* classical mutable Euclid's algorithm for computing GCD with subtractions *) @@ -13,12 +26,12 @@ Definition shape (p q : ptr) (x y : nat) := (* (`unit` is needed because `Fix` always requires a single parameter) *) Definition gcd_loopT (p q : ptr) : Type := unit -> - {x y : nat}, STsep (shape p q x y, - [vfun (_ : unit) h => - h \In shape p q (gcdn x y) (gcdn x y)]). + STsep {x y : nat} (shape p q x y, + [vfun (_ : unit) h => + h \In shape p q (gcdn x y) (gcdn x y)]). Program Definition gcd_loop (p q : ptr) := - Fix (fun (go : gcd_loopT p q) _ => + ffix (fun (go : gcd_loopT p q) _ => Do (x <-- !p; y <-- !q; if (x : nat) != y then @@ -48,11 +61,11 @@ suff {p q go m}: gcdn x (y - x) = gcdn x y by move=>->. by rewrite -gcdnDr subnK //; apply: ltnW. Qed. -(* note that there's no guarantee on termination, so we have only partial correctness *) -(* i.e. the algorithm will get stuck if u = 0 /\ v != 0 or vice versa *) +(* There's no guarantee on termination, as this is partial correctness. *) +(* The algorithm loops if u = 0 /\ v != 0 or vice versa *) Program Definition gcd u v : STsep (PredT, [vfun r _ => r = gcdn u v]) := - (* we allocate in the reverse order because the symbolic heap behaves as a stack *) + (* allocate in the reverse order because the symbolic heap behaves as a stack *) (* this way it'll match the specification perfectly, removing a bit of bureaucracy *) Do (q <-- alloc v; p <-- alloc u; diff --git a/examples/hashtab.v b/examples/hashtab.v old mode 100755 new mode 100644 index 2e7f552..bc5e7c4 --- a/examples/hashtab.v +++ b/examples/hashtab.v @@ -1,30 +1,66 @@ +(* +Copyright 2010 IMDEA Software Institute +Licensed under the Apache License, Version 2.0 (the "License"); +you may not use this file except in compliance with the License. +You may obtain a copy of the License at + http://www.apache.org/licenses/LICENSE-2.0 +Unless required by applicable law or agreed to in writing, software +distributed under the License is distributed on an "AS IS" BASIS, +WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. +See the License for the specific language governing permissions and +limitations under the License. +*) + +From HB Require Import structures. From Coq Require Import ssreflect ssrbool ssrfun. From mathcomp Require Import ssrnat eqtype seq fintype tuple finfun finset. From pcm Require Import options axioms prelude pred ordtype finmap. From pcm Require Import pcm unionmap heap autopcm. -From htt Require Import model heapauto. +From htt Require Import options model heapauto. From htt Require Import array kvmaps. -Module HashTab. -Section HashTab. - -(* a hash table is an array of buckets, i.e. KV maps *) +(* hash table is array of buckets, i.e. KV maps *) (* bucket indices are provided by the hash function *) +(* using dynaming kv-maps for buckets *) -Variables (K : ordType) (V : Type) (buckets : KVmap.Sig K V) - (n : nat) (hash : K -> 'I_n). -Definition hashtab := {array 'I_n -> KVmap.tp buckets}. -Notation KVshape := (@KVmap.shape _ _ buckets). -Notation table := (table KVshape). -Notation nil := (nil K V). +Module Type Hashtab_sig. +Parameter root : forall {K : ordType} {V : Type} (buckets : dkvm K V) + {n : nat} (hash : K -> 'I_n), Set. +Section HashTab. +Context {K : ordType} {V : Type} {buckets : dkvm K V}. +Context {n : nat} {hash : K -> 'I_n}. +Parameter null_root : root buckets hash. +Parameter shape : root buckets hash -> {finMap K -> V} -> Pred heap. +Parameter new : + STsep (emp, [vfun (x : root buckets hash) h => h \In shape x nil]). +Parameter free : forall x : root buckets hash, + STsep {s} (shape x s, [vfun _ : unit => emp]). +Parameter insert : forall (x : root buckets hash) k v, + STsep {s} (shape x s, [vfun (_ : unit) h => h \In shape x (ins k v s)]). +Parameter remove : forall (x : root buckets hash) k, + STsep {s} (shape x s, [vfun (_ : unit) h => h \In shape x (rem k s)]). +Parameter lookup : forall (x : root buckets hash) k, + STsep {s} (shape x s, [vfun y h => h \In shape x s /\ y = fnd k s]). +End HashTab. +End Hashtab_sig. -Opaque Array.write Array.new Array.free Array.read. +Module HashTab : Hashtab_sig. +Definition root K V (buckets : dkvm K V) n (hash : K -> 'I_n) : + Set := {array 'I_n -> buckets}. +Definition null_root K V buckets n hash : + @root K V buckets n hash := Array null. +Section HashTab. +Context (K : ordType) (V : Type) {buckets : dkvm K V} + {n : nat} {hash : K -> 'I_n}. +Notation KVshape := (@dkvm_shape _ _ buckets). +Notation table := (table KVshape). +Notation root := (root buckets hash). (* hash table is specified by a single finMap *) -(* which is morally a "flattening" of all buckets *) -Definition shape x (s : finMap K V) := - [Pred h | exists (tab : {ffun 'I_n -> KVmap.tp buckets}) (* array spec *) - (bucket : 'I_n -> finMap K V), (* buckets spec *) +(* which is the "flattening" of all buckets *) +Definition shape (x : root) (s : finMap K V) : Pred heap := + [Pred h | exists (tab : {ffun 'I_n -> buckets}) (* array spec *) + (bucket : 'I_n -> {finMap K -> V}), (* buckets spec *) [/\ forall k, fnd k s = fnd k (bucket (hash k)), forall i k, k \in supp (bucket i) -> hash k = i & h \In Array.shape x tab # sepit setT (table tab bucket)] ]. @@ -36,24 +72,25 @@ Definition shape x (s : finMap K V) := Definition new_loopinv x := forall k, STsep (fun h => k <= n /\ exists tab, h \In Array.shape x tab # - sepit [set x:'I_n | x < k] (table tab (fun=> nil)), - [vfun y => shape y nil]). + sepit [set x:'I_n | x < k] (table tab (fun=>nil)), + [vfun y => shape y nil]). Program Definition new : STsep (emp, [vfun y => shape y nil]) := - Do (t <-- Array.new _ (KVmap.default buckets); - let go := Fix (fun (loop : new_loopinv t) k => + Do (t <-- Array.new dkvm_null; + let go := ffix (fun (loop : new_loopinv t) k => Do (if decP (b := k < n) idP is left pf then - b <-- KVmap.new buckets; + b <-- dkvm_new buckets; Array.write t (Ordinal pf) b;; loop k.+1 else ret t)) in go 0). (* first the bucket initialization loop *) -Next Obligation. +Next Obligation. (* pull out preconditions, branch on k comparison *) -move=>/= arr loop k [] _ /= [Eleq][tab][h1][h2][-> H1]; case: decP=>[{Eleq}pf H2|]; last first. -- (* k = n, return the array pointer *) - case: ltnP Eleq (eqn_leq k n)=>// _ -> /= /eqP Ek _ H; step=>_. +move=>/= arr loop k [] _ /= [Eleq][tab][h1][h2][-> H1]. +case: decP=>[{Eleq}pf H2|]; last first. +(* k = n, return the array pointer *) +- case: ltnP Eleq (eqn_leq k n)=>// _ -> /= /eqP Ek _ H; step=>_. (* pass through the constructed table, aggregated finmap is empty *) exists tab, (fun => nil); split=>//; exists h1, h2; split=>//{h1 H1}. (* h2 holds the table *) @@ -61,33 +98,36 @@ move=>/= arr loop k [] _ /= [Eleq][tab][h1][h2][-> H1]; case: decP=>[{Eleq}pf H2 (* k < n, allocate an empty bucket by framing on Unit *) apply: [stepU]=>//= b m Hm. (* write its id to the array under index k *) -apply: [stepX tab] @ h1=>{h1 H1}//= _ m2 [E2 V2]. +apply: [stepX tab] @ h1=>{h1 H1}//= _ m2 E2. (* invoke the loop *) apply: [gE]=>//=; split=>//; rewrite joinCA. (* extend the table by the new index/bucket pair, simplify *) -exists [ffun z => if z == Ordinal pf then b else tab z], m2, (m \+ h2); split=>//{m2 E2 V2}. +exists [ffun z => if z == Ordinal pf then b else tab z], m2, (m \+ h2). +split=>//{m2 E2}. (* remove the new bucket from the heap *) -rewrite (sepitS (Ordinal pf)) in_set leqnn {1}/table ffunE eq_refl; exists m, h2; do!split=>{m Hm}//. -apply: tableP2 H2=>{h2}//. -- by case=>x Hx; rewrite !in_set -val_eqE /= ltnS (leq_eqVlt x); case: ltngtP. +rewrite (sepitS (Ordinal pf)) in_set leqnn {1}/table ffunE eq_refl. +exists m, h2; do!split=>{m Hm}//; apply: tableP2 H2=>{h2}//. +- case=>x Hx; rewrite !in_set in_set1 -val_eqE /= ltnS (leq_eqVlt x). + by case: ltngtP. (* removing k from the domain of the new table gives the old table back *) by move=>x _; rewrite in_set ffunE; case: eqP=>//->; rewrite ltnn. Qed. -(* now the outer function *) +(* the outer function *) Next Obligation. (* simplify *) move=>/= [] _ ->. (* allocate the array *) -apply: [stepE]=>//= y m [Hm Vm]. +apply: [stepE]=>//= y m Hm. (* invoke the loop with index 0 *) apply: [gE]=>//=; split=>//. (* the table is empty *) -exists [ffun => KVmap.default buckets], m, Unit; split=>//=; first by rewrite unitR. +exists [ffun => dkvm_null], m, Unit; split=>//=. +- by rewrite unitR. (* there are no buckets in the heap yet *) by rewrite (eq_sepit (s2 := set0)) // sepit0. Qed. -(* freeing a hashtable is freeing the array + buckets *) +(* freeing hashtable = freeing the array + buckets *) (* loop invariant: *) (* at iteration k, the heap still holds the array and n-k buckets *) @@ -95,39 +135,42 @@ Definition free_loopinv x := forall k, STsep (fun h => k <= n /\ exists t b, h \In Array.shape x t # sepit [set x:'I_n | x >= k] (table t b), - [vfun _ : unit => emp]). + [vfun _ : unit => emp]). -Program Definition free x : {s}, STsep (shape x s, - [vfun _ : unit => emp]) := - (* we add an extra Do here so we can derive the precondition from the loop *) - Do (Fix (fun (loop : free_loopinv x) k => +Program Definition free x : STsep {s} (shape x s, + [vfun _ : unit => emp]) := + (* the extra Do here enables deriving precondition from the loop *) + Do (ffix (fun (loop : free_loopinv x) k => Do (if decP (b := k < n) idP is left pf then b <-- Array.read x (Ordinal pf); - KVmap.free b;; + dkvm_free b;; loop k.+1 else Array.free x)) 0). (* first the loop *) Next Obligation. (* pull out the ghost + preconditions, branch on k comparison *) -move=>/= x loop k [] _ /= [Eleq][tf][bf][h1][h2][-> [H1 V1]]; case: decP=>[pf H|]; last first. -- (* k = n *) - case: ltnP Eleq (eqn_leq k n)=>// _ -> /= /eqP Ek _ H. +move=>/= x loop k [] _ /= [Eleq][tf][bf][h1][h2][-> H1]. +case: decP=>[pf H|]; last first. +(* k = n *) +- case: ltnP Eleq (eqn_leq k n)=>// _ -> /= /eqP Ek _ H. (* free the array *) apply: [gE]=>//=; exists tf. (* h2 is empty *) - move: H; rewrite (eq_sepit (s2 := set0)); first by rewrite sepit0=>->; rewrite unitR. + move: H; rewrite (eq_sepit (s2 := set0)). + - by rewrite sepit0=>->; rewrite unitR. by move=>y; rewrite Ek in_set in_set0 leqNgt ltn_ord. (* k < n, read from array *) apply: [stepX tf, h1] @ h1=>//= _ _ [->->]. (* split out the the k-th bucket *) -move: H; rewrite (sepitS (Ordinal pf)) in_set leqnn; case=>h3[h4][{h2}-> H3 H4]. +move: H; rewrite (sepitS (Ordinal pf)) in_set leqnn. +case=>h3[h4][{h2}-> H3 H4]. (* free it *) apply: [stepX (bf (Ordinal pf))] @ h3=>{h3 H3}//= _ _ ->; rewrite unitL. (* invoke loop, simplify *) apply: [gE]=>//=; split=>//; exists tf, bf, h1, h4; split=>//. (* drop the k-th entry from the table *) apply/tableP2/H4=>//. -move=>z; rewrite !in_set; case: eqP=>/=. +move=>z; rewrite !in_set in_set1; case: eqP=>/=. - by move=>->/=; rewrite ltnn. by move/eqP; rewrite -val_eqE /=; case: ltngtP. Qed. @@ -138,20 +181,19 @@ move=>/= x [fm][] h /= [tf][bf][_ _ H]. by apply: [gE]=>//=; eauto. Qed. -(* inserting into a hashmap is inserting into corresponding bucket + updating the array *) -(* returning the pointer is technically not needed, as the array is not moved *) -(* but we need to fit the KV map API *) - -Program Definition insert x k v : {s}, STsep (shape x s, - [vfun y => shape y (ins k v s)]) := +(* inserting into hashmap is inserting into *) +(* corresponding bucket + updating the array *) +(* returning the pointer is not needed *) +(* as the array is not moved *) +Program Definition insert x k v : + STsep {s} (shape x s, [vfun _ : unit => shape x (ins k v s)]) := Do (let hk := hash k in b <-- Array.read x hk; - b' <-- KVmap.insert b k v; - Array.write x hk b';; - ret x). + b' <-- dkvm_insert b k v; + Array.write x hk b'). Next Obligation. (* pull out ghost + deconstruct precondition *) -move=>/= x k v [fm][] _ /= [tf][bf][Hf Hh [h1][h2][-> [/= H1 _] H2]]. +move=>/= x k v [fm][] _ /= [tf][bf][Hf Hh [h1][h2][-> /= H1 H2]]. (* read the bucket from array *) apply/[stepX tf, h1] @ h1=>//= _ _ [->->]. (* split out the bucket in the heap *) @@ -159,41 +201,42 @@ move: H2; rewrite (sepitT1 (hash k)) /table; case=>h3[h4][{h2}-> H3 H4]. (* insert into the bucket *) apply/[stepX (bf (hash k))] @ h3=>{h3 H3}//= b' m2 H2. (* write the bucket to the array, return the pointer *) -apply/[stepX tf] @ h1=>{h1 H1}//= _ m3 [E3 V3]; step=>_. +apply: [gX tf]@h1=>{h1 H1} //= _ m3 E3 _. (* update the array and buckets' specs *) exists [ffun z => if z == hash k then b' else tf z], (fun b => if b == hash k then ins k v (bf b) else bf b); split=>/=. -- (* the new buckets spec is still a flattening *) - move=>k0; rewrite fnd_ins; case: eqP=>[->|/eqP Ek]. - - (* if the bucket was touched, we get the same value *) - by rewrite eq_refl fnd_ins eq_refl. +(* the new buckets spec is still a flattening *) +- move=>k0; rewrite fnd_ins; case: eqP=>[->|/eqP Ek]. + (* if the bucket was touched, we get the same value *) + - by rewrite eq_refl fnd_ins eq_refl. (* if not, we get the old spec back *) by case: ifP=>_ //; rewrite fnd_ins (negbTE Ek). -- (* the new buckets spec respects the hash *) - move=>i0 k0; case: eqP; last by move=>_; apply: Hh. +(* the new buckets spec respects the hash *) +- move=>i0 k0; case: eqP; last by move=>_; apply: Hh. by move=>->; rewrite supp_ins inE=>/orP; case; [move/eqP=>->|apply: Hh]. (* the shape is respected: first, the array fits *) -exists m3, (m2 \+ h4); split=>{Hf Hh m3 E3 V3}//. +exists m3, (m2 \+ h4); split=>{Hf Hh m3 E3}//. (* split out the modified bucket *) -rewrite (sepitT1 (hash k)) /table /= ffunE eq_refl; exists m2, h4; split=>{m2 H2}//. +rewrite (sepitT1 (hash k)) /table /= ffunE eq_refl. +exists m2, h4; split=>{m2 H2} //. (* the table fits too *) -by apply/tableP/H4=>/= x0; rewrite !in_set andbT ?ffunE =>/negbTE->. +apply/tableP/H4=>/= x0; +by rewrite !in_set in_set1 andbT ?ffunE =>/negbTE->. Qed. -(* removing from a hashmap is removing from corresponding bucket + updating the array *) -(* returning the pointer is again not needed except for the API fit *) - +(* removing from hashmap is removing from *) +(* corresponding bucket + updating the array *) +(* returning the pointer is again not needed *) Program Definition remove x k : - {s}, STsep (shape x s, - [vfun y => shape y (rem k s)]) := + STsep {s} (shape x s, + [vfun _ : unit => shape x (rem k s)]) := Do (let hk := hash k in b <-- Array.read x hk; - b' <-- KVmap.remove b k; - Array.write x hk b';; - ret x). + b' <-- dkvm_remove b k; + Array.write x hk b'). Next Obligation. (* pull out ghost + destructure precondition *) -move=>/= x k [fm][] _ /= [tf][bf][Hf Hh [h1][h2][-> [/= H1 _] H2]]. +move=>/= x k [fm][] _ /= [tf][bf][Hf Hh [h1][h2][-> /= H1 H2]]. (* read the bucket from array *) apply/[stepX tf, h1] @ h1=>//= _ _ [->->]. (* split out the bucket in the heap *) @@ -201,34 +244,35 @@ move: H2; rewrite (sepitT1 (hash k)); case=>h3[h4][{h2}-> H3 H4]. (* insert into the bucket *) apply/[stepX (bf (hash k))] @ h3=>{h3 H3}//= b' m2 H2. (* write the bucket to the array, return the pointer *) -apply/[stepX tf] @ h1=>{H1}//= _ m3 [E3 V3]; step=>_. +apply/[gX tf] @ h1=>{H1}//= _ m3 E3 _. (* update the array and buckets' specs *) exists [ffun z => if z == hash k then b' else tf z], (fun b => if b == hash k then rem k (bf b) else bf b); split=>/=. -- (* the new buckets spec is still a flattening *) - move=>k0; rewrite fnd_rem; case: eqP. - - (* if the bucket was touched, the value is gone *) - by move=>->; rewrite eq_refl fnd_rem eq_refl. +(* the new buckets spec is still a flattening *) +- move=>k0; rewrite fnd_rem; case: eqP. + (* if the bucket was touched, the value is gone *) + - by move=>->; rewrite eq_refl fnd_rem eq_refl. (* if not, we get the old spec back *) by move/eqP=>Ek; case: ifP=>_ //; rewrite fnd_rem (negbTE Ek). -- (* the new buckets spec respects the hash *) - move=>i0 k0; case: eqP; last by move=>_; apply: Hh. +(* the new buckets spec respects the hash *) +- move=>i0 k0; case: eqP; last by move=>_; apply: Hh. by move=>->; rewrite supp_rem !inE=>/andP [] _; apply: Hh. (* the shape is respected: first, the array fits *) -exists m3, (m2\+ h4); split=>{m3 E3 V3 Hf Hh}//. +exists m3, (m2\+ h4); split=>{m3 E3 Hf Hh}//. (* split out the modified bucket *) -rewrite (sepitT1 (hash k)) /table /= ffunE eq_refl; exists m2, h4; split=>{m2 H2}//. +rewrite (sepitT1 (hash k)) /table /= ffunE eq_refl. +exists m2, h4; split=>{m2 H2} //. (* the table fits too *) -by apply/tableP/H4=>/= x0; rewrite !in_set andbT ?ffunE =>/negbTE->. +apply/tableP/H4=>/= x0; +by rewrite !in_set in_set1 andbT ?ffunE => /negbTE->. Qed. (* looking up in a hashmap is looking up in the corresponging bucket *) - Program Definition lookup x k : - {s}, STsep (shape x s, + STsep {s} (shape x s, [vfun y m => m \In shape x s /\ y = fnd k s]) := Do (b <-- Array.read x (hash k); - KVmap.lookup b k). + dkvm_lookup b k). Next Obligation. (* pull out ghost + destructure precondition *) move=>/= x k [fm][] _ /= [tf][bf][Hf Hh [h1][h2][-> H1 H2]]. @@ -237,18 +281,24 @@ apply/[stepX tf, h1] @ h1=>//= _ _ [->->]. (* split out the bucket in the heap *) move: H2; rewrite (sepitT1 (hash k)); case=>h3[h4][{h2}-> H3 H4]. (* look up in the bucket, simplify *) -apply/[gX (bf (hash k))] @ h3=>{h3 H3}//= r m2 [H2 Hr] _; split; last by rewrite Hf. +apply/[gX (bf (hash k))] @ h3=>{h3 H3}//= r m2 [H2 Hr] _. +split; last by rewrite Hf. (* the shape is preserved, as nothing was modified *) -exists tf, bf; split=>//=; exists h1, (m2 \+ h4); split=>{h1 H1}//. +exists tf, bf; split=>//=; exists h1, (m2 \+ h4); split=>{h1 H1} //. by rewrite (sepitT1 (hash k)); vauto. Qed. +End HashTab. +End HashTab. -(* hash table is a KV map *) +(* hash table is (static) KV map *) +Notation hashtab := HashTab.root. +HB.instance Definition _ K V (buckets : dkvm K V) n (hash : K -> 'I_n) := + isKVM.Build K V (hashtab buckets hash) HashTab.null_root + HashTab.new HashTab.free HashTab.insert HashTab.remove HashTab.lookup. -Definition HashTab := KVmap.make (Array null) new free insert remove lookup. +(* htab is specific simple hash tab where buckets are association lists *) +Definition htab {K} V {n} hash := @hashtab K V (dalist K V) n hash. +HB.instance Definition _ K V n hash := KVM.on (@htab K V n hash). -End HashTab. -End HashTab. -Definition HT K V := HashTab.HashTab (AssocList.AssocList K V). \ No newline at end of file diff --git a/examples/kvmaps.v b/examples/kvmaps.v old mode 100755 new mode 100644 index 8b9f1cd..59a3c38 --- a/examples/kvmaps.v +++ b/examples/kvmaps.v @@ -1,61 +1,119 @@ +(* +Copyright 2020 IMDEA Software Institute +Licensed under the Apache License, Version 2.0 (the "License"); +you may not use this file except in compliance with the License. +You may obtain a copy of the License at + http://www.apache.org/licenses/LICENSE-2.0 +Unless required by applicable law or agreed to in writing, software +distributed under the License is distributed on an "AS IS" BASIS, +WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. +See the License for the specific language governing permissions and +limitations under the License. +*) + (******************) (* Key-Value maps *) (******************) +From HB Require Import structures. From Coq Require Import ssreflect ssrbool ssrfun. -From mathcomp Require Import eqtype seq path. -From pcm Require Import options axioms pred ordtype finmap. +From mathcomp Require Import eqtype seq path ssrnat. +From pcm Require Import options axioms pred ordtype finmap seqext. From pcm Require Import pcm unionmap heap autopcm automap. -From htt Require Import options interlude model heapauto. - -(***********************) -(* stateful KV map ADT *) -(***********************) - -Module KVmap. -Record Sig (K : ordType) (V : Type) : Type := - make {tp :> Type; - - default : tp; - - shape : tp -> {finMap K -> V} -> Pred heap; - - new : STsep (emp, [vfun x => shape x (nil K V)]); - - free : forall x, {s}, STsep (shape x s, - [vfun _ : unit => emp]); - - insert : forall x k v, - {s}, STsep (shape x s, - [vfun y => shape y (ins k v s)]); - - remove : forall x k, - {s}, STsep (shape x s, - [vfun y => shape y (rem k s)]); - - lookup : forall x k, - {s}, STsep (shape x s, - [vfun y m => m \In shape x s /\ y = fnd k s])}. -End KVmap. - -(**********************************************************) -(* KV map implemented as a sorted association linked list *) -(**********************************************************) - -Module AssocList. +From htt Require Import options model heapauto. + +(* Dynamic KV map is determined by its root pointer(s). *) +(* Functions such insert and remove may modify *) +(* the root, and will correspondingly return the new one. *) +(* root is abstracted to facilitate structures that may *) +(* have more than one root pointers. Also, it enables *) +(* passing K and V to methods thus reducing annotations *) +(* There's no deep reason to make root : Set, except that *) +(* it should be thought of a collection of pointers, hence small. *) + +HB.mixin Record isDKVM (K : ordType) (V : Type) (root : Set) := { + dkvm_null : root; + dkvm_shape : root -> {finMap K -> V} -> Pred heap; + dkvm_new : STsep (emp, [vfun x => dkvm_shape x nil]); + dkvm_free : forall x, + STsep {s} (dkvm_shape x s, [vfun _ : unit => emp]); + dkvm_insert : forall x k v, + STsep {s} (dkvm_shape x s, [vfun y => dkvm_shape y (ins k v s)]); + dkvm_remove : forall x k, + STsep {s} (dkvm_shape x s, [vfun y => dkvm_shape y (rem k s)]); + dkvm_lookup : forall x k, + STsep {s} (dkvm_shape x s, + [vfun y m => m \In dkvm_shape x s /\ y = fnd k s])}. + +#[short(type="dkvm")] +HB.structure Definition DKVM K V := {root of @isDKVM K V root}. + +Arguments dkvm_new {K V}. + +(* Static KV map (or just KV map) are those that *) +(* don't need to modify the root pointer. *) +(* Typical example will be hash tables *) +(* Another example is a structure obtained by packaging *) +(* the root pointer along with the dynamic KV map that the *) +(* root points to. *) + +HB.mixin Record isKVM (K : ordType) (V : Type) (root : Set) := { + kvm_null : root; + kvm_shape : root -> {finMap K -> V} -> Pred heap; + kvm_new : STsep (emp, [vfun x => kvm_shape x nil]); + kvm_free : forall x, + STsep {s} (kvm_shape x s, [vfun _ : unit => emp]); + kvm_insert : forall x k v, + STsep {s} (kvm_shape x s, [vfun _ : unit => kvm_shape x (ins k v s)]); + kvm_remove : forall x k, + STsep {s} (kvm_shape x s, [vfun _ : unit => kvm_shape x (rem k s)]); + kvm_lookup : forall x k, + STsep {s} (kvm_shape x s, + [vfun y m => m \In kvm_shape x s /\ y = fnd k s])}. + +#[short(type="kvm")] +HB.structure Definition KVM K V := {root of @isKVM K V root}. + +Arguments kvm_new {K V}. + +(*****************************************) +(* Association lists with DKVM interface *) +(*****************************************) + +(* module type for making the defs opaque *) +Module Type DAList_sig. +Parameter root : ordType -> Type -> Set. +Section DAlist. +Context {K : ordType} {V : Type}. +Parameter null_root : root K V. +Parameter shape : root K V -> {finMap K -> V} -> Pred heap. +Parameter new : STsep (emp, [vfun (x : root K V) h => h \In shape x nil]). +Parameter free : forall x : root K V, + STsep {s} (shape x s, [vfun _ : unit => emp]). +Parameter insert : forall (x : root K V) k v, + STsep {s} (shape x s, [vfun (y : root K V) h => h \In shape y (ins k v s)]). +Parameter remove : forall (x : root K V) k, + STsep {s} (shape x s, [vfun (y : root K V) h => h \In shape y (rem k s)]). +Parameter lookup : forall (x : root K V) k, + STsep {s} (shape x s, [vfun y m => m \In shape x s /\ y = fnd k s]). +End DAlist. +End DAList_sig. + +Module DAList : DAList_sig. +Definition root (K : ordType) (V : Type) := ptr. +Definition null_root K V : @root K V := null. Section AssocList. - -Variable (K : ordType) (V : Set). +Variables (K : ordType) (V : Type). Notation fmap := (finMap K V). -Notation nil := (nil K V). (* single entry of the map as a triple of heap cells *) Definition entry (p q : ptr) (k : K) (v : V) : heap := p :-> k \+ p.+1 :-> v \+ p.+2 :-> q. -(* similarly to plain linked list, we with a generic definition of a segment and *) -(* then specialize it to self-contained lists specified by a finMap structure *) - +(* similarly to plain linked list *) +(* development starts with generic definition of list segment *) +(* and then specializes to self-contained lists specified *) +(* by finMap structure *) Fixpoint shape_seg' (x y : ptr) (xs : seq (K * V)) : Pred heap := if xs is (k,v) :: tl then [Pred h | exists q h', @@ -66,63 +124,59 @@ Fixpoint shape_seg' (x y : ptr) (xs : seq (K * V)) : Pred heap := Definition shape_seg (x y : ptr) (s : fmap) : Pred heap := shape_seg' x y (seq_of s). +(* definition that adds dymmy dependence on D *) +(* to make the module system pass *) Definition shape (x : ptr) (s : fmap) : Pred heap := shape_seg x null s. -(* null pointer represents an empty map *) - +(* null pointer represents empty map *) Lemma shape_null (s : fmap) h : - valid h -> h \In shape null s -> - s = nil /\ h = Unit. + valid h -> + h \In shape null s -> + s = nil /\ h = Unit. Proof. -move=>D; case: s; case=>/=. +move=>W; case: s; case=>/=. - by move=>? [_ ->] //; rewrite fmapE. move=>[??]??[?][?][H]. -by rewrite H -!joinA validPtUn in D. +by rewrite H -!joinA validPtUn in W. Qed. (* non-empty well-formed region represents a non-empty map such that *) (* a head entry respecting the key order can be pulled out of the map *) - Lemma shape_cont (s : fmap) p h : - p != null -> h \In shape p s -> - exists k v q h', - [/\ s = ins k v (behd s), (* k:->v is the head entry *) - all (ord k) (supp (behd s)), - h = entry p q k v \+ h' & - h' \In shape q (behd s)]. + p != null -> + h \In shape p s -> + exists k v q h', + [/\ s = ins k v (behd s), (* k:->v is the head entry *) + all (ord k) (supp (behd s)), + h = entry p q k v \+ h' & + h' \In shape q (behd s)]. Proof. move=>E; case: s=>xs srt /=. -elim: xs srt=>/=. -- by move=>? [E0]; rewrite E0 in E. +elim: xs srt=>/=; first by move=>? [E0]; rewrite E0 in E. move=>[k v] /= l IH srt [q][h'][-> H]. -exists k,v,q,h'; split=>//; last by apply: path_all. +exists k, v, q, h'; split=>//; last by apply: order_path_min. by rewrite fmapE /= last_ins'. Qed. -(* TODO move to finmap? *) -Lemma all_path (s : fmap) k : all (ord k) (supp s) -> path ord k (supp s). -Proof. by rewrite path_sortedE // =>->/=; case: s. Qed. - (* inserting an entry with minimal key is prepending to the list *) - Lemma shape_cons (s : fmap) p q h k v : - all (ord k) (supp s) -> h \In shape q s -> - entry p q k v \+ h \In shape p (ins k v s). + all (ord k) (supp s) -> + h \In shape q s -> + entry p q k v \+ h \In shape p (ins k v s). Proof. -move/all_path=>S H. +move/all_path_supp=>S H. suff: ins k v s = @FinMap _ _ ((k,v)::seq_of s) S by move=>->; exists q,h. rewrite fmapE /=; case: s {H}S=>/= xs ??. by rewrite last_ins'. Qed. (* inserting an entry with maximal key is appending to the list *) - Lemma shape_seg_rcons (s : fmap) x p q h k v : - (* conceptually last s < k *) - all (ord^~ k) (supp s) -> - h \In shape_seg x p s -> - (h \+ entry p q k v) \In shape_seg x q (ins k v s). + (* conceptually last s < k *) + all (ord^~ k) (supp s) -> + h \In shape_seg x p s -> + (h \+ entry p q k v) \In shape_seg x q (ins k v s). Proof. case: s=>xs; elim: xs h x=>/=. - move=>??? _ [->->]; rewrite unitL. @@ -135,12 +189,12 @@ by apply: IH=>//; apply: (path_sorted S). Qed. (* disjoint maps can be concatenated if the order is respected *) - Lemma shape_fcat s1 s2 h1 h2 x y : - (* conceptually last s1 < head s2 *) - allrel ord (supp s1) (supp s2) -> - h1 \In shape_seg x y s1 -> h2 \In shape y s2 -> - h1 \+ h2 \In shape x (fcat s1 s2). + (* conceptually last s1 < head s2 *) + allrel ord (supp s1) (supp s2) -> + h1 \In shape_seg x y s1 -> + h2 \In shape y s2 -> + h1 \+ h2 \In shape x (fcat s1 s2). Proof. move=>O1 H1. case: s2 O1=>xs; elim: xs h1 y h2 s1 H1=>/=. @@ -149,7 +203,7 @@ move=>[k' v'] xs; rewrite /fcat /= => IH /= h1 y h2 s1 H1 srt O2 H2. case: H2=>z[h'][-> H']; rewrite joinA; apply: IH; first 1 last. - by apply/path_sorted/srt. - move=>H0; rewrite (allrel_in_l (xs':=k'::supp s1) _); last by apply: supp_ins. - rewrite allrel_consl path_all //=. + rewrite allrel_consl order_path_min //=. by apply/allrel_sub_r/O2=>?; rewrite inE orbC=>->. - by move=>?; apply: H'. apply: shape_seg_rcons=>//. @@ -158,30 +212,30 @@ Qed. (* main procedures *) -(* a new map is just a null pointer *) - -Program Definition new : STsep (emp, [vfun x => shape x nil]) := +(* new map is just a null pointer *) +Program Definition new : + STsep (emp, [vfun x => shape x nil]) := Do (ret null). -Next Obligation. by move=>[] /= _ ->; step. Qed. +Next Obligation. by case=>_ /= ->; step. Qed. (* freeing a map is deallocating all its elements *) - +(* and the root pointer *) Definition freeT : Type := - forall p, {fm}, STsep (shape p fm, [vfun _ : unit => emp]). + forall p, STsep {fm} (shape p fm, [vfun _ : unit => emp]). Program Definition free : freeT := - Fix (fun (loop : freeT) p => - Do (if p == null then ret tt - else pnext <-- !(p.+2); - dealloc p;; - dealloc (p.+1);; - dealloc (p.+2);; - loop pnext)). + ffix (fun (loop : freeT) p => + Do (if p == null then ret tt + else pnext <-- !p.+2; + dealloc p;; + dealloc p.+1;; + dealloc p.+2;; + loop pnext)). Next Obligation. (* pull out ghost var, precondition, branch *) move=>loop p [fm][] i /= H; case: eqP=>[|/eqP] E. -- (* reached the end, heap must be empty *) - by apply: vrfV=>D; step=>_; rewrite E in H; case: (shape_null D H). +(* reached the end, heap must be empty *) +- by apply: vrfV=>W; step=>_; rewrite E in H; case: (shape_null W H). (* pull out the head entry *) case: (shape_cont E H)=>k[v][q][h][_ _ {i H}-> H]. (* deallocate it *) @@ -193,58 +247,58 @@ Qed. (* looking up an element in the map *) Definition lookupT k : Type := - forall p, {fm}, STsep (shape p fm, - [vfun y m => m \In shape p fm /\ y = fnd k fm]). + forall p, STsep {fm} (shape p fm, + [vfun y m => m \In shape p fm /\ y = fnd k fm]). Program Definition lookup x (k : K) : - {fm}, STsep (shape x fm, - [vfun y m => m \In shape x fm /\ y = fnd k fm]) := - Fix (fun (loop : lookupT k) (cur : ptr) => + STsep {fm} (shape x fm, + [vfun y m => m \In shape x fm /\ y = fnd k fm]) := + ffix (fun (loop : lookupT k) (cur : ptr) => Do (if cur == null then ret None else k' <-- !cur; if k == k' - then v <-- !(cur.+1); - ret (Some v) - else if ord k' k - then next <-- !(cur.+2); - loop next - else ret None)) x. + then v <-- !cur.+1; + ret (Some v) + else if ord k' k + then next <-- !cur.+2; + loop next + else ret None)) x. Next Obligation. (* pull out ghost var+precondition, branch on cur being null *) move=>_ k go cur [fm][] h /= H; case: eqP=>[|/eqP] Ec. -- (* reached the end, the heap and the spec must be empty, element not found *) - by apply: vrfV=>Vh; step=>_; rewrite Ec in H; case: (shape_null Vh H)=>->. -(* pull out the head entry *) +(* reached the end, heap and spec must be empty, element not found *) +- by apply: vrfV=>Vh; step=>_; rewrite Ec in H; case: (shape_null Vh H)=>->. +(* pull out head entry *) case: (shape_cont Ec H)=>k'[v][next][h'][Ef O' Ei H']; rewrite {h}Ei in H *. (* read the head key, branch on equality comparison *) step; case: eqP=>[|/eqP] Ek. -- (* the key matches, return the head value *) - by do 2![step]=>_; split=>//; rewrite Ef fnd_ins Ek eq_refl. +(* the key matches, return the head value *) +- by do 2![step]=>_; split=>//; rewrite Ef fnd_ins Ek eq_refl. (* branch on comparison *) case: ifP=>Ho'. -- (* head key is less than the needed one, loop on tail *) - (* (we fall back to gR to preserve associativity) *) - step; apply/[gR (behd fm)] @ h'=>//= v0 h0' [H0 E0] _. +(* head key is less than the needed one, loop on tail *) +(* (fall back to gR to preserve associativity) *) +- step; apply/[gR (behd fm)] @ h'=>//= v0 h0' [H0 E0] _. by rewrite Ef fnd_ins (negbTE Ek); split=>//; exact: shape_cons. (* head key is bigger than the needed one, abort *) -move: (semiconnex Ek); rewrite {}Ho' orbC /= =>Ho'; step=>_; split=>//. +move: (connex Ek); rewrite {}Ho' orbC /= =>Ho'; step=>_; split=>//. (* k is not in the head entry *) apply/esym/fnd_supp; rewrite Ef supp_ins inE negb_or; apply/andP; split=>//. (* nor it is in the tail *) -by apply/notin_path/path_le/all_path/O'. +by apply/notin_path/path_le/all_path_supp/O'. Qed. -(* removing an element by key from the map, return the pointer to the new map *) +(* removing element by key from the map, return the pointer to the new map *) (* loop invariant: *) -(* 1. we split the list into a zipper-like structure `fml ++ [k'->v'] ++ fmr` *) -(* 2. the ordering is respected *) +(* 1. split the list into a zipper-like structure `fml ++ [k'->v'] ++ fmr` *) +(* 2. ordering is respected *) (* 3. k is not in fml nor in the focus entry k'->v' *) -(* the focus is needed so that we can connect the remainder of the map to it after removal *) +(* the focus is needed to connect the remainder of the map to it after removal *) Definition removeT p k : Type := forall (prevcur : ptr * ptr), - {fm}, STsep (fun h => exists fml fmr k' v', + STsep {fm} (fun h => exists fml fmr k' v', [/\ fm = fcat (ins k' v' fml) fmr, all (ord^~ k') (supp fml) /\ all (ord k') (supp fmr), k \notin supp fml /\ k != k' & @@ -252,96 +306,103 @@ Definition removeT p k : Type := (shape_seg p prevcur.1 fml # (fun h => h = entry prevcur.1 prevcur.2 k' v') # shape prevcur.2 fmr)], - [vfun _ : unit => shape p (rem k fm)]). + [vfun _ : unit => shape p (rem k fm)]). Program Definition remove x (k : K) : - {fm}, STsep (shape x fm, - [vfun y => shape y (rem k fm)]) := - Do (let go := Fix (fun (loop : removeT x k) '(prev, cur) => + STsep {fm} (shape x fm, + [vfun y => shape y (rem k fm)]) := + Do (let go := ffix (fun (loop : removeT x k) '(prev, cur) => Do (if cur == null then ret tt else k' <-- !cur; if k == k' - then next <-- !(cur.+2); + then next <-- !cur.+2; dealloc cur;; - dealloc (cur.+1);; - dealloc (cur.+2);; + dealloc cur.+1;; + dealloc cur.+2;; prev.+2 ::= (next : ptr);; ret tt else if ord k' k - then next <-- !(cur.+2); + then next <-- !cur.+2; loop (cur, next) else ret tt)) in if x == null then ret null else k' <-- !x; if k == k' - then next <-- !(x.+2); + then next <-- !x.+2; dealloc x;; - dealloc (x.+1);; - dealloc (x.+2);; + dealloc x.+1;; + dealloc x.+2;; ret next else if ord k' k - then next <-- !(x.+2); + then next <-- !x.+2; go (x, next);; ret x else ret x). (* first the loop *) -(* we don't return the pointer because it cannot change, as the head is fixed by fml *) +(* don't return the pointer because it cannot change *) +(* as the head is fixed by fml *) Next Obligation. (* pull out ghost var, preconditions and heap validity *) -move=>x k0 go _ prev cur [_][] _ /= [fml][fmr][k][v][-> [Ol Or][El E]][hl][_][-> Hl [_][hr][->->Hr]]. +move=>x k0 go _ prev cur [_][] _ /= [fml][fmr][k][v][-> [Ol Or] +[El E]][hl][_][-> Hl [_][hr][->->Hr]]. (* branch on cur *) case: eqP=>[|/eqP] Ec. -(* cur = null - nothing to left to process, i.e., fmr = [] *) + (* cur = null - nothing to left to process, i.e., fmr = [] *) - apply: vrfV=>Vh; step=>_; rewrite {}Ec in Hr *. (* k is not in fml ++ (k->v) *) - case: (shape_null (validX Vh) Hr)=>/=->->; rewrite fcats0 unitR rem_ins (negbTE E) (rem_supp El). - by exact: shape_seg_rcons. + case: (shape_null (validX Vh) Hr)=>/=->->. + rewrite fcats0 unitR rem_ins (negbTE E) (rem_supp El). + exact: shape_seg_rcons. (* cur <> null, pull out the head entry from fmr *) -case: (shape_cont Ec Hr)=>k'[v'][next][hr'][Efr /all_path Or' {hr Hr Ec}-> Hr']; rewrite joinA joinC. +case: (shape_cont Ec Hr)=>k'[v'][next][hr'] +[Efr /all_path_supp Or' {hr Hr Ec}-> Hr']; rewrite joinA joinC. (* derive ordering facts *) -move/all_path: (Or); rewrite {1}Efr; case/(path_supp_ins_inv Or')/andP=>Ho' Or''. +move/all_path_supp: (Or); rewrite {1}Efr; +case/(path_supp_ins_inv Or')/andP=>Ho' Or''. (* get head key, branch on comparing it to needed one *) step; case: eqP=>[|/eqP] Ek. -- (* k = k' - element is found, run the deallocations *) - do 4!step; rewrite !unitL; do 2![step]=>_. + (* k = k' - element is found, run the deallocations *) +- do 4!step; rewrite !unitL; do 2![step]=>_. (* pull out fml ++ (k->v) *) rewrite Efr -fcat_srem; last by rewrite supp_ins inE negb_or E. (* drop the element in the spec *) - rewrite rem_ins {1}Ek eq_refl rem_supp; last by rewrite Ek; apply: notin_path. + rewrite rem_ins {1}Ek eq_refl rem_supp; + last by rewrite Ek; apply: notin_path. (* heap shape is respected *) rewrite joinC; apply/shape_fcat/Hr'; last by apply: shape_seg_rcons. (* the ordering is respected as well *) rewrite (allrel_in_l (xs':=k::supp fml) _); last by apply: supp_ins. - rewrite allrel_consl path_all //=. - by apply/(allrel_trans (z:=k))/path_all=>//; exact: trans. + rewrite allrel_consl order_path_min //=. + by apply/(allrel_trans (z:=k))/order_path_min=>//=. (* k <> k', now branch on order comparison *) case: ifP=>Ho0. -- (* k' is less than k, invoke the loop, shifting the focus to the right *) - step; apply: [gE (fcat (ins k' v' (ins k v fml)) (behd fmr))]=>//=. - - (* prove that all conditions are respected *) - exists (ins k v fml), (behd fmr), k', v'; do!split=>//. - - (* new focus comes after fml ++ old focus *) - rewrite (eq_all_r (s2:=k::supp fml)) /= ?Ho' /=; last by apply: supp_ins. +(* k' is less than k, invoke the loop, shifting the focus to the right *) +- step; apply: [gE (fcat (ins k' v' (ins k v fml)) (behd fmr))]=>//=. + (* prove that all conditions are respected *) + - exists (ins k v fml), (behd fmr), k', v'; do!split=>//. + (* new focus comes after fml ++ old focus *) + - rewrite (eq_all_r (s2:=k::supp fml)) /= ?Ho' /=; + last by apply: supp_ins. by apply/sub_all/Ol=>? /trans; apply. - - (* new focus comes before the new suffix *) - by apply: path_all. - - (* the needed key is not in fml ++ old focus *) - by rewrite supp_ins inE negb_or E. + (* new focus comes before the new suffix *) + - by apply: order_path_min. + (* the needed key is not in fml ++ old focus *) + - by rewrite supp_ins inE negb_or E. (* heap shape is respected *) exists (hl \+ entry prev cur k v), (entry cur next k' v' \+ hr'). rewrite joinC; split=>//; last by vauto. by apply: shape_seg_rcons. - (* we can reassemble the spec because insertions of old and new foci commute *) + (* reassemble the spec, as insertions of old and new foci commute *) move=>_ m Hm Vm; rewrite Efr. by rewrite fcat_inss // -?fcat_sins // in Hm; apply: notin_path. (* k' is bigger than k, abort *) -move: (semiconnex Ek); rewrite {}Ho0 orbC /= =>Ho0. +move: (connex Ek); rewrite {}Ho0 orbC /= =>Ho0. step=>_; rewrite rem_supp. - (* the shape is preserved *) rewrite joinC; apply: shape_fcat; first 1 last. - by apply: shape_seg_rcons. - - by rewrite Efr; apply: shape_cons=>//; apply: path_all. + - by rewrite Efr; apply: shape_cons=>//; apply: order_path_min. (* ordering is preserved *) rewrite (allrel_in_l (xs':=k::supp fml) _); last by apply: supp_ins. rewrite allrel_consl Or /=. @@ -365,32 +426,34 @@ case: (shape_cont Ex H)=>{Ex}k[v][next][h'][Ef Of Eh H']; rewrite Eh. step; case: eqP=>[->|/eqP Ek]. - (* k = k', element found in head position, run deallocations, return new head *) do 5![step]=>_; rewrite !unitL Ef rem_ins eq_refl rem_supp //. - by apply/notin_path/all_path. + by apply/notin_path/all_path_supp. (* k <> k', now branch on order comparison *) case: ifP=>Ho0. (* k' is less than k, start the loop with the head entry as the focus *) - step; apply: [stepE fm]=>//=; last by move=>_ ??; step. (* invariants and shape are satisfied *) exists nil, (behd fm), k, v; do!split=>//. - - by rewrite fcat_inss; [rewrite fcat0s|apply/notin_path/all_path]. + - by rewrite fcat_inss; [rewrite fcat0s|apply/notin_path/all_path_supp]. by exists Unit, (entry x next k v \+ h'); split=>//; [rewrite unitL | vauto]. (* k' is bigger than k, abort *) -move: (semiconnex Ek); rewrite {}Ho0 orbC /= =>Ho0. +move: (connex Ek); rewrite {}Ho0 orbC /= =>Ho0. (* return the old head, invariants are preserved *) step=>_; rewrite -Eh rem_supp // Ef supp_ins !inE negb_or Ek /=. -by apply/notin_path/path_le/all_path/Of. +by apply/notin_path/path_le/all_path_supp/Of. Qed. -(* upserting (inserting or updating if the key is found) an entry into the map, return the pointer to the new map *) +(* upserting (inserting or updating if the key is found) *) +(* an entry into the map, return the pointer to the new map *) (* loop invariant, essentially identical to remove: *) -(* 1. as for remove, we split the list into a zipper-like structure `fml ++ [k'->v'] ++ fmr` *) +(* 1. as in remove, split the list into a zipper-like structure *) +(* `fml ++ [k'->v'] ++ fmr` *) (* 2. the ordering is respected *) (* 3. k is not in fml and is less than the key k' of focus entry *) -(* the focus is needed so that we can connect the new element to it after insertion *) +(* the focus is needed to connect the new element after insertion *) Definition insertT p k v : Type := forall (prevcur : ptr * ptr), - {fm}, STsep (fun h => exists fml fmr k' v', + STsep {fm} (fun h => exists fml fmr k' v', [/\ fm = fcat (ins k' v' fml) fmr, all (ord^~ k') (supp fml) /\ all (ord k') (supp fmr), k \notin supp fml /\ ord k' k & @@ -398,12 +461,12 @@ Definition insertT p k v : Type := (shape_seg p prevcur.1 fml # (fun h => h = entry prevcur.1 prevcur.2 k' v') # shape prevcur.2 fmr)], - [vfun _ : unit => shape p (ins k v fm)]). + [vfun _ : unit => shape p (ins k v fm)]). Program Definition insert x (k : K) (v : V) : - {fm}, STsep (shape x fm, - [vfun y => shape y (ins k v fm)]) := - Do (let go := Fix (fun (loop : insertT x k v) '(prev, cur) => + STsep {fm} (shape x fm, + [vfun y => shape y (ins k v fm)]) := + Do (let go := ffix (fun (loop : insertT x k v) '(prev, cur) => Do (if cur == null then new <-- allocb k 3; new.+1 ::= v;; @@ -415,7 +478,7 @@ Program Definition insert x (k : K) (v : V) : then cur.+1 ::= v;; ret tt else if ord k' k - then next <-- !(cur.+2); + then next <-- !cur.+2; loop (cur, next) else new <-- allocb k 3; new.+1 ::= v;; @@ -433,7 +496,7 @@ Program Definition insert x (k : K) (v : V) : then x.+1 ::= v;; ret x else if ord k' k - then next <-- !(x.+2); + then next <-- !x.+2; go (x, next);; ret x else new <-- allocb k 3; @@ -441,15 +504,17 @@ Program Definition insert x (k : K) (v : V) : new.+2 ::= x;; ret new). (* first the loop *) -(* we don't return the pointer because it cannot change, as the head is fixed by fml *) +(* don't return the pointer because it cannot change *) +(* as the head is fixed by fml *) Next Obligation. (* pull out ghost var+preconditions *) -move=>x k0 v0 loop _ prev cur [_][] _ /= [fml][fmr][k][v][-> [Ol Or][El Ho0]][hl][_][-> Hl [_][hr][->-> Hr]]. +move=>x k0 v0 loop _ prev cur [_][] _ /= [fml][fmr][k][v] +[-> [Ol Or][El Ho0]][hl][_][-> Hl [_][hr][->-> Hr]]. (* branch on cur *) case: eqP=>[|/eqP] Ec. -- (* cur = null, insert as the last element *) - rewrite {}Ec in Hr; apply: vrfV=>Vh. - step=>p; rewrite ptrA unitR; do 2!step; rewrite joinC; do 2![step]=>_. +(* cur = null, insert as the last element *) +- rewrite {}Ec in Hr; apply: vrfV=>Vh. + step=>p; rewrite unitR; do 2!step; rewrite joinC; do 2![step]=>_. (* fmr is empty *) case: (shape_null (validX Vh) Hr)=>/=->->. rewrite fcats0 unitR [X in _ \+ entry _ _ _ _ \+ X]joinA. @@ -458,13 +523,16 @@ case: eqP=>[|/eqP] Ec. rewrite (eq_all_r (s2:=k::supp fml)) /= ?Ho0 /=; last by apply: supp_ins. by apply/sub_all/Ol=>? /trans; apply. (* cur <> null, pull out the head entry from fmr *) -case: (shape_cont Ec Hr)=>k'[v'][next][hr'][Efr Or' {hr Hr Ec}-> Hr']; rewrite joinA joinC. +case: (shape_cont Ec Hr)=>k'[v'][next][hr'][Efr Or' {hr Hr Ec}-> Hr']. +rewrite joinA joinC. (* derive ordering facts *) -move/all_path: (Or); rewrite {1}Efr; case/(path_supp_ins_inv (all_path Or'))/andP=>Ho' /(path_all (@trans _)) Or''. +move/all_path_supp: (Or); rewrite {1}Efr. +case/(path_supp_ins_inv (all_path_supp Or'))/andP=>Ho'. +move/(order_path_min (@trans _))=>Or''. (* get new key, branch on equality comparison *) step; case: eqP=>[|/eqP] Ek. -- (* k = k', update the value at this position *) - do 2![step]=>_; rewrite Efr -fcat_sins ins_ins -Ek eq_refl joinC. +(* k = k', update the value at this position *) +- do 2![step]=>_; rewrite Efr -fcat_sins ins_ins -Ek eq_refl joinC. (* invariants are preserved as the key didn't change *) apply: shape_fcat; first 1 last. - by apply: shape_seg_rcons. @@ -478,31 +546,33 @@ step; case: eqP=>[|/eqP] Ek. case: ifP=>Ho'0. (* k' is less than k, invoke the loop, shifting the focus to the right *) - step; apply/[gE (fcat (ins k' v' (ins k v fml)) (behd fmr))]=>//=. - - (* prove that all conditions are respected *) - exists (ins k v fml), (behd fmr), k', v'; do!split=>//. - - (* new focus comes after fml ++ old focus *) - rewrite (eq_all_r (s2:=k::supp fml)) /= ?Ho' /=; last by apply: supp_ins. + (* prove that all conditions are respected *) + - exists (ins k v fml), (behd fmr), k', v'; do!split=>//. + (* new focus comes after fml ++ old focus *) + - rewrite (eq_all_r (s2:=k::supp fml)) /= ?Ho' /=; last by apply: supp_ins. by apply/sub_all/Ol=>? /trans; apply. - - (* the needed key is not in fml ++ old focus *) - rewrite supp_ins inE negb_or andbC El /=. + (* the needed key is not in fml ++ old focus *) + - rewrite supp_ins inE negb_or andbC El /=. by case: ordP Ho0. (* heap shape is respected *) exists (hl \+ entry prev cur k v), (entry cur next k' v' \+ hr'). by rewrite joinC; split=>//; [apply: shape_seg_rcons | vauto]. - (* we can reassemble the spec because insertions of old and new foci commute *) + (* reassemble the spec, as insertions of old and new foci commute *) move=>_ m Hm _; rewrite Efr. rewrite fcat_inss // in Hm; first by rewrite -fcat_sins in Hm. - by apply/notin_path/all_path. + by apply/notin_path/all_path_supp. (* k' is bigger than k, insert at this position *) -move: (semiconnex Ek); rewrite {}Ho'0 orbC /= =>Ho0'. +move: (connex Ek); rewrite {}Ho'0 orbC /= =>Ho0'. (* run the allocation and assignments *) -step=>new; rewrite ptrA unitR; do 2!step; rewrite joinA joinC; do 2![step]=>_. +step=>new; rewrite unitR; do 2!step; rewrite joinA joinC; do 2![step]=>_. rewrite Efr -fcat_sins; apply: shape_fcat; first 1 last. -- (* shape is respected for the prefix fml ++ old focus *) - by apply: shape_seg_rcons. -- (* shape is satisfed for new element + suffix *) - rewrite [X in X \+ (entry _ _ _ _ \+ _)]joinA; apply/shape_cons/shape_cons=>//. - by apply/path_all/path_supp_ins=>//; apply/path_le/all_path/Or'. +(* shape is respected for the prefix fml ++ old focus *) +- by apply: shape_seg_rcons. +(* shape is satisfed for new element + suffix *) +- rewrite [X in X \+ (entry _ _ _ _ \+ _)]joinA. + apply/shape_cons/shape_cons=>//. + apply/order_path_min=>//; apply/path_supp_ins=>//. + by apply/path_le/all_path_supp/Or'. (* ordering is respected *) rewrite (allrel_in_l (xs':=k::supp fml) _); last by apply: supp_ins. rewrite (allrel_in_r (ys':=k0::k'::supp (behd fmr)) _ _); last first. @@ -510,7 +580,7 @@ rewrite (allrel_in_r (ys':=k0::k'::supp (behd fmr)) _ _); last first. rewrite allrel_consl !allrel_consr /= Ho0 Ho' Or'' /=; apply/and3P; split. - by apply/sub_all/Ol=>? /trans; apply. - by apply/sub_all/Ol=>? /trans; apply. -by apply: (allrel_trans (z:=k))=>//; exact: trans. +by apply: (allrel_trans (z:=k))=>//; apply: trans. Qed. (* now the outer function, which mostly repeats the loop *) (* the first iteration is special because we don't yet have a left prefix *) @@ -518,38 +588,161 @@ Qed. Next Obligation. (* pull out ghost+precondition, branch on x *) move=>/= x k0 v0 [fm][]h /= H; case: eqP=>[|/eqP] Ex. -- (* x = null, insert as the only element, heap and spec are empty *) - apply: vrfV=>Vh; move: H; rewrite Ex=>/(shape_null Vh) [->->]. +(* x = null, insert as the only element, heap and spec are empty *) +- apply: vrfV=>Vh; move: H; rewrite Ex=>/(shape_null Vh) [->->]. (* run *) - step=>p; rewrite ptrA !unitR; do 3![step]=>_. + step=>p; rewrite !unitR; do 3![step]=>_. by exists null, Unit; rewrite unitR joinA. (* x <> null, pull out the head entry *) case: (shape_cont Ex H)=>{Ex}k[v][next][h'][Ef Of {h H}-> H]. (* read the head key, branch on equality comparison *) step; case: eqP=>[->|/eqP Ek]. -- (* k = k', exact key found in head position, update the head value *) - do 2![step]=>_; rewrite Ef ins_ins eq_refl. +(* k = k', exact key found in head position, update the head value *) +- do 2![step]=>_; rewrite Ef ins_ins eq_refl. by apply: shape_cons. (* k <> k', now branch on order comparison *) case: ifP=>Ho0. -- (* k' is less than k, run the loop with the head entry as the focus *) - step; apply: [stepE fm]=>//=; last by move=>_ ??; step. +(* k' is less than k, run the loop with the head entry as the focus *) +- step; apply: [stepE fm]=>//=; last by move=>_ ??; step. (* invariants are respected *) exists nil, (behd fm), k, v; do!split=>//. - - by rewrite fcat_inss; [rewrite fcat0s|apply/notin_path/all_path]. - by exists Unit, (entry x next k v \+ h'); split=>//; [rewrite unitL | vauto]. + - by rewrite fcat_inss; [rewrite fcat0s|apply/notin_path/all_path_supp]. + by exists Unit, (entry x next k v \+ h'); split=>//; [rewrite unitL|vauto]. (* k' is bigger than k, insert after the head *) -move: (semiconnex Ek); rewrite {}Ho0 orbC /= =>Ho0. +move: (connex Ek); rewrite {}Ho0 orbC /= =>Ho0. (* run allocation and assignments, return the old head *) -step=>q; rewrite ptrA unitR; do 3![step]=>_. +step=>q; rewrite unitR; do 3![step]=>_. (* invariants are preserved *) -rewrite Ef [X in X \+ (entry _ _ _ _ \+ _)]joinA; apply/shape_cons/shape_cons=>//. -by apply/path_all/path_supp_ins=>//; apply/path_le/all_path/Of. +rewrite Ef [X in X \+ (entry _ _ _ _ \+ _)]joinA. +apply/shape_cons/shape_cons=>//. +apply/order_path_min=>//. +apply/path_supp_ins=>//. +by apply/path_le/all_path_supp/Of. Qed. -(* ordered association list is a KV map *) +End AssocList. +End DAList. + +(* association list is dynamic KV map *) +Notation dalist := DAList.root. +HB.instance Definition _ K V := + isDKVM.Build K V (dalist K V) DAList.null_root + DAList.new DAList.free DAList.insert DAList.remove DAList.lookup. + +(***************************************) +(* Association maps with KVM interface *) +(***************************************) + +(* static variant packages the root pointer *) +(* with the dynamic part of the structure *) + +(* module type for making the defs opaque *) +Module Type AList_sig. +Parameter root : ordType -> Type -> Set. +Section AList. +Context {K : ordType} {V : Type}. +Parameter null_root : root K V. +Parameter shape : root K V -> {finMap K -> V} -> Pred heap. +Parameter new : STsep (emp, [vfun (x : root K V) h => h \In shape x nil]). +Parameter free : forall (x : root K V), + STsep {s} (shape x s, [vfun _ : unit => emp]). +Parameter insert : forall (x : root K V) k v, + STsep {s} (shape x s, [vfun (_ : unit) h => h \In shape x (ins k v s)]). +Parameter remove : forall (x : root K V) k, + STsep {s} (shape x s, [vfun (_ : unit) h => h \In shape x (rem k s)]). +Parameter lookup : forall (x : root K V) k, + STsep {s} (shape x s, [vfun y m => m \In shape x s /\ y = fnd k s]). +End AList. +End AList_sig. + +Module AList : AList_sig. +Section AssocList. +Definition root : ordType -> Type -> Set := fun _ _ => ptr. +Variables (K : ordType) (V : Type). +Notation fmap := (finMap K V). +Definition null_root : root K V := null. + +Definition shape (x : ptr) (f : fmap) : Pred heap := + [Pred h | exists (a : dalist K V) h', + h = x :-> a \+ h' /\ dkvm_shape a f h']. + +(* new structure, but the root pointer is given *) +Program Definition new0 (x : ptr) : + STsep {v : dalist K V} (fun h => h = x :-> v, + [vfun _ : unit => shape x nil]) := + Do (a <-- dkvm_new (dalist K V); + x ::= a). +Next Obligation. +move=>x [a][/= _ ->]; apply: [stepU]=>//= b h H. +by step; exists b, h; rewrite joinC. +Qed. -Definition AssocList := KVmap.make null new free insert remove lookup. +Program Definition new : + STsep (emp, [vfun x => shape x nil]) := + Do (a <-- dkvm_new (dalist K V); + alloc a). +Next Obligation. +case=>_ /= ->; apply: [stepU]=>//= a h H. +by step=>x; exists a, h; rewrite unitR. +Qed. +(* free structure, keep the root pointer *) +Program Definition free0 x : + STsep {s} (shape x s, [vfun (_ : unit) h => + exists a : dalist K V, h = x :-> a]) := + Do (a <-- !x; + dkvm_free (a : dalist K V)). +Next Obligation. +move=>x [fm][/= _ [a][h][-> H]]; step. +by apply: [gX fm]@h=>//= _ _ -> _; rewrite unitR; eauto. +Qed. + +Program Definition free x : + STsep {s} (shape x s, + [vfun _ : unit => emp]) := + Do (a <-- !x; + dealloc x;; + dkvm_free (a : dalist K V)). +Next Obligation. +by move=>x [fm][/= _ [a][h][-> H]]; step; step; apply: [gX fm]@h. +Qed. + +Program Definition insert x k v : + STsep {s} (shape x s, + [vfun (_ : unit) h => h \In shape x (ins k v s)]) := + Do (a <-- !x; + y <-- dkvm_insert (a : dalist K V) k v; + x ::= y). +Next Obligation. +move=>x k v [fm][/= _ [a][h][-> H]]; step. +by apply: [stepX fm]@h=>//= y {}h {}H; step; hhauto. +Qed. + +Program Definition remove x k : + STsep {s} (shape x s, + [vfun _ : unit => shape x (rem k s)]) := + Do (a <-- !x; + y <-- dkvm_remove (a : dalist K V) k; + x ::= y). +Next Obligation. +move=>x k [fm][/= _ [a][h][-> H]]; step. +by apply: [stepX fm]@h=>//= y {}h {}H; step; hhauto. +Qed. + +Program Definition lookup x k : + STsep {s} (shape x s, + [vfun y m => m \In shape x s /\ y = fnd k s]) := + Do (a <-- !x; + dkvm_lookup (a : dalist K V) k). +Next Obligation. +move=>x k [fm][/= _ [a][h][-> H]]; step. +by apply: [gX fm]@h=>//= y {}h {H}[]; hhauto. +Qed. End AssocList. -End AssocList. +End AList. + +(* association list is (static) KV map *) +Notation alist := AList.root. +HB.instance Definition _ K V := + isKVM.Build K V (alist K V) AList.null_root + AList.new AList.free AList.insert AList.remove AList.lookup. diff --git a/examples/llist.v b/examples/llist.v index 87eca58..95682de 100644 --- a/examples/llist.v +++ b/examples/llist.v @@ -1,3 +1,15 @@ +(* +Copyright 2010 IMDEA Software Institute +Licensed under the Apache License, Version 2.0 (the "License"); +you may not use this file except in compliance with the License. +You may obtain a copy of the License at + http://www.apache.org/licenses/LICENSE-2.0 +Unless required by applicable law or agreed to in writing, software +distributed under the License is distributed on an "AS IS" BASIS, +WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. +See the License for the specific language governing permissions and +limitations under the License. +*) From mathcomp Require Import ssreflect ssrbool ssrnat eqtype ssrfun seq. From pcm Require Import options axioms pred. @@ -7,11 +19,13 @@ From htt Require Import options model heapauto. (* Linked lists, storing a value and next pointer in consecutive locations. *) (* We start with a more general "segment" definition, where the last node's *) (* next pointer is an arbitrary q *) +(* NOTE: already defined in heap.v under the name llist, but repeated here *) +(* to make the file self-contained. *) Fixpoint lseg {A} (p q : ptr) (xs : seq A) := if xs is hd::tl then [Pred h | exists r h', - h = p :-> hd \+ (p .+ 1 :-> r \+ h') /\ h' \In lseg r q tl] + h = p :-> hd \+ (p.+1 :-> r \+ h') /\ h' \In lseg r q tl] else [Pred h | p = q /\ h = Unit]. Definition EmptyList : exn := exn_from_nat 15. @@ -19,47 +33,47 @@ Definition EmptyList : exn := exn_from_nat 15. Section LSeg. Variable A : Type. -(* appending a value to the list segment *) - +(* appending value to list segment *) Lemma lseg_rcons (xs : seq A) x p r h : h \In lseg p r (rcons xs x) <-> - exists q h', h = h' \+ (q :-> x \+ q .+ 1 :-> r) /\ h' \In lseg p q xs. + exists q h', h = h' \+ (q :-> x \+ q.+1 :-> r) /\ + h' \In lseg p q xs. Proof. move: xs x p r h; elim=>[|x xs IH] y p r h /=. - by split; case=>x [h'][->][<- ->]; [exists p | exists r]; hhauto. split. - case=>z [h1][->]; case/IH=>w [h2][->] H1. - by exists w, (p :-> x \+ (p .+ 1 :-> z \+ h2)); hhauto. + by exists w, (p :-> x \+ (p.+1 :-> z \+ h2)); hhauto. case=>q [h1][->][z][h2][->] H1. -exists z, (h2 \+ q :-> y \+ q .+ 1 :-> r). +exists z, (h2 \+ q :-> y \+ q.+1 :-> r). by rewrite -!joinA; split=>//; apply/IH; eauto. Qed. -(* null pointer represents an empty segment *) - +(* null pointer represents empty segment *) Lemma lseg_null (xs : seq A) q h : - valid h -> h \In lseg null q xs -> + valid h -> + h \In lseg null q xs -> [/\ q = null, xs = [::] & h = Unit]. Proof. case: xs=>[|x xs] D /= H; first by case: H=><- ->. case: H D=>r [h'][->] _; rewrite validPtUn; hhauto. Qed. -(* empty heap represents an empty segment *) - -Lemma lseg_empty (xs : seq A) p q : Unit \In lseg p q xs -> p = q /\ xs = [::]. +(* empty heap represents empty segment *) +Lemma lseg_empty (xs : seq A) p q : + Unit \In lseg p q xs -> + p = q /\ xs = [::]. Proof. -by case: xs=>[|x xs][] // r [h][/esym/join0E][/unitbE]; rewrite /heap_pts ptsU um_unitbU. +by case: xs=>[|x xs][] //= r [h][/esym/umap0E][/unitbP]; rewrite um_unitbU. Qed. (* reformulation of the specification *) - Lemma lseg_case (xs : seq A) p q h : h \In lseg p q xs -> [/\ p = q, xs = [::] & h = Unit] \/ exists x r h', [/\ xs = x :: behead xs, - h = p :-> x \+ (p .+ 1 :-> r \+ h') & + h = p :-> x \+ (p.+1 :-> r \+ h') & h' \In lseg r q (behead xs)]. Proof. case: xs=>[|x xs] /=; first by case=>->->; left. @@ -67,12 +81,11 @@ by case=>r [h'][->] H; right; hhauto. Qed. (* non-trivial segment represents a non-empty list *) - -Corollary lseg_neq (xs : seq A) p q h : +Lemma lseg_neq (xs : seq A) p q h : p != q -> h \In lseg p q xs -> exists x r h', [/\ xs = x :: behead xs, - h = p :-> x \+ (p .+ 1 :-> r \+ h') & + h = p :-> x \+ (p.+1 :-> r \+ h') & h' \In lseg r q (behead xs)]. Proof. move=>H /lseg_case; case=>//; case=>E. @@ -80,12 +93,11 @@ by rewrite E eq_refl in H. Qed. (* non-empty list is represented by a non-trivial segment *) - -Corollary lseg_lt0n (xs : seq A) p q h : +Lemma lseg_lt0n (xs : seq A) p q h : 0 < size xs -> h \In lseg p q xs -> exists x r h', [/\ xs = x :: behead xs, - h = p :-> x \+ (p .+ 1 :-> r \+ h') & + h = p :-> x \+ (p.+1 :-> r \+ h') & h' \In lseg r q (behead xs)]. Proof. move=>H /lseg_case; case=>//; case=>_ E. @@ -93,7 +105,6 @@ by rewrite E in H. Qed. (* splitting the list corresponds to splitting the heap *) - Lemma lseg_cat (xs ys : seq A) p q h : h \In lseg p q (xs++ys) <-> exists j, h \In lseg p j xs # lseg j q ys. @@ -103,7 +114,7 @@ elim: xs h p=>/=. by case=>j[_][h2][{h}-> [->->]]; rewrite unitL. move=>x xs IH h p; split. - case=>r[_][{h}-> /IH][j][h1][h2][-> H1 H2]. - exists j, (p :-> x \+ p.+ 1 :-> r \+ h1), h2; rewrite !joinA; split=>//. + exists j, (p :-> x \+ p.+1 :-> r \+ h1), h2; rewrite !joinA; split=>//. by exists r, h1; rewrite joinA. case=>j[_][h2][{h}-> [r][h1][-> H1 H2]]. exists r, (h1 \+ h2); rewrite !joinA; split=>//. @@ -121,21 +132,27 @@ Variable (A : Type). (* specializing the null and neq lemmas *) Lemma lseq_null (xs : seq A) h : - valid h -> h \In lseq null xs -> xs = [::] /\ h = Unit. + valid h -> + h \In lseq null xs -> + xs = [::] /\ h = Unit. Proof. by move=>D; case/(lseg_null D)=>_ ->. Qed. Lemma lseq_pos (xs : seq A) p h : - p != null -> h \In lseq p xs -> + p != null -> + h \In lseq p xs -> exists x r h', [/\ xs = x :: behead xs, - h = p :-> x \+ (p .+ 1 :-> r \+ h') & + h = p :-> x \+ (p.+1 :-> r \+ h') & h' \In lseq r (behead xs)]. Proof. by apply: lseg_neq. Qed. -(* a valid heap cannot match two different specs *) +(* valid heap cannot match two different specs *) Lemma lseq_func (l1 l2 : seq A) p h : - valid h -> h \In lseq p l1 -> h \In lseq p l2 -> l1 = l2. + valid h -> + h \In lseq p l1 -> + h \In lseq p l2 -> + l1 = l2. Proof. elim: l1 l2 p h => [|x1 xt IH] /= l2 p h V. - by case=>->->; case/lseq_null. @@ -154,9 +171,8 @@ Program Definition new : STsep (emp, [vfun x => lseq x (@nil A)]) := Next Obligation. by move=>[] /= _ ->; step. Qed. (* prepending a value *) - Program Definition insert p (x : A) : - {l}, STsep (lseq p l, [vfun p' => lseq p' (x::l)]) := + STsep {l} (lseq p l, [vfun p' => lseq p' (x::l)]) := Do (q <-- allocb p 2; q ::= x;; ret q). @@ -164,14 +180,13 @@ Next Obligation. (* pull out ghost var + precondition, run the first step *) move=>p x [l][] i /= H; step=>q. (* run the last 2 steps, guess the final pointer and heap from the goal *) -rewrite unitR -joinA; heval. +by rewrite unitR -joinA; heval. Qed. (* getting the head element *) (* an example of a partial program, doesn't modify the heap *) - Program Definition head p : - {l}, STsep (lseq p l, + STsep {l} (lseq p l, fun (y : ans A) h => h \In lseq p l /\ match y with Val v => l = v :: behead l | Exn e => e = EmptyList /\ l = [::] end) := @@ -191,13 +206,12 @@ by do 2![step]=>_; split=>//; rewrite E; hhauto. Qed. (* removing the head element, no-op for an empty list *) - Program Definition remove p : - {xs : seq A}, STsep (lseq p xs, [vfun p' => lseq p' (behead xs)]) := + STsep {xs : seq A} (lseq p xs, [vfun p' => lseq p' (behead xs)]) := Do (if p == null then ret p - else pnext <-- !(p .+ 1); + else pnext <-- !p.+1; dealloc p;; - dealloc p .+ 1;; + dealloc p.+1;; ret pnext). Next Obligation. (* pull out ghost + precondition, branch *) @@ -211,20 +225,20 @@ Qed. (* calculating the list length *) -(* the loop invariant: *) +(* loop invariant: *) (* 1. heap is unchanged *) (* 2. total length is accumulator + the length of unprocessed list *) Definition lenT : Type := forall (pl : ptr * nat), - {xs : seq A}, STsep (lseq pl.1 xs, - [vfun l h => l == pl.2 + length xs /\ lseq pl.1 xs h]). + STsep {xs : seq A} (lseq pl.1 xs, + [vfun l h => l == pl.2 + length xs /\ lseq pl.1 xs h]). Program Definition len p : - {xs : seq A}, STsep (lseq p xs, - [vfun l h => l == length xs /\ lseq p xs h]) := - Do (let len := Fix (fun (go : lenT) '(p, l) => - Do (if p == null then ret l - else pnext <-- !(p .+ 1); - go (pnext, l + 1))) + STsep {xs : seq A} (lseq p xs, + [vfun l h => l == length xs /\ lseq p xs h]) := + Do (let len := ffix (fun (go : lenT) '(p, l) => + Do (if p == null then ret l + else pnext <-- !p.+1; + go (pnext, l + 1))) in len (p, 0)). (* first, the loop *) Next Obligation. @@ -244,22 +258,22 @@ Qed. (* concatenation: modifies the first list, returning nothing *) -(* the loop invariant: *) -(* the first list should not be empty and not overlap the second *) +(* loop invariant: *) +(* first list isn't empty and doesn't overlap with the second *) Definition catT (p2 : ptr) : Type := - forall (p1 : ptr), {xs1 xs2 : seq A}, - STsep (fun h => p1 != null /\ (lseq p1 xs1 # lseq p2 xs2) h, - [vfun _ : unit => lseq p1 (xs1 ++ xs2)]). + forall (p1 : ptr), STsep {xs1 xs2 : seq A} + (fun h => p1 != null /\ (lseq p1 xs1 # lseq p2 xs2) h, + [vfun _ : unit => lseq p1 (xs1 ++ xs2)]). Program Definition concat p1 p2 : - {xs1 xs2 : seq A}, STsep (lseq p1 xs1 # lseq p2 xs2, - [vfun a => lseq a (xs1 ++ xs2)]) := - Do (let cat := Fix (fun (go : catT p2) q => - Do (next <-- !(q .+ 1); - if (next : ptr) == null - then q .+ 1 ::= p2;; - ret tt - else go next)) + STsep {xs1 xs2 : seq A} (lseq p1 xs1 # lseq p2 xs2, + [vfun a => lseq a (xs1 ++ xs2)]) := + Do (let cat := ffix (fun (go : catT p2) q => + Do (next <-- !q.+1; + if (next : ptr) == null + then q.+1 ::= p2;; + ret tt + else go next)) in if p1 == null then ret p2 else cat p1;; @@ -299,20 +313,20 @@ Qed. (* in-place reversal by pointer swinging *) -(* the loop invariant: *) -(* 1. the processed and remaining parts should not overlap *) -(* 2. the result is processed part + a reversal of remainder *) +(* loop invariant: *) +(* 1. processed and remaining parts don't overlap *) +(* 2. result = processed part + reversal of remainder *) Definition revT : Type := forall (p : ptr * ptr), - {i done : seq A}, STsep (lseq p.1 i # lseq p.2 done, - [vfun y => lseq y (catrev i done)]). + STsep {i done : seq A} (lseq p.1 i # lseq p.2 done, + [vfun y => lseq y (catrev i done)]). Program Definition reverse p : - {xs : seq A}, STsep (lseq p xs, [vfun p' => lseq p' (rev xs)]) := - Do (let reverse := Fix (fun (go : revT) '(i, done) => - Do (if i == null then ret done - else next <-- !i .+ 1; - i .+ 1 ::= done;; - go (next, i))) + STsep {xs : seq A} (lseq p xs, [vfun p' => lseq p' (rev xs)]) := + Do (let reverse := ffix (fun (go : revT) '(i, done) => + Do (if i == null then ret done + else next <-- !i.+1; + i.+1 ::= done;; + go (next, i))) in reverse (p, null)). (* first, the loop *) Next Obligation. @@ -341,16 +355,16 @@ Variable B : Type. (* the result is a mapped list *) Definition lmapT (f : A -> B) := forall (p : ptr), - {xs : seq A}, STsep (lseq p xs, + STsep {xs : seq A} (lseq p xs, [vfun _ : unit => lseq p (map f xs)]). Program Definition lmap f : lmapT f := - Fix (fun (lmap : lmapT f) p => + ffix (fun (lmap : lmapT f) p => Do (if p == null then ret tt else t <-- !p; p ::= f t;; - nxt <-- !p .+ 1; + nxt <-- !p.+1; lmap nxt)). Next Obligation. (* pull out ghost + precondition, branch *) @@ -365,5 +379,5 @@ Qed. End LList. -Arguments head {A} p. -Arguments remove {A} p. +Arguments head {A}. +Arguments remove {A}. diff --git a/examples/queue.v b/examples/queue.v old mode 100755 new mode 100644 index 3007251..2df546c --- a/examples/queue.v +++ b/examples/queue.v @@ -1,8 +1,21 @@ +(* +Copyright 2021 IMDEA Software Institute +Licensed under the Apache License, Version 2.0 (the "License"); +you may not use this file except in compliance with the License. +You may obtain a copy of the License at + http://www.apache.org/licenses/LICENSE-2.0 +Unless required by applicable law or agreed to in writing, software +distributed under the License is distributed on an "AS IS" BASIS, +WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. +See the License for the specific language governing permissions and +limitations under the License. +*) + From Coq Require Import ssreflect ssrbool ssrfun. -From mathcomp Require Import eqtype seq. +From mathcomp Require Import eqtype ssrnat seq. From pcm Require Import options axioms pred. From pcm Require Import pcm unionmap heap automap. -From htt Require Import model heapauto. +From htt Require Import options model heapauto. From htt Require Import llist. Record queue (T : Type) : Type := Queue {front: ptr; back: ptr}. @@ -13,17 +26,18 @@ Section Queue. Variable T : Type. Notation queue := (queue T). -(* a queue is specified as a singly-linked list split into an initial segment and a last node *) +(* queue is singly-linked list split into *) +(* the initial segment and the last node *) Definition is_queue (fr bq : ptr) (xs : seq T) := if fr == null then [Pred h | [/\ bq = null, xs = [::] & h = Unit]] else [Pred h | exists xt x h', [/\ xs = rcons xt x, - valid (h' \+ (bq :-> x \+ bq .+ 1 :-> null)), - h = h' \+ (bq :-> x \+ bq .+ 1 :-> null) & + valid (h' \+ (bq :-> x \+ bq.+1 :-> null)), + h = h' \+ (bq :-> x \+ bq.+1 :-> null) & h' \In lseg fr bq xt]]. -(* the structure itself is a pair of pointers to body + last node *) -(* insertion happens at the last node, and removal at the head *) +(* queue structure is a pair of pointers to body + last node *) +(* insertion happens at the last node, removal at the head *) Definition shape (q : queue) (xs : seq T) := [Pred h | exists fr bq h', [/\ valid (front q :-> fr \+ (back q :-> bq \+ h')), @@ -31,24 +45,22 @@ Definition shape (q : queue) (xs : seq T) := h' \In is_queue fr bq xs]]. (* well-formed queue is a valid heap *) - Lemma shapeD q xs h : h \In shape q xs -> valid h. Proof. by case=>h1[bq][h'] [] D ->. Qed. (* empty queue is a pair of null pointers *) - Lemma is_queue_nil fr bq h : - h \In is_queue fr bq [::] -> [/\ fr = null, bq = null & h = Unit]. + h \In is_queue fr bq [::] -> + [/\ fr = null, bq = null & h = Unit]. Proof. by rewrite /is_queue; case: eqP=>[->[-> _ ->] | _ [[|y xt][x][h'][]]]. Qed. (* restructuring the specification for combined list *) - Lemma is_queue_rcons fr bq xt x h : h \In is_queue fr bq (rcons xt x) <-> - (exists h', [/\ valid (h' \+ (bq :-> x \+ bq .+ 1 :-> null)), - h = h' \+ (bq :-> x \+ bq .+ 1 :-> null) & + (exists h', [/\ valid (h' \+ (bq :-> x \+ bq.+1 :-> null)), + h = h' \+ (bq :-> x \+ bq.+1 :-> null) & h' \In lseg fr bq xt]). Proof. rewrite /is_queue; split. @@ -59,10 +71,10 @@ move: (D)=>/[swap]; case/(lseg_null (validL D))=>->->->. by rewrite unitL validPtUn. Qed. -(* pointers should agree in a well-formed queue *) - +(* pointers agree in a well-formed queue *) Lemma backfront fr bq xs h : - h \In is_queue fr bq xs -> (fr == null) = (bq == null). + h \In is_queue fr bq xs -> + (fr == null) = (bq == null). Proof. rewrite /is_queue; case: ifP=>[E [->]_ _| E [xt][x][h'][_] D] //. by case: eqP D=>// -> /validR; rewrite validPtUn. @@ -71,7 +83,6 @@ Qed. (* main methods *) (* new queue is a pair of pointers to an empty segment *) - Program Definition new : STsep (emp, [vfun v => shape v [::]]) := Do (x <-- alloc null; @@ -85,34 +96,32 @@ by exists null, null, Unit; rewrite !unitR /= in V *; rewrite joinC. Qed. (* freeing a queue, possible only when it's empty *) - Program Definition free (q : queue) : STsep (shape q [::], [vfun _ h => h = Unit]) := Do (dealloc (front q);; dealloc (back q)). Next Obligation. (* pull out ghosts and precondition *) -move=>q [] _ /= [fr][bq][h][/[swap]->/[swap]]. +move=>q [_][fr][bq][h][/[swap] -> /[swap]]. (* both pointers are null *) case/is_queue_nil=>->->->; rewrite unitR=>V. (* run the program *) by do 2![step]=>_; rewrite unitR. Qed. -(* for enqueue/dequeue we manipulate the underlying segment directly *) - -(* enqueuing is adding a node at the end *) +(* enqueue/dequeue manipulate the underlying segment directly *) +(* enqueuing = adding a node at the end *) Program Definition enq (q : queue) (x : T) : - {xs}, STsep (shape q xs, - [vfun _ => shape q (rcons xs x)]) := + STsep {xs} (shape q xs, + [vfun _ => shape q (rcons xs x)]) := Do (next <-- allocb null 2; next ::= x;; ba <-- !back q; back q ::= next;; (if (ba : ptr) == null then front q - else ba .+ 1) ::= next). + else ba.+1) ::= next). Next Obligation. (* pull out ghosts + precondition *) move=>q x [xs][] _ /= [fr][bq][h'][D -> H]. @@ -123,37 +132,37 @@ rewrite -(backfront H) unitR; case: ifP H=>Ef; rewrite /is_queue ?Ef. - (* the queue was empty, set the front pointer to new node *) case=>_->->; step; rewrite unitR=>V. (* massage the heap and restructure the goal *) - exists next, next, (next :-> x \+ next.+ 1 :-> null). + exists next, next, (next :-> x \+ next.+1 :-> null). rewrite joinA joinC; split=>//; apply/(@is_queue_rcons _ _ [::]). by exists Unit; rewrite unitL; split=>//; exact: (validL V). (* the queue wasn't empty, link the new node to the last one *) case=>s2[x2][i2][->] {}D -> H2; step=>V. (* massage the heap and simplify the goal *) -exists fr, next, (i2 \+ bq :-> x2 \+ bq.+ 1 :-> next \+ next :-> x \+ next.+ 1 :-> null). +exists fr, next, (i2 \+ bq :-> x2 \+ bq.+1 :-> next \+ + next :-> x \+ next.+1 :-> null). split; first by apply: (validX V). - by rewrite joinC !joinA. (* the new node conforms to the queue spec *) -apply/is_queue_rcons; exists (i2 \+ bq :-> x2 \+ bq.+ 1 :-> next). +apply/is_queue_rcons; exists (i2 \+ bq :-> x2 \+ bq.+1 :-> next). rewrite joinA; split=>//; first by apply: (validX V). (* assemble the old queue back *) by apply/lseg_rcons; exists bq, i2; rewrite joinA. Qed. -(* dequeuing is removing the head node and adjusting pointers *) - +(* dequeuing = removing the head node and adjusting pointers *) Program Definition deq (q : queue) : - {xs}, STsep (shape q xs, - fun y h => shape q (behead xs) h /\ - match y with Val v => xs = v :: behead xs - | Exn e => e = EmptyQueue /\ xs = [::] end) := + STsep {xs} (shape q xs, + fun y h => shape q (behead xs) h /\ + match y with Val v => xs = v :: behead xs + | Exn e => e = EmptyQueue /\ xs = [::] end) := Do (fr <-- !front q; if (fr : ptr) == null then throw EmptyQueue else x <-- !fr; - next <-- !fr .+ 1; + next <-- !fr.+1; front q ::= next;; dealloc fr;; - dealloc fr .+ 1;; + dealloc fr.+1;; if (next : ptr) == null then back q ::= null;; ret x @@ -182,7 +191,7 @@ case: ifP H=>[/eqP ->|N] H. case/(lseg_null (validX V2)): H D=>/=-> _ _ /validR. by rewrite validPtUn. (* return the segment head and simplify *) -step=>V; split=>//; exists next, bq, (h2 \+ (bq :-> x \+ bq .+ 1 :-> null)). +step=>V; split=>//; exists next, bq, (h2 \+ (bq :-> x \+ bq.+1 :-> null)). by rewrite N; split=>//; vauto; apply: (validX V). Qed. diff --git a/examples/quicksort.v b/examples/quicksort.v index ff27146..d405474 100644 --- a/examples/quicksort.v +++ b/examples/quicksort.v @@ -1,81 +1,36 @@ +(* +Copyright 2022 IMDEA Software Institute +Licensed under the Apache License, Version 2.0 (the "License"); +you may not use this file except in compliance with the License. +You may obtain a copy of the License at + http://www.apache.org/licenses/LICENSE-2.0 +Unless required by applicable law or agreed to in writing, software +distributed under the License is distributed on an "AS IS" BASIS, +WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. +See the License for the specific language governing permissions and +limitations under the License. +*) + From mathcomp Require Import ssreflect ssrbool ssrfun. -From mathcomp Require Import ssrnat eqtype seq path fintype tuple finfun finset. -From mathcomp Require Import fingroup perm. +From mathcomp Require Import ssrnat eqtype seq path fintype tuple finfun. +From mathcomp Require Import finset fingroup perm. From mathcomp Require Import interval. -From pcm Require Import options axioms prelude pred ordtype slice. +From pcm Require Import options axioms prelude seqext pred ordtype slice. From pcm Require Import pcm unionmap heap. -From htt Require Import options interlude model heapauto. +From htt Require Import options model heapauto. From htt Require Import array. - -(* hack to avoid "_ *p _" notation clash *) From mathcomp Require order. Import order.Order.NatOrder order.Order.TTheory. +Local Open Scope order_scope. -(* TODO added to mathcomp > 1.16 *) -Lemma perm_on_id {T : finType} (u : {perm T}) (S : {set T}) : - perm_on S u -> #|S| <= 1 -> u = 1%g. -Proof. -rewrite leq_eqVlt ltnS leqn0 => pSu S10; apply/permP => t; rewrite perm1. -case/orP : S10; last first. - by move/eqP/cards0_eq => S0; apply: (out_perm pSu); rewrite S0 inE. -move=> /cards1P[x Sx]. -have [-> | ntx] := eqVneq t x; last by apply: (out_perm pSu); rewrite Sx inE. -by apply/eqP; have := perm_closed x pSu; rewrite Sx !inE => ->. -Qed. - -Lemma perm_onC {T : finType} (S1 S2 : {set T}) (u1 u2 : {perm T}) : - perm_on S1 u1 -> perm_on S2 u2 -> - [disjoint S1 & S2] -> - commute u1 u2. -Proof. -move=> pS1 pS2 S12; apply/permP => t; rewrite !permM. -case/boolP : (t \in S1) => tS1. - have /[!disjoint_subset] /subsetP {}S12 := S12. - by rewrite !(out_perm pS2) //; apply: S12; rewrite // perm_closed. -case/boolP : (t \in S2) => tS2. - have /[1!disjoint_sym] /[!disjoint_subset] /subsetP {}S12 := S12. - by rewrite !(out_perm pS1) //; apply: S12; rewrite // perm_closed. -by rewrite (out_perm pS1) // (out_perm pS2) // (out_perm pS1). -Qed. - -Lemma tperm_on {T : finType} (x y : T) : perm_on [set x; y] (tperm x y). -Proof. -by apply/subsetP => z /[!inE]; case: tpermP => [->|->|]; rewrite eqxx // orbT. -Qed. - -(* TODO added to fcsl-pcm *) -Corollary slice_oPR {A : Type} a x (s : seq A) : - 0 < x -> - &:s (Interval a (BRight x.-1)) = &:s (Interval a (BLeft x)). -Proof. by move=>Hx; rewrite -{2}(prednK Hx) slice_oSR. Qed. - -Lemma slice_FR {A : Type} (s : seq A) x : &:s (Interval x +oo) = &:s (Interval x (BLeft (size s))). -Proof. by rewrite /slice /= addn0. Qed. - -Lemma slice_extrude {A : Type} (s : seq A) (i : interval nat) : - order.Order.lt i.1 i.2 -> - s = &:s (Interval -oo i.1) ++ &:s i ++ &:s (Interval i.2 +oo). -Proof. -case: i=>[[[] i|[]][[] j|[]]] //=; rewrite bnd_simp=>H; - rewrite ?itv_minfR ?itv_pinfL /= ?cats0. -- by rewrite -{1}(slice_uu s) (slice_split _ true (x:=i)) //= - (slice_split _ true (x:=j) (i:=`[i, +oo[)) //= in_itv /= andbT; apply/ltnW. -- by rewrite -{1}(slice_uu s) (slice_split _ true (x:=i)) //= - (slice_split _ false (x:=j) (i:=`[i, +oo[)) //= in_itv /= andbT. -- by rewrite -{1}(slice_uu s) (slice_split _ true (x:=i)). -- by rewrite -{1}(slice_uu s) (slice_split _ false (x:=i)) //= - (slice_split _ true (x:=j) (i:=`]i, +oo[)) //= in_itv /= andbT. -- by rewrite -{1}(slice_uu s) (slice_split _ false (x:=i)) //= - (slice_split _ false (x:=j) (i:=`]i, +oo[)) //= in_itv /= andbT. -- by rewrite -{1}(slice_uu s) (slice_split _ false (x:=i)). -- by rewrite -{1}(slice_uu s) (slice_split _ true (x:=j)). -- by rewrite -{1}(slice_uu s) (slice_split _ false (x:=j)). -by rewrite slice_uu. -Qed. - -(****) +(* Brief mathematics of quickorting *) +(* There is some overlap with mathematics developed for bubblesort *) +(* but the development is repeated here to make the files *) +(* self-contained *) -Lemma leq_choose a b : a < b -> sumbool (a.+1 == b) (a.+1 < b). +Lemma leq_choose a b : + a < b -> + sumbool (a.+1 == b) (a.+1 < b). Proof. move=>H. case: (decP (b:=a.+1 < b) idP)=>[H2|/negP H2]; first by right. @@ -83,7 +38,6 @@ by left; rewrite eqn_leq H /= leqNgt. Qed. (* TODO copied from bubble *) -Section OrdArith. Fact ord_trans {n} (j : 'I_n) (i : 'I_n) (Hi : i < j) : (i.+1 < n)%N. Proof. @@ -91,20 +45,27 @@ case: j Hi=>j Hj /= Hi. by apply/leq_trans/Hj. Qed. +Section OrdArith. + (* increase by 1 within the bound *) Definition Sbo n (i : 'I_n) (prf : (i.+1 < n)%N) : 'I_n. Proof. case: i prf=>/= m Hm; apply: Ordinal. Defined. -Lemma Sbo_eq n (i : 'I_n) (prf : (i.+1 < n)%N) : nat_of_ord (Sbo prf) = i.+1. +Lemma Sbo_eq n (i : 'I_n) (prf : (i.+1 < n)%N) : + nat_of_ord (Sbo prf) = i.+1. Proof. by case: i prf. Qed. -Lemma Sbo_lift n (i j : 'I_n) (H1 : i < j) : i.+1 < j -> Sbo (ord_trans H1) < j. +Lemma Sbo_lift n (i j : 'I_n) (H1 : i < j) : + i.+1 < j -> + Sbo (ord_trans H1) < j. Proof. by case: i H1. Qed. -Lemma Sbo_leq n (i j k : 'I_n) (H1 : i <= j) (H2 : j < k) : Sbo (ord_trans (leq_ltn_trans H1 H2)) <= Sbo (ord_trans H2). +Lemma Sbo_leq n (i j k : 'I_n) (H1 : i <= j) (H2 : j < k) : + Sbo (ord_trans (leq_ltn_trans H1 H2)) <= Sbo (ord_trans H2). Proof. by case: j H1 H2; case: i. Qed. -Lemma Sbo_lt n (i j k : 'I_n) (H1 : i <= j) (H2 : j < k) : i <= Sbo (ord_trans H2). +Lemma Sbo_lt n (i j k : 'I_n) (H1 : i <= j) (H2 : j < k) : + i <= Sbo (ord_trans H2). Proof. by case: j H1 H2; case: i=>/=x Hx y Hy Hxy Hyk; apply/ltnW. Qed. (* increase by 1 with saturation *) @@ -115,7 +76,8 @@ case: i=>/= m Hm; case: (ltnP m.+1 n)=>[H|_]. by apply: (@Ordinal _ m Hm). Defined. -Lemma Sso_eq n (i : 'I_n) : nat_of_ord (Sso i) = if (i.+1 < n)%N then i.+1 else i. +Lemma Sso_eq n (i : 'I_n) : + nat_of_ord (Sso i) = if (i.+1 < n)%N then i.+1 else i. Proof. by case: i=>/= m prf; case: ltnP. Qed. (* decrease by 1 *) @@ -128,9 +90,10 @@ Proof. by case: i. Qed. End OrdArith. Section PermFgraph. -Variable (n : nat) (A : Type). +Variables (n : nat) (A : Type). -Lemma perm_on_notin (f : {ffun 'I_n -> A}) (p : 'S_n) (s : {set 'I_n}) (i : interval nat) : +Lemma perm_on_notin (f : {ffun 'I_n -> A}) (p : 'S_n) + (s : {set 'I_n}) (i : interval nat) : perm_on s p -> [disjoint s & [set x : 'I_n | (x : nat) \in i]] -> &:(fgraph (pffun p f)) i = &:(fgraph f) i. @@ -154,9 +117,9 @@ case: j=>[[] jx|[]]; case: i=>[[] ix|[]]; by rewrite leqNgt (ltn_ord y). Qed. -Corollary tperm_notin (f : {ffun 'I_n -> A}) (x y : 'I_n) (i : interval nat) : - (x : nat) \notin i -> (y : nat) \notin i -> - &:(fgraph (pffun (tperm x y) f)) i = &:(fgraph f) i. +Lemma tperm_notin (f : {ffun 'I_n -> A}) (x y : 'I_n) (i : interval nat) : + (x : nat) \notin i -> (y : nat) \notin i -> + &:(fgraph (pffun (tperm x y) f)) i = &:(fgraph f) i. Proof. move=>Hx0 Hx1. apply: perm_on_notin; first by exact: tperm_on. @@ -167,11 +130,11 @@ Qed. End PermFgraph. Section PermFgraphEq. -Variable (n : nat) (A : eqType). +Variables (n : nat) (A : eqType). Lemma perm_fgraph (p : 'S_n) (f : {ffun 'I_n -> A}) : - perm_eq (fgraph (pffun p f)) - (fgraph f). + perm_eq (fgraph (pffun p f)) + (fgraph f). Proof. apply/tuple_permP. exists (cast_perm (esym (card_ord n)) p). @@ -181,9 +144,9 @@ by rewrite tnth_fgraph tnth_map tnth_fgraph ffunE /= tnth_ord_tuple Qed. Lemma perm_on_fgraph (i : interval nat) (p : 'S_n) (f : {ffun 'I_n -> A}) : - perm_on [set x : 'I_n | (x : nat) \in i] p -> - perm_eq &:(fgraph (pffun p f)) i - &:(fgraph f ) i. + perm_on [set x : 'I_n | (x : nat) \in i] p -> + perm_eq &:(fgraph (pffun p f)) i + &:(fgraph f ) i. Proof. case: i=>i j H. case/boolP: (order.Order.lt i j)=>[Hij|]; last first. @@ -213,6 +176,13 @@ Qed. End PermFgraphEq. + +(*****************) +(*****************) +(* Verifications *) +(*****************) +(*****************) + Section Lomuto. Variable (n : nat) (A : ordType). @@ -255,11 +225,9 @@ Variable (n : nat) (A : ordType). (* go 0 (size a)-1 *) (***************************************************) -Opaque Array.write Array.read. - Program Definition swap (a : {array 'I_n -> A}) (i j : 'I_n) : - {f : {ffun 'I_n -> A}}, - STsep (Array.shape a f, + STsep {f : {ffun 'I_n -> A}} + (Array.shape a f, [vfun _ h => h \In Array.shape a (pffun (tperm i j) f)]) := Do (if i == j then skip else @@ -272,21 +240,20 @@ move=>a i j /= [f][] h /= H. case: ifP=>[/eqP->|Hij]. - by step=>_; rewrite tperm1 pffunE1. do 2!apply: [stepE f, h]=>//= _ _ [->->]. -apply: [stepE f]=>//= _ _ [-> V1]; set f1 := finfun [eta f with i |-> f j]. -apply: [gE f1]=>//= _ _ [-> V2]; set f2 := finfun [eta f1 with j |-> f i]. -move=>{h H}_; split=>//=; suff {V1 V2}: f2 = pffun (tperm i j) f by move=>->. +apply: [stepE f]=>//= _ {H}h H; set f1 := (finfun _) in H. +apply: [gE f1]=>//= _ {H}h H; set f2 := (finfun _) in H. +suff {H}: f2 = pffun (tperm i j) f by move=><-. rewrite {}/f2 {}/f1; apply/ffunP=>/= x; rewrite !ffunE /= ffunE /=. by case: tpermP=>[->|->|/eqP/negbTE->/eqP/negbTE->] {x}//; rewrite eqxx // Hij. Qed. -Opaque swap. - -Definition partition_lm_loop (a : {array 'I_n -> A}) (pivot : A) (lo hi : 'I_n) := +Definition partition_lm_loop (a : {array 'I_n -> A}) (pivot : A) + (lo hi : 'I_n) := forall ij : sigT (fun i : 'I_n => sig (fun j : 'I_n => i <= j /\ j < hi)), let i := projT1 ij in let j := proj1_sig (projT2 ij) in - {f : {ffun 'I_n -> A}}, - STsep (fun h => [/\ h \In Array.shape a f, + STsep {f : {ffun 'I_n -> A}} + (fun h => [/\ h \In Array.shape a f, lo <= i, all (oleq^~ pivot) (&:(fgraph f) `[lo : nat, i : nat[) & all (ord pivot) (&:(fgraph f) `[i : nat, j : nat[)], @@ -298,9 +265,10 @@ Definition partition_lm_loop (a : {array 'I_n -> A}) (pivot : A) (lo hi : 'I_n) all (oleq^~ pivot) (&:(fgraph f') `[lo : nat, v : nat[) & all (ord pivot) (&:(fgraph f') `[v : nat, hi : nat[)]]). -Program Definition partition_lm_pass (a : {array 'I_n -> A}) (pivot : A) (lo hi : 'I_n) (Hlo : lo < hi): - {f : {ffun 'I_n -> A}}, - STsep (Array.shape a f, +Program Definition partition_lm_pass (a : {array 'I_n -> A}) (pivot : A) + (lo hi : 'I_n) (Hlo : lo < hi): + STsep {f : {ffun 'I_n -> A}} + (Array.shape a f, [vfun (v : 'I_n) h => lo <= v <= hi /\ exists p, let f' := pffun p f in @@ -309,7 +277,8 @@ Program Definition partition_lm_pass (a : {array 'I_n -> A}) (pivot : A) (lo hi all (oleq^~ pivot) (&:(fgraph f') `[lo : nat, v : nat[) & all (ord pivot) (&:(fgraph f') `[v : nat, hi : nat[)]]) := Do (let go := - Fix (fun (loop : partition_lm_loop a pivot lo hi) '(existT i (exist j (conj Hi Hj))) => + ffix (fun (loop : partition_lm_loop a pivot lo hi) + '(existT i (exist j (conj Hi Hj))) => Do (x <-- Array.read a j; if oleq x pivot then swap a i j;; @@ -329,11 +298,11 @@ Next Obligation. move=>a pivot lo hi Hlo loop _ i _ j _ Hi Hj [f][] h /= [H Oli Ai Aj]. apply: [stepE f, h]=>//= _ _ [->->]. case: oleqP=>Hfp. -- (* a[j] <= pivot, make swap *) - apply: [stepE f]=>//= _ m Hm. +(* a[j] <= pivot, make swap *) +- apply: [stepE f]=>//= _ m Hm. case: (leq_choose Hj)=>Hj1. - - (* j+1 = hi, exit *) - step=>_; split. + (* j+1 = hi, exit *) + - step=>_; split. - rewrite Sbo_eq; apply/andP; split=>//. by apply/leq_ltn_trans/Hj. exists (tperm i j); rewrite Sbo_eq; split=>//. @@ -354,7 +323,7 @@ case: oleqP=>Hfp. rewrite all_rcons; apply/andP; split. - by rewrite ffunE tpermR. by rewrite tperm_notin ?slice_oSL // - in_itv /= negb_and leEnat ltEnat /= -!leqNgt leqnn // orbT. + in_itv /= negb_and leEnat ltEnat /= -!leqNgt leqnn // orbT. (* j+1 < hi, loop *) apply: [gE (pffun (tperm i j) f)]=>//=. - split=>//; rewrite !Sbo_eq; first by apply/ltnW. @@ -400,11 +369,10 @@ move=>/= a pivot lo hi Hlo [f][] i /= H. by apply: [gE f]=>//=; split=>//; rewrite slice_kk. Qed. -Opaque partition_lm_pass. - -Program Definition partition_lm (a : {array 'I_n -> A}) (lo hi : 'I_n) (Hlo : lo < hi): - {f : {ffun 'I_n -> A}}, - STsep (Array.shape a f, +Program Definition partition_lm (a : {array 'I_n -> A}) + (lo hi : 'I_n) (Hlo : lo < hi): + STsep {f : {ffun 'I_n -> A}} + (Array.shape a f, [vfun (v : 'I_n) h => lo <= v <= hi /\ exists p, let f' := pffun p f in @@ -417,7 +385,7 @@ Program Definition partition_lm (a : {array 'I_n -> A}) (lo hi : 'I_n) (Hlo : lo swap a v hi;; ret v). Next Obligation. -move=> a lo hi Hlo /= [f][] h /= [E V]. +move=> a lo hi Hlo /= [f][] h /= E. apply: [stepE f, h]=>//= _ _ [->->]. apply: [stepE f]=>//= v m [Hi][p][Pm Hm Al Ah]. apply: [stepE (pffun p f)]=>//= _ k Hj. @@ -451,14 +419,12 @@ End Lomuto. Section LomutoQsort. Variable (n : nat) (A : ordType). -Opaque partition_lm. - Definition quicksort_lm_loop (a : {array 'I_n.+1 -> A}) := forall (lohi : 'I_n.+1 * 'I_n.+1), let lo := lohi.1 in let hi := lohi.2 in - {f : {ffun 'I_n.+1 -> A}}, - STsep (Array.shape a f, + STsep {f : {ffun 'I_n.+1 -> A}} + (Array.shape a f, [vfun (_ : unit) h => exists p, let f' := pffun p f in [/\ perm_on [set ix : 'I_n.+1 | lo <= ix <= hi] p, @@ -466,20 +432,21 @@ Definition quicksort_lm_loop (a : {array 'I_n.+1 -> A}) := sorted oleq (&:(fgraph f') `[lo : nat, hi : nat])]]). Program Definition quicksort_lm (a : {array 'I_n.+1 -> A}) : - {f : {ffun 'I_n.+1 -> A}}, - STsep (Array.shape a f, + STsep {f : {ffun 'I_n.+1 -> A}} + (Array.shape a f, [vfun (_ : unit) h => exists p, let f' := pffun p f in h \In Array.shape a f' /\ sorted oleq (fgraph f')]) := Do (let go := - Fix (fun (loop : quicksort_lm_loop a) '(l,h) => + ffix (fun (loop : quicksort_lm_loop a) '(l,h) => Do (if decP (b:=(l : nat) < h) idP isn't left pf then skip else v <-- partition_lm a pf; loop (l, Po v);; (* we use saturating increment to stay under n+1 *) (* and keep the classical form of the algorithm *) - (* the overflow case will exit on next call because v = h = n-1 *) + (* the overflow case will exit on next call *) + (* because v = h = n-1 *) loop (Sso v, h))) in go (ord0, ord_max)). Next Obligation. @@ -507,10 +474,10 @@ exists (pr * (pl * p))%g; split=>//. case/andP=>->/= Hz. apply/leq_trans/Hvh/(leq_trans Hz). by exact: leq_pred. -(* we need to handle two edge cases: v=0 and v=n *) +(* need to handle two edge cases: v=0 and v=n *) case: (eqVneq v ord0)=>[Ev|Nv0]. -- (* if v=0, then l=0 and left partition is empty *) - have El: l = ord0. +(* if v=0, then l=0 and left partition is empty *) +- have El: l = ord0. - by move: Hvl; rewrite Ev leqn0 => /eqP El; apply/ord_inj. rewrite Ev El /= in Hpl. have Epl: pl = 1%g. @@ -529,8 +496,8 @@ case: (eqVneq v ord0)=>[Ev|Nv0]. rewrite pffunEM perm_sym -!slice_oSL (_ : 0 = (ord0 : 'I_n.+1)) //. by apply: perm_on_fgraph. move: (ltn_ord v); rewrite ltnS leq_eqVlt; case/orP=>[/eqP Ev|Nv]. -- (* if v=n, then h=n and right partition is empty *) - have Eh: (h : nat) = n. +(* if v=n, then h=n and right partition is empty *) +- have Eh: (h : nat) = n. - apply/eqP; rewrite eqn_leq; move: Hvh; rewrite Ev=>->; rewrite andbT. by move: (ltn_ord h); rewrite ltnS. rewrite Ev Eh /= ltnn in Hpr. @@ -541,15 +508,19 @@ move: (ltn_ord v); rewrite ltnS leq_eqVlt; case/orP=>[/eqP Ev|Nv]. by rewrite !inE -eqn_leq =>/eqP E; apply/eqP/ord_inj. move: Sl Hpl; rewrite Eh Ev Epr mul1g => Sl Hpl. rewrite slice_xR; last by rewrite bnd_simp leEnat; move: Hvl; rewrite Ev. - rewrite {22}(_ : n = (ord_max : 'I_n.+1)) // onth_codom /= sorted_rconsE //. - move: Sl; rewrite slice_oPR; last by rewrite lt0n -Ev. + rewrite {22}(_ : n = (ord_max : 'I_n.+1)) // onth_codom /= sorted_rconsE //=. + move: Sl; rewrite slice_oPR /order.Order.lt/= lt0n -{1}Ev Nv0. move=>->; rewrite andbT; move: Al; rewrite Ev. have ->: pffun (pl * p) f ord_max = pffun p f ord_max. - - by rewrite !ffunE permM (out_perm Hpl) // inE negb_and -!ltnNge /= ltn_predL lt0n -{3}Ev Nv0 orbT. - have ->: v = ord_max by apply/ord_inj. - rewrite (perm_all (s2:=&:(codom (pffun (pl * p) f)) `[l: nat, n[)) // pffunEM perm_sym. + - rewrite !ffunE permM (out_perm Hpl) // inE negb_and -!ltnNge /=. + by rewrite ltn_predL lt0n -{3}Ev Nv0 orbT. + have ->: v = ord_max by apply/ord_inj. + rewrite [in X in X -> _] + (perm_all (s2:=&:(codom (pffun (pl * p) f)) `[l: nat, n[)) //. + rewrite pffunEM perm_sym. rewrite {8 15}(_ : n = (ord_max : 'I_n.+1)) //; apply: perm_on_fgraph. - apply/(subset_trans Hpl)/subsetP=>/= z; rewrite 2!inE in_itv /= leEnat ltEnat /=. + apply/(subset_trans Hpl)/subsetP=>/= z. + rewrite 2!inE in_itv /= leEnat ltEnat /=. by case/andP=>->/= Hz; apply: (leq_ltn_trans Hz); rewrite ltn_predL lt0n -Ev. (* the general case *) rewrite Nv in Hpr Sr. @@ -559,11 +530,13 @@ rewrite (slice_xL (x:=v)) // onth_codom /=. have -> : pffun (pr * (pl * p)) f v = pffun p f v. - rewrite !ffunE mulgA; suff ->: (pr * pl * p)%g v = p v by []. rewrite permM. - have Hmul: perm_on [set ix : 'I_n.+1| (l <= ix <= v.-1) || (v < ix <= h)] (pr * pl)%g. + have Hmul : perm_on [set ix : 'I_n.+1| (l <= ix <= v.-1) || (v < ix <= h)] + (pr * pl)%g. - apply: perm_onM. - by apply/(subset_trans Hpr)/subsetP=>/= z; rewrite !inE=>->; rewrite orbT. by apply/(subset_trans Hpl)/subsetP=>/= z; rewrite !inE=>->. - by rewrite (out_perm Hmul) // inE negb_or !negb_and -leqNgt -!ltnNge leqnn /= andbT ltn_predL lt0n Nv0 orbT. + rewrite (out_perm Hmul) // inE negb_or !negb_and -leqNgt -!ltnNge leqnn /=. + by rewrite andbT ltn_predL lt0n Nv0 orbT. rewrite {1}pffunEM (perm_on_notin _ Hpr); last first. - rewrite disjoint_subset; apply/subsetP=>/= z. rewrite 3!inE in_itv /= negb_and leEnat ltEnat /= -leqNgt -ltnNge. @@ -577,34 +550,36 @@ rewrite -mulgA (pffunEM _ (pr * p)%g) (perm_on_notin _ Hpl) in Sr *; last first. - rewrite disjoint_subset; apply/subsetP=>/= z. rewrite 3!inE in_itv /= negb_and leEnat /= -leqNgt -ltnNge. case/andP=>_ Hz; apply/orP; left; apply: (leq_trans Hz). - by exact: leq_pred. -rewrite sorted_pairwise // pairwise_cat /= allrel_consr -andbA -!sorted_pairwise //. + exact: leq_pred. +rewrite sorted_pairwise // pairwise_cat /=. +rewrite allrel_consr -andbA -!sorted_pairwise //. apply/and5P; split=>//. - move: Al; congr (_ = _); apply/esym/perm_all; rewrite pffunEM. - apply/perm_on_fgraph/(subset_trans Hpl)/subsetP=>/= z; rewrite 2!inE in_itv /= leEnat ltEnat /=. + apply/perm_on_fgraph/(subset_trans Hpl)/subsetP=>/= z. + rewrite 2!inE in_itv /= leEnat ltEnat /=. by case/andP=>->/= Hz; apply: (leq_ltn_trans Hz); rewrite ltn_predL lt0n. - apply/allrelP=>x y Hx Hy; apply: (otrans (y:=pffun p f v)). - move/allP: Al=>/(_ x); apply. move: Hx; congr (_ = _); move: x; apply: perm_mem. rewrite pffunEM; apply: perm_on_fgraph. - apply/(subset_trans Hpl)/subsetP=>/= z; rewrite 2!inE in_itv /= leEnat ltEnat /=. + apply/(subset_trans Hpl)/subsetP=>/= z. + rewrite 2!inE in_itv /= leEnat ltEnat /=. by case/andP=>->/= Hz; apply: (leq_ltn_trans Hz); rewrite ltn_predL lt0n. apply/ordW; move/allP: Ah=>/(_ y); apply. move: Hy; congr (_ = _); move: y; apply: perm_mem. by rewrite pffunEM; apply: perm_on_fgraph. -- by rewrite slice_oPR // in Sl; rewrite lt0n. +- by rewrite slice_oPR /order.Order.lt/= lt0n Nv0 in Sl. have HS: subpred (ord (pffun p f v)) (oleq (pffun p f v)). - by move=>z /ordW. move/(sub_all HS): Ah; congr (_ = _); apply/esym/perm_all. by rewrite pffunEM; apply/perm_on_fgraph. Qed. Next Obligation. -move=>a [f][] i /= H. -apply: [gE f]=>//= _ m [p][Hp Hm Hs] _. -exists p; split=>//. -by rewrite -(slice_uu (codom _)) slice_0L slice_FR size_codom card_ord slice_oSR. +move=>a [f][] i /= H; apply: [gE f]=>//= _ m [p][Hp Hm Hs] _. +exists p; split=>//; rewrite -(slice_uu (codom _)) slice_0L. +by rewrite slice_FR size_codom card_ord slice_oSR. Qed. End LomutoQsort. -(* TODO Hoare variant *) + diff --git a/examples/stack.v b/examples/stack.v old mode 100755 new mode 100644 index 6d0892f..d374dc8 --- a/examples/stack.v +++ b/examples/stack.v @@ -1,3 +1,16 @@ +(* +Copyright 2022 IMDEA Software Institute +Licensed under the Apache License, Version 2.0 (the "License"); +you may not use this file except in compliance with the License. +You may obtain a copy of the License at + http://www.apache.org/licenses/LICENSE-2.0 +Unless required by applicable law or agreed to in writing, software +distributed under the License is distributed on an "AS IS" BASIS, +WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. +See the License for the specific language governing permissions and +limitations under the License. +*) + From Coq Require Import ssreflect ssrbool ssrfun. From mathcomp Require Import eqtype seq. From pcm Require Import options axioms pred. @@ -13,57 +26,49 @@ Section Stack. Variable T : Type. Notation stack := (stack T). -Opaque insert head remove. - -(* a stack is just a pointer to a singly-linked list *) +(* stack is a pointer to a singly-linked list *) Definition shape s (xs : seq T) := - [Pred h | exists p h', [ /\ valid (s :-> p \+ h'), - h = s :-> p \+ h' & + [Pred h | exists p h', [ /\ h = s :-> p \+ h' & h' \In lseq p xs]]. -(* a heap cannot match two different specs *) -Lemma shape_inv : forall s xs1 xs2 h, - h \In shape s xs1 -> h \In shape s xs2 -> xs1 = xs2. +(* heap cannot match two different specs *) +Lemma shape_inv s xs1 xs2 h : + valid h -> + h \In shape s xs1 -> + h \In shape s xs2 -> + xs1 = xs2. Proof. -move=>s xs1 xs2 _ [p][h1][D -> S][p2][h2][D2]. -case/(cancelO _ D)=><- _; rewrite !unitL=><-. -by apply: lseq_func=>//; move/validR: D. +move=>V [p][h1][E S][x][h'][]; rewrite {h}E in V *. +case/(cancelO _ V)=><- _; rewrite !unitL=><-. +by apply: lseq_func=>//; move/validR: V. Qed. -(* well-formed stack is a valid heap *) - -Lemma shapeD s xs : forall h, h \In shape s xs -> valid h. -Proof. by move=>h [p][h'][D ->]. Qed. - (* main methods *) (* new stack is a pointer to an empty heap/list *) - Program Definition new : STsep (emp, [vfun y => shape y [::]]) := Do (alloc null). Next Obligation. by move=>/= [] i ?; step=>??; vauto. Qed. (* freeing a stack, possible only when it's empty *) - Program Definition free s : STsep (shape s [::], [vfun _ h => h = Unit]) := Do (dealloc s). Next Obligation. -by move=>s [] ? /= [?][?][V -> [_ ->]]; step=>_; rewrite unitR. +by case=>i [?][?][->][_ ->]; step=>_; rewrite unitR. Qed. (* pushing to the stack is inserting into the list and updating the pointer *) - -Program Definition push s x : {xs}, STsep (shape s xs, - [vfun _ => shape s (x :: xs)]) := +Program Definition push s x : STsep {xs} (shape s xs, + [vfun _ => shape s (x :: xs)]) := Do (l <-- !s; l' <-- insert l x; s ::= l'). Next Obligation. (* pull out ghost + precondition, get the list *) -move=>s x [xs][] _ /= [l][h][V -> H]; step. +case=>xs [] _ /= [l][h][-> H]; step. (* run the insert procedure with the ghost, deconstruct the new list *) -apply: [stepX xs]@h=>//= x0 _ [r][h'][-> H']. +apply: [stepX xs]@h=>//= l' _ [r][h'][-> H']. (* store the new list *) by step=>V'; hhauto. Qed. @@ -71,12 +76,11 @@ Qed. (* popping from the stack is: *) (* 1. trying to get the head *) (* 2. removing it from the list and updating the pointer on success *) - Program Definition pop s : - {xs}, STsep (shape s xs, - fun y h => shape s (behead xs) h /\ - match y with Val v => xs = v :: behead xs - | Exn e => e = EmptyStack /\ xs = [::] end) := + STsep {xs} (shape s xs, + fun y h => shape s (behead xs) h /\ + match y with Val v => xs = v :: behead xs + | Exn e => e = EmptyStack /\ xs = [::] end) := Do (l <-- !s; try (head l) (fun x => @@ -86,7 +90,7 @@ Program Definition pop s : (fun _ => throw EmptyStack)). Next Obligation. (* pull out ghost vars and precondition *) -move=>s [xs][] _ /= [p][h0][V -> H]. +case=>xs [] _ /= [p][h0][-> H]. (* get the list and invoke head on it, deal with exception first *) step; apply/[tryX xs]@h0=>//= [x|ex] m [Hm]; last first. - (* throw the stack exception *) diff --git a/examples/tree.v b/examples/tree.v index 23c4027..2b59094 100644 --- a/examples/tree.v +++ b/examples/tree.v @@ -1,11 +1,24 @@ +(* +Copyright 2023 IMDEA Software Institute +Licensed under the Apache License, Version 2.0 (the "License"); +you may not use this file except in compliance with the License. +You may obtain a copy of the License at + http://www.apache.org/licenses/LICENSE-2.0 +Unless required by applicable law or agreed to in writing, software +distributed under the License is distributed on an "AS IS" BASIS, +WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. +See the License for the specific language governing permissions and +limitations under the License. +*) + +From HB Require Import structures. From mathcomp Require Import ssreflect ssrbool ssrfun. From mathcomp Require Import eqtype ssrnat seq bigop choice. From pcm Require Import pred prelude seqext. -From htt Require Import interlude. (* arbitrarily branching tree *) Inductive tree A := TNode of A & seq (tree A). -Arguments TNode [A]. +Arguments TNode {A}. Section Tree. Context {A : Type}. @@ -16,26 +29,18 @@ Definition ch (t : tree A) := let: TNode _ ts := t in ts. Lemma tree_ext (t : tree A) : TNode (rt t) (ch t) = t. Proof. by case: t. Qed. -(* a leaf is a node with an empty list *) +(* leaf is a node with an empty list *) Definition lf a : tree A := TNode a [::]. Lemma tree_ind' (P : tree A -> Prop) : - (forall a l, All P l -> P (TNode a l)) -> - forall t, P t. -Proof. -move=> indu; fix H 1. -elim => a l; apply indu. -by elim: l. -Qed. + (forall a l, All P l -> P (TNode a l)) -> + forall t, P t. +Proof. by move=>indu; fix H 1; elim => a l; apply indu; elim: l. Qed. Lemma tree_rec' (P : tree A -> Type) : - (forall a l, AllT P l -> P (TNode a l)) -> - forall t, P t. -Proof. -move=>indu; fix H 1. -elim => a l; apply: indu. -by elim: l. -Qed. + (forall a l, AllT P l -> P (TNode a l)) -> + forall t, P t. +Proof. by move=>indu; fix H 1; elim => a l; apply: indu; elim: l. Qed. (* custom induction principles *) @@ -52,7 +57,7 @@ Fixpoint preorder (t : tree A) : seq A := let: TNode a ts := t in foldl (fun s t => s ++ preorder t) [::a] ts. -Corollary foldl_cat {B C} z (fs : seq B) (a : B -> seq C): +Lemma foldl_cat {B C} z (fs : seq B) (a : B -> seq C): foldl (fun s t => s ++ a t) z fs = z ++ foldl (fun s t => s ++ a t) [::] fs. Proof. @@ -60,7 +65,9 @@ apply/esym/fusion_foldl; last by rewrite cats0. by move=>x y; rewrite catA. Qed. -Lemma preorderE t : preorder t = rt t :: \big[cat/[::]]_(c <- ch t) (preorder c). +Lemma preorderE t : + preorder t = + rt t :: \big[cat/[::]]_(c <- ch t) (preorder c). Proof. case: t=>a cs /=; rewrite foldl_cat /=; congr (_ :: _). elim: cs=>/= [| c cs IH]; first by rewrite big_nil. @@ -95,8 +102,8 @@ Qed. End EncodeDecodeTree. -Definition tree_eqMixin (A : eqType) := PcanEqMixin (@pcancel_tree A). -Canonical tree_eqType (A : eqType) := Eval hnf in EqType _ (@tree_eqMixin A). +HB.instance Definition _ (A : eqType) := + Equality.copy (tree A) (pcan_type (pcancel_tree A)). Section TreeEq. Context {A : eqType}. @@ -108,12 +115,11 @@ Fixpoint mem_tree (t : tree A) : pred A := Definition tree_eqclass := tree A. Identity Coercion tree_of_eqclass : tree_eqclass >-> tree. Coercion pred_of_tree (t : tree_eqclass) : {pred A} := mem_tree t. - Canonical tree_predType := ssrbool.PredType (pred_of_tree : tree A -> pred A). -(* The line below makes mem_tree a canonical instance of topred. *) -Canonical mem_tree_predType := ssrbool.PredType mem_tree. -Lemma in_tnode a t ts : (t \in TNode a ts) = (t == a) || has (fun q => t \in q) ts. +Lemma in_tnode a t ts : + (t \in TNode a ts) = + (t == a) || has (fun q => t \in q) ts. Proof. by []. Qed. (* frequently used facts about membership in a tree *) @@ -135,7 +141,7 @@ Lemma in_preorder t : preorder t =i t. Proof. elim/tree_ind1: t=>t cs IH x. rewrite preorderE in_tnode /= inE; case: eqVneq=>//= N. -rewrite big_cat_mem_seq; apply: eq_in_has=>z Hz. +rewrite big_cat_mem_has; apply: eq_in_has=>z Hz. by rewrite IH //; apply/mem_seqP. Qed. @@ -148,7 +154,7 @@ Qed. End TreeEq. -Arguments in_tnode2 [A x y a ts]. +Arguments in_tnode2 {A x y a ts}. #[export] Hint Resolve in_tnode1 : core. (* a simplified induction principle for eq types *) diff --git a/examples/union_find.v b/examples/union_find.v index 2c002aa..27041ea 100644 --- a/examples/union_find.v +++ b/examples/union_find.v @@ -1,53 +1,21 @@ +(* +Copyright 2023 IMDEA Software Institute +Licensed under the Apache License, Version 2.0 (the "License"); +you may not use this file except in compliance with the License. +You may obtain a copy of the License at + http://www.apache.org/licenses/LICENSE-2.0 +Unless required by applicable law or agreed to in writing, software +distributed under the License is distributed on an "AS IS" BASIS, +WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. +See the License for the specific language governing permissions and +limitations under the License. +*) + From mathcomp Require Import ssreflect ssrbool ssrfun fintype. From mathcomp Require Import eqtype ssrnat seq bigop choice. From pcm Require Import options axioms pred seqext. From pcm Require Import prelude pcm unionmap natmap heap autopcm automap. -From htt Require Import options interlude model heapauto tree. - -(* TODO: remove these after upgrading FCSL-PCM to 1.9+ *) -Corollary foldr_join A (U : pcm) h (s : seq A) (a : A -> U): - foldr (fun t h => h \+ a t) h s = - h \+ foldr (fun t h => h \+ a t) Unit s. -Proof. -apply/esym/fusion_foldr; last by rewrite unitR. -by move=>x y; rewrite joinA. -Qed. - -Section BigOpsUM. -Variables (K : ordType) (C : pred K) (T : Type). -Variables (U : union_map_class C T). -Variables (I : Type) (f : I -> U). - -Lemma big_find_someXP (xs : seq I) P a v : - find a (\big[PCM.join/Unit]_(i <- xs | P i) f i) = Some v -> - exists i, [/\ i \In xs, P i & find a (f i) = Some v]. -Proof. -by rewrite -big_filter=>/big_find_someX [i] /Mem_filter [H1 H2 H3]; exists i. -Qed. - -Lemma bigInXP (xs : seq I) P a v : - (a, v) \In \big[PCM.join/Unit]_(i <- xs | P i) f i -> - exists i, [/\ i \In xs, P i & (a, v) \In f i]. -Proof. by case/In_find/big_find_someXP=>x [X1 X2 /In_find]; exists x. Qed. - -End BigOpsUM. - -Section OMapBig. -Variables (K : ordType) (C : pred K) (T T' : Type). -Variables (U : union_map_class C T) (U' : union_map_class C T'). -Variables (I : Type) (f : I -> U). - -Lemma omapVUnBig (a : K * T -> option T') ts : - omap a (\big[PCM.join/Unit]_(x <- ts) f x) = - if valid (\big[PCM.join/Unit]_(x <- ts) f x) - then \big[PCM.join/Unit]_(x <- ts) omap a (f x) - else um_undef : U'. -Proof. -elim: ts=>[|t ts IH]; first by rewrite !big_nil omap0 valid_unit. -by rewrite !big_cons omapVUn IH; case: ifP=>//= /validR=>->. -Qed. - -End OMapBig. +From htt Require Import options model heapauto tree. (**************) (**************) @@ -55,7 +23,6 @@ End OMapBig. (**************) (**************) - (******************) (* inverted trees *) (******************) @@ -68,7 +35,7 @@ End OMapBig. Fixpoint tlay (t : tree ptr) (r : ptr) : heap := foldr (fun x h => h \+ tlay x (rt t)) (rt t :-> r) (ch t). -Fixpoint tset (t : tree ptr) (r : ptr) : union_mapUMC ptr_ordType ptr := +Fixpoint tset (t : tree ptr) (r : ptr) : umap ptr ptr := foldr (fun t h => h \+ tset t r) (rt t \\-> r) (ch t). (* explicit equations for expanding the defs of tlayout and tset *) @@ -77,7 +44,7 @@ Fixpoint tset (t : tree ptr) (r : ptr) : union_mapUMC ptr_ordType ptr := Lemma tlayE t r : tlay t r = - rt t :-> r \+ \big[PCM.join/Unit]_(x <- ch t) (tlay x (rt t)). + rt t :-> r \+ \big[join/Unit]_(x <- ch t) (tlay x (rt t)). Proof. case: t=>a ts /=; rewrite foldr_join; congr (_ \+ _). elim: ts=>[|t ts IH] /=; first by rewrite big_nil. @@ -86,7 +53,7 @@ Qed. Lemma tsetE t r : tset t r = - rt t \\-> r \+ \big[PCM.join/Unit]_(x <- ch t) (tset x r). + rt t \\-> r \+ \big[join/Unit]_(x <- ch t) (tset x r). Proof. case: t=>a ts /=; rewrite foldr_join; congr (_ \+ _). elim: ts=>[|t ts IH] /=; first by rewrite big_nil. @@ -105,25 +72,25 @@ Lemma valid_dom_tset (t : tree ptr) r : if valid (tset t r) then preorder t else [::]). Proof. elim/tree_ind2: t r=>a ts IH r. rewrite tsetE preorderE. -rewrite validPtUn !big_valid_dom_seq /= big_cat_mem_seq. +rewrite validPtUn !big_valid_dom_seq /= big_cat_mem_has. case: allP=>H /=; last first. - split=>[|x]; last first. - by rewrite domUn inE validPtUn !big_valid_dom_seq; case: allP. - rewrite andbC; case: uniq_big_catP=>//=; case=>H1 _ _. + rewrite andbC; case: uniq_big_catE=>//=; case=>H1 _ _. by case: H=>t T; rewrite IH // H1. case U1: (uniq _)=>/=; last first. - rewrite andbC; case U2: (uniq _)=>/=; last first. - by split=>// x; rewrite domPtUn inE validPtUn /= !big_valid_dom_seq U1 andbF. - case/uniq_big_catP: U2=>H1 H2 H3. - case: uniq_big_catP U1=>//; case; split. + case/uniq_big_catE: U2=>H1 H2 H3. + case: uniq_big_catE U1=>//; case; split. - by move=>i; rewrite uniq_dom. - move=>i k X D; apply: (H2 i k X). by rewrite (IH i X r) H in D. move=>i j k X Y Di Dj. apply: (H3 i j k)=>//; first by rewrite (IH i X r) H in Di. by rewrite (IH j Y r) H in Dj. -rewrite big_cat_mem_seq andbC; case: uniq_big_catP=>/=; last first. -- case/uniq_big_catP: U1=>H1 H2 H3; case; split. +rewrite big_cat_mem_has andbC; case: uniq_big_catE=>/=; last first. +- case/uniq_big_catE: U1=>H1 H2 H3; case; split. - by move=>i X; rewrite -(IH i X r) H. - by move=>i k X D; apply: (H2 i k X); rewrite (IH i X r) H. move=>i j k X Y Di Dj; apply: (H3 i j k X Y); first by rewrite (IH i X r) H. @@ -131,8 +98,8 @@ rewrite big_cat_mem_seq andbC; case: uniq_big_catP=>/=; last first. case=>K1 K2 K3; split=>[|x]. - by rewrite -!all_predC; apply: eq_in_all=>i X; rewrite /= IH ?H. rewrite domPtUn inE validPtUn /= !big_valid_dom_seq U1 andbT. -case: allP=>//= _; rewrite !big_cat_mem_seq -all_predC. -case: allP=>//= A; rewrite inE big_cat_mem_seq eq_sym. +case: allP=>//= _; rewrite !big_cat_mem_has -all_predC. +case: allP=>//= A; rewrite inE big_cat_mem_has eq_sym. by case: (x =P a)=>//= _; apply: eq_in_has=>i X; rewrite IH ?H. Qed. @@ -174,10 +141,10 @@ Qed. Lemma domeq_tlay_tset (t : tree ptr) r1 r2 : {in t, forall x, x != null} -> - dom_eq2 (tlay t r1) (tset t r2). + dom_eq (tlay t r1) (tset t r2). Proof. elim/tree_ind1: t r1 r2=>a ts IH r1 r2 Tn. -rewrite tlayE tsetE /= domeq2PtUn ?Tn // big_domeq2Un //. +rewrite tlayE tsetE /= domeqPtUn ?Tn // big_domeqUn //. by move=>x X; apply: IH=>// z Z; apply: Tn (in_tnode2 X Z). Qed. @@ -194,9 +161,9 @@ Lemma valid_tlay (t : tree ptr) r : valid (tset t r) && all (fun x => x != null) (preorder t). Proof. apply/idP/idP; last first. -- by case/andP=>V /tallP W; rewrite (domeq2VE (domeq_tlay_tset r r W)). +- by case/andP=>V /tallP W; rewrite (domeqVE (domeq_tlay_tset r r W)). move/[dup]=>V /valid_tlayE /[dup] N /tallP ->. -by rewrite -(domeq2VE (domeq_tlay_tset r r N)) V. +by rewrite -(domeqVE (domeq_tlay_tset r r N)) V. Qed. Lemma valid_tlayN (t : tree ptr) r : @@ -207,7 +174,7 @@ Lemma dom_tlay (t : tree ptr) r : valid (tlay t r) -> dom (tlay t r) =i t. Proof. rewrite valid_tlay=>/andP [V /tallP A] x. -by rewrite -(dom_tset V) (domeq2DE (domeq_tlay_tset r r A)). +by rewrite -(dom_tset V) (domeqDE (domeq_tlay_tset r r A)). Qed. Lemma dom_tlayE (t : tree ptr) r : @@ -242,15 +209,16 @@ case V : (valid (tset t r))=>/=; last first. rewrite -(dom_tset V); case: dom_find=>// v E _. elim/tree_ind1: t V E=>a ts IH; rewrite tsetE /= => V. rewrite !findPtUn2 //; case: (x =P a)=>//= _. -case/big_find_someX=>i X D; apply: big_find_some (validX V) (X) (IH _ X _ D). -by apply: big_validV (validR V) X. +case/big_find_someX=>i X /[dup] D /In_find/In_valid W. +by apply: IH X W (D). Qed. Lemma In_tsetP x r t: reflect ((x, r) \In tset t (rt t)) [&& valid (tset t (rt t)), x \in t & r == rt t]. Proof. -apply/(iffP idP); rewrite In_find find_tset; first by case/and3P=>->-> /eqP ->. +apply/(iffP idP); rewrite In_find find_tset. +- by case/and3P=>->-> /eqP ->. by case: ifP=>// /andP [->->][->] /=. Qed. @@ -293,10 +261,10 @@ Proof. by move/tlay_rt; rewrite eqxx. Qed. Definition flay ts := foldr (fun t h => tlay t (rt t) \+ h) Unit ts. Definition fset ts := foldr (fun t h => tset t (rt t) \+ h) Unit ts. -Lemma flayE ts : flay ts = \big[PCM.join/Unit]_(t <- ts) (tlay t (rt t)). +Lemma flayE ts : flay ts = \big[join/Unit]_(t <- ts) (tlay t (rt t)). Proof. by elim: ts=>[|t ts IH] /=; [rewrite big_nil|rewrite big_cons IH]. Qed. -Lemma fsetE ts : fset ts = \big[PCM.join/Unit]_(t <- ts) (tset t (rt t)). +Lemma fsetE ts : fset ts = \big[join/Unit]_(t <- ts) (tset t (rt t)). Proof. by elim: ts=>[|t ts IH] /=; [rewrite big_nil|rewrite big_cons IH]. Qed. Lemma find_flayTp (x : ptr) (ts : seq (tree ptr)) (p : dynamic id) : @@ -322,7 +290,7 @@ Lemma dom_fset (ts : seq (tree ptr)) x : valid (fset ts) -> x \in dom (fset ts) = has (fun t => x \in t) ts. Proof. -elim: ts=>[|t ts IH] //= V; first by rewrite dom0. +elim: ts=>[|t ts IH] //= V. by rewrite domUn V inE /= IH ?dom_tset ?(validX V). Qed. @@ -339,7 +307,7 @@ Lemma range_fsetE ts : range (fset ts) =i [pred x | valid (fset ts) && has (fun t => x == rt t) ts]. Proof. -elim: ts=>[|t ts IH] /= x; first by rewrite range0 andbF. +elim: ts=>[|t ts IH] //= x. rewrite rangeUn inE; case V : (valid _)=>//=. by rewrite range_tset ?inE ?IH ?(validX V). Qed. @@ -351,11 +319,13 @@ Proof. by move=>V; rewrite range_fsetE V inE. Qed. Lemma valid_fset_tset (ts : seq (tree ptr)) : - valid (fset ts) -> {in ts, forall i, valid (tset i (rt i))}. + valid (fset ts) -> + {in ts, forall i, valid (tset i (rt i))}. Proof. by move=>+ i Hi; rewrite fsetE big_valid_seq=>/andP [/allP /(_ i Hi)]. Qed. Lemma valid_flay_tlay (ts : seq (tree ptr)) : - valid (flay ts) -> {in ts, forall i, valid (tlay i (rt i))}. + valid (flay ts) -> + {in ts, forall i, valid (tlay i (rt i))}. Proof. by move=>+ i Hi; rewrite flayE big_valid_seq=>/andP [/allP /(_ i Hi)]. Qed. Lemma dom_flay (ts : seq (tree ptr)) x : @@ -379,7 +349,7 @@ Lemma valid_flay_fset (ts : seq (tree ptr)) : valid (flay ts) = valid (fset ts) && all (fun t => all (fun i => i != null) (preorder t)) ts. Proof. -elim: ts=>[|t ts IH] //=; first by rewrite valid_unit. +elim: ts=>[|t ts IH] //=. rewrite !validUnAE IH -!andbA valid_tlay /=. case V1 : (valid _)=>//=. case V2 : (valid _)=>/=; last by rewrite andbF. @@ -395,7 +365,7 @@ Lemma dom_flay_fset (ts : seq (tree ptr)) : all (fun t => all (fun i => i != null) (preorder t)) ts -> dom (flay ts) = dom (fset ts). Proof. -move=>A; apply/dom2E=>x; rewrite dom_flayE dom_fsetE !inE valid_flay_fset A /=. +move=>A; apply/domE=>x; rewrite dom_flayE dom_fsetE !inE valid_flay_fset A /=. by rewrite andbT. Qed. @@ -462,7 +432,7 @@ Lemma dom_flay_big (ts : seq (tree ptr)) : valid (flay ts) -> dom (flay ts) =i \big[cat/[::]]_(t <- ts) preorder t. Proof. -move=>V x; rewrite flayE big_domUn inE -flayE V big_cat_mem_seq /=. +move=>V x; rewrite flayE big_domUn inE -flayE V big_cat_mem_has /=. by apply: eq_in_has=>i H; rewrite dom_tlay ?in_preorder ?(valid_flay_tlay V H). Qed. @@ -499,7 +469,6 @@ apply/In_fsetP; rewrite Vs; apply/hasP; exists j=>//. by rewrite P E (flay_mem_eq V Xi Ti Xj Tj) eqxx. Qed. - Lemma froot_loop x r ts : (x, r) \In fset ts -> find x (flay ts) = Some (idyn x) -> @@ -509,7 +478,7 @@ elim: ts r=>[|t ts IH] r //= /In_find H1 H2. move: (dom_valid (find_some H1)) (dom_valid (find_some H2))=>V1 V2. move: H1 H2; rewrite !findUnL ?(dom_tset,dom_tlay,validL V1,validL V2) //. case A: (x \in t). - rewrite find_tset ifT; first by case=>E; move/tlay_rt_loop=>K; rewrite K -E. +- rewrite find_tset ifT; first by case=>E; move/tlay_rt_loop=>K; rewrite K -E. apply/andP; split=>//; first by apply: validL V1. by rewrite -In_find; apply: IH. Qed. @@ -528,17 +497,17 @@ Proof. by rewrite flayE big_cons big_nil unitR. Qed. Lemma flay_uniq_ts ts : valid (flay ts) -> uniq ts. Proof. -move=>/flay_uniq/uniq_big_catP [_ H _]; apply: count_mem_uniq=>t. +move=>/flay_uniq/uniq_big_catE [_ H _]; apply: count_mem_uniq=>t. case T : (t \in ts); last by apply/count_memPn; apply: negbT. by apply: (H t (rt t))=>//; rewrite in_preorder rt_in. Qed. -Lemma nochange_mapv (K : ordType) (V : eqType) (m : union_mapPCM K V) b x : +Lemma nochange_mapv (K : ordType) (V : eqType) (m : umap K V) b x : valid m -> x \notin range m -> mapv [fun v => v with x |-> b] m = m. Proof. -move=>W /negP R; apply: umem_eq=>[|//|[k v]]; first by rewrite valid_omf. +move=>W /negP R; apply: umem_eq=>[|//|[k v]]; first by rewrite pfV. rewrite In_omapX /=; split=>[[w]|H]. - by case: (w =P x)=>[->{w} /mem_range/R|_ /[swap][[]->]//]. by exists v=>//; case: (v =P x) H=>// -> /mem_range/R. @@ -549,13 +518,11 @@ Lemma change_tset ta a b: mapv [fun v => v with a |-> b] (tset ta a) = tset ta b. Proof. elim/tree_ind2: ta a=>c ts IH a //; rewrite !tsetE /= => V. -rewrite omapVUn omapPt /= eq_refl omapVUnBig !(validX V). +rewrite omapVUn omapPt /= eq_refl big_omapVUn !(validX V). congr (_ \+ _); apply: eq_big_seq=>i /[dup] K /mem_seqP K'. by rewrite IH // (big_validV (validX V) K'). Qed. - - (*******************) (* Shape predicate *) (*******************) @@ -568,10 +535,11 @@ Proof. by case=>ts [->->]; rewrite valid_flay_fset=>/andP []. Qed. (*******) (* NEW *) (*******) + (* Creates a new equivalence class with a single element *) Program Definition newT : - {m}, STsep (shape m, [vfun r => shape (r \\-> r \+ m)]) := + STsep {m} (shape m, [vfun r => shape (r \\-> r \+ m)]) := Do (p <-- alloc null; p ::= p;; ret p). @@ -580,21 +548,18 @@ case=>m [] h [ts] [->-> V]. step=>p. do !step. move=>V2. by exists (TNode p nil :: ts); split. Qed. -Opaque newT. - (********) (* FIND *) (********) + (* Returns the canonical representative of the equivalence class of an element*) Definition find_tp (x : ptr) := - {rs r}, STsep (fun h => shape rs h /\ (x, r) \In rs, + STsep {rs r} (fun h => shape rs h /\ (x, r) \In rs, [vfun res h => shape rs h /\ res = r]). - - Program Definition find1 (x : ptr) : find_tp x := - Do (let root := Fix (fun (go : forall x, find_tp x) (x : ptr) => + Do (let root := ffix (fun (go : forall x, find_tp x) (x : ptr) => Do (p <-- !x; if x == p then ret p else go p)) in root x). @@ -612,18 +577,16 @@ apply: [gE fset ts, r]=>//=. by do !split=>//=; exists ts; split. Qed. -Opaque find1. - (*********) (* UNION *) (*********) -(* Joins the equivalence classes of the two arguments *) -Definition union_tp (x y : ptr) := {rx ry m}, - STsep (fun h => [/\ shape m h, (x, rx) \In m & (y, ry) \In m], - [vfun res h => shape (mapv [fun v => v with rx |-> ry] m) h /\ - res = ry]). +(* Joins the equivalence classes of the two arguments *) +Definition union_tp (x y : ptr) := STsep {rx ry m} + (fun h => [/\ shape m h, (x, rx) \In m & (y, ry) \In m], + [vfun res h => shape (mapv [fun v => v with rx |-> ry] m) h /\ + res = ry]). Program Definition union (x y : ptr) : union_tp x y := Do (x_rt <-- find1 x; @@ -683,16 +646,12 @@ rewrite range_tset // mem_seq1; case: eqP =>// eqAB. rewrite tsetE -rtB joinA big_filter //; congr (_ \+ _). (*Case 2.3: tset of trees different than a and b don't change *) apply: nochange_mapv =>//; apply/negP; move/mem_rangeX; case=>k H. -move: (@bigInXP ptr_ordType _ _ (union_mapUMC ptr_ordType ptr) (tree ptr) (fun i0 => tset i0 (rt i0)))=>//=. -move/(_ _ _ _ _ H); case=>j [/mem_seqP X1 /andP [X2 X3]]. +case: (bigInXP H)=>j [/mem_seqP X1 /andP [X2 X3]]. move/mem_range; rewrite range_tset //; last by apply: valid_fset_tset X1. rewrite inE => /eqP X4; move/negP: X2; apply; apply/eqP. by apply: flay_mem_eq X1 B J=>//; rewrite X4 rt_in. Qed. -Opaque union. - - (*********) (* Tests *) (*********) @@ -719,8 +678,8 @@ by move/shapeV: X; rewrite invalidX. Qed. Program Definition test2 (x: ptr): - {y}, STsep (fun h => shape (x \\-> y \+ y \\-> y) h, - [vfun res h => shape (x \\-> y \+ y \\-> y) h /\ res = y]) := + STsep {y} (fun h => shape (x \\-> y \+ y \\-> y) h, + [vfun res h => shape (x \\-> y \+ y \\-> y) h /\ res = y]) := Do (res <-- find1 x; ret res). Next Obligation. diff --git a/htt/domain.v b/htt/domain.v new file mode 100644 index 0000000..8ee1d97 --- /dev/null +++ b/htt/domain.v @@ -0,0 +1,1031 @@ +(* +Copyright 2019 IMDEA Software Institute +Licensed under the Apache License, Version 2.0 (the "License"); +you may not use this file except in compliance with the License. +You may obtain a copy of the License at + http://www.apache.org/licenses/LICENSE-2.0 +Unless required by applicable law or agreed to in writing, software +distributed under the License is distributed on an "AS IS" BASIS, +WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. +See the License for the specific language governing permissions and +limitations under the License. +*) + +From HB Require Import structures. +From Coq Require Import ssreflect ssrbool ssrfun. +From mathcomp Require Import ssrnat eqtype. +From pcm Require Import options axioms pred prelude. + +(**************************************************) +(* This file develops some basic domain theory of *) +(* posets, complete lattices, complete partial *) +(* orders (CPOs). It developes the standard fixed *) +(* point theorems: Knaster-Tarski's for complete *) +(* lattices, and Kleene's for CPOs. *) +(**************************************************) + +(**********) +(* Posets *) +(**********) + +Definition poset_axiom T (leq : T -> T -> Prop) := + [/\ forall x, leq x x, + forall x y, leq x y -> leq y x -> x = y & + forall x y z, leq x y -> leq y z -> leq x z]. + +HB.mixin Record isPoset T := { + poset_leq : T -> T -> Prop; + poset_subproof : poset_axiom poset_leq}. + +#[short(type="poset")] +HB.structure Definition Poset := {T of isPoset T}. + +Notation "x <== y" := (poset_leq x y) (at level 70). + +Section Repack. +Variable T : poset. +Lemma poset_refl (x : T) : x <== x. +Proof. by case: (@poset_subproof T). Qed. +Lemma poset_asym (x y : T) : x <== y -> y <== x -> x = y. +Proof. by case: (@poset_subproof T)=>_ H _; apply: H. Qed. +Lemma poset_trans (y x z : T) : x <== y -> y <== z -> x <== z. +Proof. by case: (@poset_subproof T)=>_ _; apply. Qed. +End Repack. + +#[export] Hint Resolve poset_refl : core. + +Add Parametric Relation (T : poset) : T (@poset_leq T) + reflexivity proved by (@poset_refl _) + transitivity proved by (fun x y => @poset_trans _ y x) as poset_rel. + +(**********************) +(* monotone functions *) +(**********************) + +Definition monofun_axiom (T1 T2 : poset) (f : T1 -> T2) := + forall x y, x <== y -> f x <== f y. + +HB.mixin Record isMonoFun (T1 T2 : poset) (f : T1 -> T2) := { + monofun_subproof : monofun_axiom f}. + +#[short(type="mono_fun")] +HB.structure Definition MonoFun (T1 T2 : poset) := + {f of isMonoFun T1 T2 f}. + +Lemma monofunP (T1 T2 : poset) (f : mono_fun T1 T2) x y : + x <== y -> f x <== f y. +Proof. exact: monofun_subproof. Qed. + +(**************************) +(* some basic definitions *) +(**************************) + +Section IdealDef. +Variable T : poset. + +Structure ideal (P : T) := Ideal {id_val : T; id_pf : id_val <== P}. + +(* Changing the type of the ideal *) + +Lemma relaxP (P1 P2 : T) : P1 <== P2 -> forall p, p <== P1 -> p <== P2. +Proof. by move=>H1 p H2; apply: poset_trans H1. Qed. + +Definition relax (P1 P2 : T) (x : ideal P1) (pf : P1 <== P2) := + Ideal (relaxP pf (id_pf x)). + +End IdealDef. + +(***********************) +(* poset constructions *) +(***********************) + +Section SubPoset. +Variables (T : poset) (s : Pred T). +Local Notation tp := {x : T | x \In s}. + +Definition sub_leq (p1 p2 : tp) := sval p1 <== sval p2. + +Lemma sub_is_poset : poset_axiom sub_leq. +Proof. +split=>[x|[x Hx][y Hy]|[x Hx][y Hy][z Hz]]; rewrite /sub_leq //=. +- move=>H1 H2; rewrite -(poset_asym H1 H2) in Hy *. + by rewrite (pf_irr Hx). +by apply: poset_trans. +Qed. +HB.instance Definition _ := isPoset.Build tp sub_is_poset. +End SubPoset. + +(* pairing of posets *) + +Section ProdPoset. +Variables A B : poset. +Local Notation tp := (A * B)%type. + +Definition poset_prod_leq := + [fun p1 p2 : tp => p1.1 <== p2.1 /\ p1.2 <== p2.2]. + +Lemma prod_is_poset : poset_axiom poset_prod_leq. +Proof. +split=>[//|[x1 x2][y1 y2]|[x1 x2][y1 y2][z1 z2]] /=. +- by case=>H1 H2 [H3 H4]; congr (_, _); apply: poset_asym. +case=>H1 H2 [H3 H4]; split; +by [apply: poset_trans H3|apply: poset_trans H4]. +Qed. +HB.instance Definition _ := isPoset.Build tp prod_is_poset. +End ProdPoset. + +(* functions into poset form a poset *) + +Section FunPoset. +Variable (A : Type) (B : poset). +Local Notation tp := (A -> B). + +Definition poset_fun_leq := [fun p1 p2 : tp => forall x, p1 x <== p2 x]. + +Lemma fun_is_poset : poset_axiom poset_fun_leq. +Proof. +split=>[x|x y H1 H2|x y z H1 H2 t] //=. +- by apply: fext=>z; apply: poset_asym; [apply: H1|apply: H2]. +by apply: poset_trans (H2 t). +Qed. +HB.instance Definition _ := isPoset.Build tp fun_is_poset. +End FunPoset. + +(* dependent functions into a poset form a poset *) + +Section DFunPoset. +Variables (A : Type) (B : A -> poset). +Local Notation tp := (forall x, B x). + +Definition poset_dfun_leq := [fun p1 p2 : tp => forall x, p1 x <== p2 x]. + +Lemma dfun_is_poset : poset_axiom poset_dfun_leq. +Proof. +split=>[x|x y H1 H2|x y z H1 H2 t] //. +- by apply: fext=>z; apply: poset_asym; [apply: H1|apply: H2]. +by apply: poset_trans (H2 t). +Qed. + +(* no instance declaration, since it's keyed on forall *) +(* and forall is not a symbol *) +Definition dfun_poset_mix := isPoset.Build tp dfun_is_poset. +Definition dfun_poset : poset := Poset.pack_ dfun_poset_mix. +End DFunPoset. + +(* ideal of poset is poset *) + +Section IdealPoset. +Variable (T : poset) (P : T). + +Definition poset_ideal_leq := + [fun p1 p2 : ideal P => id_val p1 <== id_val p2]. + +Lemma ideal_is_poset : poset_axiom poset_ideal_leq. +Proof. +split=>[[x]|[x1 H1][x2 H2]|[x1 H1][x2 H2][x3 H3]] //=. +- move=>H3 H4; rewrite (poset_asym H3 H4) in H1 H2 *. + by rewrite (pf_irr H1). +by apply: poset_trans. +Qed. +HB.instance Definition _ := isPoset.Build (ideal P) ideal_is_poset. +End IdealPoset. + +(* Prop is poset *) +Definition poset_prop_leq := [fun p1 p2 : Prop => p1 -> p2]. +Lemma prop_is_poset : poset_axiom poset_prop_leq. +Proof. by split=>[x|x y H1 H2|x y z H1 H2 /H1] //; apply: pext. Qed. +HB.instance Definition _ := isPoset.Build Prop prop_is_poset. + + +(*********************) +(* Complete lattices *) +(*********************) + +Definition lattice_axiom (T : poset) (sup : Pred T -> T) := + [/\ (* sup is upper bound *) + forall (s : Pred T) x, x \In s -> x <== sup s & + (* sup is least upper bound *) + forall (s : Pred T) x, + (forall y, y \In s -> y <== x) -> sup s <== x]. + +HB.mixin Record isLattice T of Poset T := { + sup : Pred T -> T; + lattice_subproof : lattice_axiom sup}. + +#[short(type="lattice")] +HB.structure Definition Lattice := {T of Poset T & isLattice T}. + +Section Repack. +Variable T : lattice. +Lemma supP (s : Pred T) x : x \In s -> x <== sup s. +Proof. by case: (@lattice_subproof T)=>H _; apply: H. Qed. +Lemma supM (s : Pred T) x : (forall y, y \In s -> y <== x) -> sup s <== x. +Proof. by case: (@lattice_subproof T)=>_; apply. Qed. +End Repack. + +(* infima follow *) +Section Infimum. +Variable T : lattice. + +Definition inf (s : Pred T) := + sup [Pred x : T | forall y, y \In s -> x <== y]. + +Lemma infP s x : x \In s -> inf s <== x. +Proof. by move=>H; apply: supM=>y; apply. Qed. + +Lemma infM s x : (forall y, y \In s -> x <== y) -> x <== inf s. +Proof. by apply: supP. Qed. + +End Infimum. + +(* least and greatest fixed points by Knaster-Tarski construction *) +Section KnasterTarskiDefs. +Variable T : lattice. + +Definition lbot : T := sup Pred0. + +Definition tarski_lfp (f : T -> T) := inf [Pred x : T | f x <== x]. +Definition tarski_gfp (f : T -> T) := sup [Pred x : T | x <== f x]. + +Definition sup_closed (T : lattice) := + [Pred s : Pred T | forall d, d <=p s -> sup d \In s]. + +Definition sup_closure (T : lattice) (s : Pred T) := + [Pred p : T | forall t : Pred T, s <=p t -> t \In sup_closed T -> p \In t]. + +End KnasterTarskiDefs. + +Arguments lbot {T}. +Arguments sup_closed {T}. +Arguments sup_closure [T] s. +Prenex Implicits sup_closed sup_closure. + +Section BasicProperties. +Variable T : lattice. + +Lemma sup_mono (s1 s2 : Pred T) : s1 <=p s2 -> sup s1 <== sup s2. +Proof. by move=>H; apply: supM=>y /H/supP. Qed. + +Lemma supE (s1 s2 : Pred T) : s1 =p s2 -> sup s1 = sup s2. +Proof. by move=>H; apply: poset_asym; apply: supM=>y /H/supP. Qed. + +(* Knaster-Tarski theorem *) +Lemma tarski_lfp_fixed (f : mono_fun T T) : + f (tarski_lfp f) = tarski_lfp f. +Proof. +suff L : f (tarski_lfp f) <== tarski_lfp f. +- by apply: poset_asym=>//; apply/infP/monofunP. +by apply: infM=>x L; apply: poset_trans (L); apply/monofunP/infP. +Qed. + +Lemma tarski_lfp_least f (x : T) : f x <== x -> tarski_lfp f <== x. +Proof. by move=>H; apply: infP. Qed. + +Lemma tarski_gfp_fixed (f : mono_fun T T) : + f (tarski_gfp f) = tarski_gfp f. +Proof. +suff L: tarski_gfp f <== f (tarski_gfp f). +- by apply: poset_asym=>//; apply/supP/monofunP. +by apply: supM=>x L; apply: poset_trans (L) _; apply/monofunP/supP. +Qed. + +Lemma tarski_gfp_greatest f (x : T) : x <== f x -> x <== tarski_gfp f. +Proof. by move=>H; apply: supP. Qed. + +Lemma tarski_lfp_mono (f1 : T -> T) (f2 : mono_fun T T) : + f1 <== f2 -> tarski_lfp f1 <== tarski_lfp f2. +Proof. +move=>H; apply: infP; apply: poset_trans (H (tarski_lfp f2)) _. +by rewrite tarski_lfp_fixed. +Qed. + +Lemma tarski_gfp_mono (f1 : mono_fun T T) (f2 : T -> T) : + (f1 : T -> T) <== f2 -> tarski_gfp f1 <== tarski_gfp f2. +Proof. +move=>H; apply/supP/poset_trans/(H (tarski_gfp f1)). +by rewrite tarski_gfp_fixed. +Qed. + +(* closure contains s *) +Lemma sup_clos_sub (s : Pred T) : s <=p sup_closure s. +Proof. by move=>p H1 t H2 H3; apply: H2 H1. Qed. + +(* closure is smallest *) +Lemma sup_clos_min (s : Pred T) : + forall t, s <=p t -> sup_closed t -> sup_closure s <=p t. +Proof. by move=>t H1 H2 p; move/(_ _ H1 H2). Qed. + +(* closure is closed *) +Lemma sup_closP (s : Pred T) : sup_closed (sup_closure s). +Proof. +move=>d H1 t H3 H4; move: (sup_clos_min H3 H4)=>{}H3. +by apply: H4=>x; move/H1; move/H3. +Qed. + +Lemma sup_clos_idemp (s : Pred T) : sup_closed s -> sup_closure s =p s. +Proof. by move=>p; split; [apply: sup_clos_min | apply: sup_clos_sub]. Qed. + +Lemma sup_clos_mono (s1 s2 : Pred T) : + s1 <=p s2 -> sup_closure s1 <=p sup_closure s2. +Proof. +move=>H1; apply: sup_clos_min (@sup_closP s2). +by move=>t /H1; apply: sup_clos_sub. +Qed. + +End BasicProperties. + +(* lattice constructions *) + +(* subset lattice *) +Section SubLattice. +Variables (T : lattice) (s : Pred T) (C : sup_closed s). +Local Notation tp := {x : T | x \In s}. + +Definition lattice_sub_sup' (u : Pred tp) : T := + sup [Pred x : T | exists y, y \In u /\ x = sval y]. +Lemma lattice_sub_supX (u : Pred tp) : lattice_sub_sup' u \In s. +Proof. by apply: C=>x [][y H][_] ->. Qed. +Definition lattice_sub_sup (u : Pred tp) : tp := + exist _ (lattice_sub_sup' u) (lattice_sub_supX u). + +Lemma sub_is_lattice : lattice_axiom lattice_sub_sup. +Proof. +split=>u x H; first by apply: supP; exists x. +by apply: supM=>y [z][H1] ->; apply: H H1. +Qed. +HB.instance Definition _ := isLattice.Build tp sub_is_lattice. +End SubLattice. + +(* product *) +Section ProdLattice. +Variables (A B : lattice). +Local Notation tp := (A * B)%type. + +Definition lattice_prod_sup (s : Pred tp) : tp := + (sup [Pred p | exists f, p = f.1 /\ f \In s], + sup [Pred p | exists f, p = f.2 /\ f \In s]). + +Lemma prod_is_lattice : lattice_axiom lattice_prod_sup. +Proof. +split=>s p H; first by split; apply: supP; exists p. +by split; apply: supM=>y [f][->]; case/H. +Qed. +HB.instance Definition _ := isLattice.Build tp prod_is_lattice. +End ProdLattice. + +(* functions into latice form lattice *) +Section FunLattice. +Variables (A : Type) (B : lattice). +Local Notation tp := (A -> B). + +Definition lattice_fun_sup (s : Pred tp) : tp := + fun x => sup [Pred p | exists f, f \In s /\ p = f x]. + +Lemma fun_is_lattice : lattice_axiom lattice_fun_sup. +Proof. +split=>s p H x; first by apply: supP; exists p. +by apply: supM=>y [f][H1] ->; apply: H. +Qed. +HB.instance Definition _ := isLattice.Build tp fun_is_lattice. +End FunLattice. + +(* dependent functions into a lattice form a lattice *) +Section DFunLattice. +Variables (A : Type) (B : A -> lattice). +Local Notation tp := (dfun_poset B). + +Definition lattice_dfun_sup (s : Pred tp) : tp := + fun x => sup [Pred p | exists f, f \In s /\ p = f x]. + +Lemma dfun_is_lattice : lattice_axiom lattice_dfun_sup. +Proof. +split=>s p H x; first by apply: supP; exists p. +by apply: supM=>y [f][H1] ->; apply: H. +Qed. + +Definition dfun_lattice_mix := isLattice.Build tp dfun_is_lattice. +Definition dfun_lattice : lattice := Lattice.pack_ dfun_lattice_mix. +End DFunLattice. + +(* applied sup equals the sup of applications *) +Lemma sup_appE A (B : lattice) (s : Pred (A -> B)) (x : A) : + sup s x = sup [Pred y : B | exists f, f \In s /\ y = f x]. +Proof. by []. Qed. + +Lemma sup_dappE A (B : A -> lattice) (s : Pred (dfun_lattice B)) (x : A) : + sup s x = sup [Pred y : B x | exists f, f \In s /\ y = f x]. +Proof. by []. Qed. + +(* ideal of a lattice forms a lattice *) +Section IdealLattice. +Variables (T : lattice) (P : T). + +Definition lattice_ideal_sup' (s : Pred (ideal P)) := + sup [Pred x | exists p, p \In s /\ id_val p = x]. +Lemma lattice_ideal_supP' (s : Pred (ideal P)) : lattice_ideal_sup' s <== P. +Proof. by apply: supM=>y [[x]] H /= [_] <-. Qed. +Definition lattice_ideal_sup (s : Pred (ideal P)) := + Ideal (lattice_ideal_supP' s). + +Lemma ideal_is_lattice : lattice_axiom lattice_ideal_sup. +Proof. +split=>s p H; first by apply: supP; exists p. +by apply: supM=>y [q][H1] <-; apply: H. +Qed. +HB.instance Definition _ := isLattice.Build (ideal P) ideal_is_lattice. +End IdealLattice. + +(* Prop is a lattice *) +Definition lattice_prop_sup (s : Pred Prop) : Prop := + exists p, p \In s /\ p. +Lemma prop_is_lattice : lattice_axiom lattice_prop_sup. +Proof. by split=>s p; [exists p|move=>H [r][]; move/H]. Qed. +HB.instance Definition _ := isLattice.Build Prop prop_is_lattice. + +(****************) +(* Monotonicity *) +(****************) + +Section MonoCanonicals. +Variables T T1 T2 T3 S1 S2 : poset. + +Lemma idfun_is_mono : monofun_axiom (@idfun T). Proof. by []. Qed. +HB.instance Definition _ := isMonoFun.Build T T idfun idfun_is_mono. + +Definition const_fun (y : T2) (_ : T1) := y. +Lemma const_is_mono (y : T2) : monofun_axiom (const_fun y). Proof. by []. Qed. +HB.instance Definition _ y := + isMonoFun.Build T1 T2 (const_fun y) (const_is_mono y). + +Lemma fst_is_mono : monofun_axiom (@fst T1 T2). +Proof. by case=>x1 x2 [y1 y2][]. Qed. +HB.instance Definition _ := isMonoFun.Build (T1*T2)%type T1 fst fst_is_mono. + +Lemma snd_is_mono : monofun_axiom (@snd T1 T2). +Proof. by case=>x1 x2 [y1 y2][]. Qed. +HB.instance Definition _ := isMonoFun.Build (T1*T2)%type T2 snd snd_is_mono. + +Definition diag (x : T) := (x, x). +Lemma diag_is_mono : monofun_axiom diag. Proof. by []. Qed. +HB.instance Definition _ := isMonoFun.Build T (T*T)%type diag diag_is_mono. + +Lemma sval_is_mono (P : Pred T) : monofun_axiom (@sval T P). +Proof. by []. Qed. +HB.instance Definition _ P := + isMonoFun.Build {x : T | P x} T sval (@sval_is_mono P). + +Definition app A B (x : A) := fun f : A -> B => f x. +Lemma app_is_mono A (x : A) : monofun_axiom (@app A T x). +Proof. by move=>f1 f2; apply. Qed. +HB.instance Definition _ A x := + isMonoFun.Build (A -> T) T (@app A T x) (app_is_mono x). + +Definition dapp A (B : A -> poset) (x : A) := fun f : dfun_poset B => f x. +Lemma dapp_is_mono A (B : A -> poset) (x : A) : monofun_axiom (@dapp A B x). +Proof. by move=>f1 f2; apply. Qed. +HB.instance Definition _ A B x := + isMonoFun.Build (dfun_poset B) (B x) (@dapp A B x) (dapp_is_mono x). + +Section Comp. +Variables (f1 : mono_fun T2 T1) (f2 : mono_fun T3 T2). +Lemma comp_is_mono : monofun_axiom (f1 \o f2). +Proof. by move=>x y H; apply/monofunP/monofunP. Qed. +HB.instance Definition _ := isMonoFun.Build T3 T1 (f1 \o f2) comp_is_mono. +End Comp. + +Section Prod. +Variables (f1 : mono_fun S1 T1) (f2 : mono_fun S2 T2). +Lemma funprod_is_mono : monofun_axiom (f1 \* f2). +Proof. by case=>x1 x2 [y1 y2][/= H1 H2]; split=>/=; apply: monofunP. Qed. +HB.instance Definition _ := + isMonoFun.Build (S1 * S2)%type (T1 * T2)%type (f1 \* f2) funprod_is_mono. +End Prod. +End MonoCanonicals. + +Prenex Implicits diag app. + +(**********) +(* Chains *) +(**********) + +Definition chain_axiom (T : poset) (s : Pred T) := + (exists d, d \In s) /\ + (forall x y, x \In s -> y \In s -> x <== y \/ y <== x). + +HB.mixin Record isChain (T : poset) (s : Pred T) := { + chain_subproof : chain_axiom s}. + +#[short(type="chain")] +HB.structure Definition Chain (T : poset) := {s of isChain T s}. + +Arguments chain T%_type. + +(* \In notation *) +Definition chain_pred (T : poset) (x : chain T) : Pred T := x. +Canonical chainPredType (T : poset) := PropPredType (@chain_pred T). + +Lemma chainE (T : poset) (s1 s2 : chain T) : + s1 = s2 <-> chain_pred s1 =p chain_pred s2. +Proof. +split=>[->//|]; case: s1 s2=>s1 H1 [s2 H2] /= E; move: H1 H2. +suff: s1 = s2. +- move=>-> [[H1]][[H2]]; congr (Chain.Pack (Chain.Class _)). + by rewrite (pf_irr H1). +by apply: fext=>x; apply: pext; split=> /E. +Qed. + +Section ImageChain. +Variables (T1 T2 : poset) (f : mono_fun T1 T2) (s : chain T1). + +Lemma image_is_chain : chain_axiom (Image f s). +Proof. +case: s=>p [[[[d H1] H2]]] /=. +split=>[|x y]; first by exists (f d); exists d. +case=>x1 -> H3 [y1 -> H4]; rewrite -!toPredE /= in H3 H4. +by case: (H2 x1 y1 H3 H4)=>L; [left | right]; apply: monofunP L. +Qed. +HB.instance Definition _ := isChain.Build T2 (Image f s) image_is_chain. +End ImageChain. + +Lemma id_chainE T (s : chain T) f : f =1 idfun -> Image f s =p s. +Proof. +move=>H x; split; first by case=>y ->; rewrite H. +by exists x=>//; rewrite H. +Qed. + +Section ChainConst. +Variables (T1 T2 : poset) (y : T2). + +Definition const_chain : Pred _ := xPred1 y. +Lemma const_is_chain : chain_axiom const_chain. +Proof. by split; [exists y | move=>x1 x2 ->->; left]. Qed. +HB.instance Definition _ := isChain.Build T2 const_chain const_is_chain. + +Lemma const_chainE (s : chain T1) : Image (const_fun y) s =p const_chain. +Proof. +move=>z1; split; first by case=>z2 ->. +by case: s=>s [[[[d H1]] H2]] /= <-; exists d. +Qed. +End ChainConst. + +Section ChainCompose. +Variables T1 T2 T3 : poset. +Variables (f1 : mono_fun T2 T1) (f2 : mono_fun T3 T2). +Variable s : chain T3. +Lemma comp_chainE : Image f1 (Image f2 s) =p Image (f1 \o f2) s. +Proof. by move=>x; apply: ImageIm. Qed. +End ChainCompose. + +Section ProjChain. +Variables T1 T2 : poset. +Variables s : chain (T1 * T2). +Definition proj1_chain : Pred T1 := Image fst s. +HB.instance Definition _ := Chain.copy proj1_chain proj1_chain. +Definition proj2_chain : Pred T2 := Image snd s. +HB.instance Definition _ := Chain.copy proj2_chain proj2_chain. +End ProjChain. + +Section DiagChain. +Variables (T : poset) (s : chain T). +Definition diag_chain : Pred (T * T) := Image diag s. +HB.instance Definition _ := Chain.copy diag_chain diag_chain. +End DiagChain. + + +(*********) +(* CPO's *) +(*********) + +Definition cpo_axiom (T : poset) (bot : T) (lim : chain T -> T) := + [/\ forall x, bot <== x, + forall (s : chain T) x, x \In s -> x <== lim s & + forall (s : chain T) x, + (forall y, y \In s -> y <== x) -> lim s <== x]. + +HB.mixin Record isCPO T of Poset T := { + bot : T; + lim_op : chain T -> T; + cpo_subproof: cpo_axiom bot lim_op}. + +#[short(type="cpo")] +HB.structure Definition CPO := {T of Poset T & isCPO T}. + +Definition limx (D : cpo) (s : chain D) of phantom (Pred _) s := lim_op s. +Notation lim s := (limx (Phantom (Pred _) s)). + +Section Repack. +Variable D : cpo. +Lemma botP (x : D) : bot <== x. +Proof. by case: (@cpo_subproof D). Qed. +Lemma limP (s : chain D) x : x \In s -> x <== lim s. +Proof. by case: (@cpo_subproof D)=>_ H _; apply: H. Qed. +Lemma limM (s : chain D) x : (forall y, y \In s -> y <== x) -> lim s <== x. +Proof. by case: (@cpo_subproof D)=>_ _; apply. Qed. +End Repack. + +#[export] Hint Resolve botP : core. + +Lemma limE (D : cpo) (s1 s2 : chain D) : s1 =p s2 -> lim s1 = lim s2. +Proof. by move/chainE=>->. Qed. + +(* common chain constructions *) + +(* adding bot to a chain *) + +Definition lift (D : cpo) (s : Pred D) : Pred D := + fun x => x = bot \/ x \In s. + +Lemma lift_is_chain (D : cpo) (s : chain D) : chain_axiom (lift s). +Proof. +case: s=>p [[[[d H1]] H2]] /=. +split=>[|x y]; first by exists d; right. +by case=>[->|H3][->|H4]; auto. +Qed. + +HB.instance Definition _ (D : cpo) (s : chain D) := + isChain.Build D (lift s) (@lift_is_chain D s). + + +(****************************) +(* common cpo constructions *) +(****************************) + +(* products *) +Section ProdCPO. +Variables (A B : cpo). +Definition cpo_prod_bot := (@bot A, @bot B). +Definition cpo_prod_lim (s : chain (A * B)) := + (lim (Image fst s), lim (Image snd s)). + +Lemma prod_is_cpo : cpo_axiom cpo_prod_bot cpo_prod_lim. +Proof. +split=>[x|s x|s x]. +- by split; apply: botP. +- by split; apply: limP; exists x. +by split; apply: limM=>y [z ->]; case/H. +Qed. +HB.instance Definition _ := isCPO.Build (A * B)%type prod_is_cpo. +End ProdCPO. + +(* functions *) +Section FunCPO. +Variable (A : Type) (B : cpo). +Definition cpo_fun_bot (x : A) := @bot B. +Definition cpo_fun_lim (s : chain (A -> B)) x := + lim (Image (app x) s). + +Lemma fun_is_cpo : cpo_axiom cpo_fun_bot cpo_fun_lim. +Proof. +split=>[f|s f|s f]. +- by move=>y; apply: botP. +- by move=>H t; apply: limP; exists f. +by move=>H1 t; apply: limM=>y [g ->] H2; apply: H1. +Qed. +HB.instance Definition _ := isCPO.Build (A -> B) fun_is_cpo. +End FunCPO. + +(* dependent functions *) +Section DFunCPO. +Variables (A : Type) (B : A -> cpo). + +Definition cpo_dfun_bot : dfun_poset B := fun x => bot. +Definition cpo_dfun_lim (s : chain (dfun_poset B)) : dfun_poset B := + fun x => lim (Image (dapp x) s). + +Lemma dfun_is_cpo : cpo_axiom cpo_dfun_bot cpo_dfun_lim. +Proof. +split=>[x|s x|s x]. +- by move=>y; apply: botP. +- by move=>H t; apply: limP; exists x. +by move=>H1 t; apply: limM=>y [f ->] H2; apply: H1. +Qed. + +Definition dfun_cpo_mix := isCPO.Build (dfun_poset B) dfun_is_cpo. +Definition dfun_cpo : cpo := CPO.pack_ dfun_cpo_mix. +End DFunCPO. + +(* Prop *) +Definition cpo_prop_bot : Prop := False. +Definition cpo_prop_lim (s : chain Prop) : Prop := + exists p, p \In s /\ p. +Lemma prop_is_cpo : cpo_axiom cpo_prop_bot cpo_prop_lim. +Proof. by split=>[x|s p|s p H] //=; [exists p|case=>r []; move/H]. Qed. +HB.instance Definition _ := isCPO.Build Prop prop_is_cpo. + +(* Pred *) +HB.instance Definition _ A := CPO.copy (Pred A) _. + +(* every complete lattice is cpo *) +Section LatticeCPO. +Variable A : lattice. + +Definition lat_bot : A := lbot. +Definition lat_lim (s : Pred A) : A := sup s. + +Lemma lat_is_cpo : cpo_axiom lat_bot lat_lim. +Proof. by split=>[x|s x|s x]; [apply: supM|apply: supP|apply: supM]. Qed. + +(* no instance declaration as nothing to *) +(* hang the new structure onto *) +Definition lat_cpo_mix := isCPO.Build A lat_is_cpo. +Definition lat_cpo := CPO.pack_ lat_cpo_mix. +End LatticeCPO. + +(* sub-CPO's *) + +(* every chain-closed subset of a cpo is a cpo *) + +Section AdmissibleClosure. +Variable T : cpo. + +Definition chain_closed := + [Pred s : Pred T | + bot \In s /\ forall d : chain T, d <=p s -> lim d \In s]. + +(* admissible closure of s is the smallest closed set containing s *) +(* basically extends s to include the sups of chains *) +Definition chain_closure (s : Pred T) := + [Pred p : T | forall t : Pred T, s <=p t -> chain_closed t -> p \In t]. + +(* admissible closure contains s *) +Lemma chain_clos_sub (s : Pred T) : s <=p chain_closure s. +Proof. by move=>p H1 t H2 H3; apply: H2 H1. Qed. + +(* admissible closure is smallest *) +Lemma chain_clos_min (s : Pred T) t : + s <=p t -> chain_closed t -> chain_closure s <=p t. +Proof. by move=>H1 H2 p; move/(_ _ H1 H2). Qed. + +(* admissible closure is closed *) +Lemma chain_closP (s : Pred T) : chain_closed (chain_closure s). +Proof. +split; first by move=>t _ []. +move=>d H1 t H3 H4; move: (chain_clos_min H3 H4)=>{}H3. +by case: H4=>_; apply=>// x; move/H1; move/H3. +Qed. + +Lemma chain_clos_idemp (s : Pred T) : + chain_closed s -> chain_closure s =p s. +Proof. +move=>p; split; last by apply: chain_clos_sub. +by apply: chain_clos_min=>// /chain_closP. +Qed. + +Lemma chain_clos_mono (s1 s2 : Pred T) : + s1 <=p s2 -> chain_closure s1 <=p chain_closure s2. +Proof. +move=>H1; apply: chain_clos_min (chain_closP s2)=>p H2. +by apply: chain_clos_sub; apply: H1. +Qed. + +(* intersection of admissible sets is admissible *) +Lemma chain_closI (s1 s2 : Pred T) : + chain_closed s1 -> chain_closed s2 -> chain_closed (PredI s1 s2). +Proof. +move=>[H1 S1][H2 S2]; split=>// d H. +by split; [apply: S1 | apply: S2]=>// x; case/H. +Qed. + +End AdmissibleClosure. + +Arguments chain_closed {T}. +Prenex Implicits chain_closed. + +(* diagonal of an admissible set of pairs is admissible *) +Lemma chain_clos_diag (T : cpo) (s : Pred (T * T)) : + chain_closed s -> chain_closed [Pred t : T | (t, t) \In s]. +Proof. +case=>B H1; split=>// d H2; move: (H1 (Image diag d)). +rewrite /limx/=/lim_op /= /cpo_prod_lim /=. +rewrite !(limE (comp_chainE _ _ _)) !(limE (id_chainE _ _)) //=. +by apply; case=>x1 x2 [z ->]; apply: H2. +Qed. + +Section SubCPO. +Variables (D : cpo) (s : Pred D) (C : chain_closed s). +Local Notation tp := {x : D | x \In s}. + +Definition subcpo_bot := exist _ (@bot D) (proj1 C). +Lemma subcpo_limX (u : chain tp) : lim (Image sval u) \In s. +Proof. by case: C u=>P H u; apply: (H)=>t [[y H1 ->]]. Qed. +Definition subcpo_lim (u : chain tp) : tp := + exist _ (lim (Image sval u)) (subcpo_limX u). + +Lemma subcpo_is_cpo : cpo_axiom subcpo_bot subcpo_lim. +Proof. +split=>[x|u x H|u x H] //=. +- by apply: botP. +- by apply: limP; exists x. +by apply: limM=>y [z ->]; apply: H. +Qed. +HB.instance Definition _ := isCPO.Build tp subcpo_is_cpo. +End SubCPO. + +(***********************************************) +(* Continuity and Kleene's fixed point theorem *) +(***********************************************) + +Lemma lim_mono (D : cpo) (s1 s2 : chain D) : + s1 <=p s2 -> lim s1 <== lim s2. +Proof. by move=>H; apply: limM=>y /H/limP. Qed. + +Lemma lim_liftE (D : cpo) (s : chain D) : lim s = lim (lift s). +Proof. +apply: poset_asym; apply: limM=>y H; first by apply: limP; right. +by case: H; [move=>-> | apply: limP]. +Qed. + +(* applied lim equals the lim of applications *) +(* ie., part of continuity of application *) +(* but is so often used, I give it a name *) +Lemma lim_appE A (D : cpo) (s : chain (A -> D)) (x : A) : + lim s x = lim (Image (app x) s). +Proof. by []. Qed. + +Lemma lim_dappE A (D : A -> cpo) (s : chain (dfun_cpo D)) (x : A) : + lim s x = lim (Image (dapp x) s). +Proof. by []. Qed. + +(************************) +(* continuous functions *) +(************************) + +Definition contfun_axiom (D1 D2 : cpo) (f : mono_fun D1 D2) := + forall s : chain D1, f (lim s) <== lim (Image f s). + +HB.mixin Record isContFun (D1 D2 : cpo) (f : D1 -> D2) + of @MonoFun D1 D2 f := {contfun_subproof : contfun_axiom f}. + +#[short(type="cont_fun")] +HB.structure Definition ContFun (D1 D2 : cpo) := + {f of @MonoFun D1 D2 f & isContFun D1 D2 f}. + +Lemma contE (D1 D2 : cpo) (f : cont_fun D1 D2) (s : chain D1) : + f (lim s) = lim (Image f s). +Proof. +apply: poset_asym; first by apply: contfun_subproof. +by apply: limM=>y [x1 ->] /limP/monofunP. +Qed. + +(* just for this proof, we want that nat is a poset under <= *) +(* in other scopes, we'll use the equality as ordering *) +(* Thus, we encapsulate under module. *) + +Module Kleene. + +Lemma nat_is_poset : poset_axiom leq. +Proof. +split=>[|x y H1 H2|x y z] //; last by apply: leq_trans. +by apply: anti_leq; rewrite H1 H2. +Qed. +HB.instance Definition _ := isPoset.Build nat nat_is_poset. + +(* nats form a chain *) +Definition nats : Pred nat := PredT. +Lemma nats_chain_axiom : chain_axiom nats. +Proof. +split=>[|x y _ _]; first by exists 0. +rewrite /poset_leq /= [y <= x]leq_eqVlt. +by case: leqP; [left | rewrite orbT; right]. +Qed. +HB.instance Definition _ := isChain.Build nat nats nats_chain_axiom. + +Section Kleene. +Variables (D : cpo) (f : cont_fun D D). + +Fixpoint pow m := if m is n.+1 then f (pow n) else bot. + +Lemma pow_is_mono : monofun_axiom pow. +Proof. +move=>m n; elim: n m=>[|n IH] m /=; first by case: m. +rewrite {1}/poset_leq /= leq_eqVlt ltnS. +case/orP; first by move/eqP=>->. +move/IH=>{IH} H; apply: poset_trans H _. +by elim: n=>[|n IH] //=; apply: monofunP IH. +Qed. +HB.instance Definition _ := isMonoFun.Build nat D pow pow_is_mono. + +Lemma reindex : Image pow nats =p lift (Image f (Image pow nats)). +Proof. +move=>x; split. +- case; case=>[|n] -> _; first by left. + by right; exists (pow n)=>//; exists n. +case=>/=; first by move=>->; exists 0. +by case=>y -> [n ->]; exists n.+1. +Qed. + +Definition kleene_lfp := lim (Image pow nats). + +Lemma kleene_lfp_fixed : f kleene_lfp = kleene_lfp. +Proof. by rewrite /kleene_lfp contE /= (limE reindex) /= lim_liftE. Qed. + +Lemma kleene_lfp_least x : f x = x -> kleene_lfp <== x. +Proof. +move=>H; apply: limM=>y [n ->] _. +by elim: n=>[|n IH] //=; rewrite -H; apply: monofunP IH. +Qed. + +End Kleene. + +Lemma kleene_lfp_mono (D : cpo) (f1 f2 : cont_fun D D) : + (f1 : D -> D) <== f2 -> + kleene_lfp f1 <== kleene_lfp f2. +Proof. +move=>H; apply: limM=>x [n ->] X. +have N : forall n, pow f1 n <== pow f2 n. +- by elim=>//= m /(monofunP f2); apply: poset_trans (H _). +by apply: poset_trans (N n) _; apply: limP; exists n. +Qed. + +Module Exports. +Notation kleene_lfp := kleene_lfp. +Notation kleene_lfp_fixed := kleene_lfp_fixed. +Notation kleene_lfp_least := kleene_lfp_least. +Notation kleene_lfp_mono := kleene_lfp_mono. +End Exports. + +End Kleene. + +Export Kleene.Exports. + + +(**********************************) +(* Continuity of common functions *) +(**********************************) + +Lemma idfun_is_cont (D : cpo) : contfun_axiom (@idfun D). +Proof. by move=>d; rewrite (limE (id_chainE _ _)). Qed. +HB.instance Definition _ (D : cpo) := + isContFun.Build D D idfun (@idfun_is_cont D). + +Lemma const_is_cont (D1 D2 : cpo) (y : D2) : + contfun_axiom (@const_fun D1 D2 y). +Proof. +by move=>s; apply: limP; case: s=>[p][[[[d H1] H2]]]; exists d. +Qed. +HB.instance Definition _ (D1 D2 : cpo) (y : D2) := + isContFun.Build D1 D2 (const_fun y) (const_is_cont y). + +Section ContCompose. +Variables D1 D2 D3 : cpo. +Variable (f1 : cont_fun D2 D1) (f2 : cont_fun D3 D2). +Lemma comp_is_cont : contfun_axiom (f1 \o f2). +Proof. by move=>s; rewrite /= !contE (limE (comp_chainE _ _ _)). Qed. +HB.instance Definition _ := isContFun.Build D3 D1 (f1 \o f2) comp_is_cont. +End ContCompose. + +Lemma fst_is_cont (D1 D2 : cpo) : contfun_axiom (@fst D1 D2). +Proof. by []. Qed. +HB.instance Definition _ (D1 D2 : cpo) := + isContFun.Build (D1*D2)%type D1 fst (@fst_is_cont D1 D2). + +Lemma snd_is_cont (D1 D2 : cpo) : contfun_axiom (@snd D1 D2). +Proof. by []. Qed. +HB.instance Definition _ (D1 D2 : cpo) := + isContFun.Build (D1*D2)%type D2 snd (@snd_is_cont D1 D2). + +Lemma diag_is_cont (D : cpo) : contfun_axiom (@diag D). +Proof. +move=>d; split=>/=; +by rewrite (limE (comp_chainE _ _ _)) (limE (id_chainE _ _)). +Qed. +HB.instance Definition _ (D : cpo) := + isContFun.Build D (D*D)%type diag (@diag_is_cont D). + +Lemma app_is_cont A (D : cpo) x : contfun_axiom (@app A D x). +Proof. by []. Qed. +HB.instance Definition _ A (D : cpo) x := + isContFun.Build (A -> D) D (@app A D x) (app_is_cont x). + +(* can't reuse dapp and it's proof of monotonicity *) +(* because inheritance doesn't propagate properly *) +(* for dependent functions; so we duplicate *) +(* See mathcomp/finfun.v for solution (dfun vs. dffun). *) +Definition dapp2 A (B : A -> cpo) (x : A) := + fun f : dfun_cpo B => f x. +Lemma dapp2_is_mono A (B : A -> cpo) x : monofun_axiom (@dapp2 A B x). +Proof. by move=>f1 f2; apply. Qed. +HB.instance Definition _ A (B : A -> cpo) x := + isMonoFun.Build (dfun_poset B) (B x) (@dapp2 A B x) (dapp2_is_mono x). +Lemma dapp2_is_cont A (B : A -> cpo) (x : A) : contfun_axiom (@dapp2 A B x). +Proof. +move=>s; have <- // : dapp2 x (lim s) = lim [Image dapp2 x i | i <- s]. +by apply: limE=>/=. +Qed. +HB.instance Definition _ A (B : A -> cpo) x := + isContFun.Build (dfun_cpo B) (B x) (@dapp2 A B x) (dapp2_is_cont x). + +Section ProdCont. +Variables S1 S2 T1 T2 : cpo. +Variables (f1 : cont_fun S1 T1) (f2 : cont_fun S2 T2). + +Lemma prod_is_cont : contfun_axiom (f1 \* f2). +Proof. +move=>d; split=>//=; +by rewrite contE !(limE (comp_chainE _ _ _)); apply: lim_mono. +Qed. +HB.instance Definition _ := + isContFun.Build (S1*S2)%type (T1*T2)%type (f1 \* f2) prod_is_cont. +End ProdCont. + + diff --git a/theories/dune b/htt/dune similarity index 91% rename from theories/dune rename to htt/dune index 501fc38..a6eeaa3 100644 --- a/theories/dune +++ b/htt/dune @@ -2,7 +2,7 @@ (coq.theory (name htt) - (package coq-htt) + (package coq-htt-core) (synopsis "Hoare Type Theory") (flags :standard -w -notation-overridden diff --git a/theories/heapauto.v b/htt/heapauto.v similarity index 82% rename from theories/heapauto.v rename to htt/heapauto.v index 994935c..bcac768 100644 --- a/theories/heapauto.v +++ b/htt/heapauto.v @@ -1,6 +1,19 @@ +(* +Copyright 2012 IMDEA Software Institute +Licensed under the Apache License, Version 2.0 (the "License"); +you may not use this file except in compliance with the License. +You may obtain a copy of the License at + http://www.apache.org/licenses/LICENSE-2.0 +Unless required by applicable law or agreed to in writing, software +distributed under the License is distributed on an "AS IS" BASIS, +WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. +See the License for the specific language governing permissions and +limitations under the License. +*) + From mathcomp Require Import ssreflect ssrbool ssrnat eqtype seq ssrfun. From pcm Require Import options axioms pred prelude. -From pcm Require Import pcm autopcm unionmap heap. +From pcm Require Import auto pcm autopcm unionmap heap. From htt Require Import model. Import Prenex Implicits. @@ -23,7 +36,7 @@ Import Prenex Implicits. (* Currently, it doesn't deal with weak pointers. I.e. if it sees terms *) (* like x :-> v1 and x :-> v2, it will reduce to v1 = v2 only if v1, v2 *) (* are of the same type. A more general tactic would emit obligation *) -(* dyn v1 = dyn v2, but I don't bother with this now. *) +(* dyn v1 = dyn v2, but that's left for future work. *) (* *) (**************************************************************************) @@ -108,8 +121,8 @@ Lemma try_retR e1 e2 (v : A) i (r : post B) : Proof. by move=>H; apply/vrf_try/val_retR. Qed. Lemma bnd_retR e (v : A) i (r : post B) : - vrf i (e v) r -> vrf i (bind (ret v) e) r. -Proof. by move=>H; apply/vrf_bind/val_retR. Qed. + vrf i (e v) r -> vrf i (bnd (ret v) e) r. +Proof. by move=>H; apply/vrf_bnd/val_retR. Qed. End EvalRetR. @@ -126,8 +139,8 @@ Proof. by move=>H; apply/vrf_try/val_throwR. Qed. Lemma bnd_throwR e e1 i (r : post B) : (valid i -> r (Exn e) i) -> - vrf i (bind (@throw A e) e1) r. -Proof. by move=>H; apply/vrf_bind/val_throwR. Qed. + vrf i (bnd (@throw A e) e1) r. +Proof. by move=>H; apply/vrf_bnd/val_throwR. Qed. End EvalThrowR. @@ -137,7 +150,7 @@ Variables (A B : Type). Lemma val_readR v x i (f : form (x :-> v) i) (r : post A) : (valid (untag f) -> r (Val v) f) -> - vrf f (read x) r. + vrf f (read A x) r. Proof. by rewrite formE; apply: vrf_read. Qed. Lemma try_readR e1 e2 v x i (f : form (x :-> v) i) (r : post B) : @@ -147,8 +160,8 @@ Proof. by move=>H; apply/vrf_try/val_readR. Qed. Lemma bnd_readR e v x i (f : form (x :-> v) i) (r : post B) : vrf f (e v) r -> - vrf f (bind (@read A x) e) r. -Proof. by move=>H; apply/vrf_bind/val_readR. Qed. + vrf f (bnd (@read A x) e) r. +Proof. by move=>H; apply/vrf_bnd/val_readR. Qed. End EvalReadR. @@ -169,8 +182,8 @@ Proof. by move=>H; apply/vrf_try/val_writeR. Qed. Lemma bnd_writeR e (v : A) (w : B) x i (f : forall k, form k i) (r : post C) : vrf (f (x :-> v)) (e tt) r -> - vrf (f (x :-> w)) (bind (write x v) e) r. -Proof. by move=>H; apply/vrf_bind/val_writeR. Qed. + vrf (f (x :-> w)) (bnd (write x v) e) r. +Proof. by move=>H; apply/vrf_bnd/val_writeR. Qed. End EvalWriteR. @@ -187,8 +200,8 @@ Proof. by move=>H; apply/vrf_try/val_allocR. Qed. Lemma bnd_allocR e (v : A) i (r : post B) : (forall x, vrf (x :-> v \+ i) (e x) r) -> - vrf i (bind (alloc v) e) r. -Proof. by move=>H; apply/vrf_bind/val_allocR. Qed. + vrf i (bnd (alloc v) e) r. +Proof. by move=>H; apply/vrf_bnd/val_allocR. Qed. End EvalAllocR. @@ -205,8 +218,8 @@ Proof. by move=>H; apply/vrf_try/val_allocbR. Qed. Lemma bnd_allocbR e (v : A) n i (r : post B) : (forall x, vrf (updi x (nseq n v) \+ i) (e x) r) -> - vrf i (bind (allocb v n) e) r. -Proof. by move=>H; apply/vrf_bind/val_allocbR. Qed. + vrf i (bnd (allocb v n) e) r. +Proof. by move=>H; apply/vrf_bnd/val_allocbR. Qed. End EvalAllocbR. @@ -226,8 +239,8 @@ Proof. by move=>H; apply/vrf_try/val_deallocR. Qed. Lemma bnd_deallocR e (v : A) x i (f : forall k, form k i) (r : post B) : vrf (f Unit) (e tt) r -> - vrf (f (x :-> v)) (bind (dealloc x) e) r. -Proof. by move=>H; apply/vrf_bind/val_deallocR. Qed. + vrf (f (x :-> v)) (bnd (dealloc x) e) r. +Proof. by move=>H; apply/vrf_bnd/val_deallocR. Qed. End EvalDeallocR. @@ -246,7 +259,7 @@ Structure val_form A i r (p : Prop):= Structure bnd_form A B i (e : A -> ST B) r (p : Prop) := BndForm {bnd_pivot :> ST A; - _ : p -> vrf i (bind bnd_pivot e) r}. + _ : p -> vrf i (bnd bnd_pivot e) r}. Structure try_form A B i (e1 : A -> ST B) (e2 : exn -> ST B) r (p : Prop) := @@ -264,7 +277,7 @@ Proof. by case:e=>[?]; apply. Qed. (* continue searching for a val_form. *) Lemma bnd_case_pf A B i (s : A -> ST B) r p (e : bnd_form i s r p) : - p -> vrf i (bind e s) r. + p -> vrf i (bnd e s) r. Proof. by case:e=>[?]; apply. Qed. Canonical Structure @@ -407,20 +420,20 @@ End Uncons. (* These issues can be circumvented by falling back to simpler (automated) *) (* lemmas `gU`, `gR` and their variants. *) -Lemma gX G A (s : spec G A) g (m : heapPCM) m0 j tm k wh r2 - (e : STspec G s) - (fm : Syntactify.form (empx _) j tm) - (fu : uform m0 (Syntactify.untag fm)) - (f : forall q, form q r2) - (fg : PullX.rform j k tm wh (Some (untag (heap_of (f m0))))) - (Q : post A) : +Lemma gX G A (s : spec G A) g (m : heap) m0 j tm k wh r2 + (e : STspec G s) + (fm : Syntactify.form (empx _) j tm) + (fu : uform m0 (Syntactify.untag fm)) + (f : forall q, form q r2) + (fg : PullX.rform j k tm wh (Some (untag (heap_of (f m0))))) + (Q : post A) : untagU fu = m -> (valid m -> (s g).1 m) -> (forall v n, (s g).2 (Val v) n -> valid (untag (f n)) -> Q (Val v) (f n)) -> (forall x n, (s g).2 (Exn x) n -> valid (untag (f n)) -> Q (Exn x) (f n)) -> - vrf (fg : heapPCM) e Q. + vrf (PullX.unpack fg) e Q. Proof. case: e=>e /= H; case: fu=>_ ->/= Em Hp Hv Hx; rewrite -{}Em in Hp. rewrite (pullX (Syntactify.untag fm)) formE joinCA joinA. @@ -429,7 +442,7 @@ apply/vrf_frame/vrf_post/V. by case=>[v|x] n _=>[/Hv|/Hx]; rewrite formE. Qed. -Arguments gX [G A s] g m {m0 j tm k wh r2 e fm fu f fg Q} _ _ _ _. +Arguments gX {G A s} g m {m0 j tm k wh r2 e fm fu f fg Q}. Notation "[gX] @ m" := (gX tt m erefl) (at level 0). @@ -437,55 +450,57 @@ Notation "[ 'gX' x1 , .. , xn ] @ m" := (gX (existT _ x1 .. (existT _ xn tt) ..) m erefl) (at level 0, format "[ 'gX' x1 , .. , xn ] @ m"). -(* a combination of gX + vrf_bind *) -Lemma stepX G A B (s : spec G A) g (m : heapPCM) m0 j tm k wh r2 - (e : STspec G s) (e2 : A -> ST B) - (fm : Syntactify.form (empx _) j tm) - (fu : uform m0 (Syntactify.untag fm)) - (f : forall q, form q r2) - (fg : PullX.rform j k tm wh (Some (untag (heap_of (f m0))))) - (Q : post B) : +Definition heapPCM : pcm := heap. + +(* combination of gX + vrf_bind *) +Lemma stepX G A B (s : spec G A) g (m : heapPCM) (m0 : heap) (j : ctx heapPCM) tm k wh r2 + (e : STspec G s) (e2 : A -> ST B) + (fm : Syntactify.form (empx _) j tm) + (fu : uform m0 (Syntactify.untag fm)) + (f : forall q, form q r2) + (fg : PullX.rform j k tm wh (Some (untag (heap_of (f m0))))) + (Q : post B) : untagU fu = m -> (valid m -> (s g).1 m) -> (forall v n, (s g).2 (Val v) n -> vrf (f n) (e2 v) Q) -> (forall x n, (s g).2 (Exn x) n -> valid (untag (f n)) -> Q (Exn x) (f n)) -> - vrf (fg : heapPCM) (bind e e2) Q. + vrf (fg : heapPCM) (bnd e e2) Q. Proof. -move=>Hm Hp Hv Hx. -by apply/vrf_bind/(gX _ _ Hm Hp)=>[v n H V|x n H V _]; [apply: Hv| apply: Hx]. +move=>Hm Hp Hv Hx; apply/vrf_bnd/(gX _ _ Hm Hp). +- by move=>v n H V; apply: Hv. +by move=>x n H V _; apply: Hx. Qed. -Arguments stepX [G A B s] g m {m0 j tm k wh r2 e e2 fm fu f fg Q} _ _ _ _. +Arguments stepX [G A B s] g m {m0 j tm k wh r2 e e2 fm fu f fg Q}. Notation "[stepX] @ m" := (stepX tt m erefl) (at level 0). - Notation "[ 'stepX' x1 , .. , xn ] @ m" := (stepX (existT _ x1 .. (existT _ xn tt) ..) m erefl) - (at level 0, format "[ 'stepX' x1 , .. , xn ] @ m"). - -(* a combination of gX + vrf_try *) -Lemma tryX G A B (s : spec G A) g (m : heapPCM) m0 j tm k wh r2 - (e : STspec G s) (e1 : A -> ST B) (e2 : exn -> ST B) - (fm : Syntactify.form (empx _) j tm) - (fu : uform m0 (Syntactify.untag fm)) - (f : forall q, form q r2) - (fg : PullX.rform j k tm wh (Some (untag (heap_of (f m0))))) - (Q : post B) : + (at level 0, format "[ 'stepX' x1 , .. , xn ] @ m"). + +(* combination of gX + vrf_try *) +Lemma tryX G A B (s : spec G A) g (m : heap) m0 j tm k wh r2 + (e : STspec G s) (e1 : A -> ST B) (e2 : exn -> ST B) + (fm : Syntactify.form (empx _) j tm) + (fu : uform m0 (Syntactify.untag fm)) + (f : forall q, form q r2) + (fg : PullX.rform j k tm wh (Some (untag (heap_of (f m0))))) + (Q : post B) : untagU fu = m -> (valid m -> (s g).1 m) -> (forall v n, (s g).2 (Val v) n -> vrf (f n) (e1 v) Q) -> (forall x n, (s g).2 (Exn x) n -> vrf (f n) (e2 x) Q) -> - vrf (fg : heapPCM) (try e e1 e2) Q. + vrf (PullX.unpack fg) (try e e1 e2) Q. Proof. -move=>Hm Hp Hv Hx. -by apply/vrf_try/(gX _ _ Hm Hp)=>[x|ex] n H V; [apply: Hv|apply: Hx]. +move=>Hm Hp Hv Hx; apply/vrf_try/(gX _ _ Hm Hp). +- by move=>x n H V; apply: Hv. +by move=>ex n H V; apply: Hx. Qed. -Arguments tryX [G A B s] g m {m0 j tm k wh r2 e e1 e2 fm fu f fg Q} _ _ _ _. +Arguments tryX {G A B s} g m {m0 j tm k wh r2 e e1 e2 fm fu f fg Q}. Notation "[tryX] @ m" := (stepX tt m erefl) (at level 0). - Notation "[ 'tryX' x1 , .. , xn ] @ m" := (tryX (existT _ x1 .. (existT _ xn tt) ..) m erefl) (at level 0, format "[ 'tryX' x1 , .. , xn ] @ m"). @@ -494,7 +509,7 @@ Notation "[ 'tryX' x1 , .. , xn ] @ m" := (* Simplified ghost lemma automations *) (**************************************) -(* A simpler version of an automated framing+ghost lemma which only works on *) +(* simpler version of an automated framing+ghost lemma which only works on *) (* literal heap subexpressions (here `m`) but preserves associativity. *) Lemma gR G A (s : spec G A) g m r (e : STspec G s) (f : forall k, form k r) (Q : post A) : @@ -511,40 +526,38 @@ apply/vrf_frame/vrf_post/V. by case=>[x|ex] n _ =>[/Hv|/Hx]; rewrite formE. Qed. -Arguments gR [G A s] g m {r e f Q} _ _ _. +Arguments gR {G A s} g m {r e f Q}. Notation "[gR] @ m" := (gR tt m) (at level 0). - Notation "[ 'gR' x1 , .. , xn ] @ m" := (gR (existT _ x1 .. (existT _ xn tt) ..) m) (at level 0, format "[ 'gR' x1 , .. , xn ] @ m"). -(* a combination of gR + vrf_bind *) +(* combination of gR + vrf_bind *) Lemma stepR G A B (s : spec G A) g i j (e : STspec G s) (e2 : A -> ST B) - (f : forall k, form k j) (Q : post B) : + (f : forall k, form k j) (Q : post B) : (valid i -> (s g).1 i) -> (forall x m, (s g).2 (Val x) m -> vrf (f m) (e2 x) Q) -> (forall x m, (s g).2 (Exn x) m -> valid (untag (f m)) -> Q (Exn x) (f m)) -> - vrf (f i) (bind e e2) Q. + vrf (f i) (bnd e e2) Q. Proof. -move=>Hi H1 H2. -apply/vrf_bind/(gR _ _ Hi)=>[x m H V|ex m H V _]. -- by apply: H1 H. -by apply: H2. +move=>Hi H1 H2; apply/vrf_bnd/(gR _ _ Hi). +- by move=>x m H V; apply: H1 H. +by move=>ex m H V _; apply: H2. Qed. -Arguments stepR [G A B s] g i [j e e2 f Q] _ _ _. +Arguments stepR {G A B s} g i {j e e2 f Q}. Notation "[stepR] @ i" := (stepR tt i) (at level 0). - Notation "[ 'stepR' x1 , .. , xn ] @ i" := (stepR (existT _ x1 .. (existT _ xn tt) ..) i) (at level 0, format "[ 'stepR' x1 , .. , xn ] @ i"). -(* a combination of gR + vrf_try *) -Lemma tryR G A B (s : spec G A) g i j (e : STspec G s) (e1 : A -> ST B) (e2 : exn -> ST B) - (f : forall k, form k j) (Q : post B) : +(* combination of gR + vrf_try *) +Lemma tryR G A B (s : spec G A) g i j (e : STspec G s) + (e1 : A -> ST B) (e2 : exn -> ST B) + (f : forall k, form k j) (Q : post B) : (valid i -> (s g).1 i) -> (forall x m, (s g).2 (Val x) m -> vrf (f m) (e1 x) Q) -> (forall x m, (s g).2 (Exn x) m -> vrf (f m) (e2 x) Q) -> @@ -556,17 +569,16 @@ apply/vrf_try/(gR _ _ Hi)=>[x|ex] m H V. by apply: H2. Qed. -Arguments tryR [G A B s] g i [j e e1 e2 f Q] _ _ _. +Arguments tryR {G A B s} g i {j e e1 e2 f Q}. Notation "[tryR] @ i" := (tryR tt i) (at level 0). - Notation "[ 'tryR' x1 , .. , xn ] @ i" := (tryR (existT _ x1 .. (existT _ xn tt) ..) i) (at level 0, format "[ 'tryR' x1 , .. , xn ] @ i"). -(* This is really brittle, but I didn't get around yet to substitute it *) -(* with Mtac or overloaded lemmas. So for now, let's stick with the hack *) -(* in order to support the legacy proofs *) +(* The following is brittle, and should eventually be substituted *) +(* with overloaded lemmas. For now, sticking with the hack *) +(* for the purpose of supporting legacy proofs *) (* First cancelation in hypotheses *) diff --git a/theories/model.v b/htt/model.v similarity index 56% rename from theories/model.v rename to htt/model.v index c2fe817..a2c66e4 100644 --- a/theories/model.v +++ b/htt/model.v @@ -1,8 +1,30 @@ +(* +Copyright 2012 IMDEA Software Institute +Licensed under the Apache License, Version 2.0 (the "License"); +you may not use this file except in compliance with the License. +You may obtain a copy of the License at + http://www.apache.org/licenses/LICENSE-2.0 +Unless required by applicable law or agreed to in writing, software +distributed under the License is distributed on an "AS IS" BASIS, +WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. +See the License for the specific language governing permissions and +limitations under the License. +*) + +From HB Require Import structures. From mathcomp Require Import ssreflect ssrbool ssrfun ssrnat eqtype seq. From pcm Require Import options axioms pred prelude. From pcm Require Import pcm unionmap heap. From htt Require Import domain. +(*********************************************) +(* Denotational model *) +(* denoting programs as relations *) +(* over input/output heaps and output result *) +(*********************************************) + +(* Helper definitions *) + (* Exceptions are an equality type *) Inductive exn : Type := exn_from_nat of nat. @@ -15,21 +37,23 @@ Definition eqexn := Lemma eqexnP : Equality.axiom eqexn. Proof. by move=>[x][y]/=; case: eqP=>[->|*]; constructor=>//; case. Qed. -Canonical Structure exn_eqMixin := EqMixin eqexnP. -Canonical Structure exn_eqType := EqType exn exn_eqMixin. +HB.instance Definition _ := hasDecEq.Build exn eqexnP. (* Answer type *) Inductive ans (A : Type) : Type := Val of A | Exn of exn. Arguments Exn [A]. -(* A set of heaps *) -Notation pre := (Pred heap). - -(* A set of (ans A * heap) *) -(* This models the fact that programs can hang, returning nothing, *) -(* or produce nondeterministic results (e.g. alloc). *) -Notation post A := (ans A -> heap -> Prop). - +(* precondition is predicate over heaps *) +Definition pre := Pred heap. +(* postcondition is relates return result and output heap *) +(* because it's a relation, it models that programs can produce *) +(* no results (by looping forever) or produce more than one result *) +(* non-deterministically. *) +(* The latter will be used to model the alloc program *) +Definition post A := ans A -> heap -> Prop. +(* specification is a pair of precondition and postcondition *) +(* possibly parameterized by logical variable of type G (ghost) *) +(* (G may have product type, to allow for context of logical variables) *) Definition spec G A := G -> pre * post A : Type. (*************************************************************) @@ -37,12 +61,10 @@ Definition spec G A := G -> pre * post A : Type. (*************************************************************) Module Type VrfSig. - Parameter ST : Type -> Type. - Parameter ret : forall A, A -> ST A. Parameter throw : forall A, exn -> ST A. -Parameter bind : forall A B, ST A -> (A -> ST B) -> ST B. +Parameter bnd : forall A B, ST A -> (A -> ST B) -> ST B. Parameter try : forall A B, ST A -> (A -> ST B) -> (exn -> ST B) -> ST B. Parameter read : forall A, ptr -> ST A. Parameter write : forall A, ptr -> A -> ST unit. @@ -50,58 +72,64 @@ Parameter alloc : forall A, A -> ST ptr. Parameter allocb : forall A, A -> nat -> ST ptr. Parameter dealloc : ptr -> ST unit. -Arguments throw [A] e. -Arguments read [A] x. +Arguments throw {A}. +Arguments read : clear implicits. -(* we need program to come first in the argument list - so that automation can match on it *) +(* given program e : ST A, input heap i, postcondition Q *) +(* vrf' judges if running e in i produces output heap and result *) +(* that satisfy Q *) +(* related to predicate transformers *) +(* vrf' has program as its first argument to facilitate automation *) Parameter vrf' : forall A, ST A -> heap -> post A -> Prop. -(* recover the usual [pre]prog[post] order with a notation *) +(* in practice it's more convenient to order the arguments as *) +(* initial heap, program, postcondition *) Notation vrf i e Q := (vrf' e i Q). Parameter vrfV : forall A e i (Q : post A), - (valid i -> vrf i e Q) -> vrf i e Q. + (valid i -> vrf i e Q) -> vrf i e Q. Parameter vrf_post : forall A e i (Q1 Q2 : post A), - (forall y m, valid m -> Q1 y m -> Q2 y m) -> - vrf i e Q1 -> vrf i e Q2. + (forall y m, valid m -> Q1 y m -> Q2 y m) -> + vrf i e Q1 -> vrf i e Q2. Parameter vrf_frame : forall A e i j (Q : post A), - vrf i e (fun y m => valid (m \+ j) -> Q y (m \+ j)) -> - vrf (i \+ j) e Q. + vrf i e (fun y m => valid (m \+ j) -> Q y (m \+ j)) -> + vrf (i \+ j) e Q. Parameter vrf_ret : forall A x i (Q : post A), - (valid i -> Q (Val x) i) -> vrf i (ret x) Q. + (valid i -> Q (Val x) i) -> vrf i (ret x) Q. Parameter vrf_throw : forall A e i (Q : post A), - (valid i -> Q (Exn e) i) -> vrf i (throw e) Q. -Parameter vrf_bind : forall A B (e1 : ST A) (e2 : A -> ST B) i (Q : post B), - vrf i e1 (fun x m => - match x with - | Val x' => vrf m (e2 x') Q - | Exn e => valid m -> Q (Exn e) m - end) -> - vrf i (bind e1 e2) Q. -Parameter vrf_try : forall A B (e : ST A) (e1 : A -> ST B) (e2 : exn -> ST B) i (Q : post B), - vrf i e (fun x m => - match x with - | Val x' => vrf m (e1 x') Q - | Exn ex => vrf m (e2 ex) Q - end) -> - vrf i (try e e1 e2) Q. + (valid i -> Q (Exn e) i) -> vrf i (throw e) Q. +Parameter vrf_bnd : forall A B (e1 : ST A) (e2 : A -> ST B) + i (Q : post B), + vrf i e1 (fun x m => + match x with + | Val x' => vrf m (e2 x') Q + | Exn e => valid m -> Q (Exn e) m + end) -> + vrf i (bnd e1 e2) Q. +Parameter vrf_try : forall A B (e : ST A) (e1 : A -> ST B) + (e2 : exn -> ST B) i (Q : post B), + vrf i e (fun x m => + match x with + | Val x' => vrf m (e1 x') Q + | Exn ex => vrf m (e2 ex) Q + end) -> + vrf i (try e e1 e2) Q. Parameter vrf_read : forall A x j (v : A) (Q : post A), - (valid (x :-> v \+ j) -> Q (Val v) (x :-> v \+ j)) -> - vrf (x :-> v \+ j) (read x) Q. + (valid (x :-> v \+ j) -> Q (Val v) (x :-> v \+ j)) -> + vrf (x :-> v \+ j) (read A x) Q. Parameter vrf_write : forall A x (v : A) B (u : B) j (Q : post unit), - (valid (x :-> v \+ j) -> Q (Val tt) (x :-> v \+ j)) -> - vrf (x :-> u \+ j) (write x v) Q. + (valid (x :-> v \+ j) -> Q (Val tt) (x :-> v \+ j)) -> + vrf (x :-> u \+ j) (write x v) Q. Parameter vrf_alloc : forall A (v : A) i (Q : post ptr), - (forall x, valid (x :-> v \+ i) -> Q (Val x) (x :-> v \+ i)) -> - vrf i (alloc v) Q. + (forall x, valid (x :-> v \+ i) -> Q (Val x) (x :-> v \+ i)) -> + vrf i (alloc v) Q. Parameter vrf_allocb : forall A (v : A) n i (Q : post ptr), - (forall x, valid (updi x (nseq n v) \+ i) -> - Q (Val x) (updi x (nseq n v) \+ i)) -> - vrf i (allocb v n) Q. + (forall x, valid (updi x (nseq n v) \+ i) -> + Q (Val x) (updi x (nseq n v) \+ i)) -> + vrf i (allocb v n) Q. Parameter vrf_dealloc : forall x A (v : A) j (Q : post unit), - (x \notin dom j -> valid j -> Q (Val tt) j) -> - vrf (x :-> v \+ j) (dealloc x) Q. + (x \notin dom j -> valid j -> Q (Val tt) j) -> + vrf (x :-> v \+ j) (dealloc x) Q. Definition has_spec G A (s : spec G A) (e : ST A) := forall g i, (s g).1 i -> vrf i e (s g).2. @@ -112,22 +140,37 @@ Structure STspec G A (s : spec G A) := STprog { Arguments STspec G [A] s. -Notation "'Do' e" := (@STprog _ _ _ e _) (at level 80). - -Notation "x '<--' c1 ';' c2" := (bind c1 (fun x => c2)) - (at level 81, right associativity). -Notation "c1 ';;' c2" := (bind c1 (fun _ => c2)) +(* some notation *) +Notation "x '<--' c1 ';' c2" := (bnd c1 (fun x => c2)) + (at level 81, right associativity, format + "'[v' x '<--' c1 ';' '//' c2 ']'"). +Notation "' x '<--' c1 ';' c2" := (bnd c1 (fun x => c2)) + (at level 81, right associativity, x strict pattern, format + "'[v' ' x '<--' c1 ';' '//' c2 ']'"). +Notation "c1 ';;' c2" := (bnd c1 (fun _ => c2)) (at level 81, right associativity). -Notation "'!' x" := (read x) (at level 50). +Notation "'!' x" := (read _ x) (at level 50). Notation "x '::=' e" := (write x e) (at level 60). - -Parameter Fix : forall G A (B : A -> Type) (s : forall x : A, spec G (B x)), +(* using locked for Do, to make the Do definitions opaque *) +(* otherwise, clients will inline the code of Do, preventing *) +(* the use of the program with the spec that has been verified *) +(* with Do. *) +Notation "'Do' e" := (locked (@STprog _ _ _ e _)) (at level 80). + +(* Fixed point constructor *) +(* We shall make fix work over *monotone closure* of argument function. *) +(* This is a mathematical hack to avoid spurious monotonicity proofs *) +(* that's necessary in shallow embedding, as we have no guarantee *) +(* that the argument function is actually monotone. *) +(* We *do* prove later on that all constructors are monotone. *) +(* We also hide model definition so that all functions one can *) +(* apply ffix to will be built out of monotonicity-preserving *) +(* constructors (and hence will be monotone). *) +Parameter ffix : forall G A (B : A -> Type) (s : forall x : A, spec G (B x)), ((forall x : A, STspec G (s x)) -> forall x : A, STspec G (s x)) -> forall x : A, STspec G (s x). - End VrfSig. - (********************************) (* Definition of the Hoare type *) (********************************) @@ -137,8 +180,9 @@ Module Vrf : VrfSig. Section BasePrograms. Variables (P : pre) (A : Type). -(* we carve out the model out of the following base type *) -Definition prog : Type := forall i : heap, valid i -> i \In P -> post A. +(* the model is carved out of the following base type *) +Definition prog : Type := + forall i : heap, valid i -> i \In P -> post A. (* we take only preconditions and progs with special properties *) (* which we define next *) @@ -149,7 +193,7 @@ Definition safe_mono := (* defined heaps map to defined heaps *) Definition def_strict (e : prog) := - forall i p v x, Heap.Undef \Notin e i v p x. + forall i p v x, undef \Notin e i v p x. (* frame property *) Definition frameable (e : prog) := @@ -159,35 +203,48 @@ Definition frameable (e : prog) := End BasePrograms. -Section STDef. -Variable (A : Type). +(* the model considers programs of base type, given a precondition *) +(* and adds several properties: *) +(* - safety monotonicity *) +(* - strictness of definedness *) +(* - frameability *) + +Definition st_axiom A (p : pre) (e : prog p A) := + [/\ safe_mono p, def_strict e & frameable e]. -Structure ST' := Prog { +Structure ST' (A : Type) := Prog { pre_of : pre; prog_of : prog pre_of A; - _ : safe_mono pre_of; - _ : def_strict prog_of; - _ : frameable prog_of}. + _ : st_axiom prog_of}. + +Arguments prog_of {A}. (* module field must be a definition, not structure *) Definition ST := ST'. -Lemma sfm_st e : safe_mono (pre_of e). -Proof. by case: e. Qed. +(* projections *) -Arguments prog_of : clear implicits. +Lemma sfm_st A (e : ST A) : safe_mono (pre_of e). +Proof. by case: e=>?? []. Qed. -Lemma dstr_st e : def_strict (prog_of e). -Proof. by case: e. Qed. +Lemma dstr_st A (e : ST A) : def_strict (prog_of e). +Proof. by case: e=>?? []. Qed. -Corollary dstr_valid e i p v x m : - m \In prog_of e i p v x -> valid m. +Lemma dstr_valid A (e : ST A) i p v x m : + m \In prog_of e i p v x -> + valid m. Proof. by case: m=>// /dstr_st. Qed. -Lemma fr_st e : frameable (prog_of e). -Proof. by case: e. Qed. +Lemma fr_st A (e : ST A) : frameable (prog_of e). +Proof. by case: e=>?? []. Qed. + +Arguments sfm_st {A e i j}. +Arguments dstr_st {A e i}. +Arguments fr_st {A e i j}. -Arguments fr_st [e i j]. +Section STDef. +Variable A : Type. +Implicit Type e : ST A. (* poset structure on ST *) @@ -196,30 +253,24 @@ Definition st_leq e1 e2 := forall i (v : valid i) (p : i \In pre_of e2), prog_of e1 _ v (pf _ p) <== prog_of e2 _ v p. -Lemma st_refl e : st_leq e e. +Lemma st_is_poset : poset_axiom st_leq. Proof. -exists (poset_refl _)=>i V P y m. -by rewrite (pf_irr (poset_refl (pre_of e) i P) P). -Qed. - -Lemma st_asym e1 e2 : st_leq e1 e2 -> st_leq e2 e1 -> e1 = e2. -Proof. -move: e1 e2=>[p1 e1 S1 D1 F1][p2 e2 S2 D2 F2]; rewrite /st_leq /=. -case=>E1 R1 [E2 R2]. -move: (poset_asym E1 E2)=>?; subst p2. -have : e1 = e2. -- apply: fext=>i; apply: fext=>Vi; apply: fext=>Pi; apply: fext=>y; apply: fext=>m. - move: (R2 i Vi Pi y m)=>{}R2; move: (R1 i Vi Pi y m)=>{}R1. - apply: pext; split. - - by move=>H1; apply: R1; rewrite (pf_irr (E1 i Pi) Pi). - by move=>H2; apply: R2; rewrite (pf_irr (E2 i Pi) Pi). -move=>?; subst e2. -by congr Prog; apply: pf_irr. -Qed. - -Lemma st_trans e1 e2 e3 : st_leq e1 e2 -> st_leq e2 e3 -> st_leq e1 e3. -Proof. -move: e1 e2 e3=>[p1 e1 S1 D1 F1][p2 e2 S2 D2 F2][p3 e3 S3 D3 F3]. +split=>[e|e1 e2|e1 e2 e3]. +- exists (poset_refl _)=>i V P y m. + by rewrite (pf_irr (poset_refl (pre_of e) i P) P). +- case: e1 e2=>p1 e1 [S1 D1 F1][p2 e2 [S2 D2 F2]]. + rewrite /st_leq /=; case=>E1 R1 [E2 R2]. + move: (poset_asym E1 E2)=>?; subst p2. + have : e1 = e2. + - apply: fext=>i; apply: fext=>Vi; apply: fext=>Pi. + apply: fext=>y; apply: fext=>m. + move: (R2 i Vi Pi y m)=>{}R2; move: (R1 i Vi Pi y m)=>{}R1. + apply: pext; split. + - by move=>H1; apply: R1; rewrite (pf_irr (E1 i Pi) Pi). + by move=>H2; apply: R2; rewrite (pf_irr (E2 i Pi) Pi). + move=>?; subst e2. + by congr Prog; apply: pf_irr. +case: e1 e2 e3=>p1 e1 [S1 D1 F1][p2 e2 [S2 D2 F2]][p3 e3 [S3 D3 F3]]. case=>/= E1 R1 [/= E2 R2]; rewrite /st_leq /=. have E3 := poset_trans E2 E1; exists E3=>i V P y m. set P' := E2 i P. @@ -228,58 +279,33 @@ move=>H1; apply/R2/R1. by rewrite (pf_irr (E1 i P') (E3 i P)). Qed. -(* bottom is a program that can always run but never returns (an endless loop) *) - -Definition pre_bot : pre := top. - -Definition prog_bot : prog pre_bot A := - fun _ _ _ _ => bot. - -Lemma sfmono_bot : safe_mono pre_bot. -Proof. by []. Qed. - -Lemma dstrict_bot : def_strict prog_bot. -Proof. by move=>*. Qed. - -Lemma frame_bot : frameable prog_bot. -Proof. by []. Qed. - -Definition st_bot := Prog sfmono_bot dstrict_bot frame_bot. - -Lemma st_botP e : st_leq st_bot e. -Proof. by case: e=>p e S D F; exists (@pred_topP _ _)=>???; apply: botP. Qed. - -Definition stPosetMixin := PosetMixin st_botP st_refl st_asym st_trans. -Canonical stPoset := Eval hnf in Poset ST stPosetMixin. +HB.instance Definition _ := isPoset.Build (ST A) st_is_poset. (* lattice structure on ST *) +(* suprema of programs: *) +(* - precondition is intersection of preconditions of programs *) +(* - denotation is union of denotations of programs *) + (* intersection of preconditions *) -Definition pre_sup (u : Pred ST) : pre := +Definition pre_sup (u : Pred (ST A)) : pre := fun h => forall e, e \In u -> h \In pre_of e. Definition pre_sup_leq u e (pf : e \In u) : pre_sup u <== pre_of e := fun h (pf1 : pre_sup u h) => pf1 e pf. (* union of postconditions *) -Definition prog_sup (u : Pred ST) : prog (pre_sup u) A := +Definition prog_sup (u : Pred (ST A)) : prog (pre_sup u) A := fun i V P y m => exists e (pf : e \In u), prog_of e _ V (pre_sup_leq pf P) y m. Arguments prog_sup : clear implicits. -Lemma pre_sup_sfmono u : safe_mono (pre_sup u). +Lemma progsup_is_st u : st_axiom (prog_sup u). Proof. -move=>i j Pi Vij e He. -by apply: sfm_st=>//; apply: Pi. -Qed. - -Lemma prog_sup_dstrict u : def_strict (prog_sup u). -Proof. by move=>i P V y; case; case=>p e S D F [H1] /D. Qed. - -Lemma prog_sup_frame u : frameable (prog_sup u). -Proof. -move=>i j Pi Vij Pij y m [e][He]Pe. +split=>[i j Pi Vij E He|i P V y|i j Pi Vij Pij y m [e][He Pe]]. +- by apply: sfm_st=>//; apply: Pi. +- by case; case=>p e [S D F][H1] /D. have Pi' := Pi e He; have Pij' := Pij e He. move: Pe; rewrite (pf_irr (pre_sup_leq He Pij) Pij'). case/(fr_st Pi' Vij Pij')=>h [{m}-> Vhj Ph]. @@ -287,43 +313,29 @@ exists h; split=>//; exists e, He. by rewrite (pf_irr (pre_sup_leq He Pi) Pi'). Qed. -Definition st_sup u : ST := - Prog (@pre_sup_sfmono u) (@prog_sup_dstrict u) (@prog_sup_frame u). +Definition st_sup u : ST A := Prog (progsup_is_st u). -Lemma st_supP u e : e \In u -> e <== st_sup u. +Lemma st_is_lattice : lattice_axiom st_sup. Proof. -case: e=>p e' S D F R. -exists (pre_sup_leq R)=>/=p0 y m H. -by exists (Prog S D F), R. -Qed. - -Lemma st_supM u e : - (forall e1, e1 \In u -> e1 <== e) -> st_sup u <== e. -Proof. -case: e=>p e S D F R. +split=>/= u e; case: e=>p e' X R. +- exists (pre_sup_leq R)=>/=p0 y m H. + by exists (Prog X), R. have J : p <== pre_sup u. -- by move=>/= x Px e' pf; case: (R _ pf)=>/= + _; apply. +- by move=>/= x Px ex pf; case: (R _ pf)=>/= + _; apply. exists J=>i V P y m [e0][H0 Hm]. case: (R _ H0)=>/= Hx; apply. by rewrite (pf_irr (Hx i P) (pre_sup_leq H0 (J i P))). Qed. -Definition stLatticeMixin := LatticeMixin st_supP st_supM. -Canonical stLattice := Lattice ST stLatticeMixin. - +HB.instance Definition _ := isLattice.Build (ST A) st_is_lattice. End STDef. -Arguments prog_of [A]. -Arguments sfm_st [A e i j]. -Arguments dstr_st [A e i]. -Arguments fr_st [A e i j]. +(* types that embed the specs *) Section STspecDef. Variables (G A : Type) (s : spec G A). -(* strongest postcondition predicate transformer *) - -Definition vrf' (e : ST A) i (Q : post A) := +Definition vrf' (e : ST A) (i : heap) (Q : post A) := forall (V : valid i), exists (pf : i \In pre_of e), forall y m, prog_of e _ V pf y m -> Q y m. @@ -347,35 +359,23 @@ Qed. Definition stsp_leq e1 e2 := model e1 <== model e2. -Lemma stsp_refl e : stsp_leq e e. -Proof. by case: e=>e He; apply: poset_refl. Qed. - -Lemma stsp_asym e1 e2 : stsp_leq e1 e2 -> stsp_leq e2 e1 -> e1 = e2. +Lemma stspleq_is_poset : poset_axiom stsp_leq. Proof. -move: e1 e2=>[e1 H1][e2 H2]; rewrite /stsp_leq /= =>E1 E2. -have E := poset_asym E1 E2; subst e2. -by congr STprog; apply: pf_irr. -Qed. - -Lemma stsp_trans e1 e2 e3 : stsp_leq e1 e2 -> stsp_leq e2 e3 -> stsp_leq e1 e3. -Proof. -move: e1 e2 e3=>[e1 H1][e2 H2][e3 H3]. +split=>[e|e1 e2|e1 e2 e3]. +- by case: e=>e He; apply: poset_refl. +- case: e1 e2=>e1 H1 [e2 H2]; rewrite /stsp_leq /= =>E1 E2. + have E := poset_asym E1 E2; subst e2. + by congr STprog; apply: pf_irr. +case: e1 e2 e3=>e1 H1 [e2 H2][e3 H3]. by apply: poset_trans. Qed. -Lemma st_bot_has_spec : @st_bot A \In has_spec. -Proof. by move=>g i H V /=; exists I. Qed. - -Definition stsp_bot := STprog st_bot_has_spec. - -Lemma stsp_botP e : stsp_leq stsp_bot e. -Proof. by case: e=>*; apply: botP. Qed. - -Definition stspPosetMixin := PosetMixin stsp_botP stsp_refl stsp_asym stsp_trans. -Canonical stspPoset := Eval hnf in Poset STspec stspPosetMixin. +HB.instance Definition _ := isPoset.Build STspec stspleq_is_poset. (* lattice structure on STspec *) +(* denotation of the union of STspec programs is *) +(* the union of their denotations *) Definition st_sup' (u : Pred STspec) : ST A := sup [Pred p | exists e, p = model e /\ e \In u]. @@ -392,53 +392,20 @@ Qed. Definition stsp_sup u := STprog (@st_sup_has_spec' u). -Lemma stsp_supP u e : e \In u -> e <== stsp_sup u. -Proof. by case: e=>p S R; apply: supP; exists (STprog S). Qed. - -Lemma stsp_supM u e : - (forall e1, e1 \In u -> e1 <== e) -> stsp_sup u <== e. -Proof. by case: e=>p S R; apply: supM=>/= y[q][->]; apply: R. Qed. - -Definition stspLatticeMixin := LatticeMixin stsp_supP stsp_supM. -Canonical stspLattice := Lattice STspec stspLatticeMixin. - -End STspecDef. - -Notation vrf i e Q := (vrf' e i Q). - -(************************************) -(* modeling the language primitives *) -(************************************) - -(* recursion *) -Section Fix. -Variables (G A : Type) (B : A -> Type) (s : forall x, spec G (B x)). -Notation tp := (forall x, STspec (s x)). -Notation lat := (dfunLattice (fun x => [lattice of STspec (s x)])). -Variable (f : tp -> tp). - -(* we take a fixpoint not of f, but of its monotone completion f' *) - -Definition f' (e : lat) := - sup [Pred t : lat | exists e', e' <== e /\ t = f e']. - -Lemma f'_mono : monotone f'. +Lemma stspsup_is_lattice : lattice_axiom stsp_sup. Proof. -move=>x y H; apply: sup_mono=>fz; case=>z [Hz {fz}->]. -by exists z; split=>//; apply/poset_trans/H. +split=>/= u [p S R]. +- by apply: supP; exists (STprog S). +by apply: supM=>/= y[q][->]; apply: R. Qed. -Definition Fix : tp := tarski_lfp f'. - -(* fixed point constructor which requires explicit proof of monotonicity *) -Definition Fix' (pf : monotone (f : lat -> lat)) : tp := - tarski_lfp (f : lat -> lat). +HB.instance Definition _ := isLattice.Build STspec stspsup_is_lattice. -End Fix. +End STspecDef. -Arguments Fix [G A B s] f x. -Arguments Fix' [G A B s] f pf x. +Notation vrf i e Q := (vrf' e i Q). +(* required vrf lemmas *) Section VrfLemmas. Variables (A : Type) (e : ST A). @@ -480,102 +447,110 @@ Qed. End VrfLemmas. +(************************************) +(* modeling the language primitives *) +(************************************) + +(* recursion *) + +Section Fix. +Variables (G A : Type) (B : A -> Type) (pq : forall x, spec G (B x)). +Notation tp := (forall x, STspec (pq x)). +Notation lat := (dfun_lattice (fun x => STspec (pq x))). +Variable f : tp -> tp. + +(* fixed point constructor over monotone closure *) +Definition f' (c : lat) := sup [Pred t : lat | exists c', c' <== c /\ t = f c']. +Definition ffix := tarski_lfp f'. + +(* fixed point constructor which requires explicit proof of monotonicity *) +Definition monotone' := monofun_axiom (f : lat -> lat). +Definition ffix2 (pf : monotone') : tp := tarski_lfp (f : lat -> lat). +End Fix. + +Arguments ffix {G A B pq}. +Arguments ffix2 {G A B pq}. + +(* monadic unit *) Section Return. Variables (A : Type) (x : A). Definition ret_pre : pre := top. Definition ret_prog : prog ret_pre A := - fun i _ _ y m => - m = i /\ y = Val x. - -Lemma ret_sfmono : safe_mono ret_pre. -Proof. by []. Qed. + fun i _ _ y m => m = i /\ y = Val x. -Lemma ret_dstrict : def_strict ret_prog. -Proof. by move=>i [] V _ /= [E _]; rewrite -E in V. Qed. - -Lemma ret_frame : frameable ret_prog. -Proof. by move=>i j [Vij []] _ _ [-> ->]; exists i. Qed. +Lemma retprog_is_st : st_axiom ret_prog. +Proof. +split=>[//|i [] V _ /= [E _]|i j [Vij []] _ _ [->->]]. +- by rewrite -E in V. +by exists i. +Qed. -Definition ret := Prog ret_sfmono ret_dstrict ret_frame. +Definition ret := Prog retprog_is_st. Lemma vrf_ret i (Q : post A) : (valid i -> Q (Val x) i) -> vrf i ret Q. Proof. by move=>H V; exists I=>_ _ [->->]; apply: H. Qed. - End Return. +(* exception throwing *) Section Throw. Variables (A : Type) (e : exn). Definition throw_pre : pre := top. Definition throw_prog : prog throw_pre A := - fun i _ _ y m => - m = i /\ y = @Exn A e. - -Lemma throw_sfmono : safe_mono throw_pre. -Proof. by []. Qed. - -Lemma throw_dstrict : def_strict throw_prog. -Proof. by move=>i [] V _ /= [E _]; rewrite -E in V. Qed. + fun i _ _ y m => m = i /\ y = @Exn A e. -Lemma throw_frame : frameable throw_prog. -Proof. by move=>i j [Vij []] _ _ [-> ->]; exists i. Qed. +Lemma throwprog_is_st : st_axiom throw_prog. +Proof. +split=>[//|i [V] _ /= [E _]|i j [Vij []] _ _ [->->]]. +- by rewrite -E in V. +by exists i. +Qed. -Definition throw := Prog throw_sfmono throw_dstrict throw_frame. +Definition throw := Prog throwprog_is_st. Lemma vrf_throw i (Q : post A) : (valid i -> Q (Exn e) i) -> vrf i throw Q. Proof. by move=>H V; exists I=>_ _ [->->]; apply: H. Qed. - End Throw. -Section Bind. +(* monadic bind *) +Section Bnd. Variables (A B : Type). Variables (e1 : ST A) (e2 : A -> ST B). -Definition bind_pre : pre := +Definition bnd_pre : pre := fun i => exists (Vi : valid i) (Pi : i \In pre_of e1), forall x m, prog_of e1 _ Vi Pi (Val x) m -> pre_of (e2 x) m. -Definition bind_pre_proj i : i \In bind_pre -> i \In pre_of e1 := +Definition bnd_pre_proj i : i \In bnd_pre -> i \In pre_of e1 := fun '(ex_intro _ (ex_intro p _)) => p. -Definition bind_prog : prog bind_pre B := +Definition bnd_prog : prog bnd_pre B := fun i V P y m => - exists x h (Ph : h \In prog_of e1 _ V (bind_pre_proj P) x), + exists x h (Ph : h \In prog_of e1 _ V (bnd_pre_proj P) x), match x with | Val x' => exists Ph' : h \In pre_of (e2 x'), m \In prog_of (e2 x') _ (dstr_valid Ph) Ph' y | Exn e => y = Exn e /\ m = h end. -Lemma bind_sfmono : safe_mono bind_pre. -Proof. -move=>i j [Vi][Pi]P Vij. -have Pij := sfm_st Pi Vij. -exists Vij, Pij=>x m. -case/(fr_st Pi Vij Pij)=>h [{m}-> Vhj]. -rewrite (pf_irr (validL Vij) Vi)=>/P Ph. -by apply: sfm_st=>//; apply: (validL Vhj). -Qed. - -Lemma bind_dstrict : def_strict bind_prog. -Proof. -move=>i [Vi][Pi P] Vi' y[x][h][/=]. -case: x=>[x|e]Ph. -- by case=>Ph' /dstr_st. -by case=>_; move: Ph=>/[swap]<- /dstr_st. -Qed. - -Lemma bind_frame : frameable bind_prog. +Lemma bndprog_is_st : st_axiom bnd_prog. Proof. +split. +- move=>i j [Vi][Pi] P Vij; have Pij := sfm_st Pi Vij. + exists Vij, Pij=>x m; case/(fr_st Pi Vij Pij)=>h [{m}-> Vhj]. + rewrite (pf_irr (validL Vij) Vi)=>/P Ph. + by apply: sfm_st=>//; apply: (validL Vhj). +- move=>i [Vi][Pi P] Vi' y [x][h][/=]. + case: x=>[x|e]Ph; first by case=>Ph' /dstr_st. + by case=>_; move: Ph=>/[swap]<- /dstr_st. move=>i j [Vi][Pi P] Vij [_ [Pij _]] y m [x][h][/=]. -move: (fr_st Pi Vij Pij)=>H. -case: x=>[x|e] Ph. +move: (fr_st Pi Vij Pij)=>H; case: x=>[x|e] Ph. - case=>Ph'. case: (H _ _ Ph)=>h1[Eh V1 Ph1]; subst h. rewrite (pf_irr (validL Vij) Vi) in Ph1 *; move: (P _ _ Ph1)=> P21. @@ -588,27 +563,27 @@ case/H: Ph=>h1[Eh Vh Ph1]. by exists h1; split=>//; exists (Exn e), h1, Ph1. Qed. -Definition bind := Prog bind_sfmono bind_dstrict bind_frame. +Definition bnd := Prog bndprog_is_st. -Lemma vrf_bind i (Q : post B) : +Lemma vrf_bnd i (Q : post B) : vrf i e1 (fun x m => match x with | Val x' => vrf m (e2 x') Q | Exn e => valid m -> Q (Exn e) m end) -> - vrf i bind Q. + vrf i bnd Q. Proof. move=>H Vi; case: (H Vi)=>Hi {}H /=. -have Hi' : i \In bind_pre. +have Hi' : i \In bnd_pre. - by exists Vi, Hi=>x m Pm; case: (H _ _ Pm (dstr_valid Pm)). exists Hi'=>y j /= [x][m][Pm] C. -rewrite (pf_irr Hi (bind_pre_proj Hi')) in H. +rewrite (pf_irr Hi (bnd_pre_proj Hi')) in H. case: x Pm C=>[x|e] Pm; move: (H _ _ Pm (dstr_valid Pm))=>{}H. - by case=>Pm2 Pj; case: H=>Pm2'; apply; rewrite (pf_irr Pm2' Pm2). by case=>->->. Qed. -End Bind. +End Bnd. Section Try. Variables (A B : Type). @@ -633,32 +608,23 @@ Definition try_prog : prog try_pre B := m \In prog_of (e2 ex) _ (dstr_valid Ph) Ph' y end. -Lemma try_sfmono : safe_mono try_pre. +Lemma tryprog_is_st : st_axiom try_prog. Proof. -move=>i j [Vi [Pi][E1 E2]] Vij. -have Pij := sfm_st Pi Vij. -exists Vij, Pij; split. -- move=>y m. +split. +- move=>i j [Vi [Pi][E1 E2]] Vij; have Pij := sfm_st Pi Vij. + exists Vij, Pij; split. + - move=>y m. + case/(fr_st Pi Vij Pij)=>h [{m}-> Vhj]. + rewrite (pf_irr (validL Vij) Vi)=>/E1 Ph. + by apply: sfm_st=>//; apply: (validL Vhj). + move=>ex m. case/(fr_st Pi Vij Pij)=>h [{m}-> Vhj]. - rewrite (pf_irr (validL Vij) Vi)=>/E1 Ph. + rewrite (pf_irr (validL Vij) Vi)=>/E2 Ph. by apply: sfm_st=>//; apply: (validL Vhj). -move=>ex m. -case/(fr_st Pi Vij Pij)=>h [{m}-> Vhj]. -rewrite (pf_irr (validL Vij) Vi)=>/E2 Ph. -by apply: sfm_st=>//; apply: (validL Vhj). -Qed. - -Lemma try_dstrict : def_strict try_prog. -Proof. -move=>i [Vi [Pi][E1 E2]] Vi' y[x][h][/=]. -by case: x=>[x|ex] Eh; case=>Ph /dstr_st. -Qed. - -Lemma try_frame : frameable try_prog. -Proof. +- move=>i [Vi [Pi][E1 E2]] Vi' y [x][h][/=]. + by case: x=>[x|ex] Eh; case=>Ph /dstr_st. move=>i j [Vi [Pi][E1 E2]] Vij [_ [Pij _]] y m [x][h][/=]. -move: (fr_st Pi Vij Pij)=>H. -case: x=>[x|ex] Ph. +move: (fr_st Pi Vij Pij)=>H; case: x=>[x|ex] Ph. - case=>Ph'. case: (H _ _ Ph)=>h1[Eh V1 Ph1]; subst h. rewrite (pf_irr (validL Vij) Vi) in Ph1 *; move: (E1 _ _ Ph1)=>P21. @@ -675,7 +641,7 @@ exists h2; split=>//; exists (Exn ex), h1, Ph1, P21. by rewrite (pf_irr (dstr_valid Ph1) (validL V1)). Qed. -Definition try := Prog try_sfmono try_dstrict try_frame. +Definition try := Prog tryprog_is_st. Lemma vrf_try i (Q : post B) : vrf i e (fun x m => @@ -697,18 +663,18 @@ Qed. End Try. -(* don't export, just for fun *) +(* bnd follows from try; not to export, just for fun *) Lemma bnd_is_try A B (e1 : ST A) (e2 : A -> ST B) i r : vrf i (try e1 e2 (throw B)) r -> - vrf i (bind e1 e2) r. + vrf i (bnd e1 e2) r. Proof. move=>H Vi; case: (H Vi)=>[[Vi'][P1 /= [E1 E2]]] {}H. -have J : i \In pre_of (bind e1 e2). +have J : i \In pre_of (bnd e1 e2). - exists Vi, P1=>y m. by rewrite (pf_irr Vi Vi')=>/E1. exists J=>y m /= [x][h][Ph]C. apply: H; exists x, h=>/=. -rewrite (pf_irr P1 (bind_pre_proj J)) in E2 *; exists Ph. +rewrite (pf_irr P1 (bnd_pre_proj J)) in E2 *; exists Ph. move: Ph C; case: x=>// e Ph [{y}-> {m}->]. rewrite (pf_irr Vi' Vi) in E2. by exists (E2 _ _ Ph). @@ -717,8 +683,6 @@ Qed. Section Read. Variable (A : Type) (x : ptr). -Local Notation idyn v := (@dyn _ id _ v). - Definition read_pre : pre := fun i => x \in dom i /\ exists v : A, find x i = Some (idyn v). @@ -726,27 +690,23 @@ Definition read_prog : prog read_pre A := fun i _ _ y m => exists w, [/\ m = i, y = Val w & find x m = Some (idyn w)]. -Lemma read_sfmono : safe_mono read_pre. -Proof. -move=>i j [Hx [v E]] Vij; split. -- by rewrite domUn inE Vij Hx. -by exists v; rewrite findUnL // Hx. -Qed. - -Lemma read_dstrict : def_strict read_prog. -Proof. by move=>i _ Vi _ [_ [E _ _]]; rewrite -E in Vi. Qed. - -Lemma read_frame : frameable read_prog. +Lemma readprog_is_st : st_axiom read_prog. Proof. +split. +- move=>i j [Hx [v E]] Vij; split. + - by rewrite domUn inE Vij Hx. + by exists v; rewrite findUnL // Hx. +- by move=>i _ Vi _ [_ [E _ _]]; rewrite -E in Vi. move=>i j [Hx [v E]] Vij _ _ _ [w [-> -> H]]. exists i; split=>//; exists w; split=>{v E}//. by rewrite findUnL // Hx in H. Qed. -Definition read := Prog read_sfmono read_dstrict read_frame. +Definition read := Prog readprog_is_st. Lemma vrf_read j (v : A) (Q : post A) : - (valid (x :-> v \+ j) -> Q (Val v) (x :-> v \+ j)) -> + (valid (x :-> v \+ j) -> + Q (Val v) (x :-> v \+ j)) -> vrf (x :-> v \+ j) read Q. Proof. move=>H Vi /=. @@ -763,30 +723,20 @@ End Read. Section Write. Variable (A : Type) (x : ptr) (v : A). -Local Notation idyn v := (@dyn _ id _ v). - -Definition write_pre : pre := - fun i => x \in dom i. +Definition write_pre : pre := fun i => x \in dom i. Definition write_prog : prog write_pre unit := fun i _ _ y m => [/\ y = Val tt, m = upd x (idyn v) i & x \in dom i]. -Lemma write_sfmono : safe_mono write_pre. -Proof. -move=>i j; rewrite /write_pre -!toPredE /= => Hx Vij. -by rewrite domUn inE Vij Hx. -Qed. - -Lemma write_dstrict : def_strict write_prog. -Proof. -move=>i Hx _ _ [_ E _]; rewrite /write_pre -toPredE /= in Hx. -suff {E}: valid (upd x (idyn v) i) by rewrite -E. -by rewrite validU (dom_cond Hx) (dom_valid Hx). -Qed. - -Lemma write_frame : frameable write_prog. +Lemma writeprog_is_st : st_axiom write_prog. Proof. +split. +- move=>i j; rewrite /write_pre -!toPredE /= => Hx Vij. + by rewrite domUn inE Vij Hx. +- move=>i Hx _ _ [_ E _]; rewrite /write_pre -toPredE /= in Hx. + suff {E}: valid (upd x (idyn v) i) by rewrite -E. + by rewrite validU (dom_cond Hx) (dom_valid Hx). move=>i j Hx Vij _ _ _ [-> -> _]. exists (upd x (idyn v) i); split=>//; rewrite /write_pre -toPredE /= in Hx. @@ -794,10 +744,11 @@ rewrite /write_pre -toPredE /= in Hx. by rewrite validUUn. Qed. -Definition write := Prog write_sfmono write_dstrict write_frame. +Definition write := Prog writeprog_is_st. Lemma vrf_write B (u : B) j (Q : post unit) : - (valid (x :-> v \+ j) -> Q (Val tt) (x :-> v \+ j)) -> + (valid (x :-> v \+ j) -> + Q (Val tt) (x :-> v \+ j)) -> vrf (x :-> u \+ j) write Q. Proof. move=>H Vi /=. @@ -813,8 +764,6 @@ End Write. Section Allocation. Variables (A : Type) (v : A). -Local Notation idyn v := (@dyn _ id _ v). - Definition alloc_pre : pre := top. Definition alloc_prog : prog alloc_pre ptr := @@ -822,18 +771,12 @@ Definition alloc_prog : prog alloc_pre ptr := exists l, [/\ y = Val l, m = l :-> v \+ i, l != null & l \notin dom i]. -Lemma alloc_sfmono : safe_mono alloc_pre. -Proof. by []. Qed. - -Lemma alloc_dstrict : def_strict alloc_prog. -Proof. -move=>i [] Vi _ [l][_ E Hl Hl2]. -suff {E}: valid (l :-> v \+ i) by rewrite -E. -by rewrite validPtUn; apply/and3P. -Qed. - -Lemma alloc_frame : frameable alloc_prog. +Lemma allocprog_is_st : st_axiom alloc_prog. Proof. +split=>//. +- move=>i [] Vi _ [l][_ E Hl Hl2]. + suff {E}: valid (l :-> v \+ i) by rewrite -E. + by rewrite validPtUn; apply/and3P. move=>i j [] Vij [] _ _ [l][->-> Hl Hl2]. exists (l :-> v \+ i); rewrite -joinA; split=>//. - rewrite validUnAE validPt domPtK Hl Vij /= all_predC. @@ -843,10 +786,11 @@ exists l; split=>//. by apply/dom_NNL/Hl2. Qed. -Definition alloc := Prog alloc_sfmono alloc_dstrict alloc_frame. +Definition alloc := Prog allocprog_is_st. Lemma vrf_alloc i (Q : post ptr) : - (forall x, valid (x :-> v \+ i) -> Q (Val x) (x :-> v \+ i)) -> + (forall x, valid (x :-> v \+ i) -> + Q (Val x) (x :-> v \+ i)) -> vrf i alloc Q. Proof. move=>H Vi /=. @@ -865,21 +809,17 @@ Definition allocb_prog : prog allocb_pre ptr := fun i _ _ y m => exists l, [/\ y = Val l, m = updi l (nseq n v) \+ i & valid m]. -Lemma allocb_sfmono : safe_mono allocb_pre. -Proof. by []. Qed. - -Lemma allocb_dstrict : def_strict allocb_prog. -Proof. by move=>i [] Vi y [l][]. Qed. - -Lemma allocb_frame : frameable allocb_prog. +Lemma allocbprog_is_st : st_axiom allocb_prog. Proof. +split=>//. +- by move=>i [] Vi y [l][]. move=>i j [] Vij [] _ _ [l][->-> V]. exists (updi l (nseq n v) \+ i); rewrite -joinA; split=>//. exists l; split=>//. by rewrite joinA in V; apply: (validL V). Qed. -Definition allocb := Prog allocb_sfmono allocb_dstrict allocb_frame. +Definition allocb := Prog allocbprog_is_st. Lemma vrf_allocb i (Q : post ptr) : (forall x, valid (updi x (nseq n v) \+ i) -> @@ -903,33 +843,27 @@ Definition dealloc_prog : prog dealloc_pre unit := fun i _ _ y m => [/\ y = Val tt, m = free i x & x \in dom i]. -Lemma dealloc_sfmono : safe_mono dealloc_pre. -Proof. -move=>i j; rewrite /dealloc_pre -!toPredE /= => Hx Vij. -by rewrite domUn inE Vij Hx. -Qed. - -Lemma dealloc_dstrict : def_strict dealloc_prog. -Proof. -move=>i _ Vi _ [_ E _]. -suff {E}: valid (free i x) by rewrite -E. -by rewrite validF. -Qed. - -Lemma dealloc_frame : frameable dealloc_prog. +Lemma deallocprog_is_st : st_axiom dealloc_prog. Proof. +split. +- move=>i j; rewrite /dealloc_pre -!toPredE /= => Hx Vij. + by rewrite domUn inE Vij Hx. +- move=>i _ Vi _ [_ E _]. + suff {E}: valid (free i x) by rewrite -E. + by rewrite validF. move=>i j Hx Vij Hx' _ _ [->-> _]. exists (free i x); split=>//; rewrite /dealloc_pre -!toPredE /= in Hx Hx'. -- by apply/freeUnR/dom_inNL/Hx. +- by apply/freeUnL/dom_inNL/Hx. by apply: validFUn. Qed. -Definition dealloc := - Prog dealloc_sfmono dealloc_dstrict dealloc_frame. +Definition dealloc := Prog deallocprog_is_st. Lemma vrf_dealloc A (v : A) j (Q : post unit) : - (x \notin dom j -> valid j -> Q (Val tt) j) -> + (x \notin dom j -> + valid j -> + Q (Val tt) j) -> vrf (x :-> v \+ j) dealloc Q. Proof. move=>H Vi /=. @@ -945,39 +879,44 @@ End Deallocation. (* Monotonicity of the constructors *) Section Monotonicity. - -Variables (A B : Type). +Variables A B : Type. Lemma do_mono G (e1 e2 : ST A) (s : spec G A) - (pf1 : has_spec s e1) (pf2 : has_spec s e2) : - e1 <== e2 -> @STprog _ _ _ e1 pf1 <== @STprog _ _ _ e2 pf2. + (pf1 : has_spec s e1) (pf2 : has_spec s e2) : + e1 <== e2 -> + @STprog _ _ _ e1 pf1 <== @STprog _ _ _ e2 pf2. Proof. by []. Qed. Lemma bind_mono (e1 e2 : ST A) (k1 k2 : A -> ST B) : - e1 <== e2 -> k1 <== k2 -> (bind e1 k1 : ST B) <== bind e2 k2. + e1 <== e2 -> + k1 <== k2 -> + (bnd e1 k1 : ST B) <== bnd e2 k2. Proof. move=>[H1 H2] pf2. -have pf: bind_pre e2 k2 <== bind_pre e1 k1. +have pf: bnd_pre e2 k2 <== bnd_pre e1 k1. - move=>h [Vh][Pi] H. exists Vh, (H1 _ Pi)=>x m /H2/H H'. by case: (pf2 x)=>+ _; apply. exists pf=>i Vi /[dup][[Vi'][Pi']P'] Pi x h. case; case=>[a|e][h0][Ph][Ph'] H. - exists (Val a), h0. - move: (H2 i Vi (bind_pre_proj Pi))=>H'. - rewrite (pf_irr (H1 i (bind_pre_proj Pi)) (bind_pre_proj (pf i Pi))) in H'. + move: (H2 i Vi (bnd_pre_proj Pi))=>H'. + rewrite (pf_irr (H1 i (bnd_pre_proj Pi)) (bnd_pre_proj (pf i Pi))) in H'. have Ph0 := (H' _ _ Ph); exists Ph0. - move: (P' a h0); rewrite (pf_irr Vi' Vi) (pf_irr Pi' (bind_pre_proj Pi))=>H''. + move: (P' a h0); rewrite (pf_irr Vi' Vi) (pf_irr Pi' (bnd_pre_proj Pi))=>H''. exists (H'' Ph0); case: (pf2 a)=>Pr; apply. - by rewrite (pf_irr (dstr_valid Ph0) (dstr_valid Ph)) (pf_irr (Pr h0 (H'' Ph0)) Ph'). + rewrite (pf_irr (dstr_valid Ph0) (dstr_valid Ph)). + by rewrite (pf_irr (Pr h0 (H'' Ph0)) Ph'). rewrite Ph' H; exists (Exn e), h0. -move: (H2 i Vi (bind_pre_proj Pi))=>H'. -rewrite (pf_irr (H1 i (bind_pre_proj Pi)) (bind_pre_proj (pf i Pi))) in H'. +move: (H2 i Vi (bnd_pre_proj Pi))=>H'. +rewrite (pf_irr (H1 i (bnd_pre_proj Pi)) (bnd_pre_proj (pf i Pi))) in H'. by exists (H' _ _ Ph). Qed. Lemma try_mono (e1 e2 : ST A) (k1 k2 : A -> ST B) (h1 h2 : exn -> ST B) : - e1 <== e2 -> k1 <== k2 -> h1 <== h2 -> + e1 <== e2 -> + k1 <== k2 -> + h1 <== h2 -> (try e1 k1 h1 : ST B) <== try e2 k2 h2. Proof. move=>[H1 H2] pf2 pf3. @@ -994,81 +933,93 @@ case; case=>[a|e][h0][Ph][Ph'] H. move: (H2 i Vi (try_pre_proj Pi))=>H'. rewrite (pf_irr (H1 i (try_pre_proj Pi)) (try_pre_proj (pf i Pi))) in H'. have Ph1 := (H' _ _ Ph); exists Ph1. - move: (Pk0 a h0); rewrite (pf_irr Vi' Vi) (pf_irr Pi' (try_pre_proj Pi))=>H''. + move: (Pk0 a h0); rewrite (pf_irr Vi' Vi). + rewrite (pf_irr Pi' (try_pre_proj Pi))=>H''. exists (H'' Ph1); case: (pf2 a)=>Pr; apply. - by rewrite (pf_irr (dstr_valid Ph1) (dstr_valid Ph)) (pf_irr (Pr h0 (H'' Ph1)) Ph'). + rewrite (pf_irr (dstr_valid Ph1) (dstr_valid Ph)). + by rewrite (pf_irr (Pr h0 (H'' Ph1)) Ph'). exists (Exn e), h0. move: (H2 i Vi (try_pre_proj Pi))=>H'. rewrite (pf_irr (H1 i (try_pre_proj Pi)) (try_pre_proj (pf i Pi))) in H'. have Ph1 := (H' _ _ Ph); exists Ph1. move: (Ph0 e h0); rewrite (pf_irr Vi' Vi) (pf_irr Pi' (try_pre_proj Pi))=>H''. exists (H'' Ph1); case: (pf3 e)=>Pr; apply. -by rewrite (pf_irr (dstr_valid Ph1) (dstr_valid Ph)) (pf_irr (Pr h0 (H'' Ph1)) Ph'). +rewrite (pf_irr (dstr_valid Ph1) (dstr_valid Ph)). +by rewrite (pf_irr (Pr h0 (H'' Ph1)) Ph'). Qed. (* the rest of the constructors are trivial *) Lemma ret_mono (v1 v2 : A) : - v1 = v2 -> (ret v1 : ST A) <== ret v2. + v1 = v2 -> + (ret v1 : ST A) <== ret v2. Proof. by move=>->. Qed. Lemma throw_mono (e1 e2 : exn) : - e1 = e2 -> (throw A e1 : ST A) <== throw A e2. + e1 = e2 -> + (throw A e1 : ST A) <== throw A e2. Proof. by move=>->. Qed. Lemma read_mono (p1 p2 : ptr) : - p1 = p2 -> (read A p1 : ST A) <== read A p2. + p1 = p2 -> + (read A p1 : ST A) <== read A p2. Proof. by move=>->. Qed. Lemma write_mono (p1 p2 : ptr) (x1 x2 : A) : - p1 = p2 -> x1 = x2 -> (write p1 x1 : ST unit) <== write p2 x2. + p1 = p2 -> + x1 = x2 -> + (write p1 x1 : ST unit) <== write p2 x2. Proof. by move=>->->. Qed. Lemma alloc_mono (x1 x2 : A) : - x1 = x2 -> (alloc x1 : ST ptr) <== alloc x2. + x1 = x2 -> + (alloc x1 : ST ptr) <== alloc x2. Proof. by move=>->. Qed. Lemma allocb_mono (x1 x2 : A) (n1 n2 : nat) : - x1 = x2 -> n1 = n2 -> (allocb x1 n1 : ST ptr) <== allocb x2 n2. + x1 = x2 -> + n1 = n2 -> + (allocb x1 n1 : ST ptr) <== allocb x2 n2. Proof. by move=>->->. Qed. Lemma dealloc_mono (p1 p2 : ptr) : - p1 = p2 -> (dealloc p1 : ST unit) <== dealloc p2. + p1 = p2 -> + (dealloc p1 : ST unit) <== dealloc p2. Proof. by move=>->. Qed. -Variables (G : Type) (C : A -> Type) (s : forall x, spec G (C x)). -Notation lat := (dfunLattice (fun x => [lattice of STspec (s x)])). +Variables (G : Type) (C : A -> Type) (pq : forall x, spec G (C x)). +Notation lat := (dfun_lattice (fun x => STspec (pq x))). -Lemma fix_mono (f1 f2 : lat -> lat) : f1 <== f2 -> (Fix f1 : lat) <== Fix f2. +Lemma fix_mono (f1 f2 : lat -> lat) : + f1 <== f2 -> + (ffix f1 : lat) <== ffix f2. Proof. -move=>Hf; apply: tarski_lfp_mono. -- move=>x1 x2 Hx; apply: supM=>z [x][H ->]; apply: supP; exists x. +move=>Hf; rewrite /ffix. +(* f' f2 is monotone closure of f2; hence it's monotone *) +have M : monofun_axiom (f' f2). +- move=>x1 x2 Hx; apply: supM=>z [x][H ->]; apply: supP; exists x. by split=>//; apply: poset_trans H Hx. +set ff := MonoFun.pack_ (isMonoFun.Build _ _ (f' f2) M). +apply: (tarski_lfp_mono (f2:=ff)). move=>y; apply: supM=>_ [x][H1 ->]. by apply: poset_trans (Hf x) _; apply: supP; exists x. Qed. End Monotonicity. -Section MonotonicityExamples. - Notation "'Do' e" := (@STprog _ _ _ e _) (at level 80). - -Notation "x '<--' c1 ';' c2" := (bind c1 (fun x => c2)) +Notation "x '<--' c1 ';' c2" := (bnd c1 (fun x => c2)) (at level 81, right associativity). -Notation "c1 ';;' c2" := (bind c1 (fun _ => c2)) +Notation "c1 ';;' c2" := (bnd c1 (fun _ => c2)) (at level 81, right associativity). Notation "'!' x" := (read x) (at level 50). Notation "x '::=' e" := (write x e) (at level 60). -Fixpoint fact (x : nat) := if x is x'.+1 then x * fact x' else 1. - -Definition facttp := forall x : nat, @STspec unit nat - (fun => (fun i => True, - fun v m => if v is Val w then w = fact x else False)). +(* Some examples *) -(* for the example, we need a quick lemma for function calls *) +(* the examples require a lemma for instantiating ghosts *) +(* when doing function calls *) (* this need not be exported, as it follows from the signature *) -(* will reprove it outside of the module *) +(* and will be reproven outside of the Vrf module *) Lemma gE G A (s : spec G A) g i (e : @STspec G A s) (Q : post A) : (s g).1 i -> @@ -1082,34 +1033,40 @@ case: e=>e /= /[apply] Hp Hv He; apply: vrfV=>V /=. by apply/vrf_post/Hp; case=>[v|ex] m Vm H; [apply: Hv | apply: He]. Qed. -(* now we can do the example with monotonicity explicitly *) +(* factorial *) +Fixpoint fact (x : nat) := if x is x'.+1 then x * fact x' else 1. +Definition facttp := forall x : nat, @STspec unit nat + (fun => (fun i => True, + fun v m => if v is Val w then w = fact x else False)). + +(* proof using ffix2, with monotonicity proved explicitly *) Program Definition factx := - Fix' (fun (loop : facttp) (x : nat) => + ffix2 (fun (loop : facttp) (x : nat) => Do (if x is x'.+1 then t <-- loop x'; ret (x * t) else ret 1)) _. Next Obligation. case=>i /= _; case: x; first by apply: vrf_ret. -by move=>n; apply: vrf_bind=>//; apply: gE=>// x m -> _; apply: vrf_ret. +by move=>n; apply: vrf_bnd=>//; apply: gE=>// x m -> _; apply: vrf_ret. Qed. Next Obligation. move=>x1 x2 H; case=>[|n]; first by apply: poset_refl. by apply: do_mono; apply: bind_mono=>//; apply: H. Qed. -(* Monotonocity works even if we compute *) +(* monotonocity can be proved even if we compute *) (* with propositions, and deliberately invert the polarity. *) Definition proptp := unit -> @STspec unit Prop (fun => (fun i => True, fun R m => True)). Program Definition propx := - Fix' (fun (loop : proptp) (x : unit) => + ffix2 (fun (loop : proptp) (x : unit) => Do (R <-- loop tt; ret (not R))) _. Next Obligation. -case=>i _; apply: vrf_bind=>//. +case=>i _; apply: vrf_bnd=>//. by apply: gE=>//= R m _ C; apply: vrf_ret. Qed. Next Obligation. @@ -1121,10 +1078,10 @@ Definition proptp2 := Prop -> @STspec unit Prop (fun => (fun i => True, fun (Ro : ans Prop) m => True)). Program Definition propx2 := - Fix' (fun (loop : proptp2) (x : Prop) => + ffix2 (fun (loop : proptp2) (x : Prop) => Do (R <-- loop (x -> x); ret R)) _ True. Next Obligation. -case=>i /= _; apply: vrf_bind=>//. +case=>i /= _; apply: vrf_bnd=>//. by apply: gE=>//= R m _ _; apply: vrf_ret. Qed. Next Obligation. @@ -1137,10 +1094,10 @@ Definition proptp3 := forall R : Prop, @STspec unit Prop fun (Ro : ans Prop) m => if Ro is Val Ro' then Ro' else False)). Program Definition propx3 := - Fix' (fun (loop : proptp3) (x : Prop) => + ffix2 (fun (loop : proptp3) (x : Prop) => Do (R <-- loop (x -> x); ret R)) _ True. Next Obligation. -case=>i /= _; apply: vrf_bind=>//. +case=>i /= _; apply: vrf_bnd=>//. by apply: gE=>//= R m X _; apply: vrf_ret. Qed. Next Obligation. @@ -1150,16 +1107,13 @@ Qed. (* It seems safe to say that monotonicity is always easy *) (* to prove for all the programs that we expect to write. *) -(* Thus, we will export Fix over monotone closure, but not Fix'. *) -(* Exporting Fix has the advantage of eliding the always simple *) -(* and syntax-directed proofs of monotonicity, which we just don't *) -(* want to bother with. *) +(* Thus, the module will export ffix over monotone closure, but not ffix2. *) +(* Exporting ffix has the advantage of eliding the always simple *) +(* and syntax-directed proofs of monotonicity. *) (* This is in line with the existing proofs of soundness of HTT as a *) -(* standalone theory, where you can show that all functions that can be *) -(* produced using the syntax are monotone in the model. That proof *) -(* relied on a significant subset of Coq, but still a subset. *) - +(* standalone theory, which shows that all functions that can be *) +(* produced using the declared syntax are monotone in the model. *) (* The paper is: *) (* Partiality, State and Dependent Types *) (* Kasper Svendsen, Lars Birkedal and Aleksandar Nanevski *) @@ -1167,31 +1121,28 @@ Qed. (* pages 198-212, 2011. *) (* The model in that paper, and the one we use here, are quite different *) -(* however, and we can't exclude concocting a *) -(* non-monotone program using the syntax we export here, *) +(* however, and one can't exclude concocting a *) +(* non-monotone program using the syntax exported here, *) (* based on the results of that other paper. *) -(* That said, we have so far failed to do produce such a non-monotone function. *) -(* But if such a case arises, we should switch to using Fix' *) +(* That said, so far there we failed to produce such a non-monotone function. *) +(* But if such a case arises, we should switch to using ffix2 *) (* and ask for explicit proofs of monotonicity. *) -(* For all the examples we did so far, such a proof is easily *) +(* For all the examples done so far, such a proof is easily *) (* constructed, along the above lines (moreover, such proofs ought to be automatable) *) (* Perhaps a proof, in the theory of Coq, that all functions are monotone *) -(* can be produced if we worked from parametricity properties of Coq, *) +(* can be produced if one worked from parametricity properties of Coq, *) (* which have been established in the past work, and even internalized *) (* into Coq by means of parametricity axioms. But that remains future work. *) -End MonotonicityExamples. - End Vrf. - Export Vrf. +(* some helpers *) Definition skip := ret tt. -Corollary vrf_mono A (e : ST A) i : monotone (vrf' e i). +Lemma vrf_mono A (e : ST A) i : monofun_axiom (vrf' e i). Proof. by move=>/= Q1 Q2 H; apply: vrf_post=>y m _; apply: H. Qed. - (******************************************) (* Notation for logical variable postexts *) (******************************************) @@ -1203,11 +1154,17 @@ Definition logvar {B A} (G : A -> Type) (s : forall x : A, spec (G x) B) : spec {x : A & G x} B := fun '(existT x g) => s x g. -Notation "'STsep' ( p , q ) " := (STspec unit (logbase p q)) (at level 0). +Notation "'STsep' ( p , q ) " := + (STspec unit (fun => (p, q))) (at level 0, + format "'[hv ' STsep ( '[' p , '/' q ']' ) ']'"). -Notation "{ x .. y }, 'STsep' ( p , q ) " := +Notation "'STsep' { x .. y } ( p , q )" := (STspec _ (logvar (fun x => .. (logvar (fun y => logbase p q)) .. ))) - (at level 0, x binder, y binder, right associativity). + (at level 0, x binder, y binder, right associativity, + format "'[hv ' STsep { x .. y } '/ ' ( '[' p , '/' q ']' ) ']'"). + +Notation "[> h1 , .. , hn <]" := + (existT _ h1 .. (existT _ hn tt) .. ) (at level 0). (************************************************************) (* Lemmas for pulling out and instantiating ghost variables *) @@ -1216,11 +1173,11 @@ Notation "{ x .. y }, 'STsep' ( p , q ) " := (* Lemmas without framing, i.e. they pass the entire heap to the *) (* routine being invoked. *) -Lemma gE G A (s : spec G A) g i (e : STspec G s) (Q : post A) : - (s g).1 i -> - (forall v m, (s g).2 (Val v) m -> +Lemma gE G A (pq : spec G A) (e : STspec G pq) g (Q : post A) i : + (pq g).1 i -> + (forall v m, (pq g).2 (Val v) m -> valid m -> Q (Val v) m) -> - (forall x m, (s g).2 (Exn x) m -> + (forall x m, (pq g).2 (Exn x) m -> valid m -> Q (Exn x) m) -> vrf i e Q. Proof. @@ -1228,49 +1185,50 @@ case: e=>e /= /[apply] Hp Hv He; apply: vrfV=>V /=. by apply/vrf_post/Hp; case=>[v|ex] m Vm H; [apply: Hv | apply: He]. Qed. -Arguments gE [G A s] g [i e Q] _ _. - -Notation "[gE]" := (gE tt) (at level 0). +Arguments gE {G A pq e} g {Q}. +Notation "[gE]" := (gE tt) (at level 0). Notation "[ 'gE' x1 , .. , xn ]" := (gE (existT _ x1 .. (existT _ xn tt) ..)) - (at level 0, format "[ 'gE' x1 , .. , xn ]"). - -(* a combination of gE + vrf_bind, for "stepping over" the call *) -Lemma stepE G A B (s : spec G A) g i (e : STspec G s) (e2 : A -> ST B) (Q : post B) : - (s g).1 i -> - (forall x m, (s g).2 (Val x) m -> vrf m (e2 x) Q) -> - (forall x m, (s g).2 (Exn x) m -> +(at level 0, format "[ 'gE' x1 , .. , xn ]"). + +(* combination of gE + vrf_bind, for "stepping over" the call *) +Lemma stepE G A B (pq : spec G A) (e : STspec G pq) + (e2 : A -> ST B) i (Q : post B) (g : G) : + (pq g).1 i -> + (forall x m, (pq g).2 (Val x) m -> vrf m (e2 x) Q) -> + (forall x m, (pq g).2 (Exn x) m -> valid m -> Q (Exn x) m) -> - vrf i (bind e e2) Q. + vrf i (bnd e e2) Q. Proof. -move=>Hp Hv He. -by apply/vrf_bind/(gE _ Hp)=>[v m P|x m P _] V; [apply: Hv | apply: He]. +move=>Hp Hv He; apply/vrf_bnd/(gE _ _ Hp). +- by move=>v m P V; apply: Hv. +by move=>x m P _ V; apply: He. Qed. -Arguments stepE [G A B s] g [i e e2 Q] _ _. +Arguments stepE {G A B pq e e2 i Q}. Notation "[stepE]" := (stepE tt) (at level 0). - Notation "[ 'stepE' x1 , .. , xn ]" := (stepE (existT _ x1 .. (existT _ xn tt) ..)) - (at level 0, format "[ 'stepE' x1 , .. , xn ]"). - -(* a combination of gE + vrf_try *) -Lemma tryE G A B (s : spec G A) g i (e : STspec G s) (e1 : A -> ST B) (e2 : exn -> ST B) (Q : post B) : - (s g).1 i -> - (forall x m, (s g).2 (Val x) m -> vrf m (e1 x) Q) -> - (forall x m, (s g).2 (Exn x) m -> vrf m (e2 x) Q) -> +(at level 0, format "[ 'stepE' x1 , .. , xn ]"). + +(* combination of gE + vrf_try *) +Lemma tryE G A B (pq : spec G A) (e : STspec G pq) (e1 : A -> ST B) + (e2 : exn -> ST B) i (Q : post B) (g : G) : + (pq g).1 i -> + (forall x m, (pq g).2 (Val x) m -> vrf m (e1 x) Q) -> + (forall x m, (pq g).2 (Exn x) m -> vrf m (e2 x) Q) -> vrf i (try e e1 e2) Q. Proof. -move=>Hp Hv Hx. -by apply/vrf_try/(gE _ Hp)=>[x|ex] m Vm P; [apply: Hv | apply: Hx]. +move=>Hp Hv Hx; apply/vrf_try/(gE _ _ Hp). +- by move=>x m Vm P; apply: Hv. +by move=>ex m Vm P; apply: Hx. Qed. -Arguments tryE [G A B s] g [i e e1 e2 Q] _ _. +Arguments tryE {G A B pq e e1 e2 i Q}. Notation "[tryE]" := (tryE tt) (at level 0). - Notation "[ 'tryE' x1 , .. , xn ]" := (tryE (existT _ x1 .. (existT _ xn tt) ..)) (at level 0, format "[ 'tryE' x1 , .. , xn ]"). @@ -1279,11 +1237,11 @@ Notation "[ 'tryE' x1 , .. , xn ]" := (* empty heap to the routine. For more sophisticated framing *) (* variants see the `heapauto` module. *) -Lemma gU G A (s : spec G A) g i (e : STspec G s) (Q : post A) : - (s g).1 Unit -> - (forall v m, (s g).2 (Val v) m -> +Lemma gU G A (pq : spec G A) (e : STspec G pq) i (Q : post A) g : + (pq g).1 Unit -> + (forall v m, (pq g).2 (Val v) m -> valid (m \+ i) -> Q (Val v) (m \+ i)) -> - (forall x m, (s g).2 (Exn x) m -> + (forall x m, (pq g).2 (Exn x) m -> valid (m \+ i) -> Q (Exn x) (m \+ i)) -> vrf i e Q. Proof. @@ -1293,69 +1251,62 @@ by case=>[x|ex] n _ =>[/Hv|/Hx]. Qed. Notation "[gU]" := (gU tt) (at level 0). - Notation "[ 'gU' x1 , .. , xn ]" := (gU (existT _ x1 .. (existT _ xn tt) ..)) (at level 0, format "[ 'gU' x1 , .. , xn ]"). -(* a combination of gU + vrf_bind *) -Lemma stepU G A B (s : spec G A) g i (e : STspec G s) (e2 : A -> ST B) - (Q : post B) : - (s g).1 Unit -> - (forall x m, (s g).2 (Val x) m -> vrf (m \+ i) (e2 x) Q) -> - (forall x m, (s g).2 (Exn x) m -> +(* combination of gU + vrf_bind *) +Lemma stepU G A B (pq : spec G A) (e : STspec G pq) (e2 : A -> ST B) + i (Q : post B) g : + (pq g).1 Unit -> + (forall x m, (pq g).2 (Val x) m -> vrf (m \+ i) (e2 x) Q) -> + (forall x m, (pq g).2 (Exn x) m -> valid (m \+ i) -> Q (Exn x) (m \+ i)) -> - vrf i (bind e e2) Q. + vrf i (bnd e e2) Q. Proof. -move=>Hp Hv Hx. -apply/vrf_bind/(gU _ Hp)=>[x m H V|ex m H V _]. -- by apply: Hv H. -by apply: Hx. +move=>Hp Hv Hx; apply/vrf_bnd/(gU _ Hp). +- by move=>x m H V; apply: Hv H. +by move=>ex m H V _; apply: Hx. Qed. -Arguments stepU [G A B s] g i [e e2 Q] _ _ _. +Arguments stepU {G A B pq e e2 i Q}. Notation "[stepU]" := (stepU tt) (at level 0). - Notation "[ 'stepU' x1 , .. , xn ]" := (stepU (existT _ x1 .. (existT _ xn tt) ..)) (at level 0, format "[ 'stepU' x1 , .. , xn ]"). -(* a combination of gU + vrf_try *) -Lemma tryU G A B (s : spec G A) g i (e : STspec G s) - (e1 : A -> ST B) (e2 : exn -> ST B) (Q : post B) : - (s g).1 Unit -> - (forall x m, (s g).2 (Val x) m -> vrf (m \+ i) (e1 x) Q) -> - (forall x m, (s g).2 (Exn x) m -> vrf (m \+ i) (e2 x) Q) -> +(* combination of gU + vrf_try *) +Lemma tryU G A B (pq : spec G A) (e : STspec G pq) + (e1 : A -> ST B) (e2 : exn -> ST B) i (Q : post B) g : + (pq g).1 Unit -> + (forall x m, (pq g).2 (Val x) m -> vrf (m \+ i) (e1 x) Q) -> + (forall x m, (pq g).2 (Exn x) m -> vrf (m \+ i) (e2 x) Q) -> vrf i (try e e1 e2) Q. Proof. -move=>Hi H1 H2. -apply/vrf_try/(gU _ Hi)=>[x|ex] m H V. -- by apply: H1 H. -by apply: H2. +move=>Hi H1 H2; apply/vrf_try/(gU _ Hi). +- by move=>x m H V; apply: H1 H. +by move=>ex m H V; apply: H2. Qed. -Arguments tryU [G A B s] g i [e e1 e2 Q] _ _ _. +Arguments tryU {G A B pq e e1 e2 i Q}. Notation "[tryU]" := (tryU tt) (at level 0). - Notation "[ 'tryU' x1 , .. , xn ]" := (tryU (existT _ x1 .. (existT _ xn tt) ..)) (at level 0, format "[ 'tryU' x1 , .. , xn ]"). -(* some notation for writing posts that signify no exceptions are raised *) +(* notation for writing posts that signify *) +(* that no exception is raised *) Definition vfun' A (f : A -> heap -> Prop) : post A := fun y i => if y is Val v then f v i else False. Notation "[ 'vfun' x => p ]" := (vfun' (fun x => p)) (at level 0, x name, format "[ 'vfun' x => p ]") : fun_scope. - Notation "[ 'vfun' x : aT => p ]" := (vfun' (fun (x : aT) => p)) (at level 0, x name, only parsing) : fun_scope. - Notation "[ 'vfun' x i => p ]" := (vfun' (fun x i => p)) (at level 0, x name, format "[ 'vfun' x i => p ]") : fun_scope. - Notation "[ 'vfun' ( x : aT ) i => p ]" := (vfun' (fun (x : aT) i => p)) (at level 0, x name, only parsing) : fun_scope. diff --git a/htt/options.v b/htt/options.v new file mode 100644 index 0000000..94dba88 --- /dev/null +++ b/htt/options.v @@ -0,0 +1,12 @@ +(* turn off the automation of Program *) +#[export] Obligation Tactic := auto. +(* turn off other Program extensions *) +#[export] Unset Program Cases. +(* commenting this out allows a bit more convenience *) +(* when working with tuples, as one can pass a proof *) +(* of a "wrong" fact, and Program would emit obligation *) +(* that wrong fact equals the right fact. *) +(* If left uncommented, the equality will have to be *) +(* witnessed manually in the program *) +(* #[export] Unset Program Generalized Coercion. *) +#[export] Unset Program Mode. diff --git a/meta.yml b/meta.yml index 3944bf1..9459fe1 100644 --- a/meta.yml +++ b/meta.yml @@ -1,7 +1,7 @@ fullname: Hoare Type Theory shortname: htt organization: imdea-software -opam_name: coq-htt +opam_name: coq-htt-core community: false action: true dune: true @@ -58,6 +58,8 @@ authors: initial: false - name: Alexander Gryzlov initial: false +- name: Marcos Grandury + initial: false publications: - pub_url: https://software.imdea.org/~aleks/papers/reflect/reflect.pdf @@ -78,36 +80,32 @@ license: file: LICENSE supported_coq_versions: - text: Coq 8.15 to 8.17 - opam: '{ (>= "8.15" & < "8.19~") | (= "dev") }' + text: Coq 8.19 to 8.20 + opam: '{ (>= "8.19" & < "8.21~") | (= "dev") }' tested_coq_opam_versions: -- version: '1.17.0-coq-8.15' +- version: '2.2.0-coq-8.19' + repo: 'mathcomp/mathcomp' +- version: '2.2.0-coq-8.20' repo: 'mathcomp/mathcomp' -- version: '1.17.0-coq-8.18' +- version: 'latest-coq-dev' repo: 'mathcomp/mathcomp' -# - version: 'coq-dev' -# repo: 'mathcomp/mathcomp-dev' dependencies: - opam: name: coq-mathcomp-ssreflect - version: '{ (>= "1.17.0" & < "1.18~") | (= "dev") }' + version: '{ (>= "2.2.0" & < "2.3~") | (= "dev") }' description: |- - [MathComp ssreflect 1.17](https://math-comp.github.io) + [MathComp ssreflect 2.2](https://math-comp.github.io) - opam: name: coq-mathcomp-algebra description: |- [MathComp algebra](https://math-comp.github.io) -- opam: - name: coq-mathcomp-fingroup - description: |- - [MathComp fingroup](https://math-comp.github.io) - opam: name: coq-fcsl-pcm - version: '{ (>= "1.8.0" & < "1.9~") | (= "dev") }' + version: '{ (>= "2.0.0" & < "2.1~") | (= "dev") }' description: |- - [FCSL-PCM 1.8](https://github.com/imdea-software/fcsl-pcm) + [FCSL-PCM 2.0](https://github.com/imdea-software/fcsl-pcm) namespace: htt diff --git a/regenerate.sh b/regenerate.sh index 1e1d9f5..47ca0d5 100755 --- a/regenerate.sh +++ b/regenerate.sh @@ -1,7 +1,7 @@ #!/usr/bin/env bash TMP=$(mktemp -d); git clone https://github.com/coq-community/templates.git $TMP -$TMP/generate.sh README.md coq-htt.opam dune-project +$TMP/generate.sh README.md coq-htt-core.opam dune-project echo "Proceeding with customized generation..." @@ -24,7 +24,7 @@ for f in "$srcdir"/*.mustache; do mustache='{{ dune }}' bool=$(get_yaml meta.yml <<<"$mustache") if [ -n "$bool" ] && [ "$bool" != false ]; then - mkdir -p -v theories && target="theories/$target" + mkdir -p -v htt && target="htt/$target" else continue fi diff --git a/templates-extra/coq-htt-examples.opam.mustache b/templates-extra/coq-htt.opam.mustache similarity index 95% rename from templates-extra/coq-htt-examples.opam.mustache rename to templates-extra/coq-htt.opam.mustache index 2d53c31..91e12ca 100644 --- a/templates-extra/coq-htt-examples.opam.mustache +++ b/templates-extra/coq-htt.opam.mustache @@ -1,4 +1,4 @@ -# This file was generated from `meta.yml`, please do not edit manually. +; This file was generated from `meta.yml`, please do not edit manually. opam-version: "2.0" maintainer: "{{& opam-file-maintainer }}{{^ opam-file-maintainer }}palmskog@gmail.com{{/ opam-file-maintainer }}" diff --git a/templates-extra/docker-action.yml.mustache b/templates-extra/docker-action.yml.mustache index 06e35a6..89cf25b 100644 --- a/templates-extra/docker-action.yml.mustache +++ b/templates-extra/docker-action.yml.mustache @@ -1,4 +1,5 @@ -# This file was generated from `meta.yml`, please do not edit manually. +; This file was generated from `meta.yml`, please do not edit manually. + name: Docker CI on: @@ -32,7 +33,7 @@ jobs: {{/ submodule }} - uses: coq-community/docker-coq-action@v1 with: - opam_file: 'coq-htt-examples.opam' + opam_file: 'coq-htt.opam' {{! Change delimiters to avoid the next line being parsed as mustache syntax. }} {{=<% %>=}} custom_image: ${{ matrix.image }} diff --git a/templates-extra/dune.mustache b/templates-extra/dune.mustache index 0d8e5fa..351cc62 100644 --- a/templates-extra/dune.mustache +++ b/templates-extra/dune.mustache @@ -8,4 +8,4 @@ -w -notation-overridden -w -local-declaration -w -redundant-canonical-projection - -w -projection-no-head-constant)) \ No newline at end of file + -w -projection-no-head-constant)) diff --git a/theories/domain.v b/theories/domain.v deleted file mode 100644 index e287984..0000000 --- a/theories/domain.v +++ /dev/null @@ -1,1367 +0,0 @@ -From mathcomp Require Import ssreflect ssrbool ssrfun ssrnat eqtype seq. -From pcm Require Import options axioms pred prelude. - -(**********) -(* Posets *) -(**********) - -(* We put bottom in posets right away, instead of adding it later both in *) -(* lattices and in cpos. Since we never consider bottom-less posets, this *) -(* saves some tedium and name space. *) - -Module Poset. - -Section RawMixin. - -Record mixin_of (T : Type) := Mixin { - mx_leq : T -> T -> Prop; - mx_bot : T; - _ : forall x, mx_leq mx_bot x; - _ : forall x, mx_leq x x; - _ : forall x y, mx_leq x y -> mx_leq y x -> x = y; - _ : forall x y z, mx_leq x y -> mx_leq y z -> mx_leq x z}. - -End RawMixin. - -Section ClassDef. - -Record class_of T := Class {mixin : mixin_of T}. - -Structure type : Type := Pack {sort : Type; _ : class_of sort}. -Local Coercion sort : type >-> Sortclass. - -Variables (T : Type) (cT : type). -Definition class := let: Pack _ c as cT' := cT return class_of cT' in c. -Definition clone c of phant_id class c := @Pack T c. - -(* produce a poset type out of the mixin *) -(* equalize m0 and m by means of a phantom *) -Definition pack (m0 : mixin_of T) := - fun m & phant_id m0 m => Pack (@Class T m). - -Definition leq := mx_leq (mixin class). -Definition bot := mx_bot (mixin class). - -End ClassDef. - -Module Exports. -Coercion sort : type >-> Sortclass. -Notation poset := Poset.type. -Notation PosetMixin := Poset.Mixin. -Notation Poset T m := (@pack T _ m id). - -Notation "[ 'poset' 'of' T 'for' cT ]" := (@clone T cT _ id) - (at level 0, format "[ 'poset' 'of' T 'for' cT ]") : form_scope. -Notation "[ 'poset' 'of' T ]" := (@clone T _ _ id) - (at level 0, format "[ 'poset' 'of' T ]") : form_scope. - -Notation "x <== y" := (Poset.leq x y) (at level 70). -Notation bot := Poset.bot. - -Arguments Poset.bot {cT}. -Prenex Implicits Poset.bot. - -(* re-state lattice properties using the exported notation *) -Section Laws. -Variable T : poset. - -Lemma botP (x : T) : bot <== x. -Proof. by case: T x=>s [[leq b B]]. Qed. - -Lemma poset_refl (x : T) : x <== x. -Proof. by case: T x=>S [[leq b B R]]. Qed. - -Lemma poset_asym (x y : T) : x <== y -> y <== x -> x = y. -Proof. by case: T x y=>S [[l b B R A Tr]] *; apply: (A). Qed. - -Lemma poset_trans (y x z : T) : x <== y -> y <== z -> x <== z. -Proof. by case: T y x z=>S [[l b B R A Tr]] ? x y z; apply: (Tr). Qed. - -End Laws. - -#[export] -Hint Resolve botP poset_refl : core. - -Add Parametric Relation (T : poset) : T (@Poset.leq T) - reflexivity proved by (@poset_refl _) - transitivity proved by (fun x y => @poset_trans _ y x) as poset_rel. - -End Exports. - -End Poset. - -Export Poset.Exports. - -(**************************) -(* some basic definitions *) -(**************************) - -Definition monotone (T1 T2 : poset) (f : T1 -> T2) := - forall x y, x <== y -> f x <== f y. - -Section IdealDef. -Variable T : poset. - -Structure ideal (P : T) := Ideal {id_val : T; id_pf : id_val <== P}. - -(* Changing the type of the ideal *) - -Lemma relaxP (P1 P2 : T) : P1 <== P2 -> forall p, p <== P1 -> p <== P2. -Proof. by move=>H1 p H2; apply: poset_trans H1. Qed. - -Definition relax (P1 P2 : T) (x : ideal P1) (pf : P1 <== P2) := - Ideal (relaxP pf (id_pf x)). - -End IdealDef. - -(***********************) -(* poset constructions *) -(***********************) - -Section SubPoset. -Variables (T : poset) (s : Pred T) (C : bot \In s). - -Local Notation tp := {x : T | x \In s}. - -Definition sub_bot : tp := exist _ bot C. -Definition sub_leq (p1 p2 : tp) := sval p1 <== sval p2. - -Lemma sub_botP x : sub_leq sub_bot x. -Proof. by apply: botP. Qed. - -Lemma sub_refl x : sub_leq x x. -Proof. by rewrite /sub_leq. Qed. - -Lemma sub_asym x y : sub_leq x y -> sub_leq y x -> x = y. -Proof. -move: x y=>[x Hx][y Hy]; rewrite /sub_leq /= => H1 H2. -move: (poset_asym H1 H2) Hy=> <- Hy; congr exist. -by apply: pf_irr. -Qed. - -Lemma sub_trans x y z : sub_leq x y -> sub_leq y z -> sub_leq x z. -Proof. -move: x y z=>[x Hx][y Hy][z Hz]; rewrite /sub_leq /=. -by apply: poset_trans. -Qed. - -(* no need to put canonical here, because the system won't be *) -(* able to get the proof C from the {x : T | x \In s} syntax *) -Definition subPosetMixin := PosetMixin sub_botP sub_refl sub_asym sub_trans. -Definition subPoset := Eval hnf in Poset tp subPosetMixin. - -End SubPoset. - -(* pairing of posets *) - -Section PairPoset. -Variable (A B : poset). -Local Notation tp := (A * B)%type. - -Definition pair_bot : tp := (bot, bot). -Definition pair_leq (p1 p2 : tp) := p1.1 <== p2.1 /\ p1.2 <== p2.2. - -Lemma pair_botP x : pair_leq pair_bot x. -Proof. by split; apply: botP. Qed. - -Lemma pair_refl x : pair_leq x x. -Proof. by []. Qed. - -Lemma pair_asym x y : pair_leq x y -> pair_leq y x -> x = y. -Proof. -move: x y=>[x1 x2][y1 y2][/= H1 H2][/= H3 H4]. -by congr (_, _); apply: poset_asym. -Qed. - -Lemma pair_trans x y z : pair_leq x y -> pair_leq y z -> pair_leq x z. -Proof. -move: x y z=>[x1 x2][y1 y2][z1 z2][/= H1 H2][/= H3 H4]; split=>/=. -- by apply: poset_trans H3. -by apply: poset_trans H4. -Qed. - -Definition pairPosetMixin := - PosetMixin pair_botP pair_refl pair_asym pair_trans. -Canonical pairPoset := Eval hnf in Poset tp pairPosetMixin. - -End PairPoset. - -(* functions into a poset form a poset *) - -Section FunPoset. -Variable (A : Type) (B : poset). -Local Notation tp := (A -> B). - -Definition fun_bot : tp := fun x => bot. -Definition fun_leq (p1 p2 : tp) := forall x, p1 x <== p2 x. - -Lemma fun_botP x : fun_leq fun_bot x. -Proof. by move=>y; apply: botP. Qed. - -Lemma fun_refl x : fun_leq x x. -Proof. by []. Qed. - -Lemma fun_asym x y : fun_leq x y -> fun_leq y x -> x = y. -Proof. -move=>H1 H2; apply: fext=>z; -by apply: poset_asym; [apply: H1 | apply: H2]. -Qed. - -Lemma fun_trans x y z : fun_leq x y -> fun_leq y z -> fun_leq x z. -Proof. by move=>H1 H2 t; apply: poset_trans (H2 t). Qed. - -Definition funPosetMixin := PosetMixin fun_botP fun_refl fun_asym fun_trans. -Canonical funPoset := Eval hnf in Poset tp funPosetMixin. - -End FunPoset. - -(* dependent functions into a poset form a poset *) - -Section DFunPoset. -Variables (A : Type) (B : A -> poset). -Local Notation tp := (forall x, B x). - -Definition dfun_bot : tp := fun x => bot. -Definition dfun_leq (p1 p2 : tp) := forall x, p1 x <== p2 x. - -Lemma dfun_botP x : dfun_leq dfun_bot x. -Proof. by move=>y; apply: botP. Qed. - -Lemma dfun_refl x : dfun_leq x x. -Proof. by []. Qed. - -Lemma dfun_asym x y : dfun_leq x y -> dfun_leq y x -> x = y. -Proof. -move=>H1 H2; apply: fext=>z; -by apply: poset_asym; [apply: H1 | apply: H2]. -Qed. - -Lemma dfun_trans x y z : dfun_leq x y -> dfun_leq y z -> dfun_leq x z. -Proof. by move=>H1 H2 t; apply: poset_trans (H2 t). Qed. - -(* no point in declaring this canonical, since it's keyed on forall *) -(* and forall is not a symbol *) -Definition dfunPosetMixin := - PosetMixin dfun_botP dfun_refl dfun_asym dfun_trans. -Definition dfunPoset := Eval hnf in Poset tp dfunPosetMixin. - -End DFunPoset. - -(* ideal of a poset is a poset *) - -Section IdealPoset. -Variable (T : poset) (P : T). - -Definition ideal_bot := Ideal (botP P). -Definition ideal_leq (p1 p2 : ideal P) := id_val p1 <== id_val p2. - -Lemma ideal_botP x : ideal_leq ideal_bot x. -Proof. by apply: botP. Qed. - -Lemma ideal_refl x : ideal_leq x x. -Proof. by case: x=>x H; rewrite /ideal_leq. Qed. - -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: pf_irr. -Qed. - -Lemma ideal_trans x y z : ideal_leq x y -> ideal_leq y z -> ideal_leq x z. -Proof. by apply: poset_trans. Qed. - -Definition idealPosetMixin := - PosetMixin ideal_botP ideal_refl ideal_asym ideal_trans. -Canonical idealPoset := Eval hnf in Poset (ideal P) idealPosetMixin. - -End IdealPoset. - -(* Prop is a poset (Sierpinski poset) *) - -Section PropPoset. - -Definition prop_bot := False. -Definition prop_leq (p1 p2 : Prop) := p1 -> p2. - -Lemma prop_botP x : prop_leq prop_bot x. -Proof. by []. Qed. - -Lemma prop_refl x : prop_leq x x. -Proof. by []. Qed. - -Lemma prop_asym x y : prop_leq x y -> prop_leq y x -> x = y. -Proof. by move=>H1 H2; apply: pext. Qed. - -Lemma prop_trans x y z : prop_leq x y -> prop_leq y z -> prop_leq x z. -Proof. by move=>H1 H2; move/H1. Qed. - -Definition propPosetMixin := - PosetMixin prop_botP prop_refl prop_asym prop_trans. -Canonical propPoset := Eval hnf in Poset Prop propPosetMixin. - -Lemma prop_topP x : x <== True. -Proof. by []. Qed. - -End PropPoset. - -(* Pred is a poset *) - -(* Can be inherited from posets of -> and Prop, but we declare a *) -(* dedicated structure to keep the infix notation of Pred. Otherwise, *) -(* poset inference turns Pred A into A -> Prop. *) - -Section PredPoset. -Variable A : Type. - -Definition predPosetMixin : Poset.mixin_of (Pred A) := - funPosetMixin A propPoset. -Canonical predPoset := Eval hnf in Poset (Pred A) predPosetMixin. - -Lemma pred_topP (x : Pred A) : x <== PredT. -Proof. by []. Qed. - -End PredPoset. - -(* nat is a poset *) -Section NatPoset. - -Lemma nat_botP x : 0 <= x. Proof. by []. Qed. -Lemma nat_refl x : x <= x. Proof. by []. Qed. - -Lemma nat_asym x y : x <= y -> y <= x -> x = y. -Proof. by move=>H1 H2; apply: anti_leq; rewrite H1 H2. Qed. - -Lemma nat_trans x y z : x <= y -> y <= z -> x <= z. -Proof. by apply: leq_trans. Qed. - -Definition natPosetMixin := PosetMixin nat_botP nat_refl nat_asym nat_trans. -Canonical natPoset := Eval hnf in Poset nat natPosetMixin. - -End NatPoset. - - -(*********************) -(* Complete lattices *) -(*********************) - -Module Lattice. - -Section RawMixin. - -Variable T : poset. - -Record mixin_of := Mixin { - mx_sup : Pred T -> T; - _ : forall (s : Pred T) x, x \In s -> x <== mx_sup s; - _ : forall (s : Pred T) x, - (forall y, y \In s -> y <== x) -> mx_sup s <== x}. - -End RawMixin. - -Section ClassDef. - -Record class_of (T : Type) := Class { - base : Poset.class_of T; - mixin : mixin_of (Poset.Pack base)}. - -Local Coercion base : class_of >-> Poset.class_of. - -Structure type : Type := Pack {sort : Type; _ : class_of sort}. -Local Coercion sort : type >-> Sortclass. - -Variables (T : Type) (cT : type). -Definition class := let: Pack _ c as cT' := cT return class_of cT' in c. -Definition clone c of phant_id class c := @Pack T c. - -(* produce a lattice type out of the mixin *) -(* equalize m0 and m by means of a phantom *) -Definition pack b0 (m0 : mixin_of (Poset.Pack b0)) := - fun m & phant_id m0 m => Pack (@Class T b0 m). - -Definition sup (s : Pred cT) : cT := mx_sup (mixin class) s. - -Definition poset := Poset.Pack class. - -End ClassDef. - -Module Exports. -Coercion base : class_of >-> Poset.class_of. -Coercion sort : type >-> Sortclass. -Coercion poset : type >-> Poset.type. -Canonical Structure poset. - -Notation lattice := Lattice.type. -Notation LatticeMixin := Lattice.Mixin. -Notation Lattice T m := (@pack T _ _ m id). - -Notation "[ 'lattice' 'of' T 'for' cT ]" := (@clone T cT _ id) - (at level 0, format "[ 'lattice' 'of' T 'for' cT ]") : form_scope. -Notation "[ 'lattice' 'of' T ]" := (@clone T _ _ id) - (at level 0, format "[ 'lattice' 'of' T ]") : form_scope. - -Arguments Lattice.sup [cT] s. -Prenex Implicits Lattice.sup. -Notation sup := Lattice.sup. - -(* re-state lattice properties using the exported notation *) -Section Laws. -Variable T : lattice. - -Lemma supP (s : Pred T) x : x \In s -> x <== sup s. -Proof. by case: T s x=>S [[p]][/= s H1 *]; apply: H1. Qed. - -Lemma supM (s : Pred T) x : (forall y, y \In s -> y <== x) -> sup s <== x. -Proof. by case: T s x=>S [[p]][/= s H1 H2 *]; apply: (H2). Qed. - -End Laws. - -End Exports. - -End Lattice. - -Export Lattice.Exports. - -(* we have greatest lower bounds too *) -Section Infimum. -Variable T : lattice. - -Definition inf (s : Pred T) := - sup [Pred x : T | forall y, y \In s -> x <== y]. - -Lemma infP s : forall x, x \In s -> inf s <== x. -Proof. by move=>x H; apply: supM=>y; apply. Qed. - -Lemma infM s : forall x, (forall y, y \In s -> x <== y) -> x <== inf s. -Proof. by apply: supP. Qed. - -End Infimum. - -(* we can compute least and greatest fixed points *) - -Section Lat. -Variable T : lattice. - -Definition tarski_lfp (f : T -> T) := inf [Pred x : T | f x <== x]. -Definition tarski_gfp (f : T -> T) := sup [Pred x : T | x <== f x]. - -Definition sup_closed (T : lattice) := - [Pred s : Pred T | - bot \In s /\ forall d, d <=p s -> sup d \In s]. - -Definition sup_closure (T : lattice) (s : Pred T) := - [Pred p : T | forall t : Pred T, s <=p t -> t \In sup_closed T -> p \In t]. - -End Lat. - -Arguments sup_closed {T}. -Arguments sup_closure [T]. -Prenex Implicits sup_closed sup_closure. - -Section BasicProperties. -Variable T : lattice. - -Lemma sup_mono (s1 s2 : Pred T) : s1 <=p s2 -> sup s1 <== sup s2. -Proof. by move=>H; apply: supM=>y; move/H; apply: supP. Qed. - -Lemma supE (s1 s2 : Pred T) : s1 =p s2 -> sup s1 = sup s2. -Proof. by move=>H; apply: poset_asym; apply: supM=>y; move/H; apply: supP. Qed. - -(* Knaster-Tarski *) -Lemma tarski_lfp_fixed (f : T -> T) : - monotone f -> f (tarski_lfp f) = tarski_lfp f. -Proof. -move=>M; suff L: f (tarski_lfp f) <== tarski_lfp f. -- by apply: poset_asym=>//; apply: infP; apply: M L. -by apply: infM=>x L; apply: poset_trans (L); apply: M; apply: infP. -Qed. - -Lemma tarski_lfp_least f : forall x : T, f x = x -> tarski_lfp f <== x. -Proof. by move=>x H; apply: infP; rewrite InE /= H. Qed. - -Lemma tarski_gfp_fixed (f : T -> T) : - monotone f -> f (tarski_gfp f) = tarski_gfp f. -Proof. -move=>M; suff L: tarski_gfp f <== f (tarski_gfp f). -- by apply: poset_asym=>//; apply: supP; apply: M L. -by apply: supM=>x L; apply: poset_trans (L) _; apply: M; apply: supP. -Qed. - -Lemma tarski_gfp_greatest f : forall x : T, f x = x -> x <== tarski_gfp f. -Proof. by move=>x H; apply: supP; rewrite InE /= H. Qed. - -Lemma tarski_lfp_mono (f1 f2 : T -> T) : - monotone f2 -> f1 <== f2 -> tarski_lfp f1 <== tarski_lfp f2. -Proof. -move=>M H; apply/infP/(poset_trans (H (tarski_lfp f2))). -by rewrite tarski_lfp_fixed. -Qed. - -Lemma tarski_gfp_mono (f1 f2 : T -> T) : - monotone f1 -> f1 <== f2 -> tarski_gfp f1 <== tarski_gfp f2. -Proof. -move=>M H; apply/supP/poset_trans/(H (tarski_gfp f1)). -by rewrite tarski_gfp_fixed. -Qed. - -(* closure contains s *) -Lemma sup_clos_sub (s : Pred T) : s <=p sup_closure s. -Proof. by move=>p H1 t H2 H3; apply: H2 H1. Qed. - -(* closure is smallest *) -Lemma sup_clos_min (s : Pred T) : - forall t, s <=p t -> sup_closed t -> sup_closure s <=p t. -Proof. by move=>t H1 H2 p; move/(_ _ H1 H2). Qed. - -(* closure is closed *) -Lemma sup_closP (s : Pred T) : sup_closed (sup_closure s). -Proof. -split; first by move=>t _ []. -move=>d H1 t H3 H4; move: (sup_clos_min H3 H4)=>{H3} -H3. -by case: H4=>_; apply=>// x; move/H1; move/H3. -Qed. - -Lemma sup_clos_idemp (s : Pred T) : sup_closed s -> sup_closure s =p s. -Proof. by move=>p; split; [apply: sup_clos_min | apply: sup_clos_sub]. Qed. - -Lemma sup_clos_mono (s1 s2 : Pred T) : - s1 <=p s2 -> sup_closure s1 <=p sup_closure s2. -Proof. -move=>H1; apply: sup_clos_min (sup_closP s2)=>p H2. -by apply: sup_clos_sub; apply: H1. -Qed. - -End BasicProperties. - -(* lattice constructions *) - -Section SubLattice. -Variables (T : lattice) (s : Pred T) (C : sup_closed s). -Local Notation tp := (subPoset (proj1 C)). - -Definition sub_sup' (u : Pred tp) : T := - sup [Pred x : T | exists y, y \In u /\ x = sval y]. - -Lemma sub_supX (u : Pred tp) : sub_sup' u \In s. -Proof. by case: C u=>P /= H u; apply: H=>t [[y]] H1 [_] ->. Qed. - -Definition sub_sup (u : Pred tp) : tp := - exist _ (sub_sup' u) (sub_supX u). - -Lemma sub_supP (u : Pred tp) (x : tp) : x \In u -> x <== sub_sup u. -Proof. by move=>H; apply: supP; exists x. Qed. - -Lemma sub_supM (u : Pred tp) (x : tp) : - (forall y, y \In u -> y <== x) -> sub_sup u <== x. -Proof. by move=>H; apply: supM=>y [z][H1] ->; apply: H H1. Qed. - -Definition subLatticeMixin := LatticeMixin sub_supP sub_supM. -Definition subLattice := Eval hnf in Lattice {x : T | x \In s} subLatticeMixin. - -End SubLattice. - -(* pairing *) - -Section PairLattice. -Variables (A B : lattice). -Local Notation tp := (A * B)%type. - -Definition pair_sup (s : Pred tp) : tp := - (sup [Pred p | exists f, p = f.1 /\ f \In s], - sup [Pred p | exists f, p = f.2 /\ f \In s]). - -Lemma pair_supP (s : Pred tp) (p : tp) : p \In s -> p <== pair_sup s. -Proof. by move=>H; split; apply: supP; exists p. Qed. - -Lemma pair_supM (s : Pred tp) (p : tp) : - (forall q, q \In s -> q <== p) -> pair_sup s <== p. -Proof. by move=>H; split; apply: supM=>y [f][->]; case/H. Qed. - -Definition pairLatticeMixin := LatticeMixin pair_supP pair_supM. -Canonical pairLattice := Eval hnf in Lattice tp pairLatticeMixin. - -End PairLattice. - -(* functions into a latice form a lattice *) - -Section FunLattice. -Variables (A : Type) (B : lattice). -Local Notation tp := (A -> B). - -Definition fun_sup (s : Pred tp) : tp := - fun x => sup [Pred p | exists f, f \In s /\ p = f x]. - -Lemma fun_supP (s : Pred tp) (p : tp) : p \In s -> p <== fun_sup s. -Proof. by move=>H x; apply: supP; exists p. Qed. - -Lemma fun_supM (s : Pred tp) (p : tp) : - (forall q, q \In s -> q <== p) -> fun_sup s <== p. -Proof. by move=>H t; apply: supM=>x [f][H1] ->; apply: H. Qed. - -Definition funLatticeMixin := LatticeMixin fun_supP fun_supM. -Canonical funLattice := Eval hnf in Lattice tp funLatticeMixin. - -End FunLattice. - -(* dependent functions into a lattice form a lattice *) - -Section DFunLattice. -Variables (A : Type) (B : A -> lattice). -Local Notation tp := (dfunPoset B). - -Definition dfun_sup (s : Pred tp) : tp := - fun x => sup [Pred p | exists f, f \In s /\ p = f x]. - -Lemma dfun_supP (s : Pred tp) (p : tp) : - p \In s -> p <== dfun_sup s. -Proof. by move=>H x; apply: supP; exists p. Qed. - -Lemma dfun_supM (s : Pred tp) (p : tp) : - (forall q, q \In s -> q <== p) -> dfun_sup s <== p. -Proof. by move=>H t; apply: supM=>x [f][H1] ->; apply: H. Qed. - -Definition dfunLatticeMixin := LatticeMixin dfun_supP dfun_supM. -Definition dfunLattice := Eval hnf in Lattice (forall x, B x) dfunLatticeMixin. - -End DFunLattice. - -(* applied sup equals the sup of applications *) - -Lemma sup_appE A (B : lattice) (s : Pred (A -> B)) (x : A) : - sup s x = sup [Pred y : B | exists f, f \In s /\ y = f x]. -Proof. by []. Qed. - -Lemma sup_dappE A (B : A -> lattice) (s : Pred (dfunLattice B)) (x : A) : - sup s x = sup [Pred y : B x | exists f, f \In s /\ y = f x]. -Proof. by []. Qed. - -(* ideal of a lattice forms a lattice *) - -Section IdealLattice. -Variables (T : lattice) (P : T). - -Definition ideal_sup' (s : Pred (ideal P)) := - sup [Pred x | exists p, p \In s /\ id_val p = x]. - -Lemma ideal_supP' (s : Pred (ideal P)) : ideal_sup' s <== P. -Proof. by apply: supM=>y [[x]] H /= [_] <-. Qed. - -Definition ideal_sup (s : Pred (ideal P)) := Ideal (ideal_supP' s). - -Lemma ideal_supP (s : Pred (ideal P)) p : - p \In s -> p <== ideal_sup s. -Proof. by move=>H; apply: supP; exists p. Qed. - -Lemma ideal_supM (s : Pred (ideal P)) p : - (forall q, q \In s -> q <== p) -> ideal_sup s <== p. -Proof. by move=>H; apply: supM=>y [q][H1] <-; apply: H. Qed. - -Definition idealLatticeMixin := LatticeMixin ideal_supP ideal_supM. -Canonical idealLattice := Eval hnf in Lattice (ideal P) idealLatticeMixin. - -End IdealLattice. - -(* Prop is a lattice *) - -Section PropLattice. - -Definition prop_sup (s : Pred Prop) : Prop := exists p, p \In s /\ p. - -Lemma prop_supP (s : Pred Prop) p : p \In s -> p <== prop_sup s. -Proof. by exists p. Qed. - -Lemma prop_supM (s : Pred Prop) p : - (forall q, q \In s -> q <== p) -> prop_sup s <== p. -Proof. by move=>H [r][]; move/H. Qed. - -Definition propLatticeMixin := LatticeMixin prop_supP prop_supM. -Canonical propLattice := Eval hnf in Lattice Prop propLatticeMixin. - -End PropLattice. - -(* Pred is a lattice *) - -Section PredLattice. -Variable A : Type. - -Definition predLatticeMixin := funLatticeMixin A propLattice. -Canonical predLattice := Eval hnf in Lattice (Pred A) predLatticeMixin. - -End PredLattice. - -(**********) -(* Chains *) -(**********) - -Section Chains. -Variable T : poset. - -Definition chain_axiom (s : Pred T) := - (exists d, d \In s) /\ - (forall x y, x \In s -> y \In s -> x <== y \/ y <== x). - -Structure chain := Chain { - pred_of :> Pred T; - _ : chain_axiom pred_of}. - -Canonical chainPredType := @mkPredType T chain pred_of. - -End Chains. - -Lemma chainE (T : poset) (s1 s2 : chain T) : - s1 = s2 <-> pred_of s1 =p pred_of s2. -Proof. -split=>[->//|]; move: s1 s2=>[s1 H1][s2 H2] /= E; move: H1 H2. -suff: s1 = s2 by move=>-> H1 H2; congr Chain; apply: pf_irr. -by apply: fext=>x; apply: pext; split; move/E. -Qed. - -(* common chain constructions *) - -(* adding bot to a chain *) - -Section LiftChain. -Variable (T : poset) (s : chain T). - -Lemma lift_chainP : chain_axiom [Pred x : T | x = bot \/ x \In s]. -Proof. -case: s=>p [[d H1] H2] /=; split=>[|x y]; first by exists d; right. -by case=>[->|H3][->|H4]; auto. -Qed. - -Definition lift_chain := Chain lift_chainP. - -End LiftChain. - -(* mapping monotone function over a chain *) - -Section ImageChain. -Variables (T1 T2 : poset) (s : chain T1) (f : T1 -> T2) (M : monotone f). - -Lemma image_chainP : - chain_axiom [Pred x2 : T2 | exists x1, x2 = f x1 /\ x1 \In s]. -Proof. -case: s=>p [[d H1] H2]; split=>[|x y]; first by exists (f d); exists d. -case=>x1 [->] H3 [y1][->] H4; rewrite -!toPredE /= in H3 H4. -by case: (H2 x1 y1 H3 H4)=>L; [left | right]; apply: M L. -Qed. - -Definition image_chain := Chain image_chainP. - -End ImageChain. - -Notation "[ f '^^' s 'by' M ]" := (@image_chain _ _ s f M) - (at level 0, format "[ f '^^' s 'by' M ]") : form_scope. - -Section ChainId. -Variables (T : poset) (s : chain T). - -Lemma id_mono : monotone (@id T). -Proof. by []. Qed. - -Lemma id_chainE (M : monotone id) : [id ^^ s by M] = s. -Proof. by apply/chainE=>x; split; [case=>y [<-]|exists x]. Qed. - -End ChainId. - -Arguments id_mono [T]. -Prenex Implicits id_mono. - -Section ChainConst. -Variables (T1 T2 : poset) (y : T2). - -Lemma const_mono : monotone (fun x : T1 => y). -Proof. by []. Qed. - -Lemma const_chainP : chain_axiom (Pred1 y). -Proof. by split; [exists y | move=>x1 x2 ->->; left]. Qed. - -Definition const_chain := Chain const_chainP. - -Lemma const_chainE s : [_ ^^ s by const_mono] = const_chain. -Proof. -apply/chainE=>z1; split; first by case=>z2 [->]. -by case: s=>s [[d] H1] H2 <-; exists d. -Qed. - -End ChainConst. - -Arguments const_mono [T1 T2 y]. -Prenex Implicits const_mono. - -Section ChainCompose. -Variables (T1 T2 T3 : poset) (f1 : T2 -> T1) (f2 : T3 -> T2). -Variables (s : chain T3) (M1 : monotone f1) (M2 : monotone f2). - -Lemma comp_mono : monotone (f1 \o f2). -Proof. by move=>x y H; apply: M1; apply: M2. Qed. - -Lemma comp_chainE : - [f1 ^^ [f2 ^^ s by M2] by M1] = [f1 \o f2 ^^ s by comp_mono]. -Proof. -apply/chainE=>x1; split; first by case=>x2 [->][x3][->]; exists x3. -by case=>x3 [->]; exists (f2 x3); split=>//; exists x3. -Qed. - -End ChainCompose. - -Arguments comp_mono [T1 T2 T3 f1 f2]. -Prenex Implicits comp_mono. - -(* projections out of a chain *) - -Section ProjChain. -Variables (T1 T2 : poset) (s : chain [poset of T1 * T2]). - -Lemma proj1_mono : monotone (@fst T1 T2). -Proof. by case=>x1 x2 [y1 y2][]. Qed. - -Lemma proj2_mono : monotone (@snd T1 T2). -Proof. by case=>x1 x2 [y1 y2][]. Qed. - -Definition proj1_chain := [@fst _ _ ^^ s by proj1_mono]. -Definition proj2_chain := [@snd _ _ ^^ s by proj2_mono]. - -End ProjChain. - -Arguments proj1_mono [T1 T2]. -Arguments proj2_mono [T1 T2]. -Prenex Implicits proj1_mono proj2_mono. - -(* diagonal chain *) - -Section DiagChain. -Variable (T : poset) (s : chain T). - -Lemma diag_mono : monotone (fun x : T => (x, x)). -Proof. by []. Qed. - -Definition diag_chain := [_ ^^ s by diag_mono]. - -Lemma proj1_diagE : proj1_chain diag_chain = s. -Proof. by rewrite /proj1_chain /diag_chain comp_chainE id_chainE. Qed. - -Lemma proj2_diagE : proj2_chain diag_chain = s. -Proof. by rewrite /proj2_chain /diag_chain comp_chainE id_chainE. Qed. - -End DiagChain. - -Arguments diag_mono [T]. -Prenex Implicits diag_mono. - -(* applying functions from a chain of functions *) - -Section AppChain. -Variables (A : Type) (T : poset) (s : chain [poset of A -> T]). - -Lemma app_mono x : monotone (fun f : A -> T => f x). -Proof. by move=>f1 f2; apply. Qed. - -Definition app_chain x := [_ ^^ s by app_mono x]. - -End AppChain. - -Arguments app_mono [A T]. -Prenex Implicits app_mono. - -(* ditto for dependent functions *) - -Section DAppChain. -Variables (A : Type) (T : A -> poset) (s : chain (dfunPoset T)). - -Lemma dapp_mono x : monotone (fun f : dfunPoset T => f x). -Proof. by move=>f1 f2; apply. Qed. - -Definition dapp_chain x := [_ ^^ s by dapp_mono x]. - -End DAppChain. - -Arguments dapp_mono [A T]. -Prenex Implicits dapp_mono. - -(* pairing chain applications *) - -Section ProdChain. -Variables (S1 S2 T1 T2 : poset) (f1 : S1 -> T1) (f2 : S2 -> T2). -Variables (M1 : monotone f1) (M2 : monotone f2). -Variable (s : chain [poset of S1 * S2]). - -Lemma prod_mono : monotone (f1 \* f2). -Proof. by case=>x1 x2 [y1 y2][/= H1 H2]; split; [apply: M1 | apply: M2]. Qed. - -Definition prod_chain := [f1 \* f2 ^^ s by prod_mono]. - -Lemma proj1_prodE : proj1_chain prod_chain = [f1 ^^ proj1_chain s by M1]. -Proof. -rewrite /proj1_chain /prod_chain !comp_chainE !/comp /=. -by apply/chainE. -Qed. - -Lemma proj2_prodE : proj2_chain prod_chain = [f2 ^^ proj2_chain s by M2]. -Proof. -rewrite /proj2_chain /prod_chain !comp_chainE !/comp /=. -by apply/chainE. -Qed. - -End ProdChain. - -Arguments prod_mono [S1 S2 T1 T2 f1 f2]. -Prenex Implicits prod_mono. - -(* chain of all nats *) - -Section NatChain. - -Lemma nat_chain_axiom : chain_axiom (@PredT nat). -Proof. -split=>[|x y _ _]; first by exists 0. -rewrite /Poset.leq /= [y <= x]leq_eqVlt. -by case: leqP; [left | rewrite orbT; right]. -Qed. - -Definition nat_chain := Chain nat_chain_axiom. - -End NatChain. - -(*********) -(* CPO's *) -(*********) - -Module CPO. - -Section RawMixin. - -Record mixin_of (T : poset) := Mixin { - mx_lim : chain T -> T; - _ : forall (s : chain T) x, x \In s -> x <== mx_lim s; - _ : forall (s : chain T) x, - (forall y, y \In s -> y <== x) -> mx_lim s <== x}. - -End RawMixin. - -Section ClassDef. - -Record class_of (T : Type) := Class { - base : Poset.class_of T; - mixin : mixin_of (Poset.Pack base)}. - -Local Coercion base : class_of >-> Poset.class_of. - -Structure type : Type := Pack {sort; _ : 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 b0 (m0 : mixin_of (Poset.Pack b0)) := - fun m & phant_id m0 m => Pack (@Class T b0 m) T. - -Definition poset := Poset.Pack class. -Definition lim (s : chain poset) : cT := mx_lim (mixin class) s. - -End ClassDef. - -Module Import Exports. -Coercion base : class_of >-> Poset.class_of. -Coercion sort : type >-> Sortclass. -Coercion poset : type >-> Poset.type. -Canonical Structure poset. - -Notation cpo := type. -Notation CPOMixin := Mixin. -Notation CPO T m := (@pack T _ _ m id). - -Notation "[ 'cpo' 'of' T 'for' cT ]" := (@clone T cT _ idfun) - (at level 0, format "[ 'cpo' 'of' T 'for' cT ]") : form_scope. -Notation "[ 'cpo' 'of' T ]" := (@clone T _ _ id) - (at level 0, format "[ 'cpo' 'of' T ]") : form_scope. - -Arguments CPO.lim [cT]. -Prenex Implicits CPO.lim. -Notation lim := CPO.lim. - -Section Laws. -Variable D : cpo. - -Lemma limP (s : chain D) x : x \In s -> x <== lim s. -Proof. by case: D s x=>S [[p]][/= l H1 *]; apply: (H1). Qed. - -Lemma limM (s : chain D) x : (forall y, y \In s -> y <== x) -> lim s <== x. -Proof. by case: D s x=>S [[p]][/= l H1 H2 *]; apply: (H2). Qed. - -End Laws. - -End Exports. - -End CPO. - -Export CPO.Exports. - -(****************************) -(* common cpo constructions *) -(****************************) - -(* pairs *) - -Section PairCPO. -Variables (A B : cpo). -Local Notation tp := [poset of A * B]. - -Definition pair_lim (s : chain tp) : tp := - (lim (proj1_chain s), lim (proj2_chain s)). - -Lemma pair_limP (s : chain tp) x : x \In s -> x <== pair_lim s. -Proof. by split; apply: limP; exists x. Qed. - -Lemma pair_limM (s : chain tp) x : - (forall y, y \In s -> y <== x) -> pair_lim s <== x. -Proof. by split; apply: limM=>y [z][->]; case/H. Qed. - -Definition pairCPOMixin := CPOMixin pair_limP pair_limM. -Canonical pairCPO := Eval hnf in CPO (A * B) pairCPOMixin. - -End PairCPO. - -(* functions *) - -Section FunCPO. -Variable (A : Type) (B : cpo). -Local Notation tp := [poset of A -> B]. - -Definition fun_lim (s : chain tp) : tp := - fun x => lim (app_chain s x). - -Lemma fun_limP (s : chain tp) x : x \In s -> x <== fun_lim s. -Proof. by move=>H t; apply: limP; exists x. Qed. - -Lemma fun_limM (s : chain tp) x : - (forall y, y \In s -> y <== x) -> fun_lim s <== x. -Proof. by move=>H1 t; apply: limM=>y [f][->] H2; apply: H1. Qed. - -Definition funCPOMixin := CPOMixin fun_limP fun_limM. -Canonical funCPO := Eval hnf in CPO (A -> B) funCPOMixin. - -End FunCPO. - -(* dependent functions *) - -Section DFunCPO. -Variable (A : Type) (B : A -> cpo). -Local Notation tp := (dfunPoset B). - -Definition dfun_lim (s : chain tp) : tp := - fun x => lim (dapp_chain s x). - -Lemma dfun_limP (s : chain tp) x : x \In s -> x <== dfun_lim s. -Proof. by move=>H t; apply: limP; exists x. Qed. - -Lemma dfun_limM (s : chain tp) x : - (forall y, y \In s -> y <== x) -> dfun_lim s <== x. -Proof. by move=>H1 t; apply: limM=>y [f][->] H2; apply: H1. Qed. - -Definition dfunCPOMixin := CPOMixin dfun_limP dfun_limM. -Definition dfunCPO := Eval hnf in CPO (forall x, B x) dfunCPOMixin. - -End DFunCPO. - -(* Prop *) - -Section PropCPO. -Local Notation tp := [poset of Prop]. - -Definition prop_lim (s : chain tp) : tp := exists p, p \In s /\ p. - -Lemma prop_limP (s : chain tp) p : p \In s -> p <== prop_lim s. -Proof. by exists p. Qed. - -Lemma prop_limM (s : chain tp) p : - (forall q, q \In s -> q <== p) -> prop_lim s <== p. -Proof. by move=>H [r][]; move/H. Qed. - -Definition propCPOMixin := CPOMixin prop_limP prop_limM. -Canonical propCPO := Eval hnf in CPO Prop propCPOMixin. - -End PropCPO. - -(* Pred *) - -Section PredCPO. -Variable A : Type. - -Definition predCPOMixin := funCPOMixin A propCPO. -Canonical predCPO := Eval hnf in CPO (Pred A) predCPOMixin. - -End PredCPO. - -(* every complete lattice is a cpo *) - -Section LatticeCPO. -Variable A : lattice. -Local Notation tp := (Lattice.poset A). - -Definition lat_lim (s : chain tp) : tp := sup s. - -Lemma lat_limP (s : chain tp) x : x \In s -> x <== lat_lim s. -Proof. by apply: supP. Qed. - -Lemma lat_limM (s : chain tp) x : - (forall y, y \In s -> y <== x) -> lat_lim s <== x. -Proof. by apply: supM. Qed. - -Definition latCPOMixin := CPOMixin lat_limP lat_limM. -Definition latCPO := Eval hnf in CPO tp latCPOMixin. - -End LatticeCPO. - -(* sub-CPO's *) - -(* every chain-closed subset of a cpo is a cpo *) - -Section AdmissibleClosure. -Variable T : cpo. - -Definition chain_closed := - [Pred s : Pred T | - bot \In s /\ forall d : chain T, d <=p s -> lim d \In s]. - -(* admissible closure of s is the smallest closed set containing s *) -(* basically extends s to include the sups of chains *) -Definition chain_closure (s : Pred T) := - [Pred p : T | forall t : Pred T, s <=p t -> chain_closed t -> p \In t]. - -(* admissible closure contains s *) -Lemma chain_clos_sub (s : Pred T) : s <=p chain_closure s. -Proof. by move=>p H1 t H2 H3; apply: H2 H1. Qed. - -(* admissible closure is smallest *) -Lemma chain_clos_min (s : Pred T) t : - s <=p t -> chain_closed t -> chain_closure s <=p t. -Proof. by move=>H1 H2 p; move/(_ _ H1 H2). Qed. - -(* admissible closure is closed *) -Lemma chain_closP (s : Pred T) : chain_closed (chain_closure s). -Proof. -split; first by move=>t _ []. -move=>d H1 t H3 H4; move: (chain_clos_min H3 H4)=>{H3} -H3. -by case: H4=>_; apply=>// x; move/H1; move/H3. -Qed. - -Lemma chain_clos_idemp (s : Pred T) : - chain_closed s -> chain_closure s =p s. -Proof. -move=>p; split; last by apply: chain_clos_sub. -by apply: chain_clos_min=>//; apply: chain_closP. -Qed. - -Lemma chain_clos_mono (s1 s2 : Pred T) : - s1 <=p s2 -> chain_closure s1 <=p chain_closure s2. -Proof. -move=>H1; apply: chain_clos_min (chain_closP s2)=>p H2. -by apply: chain_clos_sub; apply: H1. -Qed. - -(* intersection of admissible sets is admissible *) -Lemma chain_closI (s1 s2 : Pred T) : - chain_closed s1 -> chain_closed s2 -> chain_closed (PredI s1 s2). -Proof. -move=>[H1 S1][H2 S2]; split=>// d H. -by split; [apply: S1 | apply: S2]=>// x; case/H. -Qed. - -End AdmissibleClosure. - -Arguments chain_closed {T}. -Prenex Implicits chain_closed. - -(* diagonal of an admissible set of pairs is admissible *) -Lemma chain_clos_diag (T : cpo) (s : Pred (T * T)) : - chain_closed s -> chain_closed [Pred t : T | (t, t) \In s]. -Proof. -move=>[B H1]; split=>// d H2. -rewrite InE /= -{1}(proj1_diagE d) -{2}(proj2_diagE d). -by apply: H1; case=>x1 x2 [x][[<- <-]]; apply: H2. -Qed. - -Section SubCPO. -Variables (D : cpo) (s : Pred D) (C : chain_closed s). - -Local Notation tp := (subPoset (proj1 C)). - -Lemma sval_mono : monotone (sval : tp -> D). -Proof. by move=>[x1 H1][x2 H2]; apply. Qed. - -Lemma sub_limX (u : chain tp) : lim [sval ^^ u by sval_mono] \In s. -Proof. by case: C u=>P H u; apply: (H)=>t [[y]] H1 [->]. Qed. - -Definition sub_lim (u : chain tp) : tp := - exist _ (lim [sval ^^ u by sval_mono]) (sub_limX u). - -Lemma sub_limP (u : chain tp) x : x \In u -> x <== sub_lim u. -Proof. by move=>H; apply: limP; exists x. Qed. - -Lemma sub_limM (u : chain tp) x : - (forall y, y \In u -> y <== x) -> sub_lim u <== x. -Proof. by move=>H; apply: limM=>y [z][->]; apply: H. Qed. - -Definition subCPOMixin := CPOMixin sub_limP sub_limM. -Definition subCPO := Eval hnf in CPO {x : D | x \In s} subCPOMixin. - -End SubCPO. - -(***********************************************) -(* Continuity and Kleene's fixed point theorem *) -(***********************************************) - -Lemma lim_mono (D : cpo) (s1 s2 : chain D) : - s1 <=p s2 -> lim s1 <== lim s2. -Proof. by move=>H; apply: limM=>y; move/H; apply: limP. Qed. - -Lemma limE (D : cpo) (s1 s2 : chain D) : - s1 =p s2 -> lim s1 = lim s2. -Proof. by move=>H; apply: poset_asym; apply: lim_mono=>x; rewrite H. Qed. - -Lemma lim_liftE (D : cpo) (s : chain D) : - lim s = lim (lift_chain s). -Proof. -apply: poset_asym; apply: limM=>y H; first by apply: limP; right. -by case: H; [move=>-> | apply: limP]. -Qed. - -(* applied lim equals the lim of applications *) -(* ie., part of continuity of application *) -(* but is so often used, I give it a name *) - -Lemma lim_appE A (D : cpo) (s : chain [cpo of A -> D]) (x : A) : - lim s x = lim (app_chain s x). -Proof. by []. Qed. - -Lemma lim_dappE A (D : A -> cpo) (s : chain (dfunCPO D)) (x : A) : - lim s x = lim (dapp_chain s x). -Proof. by []. Qed. - -Section Continuity. -Variables (D1 D2 : cpo) (f : D1 -> D2). - -Definition continuous := - exists M : monotone f, - forall s : chain D1, f (lim s) = lim [f ^^ s by M]. - -Lemma cont_mono : continuous -> monotone f. -Proof. by case. Qed. - -Lemma contE (s : chain D1) (C : continuous) : - f (lim s) = lim [f ^^ s by cont_mono C]. -Proof. -case: C=>M E; rewrite E; congr (lim (image_chain _ _)). -by apply: pf_irr. -Qed. - -End Continuity. - -Section Kleene. -Variables (D : cpo) (f : D -> D) (C : continuous f). - -Fixpoint pow m := if m is n.+1 then f (pow n) else bot. - -Lemma pow_mono : monotone pow. -Proof. -move=>m n; elim: n m=>[|n IH] m /=; first by case: m. -rewrite {1}/Poset.leq /= leq_eqVlt ltnS. -case/orP; first by move/eqP=>->. -move/IH=>{IH} H; apply: poset_trans H _. -by elim: n=>[|n IH] //=; apply: cont_mono IH. -Qed. - -Definition pow_chain := [pow ^^ nat_chain by pow_mono]. - -Lemma reindex : pow_chain =p lift_chain [f ^^ pow_chain by cont_mono C]. -Proof. -move=>x; split. -- case; case=>[|n][->] /=; first by left. - by right; exists (pow n); split=>//; exists n. -case=>/=; first by move=>->; exists 0. -by case=>y [->][n][->]; exists n.+1. -Qed. - -Definition kleene_lfp := lim pow_chain. - -Lemma kleene_lfp_fixed : f kleene_lfp = kleene_lfp. -Proof. by rewrite (@contE _ _ f) lim_liftE; apply: limE; rewrite reindex. Qed. - -Lemma kleene_lfp_least : forall x, f x = x -> kleene_lfp <== x. -Proof. -move=>x H; apply: limM=>y [n][->] _. -by elim: n=>[|n IH] //=; rewrite -H; apply: cont_mono IH. -Qed. - -End Kleene. - -(**********************************) -(* Continuity of common functions *) -(**********************************) - -Lemma id_cont (D : cpo) : continuous (@id D). -Proof. by exists id_mono; move=>d; rewrite id_chainE. Qed. - -Arguments id_cont {D}. -Prenex Implicits id_cont. - -Lemma const_cont (D1 D2 : cpo) (y : D2) : continuous (fun x : D1 => y). -Proof. -exists const_mono; move=>s; apply: poset_asym. -- by apply: limP; case: s=>[p][[d H1] H2]; exists d. -by apply: limM=>_ [x][->]. -Qed. - -Arguments const_cont {D1 D2 y}. -Prenex Implicits const_cont. - -Lemma comp_cont (D1 D2 D3 : cpo) (f1 : D2 -> D1) (f2 : D3 -> D2) : - continuous f1 -> continuous f2 -> continuous (f1 \o f2). -Proof. -case=>M1 H1 [M2 H2]; exists (comp_mono M1 M2); move=>d. -by rewrite /= H2 H1 comp_chainE. -Qed. - -Arguments comp_cont [D1 D2 D3 f1 f2]. -Prenex Implicits comp_cont. - -Lemma proj1_cont (D1 D2 : cpo) : continuous (@fst D1 D2). -Proof. by exists proj1_mono. Qed. - -Lemma proj2_cont (D1 D2 : cpo) : continuous (@snd D1 D2). -Proof. by exists proj2_mono. Qed. - -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)). -Proof. -exists diag_mono=>d; apply: poset_asym; -by split=>/=; [rewrite proj1_diagE | rewrite proj2_diagE]. -Qed. - -Arguments diag_cont {D}. -Prenex Implicits diag_cont. - -Lemma app_cont A (D : cpo) x : continuous (fun f : A -> D => f x). -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. - -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) : - continuous f1 -> continuous f2 -> continuous (f1 \* f2). -Proof. -case=>M1 H1 [M2 H2]; exists (prod_mono M1 M2)=>d; apply: poset_asym; -by (split=>/=; [rewrite proj1_prodE H1 | rewrite proj2_prodE H2]). -Qed. - -Arguments prod_cont [S1 S2 T1 T2 f1 f2]. -Prenex Implicits prod_cont. diff --git a/theories/interlude.v b/theories/interlude.v deleted file mode 100644 index bf75a71..0000000 --- a/theories/interlude.v +++ /dev/null @@ -1,192 +0,0 @@ -From Coq Require Import ssreflect ssrbool ssrfun. -From mathcomp Require Import ssrnat seq eqtype path fintype finfun tuple interval perm fingroup. -From pcm Require Import options prelude ordtype seqext slice. - -Lemma implyb_trans a b c : a ==> b -> b ==> c -> a ==> c. -Proof. by case: a=>//=->. Qed. - -Section Seq. -Variable (A : Type). - -Lemma rcons_nseq n (x : A) : - rcons (nseq n x) x = nseq n.+1 x. -Proof. by elim: n=>//=n ->. Qed. - -Lemma behead_rcons (xs : seq A) x : - 0 < size xs -> - behead (rcons xs x) = rcons (behead xs) x. -Proof. by case: xs. Qed. - -Lemma path_all (xs : seq A) x r : - transitive r -> - path r x xs -> all (r x) xs. -Proof. by move=>Ht; rewrite path_sortedE; [case/andP | exact: Ht]. Qed. - -Lemma sorted_rconsE (leT : rel A) xs x : - transitive leT -> - sorted leT (rcons xs x) = all (leT^~ x) xs && sorted leT xs. -Proof. -move/rev_trans=>Ht; rewrite -(revK (rcons _ _)) rev_rcons rev_sorted /=. -by rewrite path_sortedE // all_rev rev_sorted. -Qed. - -Lemma sorted1 (r : rel A) xs : size xs == 1 -> sorted r xs. -Proof. by case: xs=>// x; case. Qed. - -End Seq. - -Section SeqEq. -Variable (A : eqType). - -Lemma perm_cons2 (x y : A) s : perm_eq [:: x, y & s] [:: y, x & s]. -Proof. -by rewrite (_ : [::x,y & s] = [::x] ++ [::y] ++ s) // - (_ : [::y,x & s] = [::y] ++ [::x] ++ s) // perm_catCA. -Qed. - -(* a weaker form of in_split *) -Lemma mem_split (x : A) s : - x \in s -> exists s1 s2, s = s1 ++ [:: x] ++ s2. -Proof. by case/in_split=>s1 [s2][H _]; exists s1, s2. Qed. - -Lemma mem_split_uniq (x : A) s : - x \in s -> uniq s -> - exists s1 s2, [/\ s = s1 ++ [:: x] ++ s2, - uniq (s1 ++ s2) & - x \notin s1 ++ s2]. -Proof. -move=>/[swap] Hu /mem_split [s1 [s2 H]]. -exists s1, s2; move: Hu. -by rewrite H uniq_catCA cons_uniq; case/andP. -Qed. - -Lemma all_notin (p : pred A) xs y : - all p xs -> ~~ p y -> y \notin xs. -Proof. by move/allP=>Ha; apply/contra/Ha. Qed. - -Lemma subset_all a (s1 s2 : seq A) : {subset s1 <= s2} -> all a s2 -> all a s1. -Proof. by move=>Hs /allP Ha1; apply/allP=>s /Hs /Ha1. Qed. - -End SeqEq. - -Section Allrel. -Variables (S T : Type). - -Lemma allrel_rconsl (r : T -> S -> bool) x xs ys : - allrel r (rcons xs x) ys = allrel r xs ys && all (r x) ys. -Proof. by rewrite -cats1 allrel_catl allrel1l. Qed. - -Lemma allrel_rconsr (r : T -> S -> bool) y xs ys : - allrel r xs (rcons ys y) = allrel r xs ys && all (r^~ y) xs. -Proof. by rewrite -cats1 allrel_catr allrel1r. Qed. - -End Allrel. - -Section AllrelEq. -Variables (S T : eqType). - -Lemma allrel_in_l (r : T -> S -> bool) (xs xs' : seq T) (ys : seq S) : - xs =i xs' -> - allrel r xs ys = allrel r xs' ys. -Proof. -by move=>H; rewrite !allrel_allpairsE; apply/eq_all_r/mem_allpairs_dep. -Qed. - -Lemma allrel_in_r (r : T -> S -> bool) (xs : seq T) (ys ys' : seq S) : - ys =i ys' -> - allrel r xs ys = allrel r xs ys'. -Proof. -by move=>H; rewrite !allrel_allpairsE; apply/eq_all_r/mem_allpairs_dep. -Qed. - -Lemma allrel_sub_l (r : T -> S -> bool) (xs xs' : seq T) (ys : seq S) : - {subset xs' <= xs} -> - allrel r xs ys -> allrel r xs' ys. -Proof. -move=>H Ha; apply/allrelP=>x y Hx Hy. -by move/allrelP: Ha; apply=>//; apply: H. -Qed. - -Lemma allrel_sub_r (r : T -> S -> bool) (xs : seq T) (ys ys' : seq S) : - {subset ys' <= ys} -> - allrel r xs ys -> allrel r xs ys'. -Proof. -move=>H Ha; apply/allrelP=>x y Hx Hy. -by move/allrelP: Ha; apply=>//; apply: H. -Qed. - -Lemma allrel_trans (xs ys : seq S) z r : - transitive r -> - all (r^~ z) xs -> all (r z) ys -> allrel r xs ys. -Proof. -move=>Ht /allP Ha /allP Hp; apply/allrelP=>x y + Hy. -by move/Ha/Ht; apply; apply: Hp. -Qed. - -End AllrelEq. - -Section SeqOrd. -Variable (A : ordType). - -Lemma ord_neq (x y : A) : ord x y -> x != y. -Proof. -move=>H; apply/negP=>/eqP E. -by rewrite E irr in H. -Qed. - -Lemma sorted_cat_cons_cat (l r : seq A) x : - sorted ord (l ++ x :: r) = sorted ord (l ++ [::x]) && sorted ord (x::r). -Proof. -rewrite !(sorted_pairwise (@trans A)) cats1 pairwise_cat pairwise_rcons allrel_consr !pairwise_cons. -case/boolP: (all (ord^~ x) l)=>//= Hl; case/boolP: (all (ord x) r)=>/= [Hr|_]; last by rewrite !andbF. -by rewrite (allrel_trans (@trans A) Hl Hr). -Qed. - -End SeqOrd. - -(* surgery on finfuns: slicing & permuting *) - -Section OnthCodom. -Variable (A : Type). - -Lemma onth_tnth {n} (s : n.-tuple A) (i : 'I_n) : onth s i = Some (tnth s i). -Proof. -elim: n s i =>[|n IH] s i; first by case: i. -case/tupleP: s=>/=x s; case: (unliftP ord0 i)=>[j|]-> /=. -- by rewrite tnthS. -by rewrite tnth0. -Qed. - -Lemma onth_codom {n} (i : 'I_n) (f: {ffun 'I_n -> A}) : onth (fgraph f) i = Some (f i). -Proof. -pose i' := cast_ord (esym (card_ord n)) i. -move: (@tnth_fgraph _ _ f i'); rewrite (enum_val_ord) {2}/i' cast_ordKV=><-. -by rewrite (onth_tnth (fgraph f) i'). -Qed. - -End OnthCodom. - -Section CodomWS. -Variable (n : nat) (A : Type). - -Lemma codom_ux_rcons (f : {ffun 'I_n -> A}) (i : 'I_n) : - &:(fgraph f) `]-oo, i : nat] = rcons (&:(fgraph f) `]-oo, i : nat[) (f i). -Proof. by rewrite slice_xR // slice_uu onth_codom. Qed. - -End CodomWS. - -Section PermFfun. -Variables (I : finType) (A : Type). - -Definition pffun (p : {perm I}) (f : {ffun I -> A}) := - [ffun i => f (p i)]. - -Lemma pffunE1 (f : {ffun I -> A}) : - pffun 1%g f = f. -Proof. by apply/ffunP=>i; rewrite !ffunE permE. Qed. - -Lemma pffunEM (p p' : {perm I}) (f : {ffun I -> A}) : - pffun (p * p') f = pffun p (pffun p' f). -Proof. by apply/ffunP => i; rewrite !ffunE permM. Qed. - -End PermFfun. diff --git a/theories/options.v b/theories/options.v deleted file mode 100644 index ddcdc8f..0000000 --- a/theories/options.v +++ /dev/null @@ -1,3 +0,0 @@ -Export Unset Program Cases. -(* turn off the automation of Program *) -Obligation Tactic := auto.