Skip to content
GitLab
Projects
Groups
Snippets
Help
Loading...
Help
What's new
10
Help
Support
Community forum
Keyboard shortcuts
?
Submit feedback
Contribute to GitLab
Sign in / Register
Toggle navigation
Open sidebar
Hugo Herbelin
iris-coq
Commits
6784ac89
Commit
6784ac89
authored
Dec 06, 2018
by
Robbert Krebbers
Browse files
Options
Browse Files
Download
Email Patches
Plain Diff
Fix `iSimpl .. in ..`.
Thanks to
@Blaisorblade
for reporting.
parent
84444c51
Changes
3
Hide whitespace changes
Inline
Side-by-side
Showing
3 changed files
with
17 additions
and
2 deletions
+17
-2
tests/proofmode.ref
tests/proofmode.ref
+11
-0
tests/proofmode.v
tests/proofmode.v
+4
-0
theories/proofmode/ltac_tactics.v
theories/proofmode/ltac_tactics.v
+2
-2
No files found.
tests/proofmode.ref
View file @
6784ac89
...
...
@@ -74,6 +74,17 @@ Tactic failure: iSpecialize: cannot instantiate (⌜φ⌝ → P -∗ False)%I wi
--------------------------------------∗
▷^(S n + S m) emp
"test_iSimpl_in"
: string
1 subgoal
PROP : sbi
x, y : nat
============================
"H" : ⌜S (S (S x)) = y⌝
--------------------------------------∗
⌜S (S (S x)) = y⌝
"test_iFrame_later_1"
: string
1 subgoal
...
...
tests/proofmode.v
View file @
6784ac89
...
...
@@ -412,6 +412,10 @@ Proof.
done
.
Qed
.
Check
"test_iSimpl_in"
.
Lemma
test_iSimpl_in
x
y
:
⌜
(
3
+
x
)%
nat
=
y
⌝
-
∗
⌜
S
(
S
(
S
x
))
=
y
⌝
:
PROP
.
Proof
.
iIntros
"H"
.
iSimpl
in
"H"
.
Show
.
done
.
Qed
.
Lemma
test_iIntros_pure_neg
:
(
⌜
¬
False
⌝
:
PROP
)%
I
.
Proof
.
by
iIntros
(?).
Qed
.
...
...
theories/proofmode/ltac_tactics.v
View file @
6784ac89
...
...
@@ -119,8 +119,8 @@ Tactic Notation "iEval" tactic(t) "in" constr(H) :=
|
pm_reflexivity
|].
Tactic
Notation
"iSimpl"
:
=
iEval
simpl
.
Tactic
Notation
"iSimpl"
"in"
constr
(
H
)
:
=
iEval
simpl
in
H
.
Tactic
Notation
"iSimpl"
:
=
iEval
(
simpl
)
.
Tactic
Notation
"iSimpl"
"in"
constr
(
H
)
:
=
iEval
(
simpl
)
in
H
.
(* It would be nice to also have an `iSsrRewrite`, however, for this we need to
pass arguments to Ssreflect's `rewrite` like `/= foo /bar` in Ltac, see:
...
...
Write
Preview
Markdown
is supported
0%
Try again
or
attach a new file
.
Attach a file
Cancel
You are about to add
0
people
to the discussion. Proceed with caution.
Finish editing this message first!
Cancel
Please
register
or
sign in
to comment