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

Rename type classes in proof mode.

We are now using the prefixes Into, From, and Is (the first two are
inspired by the names of some traits in the Rust stdlib), and hopefully
doing that consistenly.
parent b5a23477
No related branches found
No related tags found
No related merge requests found
......@@ -12,14 +12,14 @@ Implicit Types P Q : iPropG heap_lang Σ.
Implicit Types Φ : val iPropG heap_lang Σ.
Implicit Types Δ : envs (iResUR heap_lang (globalF Σ)).
Global Instance sep_destruct_mapsto l q v :
SepDestruct false (l {q} v) (l {q/2} v) (l {q/2} v).
Proof. by rewrite /SepDestruct heap_mapsto_op_split. Qed.
Global Instance into_sep_mapsto l q v :
IntoSep false (l {q} v) (l {q/2} v) (l {q/2} v).
Proof. by rewrite /IntoSep heap_mapsto_op_split. Qed.
Lemma tac_wp_alloc Δ Δ' N E j e v Φ :
to_val e = Some v
(Δ heap_ctx N) nclose N E
StripLaterEnvs Δ Δ'
IntoLaterEnvs Δ Δ'
( l, Δ'',
envs_app false (Esnoc Enil j (l v)) Δ' = Some Δ''
(Δ'' Φ (LitV (LitLoc l))))
......@@ -27,60 +27,60 @@ Lemma tac_wp_alloc Δ Δ' N E j e v Φ :
Proof.
intros ???? . rewrite -wp_alloc // -always_and_sep_l.
apply and_intro; first done.
rewrite strip_later_env_sound; apply later_mono, forall_intro=> l.
rewrite into_later_env_sound; apply later_mono, forall_intro=> l.
destruct ( l) as (Δ''&?&HΔ'). rewrite envs_app_sound //; simpl.
by rewrite right_id HΔ'.
Qed.
Lemma tac_wp_load Δ Δ' N E i l q v Φ :
(Δ heap_ctx N) nclose N E
StripLaterEnvs Δ Δ'
IntoLaterEnvs Δ Δ'
envs_lookup i Δ' = Some (false, l {q} v)%I
(Δ' Φ v)
Δ WP Load (Lit (LitLoc l)) @ E {{ Φ }}.
Proof.
intros. rewrite -wp_load // -always_and_sep_l. apply and_intro; first done.
rewrite strip_later_env_sound -later_sep envs_lookup_split //; simpl.
rewrite into_later_env_sound -later_sep envs_lookup_split //; simpl.
by apply later_mono, sep_mono_r, wand_mono.
Qed.
Lemma tac_wp_store Δ Δ' Δ'' N E i l v e v' Φ :
to_val e = Some v'
(Δ heap_ctx N) nclose N E
StripLaterEnvs Δ Δ'
IntoLaterEnvs Δ Δ'
envs_lookup i Δ' = Some (false, l v)%I
envs_simple_replace i false (Esnoc Enil i (l v')) Δ' = Some Δ''
(Δ'' Φ (LitV LitUnit)) Δ WP Store (Lit (LitLoc l)) e @ E {{ Φ }}.
Proof.
intros. rewrite -wp_store // -always_and_sep_l. apply and_intro; first done.
rewrite strip_later_env_sound -later_sep envs_simple_replace_sound //; simpl.
rewrite into_later_env_sound -later_sep envs_simple_replace_sound //; simpl.
rewrite right_id. by apply later_mono, sep_mono_r, wand_mono.
Qed.
Lemma tac_wp_cas_fail Δ Δ' N E i l q v e1 v1 e2 v2 Φ :
to_val e1 = Some v1 to_val e2 = Some v2
(Δ heap_ctx N) nclose N E
StripLaterEnvs Δ Δ'
IntoLaterEnvs Δ Δ'
envs_lookup i Δ' = Some (false, l {q} v)%I v v1
(Δ' Φ (LitV (LitBool false)))
Δ WP CAS (Lit (LitLoc l)) e1 e2 @ E {{ Φ }}.
Proof.
intros. rewrite -wp_cas_fail // -always_and_sep_l. apply and_intro; first done.
rewrite strip_later_env_sound -later_sep envs_lookup_split //; simpl.
rewrite into_later_env_sound -later_sep envs_lookup_split //; simpl.
by apply later_mono, sep_mono_r, wand_mono.
Qed.
Lemma tac_wp_cas_suc Δ Δ' Δ'' N E i l v e1 v1 e2 v2 Φ :
to_val e1 = Some v1 to_val e2 = Some v2
(Δ heap_ctx N) nclose N E
StripLaterEnvs Δ Δ'
IntoLaterEnvs Δ Δ'
envs_lookup i Δ' = Some (false, l v)%I v = v1
envs_simple_replace i false (Esnoc Enil i (l v2)) Δ' = Some Δ''
(Δ'' Φ (LitV (LitBool true))) Δ WP CAS (Lit (LitLoc l)) e1 e2 @ E {{ Φ }}.
Proof.
intros; subst.
rewrite -wp_cas_suc // -always_and_sep_l. apply and_intro; first done.
rewrite strip_later_env_sound -later_sep envs_simple_replace_sound //; simpl.
rewrite into_later_env_sound -later_sep envs_simple_replace_sound //; simpl.
rewrite right_id. by apply later_mono, sep_mono_r, wand_mono.
Qed.
End heap.
......
This diff is collapsed.
......@@ -6,10 +6,10 @@ Section ghost.
Context `{inG Λ Σ A}.
Implicit Types a b : A.
Global Instance sep_destruct_own p γ a b1 b2 :
OpDestruct a b1 b2 SepDestruct p (own γ a) (own γ b1) (own γ b2).
Proof. rewrite /OpDestruct /SepDestruct => ->. by rewrite own_op. Qed.
Global Instance sep_split_own γ a b :
SepSplit (own γ (a b)) (own γ a) (own γ b) | 90.
Proof. by rewrite /SepSplit own_op. Qed.
Global Instance into_sep_own p γ a b1 b2 :
IntoOp a b1 b2 IntoSep p (own γ a) (own γ b1) (own γ b2).
Proof. rewrite /IntoOp /IntoSep => ->. by rewrite own_op. Qed.
Global Instance from_sep_own γ a b :
FromSep (own γ (a b)) (own γ a) (own γ b) | 90.
Proof. by rewrite /FromSep own_op. Qed.
End ghost.
......@@ -9,23 +9,23 @@ Implicit Types N : namespace.
Implicit Types P Q R : iProp Λ Σ.
Lemma tac_inv_fsa {A} (fsa : FSA Λ Σ A) fsaV Δ Δ' E N i P Q Φ :
FSASplit Q E fsa fsaV Φ
IsFSA Q E fsa fsaV Φ
fsaV nclose N E (of_envs Δ inv N P)
envs_app false (Esnoc Enil i ( P)) Δ = Some Δ'
(Δ' fsa (E nclose N) (λ a, P Φ a)) Δ Q.
Proof.
intros ????? HΔ'. rewrite -(fsa_split Q) -(inv_fsa fsa _ _ P) //.
intros ????? HΔ'. rewrite (is_fsa Q) -(inv_fsa fsa _ _ P) //.
rewrite // -always_and_sep_l. apply and_intro; first done.
rewrite envs_app_sound //; simpl. by rewrite right_id HΔ'.
Qed.
Lemma tac_inv_fsa_timeless {A} (fsa : FSA Λ Σ A) fsaV Δ Δ' E N i P Q Φ :
FSASplit Q E fsa fsaV Φ
IsFSA Q E fsa fsaV Φ
fsaV nclose N E (of_envs Δ inv N P) TimelessP P
envs_app false (Esnoc Enil i P) Δ = Some Δ'
(Δ' fsa (E nclose N) (λ a, P Φ a)) Δ Q.
Proof.
intros ?????? HΔ'. rewrite -(fsa_split Q) -(inv_fsa fsa _ _ P) //.
intros ?????? HΔ'. rewrite (is_fsa Q) -(inv_fsa fsa _ _ P) //.
rewrite // -always_and_sep_l. apply and_intro, wand_intro_l; first done.
trans (|={E N}=> P Δ)%I; first by rewrite pvs_timeless pvs_frame_r.
apply (fsa_strip_pvs _).
......@@ -36,7 +36,7 @@ End invariants.
Tactic Notation "iInvCore" constr(N) "as" constr(H) :=
eapply tac_inv_fsa with _ _ _ _ N H _ _;
[let P := match goal with |- FSASplit ?P _ _ _ _ => P end in
[let P := match goal with |- IsFSA ?P _ _ _ _ => P end in
apply _ || fail "iInv: cannot viewshift in goal" P
|try fast_done (* atomic *)
|done || eauto with ndisj (* [eauto with ndisj] is slow *)
......@@ -62,7 +62,7 @@ Tactic Notation "iInv" constr(N) "as" "{" simple_intropattern(x1)
Tactic Notation "iInvCore>" constr(N) "as" constr(H) :=
eapply tac_inv_fsa_timeless with _ _ _ _ N H _ _;
[let P := match goal with |- FSASplit ?P _ _ _ _ => P end in
[let P := match goal with |- IsFSA ?P _ _ _ _ => P end in
apply _ || fail "iInv: cannot viewshift in goal" P
|try fast_done (* atomic *)
|done || eauto with ndisj (* [eauto with ndisj] is slow *)
......
......@@ -7,45 +7,44 @@ Section pvs.
Context {Λ : language} {Σ : iFunctor}.
Implicit Types P Q : iProp Λ Σ.
Global Instance to_assumption_pvs E p P Q :
ToAssumption p P Q ToAssumption p P (|={E}=> Q)%I.
Proof. rewrite /ToAssumption=>->. apply pvs_intro. Qed.
Global Instance sep_split_pvs E P Q1 Q2 :
SepSplit P Q1 Q2 SepSplit (|={E}=> P) (|={E}=> Q1) (|={E}=> Q2).
Proof. rewrite /SepSplit=><-. apply pvs_sep. Qed.
Global Instance from_assumption_pvs E p P Q :
FromAssumption p P Q FromAssumption p P (|={E}=> Q)%I.
Proof. rewrite /FromAssumption=>->. apply pvs_intro. Qed.
Global Instance from_sep_pvs E P Q1 Q2 :
FromSep P Q1 Q2 FromSep (|={E}=> P) (|={E}=> Q1) (|={E}=> Q2).
Proof. rewrite /FromSep=><-. apply pvs_sep. Qed.
Global Instance or_split_pvs E1 E2 P Q1 Q2 :
OrSplit P Q1 Q2 OrSplit (|={E1,E2}=> P) (|={E1,E2}=> Q1) (|={E1,E2}=> Q2).
Proof. rewrite /OrSplit=><-. apply or_elim; apply pvs_mono; auto with I. Qed.
FromOr P Q1 Q2 FromOr (|={E1,E2}=> P) (|={E1,E2}=> Q1) (|={E1,E2}=> Q2).
Proof. rewrite /FromOr=><-. apply or_elim; apply pvs_mono; auto with I. Qed.
Global Instance exists_split_pvs {A} E1 E2 P (Φ : A iProp Λ Σ) :
ExistSplit P Φ ExistSplit (|={E1,E2}=> P) (λ a, |={E1,E2}=> Φ a)%I.
FromExist P Φ FromExist (|={E1,E2}=> P) (λ a, |={E1,E2}=> Φ a)%I.
Proof.
rewrite /ExistSplit=><-. apply exist_elim=> a. by rewrite -(exist_intro a).
rewrite /FromExist=><-. apply exist_elim=> a. by rewrite -(exist_intro a).
Qed.
Global Instance frame_pvs E1 E2 R P Q :
Frame R P Q Frame R (|={E1,E2}=> P) (|={E1,E2}=> Q).
Proof. rewrite /Frame=><-. by rewrite pvs_frame_l. Qed.
Global Instance to_wand_pvs E1 E2 R P Q :
ToWand R P Q ToWand R (|={E1,E2}=> P) (|={E1,E2}=> Q) | 100.
Proof. rewrite /ToWand=>->. apply wand_intro_l. by rewrite pvs_wand_r. Qed.
Global Instance into_wand_pvs E1 E2 R P Q :
IntoWand R P Q IntoWand R (|={E1,E2}=> P) (|={E1,E2}=> Q) | 100.
Proof. rewrite /IntoWand=>->. apply wand_intro_l. by rewrite pvs_wand_r. Qed.
Class FSASplit {A} (P : iProp Λ Σ) (E : coPset)
Class IsFSA {A} (P : iProp Λ Σ) (E : coPset)
(fsa : FSA Λ Σ A) (fsaV : Prop) (Φ : A iProp Λ Σ) := {
fsa_split : fsa E Φ ⊣⊢ P;
fsa_split_is_fsa :> FrameShiftAssertion fsaV fsa;
is_fsa : P ⊣⊢ fsa E Φ;
is_fsa_is_fsa :> FrameShiftAssertion fsaV fsa;
}.
Global Arguments fsa_split {_} _ _ _ _ _ {_}.
Global Instance fsa_split_pvs E P :
FSASplit (|={E}=> P)%I E pvs_fsa True (λ _, P).
Global Arguments is_fsa {_} _ _ _ _ _ {_}.
Global Instance is_fsa_pvs E P :
IsFSA (|={E}=> P)%I E pvs_fsa True (λ _, P).
Proof. split. done. apply _. Qed.
Global Instance fsa_split_fsa {A} (fsa : FSA Λ Σ A) E Φ :
FrameShiftAssertion fsaV fsa FSASplit (fsa E Φ) E fsa fsaV Φ.
Global Instance is_fsa_fsa {A} (fsa : FSA Λ Σ A) E Φ :
FrameShiftAssertion fsaV fsa IsFSA (fsa E Φ) E fsa fsaV Φ.
Proof. done. Qed.
Global Instance to_assert_pvs {A} P Q E (fsa : FSA Λ Σ A) fsaV Φ :
FSASplit Q E fsa fsaV Φ ToAssert P Q (|={E}=> P).
IsFSA Q E fsa fsaV Φ IntoAssert P Q (|={E}=> P).
Proof.
intros.
by rewrite /ToAssert pvs_frame_r wand_elim_r -(fsa_split Q) fsa_pvs_fsa.
intros. by rewrite /IntoAssert pvs_frame_r wand_elim_r (is_fsa Q) fsa_pvs_fsa.
Qed.
Lemma tac_pvs_intro Δ E1 E2 Q : E1 = E2 (Δ Q) Δ |={E1,E2}=> Q.
......@@ -66,11 +65,11 @@ Qed.
Lemma tac_pvs_elim_fsa {A} (fsa : FSA Λ Σ A) fsaV Δ Δ' E i p P' P Q Φ :
envs_lookup i Δ = Some (p, P') P' = (|={E}=> P)%I
FSASplit Q E fsa fsaV Φ
IsFSA Q E fsa fsaV Φ
envs_replace i p false (Esnoc Enil i P) Δ = Some Δ'
(Δ' fsa E Φ) Δ Q.
Proof.
intros ? -> ??. rewrite -(fsa_split Q) -fsa_pvs_fsa.
intros ? -> ??. rewrite (is_fsa Q) -fsa_pvs_fsa.
eapply tac_pvs_elim; set_solver.
Qed.
......@@ -85,12 +84,12 @@ Proof.
Qed.
Lemma tac_pvs_timeless_fsa {A} (fsa : FSA Λ Σ A) fsaV Δ Δ' E i p P Q Φ :
FSASplit Q E fsa fsaV Φ
IsFSA Q E fsa fsaV Φ
envs_lookup i Δ = Some (p, P)%I TimelessP P
envs_simple_replace i p (Esnoc Enil i P) Δ = Some Δ'
(Δ' fsa E Φ) Δ Q.
Proof.
intros ????. rewrite -(fsa_split Q) -fsa_pvs_fsa.
intros ????. rewrite (is_fsa Q) -fsa_pvs_fsa.
eauto using tac_pvs_timeless.
Qed.
End pvs.
......@@ -114,7 +113,7 @@ Tactic Notation "iPvsCore" constr(H) :=
[env_cbv; reflexivity || fail "iPvs:" H "not found"
|let P := match goal with |- ?P = _ => P end in
reflexivity || fail "iPvs:" H ":" P "not a pvs with the right mask"
|let P := match goal with |- FSASplit ?P _ _ _ _ => P end in
|let P := match goal with |- IsFSA ?P _ _ _ _ => P end in
apply _ || fail "iPvs:" P "not a pvs"
|env_cbv; reflexivity|simpl]
end.
......@@ -170,7 +169,7 @@ Tactic Notation "iTimeless" constr(H) :=
|env_cbv; reflexivity|simpl]
| |- _ =>
eapply tac_pvs_timeless_fsa with _ _ _ _ H _ _ _;
[let P := match goal with |- FSASplit ?P _ _ _ _ => P end in
[let P := match goal with |- IsFSA ?P _ _ _ _ => P end in
apply _ || fail "iTimeless: " P "not a pvs"
|env_cbv; reflexivity || fail "iTimeless:" H "not found"
|let P := match goal with |- TimelessP ?P => P end in
......
......@@ -8,7 +8,7 @@ Context `{stsG Λ Σ sts} (φ : sts.state sts → iPropG Λ Σ).
Implicit Types P Q : iPropG Λ Σ.
Lemma tac_sts_fsa {A} (fsa : FSA Λ _ A) fsaV Δ E N i γ S T Q Φ :
FSASplit Q E fsa fsaV Φ
IsFSA Q E fsa fsaV Φ
fsaV
envs_lookup i Δ = Some (false, sts_ownS γ S T)
(of_envs Δ sts_ctx γ N φ) nclose N E
......@@ -18,7 +18,7 @@ Lemma tac_sts_fsa {A} (fsa : FSA Λ _ A) fsaV Δ E N i γ S T Q Φ :
sts.steps (s, T) (s', T') φ s' (sts_own γ s' T' -★ Φ a))))
Δ Q.
Proof.
intros ????? HΔ'. rewrite -(fsa_split Q) -(sts_fsaS φ fsa) //.
intros ????? HΔ'. rewrite (is_fsa Q) -(sts_fsaS φ fsa) //.
rewrite // -always_and_sep_l. apply and_intro; first done.
rewrite envs_lookup_sound //; simpl; apply sep_mono_r.
apply forall_intro=>s; apply wand_intro_l.
......@@ -36,7 +36,7 @@ Tactic Notation "iSts" constr(H) "as"
| gname => eapply tac_sts_fsa with _ _ _ _ _ _ _ H _ _ _
| _ => fail "iSts:" H "not a string or gname"
end;
[let P := match goal with |- FSASplit ?P _ _ _ _ => P end in
[let P := match goal with |- IsFSA ?P _ _ _ _ => P end in
apply _ || fail "iSts: cannot viewshift in goal" P
|try fast_done (* atomic *)
|iAssumptionCore || fail "iSts:" H "not found"
......
......@@ -68,7 +68,7 @@ Tactic Notation "iClear" constr(Hs) :=
Tactic Notation "iExact" constr(H) :=
eapply tac_assumption with H _ _; (* (i:=H) *)
[env_cbv; reflexivity || fail "iExact:" H "not found"
|let P := match goal with |- ToAssumption _ ?P _ => P end in
|let P := match goal with |- FromAssumption _ ?P _ => P end in
apply _ || fail "iExact:" H ":" P "does not match goal"].
Tactic Notation "iAssumptionCore" :=
......@@ -88,7 +88,7 @@ Tactic Notation "iAssumption" :=
let rec find p Γ Q :=
match Γ with
| Esnoc ?j ?P => first
[pose proof (_ : ToAssumption p P Q) as Hass;
[pose proof (_ : FromAssumption p P Q) as Hass;
apply (tac_assumption _ j p P); [env_cbv; reflexivity|apply Hass]
|find p Γ Q]
end in
......@@ -105,20 +105,20 @@ Tactic Notation "iExFalso" := apply tac_ex_falso.
Local Tactic Notation "iPersistent" constr(H) :=
eapply tac_persistent with _ H _ _ _; (* (i:=H) *)
[env_cbv; reflexivity || fail "iPersistent:" H "not found"
|let Q := match goal with |- ToPersistentP ?Q _ => Q end in
|let Q := match goal with |- IntoPersistentP ?Q _ => Q end in
apply _ || fail "iPersistent:" H ":" Q "not persistent"
|env_cbv; reflexivity|].
Local Tactic Notation "iPure" constr(H) "as" simple_intropattern(pat) :=
eapply tac_pure with _ H _ _ _; (* (i:=H1) *)
[env_cbv; reflexivity || fail "iPure:" H "not found"
|let P := match goal with |- ToPure ?P _ => P end in
|let P := match goal with |- IsPure ?P _ => P end in
apply _ || fail "iPure:" H ":" P "not pure"
|intros pat].
Tactic Notation "iPureIntro" :=
eapply tac_pure_intro;
[let P := match goal with |- ToPure ?P _ => P end in
[let P := match goal with |- IsPure ?P _ => P end in
apply _ || fail "iPureIntro:" P "not pure"|].
(** * Specialize *)
......@@ -144,7 +144,7 @@ Local Tactic Notation "iSpecializeArgs" constr(H) open_constr(xs) :=
Local Tactic Notation "iSpecializePat" constr(H) constr(pat) :=
let solve_to_wand H1 :=
let P := match goal with |- ToWand ?P _ _ => P end in
let P := match goal with |- IntoWand ?P _ _ => P end in
apply _ || fail "iSpecialize:" H1 ":" P "not an implication/wand" in
let rec go H1 pats :=
lazymatch pats with
......@@ -154,8 +154,8 @@ Local Tactic Notation "iSpecializePat" constr(H) constr(pat) :=
eapply tac_specialize with _ _ H2 _ H1 _ _ _ _; (* (j:=H1) (i:=H2) *)
[env_cbv; reflexivity || fail "iSpecialize:" H2 "not found"
|env_cbv; reflexivity || fail "iSpecialize:" H1 "not found"
|let P := match goal with |- ToWand ?P ?Q _ => P end in
let Q := match goal with |- ToWand ?P ?Q _ => Q end in
|let P := match goal with |- IntoWand ?P ?Q _ => P end in
let Q := match goal with |- IntoWand ?P ?Q _ => Q end in
apply _ || fail "iSpecialize: cannot instantiate" H1 ":" P "with" H2 ":" Q
|env_cbv; reflexivity|go H1 pats]
| SName true ?H2 :: ?pats =>
......@@ -184,7 +184,7 @@ Local Tactic Notation "iSpecializePat" constr(H) constr(pat) :=
eapply tac_specialize_pure with _ H1 _ _ _ _ _;
[env_cbv; reflexivity || fail "iSpecialize:" H1 "not found"
|solve_to_wand H1
|let Q := match goal with |- ToPure ?Q _ => Q end in
|let Q := match goal with |- IsPure ?Q _ => Q end in
apply _ || fail "iSpecialize:" Q "not pure"
|env_cbv; reflexivity
|(*goal*)
......@@ -194,7 +194,7 @@ Local Tactic Notation "iSpecializePat" constr(H) constr(pat) :=
[env_cbv; reflexivity || fail "iSpecialize:" H1 "not found"
|solve_to_wand H1
|match k with
| GoalStd => apply to_assert_fallthrough
| GoalStd => apply into_assert_fallthrough
| GoalPvs => apply _ || fail "iSpecialize: cannot generate pvs goal"
end
|env_cbv; reflexivity || fail "iSpecialize:" Hs "not found"
......@@ -237,7 +237,7 @@ Tactic Notation "iApply" open_constr(t) :=
[iExact H
|eapply tac_apply with _ H _ _ _;
[env_cbv; reflexivity || fail 1 "iApply:" H "not found"
|let P := match goal with |- ToWand ?P _ _ => P end in
|let P := match goal with |- IntoWand ?P _ _ => P end in
apply _ || fail 1 "iApply: cannot apply" H ":" P
|lazy beta (* reduce betas created by instantiation *)]] in
let Htmp := iFresh in
......@@ -320,17 +320,17 @@ Tactic Notation "iRevert" "{" ident(x1) ident(x2) ident(x3) ident(x4)
(** * Disjunction *)
Tactic Notation "iLeft" :=
eapply tac_or_l;
[let P := match goal with |- OrSplit ?P _ _ => P end in
[let P := match goal with |- FromOr ?P _ _ => P end in
apply _ || fail "iLeft:" P "not a disjunction"|].
Tactic Notation "iRight" :=
eapply tac_or_r;
[let P := match goal with |- OrSplit ?P _ _ => P end in
[let P := match goal with |- FromOr ?P _ _ => P end in
apply _ || fail "iRight:" P "not a disjunction"|].
Local Tactic Notation "iOrDestruct" constr(H) "as" constr(H1) constr(H2) :=
eapply tac_or_destruct with _ _ H _ H1 H2 _ _ _; (* (i:=H) (j1:=H1) (j2:=H2) *)
[env_cbv; reflexivity || fail "iOrDestruct:" H "not found"
|let P := match goal with |- OrDestruct ?P _ _ => P end in
|let P := match goal with |- IntoOr ?P _ _ => P end in
apply _ || fail "iOrDestruct:" H ":" P "not a disjunction"
|env_cbv; reflexivity || fail "iOrDestruct:" H1 "not fresh"
|env_cbv; reflexivity || fail "iOrDestruct:" H2 "not fresh"| |].
......@@ -340,7 +340,7 @@ Tactic Notation "iSplit" :=
lazymatch goal with
| |- _ _ =>
eapply tac_and_split;
[let P := match goal with |- AndSplit ?P _ _ => P end in
[let P := match goal with |- FromAnd ?P _ _ => P end in
apply _ || fail "iSplit:" P "not a conjunction"| |]
| |- _ ⊣⊢ _ => apply (anti_symm ())
end.
......@@ -348,13 +348,13 @@ Tactic Notation "iSplit" :=
Tactic Notation "iSplitL" constr(Hs) :=
let Hs := words Hs in
eapply tac_sep_split with _ _ false Hs _ _; (* (js:=Hs) *)
[let P := match goal with |- SepSplit ?P _ _ => P end in
[let P := match goal with |- FromSep ?P _ _ => P end in
apply _ || fail "iSplitL:" P "not a separating conjunction"
|env_cbv; reflexivity || fail "iSplitL: hypotheses" Hs "not found"| |].
Tactic Notation "iSplitR" constr(Hs) :=
let Hs := words Hs in
eapply tac_sep_split with _ _ true Hs _ _; (* (js:=Hs) *)
[let P := match goal with |- SepSplit ?P _ _ => P end in
[let P := match goal with |- FromSep ?P _ _ => P end in
apply _ || fail "iSplitR:" P "not a separating conjunction"
|env_cbv; reflexivity || fail "iSplitR: hypotheses" Hs "not found"| |].
......@@ -364,7 +364,7 @@ Tactic Notation "iSplitR" := iSplitL "".
Local Tactic Notation "iSepDestruct" constr(H) "as" constr(H1) constr(H2) :=
eapply tac_sep_destruct with _ H _ H1 H2 _ _ _; (* (i:=H) (j1:=H1) (j2:=H2) *)
[env_cbv; reflexivity || fail "iSepDestruct:" H "not found"
|let P := match goal with |- SepDestruct _ ?P _ _ => P end in
|let P := match goal with |- IntoSep _ ?P _ _ => P end in
apply _ || fail "iSepDestruct:" H ":" P "not separating destructable"
|env_cbv; reflexivity || fail "iSepDestruct:" H1 "or" H2 " not fresh"|].
......@@ -395,15 +395,15 @@ Tactic Notation "iCombine" constr(H1) constr(H2) "as" constr(H) :=
eapply tac_combine with _ _ _ H1 _ _ H2 _ _ H _;
[env_cbv; reflexivity || fail "iCombine:" H1 "not found"
|env_cbv; reflexivity || fail "iCombine:" H2 "not found"
|let P1 := match goal with |- SepSplit _ ?P1 _ => P1 end in
let P2 := match goal with |- SepSplit _ _ ?P2 => P2 end in
|let P1 := match goal with |- FromSep _ ?P1 _ => P1 end in
let P2 := match goal with |- FromSep _ _ ?P2 => P2 end in
apply _ || fail "iCombine: cannot combine" H1 ":" P1 "and" H2 ":" P2
|env_cbv; reflexivity || fail "iCombine:" H "not fresh"|].
(** * Existential *)
Tactic Notation "iExists" uconstr(x1) :=
eapply tac_exist;
[let P := match goal with |- ExistSplit ?P _ => P end in
[let P := match goal with |- FromExist ?P _ => P end in
apply _ || fail "iExists:" P "not an existential"
|cbv beta; eexists x1].
......@@ -432,7 +432,7 @@ Local Tactic Notation "iExistDestruct" constr(H)
"as" simple_intropattern(x) constr(Hx) :=
eapply tac_exist_destruct with H _ Hx _ _; (* (i:=H) (j:=Hx) *)
[env_cbv; reflexivity || fail "iExistDestruct:" H "not found"
|let P := match goal with |- ExistDestruct ?P _ => P end in
|let P := match goal with |- IntoExist ?P _ => P end in
apply _ || fail "iExistDestruct:" H ":" P "not an existential"|];
let y := fresh in
intros y; eexists; split;
......@@ -498,18 +498,18 @@ Tactic Notation "iAlways":=
Tactic Notation "iNext":=
eapply tac_next;
[apply _
|let P := match goal with |- StripLaterL ?P _ => P end in
|let P := match goal with |- FromLater ?P _ => P end in
apply _ || fail "iNext:" P "does not contain laters"|].
(** * Introduction tactic *)
Local Tactic Notation "iIntro" "{" simple_intropattern(x) "}" := first
[ (* (∀ _, _) *) apply tac_forall_intro; intros x
| (* (?P → _) *) eapply tac_impl_intro_pure;
[let P := match goal with |- ToPure ?P _ => P end in
[let P := match goal with |- IsPure ?P _ => P end in
apply _ || fail "iIntro:" P "not pure"
|intros x]
| (* (?P -★ _) *) eapply tac_wand_intro_pure;
[let P := match goal with |- ToPure ?P _ => P end in
[let P := match goal with |- IsPure ?P _ => P end in
apply _ || fail "iIntro:" P "not pure"
|intros x]
|intros x].
......@@ -528,12 +528,12 @@ Local Tactic Notation "iIntro" constr(H) := first
Local Tactic Notation "iIntro" "#" constr(H) := first
[ (* (?P → _) *)
eapply tac_impl_intro_persistent with _ H _; (* (i:=H) *)
[let P := match goal with |- ToPersistentP ?P _ => P end in
[let P := match goal with |- IntoPersistentP ?P _ => P end in
apply _ || fail 1 "iIntro: " P " not persistent"
|env_cbv; reflexivity || fail 1 "iIntro:" H "not fresh"|]
| (* (?P -★ _) *)
eapply tac_wand_intro_persistent with _ H _; (* (i:=H) *)
[let P := match goal with |- ToPersistentP ?P _ => P end in
[let P := match goal with |- IntoPersistentP ?P _ => P end in
apply _ || fail 1 "iIntro: " P " not persistent"
|env_cbv; reflexivity || fail 1 "iIntro:" H "not fresh"|]
| fail 1 "iIntro: nothing to introduce" ].
......@@ -746,7 +746,7 @@ Tactic Notation "iAssert" open_constr(Q) "with" constr(Hs) "as" constr(pat) :=
| [SGoal ?k ?lr ?Hs] =>
eapply tac_assert with _ _ _ lr Hs H Q _; (* (js:=Hs) (j:=H) (P:=Q) *)
[match k with
| GoalStd => apply to_assert_fallthrough
| GoalStd => apply into_assert_fallthrough
| GoalPvs => apply _ || fail "iAssert: cannot generate pvs goal"
end
|env_cbv; reflexivity || fail "iAssert:" Hs "not found"
......
......@@ -11,7 +11,7 @@ Implicit Types Φ : val Λ → iProp Λ Σ.
Global Instance frame_wp E e R Φ Ψ :
( v, Frame R (Φ v) (Ψ v)) Frame R (WP e @ E {{ Φ }}) (WP e @ E {{ Ψ }}).
Proof. rewrite /Frame=> HR. rewrite wp_frame_l. apply wp_mono, HR. Qed.
Global Instance fsa_split_wp E e Φ :
FSASplit (WP e @ E {{ Φ }})%I E (wp_fsa e) (language.atomic e) Φ.
Global Instance is_fsa_wp E e Φ :
IsFSA (WP e @ E {{ Φ }})%I E (wp_fsa e) (language.atomic e) Φ.
Proof. split. done. apply _. Qed.
End weakestpre.
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