From ead5cad99e2dd336f68f964780392b9486fbeb4b Mon Sep 17 00:00:00 2001
From: Ralf Jung <jung@mpi-sws.org>
Date: Mon, 22 Feb 2016 15:36:19 +0100
Subject: [PATCH] provide a closed proof of the client

---
 barrier/client.v                | 33 ++++++++++++++++++++++++++++++---
 heap_lang/tests.v               |  4 ++--
 program_logic/ghost_ownership.v |  1 +
 3 files changed, 33 insertions(+), 5 deletions(-)

diff --git a/barrier/client.v b/barrier/client.v
index 2d11a0135..1403ea677 100644
--- a/barrier/client.v
+++ b/barrier/client.v
@@ -1,4 +1,5 @@
 From barrier Require Import barrier.
+From program_logic Require Import auth sts saved_prop hoare ownership.
 (* FIXME This needs to be imported even though barrier exports it *)
 From heap_lang Require Import notation.
 Import uPred.
@@ -17,9 +18,7 @@ Section client.
     heapN ⊥ N → heap_ctx heapN ⊑ || client {{ λ _, True }}.
   Proof.
     intros ?. rewrite /client.
-    (* FIXME: wp (eapply newchan_spec with (P:=True%I)). *)
-    wp_focus (newchan _).
-    rewrite -(newchan_spec N heapN True%I) //.
+    ewp eapply (newchan_spec N heapN True%I); last done.
     apply sep_intro_True_r; first done.
     apply forall_intro=>l. apply wand_intro_l. rewrite right_id.
     wp_let. etrans; last eapply wait_spec.
@@ -27,3 +26,31 @@ Section client.
   Qed.
 
 End client.
+
+Section ClosedProofs.
+  Definition Σ : iFunctorG := agreeF .:: constF (stsRA barrier_proto.sts) .:: authF (constF heapRA)
+                                     .:: (λ _, constF unitRA) : iFunctorG.
+  Notation iProp := (iPropG heap_lang Σ).
+
+  Instance: authG heap_lang Σ heapRA.
+  Proof. split; try apply _. by exists 2%nat. Qed.
+
+  Instance: stsG heap_lang Σ barrier_proto.sts.
+  Proof. split; try apply _. by exists 1%nat. Qed.
+
+  Instance: savedPropG heap_lang Σ.
+  Proof. by exists 0%nat. Qed.
+
+  Lemma client_safe_closed σ : {{ ownP σ : iProp }} client {{ λ v, True }}.
+  Proof.
+    apply ht_alt. rewrite (heap_alloc ⊤ (ndot nroot "Barrier")); last first.
+    { (* FIXME Really?? set_solver takes forever on "⊆ ⊤"?!? *)
+      move=>? _. exact I. }
+    apply wp_strip_pvs, exist_elim=> ?. rewrite and_elim_l.
+    rewrite -(client_safe (ndot nroot "Heap" ) (ndot nroot "Barrier" )) //.
+    (* This, too, should be automated. *)
+    apply ndot_ne_disjoint. discriminate.
+  Qed.
+
+  Print Assumptions client_safe_closed.
+End ClosedProofs.
diff --git a/heap_lang/tests.v b/heap_lang/tests.v
index 5035a8d3d..1a961c2da 100644
--- a/heap_lang/tests.v
+++ b/heap_lang/tests.v
@@ -82,12 +82,12 @@ Section ClosedProofs.
   Instance: authG heap_lang Σ heapRA.
   Proof. split; try apply _. by exists 1%nat. Qed.
 
-  Lemma heap_e_hoare σ : {{ ownP σ : iProp }} heap_e {{ λ v, v = '2 }}.
+  Lemma heap_e_closed σ : {{ ownP σ : iProp }} heap_e {{ λ v, v = '2 }}.
   Proof.
     apply ht_alt. rewrite (heap_alloc ⊤ nroot); last by rewrite nclose_nroot.
     apply wp_strip_pvs, exist_elim=> ?. rewrite and_elim_l.
     rewrite -heap_e_spec; first by eauto with I. by rewrite nclose_nroot.
   Qed.
 
-  Print Assumptions heap_e_hoare.
+  Print Assumptions heap_e_closed.
 End ClosedProofs.
diff --git a/program_logic/ghost_ownership.v b/program_logic/ghost_ownership.v
index e17065f75..ab156cae7 100644
--- a/program_logic/ghost_ownership.v
+++ b/program_logic/ghost_ownership.v
@@ -1,3 +1,4 @@
+From prelude Require Export functions.
 From algebra Require Export iprod.
 From program_logic Require Export pviewshifts.
 From program_logic Require Import ownership.
-- 
GitLab