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

more concise lambda lemmas

parent b2527d69
No related branches found
No related tags found
Loading
...@@ -164,11 +164,9 @@ Proof. ...@@ -164,11 +164,9 @@ Proof.
rewrite right_id. done. rewrite right_id. done.
Qed. Qed.
Lemma wp_rec' E e v P Q : Lemma wp_rec' E e v Q :
P wp (Σ:=Σ) E (e.[Rec e, v2e v /]) Q wp (Σ:=Σ) E (e.[Rec e, v2e v /]) Q wp (Σ:=Σ) E (App (Rec e) (v2e v)) Q.
P wp (Σ:=Σ) E (App (Rec e) (v2e v)) Q.
Proof. Proof.
intros HP.
etransitivity; last eapply wp_lift_pure_step with etransitivity; last eapply wp_lift_pure_step with
(φ := λ e', e' = e.[Rec e, v2e v /]); last first. (φ := λ e', e' = e.[Rec e, v2e v /]); last first.
- intros ? ? ? ? Hstep. inversion_clear Hstep. done. - intros ? ? ? ? Hstep. inversion_clear Hstep. done.
...@@ -178,11 +176,10 @@ Proof. ...@@ -178,11 +176,10 @@ Proof.
apply const_elim_l=>->. done. apply const_elim_l=>->. done.
Qed. Qed.
Lemma wp_lam E e v P Q : Lemma wp_lam E e v Q :
P wp (Σ:=Σ) E (e.[v2e v/]) Q wp (Σ:=Σ) E (e.[v2e v/]) Q wp (Σ:=Σ) E (App (Lam e) (v2e v)) Q.
P wp (Σ:=Σ) E (App (Lam e) (v2e v)) Q.
Proof. Proof.
intros HP. rewrite -wp_rec'; first (intros; apply later_mono; eassumption). rewrite -wp_rec'.
(* RJ: This pulls in functional extensionality. If that bothers us, we have (* RJ: This pulls in functional extensionality. If that bothers us, we have
to talk to the Autosubst guys. *) to talk to the Autosubst guys. *)
by asimpl. by asimpl.
......
...@@ -40,7 +40,7 @@ Module LiftingTests. ...@@ -40,7 +40,7 @@ Module LiftingTests.
rewrite -(wp_bind _ _ (LetCtx EmptyCtx e2)). rewrite -wp_mono. rewrite -(wp_bind _ _ (LetCtx EmptyCtx e2)). rewrite -wp_mono.
{ eapply wp_alloc; done. } { eapply wp_alloc; done. }
move=>v; apply exist_elim=>l. apply const_elim_l; move=>[-> _] {v}. move=>v; apply exist_elim=>l. apply const_elim_l; move=>[-> _] {v}.
rewrite (later_intro (ownP _)); apply wp_lam. asimpl. rewrite -wp_lam -later_intro. asimpl.
rewrite -(wp_bind _ _ (SeqCtx (StoreRCtx (LocV _) rewrite -(wp_bind _ _ (SeqCtx (StoreRCtx (LocV _)
(PlusLCtx EmptyCtx _)) (Load (Loc _)))). (PlusLCtx EmptyCtx _)) (Load (Loc _)))).
rewrite -wp_mono. rewrite -wp_mono.
...@@ -52,7 +52,7 @@ Module LiftingTests. ...@@ -52,7 +52,7 @@ Module LiftingTests.
rewrite -wp_mono. rewrite -wp_mono.
{ eapply wp_store; first reflexivity. apply: lookup_insert. } { eapply wp_store; first reflexivity. apply: lookup_insert. }
move=>v; apply const_elim_l; move=>-> {v}. move=>v; apply const_elim_l; move=>-> {v}.
rewrite (later_intro (ownP _)); apply wp_lam. asimpl. rewrite -wp_lam -later_intro. asimpl.
rewrite -wp_mono. rewrite -wp_mono.
{ eapply wp_load. apply: lookup_insert. } { eapply wp_load. apply: lookup_insert. }
move=>v; apply const_elim_l; move=>-> {v}. move=>v; apply const_elim_l; move=>-> {v}.
......
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