Skip to content
Snippets Groups Projects
Commit 9672a803 authored by Jacques-Henri Jourdan's avatar Jacques-Henri Jourdan
Browse files

Separate file updates.v

parent 1650c2a0
No related branches found
No related tags found
No related merge requests found
......@@ -29,6 +29,7 @@ theories/bi/interface.v
theories/bi/derived_connectives.v
theories/bi/derived_laws.v
theories/bi/big_op.v
theories/bi/updates.v
theories/bi/bi.v
theories/bi/tactics.v
theories/bi/fractional.v
......
From iris.algebra Require Export cmra updates.
From iris.bi Require Export derived_connectives.
From iris.bi Require Export derived_connectives updates.
From stdpp Require Import finite.
Set Default Proof Using "Type".
Local Hint Extern 1 (_ _) => etrans; [eassumption|].
......
From iris.bi Require Export derived_laws big_op.
From iris.bi Require Export derived_laws big_op updates.
Set Default Proof Using "Type".
Module Import bi.
......
From stdpp Require Import coPset.
From iris.algebra Require Export ofe.
Set Primitive Projections.
......@@ -525,34 +524,3 @@ Proof. eapply sbi_mixin_later_false_em, sbi_sbi_mixin. Qed.
End sbi_laws.
End bi.
Class BUpd (A : Type) : Type := bupd : A A.
Instance : Params (@bupd) 2.
Notation "|==> Q" := (bupd Q)
(at level 99, Q at level 200, format "|==> Q") : bi_scope.
Notation "P ==∗ Q" := (P |==> Q)
(at level 99, Q at level 200, only parsing) : stdpp_scope.
Notation "P ==∗ Q" := (P -∗ |==> Q)%I
(at level 99, Q at level 200, format "P ==∗ Q") : bi_scope.
Class FUpd (A : Type) : Type := fupd : coPset coPset A A.
Instance: Params (@fupd) 4.
Notation "|={ E1 , E2 }=> Q" := (fupd E1 E2 Q)
(at level 99, E1, E2 at level 50, Q at level 200,
format "|={ E1 , E2 }=> Q") : bi_scope.
Notation "P ={ E1 , E2 }=∗ Q" := (P -∗ |={E1,E2}=> Q)%I
(at level 99, E1,E2 at level 50, Q at level 200,
format "P ={ E1 , E2 }=∗ Q") : bi_scope.
Notation "P ={ E1 , E2 }=∗ Q" := (P -∗ |={E1,E2}=> Q)
(at level 99, E1, E2 at level 50, Q at level 200, only parsing) : stdpp_scope.
Notation "|={ E }=> Q" := (fupd E E Q)
(at level 99, E at level 50, Q at level 200,
format "|={ E }=> Q") : bi_scope.
Notation "P ={ E }=∗ Q" := (P -∗ |={E}=> Q)%I
(at level 99, E at level 50, Q at level 200,
format "P ={ E }=∗ Q") : bi_scope.
Notation "P ={ E }=∗ Q" := (P -∗ |={E}=> Q)
(at level 99, E at level 50, Q at level 200, only parsing) : stdpp_scope.
From stdpp Require Import coPset.
From iris.bi Require Import interface.
Class BUpd (A : Type) : Type := bupd : A A.
Instance : Params (@bupd) 2.
Notation "|==> Q" := (bupd Q)
(at level 99, Q at level 200, format "|==> Q") : bi_scope.
Notation "P ==∗ Q" := (P |==> Q)
(at level 99, Q at level 200, only parsing) : stdpp_scope.
Notation "P ==∗ Q" := (P -∗ |==> Q)%I
(at level 99, Q at level 200, format "P ==∗ Q") : bi_scope.
Class FUpd (A : Type) : Type := fupd : coPset coPset A A.
Instance: Params (@fupd) 4.
Notation "|={ E1 , E2 }=> Q" := (fupd E1 E2 Q)
(at level 99, E1, E2 at level 50, Q at level 200,
format "|={ E1 , E2 }=> Q") : bi_scope.
Notation "P ={ E1 , E2 }=∗ Q" := (P -∗ |={E1,E2}=> Q)%I
(at level 99, E1,E2 at level 50, Q at level 200,
format "P ={ E1 , E2 }=∗ Q") : bi_scope.
Notation "P ={ E1 , E2 }=∗ Q" := (P -∗ |={E1,E2}=> Q)
(at level 99, E1, E2 at level 50, Q at level 200, only parsing) : stdpp_scope.
Notation "|={ E }=> Q" := (fupd E E Q)
(at level 99, E at level 50, Q at level 200,
format "|={ E }=> Q") : bi_scope.
Notation "P ={ E }=∗ Q" := (P -∗ |={E}=> Q)%I
(at level 99, E at level 50, Q at level 200,
format "P ={ E }=∗ Q") : bi_scope.
Notation "P ={ E }=∗ Q" := (P -∗ |={E}=> Q)
(at level 99, E at level 50, Q at level 200, only parsing) : stdpp_scope.
0% Loading or .
You are about to add 0 people to the discussion. Proceed with caution.
Finish editing this message first!
Please register or to comment