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-coq
Project overview
Project overview
Details
Activity
Releases
Repository
Repository
Files
Commits
Branches
Tags
Contributors
Graph
Compare
Issues
0
Issues
0
List
Boards
Labels
Service Desk
Milestones
Merge Requests
0
Merge Requests
0
CI / CD
CI / CD
Pipelines
Jobs
Schedules
Operations
Operations
Incidents
Environments
Analytics
Analytics
CI / CD
Repository
Value Stream
Wiki
Wiki
Snippets
Snippets
Members
Members
Collapse sidebar
Close sidebar
Activity
Graph
Create a new issue
Jobs
Commits
Issue Boards
Open sidebar
Marianna Rapoport
iris-coq
Commits
ab7651d1
Commit
ab7651d1
authored
Nov 04, 2017
by
Robbert
Browse files
Options
Browse Files
Download
Plain Diff
Merge branch 'robbert/atomic' into 'master'
Make atomic a type class. See merge request FP/iris-coq!78
parents
97aa894b
a78f1c3b
Changes
6
Hide whitespace changes
Inline
Side-by-side
Showing
6 changed files
with
17 additions
and
22 deletions
+17
-22
theories/heap_lang/tactics.v
theories/heap_lang/tactics.v
+6
-8
theories/program_logic/ectx_language.v
theories/program_logic/ectx_language.v
+1
-1
theories/program_logic/hoare.v
theories/program_logic/hoare.v
+1
-1
theories/program_logic/language.v
theories/program_logic/language.v
+6
-7
theories/program_logic/weakestpre.v
theories/program_logic/weakestpre.v
+2
-4
theories/tests/ipm_paper.v
theories/tests/ipm_paper.v
+1
-1
No files found.
theories/heap_lang/tactics.v
View file @
ab7651d1
...
...
@@ -175,7 +175,7 @@ Proof.
f_equal
;
eauto
using
subst_is_closed_nil
,
is_closed_to_val
,
eq_sym
.
Qed
.
Definition
atomic
(
e
:
expr
)
:
=
Definition
is_
atomic
(
e
:
expr
)
:
=
match
e
with
|
Alloc
e
=>
bool_decide
(
is_Some
(
to_val
e
))
|
Load
e
=>
bool_decide
(
is_Some
(
to_val
e
))
...
...
@@ -187,7 +187,7 @@ Definition atomic (e : expr) :=
|
App
(
Rec
_
_
(
Lit
_
))
(
Lit
_
)
=>
true
|
_
=>
false
end
.
Lemma
atomic_correct
e
:
atomic
e
→
language
.
a
tomic
(
to_expr
e
).
Lemma
is_atomic_correct
e
:
is_atomic
e
→
A
tomic
(
to_expr
e
).
Proof
.
intros
He
.
apply
ectx_language_atomic
.
-
intros
σ
e'
σ
'
ef
Hstep
;
simpl
in
*.
...
...
@@ -228,13 +228,11 @@ Hint Extern 10 (AsVal _) => solve_to_val : typeclass_instances.
Ltac
solve_atomic
:
=
match
goal
with
|
|-
language
.
a
tomic
?e
=>
let
e'
:
=
W
.
of_expr
e
in
change
(
language
.
a
tomic
(
W
.
to_expr
e'
))
;
apply
W
.
atomic_correct
;
vm_compute
;
exact
I
|
|-
A
tomic
?e
=>
let
e'
:
=
W
.
of_expr
e
in
change
(
A
tomic
(
W
.
to_expr
e'
))
;
apply
W
.
is_
atomic_correct
;
vm_compute
;
exact
I
end
.
Hint
Extern
10
(
language
.
atomic
_
)
=>
solve_atomic
.
(* For the side-condition of elim_upd_fupd_wp_atomic *)
Hint
Extern
10
(
language
.
atomic
_
)
=>
solve_atomic
:
typeclass_instances
.
Hint
Extern
10
(
Atomic
_
)
=>
solve_atomic
:
typeclass_instances
.
(** Substitution *)
Ltac
simpl_subst
:
=
...
...
theories/program_logic/ectx_language.v
View file @
ab7651d1
...
...
@@ -118,7 +118,7 @@ Section ectx_language.
Lemma
ectx_language_atomic
e
:
(
∀
σ
e'
σ
'
efs
,
head_step
e
σ
e'
σ
'
efs
→
irreducible
e'
σ
'
)
→
sub_values
e
→
a
tomic
e
.
A
tomic
e
.
Proof
.
intros
Hatomic_step
Hatomic_fill
σ
e'
σ
'
efs
[
K
e1'
e2'
->
->
Hstep
].
assert
(
K
=
empty_ectx
)
as
->
by
eauto
10
using
val_stuck
.
...
...
theories/program_logic/hoare.v
View file @
ab7651d1
...
...
@@ -58,7 +58,7 @@ Proof.
Qed
.
Lemma
ht_atomic
E1
E2
P
P'
Φ
Φ
'
e
:
a
tomic
e
→
A
tomic
e
→
(
P
={
E1
,
E2
}=>
P'
)
∧
{{
P'
}}
e
@
E2
{{
Φ
'
}}
∧
(
∀
v
,
Φ
'
v
={
E2
,
E1
}=>
Φ
v
)
⊢
{{
P
}}
e
@
E1
{{
Φ
}}.
Proof
.
...
...
theories/program_logic/language.v
View file @
ab7651d1
...
...
@@ -58,15 +58,15 @@ Section language.
stronger sense defined below, and we have to be able to open invariants
around that expression. See `CasStuckS` in
[lambdaRust](https://gitlab.mpi-sws.org/FP/LambdaRust-coq/blob/master/theories/lang/lang.v). *)
Definition
a
tomic
(
e
:
expr
Λ
)
:
Prop
:
=
∀
σ
e'
σ
'
efs
,
prim_step
e
σ
e'
σ
'
efs
→
irreducible
e'
σ
'
.
Class
A
tomic
(
e
:
expr
Λ
)
:
Prop
:
=
atomic
σ
e'
σ
'
efs
:
prim_step
e
σ
e'
σ
'
efs
→
irreducible
e'
σ
'
.
(* To open invariants with a WP that does not ensure safety, we need a
stronger form of atomicity. With the above definition, in case `e` reduces
to a stuck non-value, there is no proof that the invariants have been
established again. *)
Definition
strongly_a
tomic
(
e
:
expr
Λ
)
:
Prop
:
=
∀
σ
e'
σ
'
efs
,
prim_step
e
σ
e'
σ
'
efs
→
is_Some
(
to_val
e'
).
Class
StronglyA
tomic
(
e
:
expr
Λ
)
:
Prop
:
=
strongly_atomic
σ
e'
σ
'
efs
:
prim_step
e
σ
e'
σ
'
efs
→
is_Some
(
to_val
e'
).
Inductive
step
(
ρ
1
ρ
2
:
cfg
Λ
)
:
Prop
:
=
|
step_atomic
e1
σ
1 e2
σ
2
efs
t1
t2
:
...
...
@@ -87,9 +87,8 @@ Section language.
Global
Instance
of_val_inj
:
Inj
(=)
(=)
(@
of_val
Λ
).
Proof
.
by
intros
v
v'
Hv
;
apply
(
inj
Some
)
;
rewrite
-!
to_of_val
Hv
.
Qed
.
Lemma
strongly_atomic_atomic
e
:
strongly_atomic
e
→
atomic
e
.
Proof
.
unfold
strongly_atomic
,
atomic
.
eauto
using
val_irreducible
.
Qed
.
Lemma
strongly_atomic_atomic
e
:
StronglyAtomic
e
→
Atomic
e
.
Proof
.
unfold
StronglyAtomic
,
Atomic
.
eauto
using
val_irreducible
.
Qed
.
Lemma
reducible_fill
`
{
LanguageCtx
Λ
K
}
e
σ
:
to_val
e
=
None
→
reducible
(
K
e
)
σ
→
reducible
e
σ
.
...
...
theories/program_logic/weakestpre.v
View file @
ab7651d1
...
...
@@ -147,7 +147,7 @@ Lemma wp_fupd E e Φ : WP e @ E {{ v, |={E}=> Φ v }} ⊢ WP e @ E {{ Φ }}.
Proof
.
iIntros
"H"
.
iApply
(
wp_strong_mono
E
)
;
try
iFrame
;
auto
.
Qed
.
Lemma
wp_atomic
E1
E2
e
Φ
:
a
tomic
e
→
A
tomic
e
→
(|={
E1
,
E2
}=>
WP
e
@
E2
{{
v
,
|={
E2
,
E1
}=>
Φ
v
}})
⊢
WP
e
@
E1
{{
Φ
}}.
Proof
.
iIntros
(
Hatomic
)
"H"
.
rewrite
!
wp_unfold
/
wp_pre
.
...
...
@@ -285,10 +285,8 @@ Section proofmode_classes.
(* lower precedence, if possible, it should persistently pick elim_upd_fupd_wp *)
Global
Instance
elim_modal_fupd_wp_atomic
E1
E2
e
P
Φ
:
a
tomic
e
→
A
tomic
e
→
ElimModal
(|={
E1
,
E2
}=>
P
)
P
(
WP
e
@
E1
{{
Φ
}})
(
WP
e
@
E2
{{
v
,
|={
E2
,
E1
}=>
Φ
v
}})%
I
|
100
.
Proof
.
intros
.
by
rewrite
/
ElimModal
fupd_frame_r
wand_elim_r
wp_atomic
.
Qed
.
End
proofmode_classes
.
Hint
Extern
0
(
atomic
_
)
=>
assumption
:
typeclass_instances
.
theories/tests/ipm_paper.v
View file @
ab7651d1
...
...
@@ -101,7 +101,7 @@ update modalities (which we did not cover in the paper). Normally we use these
mask changing update modalities directly in our proofs, but in this file we use
the first prove the rule as a lemma, and then use that. *)
Lemma
wp_inv_open
`
{
irisG
Λ
Σ
}
N
E
P
e
Φ
:
nclose
N
⊆
E
→
a
tomic
e
→
nclose
N
⊆
E
→
A
tomic
e
→
inv
N
P
∗
(
▷
P
-
∗
WP
e
@
E
∖
↑
N
{{
v
,
▷
P
∗
Φ
v
}})
⊢
WP
e
@
E
{{
Φ
}}.
Proof
.
iIntros
(??)
"[#Hinv Hwp]"
.
...
...
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