From 02a0929d55a1fcb4a306bd47bab3b16e3a7068da Mon Sep 17 00:00:00 2001
From: Robbert Krebbers <mail@robbertkrebbers.nl>
Date: Fri, 24 Mar 2017 15:49:25 +0100
Subject: [PATCH] Make big_opL type class opaque.
MIME-Version: 1.0
Content-Type: text/plain; charset=UTF-8
Content-Transfer-Encoding: 8bit

This commit fixes the issues that refolding of big operators did not work nicely
in the proof mode, e.g., given:

    Goal forall M (P : nat → uPred M) l,
      ([∗ list] x ∈ 10 :: l, P x) -∗ True.
    Proof. iIntros (M P l) "[H1 H2]".

We got:

    "H1" : P 10
    "H2" : (fix
            big_opL (M0 : ofeT) (o : M0 → M0 → M0) (H : Monoid o) (A : Type)
                    (f : nat → A → M0) (xs : list A) {struct xs} : M0 :=
              match xs with
              | [] => monoid_unit
              | x :: xs0 => o (f 0 x) (big_opL M0 o H A (λ n : nat, f (S n)) xs0)
              end) (uPredC M) uPred_sep uPred.uPred_sep_monoid nat
             (λ _ x : nat, P x) l
    --------------------------------------∗
    True

The problem here is that proof mode looked for an instance of `IntoAnd` for
`[∗ list] x ∈ 10 :: l, P x` and then applies the instance for separating conjunction
without folding back the fixpoint. This problem is not specific to the Iris proof
mode, but more of a general problem of Coq's `apply`, for example:

    Goal forall x l, Forall (fun _ => True) (map S (x :: l)).
    Proof.
      intros x l. constructor.

Gives:

     Forall (λ _ : nat, True)
       ((fix map (l0 : list nat) : list nat :=
          match l0 with
          | [] => []
          | a :: t => S a :: map t
          end) l)

This commit fixes this issue by making the big operators type class opaque and instead
handle them solely via corresponding type classes instances for the proof mode tactics.

Furthermore, note that we already had instances for persistence and timelessness. Those
were really needed; computation did not help to establish persistence when the list in
question was not a ground term. In fact, the sitation was worse, to establish persistence
of `[∗ list] x ∈ 10 :: l, P x` it could either use the persistence instance of big ops
directly, or use the persistency instance for `∗` first. Worst case, this can lead to an
exponential blow up because of back tracking.
---
 theories/algebra/big_op.v            |  1 +
 theories/base_logic/big_op.v         | 10 +++++-----
 theories/proofmode/class_instances.v | 23 +++++++++++++++++++++--
 3 files changed, 27 insertions(+), 7 deletions(-)

