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
Pierre-Marie Pédrot
Iris
Commits
8a9515b0
Commit
8a9515b0
authored
Dec 14, 2017
by
Jacques-Henri Jourdan
Browse files
Options
Browse Files
Download
Email Patches
Plain Diff
Fix iStartProof.
parent
8ca83760
Changes
2
Show whitespace changes
Inline
Side-by-side
Showing
2 changed files
with
11 additions
and
2 deletions
+11
-2
theories/proofmode/tactics.v
theories/proofmode/tactics.v
+2
-2
theories/tests/proofmode_iris.v
theories/tests/proofmode_iris.v
+9
-0
No files found.
theories/proofmode/tactics.v
View file @
8a9515b0
...
...
@@ -79,9 +79,9 @@ Proof. rewrite /AsValid=> ->. rewrite bi_embed_valid //. Qed.
Tactic
Notation
"iStartProof"
uconstr
(
PROP
)
:
=
lazymatch
goal
with
|
|-
@
envs_entails
?PROP'
_
_
=>
let
x
:
=
type_term
(
eq_refl
:
PROP
=
PROP'
)
in
idtac
let
x
:
=
type_term
(
eq_refl
:
@
eq
Type
PROP
PROP'
)
in
idtac
|
|-
let
_
:
=
_
in
_
=>
fail
|
|-
?
φ
=>
eapply
(@
as_valid_2
φ
PRO
P
)
;
|
|-
?
φ
=>
eapply
(
λ
P
:
PROP
,
@
as_valid_2
φ
_
P
)
;
[
apply
_
||
fail
"iStartProof: not a Bi entailment"
|
apply
tac_adequate
]
end
.
...
...
theories/tests/proofmode_iris.v
View file @
8a9515b0
...
...
@@ -36,6 +36,15 @@ Section base_logic_tests.
Lemma
test_iAssert_modality
P
:
(|==>
False
)
-
∗
|==>
P
.
Proof
.
iIntros
.
iAssert
False
%
I
with
"[> - //]"
as
%[].
Qed
.
Lemma
test_iStartProof_1
P
:
P
-
∗
P
.
Proof
.
iStartProof
.
iStartProof
.
iIntros
"$"
.
Qed
.
Lemma
test_iStartProof_2
P
:
P
-
∗
P
.
Proof
.
iStartProof
(
uPred
_
).
iStartProof
(
uPredI
_
).
iIntros
"$"
.
Qed
.
Lemma
test_iStartProof_3
P
:
P
-
∗
P
.
Proof
.
iStartProof
(
uPredI
_
).
iStartProof
(
uPredSI
_
).
iIntros
"$"
.
Qed
.
Lemma
test_iStartProof_4
P
:
P
-
∗
P
.
Proof
.
iStartProof
(
uPredSI
_
).
iStartProof
(
uPred
_
).
iIntros
"$"
.
Qed
.
End
base_logic_tests
.
Section
iris_tests
.
...
...
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