Skip to content
Snippets Groups Projects
Commit dff7eb3c authored by Robbert Krebbers's avatar Robbert Krebbers
Browse files

Merge branch 'master' of gitlab.mpi-sws.org:FP/iris-coq

parents c05f2a06 978eec47
No related branches found
No related tags found
No related merge requests found
......@@ -130,11 +130,11 @@ Lemma to_agree_car n (x : agree A) : ✓{n} x → to_agree (x n) ≡{n}≡ x.
Proof. intros [??]; split; naive_solver eauto using agree_valid_le. Qed.
(** Internalized properties *)
Lemma agree_equivI {M} a b : (to_agree a to_agree b)%I (a b : uPred M)%I.
Lemma agree_equivI {M} a b : (to_agree a to_agree b) ⊣⊢ (a b : uPred M).
Proof.
uPred.unseal. do 2 split. by intros [? Hv]; apply (Hv n). apply: to_agree_ne.
Qed.
Lemma agree_validI {M} x y : (x y) (x y : uPred M).
Lemma agree_validI {M} x y : (x y) (x y : uPred M).
Proof. uPred.unseal; split=> r n _ ?; by apply: agree_op_inv. Qed.
End agree.
......
......@@ -138,14 +138,14 @@ Admitted.
(** Internalized properties *)
Lemma auth_equivI {M} (x y : auth A) :
(x y)%I (authoritative x authoritative y own x own y : uPred M)%I.
(x y) ⊣⊢ (authoritative x authoritative y own x own y : uPred M).
Proof. by uPred.unseal. Qed.
Lemma auth_validI {M} (x : auth A) :
( x)%I (match authoritative x with
( x) ⊣⊢ (match authoritative x with
| Excl a => ( b, a own x b) a
| ExclUnit => own x
| ExclBot => False
end : uPred M)%I.
end : uPred M).
Proof. uPred.unseal. by destruct x as [[]]. Qed.
(** The notations ◯ and ● only work for CMRAs with an empty element. So, in
......
......@@ -138,16 +138,16 @@ Qed.
(** Internalized properties *)
Lemma excl_equivI {M} (x y : excl A) :
(x y)%I (match x, y with
(x y) ⊣⊢ (match x, y with
| Excl a, Excl b => a b
| ExclUnit, ExclUnit | ExclBot, ExclBot => True
| _, _ => False
end : uPred M)%I.
end : uPred M).
Proof.
uPred.unseal. do 2 split. by destruct 1. by destruct x, y; try constructor.
Qed.
Lemma excl_validI {M} (x : excl A) :
( x)%I (if x is ExclBot then False else True : uPred M)%I.
( x) ⊣⊢ (if x is ExclBot then False else True : uPred M).
Proof. uPred.unseal. by destruct x. Qed.
(** ** Local updates *)
......
......@@ -164,9 +164,9 @@ Global Instance map_cmra_discrete : CMRADiscrete A → CMRADiscrete mapR.
Proof. split; [apply _|]. intros m ? i. by apply: cmra_discrete_valid. Qed.
(** Internalized properties *)
Lemma map_equivI {M} m1 m2 : (m1 m2)%I ( i, m1 !! i m2 !! i : uPred M)%I.
Lemma map_equivI {M} m1 m2 : (m1 m2) ⊣⊢ ( i, m1 !! i m2 !! i : uPred M).
Proof. by uPred.unseal. Qed.
Lemma map_validI {M} m : ( m)%I ( i, (m !! i) : uPred M)%I.
Lemma map_validI {M} m : ( m) ⊣⊢ ( i, (m !! i) : uPred M).
Proof. by uPred.unseal. Qed.
End cmra.
......
......@@ -177,17 +177,17 @@ Proof. intros. by apply frac_validN_inv_l with 0 a, cmra_valid_validN. Qed.
(** Internalized properties *)
Lemma frac_equivI {M} (x y : frac A) :
(x y)%I (match x, y with
(x y) ⊣⊢ (match x, y with
| Frac q1 a, Frac q2 b => q1 = q2 a b
| FracUnit, FracUnit => True
| _, _ => False
end : uPred M)%I.
end : uPred M).
Proof.
uPred.unseal; do 2 split; first by destruct 1.
by destruct x, y; destruct 1; try constructor.
Qed.
Lemma frac_validI {M} (x : frac A) :
( x)%I (if x is Frac q a then (q 1)%Qc a else True : uPred M)%I.
( x) ⊣⊢ (if x is Frac q a then (q 1)%Qc a else True : uPred M).
Proof. uPred.unseal. by destruct x. Qed.
(** ** Local updates *)
......
......@@ -164,9 +164,9 @@ Section iprod_cmra.
Qed.
(** Internalized properties *)
Lemma iprod_equivI {M} g1 g2 : (g1 g2)%I ( i, g1 i g2 i : uPred M)%I.
Lemma iprod_equivI {M} g1 g2 : (g1 g2) ⊣⊢ ( i, g1 i g2 i : uPred M).
Proof. by uPred.unseal. Qed.
Lemma iprod_validI {M} g : ( g)%I ( i, g i : uPred M)%I.
Lemma iprod_validI {M} g : ( g) ⊣⊢ ( i, g i : uPred M).
Proof. by uPred.unseal. Qed.
(** Properties of iprod_insert. *)
......
......@@ -132,14 +132,14 @@ Proof. by destruct mx, my; inversion_clear 1. Qed.
(** Internalized properties *)
Lemma option_equivI {M} (x y : option A) :
(x y)%I (match x, y with
(x y) ⊣⊢ (match x, y with
| Some a, Some b => a b | None, None => True | _, _ => False
end : uPred M)%I.
end : uPred M).
Proof.
uPred.unseal. do 2 split. by destruct 1. by destruct x, y; try constructor.
Qed.
Lemma option_validI {M} (x : option A) :
( x)%I (match x with Some a => a | None => True end : uPred M)%I.
( x) ⊣⊢ (match x with Some a => a | None => True end : uPred M).
Proof. uPred.unseal. by destruct x. Qed.
(** Updates *)
......
This diff is collapsed.
......@@ -39,9 +39,9 @@ Implicit Types Ps Qs : list (uPred M).
Implicit Types A : Type.
(* Big ops *)
Global Instance big_and_proper : Proper (() ==> ()) (@uPred_big_and M).
Global Instance big_and_proper : Proper (() ==> (⊣⊢)) (@uPred_big_and M).
Proof. by induction 1 as [|P Q Ps Qs HPQ ? IH]; rewrite /= ?HPQ ?IH. Qed.
Global Instance big_sep_proper : Proper (() ==> ()) (@uPred_big_sep M).
Global Instance big_sep_proper : Proper (() ==> (⊣⊢)) (@uPred_big_sep M).
Proof. by induction 1 as [|P Q Ps Qs HPQ ? IH]; rewrite /= ?HPQ ?IH. Qed.
Global Instance big_and_ne n :
......@@ -51,19 +51,19 @@ Global Instance big_sep_ne n :
Proper (Forall2 (dist n) ==> dist n) (@uPred_big_sep M).
Proof. by induction 1 as [|P Q Ps Qs HPQ ? IH]; rewrite /= ?HPQ ?IH. Qed.
Global Instance big_and_mono' : Proper (Forall2 () ==> ()) (@uPred_big_and M).
Global Instance big_and_mono' : Proper (Forall2 () ==> ()) (@uPred_big_and M).
Proof. by induction 1 as [|P Q Ps Qs HPQ ? IH]; rewrite /= ?HPQ ?IH. Qed.
Global Instance big_sep_mono' : Proper (Forall2 () ==> ()) (@uPred_big_sep M).
Global Instance big_sep_mono' : Proper (Forall2 () ==> ()) (@uPred_big_sep M).
Proof. by induction 1 as [|P Q Ps Qs HPQ ? IH]; rewrite /= ?HPQ ?IH. Qed.
Global Instance big_and_perm : Proper (() ==> ()) (@uPred_big_and M).
Global Instance big_and_perm : Proper (() ==> (⊣⊢)) (@uPred_big_and M).
Proof.
induction 1 as [|P Ps Qs ? IH|P Q Ps|]; simpl; auto.
- by rewrite IH.
- by rewrite !assoc (comm _ P).
- etrans; eauto.
Qed.
Global Instance big_sep_perm : Proper (() ==> ()) (@uPred_big_sep M).
Global Instance big_sep_perm : Proper (() ==> (⊣⊢)) (@uPred_big_sep M).
Proof.
induction 1 as [|P Ps Qs ? IH|P Q Ps|]; simpl; auto.
- by rewrite IH.
......@@ -71,26 +71,26 @@ Proof.
- etrans; eauto.
Qed.
Lemma big_and_app Ps Qs : (Π (Ps ++ Qs))%I (Π Ps Π Qs)%I.
Lemma big_and_app Ps Qs : (Π (Ps ++ Qs)) ⊣⊢ (Π Ps Π Qs).
Proof. by induction Ps as [|?? IH]; rewrite /= ?left_id -?assoc ?IH. Qed.
Lemma big_sep_app Ps Qs : (Π (Ps ++ Qs))%I (Π Ps Π Qs)%I.
Lemma big_sep_app Ps Qs : (Π (Ps ++ Qs)) ⊣⊢ (Π Ps Π Qs).
Proof. by induction Ps as [|?? IH]; rewrite /= ?left_id -?assoc ?IH. Qed.
Lemma big_and_contains Ps Qs : Qs `contains` Ps (Π Ps)%I (Π Qs)%I.
Lemma big_and_contains Ps Qs : Qs `contains` Ps (Π Ps) (Π Qs).
Proof.
intros [Ps' ->]%contains_Permutation. by rewrite big_and_app and_elim_l.
Qed.
Lemma big_sep_contains Ps Qs : Qs `contains` Ps (Π Ps)%I (Π Qs)%I.
Lemma big_sep_contains Ps Qs : Qs `contains` Ps (Π Ps) (Π Qs).
Proof.
intros [Ps' ->]%contains_Permutation. by rewrite big_sep_app sep_elim_l.
Qed.
Lemma big_sep_and Ps : (Π Ps) (Π Ps).
Lemma big_sep_and Ps : (Π Ps) (Π Ps).
Proof. by induction Ps as [|P Ps IH]; simpl; auto with I. Qed.
Lemma big_and_elem_of Ps P : P Ps (Π Ps) P.
Lemma big_and_elem_of Ps P : P Ps (Π Ps) P.
Proof. induction 1; simpl; auto with I. Qed.
Lemma big_sep_elem_of Ps P : P Ps (Π Ps) P.
Lemma big_sep_elem_of Ps P : P Ps (Π Ps) P.
Proof. induction 1; simpl; auto with I. Qed.
(* Big ops over finite maps *)
......@@ -100,8 +100,8 @@ Section gmap.
Implicit Types Φ Ψ : K A uPred M.
Lemma big_sepM_mono Φ Ψ m1 m2 :
m2 m1 ( x k, m2 !! k = Some x Φ k x Ψ k x)
(Π{map m1} Φ) (Π{map m2} Ψ).
m2 m1 ( x k, m2 !! k = Some x Φ k x Ψ k x)
(Π{map m1} Φ) (Π{map m2} Ψ).
Proof.
intros HX . trans (Π{map m2} Φ)%I.
- by apply big_sep_contains, fmap_contains, map_to_list_contains.
......@@ -117,36 +117,36 @@ Section gmap.
apply Forall2_Forall, Forall_true=> -[i x]; apply .
Qed.
Global Instance big_sepM_proper m :
Proper (pointwise_relation _ (pointwise_relation _ ()) ==> ())
Proper (pointwise_relation _ (pointwise_relation _ (⊣⊢)) ==> (⊣⊢))
(uPred_big_sepM (M:=M) m).
Proof.
intros Φ1 Φ2 ; apply equiv_dist=> n.
apply big_sepM_ne=> k x; apply equiv_dist, .
Qed.
Global Instance big_sepM_mono' m :
Proper (pointwise_relation _ (pointwise_relation _ ()) ==> ())
Proper (pointwise_relation _ (pointwise_relation _ ()) ==> ())
(uPred_big_sepM (M:=M) m).
Proof. intros Φ1 Φ2 . apply big_sepM_mono; intros; [done|apply ]. Qed.
Lemma big_sepM_empty Φ : (Π{map } Φ)%I True%I.
Lemma big_sepM_empty Φ : (Π{map } Φ) ⊣⊢ True.
Proof. by rewrite /uPred_big_sepM map_to_list_empty. Qed.
Lemma big_sepM_insert Φ (m : gmap K A) i x :
m !! i = None (Π{map <[i:=x]> m} Φ)%I (Φ i x Π{map m} Φ)%I.
m !! i = None (Π{map <[i:=x]> m} Φ) ⊣⊢ (Φ i x Π{map m} Φ).
Proof. intros ?; by rewrite /uPred_big_sepM map_to_list_insert. Qed.
Lemma big_sepM_singleton Φ i x : (Π{map {[i := x]}} Φ)%I (Φ i x)%I.
Lemma big_sepM_singleton Φ i x : (Π{map {[i := x]}} Φ) ⊣⊢ (Φ i x).
Proof.
rewrite -insert_empty big_sepM_insert/=; last auto using lookup_empty.
by rewrite big_sepM_empty right_id.
Qed.
Lemma big_sepM_sepM Φ Ψ m :
(Π{map m} (λ i x, Φ i x Ψ i x))%I (Π{map m} Φ Π{map m} Ψ)%I.
(Π{map m} (λ i x, Φ i x Ψ i x)) ⊣⊢ (Π{map m} Φ Π{map m} Ψ).
Proof.
rewrite /uPred_big_sepM.
induction (map_to_list m) as [|[i x] l IH]; csimpl; rewrite ?right_id //.
by rewrite IH -!assoc (assoc _ (Ψ _ _)) [(Ψ _ _ _)%I]comm -!assoc.
Qed.
Lemma big_sepM_later Φ m : ( Π{map m} Φ)%I (Π{map m} (λ i x, Φ i x))%I.
Lemma big_sepM_later Φ m : ( Π{map m} Φ) ⊣⊢ (Π{map m} (λ i x, Φ i x)).
Proof.
rewrite /uPred_big_sepM.
induction (map_to_list m) as [|[i x] l IH]; csimpl; rewrite ?later_True //.
......@@ -161,7 +161,7 @@ Section gset.
Implicit Types Φ : A uPred M.
Lemma big_sepS_mono Φ Ψ X Y :
Y X ( x, x Y Φ x Ψ x) (Π{set X} Φ) (Π{set Y} Ψ).
Y X ( x, x Y Φ x Ψ x) (Π{set X} Φ) (Π{set Y} Ψ).
Proof.
intros HX . trans (Π{set Y} Φ)%I.
- by apply big_sep_contains, fmap_contains, elements_contains.
......@@ -176,38 +176,38 @@ Section gset.
apply Forall2_Forall, Forall_true=> x; apply .
Qed.
Lemma big_sepS_proper X :
Proper (pointwise_relation _ () ==> ()) (uPred_big_sepS (M:=M) X).
Proper (pointwise_relation _ (⊣⊢) ==> (⊣⊢)) (uPred_big_sepS (M:=M) X).
Proof.
intros Φ1 Φ2 ; apply equiv_dist=> n.
apply big_sepS_ne=> x; apply equiv_dist, .
Qed.
Lemma big_sepS_mono' X :
Proper (pointwise_relation _ () ==> ()) (uPred_big_sepS (M:=M) X).
Proper (pointwise_relation _ () ==> ()) (uPred_big_sepS (M:=M) X).
Proof. intros Φ1 Φ2 . apply big_sepS_mono; naive_solver. Qed.
Lemma big_sepS_empty Φ : (Π{set } Φ)%I True%I.
Lemma big_sepS_empty Φ : (Π{set } Φ) ⊣⊢ True.
Proof. by rewrite /uPred_big_sepS elements_empty. Qed.
Lemma big_sepS_insert Φ X x :
x X (Π{set {[ x ]} X} Φ)%I (Φ x Π{set X} Φ)%I.
x X (Π{set {[ x ]} X} Φ) ⊣⊢ (Φ x Π{set X} Φ).
Proof. intros. by rewrite /uPred_big_sepS elements_union_singleton. Qed.
Lemma big_sepS_delete Φ X x :
x X (Π{set X} Φ)%I (Φ x Π{set X {[ x ]}} Φ)%I.
x X (Π{set X} Φ) ⊣⊢ (Φ x Π{set X {[ x ]}} Φ).
Proof.
intros. rewrite -big_sepS_insert; last set_solver.
by rewrite -union_difference_L; last set_solver.
Qed.
Lemma big_sepS_singleton Φ x : (Π{set {[ x ]}} Φ)%I (Φ x)%I.
Lemma big_sepS_singleton Φ x : (Π{set {[ x ]}} Φ) ⊣⊢ (Φ x).
Proof. intros. by rewrite /uPred_big_sepS elements_singleton /= right_id. Qed.
Lemma big_sepS_sepS Φ Ψ X :
(Π{set X} (λ x, Φ x Ψ x))%I (Π{set X} Φ Π{set X} Ψ)%I.
(Π{set X} (λ x, Φ x Ψ x)) ⊣⊢ (Π{set X} Φ Π{set X} Ψ).
Proof.
rewrite /uPred_big_sepS.
induction (elements X) as [|x l IH]; csimpl; first by rewrite ?right_id.
by rewrite IH -!assoc (assoc _ (Ψ _)) [(Ψ _ _)%I]comm -!assoc.
Qed.
Lemma big_sepS_later Φ X : ( Π{set X} Φ)%I (Π{set X} (λ x, Φ x))%I.
Lemma big_sepS_later Φ X : ( Π{set X} Φ) ⊣⊢ (Π{set X} (λ x, Φ x)).
Proof.
rewrite /uPred_big_sepS.
induction (elements X) as [|x l IH]; csimpl; first by rewrite ?later_True.
......
......@@ -24,18 +24,18 @@ Module uPred_reflection. Section uPred_reflection.
Notation eval_list Σ l :=
(uPred_big_sep ((λ n, from_option True%I (Σ !! n)) <$> l)).
Lemma eval_flatten Σ e : eval Σ e eval_list Σ (flatten e).
Lemma eval_flatten Σ e : eval Σ e ⊣⊢ eval_list Σ (flatten e).
Proof.
induction e as [| |e1 IH1 e2 IH2];
rewrite /= ?right_id ?fmap_app ?big_sep_app ?IH1 ?IH2 //.
Qed.
Lemma flatten_entails Σ e1 e2 :
flatten e2 `contains` flatten e1 eval Σ e1 eval Σ e2.
flatten e2 `contains` flatten e1 eval Σ e1 eval Σ e2.
Proof.
intros. rewrite !eval_flatten. by apply big_sep_contains, fmap_contains.
Qed.
Lemma flatten_equiv Σ e1 e2 :
flatten e2 flatten e1 eval Σ e1 eval Σ e2.
flatten e2 flatten e1 eval Σ e1 ⊣⊢ eval Σ e2.
Proof. intros He. by rewrite !eval_flatten He. Qed.
Fixpoint prune (e : expr) : expr :=
......@@ -54,7 +54,7 @@ Module uPred_reflection. Section uPred_reflection.
induction e as [| |e1 IH1 e2 IH2]; simplify_eq/=; auto.
rewrite -IH1 -IH2. by repeat case_match; rewrite ?right_id_L.
Qed.
Lemma prune_correct Σ e : eval Σ (prune e) eval Σ e.
Lemma prune_correct Σ e : eval Σ (prune e) ⊣⊢ eval Σ e.
Proof. by rewrite !eval_flatten flatten_prune. Qed.
Fixpoint cancel_go (n : nat) (e : expr) : option expr :=
......@@ -86,7 +86,7 @@ Module uPred_reflection. Section uPred_reflection.
Qed.
Lemma cancel_entails Σ e1 e2 e1' e2' ns :
cancel ns e1 = Some e1' cancel ns e2 = Some e2'
eval Σ e1' eval Σ e2' eval Σ e1 eval Σ e2.
eval Σ e1' eval Σ e2' eval Σ e1 eval Σ e2.
Proof.
intros ??. rewrite !eval_flatten.
rewrite (flatten_cancel e1 e1' ns) // (flatten_cancel e2 e2' ns) //; csimpl.
......@@ -100,20 +100,20 @@ Module uPred_reflection. Section uPred_reflection.
| n :: l => ESep (EVar n) (to_expr l)
end.
Arguments to_expr !_ / : simpl nomatch.
Lemma eval_to_expr Σ l : eval Σ (to_expr l) eval_list Σ l.
Lemma eval_to_expr Σ l : eval Σ (to_expr l) ⊣⊢ eval_list Σ l.
Proof.
induction l as [|n1 [|n2 l] IH]; csimpl; rewrite ?right_id //.
by rewrite IH.
Qed.
Lemma split_l Σ e ns e' :
cancel ns e = Some e' eval Σ e (eval Σ (to_expr ns) eval Σ e')%I.
cancel ns e = Some e' eval Σ e ⊣⊢ (eval Σ (to_expr ns) eval Σ e').
Proof.
intros He%flatten_cancel.
by rewrite eval_flatten He fmap_app big_sep_app eval_to_expr eval_flatten.
Qed.
Lemma split_r Σ e ns e' :
cancel ns e = Some e' eval Σ e (eval Σ e' eval Σ (to_expr ns))%I.
cancel ns e = Some e' eval Σ e ⊣⊢ (eval Σ e' eval Σ (to_expr ns)).
Proof. intros. rewrite /= comm. by apply split_l. Qed.
Class Quote (Σ1 Σ2 : list (uPred M)) (P : uPred M) (e : expr) := {}.
......@@ -132,16 +132,16 @@ Module uPred_reflection. Section uPred_reflection.
Ltac quote :=
match goal with
| |- ?P1 ?P2 =>
| |- ?P1 ?P2 =>
lazymatch type of (_ : Quote [] _ P1 _) with Quote _ ?Σ2 _ ?e1 =>
lazymatch type of (_ : Quote Σ2 _ P2 _) with Quote _ ?Σ3 _ ?e2 =>
change (eval Σ3 e1 eval Σ3 e2) end end
change (eval Σ3 e1 eval Σ3 e2) end end
end.
Ltac quote_l :=
match goal with
| |- ?P1 ?P2 =>
| |- ?P1 ?P2 =>
lazymatch type of (_ : Quote [] _ P1 _) with Quote _ ?Σ2 _ ?e1 =>
change (eval Σ2 e1 P2) end
change (eval Σ2 e1 P2) end
end.
End uPred_reflection.
......@@ -162,7 +162,7 @@ Ltac close_uPreds Ps tac :=
Tactic Notation "cancel" constr(Ps) :=
uPred_reflection.quote;
let Σ := match goal with |- uPred_reflection.eval _ _ => Σ end in
let Σ := match goal with |- uPred_reflection.eval _ _ => Σ end in
let ns' := lazymatch type of (_ : uPred_reflection.QuoteArgs Σ Ps _) with
| uPred_reflection.QuoteArgs _ _ ?ns' => ns'
end in
......@@ -172,12 +172,12 @@ Tactic Notation "cancel" constr(Ps) :=
Tactic Notation "ecancel" open_constr(Ps) :=
close_uPreds Ps ltac:(fun Qs => cancel Qs).
(** [to_front [P1, P2, ..]] rewrites in the premise of such that
(** [to_front [P1, P2, ..]] rewrites in the premise of such that
the assumptions P1, P2, ... appear at the front, in that order. *)
Tactic Notation "to_front" open_constr(Ps) :=
close_uPreds Ps ltac:(fun Ps =>
uPred_reflection.quote_l;
let Σ := match goal with |- uPred_reflection.eval _ _ => Σ end in
let Σ := match goal with |- uPred_reflection.eval _ _ => Σ end in
let ns' := lazymatch type of (_ : uPred_reflection.QuoteArgs Σ Ps _) with
| uPred_reflection.QuoteArgs _ _ ?ns' => ns'
end in
......@@ -188,7 +188,7 @@ Tactic Notation "to_front" open_constr(Ps) :=
Tactic Notation "to_back" open_constr(Ps) :=
close_uPreds Ps ltac:(fun Ps =>
uPred_reflection.quote_l;
let Σ := match goal with |- uPred_reflection.eval _ _ => Σ end in
let Σ := match goal with |- uPred_reflection.eval _ _ => Σ end in
let ns' := lazymatch type of (_ : uPred_reflection.QuoteArgs Σ Ps _) with
| uPred_reflection.QuoteArgs _ _ ?ns' => ns'
end in
......@@ -205,46 +205,46 @@ Tactic Notation "sep_split" "right:" open_constr(Ps) :=
Tactic Notation "sep_split" "left:" open_constr(Ps) :=
to_front Ps; apply sep_mono.
(** Assumes a goal of the shape P ▷ Q. Alterantively, if Q
(** Assumes a goal of the shape P ▷ Q. Alterantively, if Q
is built of ★, ∧, ∨ with ▷ in all branches; that will work, too.
Will turn this goal into P Q and strip ▷ in P below ★, ∧, ∨. *)
Will turn this goal into P Q and strip ▷ in P below ★, ∧, ∨. *)
Ltac strip_later :=
let rec strip :=
lazymatch goal with
| |- (_ _) _ =>
| |- (_ _) _ =>
etrans; last (eapply equiv_entails_sym, later_sep);
apply sep_mono; strip
| |- (_ _) _ =>
| |- (_ _) _ =>
etrans; last (eapply equiv_entails_sym, later_and);
apply sep_mono; strip
| |- (_ _) _ =>
| |- (_ _) _ =>
etrans; last (eapply equiv_entails_sym, later_or);
apply sep_mono; strip
| |- _ _ => apply later_mono; reflexivity
| |- _ _ => apply later_intro; reflexivity
| |- _ _ => apply later_mono; reflexivity
| |- _ _ => apply later_intro; reflexivity
end
in let rec shape_Q :=
lazymatch goal with
| |- _ (_ _) =>
| |- _ (_ _) =>
(* Force the later on the LHS to be top-level, matching laters
below ★ on the RHS *)
etrans; first (apply equiv_entails, later_sep; reflexivity);
(* Match the arm recursively *)
apply sep_mono; shape_Q
| |- _ (_ _) =>
| |- _ (_ _) =>
etrans; first (apply equiv_entails, later_and; reflexivity);
apply sep_mono; shape_Q
| |- _ (_ _) =>
| |- _ (_ _) =>
etrans; first (apply equiv_entails, later_or; reflexivity);
apply sep_mono; shape_Q
| |- _ _ => apply later_mono; reflexivity
| |- _ _ => apply later_mono; reflexivity
(* We fail if we don't find laters in all branches. *)
end
in intros_revert ltac:(etrans; [|shape_Q];
etrans; last eapply later_mono; first solve [ strip ]).
(** Transforms a goal of the form ∀ ..., ?0... → ?1 ?2
into True ∀..., ■?0... → ?1 → ?2, applies tac, and
(** Transforms a goal of the form ∀ ..., ?0... → ?1 ?2
into True ∀..., ■?0... → ?1 → ?2, applies tac, and
the moves all the assumptions back. *)
(* TODO: this name may be a big too general *)
Ltac revert_all :=
......@@ -256,20 +256,20 @@ Ltac revert_all :=
on the sort of the argument, which is suboptimal. *)
first [ apply (const_intro_impl _ _ _ H); clear H
| revert H; apply forall_elim']
| |- _ _ => apply impl_entails
| |- _ _ => apply impl_entails
end.
(** This starts on a goal of the form ∀ ..., ?0... → ?1 ?2.
(** This starts on a goal of the form ∀ ..., ?0... → ?1 ?2.
It applies löb where all the Coq assumptions have been turned into logical
assumptions, then moves all the Coq assumptions back out to the context,
applies [tac] on the goal (now of the form _ _), and then reverts the
applies [tac] on the goal (now of the form _ _), and then reverts the
Coq assumption so that we end up with the same shape as where we started,
but with an additional assumption ★-ed to the context *)
Ltac löb tac :=
revert_all;
(* Add a box *)
etrans; last (eapply always_elim; reflexivity);
(* We now have a goal for the form True P, with the "original" conclusion
(* We now have a goal for the form True P, with the "original" conclusion
being locked. *)
apply löb_strong; etransitivity;
first (apply equiv_entails, left_id, _; reflexivity);
......@@ -278,13 +278,13 @@ Ltac löb tac :=
do the work *)
let rec go :=
lazymatch goal with
| |- _ ( _, _) =>
| |- _ ( _, _) =>
apply forall_intro; let H := fresh in intro H; go; revert H
| |- _ ( _ _) =>
| |- _ ( _ _) =>
apply impl_intro_l, const_elim_l; let H := fresh in intro H; go; revert H
(* This is the "bottom" of the goal, where we see the impl introduced
by uPred_revert_all as well as the ▷ from löb_strong and the □ we added. *)
| |- ?R (?L _) => apply impl_intro_l;
| |- ?R (?L _) => apply impl_intro_l;
trans (L R)%I;
[eapply equiv_entails, always_and_sep_r, _; reflexivity | tac]
end
......
......@@ -16,19 +16,19 @@ Section client.
Local Notation iProp := (iPropG heap_lang Σ).
Definition y_inv q y : iProp :=
( f : val, y {q} f n : Z, #> f #n {{ λ v, v = #(n + 42) }})%I.
( f : val, y {q} f n : Z, WP f #n {{ λ v, v = #(n + 42) }})%I.
Lemma y_inv_split q y :
y_inv q y (y_inv (q/2) y y_inv (q/2) y).
y_inv q y (y_inv (q/2) y y_inv (q/2) y).
Proof.
rewrite /y_inv. apply exist_elim=>f.
rewrite -!(exist_intro f). rewrite heap_mapsto_op_split.
ecancel [y {_} _; y {_} _]%I. by rewrite [X in X _]always_sep_dup.
ecancel [y {_} _; y {_} _]%I. by rewrite [X in X _]always_sep_dup.
Qed.
Lemma worker_safe q (n : Z) (b y : loc) :
(heap_ctx heapN recv heapN N b (y_inv q y))
#> worker n (%b) (%y) {{ λ _, True }}.
WP worker n (%b) (%y) {{ λ _, True }}.
Proof.
rewrite /worker. wp_lam. wp_let. ewp apply wait_spec.
rewrite comm. apply sep_mono_r. apply wand_intro_l.
......@@ -42,7 +42,7 @@ Section client.
Qed.
Lemma client_safe :
heapN N heap_ctx heapN #> client {{ λ _, True }}.
heapN N heap_ctx heapN WP client {{ λ _, True }}.
Proof.
intros ?. rewrite /client.
(ewp eapply wp_alloc); eauto with I. strip_later. apply forall_intro=>y.
......
......@@ -82,7 +82,7 @@ Lemma ress_split i i1 i2 Q R1 R2 P I :
(saved_prop_own i2 R2
saved_prop_own i1 R1 saved_prop_own i Q
(Q -★ R1 R2) ress P I)
ress P ({[i1]} ({[i2]} (I {[i]}))).
ress P ({[i1]} ({[i2]} (I {[i]}))).
Proof.
intros. rewrite /ress !sep_exist_l. apply exist_elim=>Ψ.
rewrite -(exist_intro (<[i1:=R1]> (<[i2:=R2]> Ψ))).
......@@ -112,7 +112,7 @@ Qed.
Lemma newbarrier_spec (P : iProp) (Φ : val iProp) :
heapN N
(heap_ctx heapN l, recv l P send l P -★ Φ (%l))
#> newbarrier #() {{ Φ }}.
WP newbarrier #() {{ Φ }}.
Proof.
intros HN. rewrite /newbarrier. wp_seq.
rewrite -wp_pvs. wp eapply wp_alloc; eauto with I ndisj.
......@@ -151,7 +151,7 @@ Proof.
Qed.
Lemma signal_spec l P (Φ : val iProp) :
(send l P P Φ #()) #> signal (%l) {{ Φ }}.
(send l P P Φ #()) WP signal (%l) {{ Φ }}.
Proof.
rewrite /signal /send /barrier_ctx. rewrite sep_exist_r.
apply exist_elim=>γ. rewrite -!assoc. apply const_elim_sep_l=>?. wp_let.
......@@ -176,7 +176,7 @@ Proof.
Qed.
Lemma wait_spec l P (Φ : val iProp) :
(recv l P (P -★ Φ #())) #> wait (%l) {{ Φ }}.
(recv l P (P -★ Φ #())) WP wait (%l) {{ Φ }}.
Proof.
rename P into R. wp_rec.
rewrite {1}/recv /barrier_ctx. rewrite !sep_exist_r.
......@@ -226,7 +226,7 @@ Qed.
Lemma recv_split E l P1 P2 :
nclose N E
recv l (P1 P2) |={E}=> recv l P1 recv l P2.
recv l (P1 P2) |={E}=> recv l P1 recv l P2.
Proof.
rename P1 into R1. rename P2 into R2. intros HN.
rewrite {1}/recv /barrier_ctx.
......@@ -277,7 +277,7 @@ Proof.
Qed.
Lemma recv_weaken l P1 P2 :
(P1 -★ P2) (recv l P1 -★ recv l P2).
(P1 -★ P2) (recv l P1 -★ recv l P2).
Proof.
apply wand_intro_l. rewrite /recv. rewrite sep_exist_r. apply exist_mono=>γ.
rewrite sep_exist_r. apply exist_mono=>P. rewrite sep_exist_r.
......@@ -287,7 +287,7 @@ Proof.
Qed.
Lemma recv_mono l P1 P2 :
P1 P2 recv l P1 recv l P2.
P1 P2 recv l P1 recv l P2.
Proof.
intros HP%entails_wand. apply wand_entails. rewrite HP. apply recv_weaken.
Qed.
......
......@@ -12,12 +12,12 @@ Local Notation iProp := (iPropG heap_lang Σ).
Lemma barrier_spec (heapN N : namespace) :
heapN N
recv send : loc iProp -n> iProp,
( P, heap_ctx heapN {{ True }} newbarrier #() {{ λ v,
( P, heap_ctx heapN {{ True }} newbarrier #() {{ λ v,
l : loc, v = LocV l recv l P send l P }})
( l P, {{ send l P P }} signal (LocV l) {{ λ _, True }})
( l P, {{ recv l P }} wait (LocV l) {{ λ _, P }})
( l P Q, recv l (P Q) ={N}=> recv l P recv l Q)
( l P Q, (P -★ Q) (recv l P -★ recv l Q)).
( l P Q, (P -★ Q) (recv l P -★ recv l Q)).
Proof.
intros HN.
exists (λ l, CofeMor (recv heapN N l)), (λ l, CofeMor (send heapN N l)).
......
......@@ -51,7 +51,7 @@ Note that $\COFEs$ is cartesian closed.
\begin{defn}
A \emph{CMRA} is a tuple $(\monoid : \COFEs, (\mval_n \subseteq \monoid)_{n \in \mathbb{N}}, \mcore{-}: \monoid \nfn \monoid, (\mtimes) : \monoid \times \monoid \nfn \monoid, (\mdiv) : \monoid \times \monoid \nfn \monoid)$ satisfying
\begin{align*}
\All n, \melt, \meltB.& \melt \nequiv{n} \meltB \land \melt\in\mval_n \Ra \meltB\in\mval_n \tagH{cmra-valid-nonexp} \\
\All n, \melt, \meltB.& \melt \nequiv{n} \meltB \land \melt\in\mval_n \Ra \meltB\in\mval_n \tagH{cmra-valid-ne} \\
\All n, m.& n \geq m \Ra V_n \subseteq V_m \tagH{cmra-valid-mono} \\
\All \melt, \meltB, \meltC.& (\melt \mtimes \meltB) \mtimes \meltC = \melt \mtimes (\meltB \mtimes \meltC) \tagH{cmra-assoc} \\
\All \melt, \meltB.& \melt \mtimes \meltB = \meltB \mtimes \melt \tagH{cmra-comm} \\
......@@ -67,30 +67,16 @@ Note that $\COFEs$ is cartesian closed.
\end{align*}
\end{defn}
\ralf{Copy the rest of the explanation from the paper, when that one is more polished.}
\paragraph{The division operator $\mdiv$.}
One way to describe $\mdiv$ is to say that it extracts the witness from the extension order: If $\melt \leq \meltB$, then $\melt \mdiv \meltB$ computes the difference between the two elements (\ruleref{cmra-div-op}).
Otherwise, $\mdiv$ can have arbitrary behavior.
This means that, in classical logic, the division operator can be defined for any PCM using the axiom of choice, and it will trivially satisfy \ruleref{cmra-div-op}.
However, notice that the division operator also has to be \emph{non-expansive} --- so if the carrier $\monoid$ is equipped with a non-trivial $\nequiv{n}$, there is an additional proof obligation here.
This is crucial, for the following reason:
Considering that the extension order is defined using \emph{equality}, there is a natural notion of a \emph{step-indexed extension} order using the step-indexed equivalence of the underlying COFE:
\[ \melt \mincl{n} \meltB \eqdef \Exists \meltC. \meltB \nequiv{n} \melt \mtimes \meltC \tagH{cmra-inclM} \]
One of the properties we would expect to hold is the usual correspondence between a step-indexed predicate and its non-step-indexed counterpart:
\[ \All \melt, \meltB. \melt \leq \meltB \Lra (\All n. \melt \mincl{n} \meltB) \tagH{cmra-incl-limit} \]
The right-to-left direction here is trick.
For every $n$, we obtain a proof that $\melt \mincl{n} \meltB$.
From this, we could extract a sequence of witnesses $(\meltC_m)_{m}$, and we need to arrive at a single witness $\meltC$ showing that $\melt \leq \meltB$.
Without the division operator, there is no reason to believe that such a witness exists.
However, since we can use the division operator, and since we know that this operator is \emph{non-expansive}, we can pick $\meltC \eqdef \meltB \mdiv \melt$, and then we can prove that this is indeed the desired witness.
\ralf{The only reason we actually have division is that we are working constructively \emph{and}, at the same time, remain compatible with a classic interpretation of the existential. This is pretty silly.}
This is a natural generalization of RAs over COFEs.
All operations have to be non-expansive, and the validity predicate $\mval$ can now also depend on the step-index.
\ralf{TODO: Get rid of division.}
\paragraph{The extension axiom (\ruleref{cmra-extend}).}
Notice that the existential quantification in this axiom is \emph{constructive}, \ie it is a sigma type in Coq.
The purpose of this axiom is to compute $\melt_1$, $\melt_2$ completing the following square:
\ralf{Needs some magic to fix the baseline of the $\nequiv{n}$, or so}
% RJ FIXME: Needs some magic to fix the baseline of the $\nequiv{n}$, or so
\begin{center}
\begin{tikzpicture}[every edge/.style={draw=none}]
\node (a) at (0, 0) {$\melt$};
......
......@@ -204,7 +204,7 @@ We can derive some specialized forms of the lifting axioms for the operational s
\ralf{Add these.}
\subsection{Global Functor and ghost ownership}
\subsection{Global functor and ghost ownership}
\ralf{Describe this.}
% \subsection{Global monoid}
......@@ -364,7 +364,7 @@ We can derive some specialized forms of the lifting axioms for the operational s
% \subsection{Ghost heap}
% \label{sec:ghostheap}%
% FIXME use the finmap provided by the global ghost ownership, instead of adding our own
% We define a simple ghost heap with fractional permissions.
% Some modules require a few ghost names per module instance to properly manage ghost state, but would like to expose to clients a single logical name (avoiding clutter).
% In such cases we use these ghost heaps.
......
......@@ -19,54 +19,54 @@ Implicit Types Φ : val → iProp heap_lang Σ.
(** Proof rules for the sugar *)
Lemma wp_lam E x ef e v Φ :
to_val e = Some v
#> subst' x e ef @ E {{ Φ }} #> App (Lam x ef) e @ E {{ Φ }}.
WP subst' x e ef @ E {{ Φ }} WP App (Lam x ef) e @ E {{ Φ }}.
Proof. intros. by rewrite -wp_rec. Qed.
Lemma wp_let E x e1 e2 v Φ :
to_val e1 = Some v
#> subst' x e1 e2 @ E {{ Φ }} #> Let x e1 e2 @ E {{ Φ }}.
WP subst' x e1 e2 @ E {{ Φ }} WP Let x e1 e2 @ E {{ Φ }}.
Proof. apply wp_lam. Qed.
Lemma wp_seq E e1 e2 v Φ :
to_val e1 = Some v
#> e2 @ E {{ Φ }} #> Seq e1 e2 @ E {{ Φ }}.
WP e2 @ E {{ Φ }} WP Seq e1 e2 @ E {{ Φ }}.
Proof. intros ?. by rewrite -wp_let. Qed.
Lemma wp_skip E Φ : Φ (LitV LitUnit) #> Skip @ E {{ Φ }}.
Lemma wp_skip E Φ : Φ (LitV LitUnit) WP Skip @ E {{ Φ }}.
Proof. rewrite -wp_seq // -wp_value //. Qed.
Lemma wp_match_inl E e0 v0 x1 e1 x2 e2 Φ :
to_val e0 = Some v0
#> subst' x1 e0 e1 @ E {{ Φ }} #> Match (InjL e0) x1 e1 x2 e2 @ E {{ Φ }}.
Proof. intros. by rewrite -wp_case_inl // -[X in _ X]later_intro -wp_let. Qed.
WP subst' x1 e0 e1 @ E {{ Φ }} WP Match (InjL e0) x1 e1 x2 e2 @ E {{ Φ }}.
Proof. intros. by rewrite -wp_case_inl // -[X in _ X]later_intro -wp_let. Qed.
Lemma wp_match_inr E e0 v0 x1 e1 x2 e2 Φ :
to_val e0 = Some v0
#> subst' x2 e0 e2 @ E {{ Φ }} #> Match (InjR e0) x1 e1 x2 e2 @ E {{ Φ }}.
Proof. intros. by rewrite -wp_case_inr // -[X in _ X]later_intro -wp_let. Qed.
WP subst' x2 e0 e2 @ E {{ Φ }} WP Match (InjR e0) x1 e1 x2 e2 @ E {{ Φ }}.
Proof. intros. by rewrite -wp_case_inr // -[X in _ X]later_intro -wp_let. Qed.
Lemma wp_le E (n1 n2 : Z) P Φ :
(n1 n2 P Φ (LitV (LitBool true)))
(n2 < n1 P Φ (LitV (LitBool false)))
P #> BinOp LeOp (Lit (LitInt n1)) (Lit (LitInt n2)) @ E {{ Φ }}.
(n1 n2 P Φ (LitV (LitBool true)))
(n2 < n1 P Φ (LitV (LitBool false)))
P WP BinOp LeOp (Lit (LitInt n1)) (Lit (LitInt n2)) @ E {{ Φ }}.
Proof.
intros. rewrite -wp_bin_op //; [].
destruct (bool_decide_reflect (n1 n2)); by eauto with omega.
Qed.
Lemma wp_lt E (n1 n2 : Z) P Φ :
(n1 < n2 P Φ (LitV (LitBool true)))
(n2 n1 P Φ (LitV (LitBool false)))
P #> BinOp LtOp (Lit (LitInt n1)) (Lit (LitInt n2)) @ E {{ Φ }}.
(n1 < n2 P Φ (LitV (LitBool true)))
(n2 n1 P Φ (LitV (LitBool false)))
P WP BinOp LtOp (Lit (LitInt n1)) (Lit (LitInt n2)) @ E {{ Φ }}.
Proof.
intros. rewrite -wp_bin_op //; [].
destruct (bool_decide_reflect (n1 < n2)); by eauto with omega.
Qed.
Lemma wp_eq E (n1 n2 : Z) P Φ :
(n1 = n2 P Φ (LitV (LitBool true)))
(n1 n2 P Φ (LitV (LitBool false)))
P #> BinOp EqOp (Lit (LitInt n1)) (Lit (LitInt n2)) @ E {{ Φ }}.
(n1 = n2 P Φ (LitV (LitBool true)))
(n1 n2 P Φ (LitV (LitBool false)))
P WP BinOp EqOp (Lit (LitInt n1)) (Lit (LitInt n2)) @ E {{ Φ }}.
Proof.
intros. rewrite -wp_bin_op //; [].
destruct (bool_decide_reflect (n1 = n2)); by eauto with omega.
......
......@@ -31,7 +31,7 @@ Section definitions.
Definition heap_ctx (N : namespace) : iPropG heap_lang Σ :=
auth_ctx heap_name N heap_inv.
Global Instance heap_inv_proper : Proper (() ==> ()) heap_inv.
Global Instance heap_inv_proper : Proper (() ==> (⊣⊢)) heap_inv.
Proof. solve_proper. Qed.
Global Instance heap_ctx_always_stable N : AlwaysStable (heap_ctx N).
Proof. apply _. Qed.
......@@ -97,7 +97,7 @@ Section heap.
(** Allocation *)
Lemma heap_alloc N E σ :
authG heap_lang Σ heapR nclose N E
ownP σ (|={E}=> _ : heapG Σ, heap_ctx N Π{map σ} (λ l v, l v)).
ownP σ (|={E}=> _ : heapG Σ, heap_ctx N Π{map σ} (λ l v, l v)).
Proof.
intros. rewrite -{1}(from_to_heap σ). etrans.
{ rewrite [ownP _]later_intro.
......@@ -119,30 +119,30 @@ Section heap.
(** General properties of mapsto *)
Lemma heap_mapsto_op_eq l q1 q2 v :
(l {q1} v l {q2} v)%I (l {q1+q2} v)%I.
(l {q1} v l {q2} v) ⊣⊢ (l {q1+q2} v).
Proof. by rewrite -auth_own_op map_op_singleton Frac_op dec_agree_idemp. Qed.
Lemma heap_mapsto_op l q1 q2 v1 v2 :
(l {q1} v1 l {q2} v2)%I (v1 = v2 l {q1+q2} v1)%I.
(l {q1} v1 l {q2} v2) ⊣⊢ (v1 = v2 l {q1+q2} v1).
Proof.
destruct (decide (v1 = v2)) as [->|].
{ by rewrite heap_mapsto_op_eq const_equiv // left_id. }
rewrite -auth_own_op map_op_singleton Frac_op dec_agree_ne //.
apply (anti_symm ()); last by apply const_elim_l.
apply (anti_symm ()); last by apply const_elim_l.
rewrite auth_own_valid map_validI (forall_elim l) lookup_singleton.
rewrite option_validI frac_validI discrete_valid. by apply const_elim_r.
Qed.
Lemma heap_mapsto_op_split l q v :
(l {q} v)%I (l {q/2} v l {q/2} v)%I.
(l {q} v) ⊣⊢ (l {q/2} v l {q/2} v).
Proof. by rewrite heap_mapsto_op_eq Qp_div_2. Qed.
(** Weakest precondition *)
Lemma wp_alloc N E e v P Φ :
to_val e = Some v
P heap_ctx N nclose N E
P ( l, l v -★ Φ (LocV l))
P #> Alloc e @ E {{ Φ }}.
P heap_ctx N nclose N E
P ( l, l v -★ Φ (LocV l))
P WP Alloc e @ E {{ Φ }}.
Proof.
rewrite /heap_ctx /heap_inv=> ??? HP.
trans (|={E}=> auth_own heap_name P)%I.
......@@ -165,9 +165,9 @@ Section heap.
Qed.
Lemma wp_load N E l q v P Φ :
P heap_ctx N nclose N E
P ( l {q} v (l {q} v -★ Φ v))
P #> Load (Loc l) @ E {{ Φ }}.
P heap_ctx N nclose N E
P ( l {q} v (l {q} v -★ Φ v))
P WP Load (Loc l) @ E {{ Φ }}.
Proof.
rewrite /heap_ctx /heap_inv=> ?? HPΦ.
apply (auth_fsa' heap_inv (wp_fsa _) id)
......@@ -182,9 +182,9 @@ Section heap.
Lemma wp_store N E l v' e v P Φ :
to_val e = Some v
P heap_ctx N nclose N E
P ( l v' (l v -★ Φ (LitV LitUnit)))
P #> Store (Loc l) e @ E {{ Φ }}.
P heap_ctx N nclose N E
P ( l v' (l v -★ Φ (LitV LitUnit)))
P WP Store (Loc l) e @ E {{ Φ }}.
Proof.
rewrite /heap_ctx /heap_inv=> ??? HPΦ.
apply (auth_fsa' heap_inv (wp_fsa _) (alter (λ _, Frac 1 (DecAgree v)) l))
......@@ -199,9 +199,9 @@ Section heap.
Lemma wp_cas_fail N E l q v' e1 v1 e2 v2 P Φ :
to_val e1 = Some v1 to_val e2 = Some v2 v' v1
P heap_ctx N nclose N E
P ( l {q} v' (l {q} v' -★ Φ (LitV (LitBool false))))
P #> CAS (Loc l) e1 e2 @ E {{ Φ }}.
P heap_ctx N nclose N E
P ( l {q} v' (l {q} v' -★ Φ (LitV (LitBool false))))
P WP CAS (Loc l) e1 e2 @ E {{ Φ }}.
Proof.
rewrite /heap_ctx /heap_inv=>????? HPΦ.
apply (auth_fsa' heap_inv (wp_fsa _) id)
......@@ -216,9 +216,9 @@ Section heap.
Lemma wp_cas_suc N E l e1 v1 e2 v2 P Φ :
to_val e1 = Some v1 to_val e2 = Some v2
P heap_ctx N nclose N E
P ( l v1 (l v2 -★ Φ (LitV (LitBool true))))
P #> CAS (Loc l) e1 e2 @ E {{ Φ }}.
P heap_ctx N nclose N E
P ( l v1 (l v2 -★ Φ (LitV (LitBool true))))
P WP CAS (Loc l) e1 e2 @ E {{ Φ }}.
Proof.
rewrite /heap_ctx /heap_inv=> ???? HPΦ.
apply (auth_fsa' heap_inv (wp_fsa _) (alter (λ _, Frac 1 (DecAgree v2)) l))
......
......@@ -15,14 +15,14 @@ Implicit Types ef : option (expr []).
(** Bind. *)
Lemma wp_bind {E e} K Φ :
#> e @ E {{ λ v, #> fill K (of_val v) @ E {{ Φ }}}} #> fill K e @ E {{ Φ }}.
WP e @ E {{ λ v, WP fill K (of_val v) @ E {{ Φ }}}} WP fill K e @ E {{ Φ }}.
Proof. apply weakestpre.wp_bind. Qed.
(** Base axioms for core primitives of the language: Stateful reductions. *)
Lemma wp_alloc_pst E σ e v Φ :
to_val e = Some v
( ownP σ ( l, σ !! l = None ownP (<[l:=v]>σ) -★ Φ (LocV l)))
#> Alloc e @ E {{ Φ }}.
WP Alloc e @ E {{ Φ }}.
Proof.
(* TODO RJ: This works around ssreflect bug #22. *)
intros. set (φ v' σ' ef := l,
......@@ -39,7 +39,7 @@ Qed.
Lemma wp_load_pst E σ l v Φ :
σ !! l = Some v
( ownP σ (ownP σ -★ Φ v)) #> Load (Loc l) @ E {{ Φ }}.
( ownP σ (ownP σ -★ Φ v)) WP Load (Loc l) @ E {{ Φ }}.
Proof.
intros. rewrite -(wp_lift_atomic_det_step σ v σ None) ?right_id //;
last by intros; inv_step; eauto using to_of_val.
......@@ -48,7 +48,7 @@ Qed.
Lemma wp_store_pst E σ l e v v' Φ :
to_val e = Some v σ !! l = Some v'
( ownP σ (ownP (<[l:=v]>σ) -★ Φ (LitV LitUnit)))
#> Store (Loc l) e @ E {{ Φ }}.
WP Store (Loc l) e @ E {{ Φ }}.
Proof.
intros. rewrite -(wp_lift_atomic_det_step σ (LitV LitUnit) (<[l:=v]>σ) None)
?right_id //; last by intros; inv_step; eauto.
......@@ -57,7 +57,7 @@ Qed.
Lemma wp_cas_fail_pst E σ l e1 v1 e2 v2 v' Φ :
to_val e1 = Some v1 to_val e2 = Some v2 σ !! l = Some v' v' v1
( ownP σ (ownP σ -★ Φ (LitV $ LitBool false)))
#> CAS (Loc l) e1 e2 @ E {{ Φ }}.
WP CAS (Loc l) e1 e2 @ E {{ Φ }}.
Proof.
intros. rewrite -(wp_lift_atomic_det_step σ (LitV $ LitBool false) σ None)
?right_id //; last by intros; inv_step; eauto.
......@@ -66,7 +66,7 @@ Qed.
Lemma wp_cas_suc_pst E σ l e1 v1 e2 v2 Φ :
to_val e1 = Some v1 to_val e2 = Some v2 σ !! l = Some v1
( ownP σ (ownP (<[l:=v2]>σ) -★ Φ (LitV $ LitBool true)))
#> CAS (Loc l) e1 e2 @ E {{ Φ }}.
WP CAS (Loc l) e1 e2 @ E {{ Φ }}.
Proof.
intros. rewrite -(wp_lift_atomic_det_step σ (LitV $ LitBool true)
(<[l:=v2]>σ) None) ?right_id //; last by intros; inv_step; eauto.
......@@ -74,7 +74,7 @@ Qed.
(** Base axioms for core primitives of the language: Stateless reductions *)
Lemma wp_fork E e Φ :
( Φ (LitV LitUnit) #> e {{ λ _, True }}) #> Fork e @ E {{ Φ }}.
( Φ (LitV LitUnit) WP e {{ λ _, True }}) WP Fork e @ E {{ Φ }}.
Proof.
rewrite -(wp_lift_pure_det_step (Fork e) (Lit LitUnit) (Some e)) //=;
last by intros; inv_step; eauto.
......@@ -83,8 +83,8 @@ Qed.
Lemma wp_rec E f x e1 e2 v Φ :
to_val e2 = Some v
#> subst' x e2 (subst' f (Rec f x e1) e1) @ E {{ Φ }}
#> App (Rec f x e1) e2 @ E {{ Φ }}.
WP subst' x e2 (subst' f (Rec f x e1) e1) @ E {{ Φ }}
WP App (Rec f x e1) e2 @ E {{ Φ }}.
Proof.
intros. rewrite -(wp_lift_pure_det_step (App _ _)
(subst' x e2 (subst' f (Rec f x e1) e1)) None) //= ?right_id;
......@@ -94,13 +94,13 @@ Qed.
Lemma wp_rec' E f x erec e1 e2 v2 Φ :
e1 = Rec f x erec
to_val e2 = Some v2
#> subst' x e2 (subst' f e1 erec) @ E {{ Φ }}
#> App e1 e2 @ E {{ Φ }}.
WP subst' x e2 (subst' f e1 erec) @ E {{ Φ }}
WP App e1 e2 @ E {{ Φ }}.
Proof. intros ->. apply wp_rec. Qed.
Lemma wp_un_op E op l l' Φ :
un_op_eval op l = Some l'
Φ (LitV l') #> UnOp op (Lit l) @ E {{ Φ }}.
Φ (LitV l') WP UnOp op (Lit l) @ E {{ Φ }}.
Proof.
intros. rewrite -(wp_lift_pure_det_step (UnOp op _) (Lit l') None)
?right_id -?wp_value //; intros; inv_step; eauto.
......@@ -108,21 +108,21 @@ Qed.
Lemma wp_bin_op E op l1 l2 l' Φ :
bin_op_eval op l1 l2 = Some l'
Φ (LitV l') #> BinOp op (Lit l1) (Lit l2) @ E {{ Φ }}.
Φ (LitV l') WP BinOp op (Lit l1) (Lit l2) @ E {{ Φ }}.
Proof.
intros Heval. rewrite -(wp_lift_pure_det_step (BinOp op _ _) (Lit l') None)
?right_id -?wp_value //; intros; inv_step; eauto.
Qed.
Lemma wp_if_true E e1 e2 Φ :
#> e1 @ E {{ Φ }} #> If (Lit (LitBool true)) e1 e2 @ E {{ Φ }}.
WP e1 @ E {{ Φ }} WP If (Lit (LitBool true)) e1 e2 @ E {{ Φ }}.
Proof.
rewrite -(wp_lift_pure_det_step (If _ _ _) e1 None)
?right_id //; intros; inv_step; eauto.
Qed.
Lemma wp_if_false E e1 e2 Φ :
#> e2 @ E {{ Φ }} #> If (Lit (LitBool false)) e1 e2 @ E {{ Φ }}.
WP e2 @ E {{ Φ }} WP If (Lit (LitBool false)) e1 e2 @ E {{ Φ }}.
Proof.
rewrite -(wp_lift_pure_det_step (If _ _ _) e2 None)
?right_id //; intros; inv_step; eauto.
......@@ -130,7 +130,7 @@ Qed.
Lemma wp_fst E e1 v1 e2 v2 Φ :
to_val e1 = Some v1 to_val e2 = Some v2
Φ v1 #> Fst (Pair e1 e2) @ E {{ Φ }}.
Φ v1 WP Fst (Pair e1 e2) @ E {{ Φ }}.
Proof.
intros. rewrite -(wp_lift_pure_det_step (Fst _) e1 None)
?right_id -?wp_value //; intros; inv_step; eauto.
......@@ -138,7 +138,7 @@ Qed.
Lemma wp_snd E e1 v1 e2 v2 Φ :
to_val e1 = Some v1 to_val e2 = Some v2
Φ v2 #> Snd (Pair e1 e2) @ E {{ Φ }}.
Φ v2 WP Snd (Pair e1 e2) @ E {{ Φ }}.
Proof.
intros. rewrite -(wp_lift_pure_det_step (Snd _) e2 None)
?right_id -?wp_value //; intros; inv_step; eauto.
......@@ -146,7 +146,7 @@ Qed.
Lemma wp_case_inl E e0 v0 e1 e2 Φ :
to_val e0 = Some v0
#> App e1 e0 @ E {{ Φ }} #> Case (InjL e0) e1 e2 @ E {{ Φ }}.
WP App e1 e0 @ E {{ Φ }} WP Case (InjL e0) e1 e2 @ E {{ Φ }}.
Proof.
intros. rewrite -(wp_lift_pure_det_step (Case _ _ _)
(App e1 e0) None) ?right_id //; intros; inv_step; eauto.
......@@ -154,7 +154,7 @@ Qed.
Lemma wp_case_inr E e0 v0 e1 e2 Φ :
to_val e0 = Some v0
#> App e2 e0 @ E {{ Φ }} #> Case (InjR e0) e1 e2 @ E {{ Φ }}.
WP App e2 e0 @ E {{ Φ }} WP Case (InjR e0) e1 e2 @ E {{ Φ }}.
Proof.
intros. rewrite -(wp_lift_pure_det_step (Case _ _ _)
(App e2 e0) None) ?right_id //; intros; inv_step; eauto.
......
......@@ -2,12 +2,12 @@ From iris.heap_lang Require Export derived.
Export heap_lang.
Arguments wp {_ _} _ _%E _.
Notation "#> e @ E {{ Φ } }" := (wp E e%E Φ)
Notation "'WP' e @ E {{ Φ } }" := (wp E e%E Φ)
(at level 20, e, Φ at level 200,
format "#> e @ E {{ Φ } }") : uPred_scope.
Notation "#> e {{ Φ } }" := (wp e%E Φ)
format "'WP' e @ E {{ Φ } }") : uPred_scope.
Notation "'WP' e {{ Φ } }" := (wp e%E Φ)
(at level 20, e, Φ at level 200,
format "#> e {{ Φ } }") : uPred_scope.
format "'WP' e {{ Φ } }") : uPred_scope.
Coercion LitInt : Z >-> base_lit.
Coercion LitBool : bool >-> base_lit.
......
......@@ -21,9 +21,9 @@ Local Notation iProp := (iPropG heap_lang Σ).
Lemma par_spec (Ψ1 Ψ2 : val iProp) e (f1 f2 : val) (Φ : val iProp) :
heapN N to_val e = Some (f1,f2)%V
(heap_ctx heapN #> f1 #() {{ Ψ1 }} #> f2 #() {{ Ψ2 }}
(heap_ctx heapN WP f1 #() {{ Ψ1 }} WP f2 #() {{ Ψ2 }}
v1 v2, Ψ1 v1 Ψ2 v2 -★ Φ (v1,v2)%V)
#> par e {{ Φ }}.
WP par e {{ Φ }}.
Proof.
intros. rewrite /par. ewp (by eapply wp_value). wp_let. wp_proj.
ewp (eapply spawn_spec; wp_done).
......@@ -38,9 +38,9 @@ Qed.
Lemma wp_par (Ψ1 Ψ2 : val iProp) (e1 e2 : expr []) (Φ : val iProp) :
heapN N
(heap_ctx heapN #> e1 {{ Ψ1 }} #> e2 {{ Ψ2 }}
(heap_ctx heapN WP e1 {{ Ψ1 }} WP e2 {{ Ψ2 }}
v1 v2, Ψ1 v1 Ψ2 v2 -★ Φ (v1,v2)%V)
#> ParV e1 e2 {{ Φ }}.
WP ParV e1 e2 {{ Φ }}.
Proof.
intros. rewrite -par_spec //. repeat apply sep_mono; done || by wp_seq.
Qed.
......
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