Skip to content
GitLab
Menu
Projects
Groups
Snippets
Help
Help
Support
Community forum
Keyboard shortcuts
?
Submit feedback
Contribute to GitLab
Sign in / Register
Toggle navigation
Menu
Open sidebar
Rodolphe Lepigre
Iris
Commits
7ba11ac0
Commit
7ba11ac0
authored
May 25, 2019
by
Robbert Krebbers
Browse files
More simple changes to make proof mode terms more compact.
This is a follow up of !248.
parent
a82dbaf9
Changes
2
Hide whitespace changes
Inline
Side-by-side
theories/proofmode/coq_tactics.v
View file @
7ba11ac0
...
...
@@ -252,16 +252,19 @@ Qed.
(* This is pretty much [tac_specialize_assert] with [js:=[j]] and [tac_exact],
but it is doing some work to keep the order of hypotheses preserved. *)
(* TODO: convert to not take Δ'
or Δ''
*)
Lemma
tac_specialize
remove_intuitionistic
Δ
Δ
'
Δ
''
i
p
j
q
P1
P2
R
Q
:
(* TODO: convert to not take Δ' *)
Lemma
tac_specialize
remove_intuitionistic
Δ
Δ
'
i
p
j
q
P1
P2
R
Q
:
envs_lookup_delete
remove_intuitionistic
i
Δ
=
Some
(
p
,
P1
,
Δ
'
)
→
envs_lookup
j
Δ
'
=
Some
(
q
,
R
)
→
IntoWand
q
p
R
P1
P2
→
envs_replace
j
q
(
p
&&
q
)
(
Esnoc
Enil
j
P2
)
Δ
'
=
Some
Δ
''
→
envs_entails
Δ
''
Q
→
envs_entails
Δ
Q
.
match
envs_replace
j
q
(
p
&&
q
)
(
Esnoc
Enil
j
P2
)
Δ
'
with
|
Some
Δ
''
=>
envs_entails
Δ
''
Q
|
None
=>
False
end
→
envs_entails
Δ
Q
.
Proof
.
rewrite
envs_entails_eq
/
IntoWand
.
intros
[?
->]%
envs_lookup_delete_Some
?
HR
?
<-.
intros
[?
->]%
envs_lookup_delete_Some
?
HR
?.
destruct
(
envs_replace
_
_
_
_
_
)
as
[
Δ
''
|]
eqn
:
?
;
last
done
.
rewrite
(
envs_lookup_sound'
_
remove_intuitionistic
)
//.
rewrite
envs_replace_singleton_sound
//.
destruct
p
;
simpl
in
*.
-
rewrite
-{
1
}
intuitionistically_idemp
-{
1
}
intuitionistically_if_idemp
.
...
...
@@ -327,16 +330,20 @@ Proof.
by
rewrite
intuitionistically_emp
left_id
wand_elim_r
.
Qed
.
Lemma
tac_specialize_assert_intuitionistic
Δ
Δ
'
Δ
''
j
q
P1
P1'
P2
R
Q
:
Lemma
tac_specialize_assert_intuitionistic
Δ
Δ
'
j
q
P1
P1'
P2
R
Q
:
envs_lookup_delete
true
j
Δ
=
Some
(
q
,
R
,
Δ
'
)
→
IntoWand
q
true
R
P1
P2
→
Persistent
P1
→
IntoAbsorbingly
P1'
P1
→
envs_simple_replace
j
q
(
Esnoc
Enil
j
P2
)
Δ
=
Some
Δ
''
→
envs_entails
Δ
'
P1'
→
envs_entails
Δ
''
Q
→
envs_entails
Δ
Q
.
envs_entails
Δ
'
P1'
→
match
envs_simple_replace
j
q
(
Esnoc
Enil
j
P2
)
Δ
with
|
Some
Δ
''
=>
envs_entails
Δ
''
Q
|
None
=>
False
end
→
envs_entails
Δ
Q
.
Proof
.
rewrite
envs_entails_eq
=>
/
envs_lookup_delete_Some
[?
->]
????
HP1
<-.
rewrite
envs_lookup_sound
//.
rewrite
envs_entails_eq
=>
/
envs_lookup_delete_Some
[?
->]
???
HP1
HQ
.
destruct
(
envs_simple_replace
_
_
_
_
)
as
[
Δ
''
|]
eqn
:
?
;
last
done
.
rewrite
-
HQ
envs_lookup_sound
//.
rewrite
-(
idemp
bi_and
(
of_envs
(
envs_delete
_
_
_
_
))).
rewrite
{
2
}
envs_simple_replace_singleton_sound'
//
;
simpl
.
rewrite
{
1
}
HP1
(
into_absorbingly
P1'
P1
)
(
persistent_persistently_2
P1
).
...
...
@@ -346,15 +353,19 @@ Proof.
by
rewrite
intuitionistically_if_sep_2
(
into_wand
q
true
)
wand_elim_l
wand_elim_r
.
Qed
.
Lemma
tac_specialize_intuitionistic_helper
Δ
Δ
''
j
q
P
R
R'
Q
:
Lemma
tac_specialize_intuitionistic_helper
Δ
j
q
P
R
R'
Q
:
envs_lookup
j
Δ
=
Some
(
q
,
P
)
→
(
if
q
then
TCTrue
else
BiAffine
PROP
)
→
envs_entails
Δ
(<
absorb
>
R
)
→
IntoPersistent
false
R
R'
→
envs_replace
j
q
true
(
Esnoc
Enil
j
R'
)
Δ
=
Some
Δ
''
→
envs_entails
Δ
''
Q
→
envs_entails
Δ
Q
.
match
envs_replace
j
q
true
(
Esnoc
Enil
j
R'
)
Δ
with
|
Some
Δ
''
=>
envs_entails
Δ
''
Q
|
None
=>
False
end
→
envs_entails
Δ
Q
.
Proof
.
rewrite
envs_entails_eq
=>
??
HR
??
<-.
rewrite
-(
idemp
bi_and
(
of_envs
Δ
))
{
1
}
HR
.
rewrite
envs_entails_eq
=>
??
HR
??.
destruct
(
envs_replace
_
_
_
_
_
)
as
[
Δ
'
|]
eqn
:
?
;
last
done
.
rewrite
-(
idemp
bi_and
(
of_envs
Δ
))
{
1
}
HR
.
rewrite
envs_replace_singleton_sound
//
;
destruct
q
;
simpl
.
-
by
rewrite
(
_
:
R
=
<
pers
>
?false
R
)%
I
//
(
into_persistent
_
R
)
absorbingly_elim_persistently
sep_elim_r
...
...
@@ -374,12 +385,16 @@ Proof.
rewrite
intuitionistically_if_elim
comm
.
f_equiv
;
auto
using
pure_intro
.
Qed
.
Lemma
tac_revert
Δ
Δ
'
i
p
P
Q
:
envs_lookup_delete
true
i
Δ
=
Some
(
p
,
P
,
Δ
'
)
→
envs_entails
Δ
'
((
if
p
then
□
P
else
P
)%
I
-
∗
Q
)
→
Lemma
tac_revert
Δ
i
Q
:
match
envs_lookup_delete
true
i
Δ
with
|
Some
(
p
,
P
,
Δ
'
)
=>
envs_entails
Δ
'
((
if
p
then
□
P
else
P
)%
I
-
∗
Q
)
|
None
=>
False
end
→
envs_entails
Δ
Q
.
Proof
.
rewrite
envs_entails_eq
=>
?
HQ
.
rewrite
envs_lookup_delete_sound
//=.
rewrite
envs_entails_eq
=>
HQ
.
destruct
(
envs_lookup_delete
_
_
_
)
as
[[[
p
P
]
Δ
'
]|]
eqn
:
?
;
last
done
.
rewrite
envs_lookup_delete_sound
//=.
rewrite
HQ
.
destruct
p
;
simpl
;
auto
using
wand_elim_r
.
Qed
.
...
...
theories/proofmode/ltac_tactics.v
View file @
7ba11ac0
...
...
@@ -582,15 +582,16 @@ Local Tactic Notation "iForallRevert" ident(x) :=
(** The tactic [iRevertHyp H with tac] reverts the hypothesis [H] and calls
[tac] with a Boolean that is [true] iff [H] was in the intuitionistic context. *)
Tactic
Notation
"iRevertHyp"
constr
(
H
)
"with"
tactic1
(
tac
)
:
=
(* Create a Boolean evar [p] to keep track of whether the hypothesis [H] was
in the intuitionistic context. *)
let
p
:
=
fresh
in
evar
(
p
:
bool
)
;
let
p'
:
=
eval
unfold
p
in
p
in
clear
p
;
eapply
tac_revert
with
_
H
p'
_;
[
pm_reflexivity
||
let
H
:
=
pretty_ident
H
in
fail
"iRevert:"
H
"not found"
|
pm_reduce
;
tac
p'
].
eapply
tac_revert
with
H
;
[
lazymatch
goal
with
|
|-
match
envs_lookup_delete
true
?i
?
Δ
with
_
=>
_
end
=>
lazymatch
eval
pm_eval
in
(
envs_lookup_delete
true
i
Δ
)
with
|
Some
(
?p
,
_
,
_
)
=>
pm_reduce
;
tac
p
|
None
=>
let
H
:
=
pretty_ident
H
in
fail
"iRevert:"
H
"not found"
end
end
].
Tactic
Notation
"iRevertHyp"
constr
(
H
)
:
=
iRevertHyp
H
with
(
fun
_
=>
idtac
).
...
...
@@ -783,7 +784,7 @@ Ltac iSpecializePat_go H1 pats :=
|
SIdent
?H2
[]
::
?pats
=>
(* If we not need to specialize [H2] we can avoid a lot of unncessary
context manipulation. *)
notypeclasses
refine
(
tac_specialize
false
_
_
_
H2
_
H1
_
_
_
_
_
_
_
_
_
_
)
;
notypeclasses
refine
(
tac_specialize
false
_
_
H2
_
H1
_
_
_
_
_
_
_
_
_
)
;
[
pm_reflexivity
||
let
H2
:
=
pretty_ident
H2
in
fail
"iSpecialize:"
H2
"not found"
...
...
@@ -794,7 +795,7 @@ Ltac iSpecializePat_go H1 pats :=
let
P
:
=
match
goal
with
|-
IntoWand
_
_
?P
?Q
_
=>
P
end
in
let
Q
:
=
match
goal
with
|-
IntoWand
_
_
?P
?Q
_
=>
Q
end
in
fail
"iSpecialize: cannot instantiate"
P
"with"
Q
|
pm_re
flexivity
|
iSpecializePat_go
H1
pats
]
|
pm_re
duce
;
iSpecializePat_go
H1
pats
]
|
SIdent
?H2
?pats1
::
?pats
=>
(* If [H2] is in the intuitionistic context, we copy it into a new
hypothesis [Htmp], so that it can be used multiple times. *)
...
...
@@ -810,7 +811,7 @@ Ltac iSpecializePat_go H1 pats :=
Ltac backtraces (which would otherwise include the whole closure). *)
[..
(* side-conditions of [iSpecialize] *)
|
(* Use [remove_intuitionistic = true] to remove the copy [Htmp]. *)
notypeclasses
refine
(
tac_specialize
true
_
_
_
H2tmp
_
H1
_
_
_
_
_
_
_
_
_
_
)
;
notypeclasses
refine
(
tac_specialize
true
_
_
H2tmp
_
H1
_
_
_
_
_
_
_
_
_
)
;
[
pm_reflexivity
||
let
H2tmp
:
=
pretty_ident
H2tmp
in
fail
"iSpecialize:"
H2tmp
"not found"
...
...
@@ -821,7 +822,7 @@ Ltac iSpecializePat_go H1 pats :=
let
P
:
=
match
goal
with
|-
IntoWand
_
_
?P
?Q
_
=>
P
end
in
let
Q
:
=
match
goal
with
|-
IntoWand
_
_
?P
?Q
_
=>
Q
end
in
fail
"iSpecialize: cannot instantiate"
P
"with"
Q
|
pm_re
flexivity
|
iSpecializePat_go
H1
pats
]]
|
pm_re
duce
;
iSpecializePat_go
H1
pats
]]
|
SPureGoal
?d
::
?pats
=>
notypeclasses
refine
(
tac_specialize_assert_pure
_
H1
_
_
_
_
_
_
_
_
_
_
_
_
)
;
[
pm_reflexivity
||
...
...
@@ -835,7 +836,7 @@ Ltac iSpecializePat_go H1 pats :=
|
pm_reduce
;
iSpecializePat_go
H1
pats
]
|
SGoal
(
SpecGoal
GIntuitionistic
false
?Hs_frame
[]
?d
)
::
?pats
=>
notypeclasses
refine
(
tac_specialize_assert_intuitionistic
_
_
_
H1
_
_
_
_
_
_
_
_
_
_
_
_
_
)
;
notypeclasses
refine
(
tac_specialize_assert_intuitionistic
_
_
H1
_
_
_
_
_
_
_
_
_
_
_
_
)
;
[
pm_reflexivity
||
let
H1
:
=
pretty_ident
H1
in
fail
"iSpecialize:"
H1
"not found"
...
...
@@ -844,9 +845,8 @@ Ltac iSpecializePat_go H1 pats :=
let
Q
:
=
match
goal
with
|-
Persistent
?Q
=>
Q
end
in
fail
"iSpecialize:"
Q
"not persistent"
|
iSolveTC
|
pm_reflexivity
|
iFrame
Hs_frame
;
solve_done
d
(*goal*)
|
iSpecializePat_go
H1
pats
]
|
pm_reduce
;
iSpecializePat_go
H1
pats
]
|
SGoal
(
SpecGoal
GIntuitionistic
_
_
_
_
)
::
?pats
=>
fail
"iSpecialize: cannot select hypotheses for intuitionistic premise"
|
SGoal
(
SpecGoal
?m
?lr
?Hs_frame
?Hs
?d
)
::
?pats
=>
...
...
@@ -866,7 +866,7 @@ Ltac iSpecializePat_go H1 pats :=
|
iFrame
Hs_frame
;
solve_done
d
(*goal*)
|
iSpecializePat_go
H1
pats
]
|
SAutoFrame
GIntuitionistic
::
?pats
=>
notypeclasses
refine
(
tac_specialize_assert_intuitionistic
_
_
_
H1
_
_
_
_
_
_
_
_
_
_
_
_
_
)
;
notypeclasses
refine
(
tac_specialize_assert_intuitionistic
_
_
H1
_
_
_
_
_
_
_
_
_
_
_
_
)
;
[
pm_reflexivity
||
let
H1
:
=
pretty_ident
H1
in
fail
"iSpecialize:"
H1
"not found"
...
...
@@ -874,9 +874,8 @@ Ltac iSpecializePat_go H1 pats :=
|
iSolveTC
||
let
Q
:
=
match
goal
with
|-
Persistent
?Q
=>
Q
end
in
fail
"iSpecialize:"
Q
"not persistent"
|
pm_reflexivity
|
solve
[
iFrame
"∗ #"
]
|
iSpecializePat_go
H1
pats
]
|
pm_reduce
;
iSpecializePat_go
H1
pats
]
|
SAutoFrame
?m
::
?pats
=>
notypeclasses
refine
(
tac_specialize_frame
_
_
H1
_
_
_
_
_
_
_
_
_
_
_
_
)
;
[
pm_reflexivity
||
...
...
@@ -943,7 +942,7 @@ Tactic Notation "iSpecializeCore" open_constr(H)
let
PROP
:
=
iBiOfGoal
in
lazymatch
eval
compute
in
(
q
||
tc_to_bool
(
BiAffine
PROP
))
with
|
true
=>
notypeclasses
refine
(
tac_specialize_intuitionistic_helper
_
_
H
_
_
_
_
_
_
_
_
_
_
_
)
;
notypeclasses
refine
(
tac_specialize_intuitionistic_helper
_
H
_
_
_
_
_
_
_
_
_
_
)
;
[
pm_reflexivity
(* This premise, [envs_lookup j Δ = Some (q,P)],
holds because [iTypeOf] succeeded *)
...
...
@@ -957,8 +956,7 @@ Tactic Notation "iSpecializeCore" open_constr(H)
|
iSolveTC
||
let
Q
:
=
match
goal
with
|-
IntoPersistent
_
?Q
_
=>
Q
end
in
fail
"iSpecialize:"
Q
"not persistent"
|
pm_reflexivity
|
(* goal *)
]
|
pm_reduce
(* goal *)
]
|
false
=>
iSpecializePat
H
pat
end
|
None
=>
...
...
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