Skip to content
Snippets Groups Projects
Commit 5c232f1e authored by Ralf Jung's avatar Ralf Jung
Browse files

Merge branch 'v2.0' of gitlab.mpi-sws.org:FP/iris-coq into v2.0

parents 83979416 a39468f8
No related branches found
No related tags found
No related merge requests found
......@@ -7,10 +7,11 @@ This version is known to compile with:
- Ssreflect 1.6
- Autosubst 1.4
For development, better make sure you have a version of Ssreflect that includes commit be724937
(no such version has been released so far, you'll have to fetch the development branch yourself).
Iris compiles fine even without this patch, but proof bullets will only be in 'strict' (enforcing)
mode with the fixed version of Ssreflect.
For development, better make sure you have a version of Ssreflect that includes
commit be724937 (no such version has been released so far, you will have to
fetch the development branch yourself). Iris compiles fine even without this
patch, but proof bullets will only be in 'strict' (enforcing) mode with the
fixed version of Ssreflect.
BUILDING INSTRUCTIONS
---------------------
......
......@@ -36,6 +36,8 @@ prelude/list.v
prelude/error.v
modures/option.v
modures/cmra.v
modures/cmra_big_op.v
modures/cmra_tactics.v
modures/sts.v
modures/auth.v
modures/fin_maps.v
......@@ -45,7 +47,6 @@ modures/base.v
modures/dra.v
modures/cofe_solver.v
modures/agree.v
modures/ra.v
modures/excl.v
iris/model.v
iris/adequacy.v
......@@ -56,14 +57,13 @@ iris/viewshifts.v
iris/wsat.v
iris/ownership.v
iris/weakestpre.v
iris/language.v
iris/pviewshifts.v
iris/resources.v
iris/hoare.v
iris/parameter.v
iris/language.v
iris/functor.v
iris/tests.v
barrier/heap_lang.v
barrier/parameter.v
barrier/lifting.v
barrier/sugar.v
barrier/tests.v
......@@ -390,8 +390,10 @@ Section Language.
Definition ectx_step e1 σ1 e2 σ2 (ef: option expr) :=
K e1' e2', e1 = fill K e1' e2 = fill K e2'
prim_step e1' σ1 e2' σ2 ef.
Global Program Instance heap_lang : Language expr value state := {|
Program Canonical Structure heap_lang : language := {|
language.expr := expr;
language.val := value;
language.state := state;
of_val := v2e;
to_val := e2v;
language.atomic := atomic;
......
Require Import prelude.gmap iris.lifting.
Require Export iris.weakestpre barrier.parameter.
Require Export iris.weakestpre barrier.heap_lang.
Import uPred.
(* TODO RJ: Figure out a way to to always use our Σ. *)
Section lifting.
Context {Σ : iFunctor}.
Implicit Types P : iProp heap_lang Σ.
Implicit Types Q : val heap_lang iProp heap_lang Σ.
(** Bind. *)
Lemma wp_bind {E e} K Q :
wp (Σ:=Σ) E e (λ v, wp (Σ:=Σ) E (fill K (v2e v)) Q) wp (Σ:=Σ) E (fill K e) Q.
Proof.
by apply (wp_bind (Σ:=Σ) (K := fill K)), fill_is_ctx.
Qed.
wp E e (λ v, wp E (fill K (v2e v)) Q) wp E (fill K e) Q.
Proof. apply (wp_bind (K:=fill K)), fill_is_ctx. Qed.
(** Base axioms for core primitives of the language: Stateful reductions. *)
......@@ -17,9 +18,9 @@ Lemma wp_lift_step E1 E2 (φ : expr → state → Prop) Q e1 σ1 :
E1 E2 to_val e1 = None
reducible e1 σ1
( e2 σ2 ef, prim_step e1 σ1 e2 σ2 ef ef = None φ e2 σ2)
pvs E2 E1 (ownP (Σ:=Σ) σ1 e2 σ2, ( φ e2 σ2 ownP (Σ:=Σ) σ2) -★
pvs E1 E2 (wp (Σ:=Σ) E2 e2 Q))
wp (Σ:=Σ) E2 e1 Q.
pvs E2 E1 (ownP σ1 e2 σ2, ( φ e2 σ2 ownP σ2) -★
pvs E1 E2 (wp E2 e2 Q))
wp E2 e1 Q.
Proof.
intros ? He Hsafe Hstep.
(* RJ: working around https://coq.inria.fr/bugs/show_bug.cgi?id=4536 *)
......@@ -45,8 +46,8 @@ Qed.
postcondition a predicate over a *location* *)
Lemma wp_alloc_pst E σ e v Q :
e2v e = Some v
(ownP (Σ:=Σ) σ ( l, (σ !! l = None) ownP (Σ:=Σ) (<[l:=v]>σ) -★ Q (LocV l)))
wp (Σ:=Σ) E (Alloc e) Q.
(ownP σ ( l, (σ !! l = None) ownP (<[l:=v]>σ) -★ Q (LocV l)))
wp E (Alloc e) Q.
Proof.
(* RJ FIXME (also for most other lemmas in this file): rewrite would be nicer... *)
intros Hvl. etransitivity; last eapply wp_lift_step with (σ1 := σ)
......@@ -72,7 +73,7 @@ Qed.
Lemma wp_load_pst E σ l v Q :
σ !! l = Some v
(ownP (Σ:=Σ) σ (ownP σ -★ Q v)) wp (Σ:=Σ) E (Load (Loc l)) Q.
(ownP σ (ownP σ -★ Q v)) wp E (Load (Loc l)) Q.
Proof.
intros Hl. etransitivity; last eapply wp_lift_step with (σ1 := σ)
(φ := λ e' σ', e' = v2e v σ' = σ); last first.
......@@ -93,7 +94,7 @@ Qed.
Lemma wp_store_pst E σ l e v v' Q :
e2v e = Some v
σ !! l = Some v'
(ownP (Σ:=Σ) σ (ownP (<[l:=v]>σ) -★ Q LitUnitV)) wp (Σ:=Σ) E (Store (Loc l) e) Q.
(ownP σ (ownP (<[l:=v]>σ) -★ Q LitUnitV)) wp E (Store (Loc l) e) Q.
Proof.
intros Hvl Hl. etransitivity; last eapply wp_lift_step with (σ1 := σ)
(φ := λ e' σ', e' = LitUnit σ' = <[l:=v]>σ); last first.
......@@ -114,17 +115,12 @@ Qed.
Lemma wp_cas_fail_pst E σ l e1 v1 e2 v2 v' Q :
e2v e1 = Some v1 e2v e2 = Some v2
σ !! l = Some v' v' <> v1
(ownP (Σ:=Σ) σ (ownP σ -★ Q LitFalseV)) wp (Σ:=Σ) E (Cas (Loc l) e1 e2) Q.
(ownP σ (ownP σ -★ Q LitFalseV)) wp E (Cas (Loc l) e1 e2) Q.
Proof.
intros Hvl Hl. etransitivity; last eapply wp_lift_step with (σ1 := σ)
(φ := λ e' σ', e' = LitFalse σ' = σ); last first.
- intros e2' σ2' ef Hstep. inversion_clear Hstep; first done.
(* FIXME this rewriting is rather ugly. *)
exfalso. rewrite Hvl in Hv1. case:Hv1=>?; subst v1. rewrite Hlookup in H.
case:H=>?; subst v'. done.
- do 3 eexists. eapply CasFailS; eassumption.
- reflexivity.
- reflexivity.
(φ := λ e' σ', e' = LitFalse σ' = σ) (E1:=E); auto; last first.
- by inversion_clear 1; simplify_map_equality.
- do 3 eexists; econstructor; eauto.
- rewrite -pvs_intro.
apply sep_mono; first done. apply later_mono.
apply forall_intro=>e2'. apply forall_intro=>σ2'.
......@@ -137,7 +133,7 @@ Qed.
Lemma wp_cas_suc_pst E σ l e1 v1 e2 v2 Q :
e2v e1 = Some v1 e2v e2 = Some v2
σ !! l = Some v1
(ownP (Σ:=Σ) σ (ownP (<[l:=v2]>σ) -★ Q LitTrueV)) wp (Σ:=Σ) E (Cas (Loc l) e1 e2) Q.
(ownP σ (ownP (<[l:=v2]>σ) -★ Q LitTrueV)) wp E (Cas (Loc l) e1 e2) Q.
Proof.
intros Hvl Hl. etransitivity; last eapply wp_lift_step with (σ1 := σ)
(φ := λ e' σ', e' = LitTrue σ' = <[l:=v2]>σ); last first.
......@@ -162,7 +158,8 @@ Qed.
(** Base axioms for core primitives of the language: Stateless reductions *)
Lemma wp_fork E e :
wp (Σ:=Σ) coPset_all e (λ _, True) wp (Σ:=Σ) E (Fork e) (λ v, (v = LitUnitV)).
wp coPset_all e (λ _, True : iProp heap_lang Σ)
wp E (Fork e) (λ v, (v = LitUnitV)).
Proof.
etransitivity; last eapply wp_lift_pure_step with
(φ := λ e' ef, e' = LitUnit ef = Some e);
......@@ -175,21 +172,16 @@ Proof.
eapply ForkS.
- reflexivity.
- apply later_mono.
apply forall_intro=>e2. apply forall_intro=>ef.
apply impl_intro_l. apply const_elim_l. intros [-> ->].
(* FIXME RJ This is ridicolous. *)
transitivity (True wp coPset_all e (λ _ : ival Σ, True))%I;
first by rewrite left_id.
apply sep_mono; last reflexivity.
rewrite -wp_value'; last reflexivity.
by apply const_intro.
apply forall_intro=>e2; apply forall_intro=>ef.
apply impl_intro_l, const_elim_l=>-[-> ->] /=; apply sep_intro_True_l; auto.
by rewrite -wp_value' //; apply const_intro.
Qed.
Lemma wp_lift_pure_step E (φ : expr Prop) Q e1 :
to_val e1 = None
( σ1, reducible e1 σ1)
( σ1 e2 σ2 ef, prim_step e1 σ1 e2 σ2 ef σ1 = σ2 ef = None φ e2)
( e2, φ e2 wp (Σ:=Σ) E e2 Q) wp (Σ:=Σ) E e1 Q.
( e2, φ e2 wp E e2 Q) wp E e1 Q.
Proof.
intros He Hsafe Hstep.
(* RJ: working around https://coq.inria.fr/bugs/show_bug.cgi?id=4536 *)
......@@ -209,7 +201,7 @@ Qed.
Lemma wp_rec E ef e v Q :
e2v e = Some v
wp (Σ:=Σ) E ef.[Rec ef, e /] Q wp (Σ:=Σ) E (App (Rec ef) e) Q.
wp E ef.[Rec ef, e /] Q wp E (App (Rec ef) e) Q.
Proof.
etransitivity; last eapply wp_lift_pure_step with
(φ := λ e', e' = ef.[Rec ef, e /]); last first.
......@@ -221,7 +213,7 @@ Proof.
Qed.
Lemma wp_plus n1 n2 E Q :
Q (LitNatV (n1 + n2)) wp (Σ:=Σ) E (Plus (LitNat n1) (LitNat n2)) Q.
Q (LitNatV (n1 + n2)) wp E (Plus (LitNat n1) (LitNat n2)) Q.
Proof.
etransitivity; last eapply wp_lift_pure_step with
(φ := λ e', e' = LitNat (n1 + n2)); last first.
......@@ -235,7 +227,7 @@ Qed.
Lemma wp_le_true n1 n2 E Q :
n1 n2
Q LitTrueV wp (Σ:=Σ) E (Le (LitNat n1) (LitNat n2)) Q.
Q LitTrueV wp E (Le (LitNat n1) (LitNat n2)) Q.
Proof.
intros Hle. etransitivity; last eapply wp_lift_pure_step with
(φ := λ e', e' = LitTrue); last first.
......@@ -250,7 +242,7 @@ Qed.
Lemma wp_le_false n1 n2 E Q :
n1 > n2
Q LitFalseV wp (Σ:=Σ) E (Le (LitNat n1) (LitNat n2)) Q.
Q LitFalseV wp E (Le (LitNat n1) (LitNat n2)) Q.
Proof.
intros Hle. etransitivity; last eapply wp_lift_pure_step with
(φ := λ e', e' = LitFalse); last first.
......@@ -265,7 +257,7 @@ Qed.
Lemma wp_fst e1 v1 e2 v2 E Q :
e2v e1 = Some v1 e2v e2 = Some v2
Q v1 wp (Σ:=Σ) E (Fst (Pair e1 e2)) Q.
Q v1 wp E (Fst (Pair e1 e2)) Q.
Proof.
intros Hv1 Hv2. etransitivity; last eapply wp_lift_pure_step with
(φ := λ e', e' = e1); last first.
......@@ -279,7 +271,7 @@ Qed.
Lemma wp_snd e1 v1 e2 v2 E Q :
e2v e1 = Some v1 e2v e2 = Some v2
Q v2 wp (Σ:=Σ) E (Snd (Pair e1 e2)) Q.
Q v2 wp E (Snd (Pair e1 e2)) Q.
Proof.
intros Hv1 Hv2. etransitivity; last eapply wp_lift_pure_step with
(φ := λ e', e' = e2); last first.
......@@ -293,7 +285,7 @@ Qed.
Lemma wp_case_inl e0 v0 e1 e2 E Q :
e2v e0 = Some v0
wp (Σ:=Σ) E e1.[e0/] Q wp (Σ:=Σ) E (Case (InjL e0) e1 e2) Q.
wp E e1.[e0/] Q wp E (Case (InjL e0) e1 e2) Q.
Proof.
intros Hv0. etransitivity; last eapply wp_lift_pure_step with
(φ := λ e', e' = e1.[e0/]); last first.
......@@ -306,7 +298,7 @@ Qed.
Lemma wp_case_inr e0 v0 e1 e2 E Q :
e2v e0 = Some v0
wp (Σ:=Σ) E e2.[e0/] Q wp (Σ:=Σ) E (Case (InjR e0) e1 e2) Q.
wp E e2.[e0/] Q wp E (Case (InjR e0) e1 e2) Q.
Proof.
intros Hv0. etransitivity; last eapply wp_lift_pure_step with
(φ := λ e', e' = e2.[e0/]); last first.
......@@ -322,7 +314,7 @@ Qed.
Lemma wp_le n1 n2 E P Q :
(n1 n2 P Q LitTrueV)
(n1 > n2 P Q LitFalseV)
P wp (Σ:=Σ) E (Le (LitNat n1) (LitNat n2)) Q.
P wp E (Le (LitNat n1) (LitNat n2)) Q.
Proof.
intros HPle HPgt.
assert (Decision (n1 n2)) as Hn12 by apply _.
......@@ -330,3 +322,4 @@ Proof.
- rewrite -wp_le_true; auto.
- assert (n1 > n2) by omega. rewrite -wp_le_false; auto.
Qed.
End lifting.
Require Export barrier.heap_lang.
Require Import iris.parameter.
Definition Σ := IParamConst heap_lang unitRA.
......@@ -16,10 +16,14 @@ Definition LamV (e : {bind expr}) := RecV e.[ren(+1)].
Definition LetCtx (K1 : ectx) (e2 : {bind expr}) := AppRCtx (LamV e2) K1.
Definition SeqCtx (K1 : ectx) (e2 : expr) := LetCtx K1 (e2.[ren(+1)]).
Section suger.
Context {Σ : iFunctor}.
Implicit Types P : iProp heap_lang Σ.
Implicit Types Q : val heap_lang iProp heap_lang Σ.
(** Proof rules for the sugar *)
Lemma wp_lam E ef e v Q :
e2v e = Some v
wp (Σ:=Σ) E ef.[e/] Q wp (Σ:=Σ) E (App (Lam ef) e) Q.
e2v e = Some v wp E ef.[e/] Q wp E (App (Lam ef) e) Q.
Proof.
intros Hv. rewrite -wp_rec; last eassumption.
(* RJ: This pulls in functional extensionality. If that bothers us, we have
......@@ -28,20 +32,20 @@ Proof.
Qed.
Lemma wp_let e1 e2 E Q :
wp (Σ:=Σ) E e1 (λ v, wp (Σ:=Σ) E (e2.[v2e v/]) Q) wp (Σ:=Σ) E (Let e1 e2) Q.
wp E e1 (λ v, wp E (e2.[v2e v/]) Q) wp E (Let e1 e2) Q.
Proof.
rewrite -(wp_bind (LetCtx EmptyCtx e2)). apply wp_mono=>v.
rewrite -wp_lam //. by rewrite v2v.
Qed.
Lemma wp_if_true e1 e2 E Q :
wp (Σ:=Σ) E e1 Q wp (Σ:=Σ) E (If LitTrue e1 e2) Q.
wp E e1 Q wp E (If LitTrue e1 e2) Q.
Proof.
rewrite -wp_case_inl //. by asimpl.
Qed.
Lemma wp_if_false e1 e2 E Q :
wp (Σ:=Σ) E e2 Q wp (Σ:=Σ) E (If LitFalse e1 e2) Q.
wp E e2 Q wp E (If LitFalse e1 e2) Q.
Proof.
rewrite -wp_case_inr //. by asimpl.
Qed.
......@@ -49,7 +53,7 @@ Qed.
Lemma wp_lt n1 n2 E P Q :
(n1 < n2 P Q LitTrueV)
(n1 n2 P Q LitFalseV)
P wp (Σ:=Σ) E (Lt (LitNat n1) (LitNat n2)) Q.
P wp E (Lt (LitNat n1) (LitNat n2)) Q.
Proof.
intros HPlt HPge.
rewrite -(wp_bind (LeLCtx EmptyCtx _)) -wp_plus -later_intro. simpl.
......@@ -59,7 +63,7 @@ Qed.
Lemma wp_eq n1 n2 E P Q :
(n1 = n2 P Q LitTrueV)
(n1 n2 P Q LitFalseV)
P wp (Σ:=Σ) E (Eq (LitNat n1) (LitNat n2)) Q.
P wp E (Eq (LitNat n1) (LitNat n2)) Q.
Proof.
intros HPeq HPne.
rewrite -wp_let -wp_value' // -later_intro. asimpl.
......@@ -72,3 +76,4 @@ Proof.
- asimpl. rewrite -wp_case_inr // -later_intro -wp_value' //.
apply HPne; omega.
Qed.
End suger.
......@@ -25,16 +25,16 @@ Module LangTests.
Qed.
End LangTests.
Module ParamTests.
Print Assumptions Σ.
End ParamTests.
Module LiftingTests.
Context {Σ : iFunctor}.
Implicit Types P : iProp heap_lang Σ.
Implicit Types Q : val heap_lang iProp heap_lang Σ.
(* TODO RJ: Some syntactic sugar for language expressions would be nice. *)
Definition e3 := Load (Var 0).
Definition e2 := Seq (Store (Var 0) (Plus (Load $ Var 0) (LitNat 1))) e3.
Definition e := Let (Alloc (LitNat 1)) e2.
Goal σ E, (ownP (Σ:=Σ) σ) (wp (Σ:=Σ) E e (λ v, (v = LitNatV 2))).
Goal σ E, (ownP σ : iProp heap_lang Σ) (wp E e (λ v, (v = LitNatV 2))).
Proof.
move=> σ E. rewrite /e.
rewrite -wp_let. rewrite -wp_alloc_pst; last done.
......@@ -74,7 +74,7 @@ Module LiftingTests.
Lemma FindPred_spec n1 n2 E Q :
((n1 < n2) Q (LitNatV $ pred n2))
wp (Σ:=Σ) E (App (FindPred (LitNat n2)) (LitNat n1)) Q.
wp E (App (FindPred (LitNat n2)) (LitNat n1)) Q.
Proof.
revert n1. apply löb_all_1=>n1.
rewrite -wp_rec //. asimpl.
......@@ -104,7 +104,7 @@ Module LiftingTests.
Qed.
Lemma Pred_spec n E Q :
Q (LitNatV $ pred n) wp (Σ:=Σ) E (App Pred (LitNat n)) Q.
Q (LitNatV $ pred n) wp E (App Pred (LitNat n)) Q.
Proof.
rewrite -wp_lam //. asimpl.
rewrite -(wp_bind (CaseCtx EmptyCtx _ _)).
......@@ -118,7 +118,9 @@ Module LiftingTests.
+ done.
Qed.
Goal E, True wp (Σ:=Σ) E (Let (App Pred (LitNat 42)) (App Pred (Var 0))) (λ v, (v = LitNatV 40)).
Goal E,
(True : iProp heap_lang Σ)
wp E (Let (App Pred (LitNat 42)) (App Pred (Var 0))) (λ v, (v = LitNatV 40)).
Proof.
intros E. rewrite -wp_let. rewrite -Pred_spec -!later_intro.
asimpl. (* TODO RJ: Can we somehow make it so that Pred gets folded again? *)
......
......@@ -7,9 +7,10 @@ Local Hint Extern 10 (✓{_} _) =>
solve_validN.
Section adequacy.
Context {Σ : iParam}.
Implicit Types e : iexpr Σ.
Implicit Types Q : ival Σ iProp Σ.
Context {Λ : language} {Σ : iFunctor}.
Implicit Types e : expr Λ.
Implicit Types Q : val Λ iProp Λ Σ.
Implicit Types m : iGst Λ Σ.
Transparent uPred_holds.
Notation wptp n := (Forall3 (λ e Q r, uPred_holds (wp coPset_all e Q) n r)).
......@@ -46,7 +47,7 @@ Proof.
* apply (IH (Qs1 ++ Q :: Qs2) (rs1 ++ r2 r2' :: rs2)).
{ rewrite /option_list right_id_L.
apply Forall3_app, Forall3_cons; eauto using wptp_le.
apply uPred_weaken with r2 (k + n); eauto using @ra_included_l. }
apply uPred_weaken with r2 (k + n); eauto using cmra_included_l. }
by rewrite -Permutation_middle /= big_op_app.
Qed.
Lemma ht_adequacy_steps P Q k n e1 t2 σ1 σ2 r1 :
......@@ -60,7 +61,7 @@ Proof.
intros Hht ????; apply (nsteps_wptp [pvs coPset_all coPset_all Q] k n
([e1],σ1) (t2,σ2) [r1]); rewrite /big_op ?right_id; auto.
constructor; last constructor.
apply Hht with r1 (k + n); eauto using @ra_included_unit.
apply Hht with r1 (k + n); eauto using cmra_included_unit.
by destruct (k + n).
Qed.
Lemma ht_adequacy_own Q e1 t2 σ1 m σ2 :
......@@ -70,15 +71,16 @@ Lemma ht_adequacy_own Q e1 t2 σ1 m σ2 :
rs2 Qs', wptp 3 t2 ((λ v, pvs coPset_all coPset_all (Q v)) :: Qs') rs2
wsat 3 coPset_all σ2 (big_op rs2).
Proof.
intros Hv ? [k ?]%rtc_nsteps. eapply ht_adequacy_steps with (r1 := (Res (Excl σ1) m)); eauto; [|].
- by rewrite Nat.add_comm; apply wsat_init, cmra_valid_validN.
- exists (Res (Excl σ1) ), (Res m). split_ands.
+ by rewrite /op /cmra_op /= /res_op /= !ra_empty_l ra_empty_r.
+ by rewrite /uPred_holds /=.
+ by apply ownG_spec.
intros Hv ? [k ?]%rtc_nsteps.
eapply ht_adequacy_steps with (r1 := (Res (Excl σ1) m)); eauto; [|].
{ by rewrite Nat.add_comm; apply wsat_init, cmra_valid_validN. }
exists (Res (Excl σ1) ), (Res m); split_ands.
* by rewrite Res_op ?left_id ?right_id.
* by rewrite /uPred_holds /=.
* by apply ownG_spec.
Qed.
Theorem ht_adequacy_result E φ e v t2 σ1 m σ2 :
m
m
{{ ownP σ1 ownG m }} e @ E {{ λ v', φ v' }}
rtc step ([e], σ1) (of_val v :: t2, σ2)
φ v.
......@@ -92,7 +94,7 @@ Proof.
by rewrite right_id_L.
Qed.
Lemma ht_adequacy_reducible E Q e1 e2 t2 σ1 m σ2 :
m
m
{{ ownP σ1 ownG m }} e1 @ E {{ Q }}
rtc step ([e1], σ1) (t2, σ2)
e2 t2 to_val e2 = None reducible e2 σ2.
......@@ -106,7 +108,7 @@ Proof.
rewrite ?right_id_L ?big_op_delete; auto.
Qed.
Theorem ht_adequacy_safe E Q e1 t2 σ1 m σ2 :
m
m
{{ ownP σ1 ownG m }} e1 @ E {{ Q }}
rtc step ([e1], σ1) (t2, σ2)
Forall (λ e, is_Some (to_val e)) t2 t3 σ3, step (t2, σ2) (t3, σ3).
......
Require Export modures.cmra.
Structure iFunctor := IFunctor {
ifunctor_car :> cofeT cmraT;
ifunctor_empty A : Empty (ifunctor_car A);
ifunctor_identity A : CMRAIdentity (ifunctor_car A);
ifunctor_map {A B} (f : A -n> B) : ifunctor_car A -n> ifunctor_car B;
ifunctor_map_ne {A B} n : Proper (dist n ==> dist n) (@ifunctor_map A B);
ifunctor_map_id {A : cofeT} (x : ifunctor_car A) : ifunctor_map cid x x;
ifunctor_map_compose {A B C} (f : A -n> B) (g : B -n> C) x :
ifunctor_map (g f) x ifunctor_map g (ifunctor_map f x);
ifunctor_map_mono {A B} (f : A -n> B) : CMRAMonotone (ifunctor_map f)
}.
Existing Instances ifunctor_empty ifunctor_identity.
Existing Instances ifunctor_map_ne ifunctor_map_mono.
Lemma ifunctor_map_ext (Σ : iFunctor) {A B} (f g : A -n> B) m :
( x, f x g x) ifunctor_map Σ f m ifunctor_map Σ g m.
Proof.
by intros; apply equiv_dist=> n; apply ifunctor_map_ne=> ?; apply equiv_dist.
Qed.
Program Definition iFunctor_const (icmra : cmraT) {icmra_empty : Empty icmra}
{icmra_identity : CMRAIdentity icmra} : iFunctor :=
{| ifunctor_car A := icmra; ifunctor_map A B f := cid |}.
Solve Obligations with done.
\ No newline at end of file
Require Export iris.weakestpre iris.viewshifts.
Definition ht {Σ} (E : coPset) (P : iProp Σ)
(e : iexpr Σ) (Q : ival Σ iProp Σ) : iProp Σ :=
Definition ht {Λ Σ} (E : coPset) (P : iProp Λ Σ)
(e : expr Λ) (Q : val Λ iProp Λ Σ) : iProp Λ Σ :=
( (P wp E e (λ v, pvs E E (Q v))))%I.
Instance: Params (@ht) 2.
Instance: Params (@ht) 3.
Notation "{{ P } } e @ E {{ Q } }" := (ht E P e Q)
(at level 74, format "{{ P } } e @ E {{ Q } }") : uPred_scope.
......@@ -11,27 +11,27 @@ Notation "{{ P } } e @ E {{ Q } }" := (True ⊑ ht E P e Q)
(at level 74, format "{{ P } } e @ E {{ Q } }") : C_scope.
Section hoare.
Context {Σ : iParam}.
Implicit Types P : iProp Σ.
Implicit Types Q : ival Σ iProp Σ.
Implicit Types v : ival Σ.
Context {Λ : language} {Σ : iFunctor}.
Implicit Types P : iProp Λ Σ.
Implicit Types Q : val Λ iProp Λ Σ.
Implicit Types v : val Λ.
Import uPred.
Global Instance ht_ne E n :
Proper (dist n ==> eq ==> pointwise_relation _ (dist n) ==> dist n) (@ht Σ E).
Proper (dist n ==> eq==>pointwise_relation _ (dist n) ==> dist n) (@ht Λ Σ E).
Proof. by intros P P' HP e ? <- Q Q' HQ; rewrite /ht HP; setoid_rewrite HQ. Qed.
Global Instance ht_proper E :
Proper (() ==> eq ==> pointwise_relation _ () ==> ()) (@ht Σ E).
Proper (() ==> eq ==> pointwise_relation _ () ==> ()) (@ht Λ Σ E).
Proof. by intros P P' HP e ? <- Q Q' HQ; rewrite /ht HP; setoid_rewrite HQ. Qed.
Lemma ht_mono E P P' Q Q' e :
P P' ( v, Q' v Q v) {{ P' }} e @ E {{ Q' }} {{ P }} e @ E {{ Q }}.
Proof. by intros HP HQ; rewrite /ht -HP; setoid_rewrite HQ. Qed.
Global Instance ht_mono' E :
Proper (flip () ==> eq ==> pointwise_relation _ () ==> ()) (@ht Σ E).
Proper (flip () ==> eq ==> pointwise_relation _ () ==> ()) (@ht Λ Σ E).
Proof. by intros P P' HP e ? <- Q Q' HQ; apply ht_mono. Qed.
Lemma ht_val E v :
{{ True }} of_val v @ E {{ λ v', (v = v') }}.
{{ True : iProp Λ Σ }} of_val v @ E {{ λ v', (v = v') }}.
Proof.
apply (always_intro' _ _), impl_intro_l.
by rewrite -wp_value -pvs_intro; apply const_intro.
......
......@@ -8,12 +8,14 @@ Local Notation "{{ P } } ef ?@ E {{ Q } }" :=
(at level 74, format "{{ P } } ef ?@ E {{ Q } }") : C_scope.
Section lifting.
Context {Σ : iParam}.
Implicit Types e : iexpr Σ.
Context {Λ : language} {Σ : iFunctor}.
Implicit Types e : expr Λ.
Implicit Types P : iProp Λ Σ.
Implicit Types R : val Λ iProp Λ Σ.
Import uPred.
Lemma ht_lift_step E1 E2
(φ : iexpr Σ istate Σ option (iexpr Σ) Prop) P P' Q1 Q2 R e1 σ1 :
(φ : expr Λ state Λ option (expr Λ) Prop) P P' Q1 Q2 R e1 σ1 :
E1 E2 to_val e1 = None
reducible e1 σ1
( e2 σ2 ef, prim_step e1 σ1 e2 σ2 ef φ e2 σ2 ef)
......@@ -42,8 +44,7 @@ Proof.
rewrite {1}/ht -always_wand_impl always_elim wand_elim_r; apply wp_mono=>v.
by apply const_intro.
Qed.
Lemma ht_lift_atomic E
(φ : iexpr Σ istate Σ option (iexpr Σ) Prop) P e1 σ1 :
Lemma ht_lift_atomic E (φ : expr Λ state Λ option (expr Λ) Prop) P e1 σ1 :
atomic e1
reducible e1 σ1
( e2 σ2 ef, prim_step e1 σ1 e2 σ2 ef φ e2 σ2 ef)
......@@ -68,7 +69,7 @@ Proof.
rewrite -(exist_intro σ2) -(exist_intro ef) (of_to_val e2) //.
by rewrite -always_and_sep_r'; apply and_intro; try apply const_intro.
Qed.
Lemma ht_lift_pure_step E (φ : iexpr Σ option (iexpr Σ) Prop) P P' Q e1 :
Lemma ht_lift_pure_step E (φ : expr Λ option (expr Λ) Prop) P P' Q e1 :
to_val e1 = None
( σ1, reducible e1 σ1)
( σ1 e2 σ2 ef, prim_step e1 σ1 e2 σ2 ef σ1 = σ2 φ e2 ef)
......@@ -94,7 +95,7 @@ Proof.
by apply const_intro.
Qed.
Lemma ht_lift_pure_determistic_step E
(φ : iexpr Σ option (iexpr Σ) Prop) P P' Q e1 e2 ef :
(φ : expr Λ option (expr Λ) Prop) P P' Q e1 e2 ef :
to_val e1 = None
( σ1, reducible e1 σ1)
( σ1 e2' σ2 ef', prim_step e1 σ1 e2' σ2 ef' σ1 = σ2 e2 = e2' ef = ef')
......
Require Export modures.base.
Require Export modures.cmra.
Class Language (E V St : Type) := {
of_val : V E;
to_val : E option V;
atomic : E Prop;
prim_step : E St E St option E Prop;
Structure language := Language {
expr : Type;
val : Type;
state : Type;
of_val : val expr;
to_val : expr option val;
atomic : expr Prop;
prim_step : expr state expr state option expr Prop;
to_of_val v : to_val (of_val v) = Some v;
of_to_val e v : to_val e = Some v of_val v = e;
values_stuck e σ e' σ' ef : prim_step e σ e' σ' ef to_val e = None;
......@@ -14,31 +17,43 @@ Class Language (E V St : Type) := {
prim_step e1 σ1 e2 σ2 ef
is_Some (to_val e2)
}.
Arguments of_val {_} _.
Arguments to_val {_} _.
Arguments atomic {_} _.
Arguments prim_step {_} _ _ _ _ _.
Arguments to_of_val {_} _.
Arguments of_to_val {_} _ _ _.
Arguments values_stuck {_} _ _ _ _ _ _.
Arguments atomic_not_value {_} _ _.
Arguments atomic_step {_} _ _ _ _ _ _ _.
Canonical Structure istateC Σ := leibnizC (state Σ).
Definition cfg (Λ : language) := (list (expr Λ) * state Λ)%type.
Section language.
Context `{Language E V St}.
Context {Λ : language}.
Implicit Types v : val Λ.
Definition reducible (e : E) (σ : St) :=
Definition reducible (e : expr Λ) (σ : state Λ) :=
e' σ' ef, prim_step e σ e' σ' ef.
Inductive step (ρ1 ρ2 : cfg Λ) : Prop :=
| step_atomic e1 σ1 e2 σ2 ef t1 t2 :
ρ1 = (t1 ++ e1 :: t2, σ1)
ρ2 = (t1 ++ e2 :: t2 ++ option_list ef, σ2)
prim_step e1 σ1 e2 σ2 ef
step ρ1 ρ2.
Lemma reducible_not_val e σ : reducible e σ to_val e = None.
Proof. intros (?&?&?&?); eauto using values_stuck. Qed.
Lemma atomic_of_val v : ¬atomic (of_val v).
Proof.
by intros Hat; apply atomic_not_value in Hat; rewrite to_of_val in Hat.
Qed.
Global Instance: Injective (=) (=) of_val.
Global Instance: Injective (=) (=) (@of_val Λ).
Proof. by intros v v' Hv; apply (injective Some); rewrite -!to_of_val Hv. Qed.
Definition cfg : Type := (list E * St)%type.
Inductive step (ρ1 ρ2 : cfg) : Prop :=
| step_atomic e1 σ1 e2 σ2 ef t1 t2 :
ρ1 = (t1 ++ e1 :: t2, σ1)
ρ2 = (t1 ++ e2 :: t2 ++ option_list ef, σ2)
prim_step e1 σ1 e2 σ2 ef
step ρ1 ρ2.
Record is_ctx (K : E E) := IsCtx {
Record is_ctx (K : expr Λ expr Λ) := IsCtx {
is_ctx_value e : to_val e = None to_val (K e) = None;
is_ctx_step_preserved e1 σ1 e2 σ2 ef :
prim_step e1 σ1 e2 σ2 ef prim_step (K e1) σ1 (K e2) σ2 ef;
......
......@@ -7,14 +7,15 @@ Local Hint Extern 10 (✓{_} _) =>
solve_validN.
Section lifting.
Context {Σ : iParam}.
Implicit Types v : ival Σ.
Implicit Types e : iexpr Σ.
Implicit Types σ : istate Σ.
Context {Λ : language} {Σ : iFunctor}.
Implicit Types v : val Λ.
Implicit Types e : expr Λ.
Implicit Types σ : state Λ.
Implicit Types Q : val Λ iProp Λ Σ.
Transparent uPred_holds.
Lemma wp_lift_step E1 E2
(φ : iexpr Σ istate Σ option (iexpr Σ) Prop) Q e1 σ1 :
(φ : expr Λ state Λ option (expr Λ) Prop) Q e1 σ1 :
E1 E2 to_val e1 = None
reducible e1 σ1
( e2 σ2 ef, prim_step e1 σ1 e2 σ2 ef φ e2 σ2 ef)
......@@ -35,7 +36,7 @@ Proof.
{ rewrite (commutative _ r2) -(associative _); eauto using wsat_le. }
by exists r1', r2'; split_ands; [| |by intros ? ->].
Qed.
Lemma wp_lift_pure_step E (φ : iexpr Σ option (iexpr Σ) Prop) Q e1 :
Lemma wp_lift_pure_step E (φ : expr Λ option (expr Λ) Prop) Q e1 :
to_val e1 = None
( σ1, reducible e1 σ1)
( σ1 e2 σ2 ef, prim_step e1 σ1 e2 σ2 ef σ1 = σ2 φ e2 ef)
......
......@@ -2,15 +2,17 @@ Require Export modures.logic iris.resources.
Require Import modures.cofe_solver.
Module iProp.
Definition F (Σ : iParam) (A B : cofeT) : cofeT := uPredC (resRA Σ (laterC A)).
Definition map {Σ : iParam} {A1 A2 B1 B2 : cofeT}
(f : (A2 -n> A1) * (B1 -n> B2)) : F Σ A1 B1 -n> F Σ A2 B2 :=
Definition F (Λ : language) (Σ : iFunctor) (A B : cofeT) : cofeT :=
uPredC (resRA Λ Σ (laterC A)).
Definition map {Λ : language} {Σ : iFunctor} {A1 A2 B1 B2 : cofeT}
(f : (A2 -n> A1) * (B1 -n> B2)) : F Λ Σ A1 B1 -n> F Λ Σ A2 B2 :=
uPredC_map (resRA_map (laterC_map (f.1))).
Definition result Σ : solution (F Σ).
Definition result Λ Σ : solution (F Λ Σ).
Proof.
apply (solver.result _ (@map Σ)).
* intros A B P. rewrite /map /= -{2}(uPred_map_id P). apply uPred_map_ext=>r {P} /=.
rewrite -{2}(res_map_id r). apply res_map_ext=>{r} r /=. by rewrite later_map_id.
apply (solver.result _ (@map Λ Σ)).
* intros A B P. rewrite /map /= -{2}(uPred_map_id P). apply uPred_map_ext=> r.
rewrite /= -{2}(res_map_id r); apply res_map_ext=>{r} r /=.
by rewrite later_map_id.
* intros A1 A2 A3 B1 B2 B3 f g f' g' P. rewrite /map /=.
rewrite -uPred_map_compose. apply uPred_map_ext=>{P} r /=.
rewrite -res_map_compose. apply res_map_ext=>{r} r /=.
......@@ -21,22 +23,24 @@ Qed.
End iProp.
(* Solution *)
Definition iPreProp (Σ : iParam) : cofeT := iProp.result Σ.
Notation iRes Σ := (res Σ (laterC (iPreProp Σ))).
Notation iResRA Σ := (resRA Σ (laterC (iPreProp Σ))).
Notation iWld Σ := (mapRA positive (agreeRA (laterC (iPreProp Σ)))).
Notation iPst Σ := (exclRA (istateC Σ)).
Notation iGst Σ := (icmra Σ (laterC (iPreProp Σ))).
Definition iProp (Σ : iParam) : cofeT := uPredC (iResRA Σ).
Definition iProp_unfold {Σ} : iProp Σ -n> iPreProp Σ := solution_fold _.
Definition iProp_fold {Σ} : iPreProp Σ -n> iProp Σ := solution_unfold _.
Lemma iProp_fold_unfold {Σ} (P : iProp Σ) : iProp_fold (iProp_unfold P) P.
Definition iPreProp (Λ : language) (Σ : iFunctor) : cofeT := iProp.result Λ Σ.
Notation iRes Λ Σ := (res Λ Σ (laterC (iPreProp Λ Σ))).
Notation iResRA Λ Σ := (resRA Λ Σ (laterC (iPreProp Λ Σ))).
Notation iWld Λ Σ := (mapRA positive (agreeRA (laterC (iPreProp Λ Σ)))).
Notation iPst Λ := (exclRA (istateC Λ)).
Notation iGst Λ Σ := (ifunctor_car Σ (laterC (iPreProp Λ Σ))).
Definition iProp (Λ : language) (Σ : iFunctor) : cofeT := uPredC (iResRA Λ Σ).
Definition iProp_unfold {Λ Σ} : iProp Λ Σ -n> iPreProp Λ Σ := solution_fold _.
Definition iProp_fold {Λ Σ} : iPreProp Λ Σ -n> iProp Λ Σ := solution_unfold _.
Lemma iProp_fold_unfold {Λ Σ} (P : iProp Λ Σ) : iProp_fold (iProp_unfold P) P.
Proof. apply solution_unfold_fold. Qed.
Lemma iProp_unfold_fold {Σ} (P : iPreProp Σ) : iProp_unfold (iProp_fold P) P.
Lemma iProp_unfold_fold {Λ Σ} (P : iPreProp Λ Σ) :
iProp_unfold (iProp_fold P) P.
Proof. apply solution_fold_unfold. Qed.
Bind Scope uPred_scope with iProp.
Instance iProp_fold_inj n Σ : Injective (dist n) (dist n) (@iProp_fold Σ).
Instance iProp_fold_inj n Λ Σ : Injective (dist n) (dist n) (@iProp_fold Λ Σ).
Proof. by intros X Y H; rewrite -(iProp_unfold_fold X) H iProp_unfold_fold. Qed.
Instance iProp_unfold_inj n Σ : Injective (dist n) (dist n) (@iProp_unfold Σ).
Instance iProp_unfold_inj n Λ Σ :
Injective (dist n) (dist n) (@iProp_unfold Λ Σ).
Proof. by intros X Y H; rewrite -(iProp_fold_unfold X) H iProp_fold_unfold. Qed.
Require Export iris.model.
Definition inv {Σ} (i : positive) (P : iProp Σ) : iProp Σ :=
Definition inv {Λ Σ} (i : positive) (P : iProp Λ Σ) : iProp Λ Σ :=
uPred_own (Res {[ i to_agree (Later (iProp_unfold P)) ]} ).
Arguments inv {_} _ _%I.
Definition ownP {Σ} (σ : istate Σ) : iProp Σ := uPred_own (Res (Excl σ) ).
Definition ownG {Σ} (m : iGst Σ) : iProp Σ := uPred_own (Res m).
Instance: Params (@inv) 2.
Instance: Params (@ownP) 1.
Instance: Params (@ownG) 1.
Arguments inv {_ _} _ _%I.
Definition ownP {Λ Σ} (σ: state Λ) : iProp Λ Σ := uPred_own (Res (Excl σ) ).
Definition ownG {Λ Σ} (m : iGst Λ Σ) : iProp Λ Σ := uPred_own (Res m).
Instance: Params (@inv) 3.
Instance: Params (@ownP) 2.
Instance: Params (@ownG) 2.
Typeclasses Opaque inv ownG ownP.
Section ownership.
Context {Σ : iParam}.
Implicit Types r : iRes Σ.
Implicit Types σ : istate Σ.
Implicit Types P : iProp Σ.
Implicit Types m : iGst Σ.
Context {Λ : language} {Σ : iFunctor}.
Implicit Types r : iRes Λ Σ.
Implicit Types σ : state Λ.
Implicit Types P : iProp Λ Σ.
Implicit Types m : iGst Λ Σ.
(* Invariants *)
Global Instance inv_contractive i : Contractive (@inv Σ i).
Global Instance inv_contractive i : Contractive (@inv Λ Σ i).
Proof.
intros n P Q HPQ.
apply (_: Proper (_ ==> _) iProp_unfold), Later_contractive in HPQ.
......@@ -27,7 +27,8 @@ Proof.
Qed.
Lemma always_inv i P : ( inv i P)%I inv i P.
Proof.
by apply uPred.always_own; rewrite Res_unit !ra_unit_empty map_unit_singleton.
apply uPred.always_own.
by rewrite Res_unit !cmra_unit_empty map_unit_singleton.
Qed.
Global Instance inv_always_stable i P : AlwaysStable (inv i P).
Proof. by rewrite /AlwaysStable always_inv. Qed.
......@@ -35,28 +36,29 @@ Lemma inv_sep_dup i P : inv i P ≡ (inv i P ★ inv i P)%I.
Proof. apply (uPred.always_sep_dup' _). Qed.
(* physical state *)
Lemma ownP_twice σ1 σ2 : (ownP σ1 ownP σ2 : iProp Σ) False.
Lemma ownP_twice σ1 σ2 : (ownP σ1 ownP σ2 : iProp Λ Σ) False.
Proof.
rewrite /ownP -uPred.own_op Res_op.
by apply uPred.own_invalid; intros (_&?&_).
Qed.
Global Instance ownP_timeless σ : TimelessP (ownP σ).
Global Instance ownP_timeless σ : TimelessP (@ownP Λ Σ σ).
Proof. rewrite /ownP; apply _. Qed.
(* ghost state *)
Global Instance ownG_ne n : Proper (dist n ==> dist n) (@ownG Σ).
Global Instance ownG_ne n : Proper (dist n ==> dist n) (@ownG Λ Σ).
Proof. by intros m m' Hm; unfold ownG; rewrite Hm. Qed.
Global Instance ownG_proper : Proper (() ==> ()) (@ownG Σ) := ne_proper _.
Global Instance ownG_proper : Proper (() ==> ()) (@ownG Λ Σ) := ne_proper _.
Lemma ownG_op m1 m2 : ownG (m1 m2) (ownG m1 ownG m2)%I.
Proof. by rewrite /ownG -uPred.own_op Res_op !(left_id _ _). Qed.
Lemma always_ownG_unit m : ( ownG (unit m))%I ownG (unit m).
Proof.
by apply uPred.always_own; rewrite Res_unit !ra_unit_empty ra_unit_idempotent.
apply uPred.always_own.
by rewrite Res_unit !cmra_unit_empty cmra_unit_idempotent.
Qed.
Lemma ownG_valid m : (ownG m) ( m).
Proof. by rewrite /ownG uPred.own_valid; apply uPred.valid_mono=> n [? []]. Qed.
Lemma ownG_valid_r m : (ownG m) (ownG m m).
Proof. apply uPred.always_entails_r', ownG_valid; by apply _. Qed.
Proof. apply (uPred.always_entails_r' _ _), ownG_valid. Qed.
Global Instance ownG_timeless m : Timeless m TimelessP (ownG m).
Proof. rewrite /ownG; apply _. Qed.
......@@ -69,15 +71,15 @@ Proof.
* intros [(P'&Hi&HP) _]; rewrite Hi.
by apply Some_dist, symmetry, agree_valid_includedN,
(cmra_included_includedN _ P'),HP; apply map_lookup_validN with (wld r) i.
* intros ?; split_ands; try apply cmra_empty_least; eauto.
* intros ?; split_ands; try apply cmra_empty_leastN; eauto.
Qed.
Lemma ownP_spec r n σ : {n} r (ownP σ) n r pst r ={n}= Excl σ.
Proof.
intros (?&?&?); rewrite /uPred_holds /= res_includedN /= Excl_includedN //.
naive_solver (apply cmra_empty_least).
naive_solver (apply cmra_empty_leastN).
Qed.
Lemma ownG_spec r n m : (ownG m) n r m {n} gst r.
Proof.
rewrite /uPred_holds /= res_includedN; naive_solver (apply cmra_empty_least).
rewrite /uPred_holds /= res_includedN; naive_solver (apply cmra_empty_leastN).
Qed.
End ownership.
Require Export modures.cmra iris.language.
Record iParam := IParam {
iexpr : Type;
ival : Type;
istate : Type;
ilanguage : Language iexpr ival istate;
icmra : cofeT cmraT;
icmra_empty A : Empty (icmra A);
icmra_empty_spec A : RAIdentity (icmra A);
icmra_empty_timeless A : Timeless ( : icmra A);
icmra_map {A B} (f : A -n> B) : icmra A -n> icmra B;
icmra_map_ne {A B} n : Proper (dist n ==> dist n) (@icmra_map A B);
icmra_map_id {A : cofeT} (x : icmra A) : icmra_map cid x x;
icmra_map_compose {A B C} (f : A -n> B) (g : B -n> C) x :
icmra_map (g f) x icmra_map g (icmra_map f x);
icmra_map_mono {A B} (f : A -n> B) : CMRAMonotone (icmra_map f)
}.
Arguments IParam {_ _ _} ilanguage icmra {_ _ _} _ {_ _ _ _}.
Existing Instances ilanguage.
Existing Instances icmra_empty icmra_empty_spec icmra_empty_timeless.
Existing Instances icmra_map_ne icmra_map_mono.
Lemma icmra_map_ext (Σ : iParam) {A B} (f g : A -n> B) m :
( x, f x g x) icmra_map Σ f m icmra_map Σ g m.
Proof.
by intros ?; apply equiv_dist=> n; apply icmra_map_ne=> ?; apply equiv_dist.
Qed.
Canonical Structure istateC Σ := leibnizC (istate Σ).
Definition IParamConst {iexpr ival istate : Type}
(ilanguage : Language iexpr ival istate)
(icmra : cmraT) {icmra_empty : Empty icmra}
{icmra_empty_spec : RAIdentity icmra} {icmra_empty_timeless: Timeless ( : icmra)}:
iParam.
eapply (IParam ilanguage (fun _ => icmra) (fun _ _ _ => cid)).
Unshelve.
- by intros.
- by intros.
Defined.
......@@ -7,38 +7,38 @@ Local Hint Extern 10 (✓{_} _) =>
repeat match goal with H : wsat _ _ _ _ |- _ => apply wsat_valid in H end;
solve_validN.
Program Definition pvs {Σ} (E1 E2 : coPset) (P : iProp Σ) : iProp Σ :=
Program Definition pvs {Λ Σ} (E1 E2 : coPset) (P : iProp Λ Σ) : iProp Λ Σ :=
{| uPred_holds n r1 := rf k Ef σ,
1 < k n (E1 E2) Ef =
wsat k (E1 Ef) σ (r1 rf)
r2, P k r2 wsat k (E2 Ef) σ (r2 rf) |}.
Next Obligation.
intros Σ E1 E2 P r1 r2 n HP Hr rf k Ef σ ?? Hwsat; simpl in *.
intros Λ Σ E1 E2 P r1 r2 n HP Hr rf k Ef σ ?? Hwsat; simpl in *.
apply HP; auto. by rewrite (dist_le _ _ _ _ Hr); last lia.
Qed.
Next Obligation. intros Σ E1 E2 P r rf k Ef σ; simpl in *; lia. Qed.
Next Obligation. intros Λ Σ E1 E2 P r rf k Ef σ; simpl in *; lia. Qed.
Next Obligation.
intros Σ E1 E2 P r1 r2 n1 n2 HP [r3 ?] Hn ? rf k Ef σ ?? Hws; setoid_subst.
intros Λ Σ E1 E2 P r1 r2 n1 n2 HP [r3 ?] Hn ? rf k Ef σ ?? Hws; setoid_subst.
destruct (HP (r3rf) k Ef σ) as (r'&?&Hws'); rewrite ?(associative op); auto.
exists (r' r3); rewrite -(associative _); split; last done.
apply uPred_weaken with r' k; eauto using @ra_included_l.
apply uPred_weaken with r' k; eauto using cmra_included_l.
Qed.
Arguments pvs {_} _ _ _%I : simpl never.
Instance: Params (@pvs) 3.
Arguments pvs {_ _} _ _ _%I : simpl never.
Instance: Params (@pvs) 4.
Section pvs.
Context {Σ : iParam}.
Implicit Types P Q : iProp Σ.
Implicit Types m : iGst Σ.
Context {Λ : language} {Σ : iFunctor}.
Implicit Types P Q : iProp Λ Σ.
Implicit Types m : iGst Λ Σ.
Transparent uPred_holds.
Global Instance pvs_ne E1 E2 n : Proper (dist n ==> dist n) (@pvs Σ E1 E2).
Global Instance pvs_ne E1 E2 n : Proper (dist n ==> dist n) (@pvs Λ Σ E1 E2).
Proof.
intros P Q HPQ r1 n' ??; simpl; split; intros HP rf k Ef σ ???;
destruct (HP rf k Ef σ) as (r2&?&?); auto;
exists r2; split_ands; auto; apply HPQ; eauto.
Qed.
Global Instance pvs_proper E1 E2 : Proper (() ==> ()) (@pvs Σ E1 E2).
Global Instance pvs_proper E1 E2 : Proper (() ==> ()) (@pvs Λ Σ E1 E2).
Proof. apply ne_proper, _. Qed.
Lemma pvs_intro E P : P pvs E E P.
......@@ -55,7 +55,7 @@ Lemma pvs_timeless E P : TimelessP P → (▷ P) ⊑ pvs E E P.
Proof.
rewrite uPred.timelessP_spec=> HP r [|n] ? HP' rf k Ef σ ???; first lia.
exists r; split; last done.
apply HP, uPred_weaken with r n; eauto using cmra_valid_le.
apply HP, uPred_weaken with r n; eauto using cmra_validN_le.
Qed.
Lemma pvs_trans E1 E2 E3 P :
E2 E1 E3 pvs E1 E2 (pvs E2 E3 P) pvs E1 E3 P.
......@@ -85,7 +85,7 @@ Proof.
destruct (wsat_open k Ef σ (r rf) i P) as (rP&?&?); auto.
{ rewrite lookup_wld_op_l ?Hinv; eauto; apply dist_le with (S n); eauto. }
exists (rP r); split; last by rewrite (left_id_L _ _) -(associative _).
eapply uPred_weaken with rP (S k); eauto using @ra_included_l.
eapply uPred_weaken with rP (S k); eauto using cmra_included_l.
Qed.
Lemma pvs_close i P : (inv i P P) pvs {[ i ]} True.
Proof.
......@@ -96,12 +96,12 @@ Proof.
* by rewrite -(left_id_L () Ef).
* apply uPred_weaken with r n; auto.
Qed.
Lemma pvs_updateP E m (P : iGst Σ Prop) :
Lemma pvs_updateP E m (P : iGst Λ Σ Prop) :
m ⇝: P ownG m pvs E E ( m', P m' ownG m').
Proof.
intros Hup r [|n] ? Hinv%ownG_spec rf [|k] Ef σ ???; try lia.
destruct (wsat_update_gst k (E Ef) σ r rf m P)
as (m'&?&?); eauto using cmra_included_le.
as (m'&?&?); eauto using cmra_includedN_le.
by exists (update_gst m' r); split; [exists m'; split; [|apply ownG_spec]|].
Qed.
Lemma pvs_alloc E P : ¬set_finite E P pvs E E ( i, (i E) inv i P).
......@@ -114,8 +114,9 @@ Proof.
Qed.
(* Derived rules *)
Opaque uPred_holds.
Import uPred.
Global Instance pvs_mono' E1 E2 : Proper (() ==> ()) (@pvs Σ E1 E2).
Global Instance pvs_mono' E1 E2 : Proper (() ==> ()) (@pvs Λ Σ E1 E2).
Proof. intros P Q; apply pvs_mono. Qed.
Lemma pvs_trans' E P : pvs E E (pvs E E P) pvs E E P.
Proof. apply pvs_trans; solve_elem_of. Qed.
......
Require Export modures.fin_maps modures.agree modures.excl iris.parameter.
Require Export modures.fin_maps modures.agree modures.excl.
Require Export iris.language iris.functor.
Record res (Σ : iParam) (A : cofeT) := Res {
Record res (Λ : language) (Σ : iFunctor) (A : cofeT) := Res {
wld : mapRA positive (agreeRA A);
pst : exclRA (istateC Σ);
gst : icmra Σ A;
pst : exclRA (istateC Λ);
gst : Σ A;
}.
Add Printing Constructor res.
Arguments Res {_ _} _ _ _.
Arguments wld {_ _} _.
Arguments pst {_ _} _.
Arguments gst {_ _} _.
Instance: Params (@Res) 2.
Instance: Params (@wld) 2.
Instance: Params (@pst) 2.
Instance: Params (@gst) 2.
Arguments Res {_ _ _} _ _ _.
Arguments wld {_ _ _} _.
Arguments pst {_ _ _} _.
Arguments gst {_ _ _} _.
Instance: Params (@Res) 3.
Instance: Params (@wld) 3.
Instance: Params (@pst) 3.
Instance: Params (@gst) 3.
Section res.
Context {Σ : iParam} {A : cofeT}.
Implicit Types r : res Σ A.
Context {Λ : language} {Σ : iFunctor} {A : cofeT}.
Implicit Types r : res Λ Σ A.
Inductive res_equiv' (r1 r2 : res Σ A) := Res_equiv :
Inductive res_equiv' (r1 r2 : res Λ Σ A) := Res_equiv :
wld r1 wld r2 pst r1 pst r2 gst r1 gst r2 res_equiv' r1 r2.
Instance res_equiv : Equiv (res Σ A) := res_equiv'.
Inductive res_dist' (n : nat) (r1 r2 : res Σ A) := Res_dist :
Instance res_equiv : Equiv (res Λ Σ A) := res_equiv'.
Inductive res_dist' (n : nat) (r1 r2 : res Λ Σ A) := Res_dist :
wld r1 ={n}= wld r2 pst r1 ={n}= pst r2 gst r1 ={n}= gst r2
res_dist' n r1 r2.
Instance res_dist : Dist (res Σ A) := res_dist'.
Instance res_dist : Dist (res Λ Σ A) := res_dist'.
Global Instance Res_ne n :
Proper (dist n ==> dist n ==> dist n ==> dist n) (@Res Σ A).
Proper (dist n ==> dist n ==> dist n ==> dist n) (@Res Λ Σ A).
Proof. done. Qed.
Global Instance Res_proper : Proper (() ==> () ==> () ==> ()) (@Res Σ A).
Global Instance Res_proper : Proper (() ==> () ==> () ==> ()) (@Res Λ Σ A).
Proof. done. Qed.
Global Instance wld_ne n : Proper (dist n ==> dist n) (@wld Σ A).
Global Instance wld_ne n : Proper (dist n ==> dist n) (@wld Λ Σ A).
Proof. by destruct 1. Qed.
Global Instance wld_proper : Proper (() ==> ()) (@wld Σ A).
Global Instance wld_proper : Proper (() ==> ()) (@wld Λ Σ A).
Proof. by destruct 1. Qed.
Global Instance pst_ne n : Proper (dist n ==> dist n) (@pst Σ A).
Global Instance pst_ne n : Proper (dist n ==> dist n) (@pst Λ Σ A).
Proof. by destruct 1. Qed.
Global Instance pst_ne' n : Proper (dist (S n) ==> ()) (@pst Σ A).
Global Instance pst_ne' n : Proper (dist (S n) ==> ()) (@pst Λ Σ A).
Proof.
intros σ σ' [???]; apply (timeless _), dist_le with (S n); auto with lia.
Qed.
Global Instance pst_proper : Proper (() ==> (=)) (@pst Σ A).
Global Instance pst_proper : Proper (() ==> (=)) (@pst Λ Σ A).
Proof. by destruct 1; unfold_leibniz. Qed.
Global Instance gst_ne n : Proper (dist n ==> dist n) (@gst Σ A).
Global Instance gst_ne n : Proper (dist n ==> dist n) (@gst Λ Σ A).
Proof. by destruct 1. Qed.
Global Instance gst_proper : Proper (() ==> ()) (@gst Σ A).
Global Instance gst_proper : Proper (() ==> ()) (@gst Λ Σ A).
Proof. by destruct 1. Qed.
Instance res_compl : Compl (res Σ A) := λ c,
Instance res_compl : Compl (res Λ Σ A) := λ c,
Res (compl (chain_map wld c))
(compl (chain_map pst c)) (compl (chain_map gst c)).
Definition res_cofe_mixin : CofeMixin (res Σ A).
Definition res_cofe_mixin : CofeMixin (res Λ Σ A).
Proof.
split.
* intros w1 w2; split.
......@@ -72,31 +73,30 @@ Global Instance res_timeless r :
Timeless (wld r) Timeless (gst r) Timeless r.
Proof. by destruct 3; constructor; try apply (timeless _). Qed.
Instance res_op : Op (res Σ A) := λ r1 r2,
Instance res_op : Op (res Λ Σ A) := λ r1 r2,
Res (wld r1 wld r2) (pst r1 pst r2) (gst r1 gst r2).
Global Instance res_empty : Empty (res Σ A) := Res ∅.
Instance res_unit : Unit (res Σ A) := λ r,
Global Instance res_empty : Empty (res Λ Σ A) := Res ∅.
Instance res_unit : Unit (res Λ Σ A) := λ r,
Res (unit (wld r)) (unit (pst r)) (unit (gst r)).
Instance res_valid : Valid (res Σ A) := λ r, (wld r) (pst r) (gst r).
Instance res_validN : ValidN (res Σ A) := λ n r,
Instance res_validN : ValidN (res Λ Σ A) := λ n r,
{n} (wld r) {n} (pst r) {n} (gst r).
Instance res_minus : Minus (res Σ A) := λ r1 r2,
Instance res_minus : Minus (res Λ Σ A) := λ r1 r2,
Res (wld r1 wld r2) (pst r1 pst r2) (gst r1 gst r2).
Lemma res_included (r1 r2 : res Σ A) :
Lemma res_included (r1 r2 : res Λ Σ A) :
r1 r2 wld r1 wld r2 pst r1 pst r2 gst r1 gst r2.
Proof.
split; [|by intros ([w ?]&[σ ?]&[m ?]); exists (Res w σ m)].
intros [r Hr]; split_ands;
[exists (wld r)|exists (pst r)|exists (gst r)]; apply Hr.
Qed.
Lemma res_includedN (r1 r2 : res Σ A) n :
Lemma res_includedN (r1 r2 : res Λ Σ A) n :
r1 {n} r2 wld r1 {n} wld r2 pst r1 {n} pst r2 gst r1 {n} gst r2.
Proof.
split; [|by intros ([w ?]&[σ ?]&[m ?]); exists (Res w σ m)].
intros [r Hr]; split_ands;
[exists (wld r)|exists (pst r)|exists (gst r)]; apply Hr.
Qed.
Definition res_cmra_mixin : CMRAMixin (res Σ A).
Definition res_cmra_mixin : CMRAMixin (res Λ Σ A).
Proof.
split.
* by intros n x [???] ? [???]; constructor; simpl in *; cofe_subst.
......@@ -105,26 +105,19 @@ Proof.
* by intros n [???] ? [???] [???] ? [???];
constructor; simpl in *; cofe_subst.
* done.
* by intros n ? (?&?&?); split_ands'; apply cmra_valid_S.
* intros r; unfold valid, res_valid, validN, res_validN.
rewrite !cmra_valid_validN; naive_solver.
* by intros n ? (?&?&?); split_ands'; apply cmra_validN_S.
* intros ???; constructor; simpl; apply (associative _).
* intros ??; constructor; simpl; apply (commutative _).
* intros ?; constructor; simpl; apply ra_unit_l.
* intros ?; constructor; simpl; apply ra_unit_idempotent.
* intros ?; constructor; simpl; apply cmra_unit_l.
* intros ?; constructor; simpl; apply cmra_unit_idempotent.
* intros n r1 r2; rewrite !res_includedN.
by intros (?&?&?); split_ands'; apply cmra_unit_preserving.
by intros (?&?&?); split_ands'; apply cmra_unit_preservingN.
* intros n r1 r2 (?&?&?);
split_ands'; simpl in *; eapply cmra_valid_op_l; eauto.
split_ands'; simpl in *; eapply cmra_validN_op_l; eauto.
* intros n r1 r2; rewrite res_includedN; intros (?&?&?).
by constructor; apply cmra_op_minus.
Qed.
Global Instance res_ra_empty : RAIdentity (res Σ A).
Proof.
by repeat split; simpl; repeat apply ra_empty_valid; rewrite (left_id _ _).
Qed.
Definition res_cmra_extend_mixin : CMRAExtendMixin (res Σ A).
Definition res_cmra_extend_mixin : CMRAExtendMixin (res Λ Σ A).
Proof.
intros n r r1 r2 (?&?&?) [???]; simpl in *.
destruct (cmra_extend_op n (wld r) (wld r1) (wld r2)) as ([w w']&?&?&?),
......@@ -134,14 +127,21 @@ Proof.
Qed.
Canonical Structure resRA : cmraT :=
CMRAT res_cofe_mixin res_cmra_mixin res_cmra_extend_mixin.
Global Instance res_cmra_identity : CMRAIdentity resRA.
Proof.
split.
* intros n; split_ands'; apply cmra_empty_valid.
* by split; rewrite /= (left_id _ _).
* apply _.
Qed.
Definition update_pst (σ : istate Σ) (r : res Σ A) : res Σ A :=
Definition update_pst (σ : state Λ) (r : res Λ Σ A) : res Λ Σ A :=
Res (wld r) (Excl σ) (gst r).
Definition update_gst (m : icmra Σ A) (r : res Σ A) : res Σ A :=
Definition update_gst (m : Σ A) (r : res Λ Σ A) : res Λ Σ A :=
Res (wld r) (pst r) m.
Lemma wld_validN n r : {n} r {n} (wld r).
Proof. by intros (?&?&?). Qed.
Proof. by intros (?&?&?). Qed.
Lemma gst_validN n r : {n} r {n} (gst r).
Proof. by intros (?&?&?). Qed.
Lemma Res_op w1 w2 σ1 σ2 m1 m2 :
......@@ -153,7 +153,7 @@ Lemma lookup_wld_op_l n r1 r2 i P :
{n} (r1r2) wld r1 !! i ={n}= Some P (wld r1 wld r2) !! i ={n}= Some P.
Proof.
move=>/wld_validN /(_ i) Hval Hi1P; move: Hi1P Hval; rewrite lookup_op.
destruct (wld r2 !! i) as [P'|] eqn:Hi; rewrite !Hi ?(right_id _ _) // =>-> ?.
destruct (wld r2 !! i) as [P'|] eqn:Hi; rewrite !Hi ?right_id // =>-> ?.
by constructor; rewrite (agree_op_inv P P') // agree_idempotent.
Qed.
Lemma lookup_wld_op_r n r1 r2 i P :
......@@ -166,49 +166,50 @@ Proof. by intros ? ? [???]; constructor; apply (timeless _). Qed.
End res.
Arguments resRA : clear implicits.
Definition res_map {Σ A B} (f : A -n> B) (r : res Σ A) : res Σ B :=
Definition res_map {Λ Σ A B} (f : A -n> B) (r : res Λ Σ A) : res Λ Σ B :=
Res (agree_map f <$> (wld r))
(pst r)
(icmra_map Σ f (gst r)).
Instance res_map_ne Σ (A B : cofeT) (f : A -n> B) :
(ifunctor_map Σ f (gst r)).
Instance res_map_ne Λ Σ (A B : cofeT) (f : A -n> B) :
( n, Proper (dist n ==> dist n) f)
n, Proper (dist n ==> dist n) (@res_map Σ _ _ f).
n, Proper (dist n ==> dist n) (@res_map Λ Σ _ _ f).
Proof. by intros Hf n [] ? [???]; constructor; simpl in *; cofe_subst. Qed.
Lemma res_map_id {Σ A} (r : res Σ A) : res_map cid r r.
Lemma res_map_id {Λ Σ A} (r : res Λ Σ A) : res_map cid r r.
Proof.
constructor; simpl; [|done|].
* rewrite -{2}(map_fmap_id (wld r)); apply map_fmap_setoid_ext=> i y ? /=.
by rewrite -{2}(agree_map_id y); apply agree_map_ext=> y' /=.
* by rewrite -{2}(icmra_map_id Σ (gst r)); apply icmra_map_ext=> m /=.
* by rewrite -{2}(ifunctor_map_id Σ (gst r)); apply ifunctor_map_ext=> m /=.
Qed.
Lemma res_map_compose {Σ A B C} (f : A -n> B) (g : B -n> C) (r : res Σ A) :
Lemma res_map_compose {Λ Σ A B C} (f : A -n> B) (g : B -n> C) (r : res Λ Σ A) :
res_map (g f) r res_map g (res_map f r).
Proof.
constructor; simpl; [|done|].
* rewrite -map_fmap_compose; apply map_fmap_setoid_ext=> i y _ /=.
by rewrite -agree_map_compose; apply agree_map_ext=> y' /=.
* by rewrite -icmra_map_compose; apply icmra_map_ext=> m /=.
* by rewrite -ifunctor_map_compose; apply ifunctor_map_ext=> m /=.
Qed.
Lemma res_map_ext {Σ A B} (f g : A -n> B) :
( r, f r g r) -> r : res Σ A, res_map f r res_map g r.
Lemma res_map_ext {Λ Σ A B} (f g : A -n> B) (r : res Λ Σ A) :
( x, f x g x) res_map f r res_map g r.
Proof.
move=>H r. split; simpl;
[apply map_fmap_setoid_ext; intros; apply agree_map_ext | | apply icmra_map_ext]; done.
intros Hfg; split; simpl; auto.
* by apply map_fmap_setoid_ext=>i x ?; apply agree_map_ext.
* by apply ifunctor_map_ext.
Qed.
Definition resRA_map {Σ A B} (f : A -n> B) : resRA Σ A -n> resRA Σ B :=
CofeMor (res_map f : resRA Σ A resRA Σ B).
Instance res_map_cmra_monotone {Σ} {A B : cofeT} (f : A -n> B) :
CMRAMonotone (@res_map Σ _ _ f).
Definition resRA_map {Λ Σ A B} (f : A -n> B) : resRA Λ Σ A -n> resRA Λ Σ B :=
CofeMor (res_map f : resRA Λ Σ A resRA Λ Σ B).
Instance res_map_cmra_monotone {Λ Σ} {A B : cofeT} (f : A -n> B) :
CMRAMonotone (@res_map Λ Σ _ _ f).
Proof.
split.
* by intros n r1 r2; rewrite !res_includedN;
intros (?&?&?); split_ands'; simpl; try apply includedN_preserving.
* by intros n r (?&?&?); split_ands'; simpl; try apply validN_preserving.
Qed.
Lemma resRA_map_ne {Σ A B} (f g : A -n> B) n :
f ={n}= g resRA_map (Σ := Σ) f ={n}= resRA_map g.
Instance resRA_map_ne {Λ Σ A B} n :
Proper (dist n ==> dist n) (@resRA_map Λ Σ A B).
Proof.
intros ? r. split; simpl;
[apply (mapRA_map_ne _ (agreeRA_map f) (agreeRA_map g)), agreeRA_map_ne
| | apply icmra_map_ne]; done.
intros f g Hfg r; split; simpl; auto.
* by apply (mapRA_map_ne _ (agreeRA_map f) (agreeRA_map g)), agreeRA_map_ne.
* by apply ifunctor_map_ne.
Qed.
......@@ -2,5 +2,6 @@
Require Import iris.model.
Module ModelTest. (* Make sure we got the notations right. *)
Definition iResTest (Σ : iParam) (w : iWld Σ) (p : iPst Σ) (g : iGst Σ) : iRes Σ := Res w p g.
Definition iResTest {Λ : language} {Σ : iFunctor}
(w : iWld Λ Σ) (p : iPst Λ) (g : iGst Λ Σ) : iRes Λ Σ := Res w p g.
End ModelTest.
Require Export iris.pviewshifts.
Definition vs {Σ} (E1 E2 : coPset) (P Q : iProp Σ) : iProp Σ :=
Definition vs {Λ Σ} (E1 E2 : coPset) (P Q : iProp Λ Σ) : iProp Λ Σ :=
( (P pvs E1 E2 Q))%I.
Arguments vs {_} _ _ _%I _%I.
Instance: Params (@vs) 3.
Arguments vs {_ _} _ _ _%I _%I.
Instance: Params (@vs) 4.
Notation "P >{ E1 , E2 }> Q" := (vs E1 E2 P%I Q%I)
(at level 69, E1 at level 1, format "P >{ E1 , E2 }> Q") : uPred_scope.
Notation "P >{ E1 , E2 }> Q" := (True vs E1 E2 P%I Q%I)
......@@ -14,9 +14,9 @@ Notation "P >{ E }> Q" := (True ⊑ vs E E P%I Q%I)
(at level 69, E at level 1, format "P >{ E }> Q") : C_scope.
Section vs.
Context {Σ : iParam}.
Implicit Types P Q : iProp Σ.
Implicit Types m : iGst Σ.
Context {Λ : language} {Σ : iFunctor}.
Implicit Types P Q : iProp Λ Σ.
Implicit Types m : iGst Λ Σ.
Import uPred.
Lemma vs_alt E1 E2 P Q : (P pvs E1 E2 Q) P >{E1,E2}> Q.
......@@ -25,14 +25,15 @@ Proof.
by rewrite always_const (right_id _ _).
Qed.
Global Instance vs_ne E1 E2 n :
Proper (dist n ==> dist n ==> dist n) (@vs Σ E1 E2).
Proper (dist n ==> dist n ==> dist n) (@vs Λ Σ E1 E2).
Proof. by intros P P' HP Q Q' HQ; rewrite /vs HP HQ. Qed.
Global Instance vs_proper E1 E2 : Proper (() ==> () ==> ()) (@vs Σ E1 E2).
Global Instance vs_proper E1 E2 : Proper (() ==> () ==> ()) (@vs Λ Σ E1 E2).
Proof. apply ne_proper_2, _. Qed.
Lemma vs_mono E1 E2 P P' Q Q' :
P P' Q' Q P' >{E1,E2}> Q' P >{E1,E2}> Q.
Proof. by intros HP HQ; rewrite /vs -HP HQ. Qed.
Global Instance vs_mono' E1 E2: Proper (flip () ==> () ==> ()) (@vs Σ E1 E2).
Global Instance vs_mono' E1 E2 :
Proper (flip () ==> () ==> ()) (@vs Λ Σ E1 E2).
Proof. by intros until 2; apply vs_mono. Qed.
Lemma vs_false_elim E1 E2 P : False >{E1,E2}> P.
......@@ -85,7 +86,7 @@ Proof.
intros; rewrite -{1}(left_id_L () E) -vs_mask_frame; last solve_elem_of.
apply vs_close.
Qed.
Lemma vs_updateP E m (P : iGst Σ Prop) :
Lemma vs_updateP E m (P : iGst Λ Σ Prop) :
m ⇝: P ownG m >{E}> ( m', P m' ownG m').
Proof. by intros; apply vs_alt, pvs_updateP. Qed.
Lemma vs_update E m m' : m m' ownG m >{E}> ownG m'.
......
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