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

use a gross hack to express diamond-shaped dependencies between the different parts of Iris

parent ca346a48
No related branches found
No related tags found
No related merge requests found
......@@ -85,6 +85,7 @@ VFILES:=core_lang.v\
iris_meta.v\
iris_vs.v\
iris_wp.v\
iris_wp_rules.v\
lang.v\
masks.v\
world_prop.v
......
......@@ -4,7 +4,11 @@ Require Import ModuRes.RA ModuRes.UPred ModuRes.BI ModuRes.PreoMet ModuRes.Finma
Set Bullet Behavior "Strict Subproofs".
Module IrisRes (RL : RA_T) (C : CORE_LANG) <: RA_T.
(* Because Coq has a restriction of how to apply functors, we have to hack a bit here.
The hack that involves least work, is to duplicate the definition of our final
resource type, as a module type (which is how we can use it, circumventing the
Coq restrictions) and as a module (to show the type can be instantiated). *)
Module Type IRIS_RES (RL : RA_T) (C : CORE_LANG) <: RA_T.
Instance state_type : Setoid C.state := discreteType.
Definition res := (ra_res_ex C.state * RL.res)%type.
......@@ -16,12 +20,15 @@ Module IrisRes (RL : RA_T) (C : CORE_LANG) <: RA_T.
(* The order on (ra_pos res) is inferred correctly, but this one is not *)
Instance res_pord: preoType res := ra_preo res.
End IRIS_RES.
Module IrisRes (RL : RA_T) (C : CORE_LANG) <: IRIS_RES RL C.
Include IRIS_RES RL C. (* I cannot believe Coq lets me do this... *)
End IrisRes.
Module IrisCore (RL : RA_T) (C : CORE_LANG).
Module Type IRIS_CORE (RL : RA_T) (C : CORE_LANG) (R: IRIS_RES RL C) (WP: WORLD_PROP R).
Export C.
Module Export R := IrisRes RL C.
Module Export WP := WorldProp R.
Export R.
Export WP.
Delimit Scope iris_scope with iris.
Local Open Scope ra_scope.
......@@ -441,4 +448,8 @@ Module IrisCore (RL : RA_T) (C : CORE_LANG).
Ltac wsatM H := solve [done | apply (wsatM H); solve [done | omega] ].
End IRIS_CORE.
Module IrisCore (RL : RA_T) (C : CORE_LANG) (R: IRIS_RES RL C) (WP: WORLD_PROP R) : IRIS_CORE RL C R WP.
Include IRIS_CORE RL C R WP.
End IrisCore.
Require Import ssreflect.
Require Import core_lang masks iris_wp.
Require Import core_lang masks world_prop iris_core iris_vs iris_wp.
Require Import ModuRes.RA ModuRes.UPred ModuRes.BI ModuRes.PreoMet ModuRes.Finmap.
Set Bullet Behavior "Strict Subproofs".
Module IrisMeta (RL : RA_T) (C : CORE_LANG).
Module Export WP := IrisWP RL C.
Module Type IRIS_META (RL : RA_T) (C : CORE_LANG) (R: IRIS_RES RL C) (WP: WORLD_PROP R) (CORE: IRIS_CORE RL C R WP) (VS: IRIS_VS RL C R WP CORE) (HT: IRIS_HT RL C R WP CORE).
Export VS.
Export HT.
Local Open Scope lang_scope.
Local Open Scope ra_scope.
......@@ -201,6 +202,8 @@ Module IrisMeta (RL : RA_T) (C : CORE_LANG).
Set Bullet Behavior "None". (* PDS: Ridiculous. *)
Section RobustSafety.
Definition wf_nat_ind := well_founded_induction Wf_nat.lt_wf.
Implicit Types (P : Props) (i n k : nat) (safe : bool) (m : mask) (e : expr) (Q : vPred) (r : pres) (w : Wld) (σ : state).
Program Definition restrictV (Q : expr -n> Props) : vPred :=
......@@ -385,4 +388,8 @@ Module IrisMeta (RL : RA_T) (C : CORE_LANG).
End Lifting.
End IRIS_META.
Module IrisMeta (RL : RA_T) (C : CORE_LANG) (R: IRIS_RES RL C) (WP: WORLD_PROP R) (CORE: IRIS_CORE RL C R WP) (VS: IRIS_VS RL C R WP CORE) (HT: IRIS_HT RL C R WP CORE) : IRIS_META RL C R WP CORE VS HT.
Include IRIS_META RL C R WP CORE VS HT.
End IrisMeta.
......@@ -4,8 +4,8 @@ Require Import ModuRes.RA ModuRes.UPred ModuRes.BI ModuRes.PreoMet ModuRes.Finma
Set Bullet Behavior "Strict Subproofs".
Module IrisVS (RL : RA_T) (C : CORE_LANG).
Module Export CORE := IrisCore RL C.
Module Type IRIS_VS (RL : RA_T) (C : CORE_LANG) (R: IRIS_RES RL C) (WP: WORLD_PROP R) (CORE: IRIS_CORE RL C R WP).
Export CORE.
Local Open Scope ra_scope.
Local Open Scope bi_scope.
......@@ -79,8 +79,6 @@ Module IrisVS (RL : RA_T) (C : CORE_LANG).
Section ViewShiftProps.
Definition mask_sing i := mask_set mask_emp i True.
Lemma vsTimeless m P :
timeless P vs m m (P) P.
Proof.
......@@ -332,4 +330,8 @@ Module IrisVS (RL : RA_T) (C : CORE_LANG).
End ViewShiftProps.
End IRIS_VS.
Module IrisVS (RL : RA_T) (C : CORE_LANG) (R: IRIS_RES RL C) (WP: WORLD_PROP R) (CORE: IRIS_CORE RL C R WP) : IRIS_VS RL C R WP CORE.
Include IRIS_VS RL C R WP CORE.
End IrisVS.
This diff is collapsed.
This diff is collapsed.
......@@ -130,3 +130,5 @@ Proof.
intros m1 m1' EQm1 m2 m2' EQm2; split; intros Hd k [Hm1 Hm2];
apply (Hd k); (split; [apply EQm1, Hm1 | apply EQm2, Hm2]).
Qed.
Definition mask_sing i := mask_set mask_emp i True.
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