Newer
Older
From iris.proofmode Require Import coq_tactics.
From iris.proofmode Require Import intro_patterns spec_patterns sel_patterns.
From iris.base_logic Require Export base_logic.
From iris.proofmode Require Export classes notation.
From iris.proofmode Require Import class_instances.
From iris.prelude Require Import stringmap hlist.
From iris.proofmode Require Import strings.
beq ascii_beq string_beq
env_lookup env_lookup_delete env_delete env_app env_replace env_dom
env_persistent env_spatial env_spatial_is_nil envs_dom
envs_lookup envs_lookup_delete envs_delete envs_snoc envs_app
envs_simple_replace envs_replace envs_split
envs_clear_spatial envs_clear_persistent
envs_split_go envs_split].
Ltac env_cbv :=
match goal with |- ?u => let v := eval env_cbv in u in change v end.
(* Tactic Notation tactics cannot return terms *)
Ltac iFresh' H :=
(* [vm_compute fails] if any of the hypotheses in [Δ] contain evars, so
first use [cbv] to compute the domain of [Δ] *)
let Hs := eval cbv in (envs_dom Δ) in
eval vm_compute in (fresh_string_of_set H (of_list Hs))
| _ => H
Ltac iFresh := iFresh' "~".
Tactic Notation "iTypeOf" constr(H) tactic(tac):=
let Δ := match goal with |- of_envs ?Δ ⊢ _ => Δ end in
lazymatch eval env_cbv in (envs_lookup H Δ) with
Tactic Notation "iMatchHyp" tactic1(tac) :=
Robbert Krebbers
committed
match goal with
| |- context[ environments.Esnoc _ ?x ?P ] => tac x P
end.
(** * Start a proof *)
Tactic Notation "iProof" :=
lazymatch goal with
| |- of_envs _ ⊢ _ => fail "iProof: already in Iris proofmode"
| |- ?P =>
match eval hnf in P with
Jacques-Henri Jourdan
committed
(* need to use the unfolded version of [uPred_valid] due to the hnf *)
| True ⊢ _ => apply tac_adequate
| _ ⊢ _ => apply uPred.wand_entails, tac_adequate
(* need to use the unfolded version of [⊣⊢] due to the hnf *)
| uPred_equiv' _ _ => apply uPred.iff_equiv, tac_adequate
end
end.
(** * Context manipulation *)
Tactic Notation "iRename" constr(H1) "into" constr(H2) :=
eapply tac_rename with _ H1 H2 _ _; (* (i:=H1) (j:=H2) *)
[env_cbv; reflexivity || fail "iRename:" H1 "not found"
|env_cbv; reflexivity || fail "iRename:" H2 "not fresh"|].
Local Inductive esel_pat :=
| ESelPure
| ESelName : bool → string → esel_pat.
Ltac iElaborateSelPat pat tac :=
let rec go pat Δ Hs :=
lazymatch pat with
| [] => let Hs' := eval cbv in Hs in tac Hs'
| SelPure :: ?pat => go pat Δ (ESelPure :: Hs)
| SelPersistent :: ?pat =>
let Hs' := eval env_cbv in (env_dom (env_persistent Δ)) in
let Δ' := eval env_cbv in (envs_clear_persistent Δ) in
go pat Δ' ((ESelName true <$> Hs') ++ Hs)
| SelSpatial :: ?pat =>
let Hs' := eval env_cbv in (env_dom (env_spatial Δ)) in
let Δ' := eval env_cbv in (envs_clear_spatial Δ) in
go pat Δ' ((ESelName false <$> Hs') ++ Hs)
| SelName ?H :: ?pat =>
lazymatch eval env_cbv in (envs_lookup_delete H Δ) with
| Some (?p,_,?Δ') => go pat Δ' (ESelName p H :: Hs)
| None => fail "iElaborateSelPat:" H "not found"
end
end in
lazymatch goal with
| |- of_envs ?Δ ⊢ _ =>
let pat := sel_pat.parse pat in go pat Δ (@nil esel_pat)
end.
Tactic Notation "iClear" constr(Hs) :=
let rec go Hs :=
lazymatch Hs with
| ESelPure :: ?Hs => clear; go Hs
| ESelName _ ?H :: ?Hs =>
eapply tac_clear with _ H _ _; (* (i:=H) *)
[env_cbv; reflexivity || fail "iClear:" H "not found"|go Hs]
end in
Tactic Notation "iClear" "(" ident_list(xs) ")" constr(Hs) :=
iClear Hs; clear xs.
(** * Assumptions *)
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 |- FromAssumption _ ?P _ => P end in
apply _ || fail "iExact:" H ":" P "does not match goal"].
Tactic Notation "iAssumptionCore" :=
let rec find Γ i P :=
match Γ with
| Esnoc ?Γ ?j ?Q => first [unify P Q; unify i j| find Γ i P]
end in
match goal with
| |- envs_lookup ?i (Envs ?Γp ?Γs) = Some (_, ?P) =>
first [is_evar i; fail 1 | env_cbv; reflexivity]
| |- envs_lookup ?i (Envs ?Γp ?Γs) = Some (_, ?P) =>
is_evar i; first [find Γp i P | find Γs i P]; env_cbv; reflexivity
end.
let Hass := fresh in
let rec find p Γ Q :=
match Γ with
| Esnoc ?Γ ?j ?P => first
[pose proof (_ : FromAssumption p P Q) as Hass;
apply (tac_assumption _ j p P); [env_cbv; reflexivity|apply Hass]
|find p Γ Q]
end in
match goal with
| |- of_envs (Envs ?Γp ?Γs) ⊢ ?Q =>
first [find true Γp Q | find false Γs Q
|fail "iAssumption:" Q "not found"]
end.
(** * False *)
Tactic Notation "iExFalso" := apply tac_ex_falso.
(** * Making hypotheses persistent or pure *)
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 |- IntoPersistentP ?Q _ => Q end in
apply _ || fail "iPersistent:" Q "not persistent"
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 |- IntoPure ?P _ => P end in
apply _ || fail "iPure:" P "not pure"
Tactic Notation "iPureIntro" :=
eapply tac_pure_intro;
[let P := match goal with |- FromPure ?P _ => P end in
apply _ || fail "iPureIntro:" P "not pure"|].
163
164
165
166
167
168
169
170
171
172
173
174
175
176
177
178
179
180
181
182
183
184
185
186
187
188
189
190
191
192
193
194
195
196
197
198
199
200
201
202
203
204
205
206
207
208
209
210
211
212
213
214
215
216
217
218
219
220
221
222
223
224
225
226
227
228
(** Framing *)
Local Ltac iFrameFinish :=
lazy iota beta;
try match goal with
| |- _ ⊢ True => exact (uPred.pure_intro _ _ I)
end.
Local Ltac iFramePure t :=
let φ := type of t in
eapply (tac_frame_pure _ _ _ _ t);
[apply _ || fail "iFrame: cannot frame" φ
|iFrameFinish].
Local Ltac iFrameHyp H :=
eapply tac_frame with _ H _ _ _;
[env_cbv; reflexivity || fail "iFrame:" H "not found"
|let R := match goal with |- Frame ?R _ _ => R end in
apply _ || fail "iFrame: cannot frame" R
|iFrameFinish].
Local Ltac iFrameAnyPure :=
repeat match goal with H : _ |- _ => iFramePure H end.
Local Ltac iFrameAnyPersistent :=
let rec go Hs :=
match Hs with [] => idtac | ?H :: ?Hs => repeat iFrameHyp H; go Hs end in
match goal with
| |- of_envs ?Δ ⊢ _ =>
let Hs := eval cbv in (env_dom (env_persistent Δ)) in go Hs
end.
Local Ltac iFrameAnySpatial :=
let rec go Hs :=
match Hs with [] => idtac | ?H :: ?Hs => try iFrameHyp H; go Hs end in
match goal with
| |- of_envs ?Δ ⊢ _ =>
let Hs := eval cbv in (env_dom (env_spatial Δ)) in go Hs
end.
Tactic Notation "iFrame" := iFrameAnySpatial.
Tactic Notation "iFrame" "(" constr(t1) ")" :=
iFramePure t1.
Tactic Notation "iFrame" "(" constr(t1) constr(t2) ")" :=
iFramePure t1; iFrame ( t2 ).
Tactic Notation "iFrame" "(" constr(t1) constr(t2) constr(t3) ")" :=
iFramePure t1; iFrame ( t2 t3 ).
Tactic Notation "iFrame" "(" constr(t1) constr(t2) constr(t3) constr(t4) ")" :=
iFramePure t1; iFrame ( t2 t3 t4 ).
Tactic Notation "iFrame" "(" constr(t1) constr(t2) constr(t3) constr(t4)
constr(t5) ")" :=
iFramePure t1; iFrame ( t2 t3 t4 t5 ).
Tactic Notation "iFrame" "(" constr(t1) constr(t2) constr(t3) constr(t4)
constr(t5) constr(t6) ")" :=
iFramePure t1; iFrame ( t2 t3 t4 t5 t6 ).
Tactic Notation "iFrame" "(" constr(t1) constr(t2) constr(t3) constr(t4)
constr(t5) constr(t6) constr(t7) ")" :=
iFramePure t1; iFrame ( t2 t3 t4 t5 t6 t7 ).
Tactic Notation "iFrame" "(" constr(t1) constr(t2) constr(t3) constr(t4)
constr(t5) constr(t6) constr(t7) constr(t8)")" :=
iFramePure t1; iFrame ( t2 t3 t4 t5 t6 t7 t8 ).
Tactic Notation "iFrame" constr(Hs) :=
let rec go Hs :=
match Hs with
| [] => idtac
| SelPure :: ?Hs => iFrameAnyPure; go Hs
| SelPersistent :: ?Hs => iFrameAnyPersistent; go Hs
| SelSpatial :: ?Hs => iFrameAnySpatial; go Hs
| SelName ?H :: ?Hs => iFrameHyp H; go Hs
end
in let Hs := sel_pat.parse Hs in go Hs.
Tactic Notation "iFrame" "(" constr(t1) ")" constr(Hs) :=
iFramePure t1; iFrame Hs.
Tactic Notation "iFrame" "(" constr(t1) constr(t2) ")" constr(Hs) :=
iFramePure t1; iFrame ( t2 ) Hs.
Tactic Notation "iFrame" "(" constr(t1) constr(t2) constr(t3) ")" constr(Hs) :=
iFramePure t1; iFrame ( t2 t3 ) Hs.
Tactic Notation "iFrame" "(" constr(t1) constr(t2) constr(t3) constr(t4) ")"
constr(Hs) :=
iFramePure t1; iFrame ( t2 t3 t4 ) Hs.
Tactic Notation "iFrame" "(" constr(t1) constr(t2) constr(t3) constr(t4)
constr(t5) ")" constr(Hs) :=
iFramePure t1; iFrame ( t2 t3 t4 t5 ) Hs.
Tactic Notation "iFrame" "(" constr(t1) constr(t2) constr(t3) constr(t4)
constr(t5) constr(t6) ")" constr(Hs) :=
iFramePure t1; iFrame ( t2 t3 t4 t5 t6 ) Hs.
Tactic Notation "iFrame" "(" constr(t1) constr(t2) constr(t3) constr(t4)
constr(t5) constr(t6) constr(t7) ")" constr(Hs) :=
iFramePure t1; iFrame ( t2 t3 t4 t5 t6 t7 ) Hs.
Tactic Notation "iFrame" "(" constr(t1) constr(t2) constr(t3) constr(t4)
constr(t5) constr(t6) constr(t7) constr(t8)")" constr(Hs) :=
iFramePure t1; iFrame ( t2 t3 t4 t5 t6 t7 t8 ) Hs.
Record iTrm {X As} :=
ITrm { itrm : X ; itrm_vars : hlist As ; itrm_hyps : string }.
Arguments ITrm {_ _} _ _ _.
Notation "( H $! x1 .. xn )" :=
(ITrm H (hcons x1 .. (hcons xn hnil) ..) "") (at level 0, x1, xn at level 9).
Notation "( H $! x1 .. xn 'with' pat )" :=
(ITrm H (hcons x1 .. (hcons xn hnil) ..) pat) (at level 0, x1, xn at level 9).
Notation "( H 'with' pat )" := (ITrm H hnil pat) (at level 0).
Local Tactic Notation "iSpecializeArgs" constr(H) open_constr(xs) :=
match xs with
| hnil => idtac
| _ =>
eapply tac_forall_specialize with _ H _ _ _ xs; (* (i:=H) (a:=x) *)
[env_cbv; reflexivity || fail 1 "iSpecialize:" H "not found"
|let P := match goal with |- ForallSpecialize _ ?P _ => P end in
apply _ || fail 1 "iSpecialize:" P "not a forall of the right arity or type"
Local Tactic Notation "iSpecializePat" constr(H) constr(pat) :=
let P := match goal with |- IntoWand ?P _ _ => P end in
apply _ || fail "iSpecialize:" P "not an implication/wand" in
let rec go H1 pats :=
lazymatch pats with
| [] => idtac
| SForall :: ?pats => try (iSpecializeArgs H1 (hcons _ _)); go H1 pats
| SName ?H2 :: ?pats =>
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 |- IntoWand ?P ?Q _ => P end in
let Q := match goal with |- IntoWand ?P ?Q _ => Q end in
apply _ || fail "iSpecialize: cannot instantiate" P "with" Q
| SGoalPersistent :: ?pats =>
eapply tac_specialize_assert_persistent with _ _ H1 _ _ _ _;
[env_cbv; reflexivity || fail "iSpecialize:" H1 "not found"
|solve_to_wand H1
|let Q := match goal with |- PersistentP ?Q => Q end in
apply _ || fail "iSpecialize:" Q "not persistent"
|env_cbv; reflexivity
|(*goal*)
|go H1 pats]
| SGoalPure :: ?pats =>
eapply tac_specialize_assert_pure with _ H1 _ _ _ _ _;
[env_cbv; reflexivity || fail "iSpecialize:" H1 "not found"
|solve_to_wand H1
|let Q := match goal with |- FromPure ?Q _ => Q end in
apply _ || fail "iSpecialize:" Q "not pure"
|env_cbv; reflexivity
|(*goal*)
|go H1 pats]
| SGoal (SpecGoal ?m ?lr ?Hs_frame ?Hs) :: ?pats =>
let Hs' := eval cbv in (if lr then Hs else Hs_frame ++ Hs) in
eapply tac_specialize_assert with _ _ _ H1 _ lr Hs' _ _ _ _;
[env_cbv; reflexivity || fail "iSpecialize:" H1 "not found"
|solve_to_wand H1
|match m with
| false => apply elim_modal_dummy
| true => apply _ || fail "iSpecialize: goal not a modality"
|env_cbv; reflexivity || fail "iSpecialize:" Hs "not found"
|iFrame Hs_frame (*goal*)
end in let pats := spec_pat.parse pat in go H pats.
(* p = whether the conclusion of the specialized term is persistent. It can
either be a Boolean or an introduction pattern, which will be coerced in true
when it only contains `#` or `%` patterns at the top-level. *)
Tactic Notation "iSpecializeCore" open_constr(t) "as" constr(p) tactic(tac) :=
let p :=
lazymatch type of p with
| intro_pat => eval cbv in (intro_pat_persistent p)
| string =>
let pat := intro_pat.parse_one p in
eval cbv in (intro_pat_persistent pat)
| _ => p
end in
lazymatch t with
| ITrm ?H ?xs ?pat =>
lazymatch p with
| true =>
eapply tac_specialize_persistent_helper with _ H _ _ _;
[env_cbv; reflexivity || fail "iSpecialize:" H "not found"
|iSpecializeArgs H xs; iSpecializePat H pat; last (iExact H)
|let Q := match goal with |- PersistentP ?Q => Q end in
apply _ || fail "iSpecialize:" Q "not persistent"
|env_cbv; reflexivity|tac H]
| false => iSpecializeArgs H xs; iSpecializePat H pat; last (tac H)
end
Tactic Notation "iSpecialize" open_constr(t) :=
iSpecializeCore t as false (fun _ => idtac).
Tactic Notation "iSpecialize" open_constr(t) "#" :=
iSpecializeCore t as true (fun _ => idtac).
Jacques-Henri Jourdan
committed
(* The tactic [iIntoValid] tactic solves a goal [uPred_valid Q]. The
arguments [t] is a Coq term whose type is of the following shape:
Jacques-Henri Jourdan
committed
- [∀ (x_1 : A_1) .. (x_n : A_n), uPred_valid Q]
- [∀ (x_1 : A_1) .. (x_n : A_n), P1 ⊢ P2], in which case [Q] becomes [P1 -∗ P2]
- [∀ (x_1 : A_1) .. (x_n : A_n), P1 ⊣⊢ P2], in which case [Q] becomes [P1 ↔ P2]
The tactic instantiates each dependent argument [x_i] with an evar and generates
a goal [P] for non-dependent arguments [x_i : P]. *)
Jacques-Henri Jourdan
committed
Tactic Notation "iIntoValid" open_constr(t) :=
let tT := type of t in
lazymatch eval hnf in tT with
| True ⊢ _ => apply t
| _ ⊢ _ => apply (uPred.entails_wand _ _ t)
(* need to use the unfolded version of [⊣⊢] due to the hnf *)
| uPred_equiv' _ _ => apply (uPred.equiv_iff _ _ t)
| ?P → ?Q => let H := fresh in assert P as H; [|go uconstr:(t H); clear H]
| ∀ _ : ?T, _ =>
(* Put [T] inside an [id] to avoid TC inference from being invoked. *)
(* This is a workarround for Coq bug #4969. *)
let e := fresh in evar (e:id T);
let e' := eval unfold e in e in clear e; go (t e')
end
in go t.
Tactic Notation "iPoseProofCore" open_constr(lem) "as" constr(p) tactic(tac) :=
let pose_trm t tac :=
let Htmp := iFresh in
lazymatch type of t with
| string =>
eapply tac_pose_proof_hyp with _ _ t _ Htmp _;
[env_cbv; reflexivity || fail "iPoseProof:" t "not found"
|env_cbv; reflexivity || fail "iPoseProof:" Htmp "not fresh"
|tac Htmp]
| _ =>
eapply tac_pose_proof with _ Htmp _; (* (j:=H) *)
Jacques-Henri Jourdan
committed
[iIntoValid t
|env_cbv; reflexivity || fail "iPoseProof:" Htmp "not fresh"
|tac Htmp]
end;
try (apply _) (* solve TC constraints. It is essential that this happens
after the continuation [tac] has been called. *)
in lazymatch lem with
| ITrm ?t ?xs ?pat =>
pose_trm t ltac:(fun Htmp => iSpecializeCore (ITrm Htmp xs pat) as p tac)
Tactic Notation "iPoseProof" open_constr(lem) "as" constr(H) :=
iPoseProofCore lem as false (fun Htmp => iRename Htmp into H).
Tactic Notation "iApply" open_constr(lem) :=
let finish H := first
[iExact H
|eapply tac_apply with _ H _ _ _;
[env_cbv; reflexivity || fail 1 "iApply:" H "not found"
|let P := match goal with |- IntoWand ?P _ _ => P end in
apply _ || fail 1 "iApply: cannot apply" P
|lazy beta (* reduce betas created by instantiation *)]] in
lazymatch lem with
| ITrm ?t ?xs ?pat =>
iPoseProofCore t as false (fun Htmp =>
iSpecializeArgs Htmp xs;
try (iSpecializeArgs Htmp (hcons _ _));
iSpecializePat Htmp pat; last finish Htmp)
iPoseProofCore lem as false (fun Htmp =>
try (iSpecializeArgs Htmp (hcons _ _));
Local Tactic Notation "iForallRevert" ident(x) :=
let err x :=
intros x;
iMatchHyp (fun H P =>
lazymatch P with
| context [x] => fail 2 "iRevert:" x "is used in hypothesis" H
end) in
let A := type of x in
lazymatch type of A with
| Prop => revert x; first [apply tac_pure_revert|err x]
| _ => revert x; first [apply tac_forall_revert|err x]
end.
let rec go Hs :=
lazymatch Hs with
| ESelPure :: ?Hs =>
repeat match goal with x : _ |- _ => revert x end;
go Hs
| ESelName _ ?H :: ?Hs =>
eapply tac_revert with _ H _ _; (* (i:=H2) *)
[env_cbv; reflexivity || fail "iRevert:" H "not found"
|env_cbv; go Hs]
Tactic Notation "iRevert" "(" ident(x1) ")" :=
462
463
464
465
466
467
468
469
470
471
472
473
474
475
476
477
478
479
480
481
482
483
484
485
486
487
Tactic Notation "iRevert" "(" ident(x1) ident(x2) ")" :=
iForallRevert x2; iRevert ( x1 ).
Tactic Notation "iRevert" "(" ident(x1) ident(x2) ident(x3) ")" :=
iForallRevert x3; iRevert ( x1 x2 ).
Tactic Notation "iRevert" "(" ident(x1) ident(x2) ident(x3) ident(x4) ")" :=
iForallRevert x4; iRevert ( x1 x2 x3 ).
Tactic Notation "iRevert" "(" ident(x1) ident(x2) ident(x3) ident(x4)
ident(x5) ")" :=
iForallRevert x5; iRevert ( x1 x2 x3 x4 ).
Tactic Notation "iRevert" "(" ident(x1) ident(x2) ident(x3) ident(x4)
ident(x5) ident(x6) ")" :=
iForallRevert x6; iRevert ( x1 x2 x3 x4 x5 ).
Tactic Notation "iRevert" "(" ident(x1) ident(x2) ident(x3) ident(x4)
ident(x5) ident(x6) ident(x7) ")" :=
iForallRevert x7; iRevert ( x1 x2 x3 x4 x5 x6 ).
Tactic Notation "iRevert" "(" ident(x1) ident(x2) ident(x3) ident(x4)
ident(x5) ident(x6) ident(x7) ident(x8) ")" :=
iForallRevert x8; iRevert ( x1 x2 x3 x4 x5 x6 x7 ).
Tactic Notation "iRevert" "(" ident(x1) ")" constr(Hs) :=
iRevert Hs; iRevert ( x1 ).
Tactic Notation "iRevert" "(" ident(x1) ident(x2) ")" constr(Hs) :=
iRevert Hs; iRevert ( x1 x2 ).
Tactic Notation "iRevert" "(" ident(x1) ident(x2) ident(x3) ")" constr(Hs) :=
iRevert Hs; iRevert ( x1 x2 x3 ).
Tactic Notation "iRevert" "(" ident(x1) ident(x2) ident(x3) ident(x4) ")"
iRevert Hs; iRevert ( x1 x2 x3 x4 ).
Tactic Notation "iRevert" "(" ident(x1) ident(x2) ident(x3) ident(x4)
ident(x5) ")" constr(Hs) :=
iRevert Hs; iRevert ( x1 x2 x3 x4 x5 ).
Tactic Notation "iRevert" "(" ident(x1) ident(x2) ident(x3) ident(x4)
ident(x5) ident(x6) ")" constr(Hs) :=
iRevert Hs; iRevert ( x1 x2 x3 x4 x5 x6 ).
Tactic Notation "iRevert" "(" ident(x1) ident(x2) ident(x3) ident(x4)
ident(x5) ident(x6) ident(x7) ")" constr(Hs) :=
iRevert Hs; iRevert ( x1 x2 x3 x4 x5 x6 x7 ).
Tactic Notation "iRevert" "(" ident(x1) ident(x2) ident(x3) ident(x4)
ident(x5) ident(x6) ident(x7) ident(x8) ")" constr(Hs) :=
iRevert Hs; iRevert ( x1 x2 x3 x4 x5 x6 x7 x8 ).
(** * Disjunction *)
Tactic Notation "iLeft" :=
eapply tac_or_l;
[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 |- 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 |- IntoOr ?P _ _ => P end in
Robbert Krebbers
committed
apply _ || fail "iOrDestruct: cannot destruct" P
|env_cbv; reflexivity || fail "iOrDestruct:" H1 "not fresh"
|env_cbv; reflexivity || fail "iOrDestruct:" H2 "not fresh"| |].
(** * Conjunction and separating conjunction *)
Tactic Notation "iSplit" :=
lazymatch goal with
| |- _ ⊢ _ =>
eapply tac_and_split;
[let P := match goal with |- FromAnd ?P _ _ => P end in
apply _ || fail "iSplit:" P "not a conjunction"| |]
end.
Tactic Notation "iSplitL" constr(Hs) :=
let Hs := words Hs in
eapply tac_sep_split with _ _ false Hs _ _; (* (js:=Hs) *)
[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 in the context"| |].
Tactic Notation "iSplitR" constr(Hs) :=
let Hs := words Hs in
eapply tac_sep_split with _ _ true Hs _ _; (* (js:=Hs) *)
[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 in the context"| |].
Tactic Notation "iSplitL" := iSplitR "".
Tactic Notation "iSplitR" := iSplitL "".
Robbert Krebbers
committed
Local Tactic Notation "iAndDestruct" constr(H) "as" constr(H1) constr(H2) :=
eapply tac_and_destruct with _ H _ H1 H2 _ _ _; (* (i:=H) (j1:=H1) (j2:=H2) *)
[env_cbv; reflexivity || fail "iAndDestruct:" H "not found"
|let P := match goal with |- IntoAnd _ ?P _ _ => P end in
apply _ || fail "iAndDestruct: cannot destruct" P
|env_cbv; reflexivity || fail "iAndDestruct:" H1 "or" H2 " not fresh"|].
Local Tactic Notation "iAndDestructChoice" constr(H) "as" constr(lr) constr(H') :=
eapply tac_and_destruct_choice with _ H _ lr H' _ _ _;
[env_cbv; reflexivity || fail "iAndDestruct:" H "not found"
|let P := match goal with |- IntoAnd _ ?P _ _ => P end in
apply _ || fail "iAndDestruct: cannot destruct" P
|env_cbv; reflexivity || fail "iAndDestruct:" H' " not fresh"|].
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 |- FromSep _ ?P1 _ => P1 end in
let P2 := match goal with |- FromSep _ _ ?P2 => P2 end in
apply _ || fail "iCombine: cannot combine" P1 "and" P2
|env_cbv; reflexivity || fail "iCombine:" H "not fresh"|].
(** * Existential *)
Tactic Notation "iExists" uconstr(x1) :=
eapply tac_exist;
[let P := match goal with |- FromExist ?P _ => P end in
Tactic Notation "iExists" uconstr(x1) "," uconstr(x2) :=
Tactic Notation "iExists" uconstr(x1) "," uconstr(x2) "," uconstr(x3) :=
Tactic Notation "iExists" uconstr(x1) "," uconstr(x2) "," uconstr(x3) ","
uconstr(x4) :=
Tactic Notation "iExists" uconstr(x1) "," uconstr(x2) "," uconstr(x3) ","
uconstr(x4) "," uconstr(x5) :=
Tactic Notation "iExists" uconstr(x1) "," uconstr(x2) "," uconstr(x3) ","
uconstr(x4) "," uconstr(x5) "," uconstr(x6) :=
Tactic Notation "iExists" uconstr(x1) "," uconstr(x2) "," uconstr(x3) ","
uconstr(x4) "," uconstr(x5) "," uconstr(x6) "," uconstr(x7) :=
Tactic Notation "iExists" uconstr(x1) "," uconstr(x2) "," uconstr(x3) ","
uconstr(x4) "," uconstr(x5) "," uconstr(x6) "," uconstr(x7) ","
uconstr(x8) :=
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 |- IntoExist ?P _ => P end in
Robbert Krebbers
committed
apply _ || fail "iExistDestruct: cannot destruct" P|];
let y := fresh in
intros y; eexists; split;
[env_cbv; reflexivity || fail "iExistDestruct:" Hx "not fresh"
|revert y; intros x].
(** * Always *)
Tactic Notation "iAlways":=
apply tac_always_intro;
[reflexivity || fail "iAlways: spatial context non-empty"|].
(** * Later *)
Tactic Notation "iNext" open_constr(n) :=
let P := match goal with |- _ ⊢ ?P => P end in
try lazymatch n with 0 => fail 1 "iNext: cannot strip 0 laters" end;
eapply (tac_next _ _ n);
[apply _ || fail "iNext:" P "does not contain" n "laters"
|lazymatch goal with
| |- IntoLaterNEnvs 0 _ _ => fail "iNext:" P "does not contain laters"
| _ => apply _
end|].
Tactic Notation "iNext":= iNext _.
(** * Update modality *)
Tactic Notation "iModIntro" :=
eapply tac_modal_intro;
[let P := match goal with |- IntoModal _ ?P => P end in
apply _ || fail "iModIntro:" P "not a modality"|].
Tactic Notation "iModCore" constr(H) :=
eapply tac_modal_elim with _ H _ _ _ _;
[env_cbv; reflexivity || fail "iMod:" H "not found"
|let P := match goal with |- ElimModal ?P _ _ _ => P end in
let Q := match goal with |- ElimModal _ _ ?Q _ => Q end in
apply _ || fail "iMod: cannot eliminate modality " P "in" Q
Local Tactic Notation "iDestructHyp" constr(H) "as" constr(pat) :=
let rec go Hz pat :=
lazymatch pat with
| IAnom => idtac
| IDrop => iClear Hz
| IList [[]] => iExFalso; iExact Hz
Robbert Krebbers
committed
| IList [[?pat1; IDrop]] => iAndDestructChoice Hz as true Hz; go Hz pat1
| IList [[IDrop; ?pat2]] => iAndDestructChoice Hz as false Hz; go Hz pat2
Robbert Krebbers
committed
let Hy := iFresh in iAndDestruct Hz as Hz Hy; go Hz pat1; go Hy pat2
| IList [[?pat1];[?pat2]] => iOrDestruct Hz as Hz Hz; [go Hz pat1|go Hz pat2]
| IPureElim => iPure Hz as ?
| IAlwaysElim ?pat => iPersistent Hz; go Hz pat
| IModalElim ?pat => iModCore Hz; go Hz pat
| _ => fail "iDestruct:" pat "invalid"
end
in let pat := intro_pat.parse_one pat in go H pat.
Local Tactic Notation "iDestructHyp" constr(H) "as" "(" simple_intropattern(x1) ")"
constr(pat) :=
iExistDestruct H as x1 H; iDestructHyp H as @ pat.
Local Tactic Notation "iDestructHyp" constr(H) "as" "(" simple_intropattern(x1)
simple_intropattern(x2) ")" constr(pat) :=
iExistDestruct H as x1 H; iDestructHyp H as ( x2 ) pat.
Local Tactic Notation "iDestructHyp" constr(H) "as" "(" simple_intropattern(x1)
simple_intropattern(x2) simple_intropattern(x3) ")" constr(pat) :=
iExistDestruct H as x1 H; iDestructHyp H as ( x2 x3 ) pat.
Local Tactic Notation "iDestructHyp" constr(H) "as" "(" simple_intropattern(x1)
simple_intropattern(x2) simple_intropattern(x3) simple_intropattern(x4) ")"
iExistDestruct H as x1 H; iDestructHyp H as ( x2 x3 x4 ) pat.
Local Tactic Notation "iDestructHyp" constr(H) "as" "(" simple_intropattern(x1)
simple_intropattern(x2) simple_intropattern(x3) simple_intropattern(x4)
simple_intropattern(x5) ")" constr(pat) :=
iExistDestruct H as x1 H; iDestructHyp H as ( x2 x3 x4 x5 ) pat.
Local Tactic Notation "iDestructHyp" constr(H) "as" "(" simple_intropattern(x1)
simple_intropattern(x2) simple_intropattern(x3) simple_intropattern(x4)
simple_intropattern(x5) simple_intropattern(x6) ")" constr(pat) :=
iExistDestruct H as x1 H; iDestructHyp H as ( x2 x3 x4 x5 x6 ) pat.
Local Tactic Notation "iDestructHyp" constr(H) "as" "(" simple_intropattern(x1)
simple_intropattern(x2) simple_intropattern(x3) simple_intropattern(x4)
simple_intropattern(x5) simple_intropattern(x6) simple_intropattern(x7) ")"
iExistDestruct H as x1 H; iDestructHyp H as ( x2 x3 x4 x5 x6 x7 ) pat.
Local Tactic Notation "iDestructHyp" constr(H) "as" "(" simple_intropattern(x1)
simple_intropattern(x2) simple_intropattern(x3) simple_intropattern(x4)
simple_intropattern(x5) simple_intropattern(x6) simple_intropattern(x7)
simple_intropattern(x8) ")" constr(pat) :=
iExistDestruct H as x1 H; iDestructHyp H as ( x2 x3 x4 x5 x6 x7 x8 ) pat.
Local Tactic Notation "iIntro" "(" simple_intropattern(x) ")" :=
try first
[(* (∀ _, _) *) apply tac_forall_intro
|(* (?P → _) *) eapply tac_impl_intro_pure;
[let P := match goal with |- IntoPure ?P _ => P end in
apply _ || fail "iIntro:" P "not pure"|]
|(* (?P -∗ _) *) eapply tac_wand_intro_pure;
[let P := match goal with |- IntoPure ?P _ => P end in
apply _ || fail "iIntro:" P "not pure"|]
|(* ⌜∀ _, _⌝ *) apply tac_pure_forall_intro
|(* ⌜_ → _⌝ *) apply tac_pure_impl_intro];
Local Tactic Notation "iIntro" constr(H) := first
[ (* (?Q → _) *)
eapply tac_impl_intro with _ H; (* (i:=H) *)
[reflexivity || fail 1 "iIntro: introducing" H
"into non-empty spatial context"
|env_cbv; reflexivity || fail "iIntro:" H "not fresh"|]
eapply tac_wand_intro with _ H; (* (i:=H) *)
[env_cbv; reflexivity || fail 1 "iIntro:" H "not fresh"|]
| fail 1 "iIntro: nothing to introduce" ].
Local Tactic Notation "iIntro" "#" constr(H) := first
[ (* (?P → _) *)
eapply tac_impl_intro_persistent with _ H _; (* (i:=H) *)
[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"|]
eapply tac_wand_intro_persistent with _ H _; (* (i:=H) *)
[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" ].
Local Tactic Notation "iIntroForall" :=
lazymatch goal with
| |- ∀ _, ?P => fail
| |- ∀ _, _ => intro
| |- _ ⊢ (∀ x : _, _) => iIntro (x)
end.
Local Tactic Notation "iIntro" :=
lazymatch goal with
| |- _ → ?P => intro
| |- _ ⊢ (_ -∗ _) => iIntro (?) || let H := iFresh in iIntro #H || iIntro H
| |- _ ⊢ (_ → _) => iIntro (?) || let H := iFresh in iIntro #H || iIntro H
Tactic Notation "iIntros" constr(pat) :=
let rec go pats :=
lazymatch pats with
| [] => idtac
| IPureElim :: ?pats => iIntro (?); go pats
| IAlwaysElim IAnom :: ?pats => let H := iFresh in iIntro #H; go pats
| IAnom :: ?pats => let H := iFresh in iIntro H; go pats
| IAlwaysElim (IName ?H) :: ?pats => iIntro #H; go pats
| IName ?H :: ?pats => iIntro H; go pats
| IPureIntro :: ?pats => iPureIntro; go pats
| IAlwaysIntro :: ?pats => iAlways; go pats
| IModalIntro :: ?pats => iModIntro; go pats
| IForall :: ?pats => repeat iIntroForall; go pats
| IAll :: ?pats => repeat (iIntroForall || iIntro); go pats
Robbert Krebbers
committed
| 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
let H := iFresh in iIntro #H; iDestructHyp H as pat; go pats
| ?pat :: ?pats =>
let H := iFresh in iIntro H; iDestructHyp H as pat; go pats
end
in let pats := intro_pat.parse pat in try iProof; go pats.
Tactic Notation "iIntros" := iIntros [IAll].
Tactic Notation "iIntros" "(" simple_intropattern(x1) ")" :=
try iProof; iIntro ( x1 ).
Tactic Notation "iIntros" "(" simple_intropattern(x1)
simple_intropattern(x2) ")" :=
iIntros ( x1 ); iIntro ( x2 ).
Tactic Notation "iIntros" "(" simple_intropattern(x1) simple_intropattern(x2)
simple_intropattern(x3) ")" :=
iIntros ( x1 x2 ); iIntro ( x3 ).
Tactic Notation "iIntros" "(" simple_intropattern(x1) simple_intropattern(x2)
simple_intropattern(x3) simple_intropattern(x4) ")" :=
iIntros ( x1 x2 x3 ); iIntro ( x4 ).
Tactic Notation "iIntros" "(" simple_intropattern(x1) simple_intropattern(x2)
simple_intropattern(x3) simple_intropattern(x4) simple_intropattern(x5) ")" :=
iIntros ( x1 x2 x3 x4 ); iIntro ( x5 ).
Tactic Notation "iIntros" "(" simple_intropattern(x1) simple_intropattern(x2)
simple_intropattern(x3) simple_intropattern(x4) simple_intropattern(x5)
simple_intropattern(x6) ")" :=
iIntros ( x1 x2 x3 x4 x5 ); iIntro ( x6 ).
Tactic Notation "iIntros" "(" simple_intropattern(x1) simple_intropattern(x2)
simple_intropattern(x3) simple_intropattern(x4) simple_intropattern(x5)
simple_intropattern(x6) simple_intropattern(x7) ")" :=
iIntros ( x1 x2 x3 x4 x5 x6 ); iIntro ( x7 ).
Tactic Notation "iIntros" "(" simple_intropattern(x1) simple_intropattern(x2)
simple_intropattern(x3) simple_intropattern(x4) simple_intropattern(x5)
simple_intropattern(x6) simple_intropattern(x7) simple_intropattern(x8) ")" :=
iIntros ( x1 x2 x3 x4 x5 x6 x7 ); iIntro ( x8 ).
Tactic Notation "iIntros" "(" simple_intropattern(x1) ")" constr(p) :=
iIntros ( x1 ); iIntros p.
Tactic Notation "iIntros" "(" simple_intropattern(x1) simple_intropattern(x2)
")" constr(p) :=
iIntros ( x1 x2 ); iIntros p.
Tactic Notation "iIntros" "(" simple_intropattern(x1) simple_intropattern(x2)
simple_intropattern(x3) ")" constr(p) :=
iIntros ( x1 x2 x3 ); iIntros p.
Tactic Notation "iIntros" "(" simple_intropattern(x1) simple_intropattern(x2)
simple_intropattern(x3) simple_intropattern(x4) ")" constr(p) :=
iIntros ( x1 x2 x3 x4 ); iIntros p.
Tactic Notation "iIntros" "(" simple_intropattern(x1) simple_intropattern(x2)
simple_intropattern(x3) simple_intropattern(x4) simple_intropattern(x5)
")" constr(p) :=
iIntros ( x1 x2 x3 x4 x5 ); iIntros p.
Tactic Notation "iIntros" "("simple_intropattern(x1) simple_intropattern(x2)
simple_intropattern(x3) simple_intropattern(x4) simple_intropattern(x5)
simple_intropattern(x6) ")" constr(p) :=
iIntros ( x1 x2 x3 x4 x5 x6 ); iIntros p.
Tactic Notation "iIntros" "(" simple_intropattern(x1) simple_intropattern(x2)
simple_intropattern(x3) simple_intropattern(x4) simple_intropattern(x5)
simple_intropattern(x6) simple_intropattern(x7) ")" constr(p) :=
iIntros ( x1 x2 x3 x4 x5 x6 x7 ); iIntros p.
Tactic Notation "iIntros" "(" simple_intropattern(x1) simple_intropattern(x2)
simple_intropattern(x3) simple_intropattern(x4) simple_intropattern(x5)
simple_intropattern(x6) simple_intropattern(x7) simple_intropattern(x8)
")" constr(p) :=
iIntros ( x1 x2 x3 x4 x5 x6 x7 x8 ); iIntros p.
(* Used for generalization in iInduction and iLöb *)
Tactic Notation "iRevertIntros" constr(Hs) "with" tactic(tac) :=
let rec go Hs :=
lazymatch Hs with
| [] => tac
| ESelPure :: ?Hs => fail "iRevertIntros: % not supported"
| ESelName ?p ?H :: ?Hs =>
iRevert H; go Hs;
let H' :=
match p with true => constr:([IAlwaysElim (IName H)]) | false => H end in
843
844
845
846
847
848
849
850
851
852
853
854
855
856
857
858
859
860
861
862
863
864
865
866
867
868
869
870
871
872
873
iIntros H'
end in
iElaborateSelPat Hs go.
Tactic Notation "iRevertIntros" "(" ident(x1) ")" constr(Hs) "with" tactic(tac):=
iRevertIntros Hs with (iRevert (x1); tac; iIntros (x1)).
Tactic Notation "iRevertIntros" "(" ident(x1) ident(x2) ")" constr(Hs)
"with" tactic(tac):=
iRevertIntros Hs with (iRevert (x1 x2); tac; iIntros (x1 x2)).
Tactic Notation "iRevertIntros" "(" ident(x1) ident(x2) ident(x3) ")" constr(Hs)
"with" tactic(tac):=
iRevertIntros Hs with (iRevert (x1 x2 x3); tac; iIntros (x1 x2 x3)).
Tactic Notation "iRevertIntros" "(" ident(x1) ident(x2) ident(x3) ident(x4) ")"
constr(Hs) "with" tactic(tac):=
iRevertIntros Hs with (iRevert (x1 x2 x3 x4); tac; iIntros (x1 x2 x3 x4)).
Tactic Notation "iRevertIntros" "(" ident(x1) ident(x2) ident(x3) ident(x4)
ident(x5) ")" constr(Hs) "with" tactic(tac):=
iRevertIntros Hs with (iRevert (x1 x2 x3 x4 x5); tac; iIntros (x1 x2 x3 x4 x5)).
Tactic Notation "iRevertIntros" "(" ident(x1) ident(x2) ident(x3) ident(x4)
ident(x5) ident(x6) ")" constr(Hs) "with" tactic(tac):=
iRevertIntros Hs with (iRevert (x1 x2 x3 x4 x5 x6);
tac; iIntros (x1 x2 x3 x4 x5 x6)).
Tactic Notation "iRevertIntros" "(" ident(x1) ident(x2) ident(x3) ident(x4)
ident(x5) ident(x6) ident(x7) ")" constr(Hs) "with" tactic(tac):=
iRevertIntros Hs with (iRevert (x1 x2 x3 x4 x5 x6 x7);
tac; iIntros (x1 x2 x3 x4 x5 x6 x7)).
Tactic Notation "iRevertIntros" "(" ident(x1) ident(x2) ident(x3) ident(x4)
ident(x5) ident(x6) ident(x7) ident(x8) ")" constr(Hs) "with" tactic(tac):=
iRevertIntros Hs with (iRevert (x1 x2 x3 x4 x5 x6 x7 x8);
tac; iIntros (x1 x2 x3 x4 x5 x6 x7 x8)).
Tactic Notation "iRevertIntros" "with" tactic(tac) :=
Tactic Notation "iRevertIntros" "(" ident(x1) ")" "with" tactic(tac):=
iRevertIntros (x1) "" with tac.
Tactic Notation "iRevertIntros" "(" ident(x1) ident(x2) ")" "with" tactic(tac):=
iRevertIntros (x1 x2) "" with tac.
Tactic Notation "iRevertIntros" "(" ident(x1) ident(x2) ident(x3) ")"
"with" tactic(tac):=
iRevertIntros (x1 x2 x3) "" with tac.
Tactic Notation "iRevertIntros" "(" ident(x1) ident(x2) ident(x3) ident(x4) ")"
"with" tactic(tac):=
iRevertIntros (x1 x2 x3 x4) "" with tac.
Tactic Notation "iRevertIntros" "(" ident(x1) ident(x2) ident(x3) ident(x4)
ident(x5) ")" "with" tactic(tac):=
iRevertIntros (x1 x2 x3 x4 x5) "" with tac.
Tactic Notation "iRevertIntros" "(" ident(x1) ident(x2) ident(x3) ident(x4)
ident(x5) ident(x6) ")" "with" tactic(tac):=
iRevertIntros (x1 x2 x3 x4 x5 x6) "" with tac.
Tactic Notation "iRevertIntros" "(" ident(x1) ident(x2) ident(x3) ident(x4)
ident(x5) ident(x6) ident(x7) ")" "with" tactic(tac):=
iRevertIntros (x1 x2 x3 x4 x5 x6 x7) "" with tac.
Tactic Notation "iRevertIntros" "(" ident(x1) ident(x2) ident(x3) ident(x4)
ident(x5) ident(x6) ident(x7) ident(x8) ")" "with" tactic(tac):=
iRevertIntros (x1 x2 x3 x4 x5 x6 x7 x8) "" with tac.
Tactic Notation "iDestructCore" open_constr(lem) "as" constr(p) tactic(tac) :=
let intro_destruct n :=
let rec go n' :=
lazymatch n' with
| 0 => fail "iDestruct: cannot introduce" n "hypotheses"
| 1 => repeat iIntroForall; let H := iFresh in iIntro H; tac H
| S ?n' => repeat iIntroForall; let H := iFresh in iIntro H; go n'
end in intros; try iProof; go n in
lazymatch type of lem with
| nat => intro_destruct lem
| Z => (* to make it work in Z_scope. We should just be able to bind
tactic notation arguments to notation scopes. *)
let n := eval compute in (Z.to_nat lem) in intro_destruct n
| string => tac lem
| iTrm =>
(* only copy the hypothesis when universal quantifiers are instantiated *)
| @iTrm string ?H _ hnil ?pat => iSpecializeCore lem as p tac
| _ => iPoseProofCore lem as p tac
| _ => iPoseProofCore lem as p tac
Tactic Notation "iDestruct" open_constr(lem) "as" constr(pat) :=
iDestructCore lem as pat (fun H => iDestructHyp H as pat).
Tactic Notation "iDestruct" open_constr(lem) "as" "(" simple_intropattern(x1) ")"
iDestructCore lem as pat (fun H => iDestructHyp H as ( x1 ) pat).
Tactic Notation "iDestruct" open_constr(lem) "as" "(" simple_intropattern(x1)
simple_intropattern(x2) ")" constr(pat) :=
iDestructCore lem as pat (fun H => iDestructHyp H as ( x1 x2 ) pat).
Tactic Notation "iDestruct" open_constr(lem) "as" "(" simple_intropattern(x1)
simple_intropattern(x2) simple_intropattern(x3) ")" constr(pat) :=
iDestructCore lem as pat (fun H => iDestructHyp H as ( x1 x2 x3 ) pat).
Tactic Notation "iDestruct" open_constr(lem) "as" "(" simple_intropattern(x1)
simple_intropattern(x2) simple_intropattern(x3) simple_intropattern(x4) ")"
iDestructCore lem as pat (fun H => iDestructHyp H as ( x1 x2 x3 x4 ) pat).
Tactic Notation "iDestruct" open_constr(lem) "as" "(" simple_intropattern(x1)
simple_intropattern(x2) simple_intropattern(x3) simple_intropattern(x4)
simple_intropattern(x5) ")" constr(pat) :=
iDestructCore lem as pat (fun H => iDestructHyp H as ( x1 x2 x3 x4 x5 ) pat).
Tactic Notation "iDestruct" open_constr(lem) "as" "(" simple_intropattern(x1)
simple_intropattern(x2) simple_intropattern(x3) simple_intropattern(x4)
simple_intropattern(x5) simple_intropattern(x6) ")" constr(pat) :=
iDestructCore lem as pat (fun H => iDestructHyp H as ( x1 x2 x3 x4 x5 x6 ) pat).
Tactic Notation "iDestruct" open_constr(lem) "as" "(" simple_intropattern(x1)
simple_intropattern(x2) simple_intropattern(x3) simple_intropattern(x4)
simple_intropattern(x5) simple_intropattern(x6) simple_intropattern(x7) ")"
iDestructCore lem as pat (fun H => iDestructHyp H as ( x1 x2 x3 x4 x5 x6 x7 ) pat).
Tactic Notation "iDestruct" open_constr(lem) "as" "(" simple_intropattern(x1)
simple_intropattern(x2) simple_intropattern(x3) simple_intropattern(x4)
simple_intropattern(x5) simple_intropattern(x6) simple_intropattern(x7)
simple_intropattern(x8) ")" constr(pat) :=
iDestructCore lem as pat (fun H => iDestructHyp H as ( x1 x2 x3 x4 x5 x6 x7 x8 ) pat).
Tactic Notation "iDestruct" open_constr(lem) "as" "%" simple_intropattern(pat) :=
iDestructCore lem as true (fun H => iPure H as pat).
(** * Induction *)
Tactic Notation "iInductionCore" constr(x)
"as" simple_intropattern(pat) constr(IH) :=
let rec fix_ihs :=
lazymatch goal with
| H : coq_tactics.of_envs _ ⊢ _ |- _ =>
eapply tac_revert_ih;
[reflexivity || fail "iInduction: persistent context not empty"
|apply H|];
clear H; fix_ihs;
let IH' := iFresh' IH in iIntros [IAlwaysElim (IName IH')]
| _ => idtac
end in
induction x as pat; fix_ihs.
Tactic Notation "iInduction" constr(x) "as" simple_intropattern(pat) constr(IH) :=
iRevertIntros "∗" with (iInductionCore x as pat IH).
Tactic Notation "iInduction" constr(x) "as" simple_intropattern(pat) constr(IH)
"forall" "(" ident(x1) ")" :=
iRevertIntros(x1) "∗" with (iInductionCore x as pat IH).
Tactic Notation "iInduction" constr(x) "as" simple_intropattern(pat) constr(IH)
"forall" "(" ident(x1) ident(x2) ")" :=
iRevertIntros(x1 x2) "∗" with (iInductionCore x as pat IH).
Tactic Notation "iInduction" constr(x) "as" simple_intropattern(pat) constr(IH)
"forall" "(" ident(x1) ident(x2) ident(x3) ")" :=
iRevertIntros(x1 x2 x3) "∗" with (iInductionCore x as pat IH).
Tactic Notation "iInduction" constr(x) "as" simple_intropattern(pat) constr(IH)
"forall" "(" ident(x1) ident(x2) ident(x3) ident(x4) ")" :=
iRevertIntros(x1 x2 x3 x4) "∗" with (iInductionCore x as pat IH).
Tactic Notation "iInduction" constr(x) "as" simple_intropattern(pat) constr(IH)
"forall" "(" ident(x1) ident(x2) ident(x3) ident(x4) ident(x5) ")" :=
iRevertIntros(x1 x2 x3 x4 x5) "∗" with (iInductionCore x as aat IH).
Tactic Notation "iInduction" constr(x) "as" simple_intropattern(pat) constr(IH)
"forall" "(" ident(x1) ident(x2) ident(x3) ident(x4) ident(x5) ident(x6) ")" :=
iRevertIntros(x1 x2 x3 x4 x5 x6) "∗" with (iInductionCore x as pat IH).
Tactic Notation "iInduction" constr(x) "as" simple_intropattern(pat) constr(IH)
"forall" "(" ident(x1) ident(x2) ident(x3) ident(x4) ident(x5) ident(x6)
ident(x7) ")" :=
iRevertIntros(x1 x2 x3 x4 x5 x6 x7) "∗" with (iInductionCore x as pat IH).
Tactic Notation "iInduction" constr(x) "as" simple_intropattern(pat) constr(IH)
"forall" "(" ident(x1) ident(x2) ident(x3) ident(x4) ident(x5) ident(x6)