Forked from
Iris / Iris
3839 commits behind the upstream repository.
-
Robbert Krebbers authoredRobbert Krebbers authored
proofmode.v 24.18 KiB
From iris.proofmode Require Import tactics intro_patterns.
From stdpp Require Import gmap hlist.
Set Default Proof Using "Type".
Section tests.
Context {PROP : sbi}.
Implicit Types P Q R : PROP.
Check "demo_0".
Lemma demo_0 P Q : □ (P ∨ Q) -∗ (∀ x, ⌜x = 0⌝ ∨ ⌜x = 1⌝) → (Q ∨ P).
Proof.
iIntros "H #H2". Show. iDestruct "H" as "###H".
(* should remove the disjunction "H" *)
iDestruct "H" as "[#?|#?]"; last by iLeft. Show.
(* should keep the disjunction "H" because it is instantiated *)
iDestruct ("H2" $! 10) as "[%|%]". done. done.
Qed.
Lemma demo_2 P1 P2 P3 P4 Q (P5 : nat → PROP) `{!Affine P4, !Absorbing P2} :
P2 ∗ (P3 ∗ Q) ∗ True ∗ P1 ∗ P2 ∗ (P4 ∗ (∃ x:nat, P5 x ∨ P3)) ∗ emp -∗
P1 -∗ (True ∗ True) -∗
(((P2 ∧ False ∨ P2 ∧ ⌜0 = 0⌝) ∗ P3) ∗ Q ∗ P1 ∗ True) ∧
(P2 ∨ False) ∧ (False → P5 0).
Proof.
(* Intro-patterns do something :) *)
iIntros "[H2 ([H3 HQ]&?&H1&H2'&foo&_)] ? [??]".
(* To test destruct: can also be part of the intro-pattern *)
iDestruct "foo" as "[_ meh]".
repeat iSplit; [|by iLeft|iIntros "#[]"].
iFrame "H2".
(* split takes a list of hypotheses just for the LHS *)
iSplitL "H3".
- iFrame "H3". iRight. auto.
- iSplitL "HQ". iAssumption. by iSplitL "H1".
Qed.
Lemma demo_3 P1 P2 P3 :
P1 ∗ P2 ∗ P3 -∗ P1 ∗ ▷ (P2 ∗ ∃ x, (P3 ∧ ⌜x = 0⌝) ∨ P3).
Proof. iIntros "($ & $ & $)". iNext. by iExists 0. Qed.
Definition foo (P : PROP) := (P -∗ P)%I.
Definition bar : PROP := (∀ P, foo P)%I.
Lemma test_unfold_constants : bar.
Proof. iIntros (P) "HP //". Qed.
Lemma test_iRewrite {A : ofeT} (x y : A) P :
□ (∀ z, P -∗ <affine> (z ≡ y)) -∗ (P -∗ P ∧ (x,x) ≡ (y,x)).
Proof.
iIntros "#H1 H2".
iRewrite (bi.internal_eq_sym x x with "[# //]").
iRewrite -("H1" $! _ with "[- //]").
auto.
Qed.
Check "test_iDestruct_and_emp".
Lemma test_iDestruct_and_emp P Q `{!Persistent P, !Persistent Q} :
P ∧ emp -∗ emp ∧ Q -∗ <affine> (P ∗ Q).
Proof. iIntros "[#? _] [_ #?]". Show. auto. Qed.
Lemma test_iIntros_persistent P Q `{!Persistent Q} : (P → Q → P ∧ Q)%I.
Proof. iIntros "H1 #H2". by iFrame "∗#". Qed.
Lemma test_iIntros_pure (ψ φ : Prop) P : ψ → (⌜ φ ⌝ → P → ⌜ φ ∧ ψ ⌝ ∧ P)%I.
Proof. iIntros (??) "H". auto. Qed.
Lemma test_iIntros_pure_not : (⌜ ¬False ⌝ : PROP)%I.
Proof. by iIntros (?). Qed.
Lemma test_fast_iIntros P Q :