diff --git a/theories/algebra/big_op.v b/theories/algebra/big_op.v
index fcb0367ff..26e043615 100644
--- a/theories/algebra/big_op.v
+++ b/theories/algebra/big_op.v
@@ -25,6 +25,7 @@ Fixpoint big_opL `{Monoid M o} {A} (f : nat → A → M) (xs : list A) : M :=
   end.
 Instance: Params (@big_opL) 4.
 Arguments big_opL {M} o {_ A} _ !_ /.
+Typeclasses Opaque big_opL.
 Notation "'[^' o 'list]' k ↦ x ∈ l , P" := (big_opL o (λ k x, P) l)
   (at level 200, o at level 1, l at level 10, k, x at level 1, right associativity,
    format "[^ o  list]  k ↦ x  ∈  l ,  P") : C_scope.
diff --git a/theories/base_logic/big_op.v b/theories/base_logic/big_op.v
index 43b745d9d..72b8e10d8 100644
--- a/theories/base_logic/big_op.v
+++ b/theories/base_logic/big_op.v
@@ -151,21 +151,21 @@ Section list.
 
   Global Instance big_sepL_nil_persistent Φ :
     PersistentP ([∗ list] k↦x ∈ [], Φ k x).
-  Proof. apply _. Qed.
+  Proof. simpl; apply _. Qed.
   Global Instance big_sepL_persistent Φ l :
     (∀ k x, PersistentP (Φ k x)) → PersistentP ([∗ list] k↦x ∈ l, Φ k x).
   Proof. revert Φ. induction l as [|x l IH]=> Φ ? /=; apply _. Qed.
   Global Instance big_sepL_persistent_id Ps : PersistentL Ps → PersistentP ([∗] Ps).
-  Proof. induction 1; apply _. Qed.
+  Proof. induction 1; simpl; apply _. Qed.
 
   Global Instance big_sepL_nil_timeless Φ :
     TimelessP ([∗ list] k↦x ∈ [], Φ k x).
-  Proof. apply _. Qed.
+  Proof. simpl; apply _. Qed.
   Global Instance big_sepL_timeless Φ l :
     (∀ k x, TimelessP (Φ k x)) → TimelessP ([∗ list] k↦x ∈ l, Φ k x).
   Proof. revert Φ. induction l as [|x l IH]=> Φ ? /=; apply _. Qed.
   Global Instance big_sepL_timeless_id Ps : TimelessL Ps → TimelessP ([∗] Ps).
-  Proof. induction 1; apply _. Qed.
+  Proof. induction 1; simpl; apply _. Qed.
 End list.
 
 Section list2.
@@ -589,4 +589,4 @@ Section gmultiset.
 End gmultiset.
 End big_op.
 
-Hint Resolve big_sepM_empty' big_sepS_empty' big_sepMS_empty' | 0.
+Hint Resolve big_sepL_nil' big_sepM_empty' big_sepS_empty' big_sepMS_empty' | 0.
diff --git a/theories/proofmode/class_instances.v b/theories/proofmode/class_instances.v
index e24a45acf..b5176725d 100644
--- a/theories/proofmode/class_instances.v
+++ b/theories/proofmode/class_instances.v
@@ -343,6 +343,14 @@ Global Instance from_sep_bupd P Q1 Q2 :
   FromAnd false P Q1 Q2 → FromAnd false (|==> P) (|==> Q1) (|==> Q2).
 Proof. rewrite /FromAnd=><-. apply bupd_sep. Qed.
 
+Global Instance from_and_big_sepL_cons {A} (Φ : nat → A → uPred M) x l :
+  FromAnd false ([∗ list] k ↦ y ∈ x :: l, Φ k y) (Φ 0 x) ([∗ list] k ↦ y ∈ l, Φ (S k) y).
+Proof. by rewrite /FromAnd big_sepL_cons. Qed.
+Global Instance from_and_big_sepL_cons_persistent {A} (Φ : nat → A → uPred M) x l :
+  PersistentP (Φ 0 x) →
+  FromAnd true ([∗ list] k ↦ y ∈ x :: l, Φ k y) (Φ 0 x) ([∗ list] k ↦ y ∈ l, Φ (S k) y).
+Proof. intros. by rewrite /FromAnd big_opL_cons always_and_sep_l. Qed.
+
 Global Instance from_and_big_sepL_app {A} (Φ : nat → A → uPred M) l1 l2 :
   FromAnd false ([∗ list] k ↦ y ∈ l1 ++ l2, Φ k y)
     ([∗ list] k ↦ y ∈ l1, Φ k y) ([∗ list] k ↦ y ∈ l2, Φ (length l1 + k) y).
@@ -423,8 +431,14 @@ Global Instance into_and_laterN n p P Q1 Q2 :
   IntoAnd p P Q1 Q2 → IntoAnd p (▷^n P) (▷^n Q1) (▷^n Q2).
 Proof. rewrite /IntoAnd=>->. destruct p; by rewrite ?laterN_and ?laterN_sep. Qed.
 
-(* We use [IsApp] to make sure that [frame_big_sepL_app] cannot be applied
-repeatedly often when having [ [∗ list] k ↦ x ∈ ?e, Φ k x] with [?e] an evar. *)
+(* We use [IsCons] and [IsApp] to make sure that [frame_big_sepL_cons] and
+[frame_big_sepL_app] cannot be applied repeatedly often when having
+[ [∗ list] k ↦ x ∈ ?e, Φ k x] with [?e] an evar. *)
+Global Instance into_and_big_sepL_cons {A} p (Φ : nat → A → uPred M) l x l' :
+  IsCons l x l' →
+  IntoAnd p ([∗ list] k ↦ y ∈ l, Φ k y)
+    (Φ 0 x) ([∗ list] k ↦ y ∈ l', Φ (S k) y).
+Proof. rewrite /IsCons=>->. apply mk_into_and_sep. by rewrite big_sepL_cons. Qed.
 Global Instance into_and_big_sepL_app {A} p (Φ : nat → A → uPred M) l l1 l2 :
   IsApp l l1 l2 →
   IntoAnd p ([∗ list] k ↦ y ∈ l, Φ k y)
@@ -459,6 +473,11 @@ Global Instance frame_sep_r p R P1 P2 Q Q' :
   Frame p R P2 Q → MakeSep P1 Q Q' → Frame p R (P1 ∗ P2) Q' | 10.
 Proof. rewrite /Frame /MakeSep => <- <-. by rewrite assoc -(comm _ P1) assoc. Qed.
 
+Global Instance frame_big_sepL_cons {A} p (Φ : nat → A → uPred M) R Q l x l' :
+  IsCons l x l' →
+  Frame p R (Φ 0 x ∗ [∗ list] k ↦ y ∈ l', Φ (S k) y) Q →
+  Frame p R ([∗ list] k ↦ y ∈ l, Φ k y) Q.
+Proof. rewrite /IsCons=>->. by rewrite /Frame big_sepL_cons. Qed.
 Global Instance frame_big_sepL_app {A} p (Φ : nat → A → uPred M) R Q l l1 l2 :
   IsApp l l1 l2 →
   Frame p R (([∗ list] k ↦ y ∈ l1, Φ k y) ∗
-- 
GitLab