From 7e480ad465f00feeae38365a124235b0a1d3192a Mon Sep 17 00:00:00 2001
From: Robbert Krebbers <mail@robbertkrebbers.nl>
Date: Mon, 25 Jul 2016 10:16:51 +0200
Subject: [PATCH] Move local updates to separate file.

---
 _CoqProject                 |   1 +
 algebra/auth.v              |   4 +-
 algebra/coPset.v            |   2 +-
 algebra/csum.v              |   4 +-
 algebra/gmap.v              |   4 +-
 algebra/gset.v              |   2 +-
 algebra/list.v              |   4 +-
 algebra/local_updates.v     | 100 ++++++++++++++++++++++++++++++++++++
 algebra/updates.v           |  99 +----------------------------------
 program_logic/pviewshifts.v |   2 +-
 program_logic/wsat.v        |   1 +
 11 files changed, 115 insertions(+), 108 deletions(-)
 create mode 100644 algebra/local_updates.v

diff --git a/_CoqProject b/_CoqProject
index 01ac4d7aa..8164834a3 100644
--- a/_CoqProject
+++ b/_CoqProject
@@ -59,6 +59,7 @@ algebra/frac.v
 algebra/csum.v
 algebra/list.v
 algebra/updates.v
+algebra/local_updates.v
 algebra/gset.v
 algebra/coPset.v
 program_logic/model.v
diff --git a/algebra/auth.v b/algebra/auth.v
index 4e33a44a7..f9cfc7c35 100644
--- a/algebra/auth.v
+++ b/algebra/auth.v
@@ -1,5 +1,5 @@
-From iris.algebra Require Export excl updates.
-From iris.algebra Require Import upred.
+From iris.algebra Require Export excl local_updates.
+From iris.algebra Require Import upred updates.
 Local Arguments valid _ _ !_ /.
 Local Arguments validN _ _ _ !_ /.
 
diff --git a/algebra/coPset.v b/algebra/coPset.v
index 5efa7e7ee..adaa01d0a 100644
--- a/algebra/coPset.v
+++ b/algebra/coPset.v
@@ -1,5 +1,5 @@
 From iris.algebra Require Export cmra.
