diff --git a/CHANGELOG.md b/CHANGELOG.md index bbe665ed4eac8dbd600f4a743500c997da748d59..9a1bebc64bef8b690fc5d6cc0d0845c894c7b22c 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -7,6 +7,7 @@ Coq development, but not every API-breaking change is listed. Changes marked Changes in and extensions of the theory: +* [#] Add new modality: â– ("plainly"). * [#] Camera morphisms have to be homomorphisms, not just monotone functions. * [#] A proof that `f` has a fixed point if `f^k` is contractive. * Constructions for least and greatest fixed points over monotone predicates diff --git a/ProofMode.md b/ProofMode.md index 8bc2bf2dd40fb5bd8a2b0c83e90e1fafd814ab24..cc19fdf3d8368ba37d0654b94d7e772a3b5c4726 100644 --- a/ProofMode.md +++ b/ProofMode.md @@ -121,6 +121,13 @@ Modalities an instance of the `ElimModal` type class. Instances include: later, except 0, basic update and fancy update. +The persistence and plainness modalities +---------------------------------------- + +- `iAlways` : introduce a persistence or plainness modality and the spatial + context. In case of a plainness modality, the tactic will prune all persistent + hypotheses that are not plain. + The later modality ------------------ @@ -197,6 +204,7 @@ _introduction patterns_: - `[ipat1|ipat2]` : disjunction elimination. - `[]` : false elimination. - `%` : move the hypothesis to the pure Coq context (anonymously). +- `->` and `<-` : rewrite using a pure Coq equality - `# ipat` : move the hypothesis to the persistent context. - `> ipat` : eliminate a modality (if the goal permits). @@ -207,10 +215,8 @@ appear at the top level: Items of the selection pattern can be prefixed with `$`, which cause them to be framed instead of cleared. - `!%` : introduce a pure goal (and leave the proof mode). -- `!#` : introduce an persistence or plainness modality and clear the spatial - context. In case of a plainness modality, it will prune all persistent - hypotheses that are not plain. -- `!>` : introduce a modality. +- `!#` : introduce an persistence or plainness modality (by calling `iAlways`). +- `!>` : introduce a modality (by calling `iModIntro`). - `/=` : perform `simpl`. - `//` : perform `try done` on all goals. - `//=` : syntactic sugar for `/= //` diff --git a/_CoqProject b/_CoqProject index 8f1e85df39ef852e04ad5bf10d05a7ed739c9db6..928dd91a660afed1619b5d3ef95aedaf73cabb08 100644 --- a/_CoqProject +++ b/_CoqProject @@ -78,7 +78,7 @@ theories/heap_lang/lib/ticket_lock.v theories/heap_lang/lib/counter.v theories/heap_lang/proofmode.v theories/heap_lang/adequacy.v -theories/proofmode/strings.v +theories/proofmode/base.v theories/proofmode/tokens.v theories/proofmode/coq_tactics.v theories/proofmode/environments.v diff --git a/theories/base_logic/upred.v b/theories/base_logic/upred.v index a52cce8a6492e1355456c7b4042395c50d577805..46b102d5801faecbca24f4fb2c294b8b5d9a4d93 100644 --- a/theories/base_logic/upred.v +++ b/theories/base_logic/upred.v @@ -248,6 +248,9 @@ Definition uPred_wand {M} := unseal uPred_wand_aux M. Definition uPred_wand_eq : @uPred_wand = @uPred_wand_def := seal_eq uPred_wand_aux. +(* Equivalently, this could be `∀ y, P n y`. That's closer to the intuition + of "embedding the step-indexed logic in Iris", but the two are equivalent + because Iris is afine. The following is easier to work with. *) Program Definition uPred_plainly_def {M} (P : uPred M) : uPred M := {| uPred_holds n x := P n ε |}. Solve Obligations with naive_solver eauto using uPred_closed, ucmra_unit_validN. diff --git a/theories/bi/interface.v b/theories/bi/interface.v index 3d1dbb6ea986ff58958b7221382d0add9df11fcf..fc47bb496f15a977c2a8f513367b0961a68ceedf 100644 --- a/theories/bi/interface.v +++ b/theories/bi/interface.v @@ -112,6 +112,9 @@ Section bi_mixin. bi_mixin_prop_ext P Q : bi_plainly ((P → Q) ∧ (Q → P)) ⊢ bi_internal_eq (OfeT PROP prop_ofe_mixin) P Q; + (* The following two laws are very similar, and indeed they hold + not just for â–¡ and â– , but for any modality defined as + `M P n x := ∀ y, R x y → P n y`. *) bi_mixin_persistently_impl_plainly P Q : (bi_plainly P → bi_persistently Q) ⊢ bi_persistently (bi_plainly P → Q); bi_mixin_plainly_impl_plainly P Q : diff --git a/theories/proofmode/strings.v b/theories/proofmode/base.v similarity index 80% rename from theories/proofmode/strings.v rename to theories/proofmode/base.v index bde47e5bfc932adce078ccbc0c4bbaf03cde8b4e..1e185e2cdf4d4b2852b6c2b4b043b47a538d128c 100644 --- a/theories/proofmode/strings.v +++ b/theories/proofmode/base.v @@ -1,8 +1,14 @@ -From stdpp Require Import strings. -From iris.algebra Require Import base. +From stdpp Require Export strings. +From iris.algebra Require Export base. From Coq Require Import Ascii. Set Default Proof Using "Type". +(* Directions of rewrites *) +Inductive direction := Left | Right. + +(* Some specific versions of operations on strings for the proof mode. We need +those so that we can make [cbv] unfold just them, but not the actual operations +that may appear in users' proofs. *) Local Notation "b1 && b2" := (if b1 then b2 else false) : bool_scope. Lemma lazy_andb_true (b1 b2 : bool) : b1 && b2 = true ↔ b1 = true ∧ b2 = true. diff --git a/theories/proofmode/coq_tactics.v b/theories/proofmode/coq_tactics.v index 67c483d10415c8a877fcc60dbd7c74a47817c2fc..66fc81fb1bc579524ee8599cf3020fce711d3887 100644 --- a/theories/proofmode/coq_tactics.v +++ b/theories/proofmode/coq_tactics.v @@ -1,6 +1,6 @@ From iris.bi Require Export bi. From iris.bi Require Import tactics. -From iris.proofmode Require Export environments classes. +From iris.proofmode Require Export base environments classes. From stdpp Require Import stringmap hlist. Set Default Proof Using "Type". Import bi. @@ -111,12 +111,12 @@ Fixpoint envs_split_go {PROP} if p then envs_split_go js Δ1 Δ2 else envs_split_go js Δ1' (envs_snoc Δ2 false j P) end. -(* if [lr = true] then [result = (remaining hyps, hyps named js)] and - if [lr = false] then [result = (hyps named js, remaining hyps)] *) -Definition envs_split {PROP} (lr : bool) +(* if [d = Right] then [result = (remaining hyps, hyps named js)] and + if [d = Left] then [result = (hyps named js, remaining hyps)] *) +Definition envs_split {PROP} (d : direction) (js : list string) (Δ : envs PROP) : option (envs PROP * envs PROP) := '(Δ1,Δ2) ↠envs_split_go js Δ (envs_clear_spatial Δ); - if lr then Some (Δ1,Δ2) else Some (Δ2,Δ1). + if d is Right then Some (Δ1,Δ2) else Some (Δ2,Δ1). (* Coq versions of the tactics *) Section bi_tactics. @@ -216,11 +216,11 @@ Proof. apply wand_intro_l; destruct p; simpl. - apply and_intro; [apply pure_intro|]. + destruct Hwf; constructor; simpl; eauto using Esnoc_wf. - intros j; destruct (strings.string_beq_reflect j i); naive_solver. + intros j; destruct (string_beq_reflect j i); naive_solver. + by rewrite affinely_persistently_and and_sep_affinely_persistently assoc. - apply and_intro; [apply pure_intro|]. + destruct Hwf; constructor; simpl; eauto using Esnoc_wf. - intros j; destruct (strings.string_beq_reflect j i); naive_solver. + intros j; destruct (string_beq_reflect j i); naive_solver. + solve_sep_entails. Qed. @@ -380,8 +380,8 @@ Proof. - rewrite envs_lookup_envs_delete_ne // envs_lookup_snoc_ne //. eauto. } rewrite (envs_snoc_sound Δ2 false j P) /= ?wand_elim_r; eauto. Qed. -Lemma envs_split_sound Δ lr js Δ1 Δ2 : - envs_split lr js Δ = Some (Δ1,Δ2) → Δ ⊢ Δ1 ∗ Δ2. +Lemma envs_split_sound Δ d js Δ1 Δ2 : + envs_split d js Δ = Some (Δ1,Δ2) → Δ ⊢ Δ1 ∗ Δ2. Proof. rewrite /envs_split=> ?. rewrite -(idemp bi_and Δ). rewrite {2}envs_clear_spatial_sound. @@ -392,7 +392,7 @@ Proof. destruct (envs_split_go _ _) as [[Δ1' Δ2']|] eqn:HΔ; [|done]. apply envs_split_go_sound in HΔ as ->; last first. { intros j P. by rewrite envs_lookup_envs_clear_spatial=> ->. } - destruct lr; simplify_eq; solve_sep_entails. + destruct d; simplify_eq; solve_sep_entails. Qed. Global Instance envs_Forall2_refl (R : relation PROP) : @@ -705,11 +705,11 @@ Proof. by rewrite into_wand /= assoc wand_elim_r wand_elim_r. Qed. -Lemma tac_specialize_assert Δ Δ' Δ1 Δ2' j q lr js R P1 P2 P1' Q : +Lemma tac_specialize_assert Δ Δ' Δ1 Δ2' j q neg js R P1 P2 P1' Q : envs_lookup_delete j Δ = Some (q, R, Δ') → IntoWand q false R P1 P2 → ElimModal P1' P1 Q Q → - ('(Δ1,Δ2) ↠envs_split lr js Δ'; + ('(Δ1,Δ2) ↠envs_split (if neg is true then Right else Left) js Δ'; Δ2' ↠envs_app false (Esnoc Enil j P2) Δ2; Some (Δ1,Δ2')) = Some (Δ1,Δ2') → (* does not preserve position of [j] *) (Δ1 ⊢ P1') → (Δ2' ⊢ Q) → Δ ⊢ Q. @@ -819,9 +819,9 @@ Proof. by rewrite -{1}persistently_idemp !affinely_persistently_elim impl_elim_r. Qed. -Lemma tac_assert Δ Δ1 Δ2 Δ2' lr js j P P' Q : +Lemma tac_assert Δ Δ1 Δ2 Δ2' neg js j P P' Q : ElimModal P' P Q Q → - envs_split lr js Δ = Some (Δ1,Δ2) → + envs_split (if neg is true then Right else Left) js Δ = Some (Δ1,Δ2) → envs_app false (Esnoc Enil j P) Δ2 = Some Δ2' → (Δ1 ⊢ P') → (Δ2' ⊢ Q) → Δ ⊢ Q. Proof. @@ -829,8 +829,8 @@ Proof. rewrite (envs_app_singleton_sound Δ2) //; simpl. by rewrite HP HQ. Qed. -Lemma tac_assert_persistent Δ Δ1 Δ2 Δ' lr js j P P' Q : - envs_split lr js Δ = Some (Δ1,Δ2) → +Lemma tac_assert_persistent Δ Δ1 Δ2 Δ' neg js j P P' Q : + envs_split (if neg is true then Right else Left) js Δ = Some (Δ1,Δ2) → Persistent P → FromAffinely P' P → envs_app false (Esnoc Enil j P') Δ = Some Δ' → @@ -885,28 +885,28 @@ Proof. Qed. (** * Rewriting *) -Lemma tac_rewrite Δ i p Pxy (lr : bool) Q : +Lemma tac_rewrite Δ i p Pxy d Q : envs_lookup i Δ = Some (p, Pxy) → ∀ {A : ofeT} (x y : A) (Φ : A → PROP), IntoInternalEq Pxy x y → - (Q ⊣⊢ Φ (if lr then y else x)) → + (Q ⊣⊢ Φ (if d is Left then y else x)) → NonExpansive Φ → - (Δ ⊢ Φ (if lr then x else y)) → Δ ⊢ Q. + (Δ ⊢ Φ (if d is Left then x else y)) → Δ ⊢ Q. Proof. intros ? A x y ? HPxy -> ?; apply internal_eq_rewrite'; auto. rewrite {1}envs_lookup_sound //. rewrite HPxy affinely_persistently_if_elim sep_elim_l. - destruct lr; auto using internal_eq_sym. + destruct d; auto using internal_eq_sym. Qed. -Lemma tac_rewrite_in Δ i p Pxy j q P (lr : bool) Q : +Lemma tac_rewrite_in Δ i p Pxy j q P d Q : envs_lookup i Δ = Some (p, Pxy) → envs_lookup j Δ = Some (q, P) → ∀ {A : ofeT} Δ' (x y : A) (Φ : A → PROP), IntoInternalEq Pxy x y → - (P ⊣⊢ Φ (if lr then y else x)) → + (P ⊣⊢ Φ (if d is Left then y else x)) → NonExpansive Φ → - envs_simple_replace j q (Esnoc Enil j (Φ (if lr then x else y))) Δ = Some Δ' → + envs_simple_replace j q (Esnoc Enil j (Φ (if d is Left then x else y))) Δ = Some Δ' → (Δ' ⊢ Q) → Δ ⊢ Q. Proof. intros ?? A Δ' x y Φ HPxy HP ?? <-. @@ -914,7 +914,7 @@ Proof. rewrite (envs_simple_replace_singleton_sound _ _ j) //; simpl. rewrite HP HPxy (affinely_persistently_if_elim _ (_ ≡ _)%I) sep_elim_l. rewrite persistent_and_affinely_sep_r -assoc. apply wand_elim_r'. - rewrite -persistent_and_affinely_sep_r. apply impl_elim_r'. destruct lr. + rewrite -persistent_and_affinely_sep_r. apply impl_elim_r'. destruct d. - apply (internal_eq_rewrite x y (λ y, â–¡?q Φ y -∗ Δ')%I). solve_proper. - rewrite internal_eq_sym. eapply (internal_eq_rewrite y x (λ y, â–¡?q Φ y -∗ Δ')%I). solve_proper. @@ -926,9 +926,9 @@ Lemma tac_and_split Δ P Q1 Q2 : Proof. intros. rewrite -(from_and P). by apply and_intro. Qed. (** * Separating conjunction splitting *) -Lemma tac_sep_split Δ Δ1 Δ2 lr js P Q1 Q2 : +Lemma tac_sep_split Δ Δ1 Δ2 d js P Q1 Q2 : FromSep P Q1 Q2 → - envs_split lr js Δ = Some (Δ1,Δ2) → + envs_split d js Δ = Some (Δ1,Δ2) → (Δ1 ⊢ Q1) → (Δ2 ⊢ Q2) → Δ ⊢ P. Proof. intros ?? HQ1 HQ2. rewrite envs_split_sound //. by rewrite HQ1 HQ2. Qed. @@ -972,25 +972,25 @@ Qed. (* Using this tactic, one can destruct a (non-separating) conjunction in the spatial context as long as one of the conjuncts is thrown away. It corresponds to the principle of "external choice" in linear logic. *) -Lemma tac_and_destruct_choice Δ Δ' i p (lr : bool) j P P1 P2 Q : +Lemma tac_and_destruct_choice Δ Δ' i p d j P P1 P2 Q : envs_lookup i Δ = Some (p, P) → (if p then IntoAnd p P P1 P2 : Type else TCOr (IntoAnd p P P1 P2) (TCAnd (IntoSep P P1 P2) - (if lr then TCOr (Affine P2) (Absorbing Q) + (if d is Left then TCOr (Affine P2) (Absorbing Q) else TCOr (Affine P1) (Absorbing Q)))) → - envs_simple_replace i p (Esnoc Enil j (if lr then P1 else P2)) Δ = Some Δ' → + envs_simple_replace i p (Esnoc Enil j (if d is Left then P1 else P2)) Δ = Some Δ' → (Δ' ⊢ Q) → Δ ⊢ Q. Proof. intros ? HP ? HQ. rewrite envs_simple_replace_singleton_sound //; simpl. destruct p. - { rewrite (into_and _ P) HQ. destruct lr; simpl. + { rewrite (into_and _ P) HQ. destruct d; simpl. - by rewrite and_elim_l wand_elim_r. - by rewrite and_elim_r wand_elim_r. } destruct HP as [?|[??]]. - { rewrite (into_and _ P) HQ. destruct lr; simpl. + { rewrite (into_and _ P) HQ. destruct d; simpl. - by rewrite and_elim_l wand_elim_r. - by rewrite and_elim_r wand_elim_r. } - rewrite (into_sep P) HQ. destruct lr; simpl. + rewrite (into_sep P) HQ. destruct d; simpl. - by rewrite (comm _ P1) -assoc wand_elim_r sep_elim_r. - by rewrite -assoc wand_elim_r sep_elim_r. Qed. diff --git a/theories/proofmode/environments.v b/theories/proofmode/environments.v index 31164e0be062b55b1f53cfde17e2511541271e5e..53ceca1bd781f96fb7c07605633ea069fdc7b5e8 100644 --- a/theories/proofmode/environments.v +++ b/theories/proofmode/environments.v @@ -1,5 +1,5 @@ From stdpp Require Export strings. -From iris.proofmode Require Import strings. +From iris.proofmode Require Import base. From iris.algebra Require Export base. From stdpp Require Import stringmap. Set Default Proof Using "Type". diff --git a/theories/proofmode/intro_patterns.v b/theories/proofmode/intro_patterns.v index 7d1c308a4d5cb15c612fc21f1ee40456a218ff43..d36317e05e53835f8157b24c214dae57b9531f16 100644 --- a/theories/proofmode/intro_patterns.v +++ b/theories/proofmode/intro_patterns.v @@ -1,5 +1,5 @@ From stdpp Require Export strings. -From iris.proofmode Require Import tokens sel_patterns. +From iris.proofmode Require Import base tokens sel_patterns. Set Default Proof Using "Type". Inductive intro_pat := @@ -11,6 +11,7 @@ Inductive intro_pat := | IPureElim : intro_pat | IAlwaysElim : intro_pat → intro_pat | IModalElim : intro_pat → intro_pat + | IRewrite : direction → intro_pat | IPureIntro : intro_pat | IAlwaysIntro : intro_pat | IModalIntro : intro_pat @@ -84,6 +85,7 @@ Fixpoint parse_go (ts : list token) (k : stack) : option stack := | TPure :: ts => parse_go ts (SPat IPureElim :: k) | TAlways :: ts => parse_go ts (SAlwaysElim :: k) | TModal :: ts => parse_go ts (SModalElim :: k) + | TArrow d :: ts => parse_go ts (SPat (IRewrite d) :: k) | TPureIntro :: ts => parse_go ts (SPat IPureIntro :: k) | TAlwaysIntro :: ts => parse_go ts (SPat IAlwaysIntro :: k) | TModalIntro :: ts => parse_go ts (SPat IModalIntro :: k) diff --git a/theories/proofmode/sel_patterns.v b/theories/proofmode/sel_patterns.v index 5578d17e222562e2b933cd9532cb7097223f86f4..6fa4129af1736a2e6473014339bb2f716d5f5522 100644 --- a/theories/proofmode/sel_patterns.v +++ b/theories/proofmode/sel_patterns.v @@ -1,5 +1,5 @@ From stdpp Require Export strings. -From iris.proofmode Require Import tokens. +From iris.proofmode Require Import base tokens. Set Default Proof Using "Type". Inductive sel_pat := diff --git a/theories/proofmode/tactics.v b/theories/proofmode/tactics.v index f6ca7d33082c5491c3228157550958eee3f2e6bb..e27d3b1950e7205711ff2ba91b3df11eec19ddbb 100644 --- a/theories/proofmode/tactics.v +++ b/theories/proofmode/tactics.v @@ -1,10 +1,9 @@ From iris.proofmode Require Import coq_tactics. -From iris.proofmode Require Import intro_patterns spec_patterns sel_patterns. +From iris.proofmode Require Import base intro_patterns spec_patterns sel_patterns. From iris.bi Require Export bi. From iris.proofmode Require Export classes notation. From iris.proofmode Require Import class_instances. From stdpp Require Import stringmap hlist. -From iris.proofmode Require Import strings. Set Default Proof Using "Type". Declare Reduction env_cbv := cbv [ @@ -748,7 +747,7 @@ Tactic Notation "iSplit" := Tactic Notation "iSplitL" constr(Hs) := iStartProof; let Hs := words Hs in - eapply tac_sep_split with _ _ false Hs _ _; (* (js:=Hs) *) + eapply tac_sep_split with _ _ Left Hs _ _; (* (js:=Hs) *) [apply _ || let P := match goal with |- FromSep _ ?P _ _ => P end in fail "iSplitL:" P "not a separating conjunction" @@ -760,7 +759,7 @@ Tactic Notation "iSplitL" constr(Hs) := Tactic Notation "iSplitR" constr(Hs) := iStartProof; let Hs := words Hs in - eapply tac_sep_split with _ _ true Hs _ _; (* (js:=Hs) *) + eapply tac_sep_split with _ _ Right Hs _ _; (* (js:=Hs) *) [apply _ || let P := match goal with |- FromSep _ ?P _ _ => P end in fail "iSplitR:" P "not a separating conjunction" @@ -784,8 +783,8 @@ Local Tactic Notation "iAndDestruct" constr(H) "as" constr(H1) constr(H2) := fail "iAndDestruct: cannot destruct" P |env_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' _ _ _; +Local Tactic Notation "iAndDestructChoice" constr(H) "as" constr(d) constr(H') := + eapply tac_and_destruct_choice with _ H _ d H' _ _ _; [env_reflexivity || fail "iAndDestructChoice:" H "not found" |env_cbv; apply _ || let P := match goal with |- TCOr (IntoAnd _ ?P _ _) _ => P end in @@ -898,12 +897,14 @@ Tactic Notation "iDestructHyp" constr(H) "as" constr(pat) := | IFrame => iFrame Hz | IName ?y => iRename Hz into y | IList [[]] => iExFalso; iExact Hz - | IList [[?pat1; IDrop]] => iAndDestructChoice Hz as true Hz; go Hz pat1 - | IList [[IDrop; ?pat2]] => iAndDestructChoice Hz as false Hz; go Hz pat2 + | IList [[?pat1; IDrop]] => iAndDestructChoice Hz as Left Hz; go Hz pat1 + | IList [[IDrop; ?pat2]] => iAndDestructChoice Hz as Right Hz; go Hz pat2 | IList [[?pat1; ?pat2]] => 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 ? + | IRewrite Right => iPure Hz as -> + | IRewrite Left => iPure Hz as <- | IAlwaysElim ?pat => iPersistent Hz; go Hz pat | IModalElim ?pat => iModCore Hz; go Hz pat | _ => fail "iDestruct:" pat "invalid" @@ -1689,8 +1690,8 @@ Local Tactic Notation "iRewriteCore" constr(lr) open_constr(lem) := |iRewriteFindPred |intros ??? ->; reflexivity|lazy beta; iClear Heq]). -Tactic Notation "iRewrite" open_constr(lem) := iRewriteCore false lem. -Tactic Notation "iRewrite" "-" open_constr(lem) := iRewriteCore true lem. +Tactic Notation "iRewrite" open_constr(lem) := iRewriteCore Right lem. +Tactic Notation "iRewrite" "-" open_constr(lem) := iRewriteCore Left lem. Local Tactic Notation "iRewriteCore" constr(lr) open_constr(lem) "in" constr(H) := iPoseProofCore lem as true true (fun Heq => @@ -1705,9 +1706,9 @@ Local Tactic Notation "iRewriteCore" constr(lr) open_constr(lem) "in" constr(H) |env_reflexivity|lazy beta; iClear Heq]). Tactic Notation "iRewrite" open_constr(lem) "in" constr(H) := - iRewriteCore false lem in H. + iRewriteCore Right lem in H. Tactic Notation "iRewrite" "-" open_constr(lem) "in" constr(H) := - iRewriteCore true lem in H. + iRewriteCore Left lem in H. Ltac iSimplifyEq := repeat ( iMatchHyp (fun H P => match P with ⌜_ = _âŒ%I => iDestruct H as %? end) diff --git a/theories/proofmode/tokens.v b/theories/proofmode/tokens.v index c45c14862fc2328289f7a16bc26c7f0d49b2676d..47e04bec35747fe904bf5da01c2687ce41aecdee 100644 --- a/theories/proofmode/tokens.v +++ b/theories/proofmode/tokens.v @@ -1,4 +1,4 @@ -From stdpp Require Export strings. +From iris.proofmode Require Import base. Set Default Proof Using "Type". Inductive token := @@ -25,7 +25,8 @@ Inductive token := | TForall : token | TAll : token | TMinus : token - | TSep : token. + | TSep : token + | TArrow : direction → token. Inductive state := | SName : string → state @@ -64,6 +65,8 @@ Fixpoint tokenize_go (s : string) (k : list token) (kn : state) : list token := | String "/" (String "=" s) => tokenize_go s (TSimpl :: cons_state kn k) SNone | String "*" (String "*" s) => tokenize_go s (TAll :: cons_state kn k) SNone | String "*" s => tokenize_go s (TForall :: cons_state kn k) SNone + | String "-" (String ">" s) => tokenize_go s (TArrow Right :: cons_state kn k) SNone + | String "<" (String "-" s) => tokenize_go s (TArrow Left :: cons_state kn k) SNone | String "-" s => tokenize_go s (TMinus :: cons_state kn k) SNone | String (Ascii.Ascii false true false false false true true true) (* unicode ∗ *) (String (Ascii.Ascii false false false true false false false true) diff --git a/theories/tests/proofmode.v b/theories/tests/proofmode.v index f1feeb821c73892ce128f3e9b0314addbbd820a4..449f4e9b17b79b7f0ea603beffa9733e73455d6b 100644 --- a/theories/tests/proofmode.v +++ b/theories/tests/proofmode.v @@ -235,6 +235,10 @@ Proof. iIntros ([]). Qed. +Lemma test_iIntros_rewrite P (x1 x2 x3 x4 : nat) : + x1 = x2 → (⌜ x2 = x3 ⌠∗ ⌜ x3 ≡ x4 ⌠∗ P) -∗ ⌜ x1 = x4 ⌠∗ P. +Proof. iIntros (?) "(-> & -> & $)"; auto. Qed. + Lemma test_iNext_affine P Q : bi_affinely (â–· (Q ≡ P)) -∗ bi_affinely (â–· Q) -∗ bi_affinely (â–· P). Proof. iIntros "#HPQ HQ !#". iNext. by iRewrite "HPQ" in "HQ". Qed. @@ -242,4 +246,5 @@ Proof. iIntros "#HPQ HQ !#". iNext. by iRewrite "HPQ" in "HQ". Qed. Lemma test_iAlways P Q R : â–¡ P -∗ bi_persistently Q → R -∗ bi_persistently (bi_affinely (bi_affinely P)) ∗ â–¡ Q. Proof. iIntros "#HP #HQ HR". iSplitL. iAlways. done. iAlways. done. Qed. + End tests.