Skip to content
GitLab
Projects
Groups
Snippets
Help
Loading...
Help
Help
Support
Community forum
Keyboard shortcuts
?
Submit feedback
Contribute to GitLab
Sign in / Register
Toggle navigation
I
Iris
Project overview
Project overview
Details
Activity
Releases
Repository
Repository
Files
Commits
Branches
Tags
Contributors
Graph
Compare
Issues
116
Issues
116
List
Boards
Labels
Service Desk
Milestones
Merge Requests
22
Merge Requests
22
CI / CD
CI / CD
Pipelines
Jobs
Schedules
Operations
Operations
Incidents
Environments
Analytics
Analytics
CI / CD
Repository
Value Stream
Wiki
Wiki
Members
Members
Collapse sidebar
Close sidebar
Activity
Graph
Create a new issue
Jobs
Commits
Issue Boards
Open sidebar
Iris
Iris
Commits
86dec2b6
Commit
86dec2b6
authored
Mar 05, 2016
by
Robbert Krebbers
Browse files
Options
Browse Files
Download
Email Patches
Plain Diff
Implement sep_split tactics using reflection.
parent
b052b2a2
Changes
4
Hide whitespace changes
Inline
Side-by-side
Showing
4 changed files
with
96 additions
and
88 deletions
+96
-88
algebra/upred.v
algebra/upred.v
+4
-0
algebra/upred_tactics.v
algebra/upred_tactics.v
+91
-86
program_logic/hoare_lifting.v
program_logic/hoare_lifting.v
+1
-1
program_logic/sts.v
program_logic/sts.v
+0
-1
No files found.
algebra/upred.v
View file @
86dec2b6
...
@@ -336,6 +336,10 @@ Proof.
...
@@ -336,6 +336,10 @@ Proof.
-
by
trans
P1
;
[|
trans
Q1
].
-
by
trans
P1
;
[|
trans
Q1
].
-
by
trans
P2
;
[|
trans
Q2
].
-
by
trans
P2
;
[|
trans
Q2
].
Qed
.
Qed
.
Lemma
entails_equiv_l
(
P
Q
R
:
uPred
M
)
:
P
≡
Q
→
Q
⊑
R
→
P
⊑
R
.
Proof
.
by
intros
->.
Qed
.
Lemma
entails_equiv_r
(
P
Q
R
:
uPred
M
)
:
P
⊑
Q
→
Q
≡
R
→
P
⊑
R
.
Proof
.
by
intros
?
<-.
Qed
.
(** Non-expansiveness and setoid morphisms *)
(** Non-expansiveness and setoid morphisms *)
Global
Instance
const_proper
:
Proper
(
iff
==>
(
≡
))
(@
uPred_const
M
).
Global
Instance
const_proper
:
Proper
(
iff
==>
(
≡
))
(@
uPred_const
M
).
...
...
algebra/upred_tactics.v
View file @
86dec2b6
...
@@ -2,7 +2,7 @@ From algebra Require Export upred.
...
@@ -2,7 +2,7 @@ From algebra Require Export upred.
From
algebra
Require
Export
upred_big_op
.
From
algebra
Require
Export
upred_big_op
.
Import
uPred
.
Import
uPred
.
Module
u
pred_reflection
.
Section
up
red_reflection
.
Module
u
Pred_reflection
.
Section
uP
red_reflection
.
Context
{
M
:
cmraT
}.
Context
{
M
:
cmraT
}.
Inductive
expr
:
=
Inductive
expr
:
=
...
@@ -93,6 +93,45 @@ Module upred_reflection. Section upred_reflection.
...
@@ -93,6 +93,45 @@ Module upred_reflection. Section upred_reflection.
rewrite
!
fmap_app
!
big_sep_app
.
apply
sep_mono_r
.
rewrite
!
fmap_app
!
big_sep_app
.
apply
sep_mono_r
.
Qed
.
Qed
.
Fixpoint
to_expr
(
l
:
list
nat
)
:
expr
:
=
match
l
with
|
[]
=>
ETrue
|
[
n
]
=>
EVar
n
|
n
::
l
=>
ESep
(
EVar
n
)
(
to_expr
l
)
end
.
Arguments
to_expr
!
_
/
:
simpl
nomatch
.
Lemma
eval_to_expr
Σ
l
:
eval
Σ
(
to_expr
l
)
≡
eval_list
Σ
l
.
Proof
.
induction
l
as
[|
n1
[|
n2
l
]
IH
]
;
csimpl
;
rewrite
?right_id
//.
by
rewrite
IH
.
Qed
.
Fixpoint
remove_all
(
l
k
:
list
nat
)
:
option
(
list
nat
)
:
=
match
l
with
|
[]
=>
Some
k
|
n
::
l
=>
'
(
i
,
_
)
←
list_find
(
n
=)
k
;
remove_all
l
(
delete
i
k
)
end
.
Lemma
remove_all_permutation
l
k
k'
:
remove_all
l
k
=
Some
k'
→
k
≡
ₚ
l
++
k'
.
Proof
.
revert
k
k'
;
induction
l
as
[|
n
l
IH
]
;
simpl
;
intros
k
k'
Hk
.
{
by
simplify_eq
.
}
destruct
(
list_find
_
_
)
as
[[
i
?]|]
eqn
:
?Hk'
;
simplify_eq
/=.
move
:
Hk'
;
intros
[?
<-]%
list_find_Some
.
rewrite
-(
IH
(
delete
i
k
)
k'
)
//
-
delete_Permutation
//.
Qed
.
Lemma
split_l
Σ
e
l
k
:
remove_all
l
(
flatten
e
)
=
Some
k
→
eval
Σ
e
≡
eval
Σ
(
ESep
(
to_expr
l
)
(
to_expr
k
)).
Proof
.
intros
He
%
remove_all_permutation
.
by
rewrite
eval_flatten
He
fmap_app
big_sep_app
/=
!
eval_to_expr
.
Qed
.
Lemma
split_r
Σ
e
l
k
:
remove_all
l
(
flatten
e
)
=
Some
k
→
eval
Σ
e
≡
eval
Σ
(
ESep
(
to_expr
k
)
(
to_expr
l
)).
Proof
.
intros
.
rewrite
/=
comm
.
by
apply
split_l
.
Qed
.
Class
Quote
(
Σ
1
Σ
2
:
list
(
uPred
M
))
(
P
:
uPred
M
)
(
e
:
expr
)
:
=
{}.
Class
Quote
(
Σ
1
Σ
2
:
list
(
uPred
M
))
(
P
:
uPred
M
)
(
e
:
expr
)
:
=
{}.
Global
Instance
quote_True
Σ
:
Quote
Σ
Σ
True
ETrue
.
Global
Instance
quote_True
Σ
:
Quote
Σ
Σ
True
ETrue
.
Global
Instance
quote_var
Σ
1
Σ
2
P
i
:
Global
Instance
quote_var
Σ
1
Σ
2
P
i
:
...
@@ -105,116 +144,82 @@ Module upred_reflection. Section upred_reflection.
...
@@ -105,116 +144,82 @@ Module upred_reflection. Section upred_reflection.
Global
Instance
quote_args_cons
Σ
Ps
P
ns
n
:
Global
Instance
quote_args_cons
Σ
Ps
P
ns
n
:
rlist
.
QuoteLookup
Σ
Σ
P
n
→
rlist
.
QuoteLookup
Σ
Σ
P
n
→
QuoteArgs
Σ
Ps
ns
→
QuoteArgs
Σ
(
P
::
Ps
)
(
n
::
ns
).
QuoteArgs
Σ
Ps
ns
→
QuoteArgs
Σ
(
P
::
Ps
)
(
n
::
ns
).
End
uPred_reflection
.
End
upred_reflection
.
Ltac
quote
:
=
Ltac
quote
:
=
match
goal
with
match
goal
with
|
|-
?P1
⊑
?P2
=>
|
|-
?P1
⊑
?P2
=>
lazymatch
type
of
(
_
:
Quote
[]
_
P1
_
)
with
Quote
_
?
Σ
2
_
?e1
=>
lazymatch
type
of
(
_
:
Quote
[]
_
P1
_
)
with
Quote
_
?
Σ
2
_
?e1
=>
lazymatch
type
of
(
_
:
Quote
Σ
2
_
P2
_
)
with
Quote
_
?
Σ
3
_
?e2
=>
lazymatch
type
of
(
_
:
Quote
Σ
2
_
P2
_
)
with
Quote
_
?
Σ
3
_
?e2
=>
change
(
eval
Σ
3 e1
⊑
eval
Σ
3 e2
)
change
(
eval
Σ
3 e1
⊑
eval
Σ
3 e2
)
end
end
end
end
end
.
Ltac
quote_l
:
=
match
goal
with
|
|-
?P1
⊑
?P2
=>
lazymatch
type
of
(
_
:
Quote
[]
_
P1
_
)
with
Quote
_
?
Σ
2
_
?e1
=>
change
(
eval
Σ
2 e1
⊑
P2
)
end
end
.
end
.
End
u
p
red_reflection
.
End
u
P
red_reflection
.
Tactic
Notation
"solve_sep_entails"
:
=
Tactic
Notation
"solve_sep_entails"
:
=
u
pred_reflection
.
quote
;
apply
up
red_reflection
.
flatten_entails
;
u
Pred_reflection
.
quote
;
apply
uP
red_reflection
.
flatten_entails
;
apply
(
bool_decide_unpack
_
)
;
vm_compute
;
exact
Logic
.
I
.
apply
(
bool_decide_unpack
_
)
;
vm_compute
;
exact
Logic
.
I
.
Ltac
close_uPreds
Ps
tac
:
=
let
M
:
=
match
goal
with
|-
@
uPred_entails
?M
_
_
=>
M
end
in
let
rec
go
Ps
Qs
:
=
lazymatch
Ps
with
|
[]
=>
let
Qs'
:
=
eval
cbv
[
reverse
rev_append
]
in
(
reverse
Qs
)
in
tac
Qs'
|
?P
::
?Ps
=>
find_pat
P
ltac
:
(
fun
Q
=>
go
Ps
(
Q
::
Qs
))
end
in
(* avoid evars in case Ps = @nil ?A *)
try
match
Ps
with
[]
=>
unify
Ps
(@
nil
(
uPred
M
))
end
;
go
Ps
(@
nil
(
uPred
M
)).
Tactic
Notation
"cancel"
constr
(
Ps
)
:
=
Tactic
Notation
"cancel"
constr
(
Ps
)
:
=
upred_reflection
.
quote
;
uPred_reflection
.
quote
;
match
goal
with
let
Σ
:
=
match
goal
with
|-
uPred_reflection
.
eval
?
Σ
_
⊑
_
=>
Σ
end
in
|
|-
upred_reflection
.
eval
?
Σ
_
⊑
upred_reflection
.
eval
_
_
=>
let
ns'
:
=
lazymatch
type
of
(
_
:
uPred_reflection
.
QuoteArgs
Σ
Ps
_
)
with
lazymatch
type
of
(
_
:
upred_reflection
.
QuoteArgs
Σ
Ps
_
)
with
|
uPred_reflection
.
QuoteArgs
_
_
?ns'
=>
ns'
upred_reflection
.
QuoteArgs
_
_
?ns'
=>
end
in
eapply
upred_reflection
.
cancel_entails
with
(
ns
:
=
ns'
)
;
eapply
uPred_reflection
.
cancel_entails
with
(
ns
:
=
ns'
)
;
[
cbv
;
reflexivity
|
cbv
;
reflexivity
|
simpl
]
[
cbv
;
reflexivity
|
cbv
;
reflexivity
|
simpl
].
end
end
.
Tactic
Notation
"ecancel"
open_constr
(
Ps
)
:
=
Tactic
Notation
"ecancel"
open_constr
(
Ps
)
:
=
let
rec
close
Ps
Qs
tac
:
=
close_uPreds
Ps
ltac
:
(
fun
Qs
=>
cancel
Qs
).
lazymatch
Ps
with
|
[]
=>
tac
Qs
|
?P
::
?Ps
=>
find_pat
P
ltac
:
(
fun
Q
=>
close
Ps
(
Q
::
Qs
)
tac
)
end
in
lazymatch
goal
with
|
|-
@
uPred_entails
?M
_
_
=>
close
Ps
(@
nil
(
uPred
M
))
ltac
:
(
fun
Qs
=>
cancel
Qs
)
end
.
(** [to_front [P1, P2, ..]] rewrites in the premise of ⊑ such that
(** [to_front [P1, P2, ..]] rewrites in the premise of ⊑ such that
the assumptions P1, P2, ... appear at the front, in that order. *)
the assumptions P1, P2, ... appear at the front, in that order. *)
Tactic
Notation
"to_front"
open_constr
(
Ps
)
:
=
Tactic
Notation
"to_front"
open_constr
(
Ps
)
:
=
let
rec
tofront
Ps
:
=
close_uPreds
Ps
ltac
:
(
fun
Ps
=>
lazymatch
eval
hnf
in
Ps
with
uPred_reflection
.
quote_l
;
|
[]
=>
idtac
let
Σ
:
=
match
goal
with
|-
uPred_reflection
.
eval
?
Σ
_
⊑
_
=>
Σ
end
in
|
?P
::
?Ps
=>
let
ns'
:
=
lazymatch
type
of
(
_
:
uPred_reflection
.
QuoteArgs
Σ
Ps
_
)
with
rewrite
?(
assoc
(
★
)%
I
)
;
|
uPred_reflection
.
QuoteArgs
_
_
?ns'
=>
ns'
match
goal
with
end
in
|
|-
(
?Q
★
_
)%
I
⊑
_
=>
(* check if it is already at front. *)
eapply
entails_equiv_l
;
unify
P
Q
with
typeclass_instances
first
(
apply
uPred_reflection
.
split_l
with
(
l
:
=
ns'
)
;
cbv
;
reflexivity
)
;
|
|-
_
=>
find_pat
P
ltac
:
(
fun
P
=>
rewrite
{
1
}[(
_
★
P
)%
I
]
comm
)
simpl
).
end
;
tofront
Ps
Tactic
Notation
"to_back"
open_constr
(
Ps
)
:
=
end
close_uPreds
Ps
ltac
:
(
fun
Ps
=>
in
uPred_reflection
.
quote_l
;
rewrite
[
X
in
_
⊑
X
]
lock
;
let
Σ
:
=
match
goal
with
|-
uPred_reflection
.
eval
?
Σ
_
⊑
_
=>
Σ
end
in
tofront
(
rev
Ps
)
;
let
ns'
:
=
lazymatch
type
of
(
_
:
uPred_reflection
.
QuoteArgs
Σ
Ps
_
)
with
rewrite
-[
X
in
_
⊑
X
]
lock
.
|
uPred_reflection
.
QuoteArgs
_
_
?ns'
=>
ns'
end
in
eapply
entails_equiv_l
;
first
(
apply
uPred_reflection
.
split_r
with
(
l
:
=
ns'
)
;
cbv
;
reflexivity
)
;
simpl
).
(** [sep_split] is used to introduce a (★).
(** [sep_split] is used to introduce a (★).
Use [sep_split left: [P1, P2, ...]] to define which assertions will be
Use [sep_split left: [P1, P2, ...]] to define which assertions will be
taken to the left; the rest will be available on the right.
taken to the left; the rest will be available on the right.
[sep_split right: [P1, P2, ...]] works the other way around. *)
[sep_split right: [P1, P2, ...]] works the other way around. *)
(* TODO: These tactics fail if the list contains *all* assumptions.
However, in these cases using the "other" variant with the emtpy list
is much more convenient anyway. *)
(* TODO: These tactics are pretty slow, can we use the reflection stuff
above to speed them up? *)
Tactic
Notation
"sep_split"
"right:"
open_constr
(
Ps
)
:
=
Tactic
Notation
"sep_split"
"right:"
open_constr
(
Ps
)
:
=
match
goal
with
to_back
Ps
;
apply
sep_mono
.
|
|-
?P
⊑
_
=>
let
iProp
:
=
type
of
P
in
(* ascribe a type to the list, to prevent evars from appearing *)
lazymatch
eval
hnf
in
(
Ps
:
list
iProp
)
with
|
[]
=>
apply
sep_intro_True_r
|
?P
::
?Ps
=>
to_front
(
P
::
Ps
)
;
(* Run assoc length (Ps) times -- that is 1 - length(original Ps),
and it will turn the goal in just the right shape for sep_mono. *)
let
rec
nassoc
Ps
:
=
lazymatch
eval
hnf
in
Ps
with
|
[]
=>
idtac
|
_
::
?Ps
=>
rewrite
(
assoc
(
★
)%
I
)
;
nassoc
Ps
end
in
rewrite
[
X
in
_
⊑
X
]
lock
-?(
assoc
(
★
)%
I
)
;
nassoc
Ps
;
rewrite
[
X
in
X
⊑
_
]
comm
-[
X
in
_
⊑
X
]
lock
;
apply
sep_mono
end
end
.
Tactic
Notation
"sep_split"
"left:"
open_constr
(
Ps
)
:
=
Tactic
Notation
"sep_split"
"left:"
open_constr
(
Ps
)
:
=
match
goal
with
to_front
Ps
;
apply
sep_mono
.
|
|-
?P
⊑
_
=>
let
iProp
:
=
type
of
P
in
(* ascribe a type to the list, to prevent evars from appearing *)
lazymatch
eval
hnf
in
(
Ps
:
list
iProp
)
with
|
[]
=>
apply
sep_intro_True_l
|
?P
::
?Ps
=>
to_front
(
P
::
Ps
)
;
(* Run assoc length (Ps) times -- that is 1 - length(original Ps),
and it will turn the goal in just the right shape for sep_mono. *)
let
rec
nassoc
Ps
:
=
lazymatch
eval
hnf
in
Ps
with
|
[]
=>
idtac
|
_
::
?Ps
=>
rewrite
(
assoc
(
★
)%
I
)
;
nassoc
Ps
end
in
rewrite
[
X
in
_
⊑
X
]
lock
-?(
assoc
(
★
)%
I
)
;
nassoc
Ps
;
rewrite
-[
X
in
_
⊑
X
]
lock
;
apply
sep_mono
end
end
.
(** Assumes a goal of the shape P ⊑ ▷ Q. Alterantively, if Q
(** Assumes a goal of the shape P ⊑ ▷ Q. Alterantively, if Q
is built of ★, ∧, ∨ with ▷ in all branches; that will work, too.
is built of ★, ∧, ∨ with ▷ in all branches; that will work, too.
...
...
program_logic/hoare_lifting.v
View file @
86dec2b6
...
@@ -91,7 +91,7 @@ Proof.
...
@@ -91,7 +91,7 @@ Proof.
rewrite
(
forall_elim
e2
)
(
forall_elim
ef
).
rewrite
(
forall_elim
e2
)
(
forall_elim
ef
).
rewrite
always_and_sep_l
!
always_and_sep_r
{
1
}(
always_sep_dup
(
■
_
)).
rewrite
always_and_sep_l
!
always_and_sep_r
{
1
}(
always_sep_dup
(
■
_
)).
sep_split
left
:
[
■
φ
_
_;
P
;
{{
■
φ
_
_
★
P
}}
e2
@
E
{{
Ψ
}}]%
I
.
sep_split
left
:
[
■
φ
_
_;
P
;
{{
■
φ
_
_
★
P
}}
e2
@
E
{{
Ψ
}}]%
I
.
-
rewrite
{
1
}/
ht
-
always_wand_impl
always_elim
wand_elim_r
//.
-
rewrite
assoc
{
1
}/
ht
-
always_wand_impl
always_elim
wand_elim_r
//.
-
destruct
ef
as
[
e'
|]
;
simpl
;
[|
by
apply
const_intro
].
-
destruct
ef
as
[
e'
|]
;
simpl
;
[|
by
apply
const_intro
].
rewrite
assoc
{
1
}/
ht
-
always_wand_impl
always_elim
wand_elim_r
//.
rewrite
assoc
{
1
}/
ht
-
always_wand_impl
always_elim
wand_elim_r
//.
Qed
.
Qed
.
...
...
program_logic/sts.v
View file @
86dec2b6
...
@@ -100,7 +100,6 @@ Section sts.
...
@@ -100,7 +100,6 @@ Section sts.
rewrite
/
sts_inv
later_exist
sep_exist_r
.
apply
exist_elim
=>
s
.
rewrite
/
sts_inv
later_exist
sep_exist_r
.
apply
exist_elim
=>
s
.
rewrite
later_sep
pvs_timeless
!
pvs_frame_r
.
apply
pvs_mono
.
rewrite
later_sep
pvs_timeless
!
pvs_frame_r
.
apply
pvs_mono
.
rewrite
-(
exist_intro
s
).
ecancel
[
▷
φ
_
]%
I
.
rewrite
-(
exist_intro
s
).
ecancel
[
▷
φ
_
]%
I
.
to_front
[
own
_
_;
sts_ownS
_
_
_
].
rewrite
-
own_op
own_valid_l
discrete_valid
.
rewrite
-
own_op
own_valid_l
discrete_valid
.
apply
const_elim_sep_l
=>
Hvalid
.
apply
const_elim_sep_l
=>
Hvalid
.
assert
(
s
∈
S
)
by
eauto
using
sts_auth_frag_valid_inv
.
assert
(
s
∈
S
)
by
eauto
using
sts_auth_frag_valid_inv
.
...
...
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