-From iris.algebra Require Import updates.
+From iris.algebra Require Import updates local_updates.
 From iris.prelude Require Export collections coPset.
 
 (** This is pretty much the same as algebra/gset, but I was not able to
diff --git a/algebra/csum.v b/algebra/csum.v
index c4f87af3d..341aa6bed 100644
--- a/algebra/csum.v
+++ b/algebra/csum.v
@@ -1,5 +1,5 @@
-From iris.algebra Require Export cmra updates.
-From iris.algebra Require Import upred.
+From iris.algebra Require Export cmra.
+From iris.algebra Require Import upred updates local_updates.
 Local Arguments pcore _ _ !_ /.
 Local Arguments cmra_pcore _ !_ /.
 Local Arguments validN _ _ _ !_ /.
diff --git a/algebra/gmap.v b/algebra/gmap.v
index 8b2ce73d1..2b074fb79 100644
--- a/algebra/gmap.v
+++ b/algebra/gmap.v
@@ -1,6 +1,6 @@
-From iris.algebra Require Export cmra updates.
+From iris.algebra Require Export cmra.
 From iris.prelude Require Export gmap.
-From iris.algebra Require Import upred.
+From iris.algebra Require Import upred updates local_updates.
 
 Section cofe.
 Context `{Countable K} {A : cofeT}.
diff --git a/algebra/gset.v b/algebra/gset.v
index 8701a4744..cca312c3a 100644
--- a/algebra/gset.v
+++ b/algebra/gset.v
@@ -1,5 +1,5 @@
 From iris.algebra Require Export cmra.
-From iris.algebra Require Import updates.
+From iris.algebra Require Import updates local_updates.
 From iris.prelude Require Export collections gmap.
 
 Inductive gset_disj K `{Countable K} :=
diff --git a/algebra/list.v b/algebra/list.v
index 44f4c4406..b4f0803a3 100644
--- a/algebra/list.v
+++ b/algebra/list.v
@@ -1,6 +1,6 @@
-From iris.algebra Require Export cmra updates.
+From iris.algebra Require Export cmra.
 From iris.prelude Require Export list.
-From iris.algebra Require Import upred.
+From iris.algebra Require Import upred updates local_updates.
 
 Section cofe.
 Context {A : cofeT}.
diff --git a/algebra/local_updates.v b/algebra/local_updates.v
new file mode 100644
index 000000000..efc3df3de
--- /dev/null
+++ b/algebra/local_updates.v
@@ -0,0 +1,100 @@
+From iris.algebra Require Export cmra.
+
+(** * Local updates *)
+Record local_update {A : cmraT} (mz : option A) (x y : A) := {
+  local_update_valid n : ✓{n} (x ⋅? mz) → ✓{n} (y ⋅? mz);
+  local_update_go n mz' :
+    ✓{n} (x ⋅? mz) → x ⋅? mz ≡{n}≡ x ⋅? mz' → y ⋅? mz ≡{n}≡ y ⋅? mz'
+}.
+Notation "x ~l~> y @ mz" := (local_update mz x y) (at level 70).
+Instance: Params (@local_update) 1.
+
+Section updates.
+Context {A : cmraT}.
+Implicit Types x y : A.
+
+Global Instance local_update_proper :
+  Proper ((≡) ==> (≡) ==> (≡) ==> iff) (@local_update A).
+Proof.
+  cut (Proper ((≡) ==> (≡) ==> (≡) ==> impl) (@local_update A)).
+  { intros Hproper; split; by apply Hproper. }
+  intros mz mz' Hmz x x' Hx y y' Hy [Hv Hup]; constructor; setoid_subst; auto.
+Qed.
+
+Global Instance local_update_preorder mz : PreOrder (@local_update A mz).
+Proof.
+  split.
+  - intros x; by split.
+  - intros x1 x2 x3 [??] [??]; split; eauto.
+Qed.
+
+Lemma exclusive_local_update `{!Exclusive x} y mz : ✓ y → x ~l~> y @ mz.
+Proof.
+  split; intros n.
+  - move=> /exclusiveN_opM ->. by apply cmra_valid_validN.
+  - intros mz' ? Hmz.
+    by rewrite (exclusiveN_opM n x mz) // (exclusiveN_opM n x mz') -?Hmz.
+Qed.
+
+Lemma op_local_update x1 x2 y mz :
+  x1 ~l~> x2 @ Some (y ⋅? mz) → x1 ⋅ y ~l~> x2 ⋅ y @ mz.
+Proof.
+  intros [Hv1 H1]; split.
+  - intros n. rewrite !cmra_opM_assoc. move=> /Hv1 /=; auto.
+  - intros n mz'. rewrite !cmra_opM_assoc. move=> Hv /(H1 _ (Some _) Hv) /=; auto.
+Qed.
+
+Lemma alloc_local_update x y mz :
+  (∀ n, ✓{n} (x ⋅? mz) → ✓{n} (x ⋅ y ⋅? mz)) → x ~l~> x ⋅ y @ mz.
+Proof.
+  split; first done.
+  intros n mz' _. by rewrite !(comm _ x) !cmra_opM_assoc=> ->.
+Qed.
+
+Lemma local_update_total `{CMRADiscrete A} x y mz :
+  x ~l~> y @ mz ↔
+    (✓ (x ⋅? mz) → ✓ (y ⋅? mz)) ∧
+    (∀ mz', ✓ (x ⋅? mz) → x ⋅? mz ≡ x ⋅? mz' → y ⋅? mz ≡ y ⋅? mz').
+Proof.
+  split.
+  - destruct 1. split; intros until 0;
+      rewrite !(cmra_discrete_valid_iff 0) ?(timeless_iff 0); auto.
+  - intros [??]; split; intros until 0;
+      rewrite -!cmra_discrete_valid_iff -?timeless_iff; auto.
+Qed.
+End updates.
+
+(** * Product *)
+Lemma prod_local_update {A B : cmraT} (x y : A * B) mz :
+  x.1 ~l~> y.1 @ fst <$> mz → x.2 ~l~> y.2 @ snd <$> mz →
+  x ~l~> y @ mz.
+Proof.
+  intros [Hv1 H1] [Hv2 H2]; split.
+  - intros n [??]; destruct mz; split; auto.
+  - intros n mz' [??] [??].
+    specialize (H1 n (fst <$> mz')); specialize (H2 n (snd <$> mz')).
+    destruct mz, mz'; split; naive_solver.
+Qed.
+
+(** * Option *)
+Lemma option_local_update {A : cmraT} (x y : A) mmz :
+  x ~l~> y @ mjoin mmz →
+  Some x ~l~> Some y @ mmz.
+Proof.
+  intros [Hv H]; split; first destruct mmz as [[?|]|]; auto.
+  intros n mmz'. specialize (H n (mjoin mmz')).
+  destruct mmz as [[]|], mmz' as [[]|]; inversion_clear 2; constructor; auto.
+Qed.
+
+(** * Natural numbers  *)
+Lemma nat_local_update (x y : nat) mz : x ~l~> y @ mz.
+Proof.
+  split; first done.
+  compute -[plus]; destruct mz as [z|]; intros n [z'|]; lia.
+Qed.
+
+Lemma mnat_local_update (x y : mnat) mz : x ≤ y → x ~l~> y @ mz.
+Proof.
+  split; first done.
+  compute -[max]; destruct mz as [z|]; intros n [z'|]; lia.
+Qed.
diff --git a/algebra/updates.v b/algebra/updates.v
index c856ab53e..012b53aaf 100644
--- a/algebra/updates.v
+++ b/algebra/updates.v
@@ -1,14 +1,5 @@
 From iris.algebra Require Export cmra.
 
-(** * Local updates *)
-Record local_update {A : cmraT} (mz : option A) (x y : A) := {
-  local_update_valid n : ✓{n} (x ⋅? mz) → ✓{n} (y ⋅? mz);
-  local_update_go n mz' :
-    ✓{n} (x ⋅? mz) → x ⋅? mz ≡{n}≡ x ⋅? mz' → y ⋅? mz ≡{n}≡ y ⋅? mz'
-}.
-Notation "x ~l~> y @ mz" := (local_update mz x y) (at level 70).
-Instance: Params (@local_update) 1.
-
 (** * Frame preserving updates *)
 (* This quantifies over [option A] for the frame.  That is necessary to
    make the following hold:
@@ -24,18 +15,10 @@ Definition cmra_update {A : cmraT} (x y : A) := ∀ n mz,
 Infix "~~>" := cmra_update (at level 70).
 Instance: Params (@cmra_update) 1.
 
-(** ** CMRAs *)
-Section cmra.
+Section updates.
 Context {A : cmraT}.
 Implicit Types x y : A.
 
-Global Instance local_update_proper :
-  Proper ((≡) ==> (≡) ==> (≡) ==> iff) (@local_update A).
-Proof.
-  cut (Proper ((≡) ==> (≡) ==> (≡) ==> impl) (@local_update A)).
-  { intros Hproper; split; by apply Hproper. }
-  intros mz mz' Hmz x x' Hx y y' Hy [Hv Hup]; constructor; setoid_subst; auto.
-Qed.
 Global Instance cmra_updateP_proper :
   Proper ((≡) ==> pointwise_relation _ iff ==> iff) (@cmra_updateP A).
 Proof.
@@ -48,51 +31,6 @@ Proof.
   rewrite /cmra_update=> x x' Hx y y' Hy; split=> ? n mz ?; setoid_subst; auto.
 Qed.
 
-(** ** Local updates *)
-Global Instance local_update_preorder mz : PreOrder (@local_update A mz).
-Proof.
-  split.
-  - intros x; by split.
-  - intros x1 x2 x3 [??] [??]; split; eauto.
-Qed.
-
-Lemma exclusive_local_update `{!Exclusive x} y mz : ✓ y → x ~l~> y @ mz.
-Proof.
-  split; intros n.
-  - move=> /exclusiveN_opM ->. by apply cmra_valid_validN.
-  - intros mz' ? Hmz.
-    by rewrite (exclusiveN_opM n x mz) // (exclusiveN_opM n x mz') -?Hmz.
-Qed.
-
-Lemma op_local_update x1 x2 y mz :
-  x1 ~l~> x2 @ Some (y ⋅? mz) → x1 ⋅ y ~l~> x2 ⋅ y @ mz.
-Proof.
-  intros [Hv1 H1]; split.
-  - intros n. rewrite !cmra_opM_assoc. move=> /Hv1 /=; auto.
-  - intros n mz'. rewrite !cmra_opM_assoc. move=> Hv /(H1 _ (Some _) Hv) /=; auto.
-Qed.
-
-Lemma alloc_local_update x y mz :
-  (∀ n, ✓{n} (x ⋅? mz) → ✓{n} (x ⋅ y ⋅? mz)) → x ~l~> x ⋅ y @ mz.
-Proof.
-  split; first done.
-  intros n mz' _. by rewrite !(comm _ x) !cmra_opM_assoc=> ->.
-Qed.
-
-(** ** Local updates for discrete CMRAs *)
-Lemma local_update_total `{CMRADiscrete A} x y mz :
-  x ~l~> y @ mz ↔
-    (✓ (x ⋅? mz) → ✓ (y ⋅? mz)) ∧
-    (∀ mz', ✓ (x ⋅? mz) → x ⋅? mz ≡ x ⋅? mz' → y ⋅? mz ≡ y ⋅? mz').
-Proof.
-  split.
-  - destruct 1. split; intros until 0;
-      rewrite !(cmra_discrete_valid_iff 0) ?(timeless_iff 0); auto.
-  - intros [??]; split; intros until 0;
-      rewrite -!cmra_discrete_valid_iff -?timeless_iff; auto.
-Qed.
-
-(** ** Frame preserving updates *)
 Lemma cmra_update_updateP x y : x ~~> y ↔ x ~~>: (y =).
 Proof. split=> Hup n z ?; eauto. destruct (Hup n z) as (?&<-&?); auto. Qed.
 Lemma cmra_updateP_id (P : A → Prop) x : P x → x ~~>: P.
@@ -176,7 +114,7 @@ Section total_updates.
     naive_solver eauto using 0.
   Qed.
 End total_updates.
-End cmra.
+End updates.
 
 (** * Transport *)
 Section cmra_transport.
@@ -195,17 +133,6 @@ Section prod.
   Context {A B : cmraT}.
   Implicit Types x : A * B.
 
-  Lemma prod_local_update x y mz :
-    x.1 ~l~> y.1 @ fst <$> mz → x.2 ~l~> y.2 @ snd <$> mz →
-    x ~l~> y @ mz.
-  Proof.
-    intros [Hv1 H1] [Hv2 H2]; split.
-    - intros n [??]; destruct mz; split; auto.
-    - intros n mz' [??] [??].
-      specialize (H1 n (fst <$> mz')); specialize (H2 n (snd <$> mz')).
-      destruct mz, mz'; split; naive_solver.
-  Qed.
-
   Lemma prod_updateP P1 P2 (Q : A * B → Prop) x :
     x.1 ~~>: P1 → x.2 ~~>: P2 → (∀ a b, P1 a → P2 b → Q (a,b)) → x ~~>: Q.
   Proof.
@@ -229,15 +156,6 @@ Section option.
   Context {A : cmraT}.
   Implicit Types x y : A.
 
-  Lemma option_local_update x y mmz :
-    x ~l~> y @ mjoin mmz →
-    Some x ~l~> Some y @ mmz.
-  Proof.
-    intros [Hv H]; split; first destruct mmz as [[?|]|]; auto.
-    intros n mmz'. specialize (H n (mjoin mmz')).
-    destruct mmz as [[]|], mmz' as [[]|]; inversion_clear 2; constructor; auto.
-  Qed.
-
   Lemma option_updateP (P : A → Prop) (Q : option A → Prop) x :
     x ~~>: P → (∀ y, P y → Q (Some y)) → Some x ~~>: Q.
   Proof.
@@ -252,16 +170,3 @@ Section option.
   Lemma option_update x y : x ~~> y → Some x ~~> Some y.
   Proof. rewrite !cmra_update_updateP; eauto using option_updateP with subst. Qed.
 End option.
-
-(** * Natural numbers  *)
-Lemma nat_local_update (x y : nat) mz : x ~l~> y @ mz.
-Proof.
-  split; first done.
-  compute -[plus]; destruct mz as [z|]; intros n [z'|]; lia.
-Qed.
-
-Lemma mnat_local_update (x y : mnat) mz : x ≤ y → x ~l~> y @ mz.
-Proof.
-  split; first done.
-  compute -[max]; destruct mz as [z|]; intros n [z'|]; lia.
-Qed.
diff --git a/program_logic/pviewshifts.v b/program_logic/pviewshifts.v
index 7df781061..9cbab38f8 100644
--- a/program_logic/pviewshifts.v
+++ b/program_logic/pviewshifts.v
@@ -1,5 +1,5 @@
 From iris.prelude Require Export coPset.
-From iris.algebra Require Export upred_big_op.
+From iris.algebra Require Export upred_big_op updates.
 From iris.program_logic Require Export model.
 From iris.program_logic Require Import ownership wsat.
 Local Hint Extern 10 (_ ≤ _) => omega.
diff --git a/program_logic/wsat.v b/program_logic/wsat.v
index 8e818b8c3..53f94b2d2 100644
--- a/program_logic/wsat.v
+++ b/program_logic/wsat.v
@@ -1,6 +1,7 @@
 From iris.prelude Require Export coPset.
 From iris.program_logic Require Export model.
 From iris.algebra Require Export cmra_big_op cmra_tactics.
+From iris.algebra Require Import updates.
 Local Hint Extern 10 (_ ≤ _) => omega.
 Local Hint Extern 10 (✓{_} _) => solve_validN.
 Local Hint Extern 1 (✓{_} gst _) => apply gst_validN.
-- 
GitLab