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

uninit : a few lemmas, a proper file.

parent c9a7d94e
No related branches found
No related tags found
No related merge requests found
Pipeline #
......@@ -30,6 +30,7 @@ theories/typing/typing.v
theories/typing/lft_contexts.v
theories/typing/type_context.v
theories/typing/cont_context.v
theories/typing/uninit.v
theories/typing/own.v
theories/typing/uniq_bor.v
theories/typing/shr_bor.v
......
......@@ -4,17 +4,11 @@ From lrust.lifetime Require Import borrow frac_borrow.
From lrust.lang Require Export new_delete.
From lrust.lang Require Import heap.
From lrust.typing Require Export type.
From lrust.typing Require Import typing product perm.
From lrust.typing Require Import typing product perm uninit.
Section own.
Context `{typeG Σ}.
(* Even though it does not seem too natural to put this here, it is
the only place where it is used. *)
Program Definition uninit : type :=
{| st_own tid vl := length vl = 1%nat⌝%I |}.
Next Obligation. done. Qed.
Program Definition freeable_sz (n : nat) (sz : nat) (l : loc) : iProp Σ :=
match sz, n return _ with
| 0%nat, _ => True
......@@ -122,7 +116,7 @@ Section own.
Proof. intros ?? Heq. split; f_equiv; apply Heq. Qed.
Lemma typed_new ρ (n : nat):
0 n typed_step_ty ρ (new [ #n]%E) (own n (Π(replicate n uninit))).
0 n typed_step_ty ρ (new [ #n]%E) (own n (uninit n)).
Proof.
iIntros (Hn tid) "!#(#HEAP&_&_&$)". iApply (wp_new with "HEAP"); try done.
iIntros "!>*(% & H† & H↦)". iExists _. iSplit. done. iNext.
......@@ -132,9 +126,7 @@ Section own.
clear Hn. apply (inj Z.of_nat) in Hlen. subst.
iInduction vl as [|v vl] "IH". done.
iExists [v], vl. iSplit. done. by iSplit.
- assert (ty_size (Π (replicate n uninit)) = n) as ->.
{ clear. induction n; rewrite //= IHn //. }
by rewrite freeable_sz_full.
- by rewrite uninit_sz freeable_sz_full.
Qed.
Lemma typed_delete ty (ν : expr):
......@@ -177,8 +169,7 @@ Section own.
Qed.
Lemma consumes_move ty n:
consumes ty (λ ν, ν own n ty)%P
(λ ν, ν own n (Π(replicate ty.(ty_size) uninit)))%P.
consumes ty (λ ν, ν own n ty)%P (λ ν, ν own n (uninit ty.(ty_size)))%P.
Proof.
iIntros (ν tid Φ E ?) "_ H◁ Htl HΦ". iApply (has_type_wp with "H◁").
iIntros (v) "Hνv H◁". iDestruct "Hνv" as %Hνv.
......@@ -191,7 +182,6 @@ Section own.
- rewrite -Hlen. iExists vl. iIntros "{$H↦}!>". clear.
iInduction vl as [|v vl] "IH". done.
iExists [v], vl. iSplit. done. by iSplit.
- assert (ty_size (Π (replicate (ty_size ty) uninit)) = ty_size ty) as ->; last by auto.
clear. induction ty.(ty_size). done. by apply (f_equal S).
- rewrite uninit_sz; auto.
Qed.
End own.
......@@ -117,8 +117,7 @@ Section product.
iMod ("Hclose2" with "[H2 H↦2]") as "[$$]". by iExists _; iFrame. done.
Qed.
Definition product := fold_right product2 unit.
Definition product := foldr product2 unit.
(* Given that in practice, product will be used with concrete lists,
there should be no need to declare [Copy] and [Proper] instances
for [product]. *)
......@@ -169,7 +168,7 @@ Section typing.
- iExists F, ∅. iFrame. by iPureIntro; set_solver.
Qed.
Lemma subtype_prod_flatten E L tyl1 tyl2 tyl3 :
Lemma eqtype_prod_flatten E L tyl1 tyl2 tyl3 :
eqtype E L (Π(tyl1 ++ Π tyl2 :: tyl3)) (Π(tyl1 ++ tyl2 ++ tyl3)).
Proof.
unfold product. induction tyl1; simpl; last by f_equiv.
......
From lrust.typing Require Export type.
From lrust.typing Require Import product.
Section uninit.
Context `{typeG Σ}.
Program Definition uninit_1 : type :=
{| st_own tid vl := length vl = 1%nat⌝%I |}.
Next Obligation. done. Qed.
Definition uninit (n : nat) : type :=
Π (replicate n uninit_1).
Lemma eqtype_uninit_product E L ns :
eqtype E L (uninit (foldr plus 0%nat ns)) (Π(uninit <$> ns)).
Proof.
induction ns as [|n ns IH]. done.
rewrite /= /uninit replicate_plus (eqtype_prod_flatten E L []).
induction n. done. rewrite /product /=. by f_equiv.
Qed.
Lemma uninit_sz n : ty_size (uninit n) = n.
Proof. induction n. done. simpl. by f_equal. Qed.
End uninit.
\ No newline at end of file
0% Loading or .
You are about to add 0 people to the discussion. Proceed with caution.
Please register or to comment