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

make iAuIntro smarter

parent 01557046
No related branches found
No related tags found
No related merge requests found
...@@ -456,15 +456,18 @@ Section proof_mode. ...@@ -456,15 +456,18 @@ Section proof_mode.
Proof. rewrite make_laterable_id_eq. done. Qed. Proof. rewrite make_laterable_id_eq. done. Qed.
Lemma tac_aupd_intro Γp Γs n α β Eo Ei Φ P : Lemma tac_aupd_intro Γp Γs n α β Eo Ei Φ P :
Laterable (PROP:=PROP) emp match Γs with Enil => Laterable (PROP:=PROP) emp | _ => TCTrue end
TCForall Laterable (env_to_list Γs) TCForall Laterable (env_to_list Γs)
P = env_to_prop Γs P = env_to_prop Γs
envs_entails (Envs Γp Γs n) (atomic_acc Eo Ei α P β Φ) envs_entails (Envs Γp Γs n) (atomic_acc Eo Ei α P β Φ)
envs_entails (Envs Γp Γs n) (atomic_update Eo Ei α β Φ). envs_entails (Envs Γp Γs n) (atomic_update Eo Ei α β Φ).
Proof. Proof.
intros ? HΓs ->. rewrite envs_entails_unseal of_envs_eq /atomic_acc /=. intros Hemp HΓs ->. rewrite envs_entails_unseal of_envs_eq /atomic_acc /=.
setoid_rewrite env_to_prop_sound =>HAU. setoid_rewrite env_to_prop_sound =>HAU.
rewrite assoc. apply: aupd_intro. by rewrite -assoc. rewrite assoc. apply: aupd_intro.
{ destruct Γs as [|Γs i P]; first done.
inversion HΓs. simpl. apply big_sep_sepL_laterable; done. }
by rewrite -assoc.
Qed. Qed.
End proof_mode. End proof_mode.
...@@ -475,18 +478,23 @@ Local Ltac iMakeLaterable := ...@@ -475,18 +478,23 @@ Local Ltac iMakeLaterable :=
iApply make_laterable_id_elim; iModIntro. iApply make_laterable_id_elim; iModIntro.
Tactic Notation "iAuIntro" := Tactic Notation "iAuIntro" :=
iMakeLaterable; notypeclasses refine (tac_aupd_intro _ _ _ _ _ _ _ _ _ _ _ _ _); [ iMakeLaterable;
iSolveTC || fail "iAuIntro: emp not laterable" match goal with
| iSolveTC || fail "iAuIntro: context not laterable; this should not happen, please report a bug" | |- envs_entails (Envs ?Γp ?Γs _) (atomic_update _ _ _ _ ) =>
| (* P = ...: make the P pretty *) pm_reflexivity notypeclasses refine (tac_aupd_intro Γp Γs _ _ _ _ _ Φ _ _ _ _ _); [
| (* the new proof mode goal *) ]. (* The [match Γs] precondition *)
iSolveTC || fail "iAuIntro: spacial context empty, and emp not laterable"
| iSolveTC || fail "iAuIntro: context not laterable; this should not happen, please report a bug"
| (* P = ...: make the P pretty *) pm_reflexivity
| (* the new proof mode goal *) ]
end.
(** Tactic to apply [aacc_intro]. This only really works well when you have (** Tactic to apply [aacc_intro]. This only really works well when you have
[α ?] already and pass it as [iAaccIntro with "Hα"]. Doing [α ?] already and pass it as [iAaccIntro with "Hα"]. Doing
[rewrite /atomic_acc /=] is an entirely legitimate alternative. *) [rewrite /atomic_acc /=] is an entirely legitimate alternative. *)
Tactic Notation "iAaccIntro" "with" constr(sel) := Tactic Notation "iAaccIntro" "with" constr(sel) :=
iStartProof; lazymatch goal with iStartProof; lazymatch goal with
| |- environments.envs_entails _ (@atomic_acc ?PROP ?H ?TA ?TB ?Eo ?Ei ?P ) => | |- envs_entails _ (@atomic_acc ?PROP ?H ?TA ?TB ?Eo ?Ei ?P ) =>
iApply (@aacc_intro PROP H TA TB Eo Ei α P β Φ with sel); iApply (@aacc_intro PROP H TA TB Eo Ei α P β Φ with sel);
first try solve_ndisj; last iSplit first try solve_ndisj; last iSplit
| _ => fail "iAAccIntro: Goal is not an atomic accessor" | _ => fail "iAAccIntro: Goal is not an atomic accessor"
......
...@@ -88,11 +88,25 @@ Section instances. ...@@ -88,11 +88,25 @@ Section instances.
iExists Q. iIntros "{$HQ} !> HQ". iExists x. by iApply "HΦ". iExists Q. iIntros "{$HQ} !> HQ". iExists x. by iApply "HΦ".
Qed. Qed.
Lemma big_sep_sepL_laterable Q Ps :
Laterable Q
TCForall Laterable Ps
Laterable (Q [] Ps).
Proof.
intros HQ HPs. revert Q HQ. induction HPs as [|P Ps ?? IH]; intros Q HQ.
{ simpl. rewrite right_id. done. }
simpl. rewrite assoc. apply IH; by apply _.
Qed.
Global Instance big_sepL_laterable Ps : Global Instance big_sepL_laterable Ps :
Laterable (PROP:=PROP) emp Laterable (PROP:=PROP) emp
TCForall Laterable Ps TCForall Laterable Ps
Laterable ([] Ps). Laterable ([] Ps).
Proof. induction 2; simpl; apply _. Qed. Proof.
intros. assert (Laterable (emp [] Ps)) as Hlater.
{ apply big_sep_sepL_laterable; done. }
rewrite ->left_id in Hlater; last by apply _. done.
Qed.
(** A wrapper to obtain a weaker, laterable form of any assertion. (** A wrapper to obtain a weaker, laterable form of any assertion.
Alternatively: the modality corresponding to [Laterable]. Alternatively: the modality corresponding to [Laterable].
......
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