Skip to content
GitLab
Projects
Groups
Snippets
Help
Loading...
Help
What's new
7
Help
Support
Community forum
Keyboard shortcuts
?
Submit feedback
Contribute to GitLab
Sign in / Register
Toggle navigation
Open sidebar
Iris
Fairis
Commits
d350e4f6
Commit
d350e4f6
authored
Feb 23, 2016
by
Robbert Krebbers
Browse files
Options
Browse Files
Download
Email Patches
Plain Diff
A very simple separation logic canceler by reflection.
parent
457a11d9
Changes
3
Hide whitespace changes
Inline
Side-by-side
Showing
3 changed files
with
169 additions
and
59 deletions
+169
-59
_CoqProject
_CoqProject
+1
-0
algebra/upred_tactics.v
algebra/upred_tactics.v
+123
-0
barrier/barrier.v
barrier/barrier.v
+45
-59
No files found.
_CoqProject
View file @
d350e4f6
...
@@ -51,6 +51,7 @@ algebra/excl.v
...
@@ -51,6 +51,7 @@ algebra/excl.v
algebra/iprod.v
algebra/iprod.v
algebra/functor.v
algebra/functor.v
algebra/upred.v
algebra/upred.v
algebra/upred_tactics.v
algebra/upred_big_op.v
algebra/upred_big_op.v
program_logic/model.v
program_logic/model.v
program_logic/adequacy.v
program_logic/adequacy.v
...
...
algebra/upred_tactics.v
0 → 100644
View file @
d350e4f6
From
algebra
Require
Export
upred
.
From
algebra
Require
Export
upred_big_op
.
Module
upred_reflection
.
Section
upred_reflection
.
Context
{
M
:
cmraT
}
.
Inductive
expr
:=
|
ETrue
:
expr
|
EVar
:
nat
→
expr
|
ESep
:
expr
→
expr
→
expr
.
Fixpoint
eval
(
Σ
:
list
(
uPred
M
))
(
e
:
expr
)
:
uPred
M
:=
match
e
with
|
ETrue
=>
True
|
EVar
n
=>
from_option
True
%
I
(
Σ
!!
n
)
|
ESep
e1
e2
=>
eval
Σ
e1
★
eval
Σ
e2
end
.
Fixpoint
flatten
(
e
:
expr
)
:
list
nat
:=
match
e
with
|
ETrue
=>
[]
|
EVar
n
=>
[
n
]
|
ESep
e1
e2
=>
flatten
e1
++
flatten
e2
end
.
Notation
eval_list
Σ
l
:=
(
uPred_big_sep
((
λ
n
,
from_option
True
%
I
(
Σ
!!
n
))
<
$
>
l
)).
Lemma
eval_flatten
Σ
e
:
eval
Σ
e
≡
eval_list
Σ
(
flatten
e
).
Proof
.
induction
e
as
[
|
|
e1
IH1
e2
IH2
];
rewrite
/=
?
right_id
?
fmap_app
?
big_sep_app
?
IH1
?
IH2
//.
Qed
.
Lemma
flatten_entails
Σ
e1
e2
:
flatten
e2
`contains
`
flatten
e1
→
eval
Σ
e1
⊑
eval
Σ
e2
.
Proof
.
intros
.
rewrite
!
eval_flatten
.
by
apply
big_sep_contains
,
fmap_contains
.
Qed
.
Lemma
flatten_equiv
Σ
e1
e2
:
flatten
e2
≡ₚ
flatten
e1
→
eval
Σ
e1
≡
eval
Σ
e2
.
Proof
.
intros
He
.
by
rewrite
!
eval_flatten
He
.
Qed
.
Fixpoint
prune
(
e
:
expr
)
:
expr
:=
match
e
with
|
ETrue
=>
ETrue
|
EVar
n
=>
EVar
n
|
ESep
e1
e2
=>
match
prune
e1
,
prune
e2
with
|
ETrue
,
e2
'
=>
e2
'
|
e1
'
,
ETrue
=>
e1
'
|
e1
'
,
e2
'
=>
ESep
e1
'
e2
'
end
end
.
Lemma
flatten_prune
e
:
flatten
(
prune
e
)
=
flatten
e
.
Proof
.
induction
e
as
[
|
|
e1
IH1
e2
IH2
];
simplify_eq
/=
;
auto
.
rewrite
-
IH1
-
IH2
.
by
repeat
case_match
;
rewrite
?
right_id_L
.
Qed
.
Lemma
prune_correct
Σ
e
:
eval
Σ
(
prune
e
)
≡
eval
Σ
e
.
Proof
.
by
rewrite
!
eval_flatten
flatten_prune
.
Qed
.
Fixpoint
cancel_go
(
n
:
nat
)
(
e
:
expr
)
:
option
expr
:=
match
e
with
|
ETrue
=>
None
|
EVar
n
'
=>
if
decide
(
n
=
n
'
)
then
Some
ETrue
else
None
|
ESep
e1
e2
=>
match
cancel_go
n
e1
with
|
Some
e1
'
=>
Some
(
ESep
e1
'
e2
)
|
None
=>
ESep
e1
<
$
>
cancel_go
n
e2
end
end
.
Definition
cancel
(
n
:
nat
)
(
e
:
expr
)
:
option
expr
:=
prune
<
$
>
cancel_go
n
e
.
Lemma
flatten_cancel_go
e
e
'
n
:
cancel_go
n
e
=
Some
e
'
→
flatten
e
≡ₚ
n
::
flatten
e
'
.
Proof
.
revert
e
'
;
induction
e
as
[
|
|
e1
IH1
e2
IH2
];
intros
;
repeat
(
simplify_option_eq
||
case_match
);
auto
.
*
by
rewrite
IH1
//.
*
by
rewrite
IH2
// Permutation_middle.
Qed
.
Lemma
flatten_cancel
e
e
'
n
:
cancel
n
e
=
Some
e
'
→
flatten
e
≡ₚ
n
::
flatten
e
'
.
Proof
.
rewrite
/
cancel
fmap_Some
=>
-
[
e
''
[
?
->
]].
by
rewrite
flatten_prune
-
flatten_cancel_go
.
Qed
.
Lemma
cancel_entails
Σ
e1
e2
e1
'
e2
'
n
:
cancel
n
e1
=
Some
e1
'
→
cancel
n
e2
=
Some
e2
'
→
eval
Σ
e1
'
⊑
eval
Σ
e2
'
→
eval
Σ
e1
⊑
eval
Σ
e2
.
Proof
.
intros
??
.
rewrite
!
eval_flatten
.
rewrite
(
flatten_cancel
e1
e1
'
n
)
// (flatten_cancel e2 e2' n) //; csimpl.
apply
uPred
.
sep_mono_r
.
Qed
.
Class
Quote
(
Σ
1
Σ
2
:
list
(
uPred
M
))
(
P
:
uPred
M
)
(
e
:
expr
)
:=
{}
.
Global
Instance
quote_True
Σ
:
Quote
Σ
Σ
True
ETrue
.
Global
Instance
quote_var
Σ
1
Σ
2
P
i
:
rlist
.
QuoteLookup
Σ
1
Σ
2
P
i
→
Quote
Σ
1
Σ
2
P
(
EVar
i
)
|
1000.
Global
Instance
quote_sep
Σ
1
Σ
2
Σ
3
P1
P2
e1
e2
:
Quote
Σ
1
Σ
2
P1
e1
→
Quote
Σ
2
Σ
3
P2
e2
→
Quote
Σ
1
Σ
3
(
P1
★
P2
)
(
ESep
e1
e2
).
End
upred_reflection
.
Ltac
quote
:=
match
goal
with
|
|-
?
P1
⊑
?
P2
=>
lazymatch
type
of
(
_
:
Quote
[]
_
P1
_
)
with
Quote
_
?
Σ
2
_
?
e1
=>
lazymatch
type
of
(
_
:
Quote
Σ
2
_
P2
_
)
with
Quote
_
?
Σ
3
_
?
e2
=>
change
(
eval
Σ
3
e1
⊑
eval
Σ
3
e2
)
end
end
end
.
End
upred_reflection
.
Tactic
Notation
"cancel"
constr
(
P
)
:=
let
rec
lookup
Σ
n
:=
match
Σ
with
|
P
::
_
=>
n
|
_
::
?
Σ
=>
lookup
Σ
(
S
n
)
end
in
upred_reflection
.
quote
;
match
goal
with
|
|-
upred_reflection
.
eval
?
Σ
_
⊑
upred_reflection
.
eval
_
_
=>
let
n
'
:=
lookup
Σ
0
%
nat
in
eapply
upred_reflection
.
cancel_entails
with
(
n
:=
n
'
);
[
cbv
;
reflexivity
|
cbv
;
reflexivity
|
simpl
]
end
.
barrier/barrier.v
View file @
d350e4f6
From
prelude
Require
Export
functions
.
From
prelude
Require
Export
functions
.
From
algebra
Require
Export
upred_big_op
.
From
algebra
Require
Export
upred_big_op
upred_tactics
.
From
program_logic
Require
Export
sts
saved_prop
.
From
program_logic
Require
Export
sts
saved_prop
.
From
program_logic
Require
Import
hoare
.
From
program_logic
Require
Import
hoare
.
From
heap_lang
Require
Export
derived
heap
wp_tactics
notation
.
From
heap_lang
Require
Export
derived
heap
wp_tactics
notation
.
...
@@ -407,37 +407,30 @@ Section proof.
...
@@ -407,37 +407,30 @@ Section proof.
rewrite
left_id
-
later_intro
{
1
3
}/
barrier_inv
.
rewrite
left_id
-
later_intro
{
1
3
}/
barrier_inv
.
(
*
FIXME
ssreflect
rewrite
fails
if
there
are
evars
around
.
Also
,
this
is
very
slow
because
we
don
'
t
have
a
proof
mode
.
*
)
(
*
FIXME
ssreflect
rewrite
fails
if
there
are
evars
around
.
Also
,
this
is
very
slow
because
we
don
'
t
have
a
proof
mode
.
*
)
rewrite
-
(
waiting_split
_
_
_
Q
R1
R2
);
[
|
done
..].
rewrite
-
(
waiting_split
_
_
_
Q
R1
R2
);
[
|
done
..].
match
goal
with
|
|-
_
⊑
?
G
=>
rewrite
[
G
]
lock
end
.
rewrite
{
1
}
[
saved_prop_own
i1
_
]
always_sep_dup
.
rewrite
{
1
}
[
saved_prop_own
i1
_
]
always_sep_dup
.
rewrite
{
1
}
[
saved_prop_own
i2
_
]
always_sep_dup
.
rewrite
{
1
}
[
saved_prop_own
i2
_
]
always_sep_dup
.
rewrite
!
assoc
[(
_
★
_
i1
_
)
%
I
]
comm
.
cancel
(
saved_prop_own
i1
R1
).
rewrite
!
assoc
[(
_
★
_
i
_
)
%
I
]
comm
.
cancel
(
saved_prop_own
i2
R2
).
rewrite
!
assoc
[(
_
★
(
l
↦
_
))
%
I
]
comm
.
cancel
(
l
↦
'0
)
%
I
.
rewrite
!
assoc
[(
_
★
(
waiting
_
_
))
%
I
]
comm
.
cancel
(
waiting
P
I
).
rewrite
!
assoc
[(
_
★
(
Q
-
★
_
))
%
I
]
comm
-!
assoc
5
!
assoc
.
cancel
(
Q
-
★
R1
★
R2
)
%
I
.
unlock
.
apply
sep_mono
.
cancel
(
saved_prop_own
i
Q
).
+
(
*
This
should
really
all
be
handled
automatically
.
*
)
apply
wand_intro_l
.
rewrite
!
assoc
.
eapply
pvs_wand_r
.
rewrite
/
recv
.
rewrite
!
assoc
[(
_
★
(
l
↦
_
))
%
I
]
comm
-!
assoc
.
apply
sep_mono_r
.
rewrite
-
(
exist_intro
γ
)
-
(
exist_intro
P
)
-
(
exist_intro
R1
)
-
(
exist_intro
i1
).
rewrite
!
assoc
[(
_
★
_
i2
_
)
%
I
]
comm
-!
assoc
.
apply
sep_mono_r
.
rewrite
-
(
exist_intro
γ
)
-
(
exist_intro
P
)
-
(
exist_intro
R2
)
-
(
exist_intro
i2
).
rewrite
!
assoc
[(
_
★
_
i1
_
)
%
I
]
comm
-!
assoc
.
apply
sep_mono_r
.
do
2
rewrite
!
(
assoc
(
★
)
%
I
)
[(
_
★
sts_ownS
_
_
_
)
%
I
]
comm
.
rewrite
!
assoc
[(
_
★
_
i
_
)
%
I
]
comm
-!
assoc
.
apply
sep_mono_r
.
rewrite
-!
assoc
.
rewrite
[(
sts_ownS
_
_
_
★
_
★
_
)
%
I
]
assoc
-
pvs_frame_r
.
done
.
apply
sep_mono
.
+
apply
wand_intro_l
.
rewrite
!
assoc
.
eapply
pvs_wand_r
.
rewrite
/
recv
.
*
rewrite
-
sts_ownS_op
;
by
eauto
using
sts_own_weaken
with
sts
.
rewrite
-
(
exist_intro
γ
)
-
(
exist_intro
P
)
-
(
exist_intro
R1
)
-
(
exist_intro
i1
).
*
rewrite
const_equiv
// !left_id.
rewrite
-
(
exist_intro
γ
)
-
(
exist_intro
P
)
-
(
exist_intro
R2
)
-
(
exist_intro
i2
).
rewrite
{
1
}
[
heap_ctx
_
]
always_sep_dup
!
assoc
[(
_
★
heap_ctx
_
)
%
I
]
comm
-!
assoc
.
apply
sep_mono_r
.
do
2
rewrite
!
(
assoc
(
★
)
%
I
)
[(
_
★
sts_ownS
_
_
_
)
%
I
]
comm
.
rewrite
!
assoc
!
[(
_
★
heap_ctx
_
)
%
I
]
comm
-!
assoc
.
apply
sep_mono_r
.
rewrite
-!
assoc
.
rewrite
[(
sts_ownS
_
_
_
★
_
★
_
)
%
I
]
assoc
-
pvs_frame_r
.
rewrite
{
1
}
[
sts_ctx
_
_
_
]
always_sep_dup
!
assoc
[(
_
★
sts_ctx
_
_
_
)
%
I
]
comm
-!
assoc
.
apply
sep_mono_r
.
apply
sep_mono
.
rewrite
!
assoc
!
[(
_
★
sts_ctx
_
_
_
)
%
I
]
comm
-!
assoc
.
apply
sep_mono_r
.
*
rewrite
-
sts_ownS_op
;
by
eauto
using
sts_own_weaken
with
sts
.
rewrite
comm
.
apply
sep_mono_r
.
apply
sep_intro_True_l
.
*
rewrite
const_equiv
// !left_id.
{
rewrite
-
later_intro
.
apply
wand_intro_l
.
by
rewrite
right_id
.
}
rewrite
{
1
}
[
heap_ctx
_
]
always_sep_dup
!
assoc
[(
_
★
heap_ctx
_
)
%
I
]
comm
-!
assoc
.
apply
sep_mono_r
.
apply
sep_intro_True_r
;
first
done
.
rewrite
!
assoc
!
[(
_
★
heap_ctx
_
)
%
I
]
comm
-!
assoc
.
apply
sep_mono_r
.
{
rewrite
-
later_intro
.
apply
wand_intro_l
.
by
rewrite
right_id
.
}
rewrite
{
1
}
[
sts_ctx
_
_
_
]
always_sep_dup
!
assoc
[(
_
★
sts_ctx
_
_
_
)
%
I
]
comm
-!
assoc
.
apply
sep_mono_r
.
rewrite
!
assoc
!
[(
_
★
sts_ctx
_
_
_
)
%
I
]
comm
-!
assoc
.
apply
sep_mono_r
.
rewrite
comm
.
apply
sep_mono_r
.
apply
sep_intro_True_l
.
{
rewrite
-
later_intro
.
apply
wand_intro_l
.
by
rewrite
right_id
.
}
apply
sep_intro_True_r
;
first
done
.
{
rewrite
-
later_intro
.
apply
wand_intro_l
.
by
rewrite
right_id
.
}
(
*
Case
II
:
High
state
.
TODO
:
Lots
of
this
script
is
just
copy
-
n
-
paste
of
the
previous
one
.
(
*
Case
II
:
High
state
.
TODO
:
Lots
of
this
script
is
just
copy
-
n
-
paste
of
the
previous
one
.
Most
of
that
is
because
the
goals
are
fairly
similar
in
structure
,
and
the
proof
scripts
Most
of
that
is
because
the
goals
are
fairly
similar
in
structure
,
and
the
proof
scripts
are
mostly
concerned
with
manually
managaing
the
structure
(
assoc
,
comm
,
dup
)
of
are
mostly
concerned
with
manually
managaing
the
structure
(
assoc
,
comm
,
dup
)
of
...
@@ -447,37 +440,30 @@ Section proof.
...
@@ -447,37 +440,30 @@ Section proof.
rewrite
const_equiv
;
last
by
eauto
with
sts
.
rewrite
const_equiv
;
last
by
eauto
with
sts
.
rewrite
left_id
-
later_intro
{
1
3
}/
barrier_inv
.
rewrite
left_id
-
later_intro
{
1
3
}/
barrier_inv
.
rewrite
-
(
ress_split
_
_
_
Q
R1
R2
);
[
|
done
..].
rewrite
-
(
ress_split
_
_
_
Q
R1
R2
);
[
|
done
..].
match
goal
with
|
|-
_
⊑
?
G
=>
rewrite
[
G
]
lock
end
.
rewrite
{
1
}
[
saved_prop_own
i1
_
]
always_sep_dup
.
rewrite
{
1
}
[
saved_prop_own
i1
_
]
always_sep_dup
.
rewrite
{
1
}
[
saved_prop_own
i2
_
]
always_sep_dup
.
rewrite
{
1
}
[
saved_prop_own
i2
_
]
always_sep_dup
.
rewrite
!
assoc
[(
_
★
_
i1
_
)
%
I
]
comm
.
cancel
(
saved_prop_own
i1
R1
).
rewrite
!
assoc
[(
_
★
_
i
_
)
%
I
]
comm
.
cancel
(
saved_prop_own
i2
R2
).
rewrite
!
assoc
[(
_
★
(
l
↦
_
))
%
I
]
comm
.
cancel
(
l
↦
'1
)
%
I
.
rewrite
!
assoc
[(
_
★
(
ress
_
))
%
I
]
comm
.
cancel
(
Q
-
★
R1
★
R2
)
%
I
.
rewrite
!
assoc
[(
_
★
(
Q
-
★
_
))
%
I
]
comm
-!
assoc
5
!
assoc
.
cancel
(
saved_prop_own
i
Q
).
unlock
.
apply
sep_mono
.
cancel
(
ress
I
).
+
(
*
This
should
really
all
be
handled
automatically
.
*
)
apply
wand_intro_l
.
rewrite
!
assoc
.
eapply
pvs_wand_r
.
rewrite
/
recv
.
rewrite
!
assoc
[(
_
★
(
l
↦
_
))
%
I
]
comm
-!
assoc
.
apply
sep_mono_r
.
rewrite
-
(
exist_intro
γ
)
-
(
exist_intro
P
)
-
(
exist_intro
R1
)
-
(
exist_intro
i1
).
rewrite
!
assoc
[(
_
★
_
i2
_
)
%
I
]
comm
-!
assoc
.
apply
sep_mono_r
.
rewrite
-
(
exist_intro
γ
)
-
(
exist_intro
P
)
-
(
exist_intro
R2
)
-
(
exist_intro
i2
).
rewrite
!
assoc
[(
_
★
_
i1
_
)
%
I
]
comm
-!
assoc
.
apply
sep_mono_r
.
do
2
rewrite
!
(
assoc
(
★
)
%
I
)
[(
_
★
sts_ownS
_
_
_
)
%
I
]
comm
.
rewrite
!
assoc
[(
_
★
_
i
_
)
%
I
]
comm
-!
assoc
.
apply
sep_mono_r
.
rewrite
-!
assoc
.
rewrite
[(
sts_ownS
_
_
_
★
_
★
_
)
%
I
]
assoc
-
pvs_frame_r
.
done
.
apply
sep_mono
.
+
apply
wand_intro_l
.
rewrite
!
assoc
.
eapply
pvs_wand_r
.
rewrite
/
recv
.
*
rewrite
-
sts_ownS_op
;
by
eauto
using
sts_own_weaken
with
sts
.
rewrite
-
(
exist_intro
γ
)
-
(
exist_intro
P
)
-
(
exist_intro
R1
)
-
(
exist_intro
i1
).
*
rewrite
const_equiv
// !left_id.
rewrite
-
(
exist_intro
γ
)
-
(
exist_intro
P
)
-
(
exist_intro
R2
)
-
(
exist_intro
i2
).
rewrite
{
1
}
[
heap_ctx
_
]
always_sep_dup
!
assoc
[(
_
★
heap_ctx
_
)
%
I
]
comm
-!
assoc
.
apply
sep_mono_r
.
do
2
rewrite
!
(
assoc
(
★
)
%
I
)
[(
_
★
sts_ownS
_
_
_
)
%
I
]
comm
.
rewrite
!
assoc
!
[(
_
★
heap_ctx
_
)
%
I
]
comm
-!
assoc
.
apply
sep_mono_r
.
rewrite
-!
assoc
.
rewrite
[(
sts_ownS
_
_
_
★
_
★
_
)
%
I
]
assoc
-
pvs_frame_r
.
rewrite
{
1
}
[
sts_ctx
_
_
_
]
always_sep_dup
!
assoc
[(
_
★
sts_ctx
_
_
_
)
%
I
]
comm
-!
assoc
.
apply
sep_mono_r
.
apply
sep_mono
.
rewrite
!
assoc
!
[(
_
★
sts_ctx
_
_
_
)
%
I
]
comm
-!
assoc
.
apply
sep_mono_r
.
*
rewrite
-
sts_ownS_op
;
by
eauto
using
sts_own_weaken
with
sts
.
rewrite
comm
.
apply
sep_mono_r
.
apply
sep_intro_True_l
.
*
rewrite
const_equiv
// !left_id.
{
rewrite
-
later_intro
.
apply
wand_intro_l
.
by
rewrite
right_id
.
}
rewrite
{
1
}
[
heap_ctx
_
]
always_sep_dup
!
assoc
[(
_
★
heap_ctx
_
)
%
I
]
comm
-!
assoc
.
apply
sep_mono_r
.
apply
sep_intro_True_r
;
first
done
.
rewrite
!
assoc
!
[(
_
★
heap_ctx
_
)
%
I
]
comm
-!
assoc
.
apply
sep_mono_r
.
{
rewrite
-
later_intro
.
apply
wand_intro_l
.
by
rewrite
right_id
.
}
rewrite
{
1
}
[
sts_ctx
_
_
_
]
always_sep_dup
!
assoc
[(
_
★
sts_ctx
_
_
_
)
%
I
]
comm
-!
assoc
.
apply
sep_mono_r
.
rewrite
!
assoc
!
[(
_
★
sts_ctx
_
_
_
)
%
I
]
comm
-!
assoc
.
apply
sep_mono_r
.
rewrite
comm
.
apply
sep_mono_r
.
apply
sep_intro_True_l
.
{
rewrite
-
later_intro
.
apply
wand_intro_l
.
by
rewrite
right_id
.
}
apply
sep_intro_True_r
;
first
done
.
{
rewrite
-
later_intro
.
apply
wand_intro_l
.
by
rewrite
right_id
.
}
Qed
.
Qed
.
Lemma
recv_strengthen
l
P1
P2
:
Lemma
recv_strengthen
l
P1
P2
:
...
...
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