diff --git a/theories/proofmode/coq_tactics.v b/theories/proofmode/coq_tactics.v index 775d5655818cc39353fdfe74a402c935e22a3000..f52984b25fff27deac0e4c06823630acbe06e596 100644 --- a/theories/proofmode/coq_tactics.v +++ b/theories/proofmode/coq_tactics.v @@ -647,14 +647,25 @@ Proof. by rewrite HQ /uPred_always_if wand_elim_r. Qed. -Lemma tac_revert_ih Δ P Q : +Class IntoIH (φ : Prop) (Δ : envs M) (Q : uPred M) := + into_ih : φ → Δ ⊢ Q. +Global Instance into_ih_entails Δ Q : IntoIH (of_envs Δ ⊢ Q) Δ Q. +Proof. by rewrite /IntoIH. Qed. +Global Instance into_ih_forall {A} (φ : A → Prop) Δ Φ : + (∀ x, IntoIH (φ x) Δ (Φ x)) → IntoIH (∀ x, φ x) Δ (∀ x, Φ x) | 2. +Proof. rewrite /IntoIH=> HΔ ?. apply forall_intro=> x. by rewrite (HΔ x). Qed. +Global Instance into_ih_impl (φ ψ : Prop) Δ Q : + IntoIH φ Δ Q → IntoIH (ψ → φ) Δ (⌜ψ⌠→ Q) | 1. +Proof. rewrite /IntoIH=> HΔ ?. apply impl_intro_l, pure_elim_l. auto. Qed. + +Lemma tac_revert_ih Δ P Q {φ : Prop} (Hφ : φ) : + IntoIH φ Δ P → 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. + rewrite /IntoIH. 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 P' Q : diff --git a/theories/proofmode/tactics.v b/theories/proofmode/tactics.v index c9379be107c5bb373c05138e046c51677042acc4..7d01214b8cd3ad2e3989951e497fd488ff404fbb 100644 --- a/theories/proofmode/tactics.v +++ b/theories/proofmode/tactics.v @@ -1360,10 +1360,10 @@ result in the following actions: 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: spatial context not empty, this should not happen" - |apply H|]; + | H : context [coq_tactics.of_envs _ ⊢ _] |- _ => + eapply (tac_revert_ih _ _ _ H _); + [reflexivity + || fail "iInduction: spatial context not empty, this should not happen"|]; clear H; fix_ihs; let IH' := iFresh' IH in iIntros [IAlwaysElim (IName IH')] | _ => idtac diff --git a/theories/tests/proofmode.v b/theories/tests/proofmode.v index e09f3f234346ce9450c1a9d04211873979c232e3..28e6152235d7fe8621d91f05e44915c1cf811b57 100644 --- a/theories/tests/proofmode.v +++ b/theories/tests/proofmode.v @@ -190,6 +190,14 @@ Proof. by iExists (S n). Qed. +Lemma test_iInduction_wf (x : nat) P Q : + â–¡ P -∗ Q -∗ ⌜ (x + 0 = x)%nat âŒ. +Proof. + iIntros "#HP HQ". + iInduction (lt_wf x) as [[|x] _] "IH"; simpl; first done. + rewrite (inj_iff S). by iApply ("IH" with "[%]"); first omega. +Qed. + Lemma test_iIntros_start_proof : (True : uPred M)%I. Proof.