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

Enable $-prefixes in ssr-like {H1 .. Hn} intro-patterns to perform framing.

For example iIntros "{$H1 H2} H1" frames H1, clears H2, and introduces H1.
parent f1c83a61
No related branches found
No related tags found
No related merge requests found
......@@ -13,7 +13,7 @@ Inductive intro_pat :=
| INext : intro_pat
| IForall : intro_pat
| IAll : intro_pat
| IClear : list string intro_pat.
| IClear : list (bool * string) intro_pat. (* true = frame, false = clear *)
Module intro_pat.
Inductive token :=
......@@ -71,8 +71,7 @@ Inductive stack_item :=
| SList : stack_item
| SConjList : stack_item
| SBar : stack_item
| SAmp : stack_item
| SClear : stack_item.
| SAmp : stack_item.
Notation stack := (list stack_item).
Fixpoint close_list (k : stack)
......@@ -108,13 +107,6 @@ Fixpoint close_conj_list (k : stack)
| _ => None
end.
Fixpoint close_clear (k : stack) (ss : list string) : option stack :=
match k with
| SPat (IName s) :: k => close_clear k (s :: ss)
| SClear :: k => Some (SPat (IClear (reverse ss)) :: k)
| _ => None
end.
Fixpoint parse_go (ts : list token) (k : stack) : option stack :=
match ts with
| [] => Some k
......@@ -135,9 +127,18 @@ Fixpoint parse_go (ts : list token) (k : stack) : option stack :=
| TNext :: ts => parse_go ts (SPat INext :: k)
| TAll :: ts => parse_go ts (SPat IAll :: k)
| TForall :: ts => parse_go ts (SPat IForall :: k)
| TClearL :: ts => parse_go ts (SClear :: k)
| TClearR :: ts => close_clear k [] ≫= parse_go ts
| TClearL :: ts => parse_clear ts [] k
| _ => None
end
with parse_clear (ts : list token)
(ss : list (bool * string)) (k : stack) : option stack :=
match ts with
| TFrame :: TName s :: ts => parse_clear ts ((true,s) :: ss) k
| TName s :: ts => parse_clear ts ((false,s) :: ss) k
| TClearR :: ts => parse_go ts (SPat (IClear (reverse ss)) :: k)
| _ => None
end.
Definition parse (s : string) : option (list intro_pat) :=
match k parse_go (tokenize s) [SList]; close_list k [] [] with
| Some [SPat (IList [ps])] => Some ps
......
......@@ -560,7 +560,13 @@ Tactic Notation "iIntros" constr(pat) :=
| ISimpl :: ?pats => simpl; go pats
| IAlways :: ?pats => iAlways; go pats
| INext :: ?pats => iNext; go pats
| IClear ?Hs :: ?pats => iClear Hs; go pats
| IClear ?cpats :: ?pats =>
let rec clr cpats :=
match cpats with
| [] => go pats
| (false,?H) :: ?cpats => iClear H; clr cpats
| (true,?H) :: ?cpats => iFrame H; clr cpats
end in clr cpats
| IPersistent (IName ?H) :: ?pats => iIntro #H; go pats
| IName ?H :: ?pats => iIntro H; go pats
| IPersistent IAnom :: ?pats => let H := iFresh in iIntro #H; go pats
......
......@@ -28,7 +28,7 @@ Proof.
iRevert {a b} "Ha Hb". iIntros {b a} "Hb {foo} Ha".
iAssert (uPred_ownM (a core a)) with "[Ha]" as "[Ha #Hac]".
{ by rewrite cmra_core_r. }
iFrame "Ha Hac".
iIntros "{$Hac $Ha}".
iExists (S j + z1), z2.
iNext.
iApply ("H3" $! _ 0 with "H1 []").
......
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