diff --git a/ProofMode.md b/ProofMode.md index 74e0b0d1db1483754622fd1ca701d6a1bc2dd485..34ca2dd1e9475d740bb25353cbf0c2fa1badd766 100644 --- a/ProofMode.md +++ b/ProofMode.md @@ -14,8 +14,8 @@ Applying hypotheses and lemmas - `iExact "H"` : finish the goal if the conclusion matches the hypothesis `H` - `iAssumption` : finish the goal if the conclusion matches any hypothesis - `iApply pm_trm` : match the conclusion of the current goal against the - conclusion of `pm_trm` and generates goals for the premises of `pm_trm`. See - proof mode terms below. + conclusion of `pm_trm` and generates goals for the premises of `pm_trm`. See + proof mode terms below. Context management ------------------ @@ -23,13 +23,12 @@ Context management - `iIntros (x1 ... xn) "ipat1 ... ipatn"` : introduce universal quantifiers using Coq introduction patterns `x1 ... xn` and implications/wands using proof mode introduction patterns `ipat1 ... ipatn`. -- `iClear (x1 ... xn) "H1 ... Hn"` : clear the hypothesis `H1 ... Hn` as well as - the Coq level hypotheses/variables `x1 ... xn`. The symbol `★` can be used to - clear entire spatial context. -- `iRevert (x1 ... xn) "H1 ... Hn"` : revert the proof mode hypotheses - `H1 ... Hn` into wands, as well as the Coq level hypotheses/variables - `x1 ... xn` into universal quantifiers. The symbol `★` can be used to revert - the entire spatial context. +- `iClear (x1 ... xn) "selpat"` : clear the hypotheses given by the selection + pattern `selpat` and the Coq level hypotheses/variables `x1 ... xn`. +- `iRevert (x1 ... xn) "selpat"` : revert the hypotheses given by the selection + pattern `selpat` into wands, and the Coq level hypotheses/variables + `x1 ... xn` into universal quantifiers. Persistent hypotheses are wrapped into + the always modality. - `iRename "H1" into "H2"` : rename the hypothesis `H1` into `H2`. - `iSpecialize pm_trm` : instantiate universal quantifiers and eliminate implications/wands of a hypothesis `pm_trm`. See proof mode terms below. @@ -83,9 +82,9 @@ Elimination of logical connectives Separating logic specific tactics --------------------------------- -- `iFrame (t1 .. tn) "H0 ... Hn"` : cancel the Coq terms (or Coq hypotheses) - `t1 ... tn` and Iris hypotheses `H0 ... Hn` in the goal. Apart from - hypotheses, the following symbols can be used: +- `iFrame (t1 .. tn) "selpat"` : cancel the Coq terms (or Coq hypotheses) + `t1 ... tn` and Iris hypotheses given by `selpat` in the goal. The constructs + of the selection pattern have the following meaning: + `%` : repeatedly frame hypotheses from the Coq context. + `#` : repeatedly frame hypotheses from the persistent context. @@ -102,16 +101,19 @@ Separating logic specific tactics The later modality ------------------ - `iNext` : introduce a later by stripping laters from all hypotheses. -- `iLöb as "IH" forall (x1 ... xn)` : perform Löb induction while generalizing - over the Coq level variables `x1 ... xn` and the entire spatial context. +- `iLöb as "IH" forall (x1 ... xn)` : perform Löb induction by generating a + hypothesis `IH : ▷ goal`. The tactic generalizes over the Coq level variables + `x1 ... xn`, the hypotheses given by the selection pattern `selpat`, and the + spatial context. Induction --------- -- `iInduction x as cpat "IH" forall (x1 ... xn)` : perform induction on the Coq - term `x`. The Coq introduction pattern is used to name the introduced +- `iInduction x as cpat "IH" forall (x1 ... xn) "selpat"` : perform induction on + the Coq term `x`. The Coq introduction pattern is used to name the introduced variables. The induction hypotheses are inserted into the persistent context and given fresh names prefixed `IH`. The tactic generalizes over the Coq level - variables `x1 ... xn` and the entire spatial context. + variables `x1 ... xn`, the hypotheses given by the selection pattern `selpat`, + and the spatial context. Rewriting --------- @@ -125,11 +127,11 @@ Iris - `iVsIntro` : introduction of a raw or primitive view shift. - `iVs pm_trm as (x1 ... xn) "ipat"` : run a raw or primitive view shift `pm_trm` (if the goal permits, i.e. it is a raw or primitive view shift, or - a weakest precondition). + a weakest precondition). - `iInv N as (x1 ... xn) "ipat"` : open the invariant `N`. - `iTimeless "H"` : strip a later of a timeless hypothesis `H` (if the goal - permits, i.e. it is a later, True now, raw or primitive view shift, or a - weakest precondition). + permits, i.e. it is a later, True now, raw or primitive view shift, or a + weakest precondition). Miscellaneous ------------- @@ -141,14 +143,26 @@ Miscellaneous existential quantifiers, implications and wand, always and later modalities, primitive view shifts, and pure connectives. +Selection patterns +================== + +Selection patterns are used to select hypotheses in the tactics `iRevert`, +`iClear`, `iFrame`, `iLöb` and `iInduction`. The proof mode supports the +following _selection patterns_: + +- `H` : select the hypothesis named `H`. +- `%` : select the entire pure/Coq context. +- `#` : select the entire persistent context. +- `★` : select the entire spatial context. + Introduction patterns ===================== Introduction patterns are used to perform introductions and eliminations of multiple connectives on the fly. The proof mode supports the following -introduction patterns: +_introduction patterns_: -- `H` : create a hypothesis named H. +- `H` : create a hypothesis named `H`. - `?` : create an anonymous hypothesis. - `_` : remove the hypothesis. - `$` : frame the hypothesis in the goal. @@ -197,9 +211,9 @@ Specialization patterns ======================= Since we are reasoning in a spatial logic, when eliminating a lemma or -hypotheses of type ``P_0 -★ ... -★ P_n -★ R`` one has to specify how the +hypothesis of type ``P_0 -★ ... -★ P_n -★ R``, one has to specify how the hypotheses are split between the premises. The proof mode supports the following -so called specification patterns to express this splitting: +_specification patterns_ to express splitting of hypotheses: - `H` : use the hypothesis `H` (it should match the premise exactly). If `H` is spatial, it will be consumed. diff --git a/_CoqProject b/_CoqProject index ed943035c2fbdf70e573ad252b73cdf188482637..1a100843dc5b2fc1c01ce77e4da3220ad3c04994 100644 --- a/_CoqProject +++ b/_CoqProject @@ -120,6 +120,7 @@ proofmode/pviewshifts.v proofmode/environments.v proofmode/intro_patterns.v proofmode/spec_patterns.v +proofmode/sel_patterns.v proofmode/tactics.v proofmode/notation.v proofmode/invariants.v diff --git a/program_logic/counter_examples.v b/program_logic/counter_examples.v index d51505fa8da76e1d533528519751a3248e0c28c8..aa3ceb25a82e44997fd4fe2f0c062da1d88e383f 100644 --- a/program_logic/counter_examples.v +++ b/program_logic/counter_examples.v @@ -152,7 +152,7 @@ Module inv. Section inv. iDestruct (finished_dup with "Hf") as "[Hf Hf']". iApply pvs_intro. iSplitL "Hf'"; first by eauto. (* Step 2: Open the Q-invariant. *) - iClear "HiP". clear i. iDestruct "HsQ" as (i) "HiQ". + iClear (i) "HiP ". iDestruct "HsQ" as (i) "HiQ". iApply (inv_open' i). iSplit; first done. iIntros "[HaQ | [_ #HQ]]". { iExFalso. iApply finished_not_start. by iFrame. } diff --git a/proofmode/coq_tactics.v b/proofmode/coq_tactics.v index 2d8e234ac739ca9cf6383db51610a557158ba371..fd58f1aefd8d4af071b805ec246d01b0e92ad95c 100644 --- a/proofmode/coq_tactics.v +++ b/proofmode/coq_tactics.v @@ -82,6 +82,9 @@ Definition env_spatial_is_nil {M} (Δ : envs M) := Definition envs_clear_spatial {M} (Δ : envs M) : envs M := Envs (env_persistent Δ) Enil. +Definition envs_clear_persistent {M} (Δ : envs M) : envs M := + Envs Enil (env_spatial Δ). + Fixpoint envs_split_go {M} (js : list string) (Δ1 Δ2 : envs M) : option (envs M * envs M) := match js with @@ -272,12 +275,6 @@ Proof. destruct Hwf; constructor; simpl; auto using Enil_wf. Qed. -Lemma env_fold_wand Γ Q : env_fold uPred_wand Q Γ ⊣⊢ ([★] Γ -★ Q). -Proof. - revert Q; induction Γ as [|Γ IH i P]=> Q /=; [by rewrite wand_True|]. - by rewrite IH wand_curry (comm uPred_sep). -Qed. - Lemma env_spatial_is_nil_persistent Δ : env_spatial_is_nil Δ = true → PersistentP Δ. Proof. intros; destruct Δ as [? []]; simplify_eq/=; apply _. Qed. @@ -385,9 +382,6 @@ Qed. Lemma tac_clear Δ Δ' i p P Q : envs_lookup_delete i Δ = Some (p,P,Δ') → (Δ' ⊢ Q) → Δ ⊢ Q. Proof. intros. by rewrite envs_lookup_delete_sound // sep_elim_r. Qed. -Lemma tac_clear_spatial Δ Δ' Q : - envs_clear_spatial Δ = Δ' → (Δ' ⊢ Q) → Δ ⊢ Q. -Proof. intros <- ?. by rewrite envs_clear_spatial_sound // sep_elim_l. Qed. (** * False *) Lemma tac_ex_falso Δ Q : (Δ ⊢ False) → Δ ⊢ Q. @@ -612,12 +606,6 @@ Proof. - by rewrite HQ wand_elim_r. Qed. -Lemma tac_revert_spatial Δ Q : - (envs_clear_spatial Δ ⊢ env_fold uPred_wand Q (env_spatial Δ)) → Δ ⊢ Q. -Proof. - intros HΔ. by rewrite envs_clear_spatial_sound HΔ env_fold_wand wand_elim_l. -Qed. - Lemma tac_revert_ih Δ P Q : env_spatial_is_nil Δ = true → (of_envs Δ ⊢ P) → diff --git a/proofmode/environments.v b/proofmode/environments.v index a262ae95d181732d34f5a9ae14a132ca1efa0a2d..b6838b683552985623b337e81def165c60e088f6 100644 --- a/proofmode/environments.v +++ b/proofmode/environments.v @@ -38,12 +38,6 @@ Instance: Params (@env_to_list) 1. Fixpoint env_dom {A} (Γ : env A) : list string := match Γ with Enil => [] | Esnoc Γ i _ => i :: env_dom Γ end. -Fixpoint env_fold {A B} (f : B → A → A) (x : A) (Γ : env B) : A := - match Γ with - | Enil => x - | Esnoc Γ _ y => env_fold f (f y x) Γ - end. - Fixpoint env_app {A} (Γapp : env A) (Γ : env A) : option (env A) := match Γapp with | Enil => Some Γ diff --git a/proofmode/sel_patterns.v b/proofmode/sel_patterns.v new file mode 100644 index 0000000000000000000000000000000000000000..583b49dfb9a9de124678ac14c646b32050fefe5a --- /dev/null +++ b/proofmode/sel_patterns.v @@ -0,0 +1,40 @@ +From iris.prelude Require Export strings. + +Inductive sel_pat := + | SelPure + | SelPersistent + | SelSpatial + | SelName : string → sel_pat. + +Fixpoint sel_pat_pure (ps : list sel_pat) : bool := + match ps with + | [] => false + | SelPure :: ps => true + | _ :: ps => sel_pat_pure ps + end. + +Module sel_pat. +Fixpoint cons_name (kn : string) (k : list sel_pat) : list sel_pat := + match kn with "" => k | _ => SelName (string_rev kn) :: k end. + +Fixpoint parse_go (s : string) (k : list sel_pat) (kn : string) : list sel_pat := + match s with + | "" => rev (cons_name kn k) + | String " " s => parse_go s (cons_name kn k) "" + | String "%" s => parse_go s (SelPure :: cons_name kn k) "" + | String "#" s => parse_go s (SelPersistent :: cons_name kn k) "" + | String (Ascii.Ascii false true false false false true true true) (* unicode ★ *) + (String (Ascii.Ascii false false false true true false false true) + (String (Ascii.Ascii true false true false false false false true) s)) => + parse_go s (SelSpatial :: cons_name kn k) "" + | String a s => parse_go s k (String a kn) + end. +Definition parse (s : string) : list sel_pat := parse_go s [] "". + +Ltac parse s := + lazymatch type of s with + | list sel_pat => s + | list string => eval vm_compute in (SelName <$> s) + | string => eval vm_compute in (parse s) + end. +End sel_pat. diff --git a/proofmode/tactics.v b/proofmode/tactics.v index 7ee2bb6ebc963a9666cf8c74d78c7075bd8716eb..217ec83a62041863e05ce6cbaff8d83bedaa694c 100644 --- a/proofmode/tactics.v +++ b/proofmode/tactics.v @@ -1,19 +1,21 @@ -From iris.proofmode Require Import coq_tactics intro_patterns spec_patterns. +From iris.proofmode Require Import coq_tactics. +From iris.proofmode Require Import intro_patterns spec_patterns sel_patterns. From iris.algebra Require Export upred. From iris.proofmode Require Export classes notation. From iris.proofmode Require Import class_instances. From iris.prelude Require Import stringmap hlist. Declare Reduction env_cbv := cbv [ - env_lookup env_fold env_lookup_delete env_delete env_app env_replace + env_lookup env_lookup_delete env_delete env_app env_replace env_dom decide (* operational classes *) sumbool_rec sumbool_rect (* sumbool *) bool_eq_dec bool_rec bool_rect bool_dec eqb andb (* bool *) assci_eq_dec ascii_to_digits Ascii.ascii_dec Ascii.ascii_rec Ascii.ascii_rect string_eq_dec string_rec string_rect (* strings *) - env_persistent env_spatial env_spatial_is_nil + 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_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. @@ -33,7 +35,7 @@ Ltac iFresh := iFresh' "~". Tactic Notation "iTypeOf" constr(H) tactic(tac):= let Δ := match goal with |- of_envs ?Δ ⊢ _ => Δ end in - match eval env_cbv in (envs_lookup H Δ) with + lazymatch eval env_cbv in (envs_lookup H Δ) with | Some (?p,?P) => tac p P end. @@ -56,16 +58,45 @@ Tactic Notation "iRename" constr(H1) "into" constr(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 | [] => idtac - | "★" :: ?Hs => eapply tac_clear_spatial; [env_cbv; reflexivity|go Hs] - | ?H :: ?Hs => + | 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 - let Hs := words Hs in go Hs. + iElaborateSelPat Hs go. + Tactic Notation "iClear" "(" ident_list(xs) ")" constr(Hs) := iClear Hs; clear xs. @@ -192,12 +223,12 @@ Tactic Notation "iFrame" constr(Hs) := let rec go Hs := match Hs with | [] => idtac - | "%" :: ?Hs => iFrameAnyPure; go Hs - | "#" :: ?Hs => iFrameAnyPersistent; go Hs - | "★" :: ?Hs => iFrameAnySpatial; go Hs - | ?H :: ?Hs => iFrameHyp H; go Hs + | 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 := words Hs in go Hs. + 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) := @@ -403,17 +434,18 @@ Local Tactic Notation "iForallRevert" ident(x) := end || fail "iRevert: cannot revert" x. Tactic Notation "iRevert" constr(Hs) := - let rec go H2s := - match H2s with + let rec go Hs := + lazymatch Hs with | [] => idtac - | "★" :: ?H2s => go H2s; eapply tac_revert_spatial; env_cbv - | ?H2 :: ?H2s => - go H2s; - eapply tac_revert with _ H2 _ _; (* (i:=H2) *) - [env_cbv; reflexivity || fail "iRevert:" H2 "not found" - |env_cbv] + | 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] end in - let Hs := words Hs in go Hs. + iElaborateSelPat Hs go. Tactic Notation "iRevert" "(" ident(x1) ")" := iForallRevert x1. @@ -793,37 +825,71 @@ Tactic Notation "iIntros" "(" simple_intropattern(x1) simple_intropattern(x2) ")" 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 + 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) := - match goal with - | |- of_envs ?Δ ⊢ _ => - let Hs := eval cbv in (reverse (env_dom (env_spatial Δ))) in - iRevert ["★"]; tac; iIntros Hs - end. + iRevertIntros "" with tac. Tactic Notation "iRevertIntros" "(" ident(x1) ")" "with" tactic(tac):= - iRevertIntros with (iRevert (x1); tac; iIntros (x1)). + iRevertIntros (x1) "" with tac. Tactic Notation "iRevertIntros" "(" ident(x1) ident(x2) ")" "with" tactic(tac):= - iRevertIntros with (iRevert (x1 x2); tac; iIntros (x1 x2)). + iRevertIntros (x1 x2) "" with tac. Tactic Notation "iRevertIntros" "(" ident(x1) ident(x2) ident(x3) ")" "with" tactic(tac):= - iRevertIntros with (iRevert (x1 x2 x3); tac; iIntros (x1 x2 x3)). + iRevertIntros (x1 x2 x3) "" with tac. Tactic Notation "iRevertIntros" "(" ident(x1) ident(x2) ident(x3) ident(x4) ")" "with" tactic(tac):= - iRevertIntros with (iRevert (x1 x2 x3 x4); tac; iIntros (x1 x2 x3 x4)). + iRevertIntros (x1 x2 x3 x4) "" with tac. Tactic Notation "iRevertIntros" "(" ident(x1) ident(x2) ident(x3) ident(x4) ident(x5) ")" "with" tactic(tac):= - iRevertIntros with (iRevert (x1 x2 x3 x4 x5); tac; iIntros (x1 x2 x3 x4 x5)). + 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 with (iRevert (x1 x2 x3 x4 x5 x6); - tac; iIntros (x1 x2 x3 x4 x5 x6)). + 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 with (iRevert (x1 x2 x3 x4 x5 x6 x7); - tac; iIntros (x1 x2 x3 x4 x5 x6 x7)). + 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 with (iRevert (x1 x2 x3 x4 x5 x6 x7 x8); - tac; iIntros (x1 x2 x3 x4 x5 x6 x7 x8)). + iRevertIntros (x1 x2 x3 x4 x5 x6 x7 x8) "" with tac. (** * Destruct tactic *) Tactic Notation "iDestructCore" open_constr(lem) "as" constr(p) tactic(tac) := @@ -893,7 +959,7 @@ Tactic Notation "iInductionCore" constr(x) lazymatch goal with | H : coq_tactics.of_envs _ ⊢ _ |- _ => eapply tac_revert_ih; - [env_cbv; reflexivity + [reflexivity || fail "iInduction: persistent context not empty" |apply H|]; clear H; fix_ihs; let IH' := iFresh' IH in iIntros [IAlwaysElim (IName IH')] @@ -902,64 +968,122 @@ Tactic Notation "iInductionCore" constr(x) induction x as pat; fix_ihs. Tactic Notation "iInduction" constr(x) "as" simple_intropattern(pat) constr(IH) := - iRevertIntros with (iInductionCore x as pat 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). + 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). + 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). + 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). + 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). + "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). + 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). + 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) ident(x7) ident(x8) ")" := - iRevertIntros(x1 x2 x3 x4 x5 x6 x7 x8) with (iInductionCore x as pat IH). + iRevertIntros(x1 x2 x3 x4 x5 x6 x7 x8) "★" with (iInductionCore x as pat IH). + +Tactic Notation "iInduction" constr(x) "as" simple_intropattern(pat) constr(IH) + "forall" constr(Hs) := + iRevertIntros Hs with (iInductionCore x as pat IH). +Tactic Notation "iInduction" constr(x) "as" simple_intropattern(pat) constr(IH) + "forall" "(" ident(x1) ")" constr(Hs) := + iRevertIntros(x1) Hs with (iInductionCore x as pat IH). +Tactic Notation "iInduction" constr(x) "as" simple_intropattern(pat) constr(IH) + "forall" "(" ident(x1) ident(x2) ")" constr(Hs) := + iRevertIntros(x1 x2) Hs with (iInductionCore x as pat IH). +Tactic Notation "iInduction" constr(x) "as" simple_intropattern(pat) constr(IH) + "forall" "(" ident(x1) ident(x2) ident(x3) ")" constr(Hs) := + iRevertIntros(x1 x2 x3) Hs 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) ")" constr(Hs) := + iRevertIntros(x1 x2 x3 x4) Hs 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) ")" + constr(Hs) := + iRevertIntros(x1 x2 x3 x4 x5) Hs 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) ")" + constr(Hs) := + iRevertIntros(x1 x2 x3 x4 x5 x6) Hs 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) ")" constr(Hs) := + iRevertIntros(x1 x2 x3 x4 x5 x6 x7) Hs 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) ident(x8) ")" constr(Hs) := + iRevertIntros(x1 x2 x3 x4 x5 x6 x7 x8) Hs with (iInductionCore x as pat IH). (** * Löb Induction *) Tactic Notation "iLöbCore" "as" constr (IH) := eapply tac_löb with _ IH; - [reflexivity + [reflexivity || fail "iLöb: persistent context not empty" |env_cbv; reflexivity || fail "iLöb:" IH "not fresh"|]. Tactic Notation "iLöb" "as" constr (IH) := - iRevertIntros with (iLöbCore as IH). + iRevertIntros "★" with (iLöbCore as IH). Tactic Notation "iLöb" "as" constr (IH) "forall" "(" ident(x1) ")" := - iRevertIntros(x1) with (iLöbCore as IH). + iRevertIntros(x1) "★" with (iLöbCore as IH). Tactic Notation "iLöb" "as" constr (IH) "forall" "(" ident(x1) ident(x2) ")" := - iRevertIntros(x1 x2) with (iLöbCore as IH). + iRevertIntros(x1 x2) "★" with (iLöbCore as IH). Tactic Notation "iLöb" "as" constr (IH) "forall" "(" ident(x1) ident(x2) ident(x3) ")" := - iRevertIntros(x1 x2 x3) with (iLöbCore as IH). + iRevertIntros(x1 x2 x3) "★" with (iLöbCore as IH). Tactic Notation "iLöb" "as" constr (IH) "forall" "(" ident(x1) ident(x2) ident(x3) ident(x4) ")" := - iRevertIntros(x1 x2 x3 x4) with (iLöbCore as IH). + iRevertIntros(x1 x2 x3 x4) "★" with (iLöbCore as IH). Tactic Notation "iLöb" "as" constr (IH) "forall" "(" ident(x1) ident(x2) ident(x3) ident(x4) ident(x5) ")" := - iRevertIntros(x1 x2 x3 x4 x5) with (iLöbCore as IH). + iRevertIntros(x1 x2 x3 x4 x5) "★" with (iLöbCore as IH). Tactic Notation "iLöb" "as" constr (IH) "forall" "(" ident(x1) ident(x2) ident(x3) ident(x4) ident(x5) ident(x6) ")" := - iRevertIntros(x1 x2 x3 x4 x5 x6) with (iLöbCore as IH). + iRevertIntros(x1 x2 x3 x4 x5 x6) "★" with (iLöbCore as IH). Tactic Notation "iLöb" "as" 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 (iLöbCore as IH). + iRevertIntros(x1 x2 x3 x4 x5 x6 x7) "★" with (iLöbCore as IH). Tactic Notation "iLöb" "as" constr (IH) "forall" "(" ident(x1) ident(x2) ident(x3) ident(x4) ident(x5) ident(x6) ident(x7) ident(x8) ")" := - iRevertIntros(x1 x2 x3 x4 x5 x6 x7 x8) with (iLöbCore as IH). + iRevertIntros(x1 x2 x3 x4 x5 x6 x7 x8) "★" with (iLöbCore as IH). + +Tactic Notation "iLöb" "as" constr (IH) "forall" constr(Hs) := + iRevertIntros Hs with (iLöbCore as IH). +Tactic Notation "iLöb" "as" constr (IH) "forall" "(" ident(x1) ")" constr(Hs) := + iRevertIntros(x1) Hs with (iLöbCore as IH). +Tactic Notation "iLöb" "as" constr (IH) "forall" "(" ident(x1) ident(x2) ")" + constr(Hs) := + iRevertIntros(x1 x2) Hs with (iLöbCore as IH). +Tactic Notation "iLöb" "as" constr (IH) "forall" "(" ident(x1) ident(x2) + ident(x3) ")" constr(Hs) := + iRevertIntros(x1 x2 x3) Hs with (iLöbCore as IH). +Tactic Notation "iLöb" "as" constr (IH) "forall" "(" ident(x1) ident(x2) + ident(x3) ident(x4) ")" constr(Hs) := + iRevertIntros(x1 x2 x3 x4) Hs with (iLöbCore as IH). +Tactic Notation "iLöb" "as" constr (IH) "forall" "(" ident(x1) ident(x2) + ident(x3) ident(x4) ident(x5) ")" constr(Hs) := + iRevertIntros(x1 x2 x3 x4 x5) Hs with (iLöbCore as IH). +Tactic Notation "iLöb" "as" constr (IH) "forall" "(" ident(x1) ident(x2) + ident(x3) ident(x4) ident(x5) ident(x6) ")" constr(Hs) := + iRevertIntros(x1 x2 x3 x4 x5 x6) Hs with (iLöbCore as IH). +Tactic Notation "iLöb" "as" constr (IH) "forall" "(" ident(x1) ident(x2) + ident(x3) ident(x4) ident(x5) ident(x6) ident(x7) ")" constr(Hs) := + iRevertIntros(x1 x2 x3 x4 x5 x6 x7) Hs with (iLöbCore as IH). +Tactic Notation "iLöb" "as" constr (IH) "forall" "(" ident(x1) ident(x2) + ident(x3) ident(x4) ident(x5) ident(x6) ident(x7) ident(x8) ")" constr(Hs) := + iRevertIntros(x1 x2 x3 x4 x5 x6 x7 x8) Hs with (iLöbCore as IH). (** * Assert *) Tactic Notation "iAssertCore" open_constr(Q) "with" constr(Hs) "as" tactic(tac) :=