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

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

parents c9ea799b 21419737
No related branches found
No related tags found
No related merge requests found
...@@ -25,14 +25,11 @@ matching it against other forms of ownership. *) ...@@ -25,14 +25,11 @@ matching it against other forms of ownership. *)
Definition heap_mapsto_def `{heapG Σ} (l : loc) (v: val) : iPropG heap_lang Σ := Definition heap_mapsto_def `{heapG Σ} (l : loc) (v: val) : iPropG heap_lang Σ :=
auth_own heap_name {[ l := Excl v ]}. auth_own heap_name {[ l := Excl v ]}.
(* Perform sealing *) (* Perform sealing *)
Module Type HeapMapstoSig. Definition heap_mapsto_aux : { x : _ & x = @heap_mapsto_def }.
Parameter heap_mapsto : `{heapG Σ} (l : loc) (v: val), iPropG heap_lang Σ. exact (existT _ Logic.eq_refl). Qed.
Axiom heap_mapsto_eq : @heap_mapsto = @heap_mapsto_def. Definition heap_mapsto := projT1 heap_mapsto_aux.
End HeapMapstoSig. Definition heap_mapsto_eq : @heap_mapsto = @heap_mapsto_def := projT2 heap_mapsto_aux.
Module Export HeapMapsto : HeapMapstoSig. Arguments heap_mapsto {_ _} _ _.
Definition heap_mapsto := @heap_mapsto_def.
Definition heap_mapsto_eq := Logic.eq_refl (@heap_mapsto).
End HeapMapsto.
Definition heap_inv `{i : heapG Σ} (h : heapRA) : iPropG heap_lang Σ := Definition heap_inv `{i : heapG Σ} (h : heapRA) : iPropG heap_lang Σ :=
ownP (of_heap h). ownP (of_heap h).
......
...@@ -16,14 +16,11 @@ Proof. split; try apply _. apply: inGF_inG. Qed. ...@@ -16,14 +16,11 @@ Proof. split; try apply _. apply: inGF_inG. Qed.
Definition auth_own_def `{authG Λ Σ A} (γ : gname) (a : A) : iPropG Λ Σ := Definition auth_own_def `{authG Λ Σ A} (γ : gname) (a : A) : iPropG Λ Σ :=
own γ ( a). own γ ( a).
(* Perform sealing *) (* Perform sealing *)
Module Type AuthOwnSig. Definition auth_own_aux : { x : _ & x = @auth_own_def }.
Parameter auth_own : `{authG Λ Σ A} (γ : gname) (a : A), iPropG Λ Σ. exact (existT _ Logic.eq_refl). Qed.
Axiom auth_own_eq : @auth_own = @auth_own_def. Definition auth_own := projT1 auth_own_aux.
End AuthOwnSig. Definition auth_own_eq : @auth_own = @auth_own_def := projT2 auth_own_aux.
Module Export AuthOwn : AuthOwnSig. Arguments auth_own {_ _ _ _ _} _ _.
Definition auth_own := @auth_own_def.
Definition auth_own_eq := Logic.eq_refl (@auth_own).
End AuthOwn.
Definition auth_inv `{authG Λ Σ A} Definition auth_inv `{authG Λ Σ A}
(γ : gname) (φ : A iPropG Λ Σ) : iPropG Λ Σ := (γ : gname) (φ : A iPropG Λ Σ) : iPropG Λ Σ :=
......
...@@ -11,16 +11,12 @@ Proof. apply: inGF_inG. Qed. ...@@ -11,16 +11,12 @@ Proof. apply: inGF_inG. Qed.
Definition saved_prop_own_def `{savedPropG Λ Σ} Definition saved_prop_own_def `{savedPropG Λ Σ}
(γ : gname) (P : iPropG Λ Σ) : iPropG Λ Σ := (γ : gname) (P : iPropG Λ Σ) : iPropG Λ Σ :=
own γ (to_agree (Next (iProp_unfold P))). own γ (to_agree (Next (iProp_unfold P))).
(* Perform sealing. *) (* Perform sealing *)
Module Type SavedPropOwnSig. Definition saved_prop_own_aux : { x : _ & x = @saved_prop_own_def }.
Parameter saved_prop_own : `{savedPropG Λ Σ} (γ : gname) (P : iPropG Λ Σ), exact (existT _ Logic.eq_refl). Qed.
iPropG Λ Σ. Definition saved_prop_own := projT1 saved_prop_own_aux.
Axiom saved_prop_own_eq : @saved_prop_own = @saved_prop_own_def. Definition saved_prop_own_eq : @saved_prop_own = @saved_prop_own_def := projT2 saved_prop_own_aux.
End SavedPropOwnSig. Arguments saved_prop_own {_ _ _} _ _.
Module Export SavedPropOwn : SavedPropOwnSig.
Definition saved_prop_own := @saved_prop_own_def.
Definition saved_prop_own_eq := Logic.eq_refl (@saved_prop_own).
End SavedPropOwn.
Instance: Params (@saved_prop_own) 4. Instance: Params (@saved_prop_own) 4.
Section saved_prop. Section saved_prop.
......
...@@ -16,24 +16,22 @@ Proof. split; try apply _. apply: inGF_inG. Qed. ...@@ -16,24 +16,22 @@ Proof. split; try apply _. apply: inGF_inG. Qed.
Definition sts_ownS_def `{i : stsG Λ Σ sts} (γ : gname) Definition sts_ownS_def `{i : stsG Λ Σ sts} (γ : gname)
(S : sts.states sts) (T : sts.tokens sts) : iPropG Λ Σ:= (S : sts.states sts) (T : sts.tokens sts) : iPropG Λ Σ:=
own γ (sts_frag S T). own γ (sts_frag S T).
(* Perform sealing. *)
Definition sts_ownS_aux : { x : _ & x = @sts_ownS_def }.
exact (existT _ Logic.eq_refl). Qed.
Definition sts_ownS := projT1 sts_ownS_aux.
Definition sts_ownS_eq : @sts_ownS = @sts_ownS_def := projT2 sts_ownS_aux.
Arguments sts_ownS {_ _ _ _} _ _ _.
Definition sts_own_def `{i : stsG Λ Σ sts} (γ : gname) Definition sts_own_def `{i : stsG Λ Σ sts} (γ : gname)
(s : sts.state sts) (T : sts.tokens sts) : iPropG Λ Σ := (s : sts.state sts) (T : sts.tokens sts) : iPropG Λ Σ :=
own γ (sts_frag_up s T). own γ (sts_frag_up s T).
(* Perform sealing. *) (* Perform sealing. *)
Module Type StsOwnSig. Definition sts_own_aux : { x : _ & x = @sts_own_def }.
Parameter sts_ownS : `{i : stsG Λ Σ sts} (γ : gname) exact (existT _ Logic.eq_refl). Qed.
(S : sts.states sts) (T : sts.tokens sts), iPropG Λ Σ. Definition sts_own := projT1 sts_own_aux.
Parameter sts_own : `{i : stsG Λ Σ sts} (γ : gname) Definition sts_own_eq : @sts_own = @sts_own_def := projT2 sts_own_aux.
(s : sts.state sts) (T : sts.tokens sts), iPropG Λ Σ. Arguments sts_own {_ _ _ _} _ _ _.
Axiom sts_ownS_eq : @sts_ownS = @sts_ownS_def.
Axiom sts_own_eq : @sts_own = @sts_own_def.
End StsOwnSig.
Module Export StsOwn : StsOwnSig.
Definition sts_ownS := @sts_ownS_def.
Definition sts_own := @sts_own_def.
Definition sts_ownS_eq := Logic.eq_refl (@sts_ownS_def).
Definition sts_own_eq := Logic.eq_refl (@sts_own_def).
End StsOwn.
Definition sts_inv `{i : stsG Λ Σ sts} (γ : gname) Definition sts_inv `{i : stsG Λ Σ sts} (γ : gname)
(φ : sts.state sts iPropG Λ Σ) : iPropG Λ Σ := (φ : sts.state sts iPropG Λ Σ) : iPropG Λ Σ :=
......
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