From 9eb50174027be945f7e41a7f8dad7d5d114d51b2 Mon Sep 17 00:00:00 2001 From: Robbert Krebbers <mail@robbertkrebbers.nl> Date: Mon, 19 Sep 2016 23:17:29 +0200 Subject: [PATCH] Attempt at an iInduction tactic. MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This comment mostly addresses issue #34. There are still some issues: - For iLöb we can write `iLöb (x1 .. xn) as "IH"` to revert x1 .. xn before performing Löb induction. An analogue notation for iInduction results in parsing conflicts. - The names of the induction hypotheses in the Coq intro pattern are ignored. Instead, when using `iInduction x as pat "IH"` the induction hypotheses are given fresh names starting with "IH". The problem here is that the names in the introduction pattern are idents, whereas the induction hypotheses are inserted into the proof mode context, and thus need to have strings as names. --- proofmode/coq_tactics.v | 10 ++++++++++ proofmode/tactics.v | 21 ++++++++++++++++++++- 2 files changed, 30 insertions(+), 1 deletion(-) diff --git a/proofmode/coq_tactics.v b/proofmode/coq_tactics.v index e271c84dd..408ca2844 100644 --- a/proofmode/coq_tactics.v +++ b/proofmode/coq_tactics.v @@ -541,6 +541,16 @@ 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) → + (of_envs Δ ⊢ □ P → Q) → + (of_envs Δ ⊢ Q). +Proof. + intros ? HP HPQ. + by rewrite -(idemp uPred_and Δ) {1}(persistentP Δ) {1}HP HPQ impl_elim_r. +Qed. + Lemma tac_assert Δ Δ1 Δ2 Δ2' lr js j P Q R : IntoAssert P Q R → envs_split lr js Δ = Some (Δ1,Δ2) → diff --git a/proofmode/tactics.v b/proofmode/tactics.v index c06a27dbf..3411a4361 100644 --- a/proofmode/tactics.v +++ b/proofmode/tactics.v @@ -877,12 +877,31 @@ Tactic Notation "iDestruct" open_constr(lem) "as" "(" simple_intropattern(x1) 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; + [env_cbv; reflexivity + |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). + +(** * Löb Induction *) Tactic Notation "iLöbCore" "as" constr (IH) := eapply tac_löb with _ IH; [reflexivity |env_cbv; reflexivity || fail "iLöb:" IH "not fresh"|]. -(** * Löb induction *) Tactic Notation "iLöb" "as" constr (IH) := iRevertIntros with (iLöbCore as IH). Tactic Notation "iLöb" "(" ident(x1) ")" "as" constr (IH) := -- GitLab