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
George Pirlea
Iris
Commits
37675f7b
Commit
37675f7b
authored
May 04, 2016
by
Robbert Krebbers
Browse files
Options
Browse Files
Download
Email Patches
Plain Diff
Make iPureIntro also work with uPred_valid and uPred_eq.
parent
677c3e3a
Changes
3
Hide whitespace changes
Inline
Side-by-side
Showing
3 changed files
with
12 additions
and
8 deletions
+12
-8
proofmode/coq_tactics.v
proofmode/coq_tactics.v
+4
-4
proofmode/tactics.v
proofmode/tactics.v
+5
-2
tests/proofmode.v
tests/proofmode.v
+3
-2
No files found.
proofmode/coq_tactics.v
View file @
37675f7b
...
...
@@ -292,10 +292,7 @@ Lemma tac_ex_falso Δ Q : Δ ⊢ False → Δ ⊢ Q.
Proof
.
by
rewrite
-(
False_elim
Q
).
Qed
.
(** * Pure *)
Lemma
tac_pure_intro
Δ
(
φ
:
Prop
)
:
φ
→
Δ
⊢
■
φ
.
Proof
.
apply
const_intro
.
Qed
.
Class
ToPure
(
P
:
uPred
M
)
(
φ
:
Prop
)
:
=
to_pure
:
P
⊢
■
φ
.
Class
ToPure
(
P
:
uPred
M
)
(
φ
:
Prop
)
:
=
to_pure
:
P
⊣
⊢
■
φ
.
Arguments
to_pure
:
clear
implicits
.
Global
Instance
to_pure_const
φ
:
ToPure
(
■
φ
)
φ
.
Proof
.
done
.
Qed
.
...
...
@@ -305,6 +302,9 @@ Proof. intros; red. by rewrite timeless_eq. Qed.
Global
Instance
to_pure_valid
`
{
CMRADiscrete
A
}
(
a
:
A
)
:
ToPure
(
✓
a
)
(
✓
a
).
Proof
.
intros
;
red
.
by
rewrite
discrete_valid
.
Qed
.
Lemma
tac_pure_intro
Δ
Q
(
φ
:
Prop
)
:
ToPure
Q
φ
→
φ
→
Δ
⊢
Q
.
Proof
.
intros
->.
apply
const_intro
.
Qed
.
Lemma
tac_pure
Δ
Δ
'
i
p
P
φ
Q
:
envs_lookup_delete
i
Δ
=
Some
(
p
,
P
,
Δ
'
)
→
ToPure
P
φ
→
(
φ
→
Δ
'
⊢
Q
)
→
Δ
⊢
Q
.
...
...
proofmode/tactics.v
View file @
37675f7b
...
...
@@ -113,7 +113,10 @@ Local Tactic Notation "iPure" constr(H) "as" simple_intropattern(pat) :=
apply
_
||
fail
"iPure:"
H
":"
P
"not pure"
|
intros
pat
].
Tactic
Notation
"iPureIntro"
:
=
apply
uPred
.
const_intro
.
Tactic
Notation
"iPureIntro"
:
=
eapply
tac_pure_intro
;
[
let
P
:
=
match
goal
with
|-
ToPure
?P
_
=>
P
end
in
apply
_
||
fail
"iPureIntro:"
P
"not pure"
|].
(** * Specialize *)
Record
iTrm
{
X
As
}
:
=
...
...
@@ -751,5 +754,5 @@ Tactic Notation "iRewrite" "-" open_constr(t) "in" constr(H) :=
iRewriteCore
true
t
in
H
.
(* Make sure that by and done solve trivial things in proof mode *)
Hint
Extern
0
(
of_envs
_
⊢
_
)
=>
by
apply
tac_p
ure
_i
ntro
.
Hint
Extern
0
(
of_envs
_
⊢
_
)
=>
by
iP
ure
I
ntro
.
Hint
Extern
0
(
of_envs
_
⊢
_
)
=>
iAssumption
.
tests/proofmode.v
View file @
37675f7b
...
...
@@ -72,9 +72,10 @@ Proof.
Qed
.
Lemma
demo_6
(
M
:
cmraT
)
(
P
Q
:
uPred
M
)
:
True
⊢
(
∀
x
y
z
,
x
=
0
→
y
=
0
→
z
=
0
→
P
→
□
Q
→
foo
False
).
True
⊢
(
∀
x
y
z
:
nat
,
x
=
plus
0
x
→
y
=
0
→
z
=
0
→
P
→
□
Q
→
foo
(
x
≡
x
)).
Proof
.
iIntros
{
a
}
"*"
.
iIntros
"#Hfoo **"
.
by
iIntros
"
%
"
.
by
iIntros
"
# _
"
.
Qed
.
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