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
Simon Spies
Iris
Commits
08b69a69
Commit
08b69a69
authored
Feb 20, 2016
by
Ralf Jung
Browse files
add tactic aliases: trans -> transitivity; etrans -> etransitivity
parent
709da735
Changes
31
Hide whitespace changes
Inline
Side-by-side
algebra/agree.v
View file @
08b69a69
...
...
@@ -40,8 +40,8 @@ Proof.
+
by
split
.
+
by
intros
x
y
Hxy
;
split
;
intros
;
symmetry
;
apply
Hxy
;
auto
;
apply
Hxy
.
+
intros
x
y
z
Hxy
Hyz
;
split
;
intros
n'
;
intros
.
*
trans
itivity
(
agree_is_valid
y
n'
).
by
apply
Hxy
.
by
apply
Hyz
.
*
trans
itivity
(
y
n'
).
by
apply
Hxy
.
by
apply
Hyz
,
Hxy
.
*
trans
(
agree_is_valid
y
n'
).
by
apply
Hxy
.
by
apply
Hyz
.
*
trans
(
y
n'
).
by
apply
Hxy
.
by
apply
Hyz
,
Hxy
.
-
intros
n
x
y
Hxy
;
split
;
intros
;
apply
Hxy
;
auto
.
-
intros
n
c
;
apply
and_wlog_r
;
intros
;
symmetry
;
apply
(
chain_cauchy
c
)
;
naive_solver
.
...
...
@@ -74,8 +74,8 @@ Proof.
intros
n
x
y1
y2
[
Hy'
Hy
]
;
split
;
[|
done
].
split
;
intros
(?&?&
Hxy
)
;
repeat
(
intro
||
split
)
;
try
apply
Hy'
;
eauto
using
agree_valid_le
.
-
etrans
itivity
;
[
apply
Hxy
|
apply
Hy
]
;
eauto
using
agree_valid_le
.
-
etrans
itivity
;
[
apply
Hxy
|
symmetry
;
apply
Hy
,
Hy'
]
;
-
etrans
;
[
apply
Hxy
|
apply
Hy
]
;
eauto
using
agree_valid_le
.
-
etrans
;
[
apply
Hxy
|
symmetry
;
apply
Hy
,
Hy'
]
;
eauto
using
agree_valid_le
.
Qed
.
Instance
:
Proper
(
dist
n
==>
dist
n
==>
dist
n
)
(@
op
(
agree
A
)
_
).
...
...
algebra/auth.v
View file @
08b69a69
...
...
@@ -44,7 +44,7 @@ Proof.
-
intros
n
;
split
.
+
by
intros
?
;
split
.
+
by
intros
??
[??]
;
split
;
symmetry
.
+
intros
???
[??]
[??]
;
split
;
etrans
itivity
;
eauto
.
+
intros
???
[??]
[??]
;
split
;
etrans
;
eauto
.
-
by
intros
?
[??]
[??]
[??]
;
split
;
apply
dist_S
.
-
intros
n
c
;
split
.
apply
(
conv_compl
n
(
chain_map
authoritative
c
)).
apply
(
conv_compl
n
(
chain_map
own
c
)).
...
...
algebra/cmra.v
View file @
08b69a69
...
...
@@ -250,7 +250,7 @@ Qed.
Global
Instance
cmra_included_preorder
:
PreOrder
(@
included
A
_
_
).
Proof
.
split
;
red
;
intros
until
0
;
rewrite
!
cmra_included_includedN
;
first
done
.
intros
;
etrans
itivity
;
eauto
.
intros
;
etrans
;
eauto
.
Qed
.
Lemma
cmra_validN_includedN
n
x
y
:
✓
{
n
}
y
→
x
≼
{
n
}
y
→
✓
{
n
}
x
.
Proof
.
intros
Hyv
[
z
?]
;
cofe_subst
y
;
eauto
using
cmra_validN_op_l
.
Qed
.
...
...
@@ -391,7 +391,7 @@ Section identity_updates.
Lemma
cmra_update_empty
x
:
x
~~>
∅
.
Proof
.
intros
n
z
;
rewrite
left_id
;
apply
cmra_validN_op_r
.
Qed
.
Lemma
cmra_update_empty_alt
y
:
∅
~~>
y
↔
∀
x
,
x
~~>
y
.
Proof
.
split
;
[
intros
;
trans
itivity
∅
|]
;
auto
using
cmra_update_empty
.
Qed
.
Proof
.
split
;
[
intros
;
trans
∅
|]
;
auto
using
cmra_update_empty
.
Qed
.
End
identity_updates
.
End
cmra
.
...
...
algebra/cmra_big_op.v
View file @
08b69a69
...
...
@@ -23,7 +23,7 @@ Proof.
induction
1
as
[|
x
xs1
xs2
?
IH
|
x
y
xs
|
xs1
xs2
xs3
]
;
simpl
;
auto
.
-
by
rewrite
IH
.
-
by
rewrite
!
assoc
(
comm
_
x
).
-
by
trans
itivity
(
big_op
xs2
).
-
by
trans
(
big_op
xs2
).
Qed
.
Global
Instance
big_op_proper
:
Proper
((
≡
)
==>
(
≡
))
big_op
.
Proof
.
by
induction
1
;
simpl
;
repeat
apply
(
_
:
Proper
(
_
==>
_
==>
_
)
op
).
Qed
.
...
...
algebra/cofe.v
View file @
08b69a69
...
...
@@ -99,13 +99,13 @@ Section cofe.
split
.
-
by
intros
x
;
rewrite
equiv_dist
.
-
by
intros
x
y
;
rewrite
!
equiv_dist
.
-
by
intros
x
y
z
;
rewrite
!
equiv_dist
;
intros
;
trans
itivity
y
.
-
by
intros
x
y
z
;
rewrite
!
equiv_dist
;
intros
;
trans
y
.
Qed
.
Global
Instance
dist_ne
n
:
Proper
(
dist
n
==>
dist
n
==>
iff
)
(@
dist
A
_
n
).
Proof
.
intros
x1
x2
?
y1
y2
?
;
split
;
intros
.
-
by
trans
itivity
x1
;
[|
trans
itivity
y1
].
-
by
trans
itivity
x2
;
[|
trans
itivity
y2
].
-
by
trans
x1
;
[|
trans
y1
].
-
by
trans
x2
;
[|
trans
y2
].
Qed
.
Global
Instance
dist_proper
n
:
Proper
((
≡
)
==>
(
≡
)
==>
iff
)
(@
dist
A
_
n
).
Proof
.
...
...
@@ -217,7 +217,7 @@ Section cofe_mor.
-
intros
n
;
split
.
+
by
intros
f
x
.
+
by
intros
f
g
?
x
.
+
by
intros
f
g
h
??
x
;
trans
itivity
(
g
x
).
+
by
intros
f
g
h
??
x
;
trans
(
g
x
).
-
by
intros
n
f
g
?
x
;
apply
dist_S
.
-
intros
n
c
x
;
simpl
.
by
rewrite
(
conv_compl
n
(
fun_chain
c
x
))
/=.
...
...
@@ -352,7 +352,7 @@ Section later.
-
intros
[|
n
]
;
[
by
split
|
split
]
;
unfold
dist
,
later_dist
.
+
by
intros
[
x
].
+
by
intros
[
x
]
[
y
].
+
by
intros
[
x
]
[
y
]
[
z
]
??
;
trans
itivity
y
.
+
by
intros
[
x
]
[
y
]
[
z
]
??
;
trans
y
.
-
intros
[|
n
]
[
x
]
[
y
]
?
;
[
done
|]
;
unfold
dist
,
later_dist
;
by
apply
dist_S
.
-
intros
[|
n
]
c
;
[
done
|
by
apply
(
conv_compl
n
(
later_chain
c
))].
Qed
.
...
...
algebra/cofe_solver.v
View file @
08b69a69
...
...
@@ -71,7 +71,7 @@ Proof.
-
intros
k
;
split
.
+
by
intros
X
n
.
+
by
intros
X
Y
?
n
.
+
by
intros
X
Y
Z
??
n
;
trans
itivity
(
Y
n
).
+
by
intros
X
Y
Z
??
n
;
trans
(
Y
n
).
-
intros
k
X
Y
HXY
n
;
apply
dist_S
.
by
rewrite
-(
g_tower
X
)
(
HXY
(
S
n
))
g_tower
.
-
intros
n
c
k
;
rewrite
/=
(
conv_compl
n
(
tower_chain
c
k
)).
...
...
@@ -209,7 +209,7 @@ Proof.
-
move
=>
X
/=.
rewrite
equiv_dist
;
intros
n
k
;
unfold
unfold
,
fold
;
simpl
.
rewrite
-
g_tower
-(
gg_tower
_
n
)
;
apply
(
_
:
Proper
(
_
==>
_
)
(
g
_
)).
trans
itivity
(
map
(
ff
n
,
gg
n
)
(
X
(
S
(
n
+
k
)))).
trans
(
map
(
ff
n
,
gg
n
)
(
X
(
S
(
n
+
k
)))).
{
rewrite
/
unfold
(
conv_compl
n
(
unfold_chain
X
)).
rewrite
-(
chain_cauchy
(
unfold_chain
X
)
n
(
S
(
n
+
k
)))
/=
;
last
lia
.
rewrite
-(
dist_le
_
_
_
_
(
f_tower
(
n
+
k
)
_
))
;
last
lia
.
...
...
@@ -234,6 +234,6 @@ Proof.
apply
(
contractive_ne
map
)
;
split
=>
Y
/=.
+
apply
dist_le
with
n
;
last
omega
.
rewrite
f_tower
.
apply
dist_S
.
by
rewrite
embed_tower
.
+
etrans
itivity
;
[
apply
embed_ne
,
equiv_dist
,
g_tower
|
apply
embed_tower
].
+
etrans
;
[
apply
embed_ne
,
equiv_dist
,
g_tower
|
apply
embed_tower
].
Qed
.
End
solver
.
End
solver
.
algebra/dra.v
View file @
08b69a69
...
...
@@ -60,7 +60,7 @@ Proof.
-
by
intros
[
x
px
?]
;
simpl
.
-
intros
[
x
px
?]
[
y
py
?]
;
naive_solver
.
-
intros
[
x
px
?]
[
y
py
?]
[
z
pz
?]
[?
Hxy
]
[?
Hyz
]
;
simpl
in
*.
split
;
[|
intros
;
trans
itivity
y
]
;
tauto
.
split
;
[|
intros
;
trans
y
]
;
tauto
.
Qed
.
Instance
dra_valid_proper'
:
Proper
((
≡
)
==>
iff
)
(
valid
:
A
→
Prop
).
Proof
.
by
split
;
apply
dra_valid_proper
.
Qed
.
...
...
algebra/excl.v
View file @
08b69a69
...
...
@@ -56,7 +56,7 @@ Proof.
-
intros
n
;
split
.
+
by
intros
[
x
|
|]
;
constructor
.
+
by
destruct
1
;
constructor
.
+
destruct
1
;
inversion_clear
1
;
constructor
;
etrans
itivity
;
eauto
.
+
destruct
1
;
inversion_clear
1
;
constructor
;
etrans
;
eauto
.
-
by
inversion_clear
1
;
constructor
;
apply
dist_S
.
-
intros
n
c
;
unfold
compl
,
excl_compl
.
destruct
(
Some_dec
(
maybe
Excl
(
c
1
)))
as
[[
x
Hx
]|].
...
...
algebra/fin_maps.v
View file @
08b69a69
...
...
@@ -22,7 +22,7 @@ Proof.
-
intros
n
;
split
.
+
by
intros
m
k
.
+
by
intros
m1
m2
?
k
.
+
by
intros
m1
m2
m3
??
k
;
trans
itivity
(
m2
!!
k
).
+
by
intros
m1
m2
m3
??
k
;
trans
(
m2
!!
k
).
-
by
intros
n
m1
m2
?
k
;
apply
dist_S
.
-
intros
n
c
k
;
rewrite
/
compl
/
map_compl
lookup_imap
.
feed
inversion
(
λ
H
,
chain_cauchy
c
0
(
S
n
)
H
k
)
;
simpl
;
auto
with
lia
.
...
...
algebra/iprod.v
View file @
08b69a69
...
...
@@ -35,7 +35,7 @@ Section iprod_cofe.
-
intros
n
;
split
.
+
by
intros
f
x
.
+
by
intros
f
g
?
x
.
+
by
intros
f
g
h
??
x
;
trans
itivity
(
g
x
).
+
by
intros
f
g
h
??
x
;
trans
(
g
x
).
-
intros
n
f
g
Hfg
x
;
apply
dist_S
,
Hfg
.
-
intros
n
c
x
.
rewrite
/
compl
/
iprod_compl
(
conv_compl
n
(
iprod_chain
c
x
)).
...
...
algebra/option.v
View file @
08b69a69
...
...
@@ -29,7 +29,7 @@ Proof.
-
intros
n
;
split
.
+
by
intros
[
x
|]
;
constructor
.
+
by
destruct
1
;
constructor
.
+
destruct
1
;
inversion_clear
1
;
constructor
;
etrans
itivity
;
eauto
.
+
destruct
1
;
inversion_clear
1
;
constructor
;
etrans
;
eauto
.
-
by
inversion_clear
1
;
constructor
;
apply
dist_S
.
-
intros
n
c
;
unfold
compl
,
option_compl
.
destruct
(
Some_dec
(
c
1
))
as
[[
x
Hx
]|].
...
...
algebra/sts.v
View file @
08b69a69
...
...
@@ -224,7 +224,7 @@ Proof.
split
.
-
by
intros
[]
;
constructor
.
-
by
destruct
1
;
constructor
.
-
destruct
1
;
inversion_clear
1
;
constructor
;
etrans
itivity
;
eauto
.
-
destruct
1
;
inversion_clear
1
;
constructor
;
etrans
;
eauto
.
Qed
.
Global
Instance
sts_dra
:
DRA
(
car
sts
).
Proof
.
...
...
@@ -366,7 +366,7 @@ Lemma sts_op_frag S1 S2 T1 T2 :
Proof
.
intros
HT
HS1
HS2
.
rewrite
/
sts_frag
.
(* FIXME why does rewrite not work?? *)
etrans
itivity
;
last
eapply
to_validity_op
;
try
done
;
[].
etrans
;
last
eapply
to_validity_op
;
try
done
;
[].
intros
Hval
.
constructor
;
last
set_solver
.
eapply
closed_ne
,
Hval
.
Qed
.
...
...
algebra/upred.v
View file @
08b69a69
From
algebra
Require
Export
cmra
.
Local
Hint
Extern
1
(
_
≼
_
)
=>
etrans
itivity
;
[
eassumption
|].
Local
Hint
Extern
1
(
_
≼
_
)
=>
etrans
itivity
;
[|
eassumption
].
Local
Hint
Extern
1
(
_
≼
_
)
=>
etrans
;
[
eassumption
|].
Local
Hint
Extern
1
(
_
≼
_
)
=>
etrans
;
[|
eassumption
].
Local
Hint
Extern
10
(
_
≤
_
)
=>
omega
.
Record
uPred
(
M
:
cmraT
)
:
Type
:
=
IProp
{
...
...
@@ -40,7 +40,7 @@ Section cofe.
-
intros
n
;
split
.
+
by
intros
P
x
i
.
+
by
intros
P
Q
HPQ
x
i
??
;
symmetry
;
apply
HPQ
.
+
by
intros
P
Q
Q'
HP
HQ
i
x
??
;
trans
itivity
(
Q
i
x
)
;
[
apply
HP
|
apply
HQ
].
+
by
intros
P
Q
Q'
HP
HQ
i
x
??
;
trans
(
Q
i
x
)
;
[
apply
HP
|
apply
HQ
].
-
intros
n
P
Q
HPQ
i
x
??
;
apply
HPQ
;
auto
.
-
intros
n
c
i
x
??
;
symmetry
;
apply
(
chain_cauchy
c
i
(
S
n
))
;
auto
.
Qed
.
...
...
@@ -243,8 +243,8 @@ Global Instance entails_proper :
Proper
((
≡
)
==>
(
≡
)
==>
iff
)
((
⊑
)
:
relation
(
uPred
M
)).
Proof
.
move
=>
P1
P2
/
equiv_spec
[
HP1
HP2
]
Q1
Q2
/
equiv_spec
[
HQ1
HQ2
]
;
split
;
intros
.
-
by
trans
itivity
P1
;
[|
trans
itivity
Q1
].
-
by
trans
itivity
P2
;
[|
trans
itivity
Q2
].
-
by
trans
P1
;
[|
trans
Q1
].
-
by
trans
P2
;
[|
trans
Q2
].
Qed
.
(** Non-expansiveness and setoid morphisms *)
...
...
@@ -734,7 +734,7 @@ Proof. by rewrite /uPred_iff later_and !later_impl. Qed.
Lemma
l
ö
b_strong
P
Q
:
(
P
∧
▷
Q
)
⊑
Q
→
P
⊑
Q
.
Proof
.
intros
Hl
ö
b
.
apply
impl_entails
.
etrans
itivity
;
last
by
eapply
l
ö
b
.
etrans
;
last
by
eapply
l
ö
b
.
apply
impl_intro_l
,
impl_intro_l
.
rewrite
right_id
-{
2
}
Hl
ö
b
.
apply
and_intro
;
first
by
eauto
.
by
rewrite
{
1
}(
later_intro
P
)
later_impl
impl_elim_r
.
...
...
algebra/upred_big_op.v
View file @
08b69a69
...
...
@@ -61,14 +61,14 @@ Proof.
induction
1
as
[|
P
Ps
Qs
?
IH
|
P
Q
Ps
|]
;
simpl
;
auto
.
-
by
rewrite
IH
.
-
by
rewrite
!
assoc
(
comm
_
P
).
-
etrans
itivity
;
eauto
.
-
etrans
;
eauto
.
Qed
.
Global
Instance
big_sep_perm
:
Proper
((
≡
ₚ
)
==>
(
≡
))
(@
uPred_big_sep
M
).
Proof
.
induction
1
as
[|
P
Ps
Qs
?
IH
|
P
Q
Ps
|]
;
simpl
;
auto
.
-
by
rewrite
IH
.
-
by
rewrite
!
assoc
(
comm
_
P
).
-
etrans
itivity
;
eauto
.
-
etrans
;
eauto
.
Qed
.
Lemma
big_and_app
Ps
Qs
:
(
Π
∧
(
Ps
++
Qs
))%
I
≡
(
Π
∧
Ps
∧
Π
∧
Qs
)%
I
.
...
...
@@ -103,7 +103,7 @@ Section gmap.
m2
⊆
m1
→
(
∀
x
k
,
m2
!!
k
=
Some
x
→
Φ
k
x
⊑
Ψ
k
x
)
→
(
Π★
{
map
m1
}
Φ
)
⊑
(
Π★
{
map
m2
}
Ψ
).
Proof
.
intros
HX
H
Φ
.
trans
itivity
(
Π★
{
map
m2
}
Φ
)%
I
.
intros
HX
H
Φ
.
trans
(
Π★
{
map
m2
}
Φ
)%
I
.
-
by
apply
big_sep_contains
,
fmap_contains
,
map_to_list_contains
.
-
apply
big_sep_mono'
,
Forall2_fmap
,
Forall2_Forall
.
apply
Forall_forall
=>
-[
i
x
]
?
/=.
by
apply
H
Φ
,
elem_of_map_to_list
.
...
...
@@ -163,7 +163,7 @@ Section gset.
Lemma
big_sepS_mono
Φ
Ψ
X
Y
:
Y
⊆
X
→
(
∀
x
,
x
∈
Y
→
Φ
x
⊑
Ψ
x
)
→
(
Π★
{
set
X
}
Φ
)
⊑
(
Π★
{
set
Y
}
Ψ
).
Proof
.
intros
HX
H
Φ
.
trans
itivity
(
Π★
{
set
Y
}
Φ
)%
I
.
intros
HX
H
Φ
.
trans
(
Π★
{
set
Y
}
Φ
)%
I
.
-
by
apply
big_sep_contains
,
fmap_contains
,
elements_contains
.
-
apply
big_sep_mono'
,
Forall2_fmap
,
Forall2_Forall
.
apply
Forall_forall
=>
x
?
/=.
by
apply
H
Φ
,
elem_of_elements
.
...
...
barrier/barrier.v
View file @
08b69a69
...
...
@@ -157,7 +157,7 @@ Section proof.
{
by
eapply
(
saved_prop_alloc
_
P
).
}
rewrite
pvs_frame_l
.
apply
pvs_strip_pvs
.
rewrite
sep_exist_l
.
apply
exist_elim
=>
i
.
trans
itivity
(
pvs
⊤
⊤
(
heap_ctx
heapN
★
▷
(
barrier_inv
l
P
(
State
Low
{[
i
]}))
★
saved_prop_own
i
P
)).
trans
(
pvs
⊤
⊤
(
heap_ctx
heapN
★
▷
(
barrier_inv
l
P
(
State
Low
{[
i
]}))
★
saved_prop_own
i
P
)).
-
rewrite
-
pvs_intro
.
rewrite
[(
_
★
heap_ctx
_
)%
I
]
comm
-!
assoc
.
apply
sep_mono_r
.
rewrite
{
1
}[
saved_prop_own
_
_
]
always_sep_dup
!
assoc
.
apply
sep_mono_l
.
rewrite
/
barrier_inv
/
waiting
-
later_intro
.
apply
sep_mono_r
.
...
...
@@ -215,7 +215,7 @@ Section proof.
apply
const_elim_sep_l
=>
Hs
.
destruct
p
;
last
done
.
rewrite
{
1
}/
barrier_inv
=>/={
Hs
}.
rewrite
later_sep
.
eapply
wp_store
;
eauto
with
I
ndisj
.
rewrite
-!
assoc
.
apply
sep_mono_r
.
etrans
itivity
;
last
eapply
later_mono
.
rewrite
-!
assoc
.
apply
sep_mono_r
.
etrans
;
last
eapply
later_mono
.
{
(* Is this really the best way to strip the later? *)
erewrite
later_sep
.
apply
sep_mono_r
.
apply
later_intro
.
}
apply
wand_intro_l
.
rewrite
-(
exist_intro
(
State
High
I
)).
...
...
@@ -256,7 +256,7 @@ Section proof.
apply
const_elim_sep_l
=>
Hs
.
rewrite
{
1
}/
barrier_inv
=>/=.
rewrite
later_sep
.
eapply
wp_load
;
eauto
with
I
ndisj
.
rewrite
-!
assoc
.
apply
sep_mono_r
.
etrans
itivity
;
last
eapply
later_mono
.
rewrite
-!
assoc
.
apply
sep_mono_r
.
etrans
;
last
eapply
later_mono
.
{
(* Is this really the best way to strip the later? *)
erewrite
later_sep
.
apply
sep_mono_r
.
rewrite
!
assoc
.
erewrite
later_sep
.
apply
sep_mono_l
,
later_intro
.
}
...
...
@@ -294,7 +294,7 @@ Section proof.
rewrite
[(
sts_own
_
_
_
★
_
)%
I
]
sep_elim_r
[(
sts_ctx
_
_
_
★
_
)%
I
]
sep_elim_r
.
rewrite
!
assoc
[(
_
★
saved_prop_own
i
Q
)%
I
]
comm
!
assoc
saved_prop_agree
.
wp_op
>
;
last
done
.
intros
_
.
etrans
itivity
;
last
eapply
later_mono
.
etrans
;
last
eapply
later_mono
.
{
(* Is this really the best way to strip the later? *)
erewrite
later_sep
.
apply
sep_mono
;
last
apply
later_intro
.
rewrite
->
later_sep
.
apply
sep_mono_l
.
rewrite
->
later_sep
.
done
.
}
...
...
heap_lang/heap.v
View file @
08b69a69
...
...
@@ -67,7 +67,7 @@ Section heap.
authG
heap_lang
Σ
heapRA
→
nclose
N
⊆
E
→
ownP
σ
⊑
(|={
E
}=>
∃
_
:
heapG
Σ
,
heap_ctx
N
∧
Π★
{
map
σ
}
heap_mapsto
).
Proof
.
intros
.
rewrite
-{
1
}(
from_to_heap
σ
).
etrans
itivity
.
intros
.
rewrite
-{
1
}(
from_to_heap
σ
).
etrans
.
{
rewrite
[
ownP
_
]
later_intro
.
apply
(
auth_alloc
(
ownP
∘
of_heap
)
E
N
(
to_heap
σ
))
;
last
done
.
apply
to_heap_valid
.
}
...
...
@@ -103,7 +103,7 @@ Section heap.
P
⊑
||
Alloc
e
@
E
{{
Φ
}}.
Proof
.
rewrite
/
heap_ctx
/
heap_inv
/
heap_mapsto
=>
??
Hctx
HP
.
trans
itivity
(|={
E
}=>
auth_own
heap_name
∅
★
P
)%
I
.
trans
(|={
E
}=>
auth_own
heap_name
∅
★
P
)%
I
.
{
by
rewrite
-
pvs_frame_r
-(
auth_empty
_
E
)
left_id
.
}
apply
wp_strip_pvs
,
(
auth_fsa
heap_inv
(
wp_fsa
(
Alloc
e
)))
with
N
heap_name
∅
;
simpl
;
eauto
with
I
.
...
...
heap_lang/wp_tactics.v
View file @
08b69a69
...
...
@@ -13,19 +13,19 @@ Ltac wp_strip_later :=
|
|-
_
⊑
▷
_
=>
apply
later_intro
|
|-
_
⊑
_
=>
reflexivity
end
in
revert_intros
ltac
:
(
etrans
itivity
;
[|
go
]).
in
revert_intros
ltac
:
(
etrans
;
[|
go
]).
Ltac
wp_bind
K
:
=
lazymatch
eval
hnf
in
K
with
|
[]
=>
idtac
|
_
=>
etrans
itivity
;
[|
solve
[
apply
(
wp_bind
K
)
]]
;
simpl
|
_
=>
etrans
;
[|
solve
[
apply
(
wp_bind
K
)
]]
;
simpl
end
.
Ltac
wp_finish
:
=
let
rec
go
:
=
match
goal
with
|
|-
_
⊑
▷
_
=>
etrans
itivity
;
[|
apply
later_mono
;
go
;
reflexivity
]
|
|-
_
⊑
▷
_
=>
etrans
;
[|
apply
later_mono
;
go
;
reflexivity
]
|
|-
_
⊑
wp
_
_
_
=>
etrans
itivity
;
[|
eapply
wp_value_pvs
;
reflexivity
]
;
etrans
;
[|
eapply
wp_value_pvs
;
reflexivity
]
;
(* sometimes, we will have to do a final view shift, so only apply
wp_value if we obtain a consecutive wp *)
try
(
eapply
pvs_intro
;
...
...
@@ -38,7 +38,7 @@ Tactic Notation "wp_rec" ">" :=
|
|-
_
⊑
wp
?E
?e
?Q
=>
reshape_expr
e
ltac
:
(
fun
K
e'
=>
match
eval
cbv
in
e'
with
|
App
(
Rec
_
_
_
)
_
=>
wp_bind
K
;
etrans
itivity
;
[|
eapply
wp_rec
;
reflexivity
]
;
wp_finish
wp_bind
K
;
etrans
;
[|
eapply
wp_rec
;
reflexivity
]
;
wp_finish
end
)
end
.
Tactic
Notation
"wp_rec"
:
=
wp_rec
>
;
wp_strip_later
.
...
...
@@ -48,7 +48,7 @@ Tactic Notation "wp_lam" ">" :=
|
|-
_
⊑
wp
?E
?e
?Q
=>
reshape_expr
e
ltac
:
(
fun
K
e'
=>
match
eval
cbv
in
e'
with
|
App
(
Rec
""
_
_
)
_
=>
wp_bind
K
;
etrans
itivity
;
[|
eapply
wp_lam
;
reflexivity
]
;
wp_finish
wp_bind
K
;
etrans
;
[|
eapply
wp_lam
;
reflexivity
]
;
wp_finish
end
)
end
.
Tactic
Notation
"wp_lam"
:
=
wp_lam
>
;
wp_strip_later
.
...
...
@@ -66,9 +66,9 @@ Tactic Notation "wp_op" ">" :=
|
BinOp
LeOp
_
_
=>
wp_bind
K
;
apply
wp_le
;
wp_finish
|
BinOp
EqOp
_
_
=>
wp_bind
K
;
apply
wp_eq
;
wp_finish
|
BinOp
_
_
_
=>
wp_bind
K
;
etrans
itivity
;
[|
eapply
wp_bin_op
;
reflexivity
]
;
wp_finish
wp_bind
K
;
etrans
;
[|
eapply
wp_bin_op
;
reflexivity
]
;
wp_finish
|
UnOp
_
_
=>
wp_bind
K
;
etrans
itivity
;
[|
eapply
wp_un_op
;
reflexivity
]
;
wp_finish
wp_bind
K
;
etrans
;
[|
eapply
wp_un_op
;
reflexivity
]
;
wp_finish
end
)
end
.
Tactic
Notation
"wp_op"
:
=
wp_op
>
;
wp_strip_later
.
...
...
@@ -79,7 +79,7 @@ Tactic Notation "wp_if" ">" :=
match
eval
cbv
in
e'
with
|
If
_
_
_
=>
wp_bind
K
;
etrans
itivity
;
[|
apply
wp_if_true
||
apply
wp_if_false
]
;
wp_finish
etrans
;
[|
apply
wp_if_true
||
apply
wp_if_false
]
;
wp_finish
end
)
end
.
Tactic
Notation
"wp_if"
:
=
wp_if
>
;
wp_strip_later
.
...
...
@@ -97,5 +97,5 @@ Tactic Notation "wp" ">" tactic(tac) :=
Tactic
Notation
"wp"
tactic
(
tac
)
:
=
(
wp
>
tac
)
;
wp_strip_later
.
(* In case the precondition does not match *)
Tactic
Notation
"ewp"
tactic
(
tac
)
:
=
wp
(
etrans
itivity
;
[|
tac
]).
Tactic
Notation
"ewp"
">"
tactic
(
tac
)
:
=
wp
>
(
etrans
itivity
;
[|
tac
]).
Tactic
Notation
"ewp"
tactic
(
tac
)
:
=
wp
(
etrans
;
[|
tac
]).
Tactic
Notation
"ewp"
">"
tactic
(
tac
)
:
=
wp
>
(
etrans
;
[|
tac
]).
prelude/fin_map_dom.v
View file @
08b69a69
...
...
@@ -60,7 +60,7 @@ Lemma dom_insert_subseteq {A} (m : M A) i x : dom D m ⊆ dom D (<[i:=x]>m).
Proof
.
rewrite
(
dom_insert
_
).
set_solver
.
Qed
.
Lemma
dom_insert_subseteq_compat_l
{
A
}
(
m
:
M
A
)
i
x
X
:
X
⊆
dom
D
m
→
X
⊆
dom
D
(<[
i
:
=
x
]>
m
).
Proof
.
intros
.
trans
itivity
(
dom
D
m
)
;
eauto
using
dom_insert_subseteq
.
Qed
.
Proof
.
intros
.
trans
(
dom
D
m
)
;
eauto
using
dom_insert_subseteq
.
Qed
.
Lemma
dom_singleton
{
A
}
(
i
:
K
)
(
x
:
A
)
:
dom
D
{[
i
:
=
x
]}
≡
{[
i
]}.
Proof
.
rewrite
<-
insert_empty
,
dom_insert
,
dom_empty
;
set_solver
.
Qed
.
Lemma
dom_delete
{
A
}
(
m
:
M
A
)
i
:
dom
D
(
delete
i
m
)
≡
dom
D
m
∖
{[
i
]}.
...
...
prelude/fin_maps.v
View file @
08b69a69
...
...
@@ -123,7 +123,7 @@ Section setoid.
split
.
-
by
intros
m
i
.
-
by
intros
m1
m2
?
i
.
-
by
intros
m1
m2
m3
??
i
;
trans
itivity
(
m2
!!
i
).
-
by
intros
m1
m2
m3
??
i
;
trans
(
m2
!!
i
).
Qed
.
Global
Instance
lookup_proper
(
i
:
K
)
:
Proper
((
≡
)
==>
(
≡
))
(
lookup
(
M
:
=
M
A
)
i
).
...
...
@@ -199,7 +199,7 @@ Proof.
split
;
[
intros
m
i
;
by
destruct
(
m
!!
i
)
;
simpl
|].
intros
m1
m2
m3
Hm12
Hm23
i
;
specialize
(
Hm12
i
)
;
specialize
(
Hm23
i
).
destruct
(
m1
!!
i
),
(
m2
!!
i
),
(
m3
!!
i
)
;
simplify_eq
/=
;
done
||
etrans
itivity
;
eauto
.
done
||
etrans
;
eauto
.
Qed
.
Global
Instance
:
PartialOrder
((
⊆
)
:
relation
(
M
A
)).
Proof
.
...
...
@@ -1182,10 +1182,10 @@ Proof.
intros
.
rewrite
map_union_comm
by
done
.
by
apply
map_union_subseteq_l
.
Qed
.
Lemma
map_union_subseteq_l_alt
{
A
}
(
m1
m2
m3
:
M
A
)
:
m1
⊆
m2
→
m1
⊆
m2
∪
m3
.
Proof
.
intros
.
trans
itivity
m2
;
auto
using
map_union_subseteq_l
.
Qed
.
Proof
.
intros
.
trans
m2
;
auto
using
map_union_subseteq_l
.
Qed
.
Lemma
map_union_subseteq_r_alt
{
A
}
(
m1
m2
m3
:
M
A
)
:
m2
⊥
ₘ
m3
→
m1
⊆
m3
→
m1
⊆
m2
∪
m3
.
Proof
.
intros
.
trans
itivity
m3
;
auto
using
map_union_subseteq_r
.
Qed
.
Proof
.
intros
.
trans
m3
;
auto
using
map_union_subseteq_r
.
Qed
.
Lemma
map_union_preserving_l
{
A
}
(
m1
m2
m3
:
M
A
)
:
m1
⊆
m2
→
m3
∪
m1
⊆
m3
∪
m2
.
Proof
.
rewrite
!
map_subseteq_spec
.
intros
???.
...
...
prelude/lexico.v
View file @
08b69a69
...
...
@@ -42,7 +42,7 @@ Lemma prod_lexico_transitive `{Lexico A, Lexico B, !Transitive (@lexico A _)}
Proof
.
intros
Hx12
Hx23
?
;
revert
Hx12
Hx23
.
unfold
lexico
,
prod_lexico
.
intros
[|[??]]
[?|[??]]
;
simplify_eq
/=
;
auto
.
by
left
;
trans
itivity
x2
.
by
left
;
trans
x2
.
Qed
.
Instance
prod_lexico_po
`
{
Lexico
A
,
Lexico
B
,
!
StrictOrder
(@
lexico
A
_
)}
...
...
@@ -52,7 +52,7 @@ Proof.
-
intros
[
x
y
].
apply
prod_lexico_irreflexive
.
by
apply
(
irreflexivity
lexico
y
).
-
intros
[??]
[??]
[??]
??.
eapply
prod_lexico_transitive
;
eauto
.
apply
trans
itivity
.
eapply
prod_lexico_transitive
;
eauto
.
apply
trans
.
Qed
.
Instance
prod_lexico_trichotomyT
`
{
Lexico
A
,
tA
:
!
TrichotomyT
(@
lexico
A
_
)}
`
{
Lexico
B
,
tB
:
!
TrichotomyT
(@
lexico
B
_
)}
:
TrichotomyT
(@
lexico
(
A
*
B
)
_
).
...
...
@@ -143,7 +143,7 @@ Instance sig_lexico_po `{Lexico A, !StrictOrder (@lexico A _)}
Proof
.
unfold
lexico
,
sig_lexico
.
split
.
-
intros
[
x
?]
?.
by
apply
(
irreflexivity
lexico
x
).
-
intros
[
x1
?]
[
x2
?]
[
x3
?]
??.
by
trans
itivity
x2
.
-
intros
[
x1
?]
[
x2
?]
[
x3
?]
??.
by
trans
x2
.
Qed
.
Instance
sig_lexico_trichotomy
`
{
Lexico
A
,
tA
:
!
TrichotomyT
(@
lexico
A
_
)}
(
P
:
A
→
Prop
)
`
{
∀
x
,
ProofIrrel
(
P
x
)}
:
TrichotomyT
(@
lexico
(
sig
P
)
_
).
...
...
Prev
1
2
Next
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