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

Product flattening.

parent 00c7245f
No related branches found
No related tags found
No related merge requests found
Pipeline #
...@@ -6,7 +6,7 @@ Section bool. ...@@ -6,7 +6,7 @@ Section bool.
Context `{typeG Σ}. Context `{typeG Σ}.
Program Definition bool : type := Program Definition bool : type :=
{| st_size := 1; st_own tid vl := ( z:bool, vl = [ #z ])%I |}. {| st_own tid vl := ( z:bool, vl = [ #z ])%I |}.
Next Obligation. iIntros (tid vl) "H". iDestruct "H" as (z) "%". by subst. Qed. Next Obligation. iIntros (tid vl) "H". iDestruct "H" as (z) "%". by subst. Qed.
Lemma typed_bool ρ (b:Datatypes.bool): typed_step_ty ρ #b bool. Lemma typed_bool ρ (b:Datatypes.bool): typed_step_ty ρ #b bool.
......
...@@ -10,8 +10,7 @@ Section fn. ...@@ -10,8 +10,7 @@ Section fn.
Program Definition fn {A n} (E : A elctx) Program Definition fn {A n} (E : A elctx)
(tys : A vec type n) (ty : A type) : type := (tys : A vec type n) (ty : A type) : type :=
{| st_size := 1; {| st_own tid vl := ( f, vl = [f] (x : A) (args : vec val n) (k : val),
st_own tid vl := ( f, vl = [f] (x : A) (args : vec val n) (k : val),
typed_body (E x) [] typed_body (E x) []
[CctxElt k [] 1 (λ v, [TCtx_holds (v!!!0) (ty x)])] [CctxElt k [] 1 (λ v, [TCtx_holds (v!!!0) (ty x)])]
(zip_with (TCtx_holds of_val) args (tys x)) (zip_with (TCtx_holds of_val) args (tys x))
......
...@@ -6,7 +6,7 @@ Section int. ...@@ -6,7 +6,7 @@ Section int.
Context `{typeG Σ}. Context `{typeG Σ}.
Program Definition int : type := Program Definition int : type :=
{| st_size := 1; st_own tid vl := ( z:Z, vl = [ #z ])%I |}. {| st_own tid vl := ( z:Z, vl = [ #z ])%I |}.
Next Obligation. iIntros (tid vl) "H". iDestruct "H" as (z) "%". by subst. Qed. Next Obligation. iIntros (tid vl) "H". iDestruct "H" as (z) "%". by subst. Qed.
Lemma typed_int ρ (z:Datatypes.nat) : Lemma typed_int ρ (z:Datatypes.nat) :
......
...@@ -12,7 +12,7 @@ Section own. ...@@ -12,7 +12,7 @@ Section own.
(* Even though it does not seem too natural to put this here, it is (* Even though it does not seem too natural to put this here, it is
the only place where it is used. *) the only place where it is used. *)
Program Definition uninit : type := Program Definition uninit : type :=
{| st_size := 1; st_own tid vl := length vl = 1%nat⌝%I |}. {| st_own tid vl := length vl = 1%nat⌝%I |}.
Next Obligation. done. Qed. Next Obligation. done. Qed.
Program Definition freeable_sz (n : nat) (sz : nat) (l : loc) : iProp Σ := Program Definition freeable_sz (n : nat) (sz : nat) (l : loc) : iProp Σ :=
......
...@@ -7,8 +7,17 @@ Section product. ...@@ -7,8 +7,17 @@ Section product.
Context `{typeG Σ}. Context `{typeG Σ}.
Program Definition unit : type := Program Definition unit : type :=
{| st_size := 0; st_own tid vl := vl = []⌝%I |}. {| ty_size := 0; ty_own tid vl := vl = []⌝%I; ty_shr κ tid E l := True%I |}.
Next Obligation. iIntros (tid vl) "%". simpl. by subst. Qed. Next Obligation. iIntros (tid vl) "%". by subst. Qed.
Next Obligation. by iIntros (???????) "_ _". Qed.
Next Obligation. by iIntros (???????) "_ _ $". Qed.
Global Instance unit_copy : Copy unit.
Proof.
split. by apply _.
iIntros (???????) "_ _ $". iExists 1%Qp. iSplitL; last by auto.
iExists []. iSplitL; last by auto. rewrite heap_mapsto_vec_nil. auto.
Qed.
Lemma split_prod_mt tid ty1 ty2 q l : Lemma split_prod_mt tid ty1 ty2 q l :
(l ↦∗{q}: λ vl, (l ↦∗{q}: λ vl,
...@@ -121,71 +130,49 @@ Notation Π := product. ...@@ -121,71 +130,49 @@ Notation Π := product.
Section typing. Section typing.
Context `{typeG Σ}. Context `{typeG Σ}.
(* FIXME : do we still need this (flattening and unflattening)? *) Global Instance prod2_assoc E L : Assoc (eqtype E L) product2.
Lemma ty_incl_prod2_assoc1 ρ ty1 ty2 ty3 :
ty_incl ρ (product2 ty1 (product2 ty2 ty3)) (product2 (product2 ty1 ty2) ty3).
Proof. Proof.
iIntros (tid) "_ _!>". iSplit; iIntros "!#/=*H". split; (split; simpl; [by rewrite assoc| |]; intros; iIntros "_ _ _ H").
- iDestruct "H" as (vl1 vl') "(% & Ho1 & H)". - iDestruct "H" as (vl1 vl') "(% & Ho1 & H)".
iDestruct "H" as (vl2 vl3) "(% & Ho2 & Ho3)". subst. iDestruct "H" as (vl2 vl3) "(% & Ho2 & Ho3)". subst.
iExists _, _. iSplit. by rewrite assoc. iFrame. iExists _, _. by iFrame. iExists _, _. iSplit. by rewrite assoc. iFrame. iExists _, _. by iFrame.
- iDestruct "H" as (E1 E2') "(% & Hs1 & H)". - iDestruct "H" as (E1 E2') "(% & Hs1 & H)".
iDestruct "H" as (E2 E3) "(% & Hs2 & Hs3)". rewrite shift_loc_assoc_nat. iDestruct "H" as (E2 E3) "(% & Hs2 & Hs3)". rewrite shift_loc_assoc_nat.
iSplit; last by rewrite assoc.
iExists (E1 E2), E3. iSplit. by iPureIntro; set_solver. iFrame. iExists (E1 E2), E3. iSplit. by iPureIntro; set_solver. iFrame.
iExists E1, E2. iSplit. by iPureIntro; set_solver. by iFrame. iExists E1, E2. iSplit. by iPureIntro; set_solver. by iFrame.
Qed.
Lemma ty_incl_prod2_assoc2 ρ ty1 ty2 ty3 :
ty_incl ρ (product2 (product2 ty1 ty2) ty3) (product2 ty1 (product2 ty2 ty3)).
Proof.
iIntros (tid) "_ _!>". iSplit; iIntros "!#/=*H".
- iDestruct "H" as (vl1 vl') "(% & H & Ho3)". - iDestruct "H" as (vl1 vl') "(% & H & Ho3)".
iDestruct "H" as (vl2 vl3) "(% & Ho1 & Ho2)". subst. iDestruct "H" as (vl2 vl3) "(% & Ho1 & Ho2)". subst.
iExists _, _. iSplit. by rewrite -assoc. iFrame. iExists _, _. by iFrame. iExists _, _. iSplit. by rewrite -assoc. iFrame. iExists _, _. by iFrame.
- iDestruct "H" as (E1' E3) "(% & H & Hs3)". - iDestruct "H" as (E1' E3) "(% & H & Hs3)".
iDestruct "H" as (E1 E2) "(% & Hs1 & Hs2)". rewrite shift_loc_assoc_nat. iDestruct "H" as (E1 E2) "(% & Hs1 & Hs2)". rewrite shift_loc_assoc_nat.
iSplit; last by rewrite assoc.
iExists E1, (E2 E3). iSplit. by iPureIntro; set_solver. iFrame. iExists E1, (E2 E3). iSplit. by iPureIntro; set_solver. iFrame.
iExists E2, E3. iSplit. by iPureIntro; set_solver. by iFrame. iExists E2, E3. iSplit. by iPureIntro; set_solver. by iFrame.
Qed. Qed.
Lemma ty_incl_prod_flatten ρ tyl1 tyl2 tyl3 : Global Instance prod2_unit_leftid E L : LeftId (eqtype E L) unit product2.
ty_incl ρ (Π(tyl1 ++ Π tyl2 :: tyl3)) Proof.
(Π(tyl1 ++ tyl2 ++ tyl3)). intros ty. split; (split; [done| |]); intros; iIntros "#LFT _ _ H".
- iDestruct "H" as (? ?) "(% & % & ?)". by subst.
- iDestruct "H" as (? ?) "(% & ? & ?)". rewrite shift_loc_0.
iApply (ty_shr_mono with "LFT []"); last done. set_solver. iApply lft_incl_refl.
- iExists [], _. eauto.
- iExists , F. rewrite shift_loc_0. iFrame. by iPureIntro; set_solver.
Qed.
Global Instance prod2_unit_rightid E L : RightId (eqtype E L) unit product2.
Proof. Proof.
Admitted. intros ty. split; (split; [by rewrite /= -plus_n_O| |]); intros; iIntros "#LFT _ _ H".
(* apply (ty_incl_weaken _ ⊤). apply perm_incl_top. *) - iDestruct "H" as (? ?) "(% & ? & %)". subst. by rewrite app_nil_r.
(* induction tyl1; last by apply (ty_incl_prod2 _ _ _ _ _ _). *) - iDestruct "H" as (? ?) "(% & ? & _)".
(* induction tyl2 as [|ty tyl2 IH]; simpl. *) iApply (ty_shr_mono with "LFT []"); last done. set_solver. iApply lft_incl_refl.
(* - iIntros (tid) "#LFT _". iSplitL; iIntros "/=!>!#*H". *) - iExists _, []. rewrite app_nil_r. eauto.
(* + iDestruct "H" as (vl1 vl2) "(% & % & Ho)". subst. done. *) - iExists F, ∅. iFrame. by iPureIntro; set_solver.
(* + iDestruct "H" as (E1 E2) "(% & H1 & Ho)". iSplit; last done. *) Qed.
(* rewrite shift_loc_0. iApply (ty_shr_mono with "LFT [] Ho"). set_solver. *)
(* iApply lft_incl_refl. *)
(* - etransitivity. apply ty_incl_prod2_assoc2. *)
(* eapply (ty_incl_prod2 _ _ _ _ _ _). done. apply IH. *)
(* Qed. *)
Lemma ty_incl_prod_unflatten ρ tyl1 tyl2 tyl3 : Lemma ty_incl_prod_flatten E L tyl1 tyl2 tyl3 :
ty_incl ρ (Π(tyl1 ++ tyl2 ++ tyl3)) eqtype E L (Π(tyl1 ++ Π tyl2 :: tyl3)) (Π(tyl1 ++ tyl2 ++ tyl3)).
(Π(tyl1 ++ Π tyl2 :: tyl3)).
Proof. Proof.
Admitted. unfold product. induction tyl1; simpl; last by f_equiv.
(* apply (ty_incl_weaken _ ⊤). apply perm_incl_top. *) induction tyl2. by rewrite left_id. by rewrite /= -assoc; f_equiv.
(* induction tyl1; last by apply (ty_incl_prod2 _ _ _ _ _ _). *) Qed.
(* induction tyl2 as [|ty tyl2 IH]; simpl. *)
(* - iIntros (tid) "#LFT _". iMod (bor_create with "LFT []") as "[Hbor _]". *)
(* done. instantiate (1:=True%I). by auto. instantiate (1:=static). *)
(* iMod (bor_fracture (λ _, True%I) with "LFT Hbor") as "#Hbor". done. *)
(* iSplitL; iIntros "/=!>!#*H". *)
(* + iExists [], vl. iFrame. auto. *)
(* + iSplit; last done. iExists ∅, E. iSplit. iPureIntro; set_solver. *)
(* rewrite shift_loc_0. iFrame. iExists []. iSplit; last auto. *)
(* setoid_rewrite heap_mapsto_vec_nil. *)
(* iApply (frac_bor_shorten with "[] Hbor"). iApply lft_incl_static. *)
(* - etransitivity; last apply ty_incl_prod2_assoc1. *)
(* eapply (ty_incl_prod2 _ _ _ _ _ _). done. apply IH. *)
(* Qed. *)
End typing. End typing.
...@@ -8,8 +8,7 @@ Section shr_bor. ...@@ -8,8 +8,7 @@ Section shr_bor.
Context `{typeG Σ}. Context `{typeG Σ}.
Program Definition shr_bor (κ : lft) (ty : type) : type := Program Definition shr_bor (κ : lft) (ty : type) : type :=
{| st_size := 1; {| st_own tid vl :=
st_own tid vl :=
( (l:loc), vl = [ #l ] ty.(ty_shr) κ tid (lrustN) l)%I |}. ( (l:loc), vl = [ #l ] ty.(ty_shr) κ tid (lrustN) l)%I |}.
Next Obligation. Next Obligation.
iIntros (κ ty tid vl) "H". iDestruct "H" as (l) "[% _]". by subst. iIntros (κ ty tid vl) "H". iDestruct "H" as (l) "[% _]". by subst.
...@@ -18,7 +17,7 @@ Section shr_bor. ...@@ -18,7 +17,7 @@ Section shr_bor.
Global Instance subtype_shr_bor_mono E L : Global Instance subtype_shr_bor_mono E L :
Proper (lctx_lft_incl E L --> subtype E L ==> subtype E L) shr_bor. Proper (lctx_lft_incl E L --> subtype E L ==> subtype E L) shr_bor.
Proof. Proof.
intros κ1 κ2 ty1 ty2 Hty. apply subtype_simple_type. done. intros κ1 κ2 ty1 ty2 Hty. apply subtype_simple_type.
iIntros (??) "#LFT #HE #HL H". iDestruct ( with "HE HL") as "#Hκ". iIntros (??) "#LFT #HE #HL H". iDestruct ( with "HE HL") as "#Hκ".
iDestruct "H" as (l) "(% & H)". subst. iExists _. iSplit. done. iDestruct "H" as (l) "(% & H)". subst. iExists _. iSplit. done.
iApply (ty2.(ty_shr_mono) with "LFT Hκ"). reflexivity. iApply (ty2.(ty_shr_mono) with "LFT Hκ"). reflexivity.
......
...@@ -6,7 +6,7 @@ From lrust.typing Require Import type_incl. ...@@ -6,7 +6,7 @@ From lrust.typing Require Import type_incl.
Section sum. Section sum.
Context `{typeG Σ}. Context `{typeG Σ}.
Program Definition emp : type := {| st_size := 0; st_own tid vl := False%I |}. Program Definition emp : type := {| st_own tid vl := False%I |}.
Next Obligation. iIntros (tid vl) "[]". Qed. Next Obligation. iIntros (tid vl) "[]". Qed.
Global Instance emp_empty : Empty type := emp. Global Instance emp_empty : Empty type := emp.
...@@ -101,7 +101,6 @@ Existing Instance LstTySize_nil. ...@@ -101,7 +101,6 @@ Existing Instance LstTySize_nil.
Hint Extern 1 (LstTySize _ (_ :: _)) => Hint Extern 1 (LstTySize _ (_ :: _)) =>
apply LstTySize_cons; [compute; reflexivity|] : typeclass_instances. apply LstTySize_cons; [compute; reflexivity|] : typeclass_instances.
(* Σ is commonly used for the current functor. So it cannot be defined (* Σ is commonly used for the current functor. So it cannot be defined
as Π for products. We stick to the following form. *) as Π for products. We stick to the following form. *)
Notation "Σ[ ty1 ; .. ; tyn ]" := Notation "Σ[ ty1 ; .. ; tyn ]" :=
...@@ -110,11 +109,6 @@ Notation "Σ[ ty1 ; .. ; tyn ]" := ...@@ -110,11 +109,6 @@ Notation "Σ[ ty1 ; .. ; tyn ]" :=
Section incl. Section incl.
Context `{typeG Σ}. Context `{typeG Σ}.
(* FIXME : do we need that anywhere?. *)
Lemma ty_incl_emp ρ ty : ty_incl ρ ty.
Proof.
Admitted.
(* TODO *) (* TODO *)
(* Lemma ty_incl_sum ρ n tyl1 tyl2 (_ : LstTySize n tyl1) (_ : LstTySize n tyl2) : *) (* Lemma ty_incl_sum ρ n tyl1 tyl2 (_ : LstTySize n tyl1) (_ : LstTySize n tyl2) : *)
(* Duplicable ρ → Forall2 (ty_incl ρ) tyl1 tyl2 → *) (* Duplicable ρ → Forall2 (ty_incl ρ) tyl1 tyl2 → *)
......
...@@ -56,16 +56,15 @@ Section type. ...@@ -56,16 +56,15 @@ Section type.
that simple_type does depend on it. Otherwise, the coercion defined that simple_type does depend on it. Otherwise, the coercion defined
bellow will not be acceptable by Coq. *) bellow will not be acceptable by Coq. *)
Record simple_type `{typeG Σ} := Record simple_type `{typeG Σ} :=
{ st_size : nat; { st_own : thread_id list val iProp Σ;
st_own : thread_id list val iProp Σ;
st_size_eq tid vl : st_own tid vl -∗ length vl = st_size; st_size_eq tid vl : st_own tid vl -∗ length vl = 1%nat;
st_own_persistent tid vl : PersistentP (st_own tid vl) }. st_own_persistent tid vl : PersistentP (st_own tid vl) }.
Global Existing Instance st_own_persistent. Global Existing Instance st_own_persistent.
Program Definition ty_of_st (st : simple_type) : type := Program Definition ty_of_st (st : simple_type) : type :=
{| ty_size := st.(st_size); ty_own := st.(st_own); {| ty_size := 1; ty_own := st.(st_own);
(* [st.(st_own) tid vl] needs to be outside of the fractured (* [st.(st_own) tid vl] needs to be outside of the fractured
borrow, otherwise I do not know how to prove the shr part of borrow, otherwise I do not know how to prove the shr part of
...@@ -154,12 +153,11 @@ Section subtyping. ...@@ -154,12 +153,11 @@ Section subtyping.
Qed. Qed.
Lemma subtype_simple_type (st1 st2 : simple_type) : Lemma subtype_simple_type (st1 st2 : simple_type) :
st1.(st_size) = st2.(st_size)
( tid vl, lft_ctx -∗ elctx_interp_0 E -∗ llctx_interp_0 L -∗ ( tid vl, lft_ctx -∗ elctx_interp_0 E -∗ llctx_interp_0 L -∗
st1.(st_own) tid vl -∗ st2.(st_own) tid vl) st1.(st_own) tid vl -∗ st2.(st_own) tid vl)
subtype st1 st2. subtype st1 st2.
Proof. Proof.
intros Hsz Hst. split; [done|by apply Hst|]. intros Hst. split; [done|by apply Hst|].
iIntros (????) "#LFT #HE #HL H /=". iIntros (????) "#LFT #HE #HL H /=".
iDestruct "H" as (vl) "[Hf [Hown|H†]]"; iExists vl; iFrame "Hf"; last by auto. iDestruct "H" as (vl) "[Hf [Hown|H†]]"; iExists vl; iFrame "Hf"; last by auto.
iLeft. by iApply (Hst with "LFT HE HL *"). iLeft. by iApply (Hst with "LFT HE HL *").
......
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