Skip to content
GitLab
Projects
Groups
Snippets
/
Help
Help
Support
Community forum
Keyboard shortcuts
?
Submit feedback
Contribute to GitLab
Sign in / Register
Toggle navigation
Menu
Open sidebar
Marianna Rapoport
iris-coq
Commits
c43eb936
Commit
c43eb936
authored
Jul 27, 2016
by
Robbert Krebbers
Browse files
Hack to delay type class inference in iPoseProof.
parent
3e1e5e0f
Changes
2
Hide whitespace changes
Inline
Side-by-side
proofmode/tactics.v
View file @
c43eb936
...
...
@@ -208,6 +208,11 @@ Tactic Notation "iSpecialize" open_constr(t) :=
end
.
(** * Pose proof *)
(* We use [class_apply] in the below to delay type class inference, this is
useful when difference [inG] instances are arround. See [tests/proofmode] for
a simple but artificial example.
Note that this only works when the posed lemma is prefixed with an [@]. *)
Local
Tactic
Notation
"iPoseProofCore"
open_constr
(
H1
)
"as"
constr
(
H2
)
:
=
lazymatch
type
of
H1
with
|
string
=>
...
...
@@ -216,9 +221,10 @@ Local Tactic Notation "iPoseProofCore" open_constr(H1) "as" constr(H2) :=
|
env_cbv
;
reflexivity
||
fail
"iPoseProof:"
H2
"not fresh"
|]
|
_
=>
eapply
tac_pose_proof
with
_
H2
_
_
_;
(* (j:=H) *)
[
first
[
eapply
H1
|
apply
uPred
.
equiv_iff
;
eapply
H1
]
|
apply
_
|
env_cbv
;
reflexivity
||
fail
"iPoseProof:"
H2
"not fresh"
|]
[
first
[
class_apply
H1
|
class_apply
uPred
.
equiv_iff
;
eapply
H1
]
|
(* [apply _] invokes TC inference on shelved goals, why ...? *)
typeclasses
eauto
|
env_cbv
;
class_apply
eq_refl
||
fail
"iPoseProof:"
H2
"not fresh"
|]
end
.
Tactic
Notation
"iPoseProof"
open_constr
(
t
)
"as"
constr
(
H
)
:
=
...
...
tests/proofmode.v
View file @
c43eb936
...
...
@@ -106,3 +106,17 @@ Section iris.
-
done
.
Qed
.
End
iris
.
Section
classes
.
Class
C
A
:
=
c
:
A
.
Instance
nat_C
:
C
nat
:
=
0
.
Instance
bool_C
:
C
bool
:
=
true
.
Lemma
demo_9
{
M
:
ucmraT
}
(
P
:
uPred
M
)
(
H
:
∀
`
{
C
A
},
True
⊢
(
c
=
c
:
uPred
M
))
:
P
⊢
(
c
:
nat
)
=
c
∧
(
c
:
bool
)
=
c
.
Proof
.
iIntros
"_"
.
iPoseProof
(@
H
_
)
as
"foo_bool"
.
iPoseProof
(@
H
_
)
as
"foo_nat"
.
iSplit
.
iApply
"foo_nat"
.
iApply
"foo_bool"
.
Qed
.
End
classes
.
Write
Preview
Supports
Markdown
0%
Try again
or
attach a new 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