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
George Pirlea
Iris
Commits
ec27d524
Commit
ec27d524
authored
Mar 03, 2019
by
Ralf Jung
Browse files
Options
Browse Files
Download
Email Patches
Plain Diff
add test with a closed proof
parent
d04c9f7e
Changes
3
Hide whitespace changes
Inline
Side-by-side
Showing
3 changed files
with
36 additions
and
1 deletion
+36
-1
tests/one_shot.ref
tests/one_shot.ref
+1
-0
tests/one_shot.v
tests/one_shot.v
+34
-1
theories/algebra/ofe.v
theories/algebra/ofe.v
+1
-0
No files found.
tests/one_shot.ref
View file @
ec27d524
...
@@ -40,3 +40,4 @@
...
@@ -40,3 +40,4 @@
| InjR "m" => assert: #m = "m"
| InjR "m" => assert: #m = "m"
end {{ _, True }}
end {{ _, True }}
Closed under the global context
tests/one_shot.v
View file @
ec27d524
From
iris
.
program_logic
Require
Export
weakestpre
hoare
.
From
iris
.
program_logic
Require
Export
weakestpre
hoare
.
From
iris
.
heap_lang
Require
Export
lang
.
From
iris
.
heap_lang
Require
Export
lang
.
From
iris
.
algebra
Require
Import
excl
agree
csum
.
From
iris
.
algebra
Require
Import
excl
agree
csum
.
From
iris
.
heap_lang
Require
Import
assert
proofmode
notation
.
From
iris
.
heap_lang
Require
Import
assert
proofmode
notation
adequacy
.
From
iris
.
proofmode
Require
Import
tactics
.
From
iris
.
proofmode
Require
Import
tactics
.
From
iris
.
heap_lang
.
lib
Require
Import
par
.
Set
Default
Proof
Using
"Type"
.
Set
Default
Proof
Using
"Type"
.
Definition
one_shot_example
:
val
:
=
λ
:
<>,
Definition
one_shot_example
:
val
:
=
λ
:
<>,
...
@@ -95,3 +96,35 @@ Proof.
...
@@ -95,3 +96,35 @@ Proof.
-
iIntros
"!# _"
.
wp_apply
(
wp_wand
with
"Hf2"
).
by
iIntros
(
v
)
"#? !# _"
.
-
iIntros
"!# _"
.
wp_apply
(
wp_wand
with
"Hf2"
).
by
iIntros
(
v
)
"#? !# _"
.
Qed
.
Qed
.
End
proof
.
End
proof
.
(* Have a client with a closed proof. *)
Definition
client
:
expr
:
=
let
:
"ff"
:
=
one_shot_example
#()
in
(
Fst
"ff"
#
5
|||
let
:
"check"
:
=
Snd
"ff"
#()
in
"check"
#()).
Section
client
.
Local
Set
Default
Proof
Using
"Type*"
.
Context
`
{!
heapG
Σ
,
!
one_shotG
Σ
,
!
spawnG
Σ
}.
Lemma
client_safe
:
WP
client
{{
_
,
True
}}%
I
.
Proof
.
rewrite
/
client
.
wp_apply
wp_one_shot
.
iIntros
(
f1
f2
)
"[#Hf1 #Hf2]"
.
wp_let
.
wp_apply
wp_par
.
-
wp_apply
"Hf1"
.
-
wp_proj
.
wp_bind
(
f2
_
)%
E
.
iApply
wp_wand
;
first
by
iExact
"Hf2"
.
iIntros
(
check
)
"Hcheck"
.
wp_pures
.
iApply
"Hcheck"
.
-
auto
.
Qed
.
End
client
.
(** Put together all library functors. *)
Definition
client
Σ
:
gFunctors
:
=
#[
heap
Σ
;
one_shot
Σ
;
spawn
Σ
].
(** This lemma implicitly shows that these functors are enough to meet
all library assumptions. *)
Lemma
client_adequate
σ
:
adequate
NotStuck
client
σ
(
λ
_
_
,
True
).
Proof
.
apply
(
heap_adequacy
client
Σ
)=>
?.
apply
client_safe
.
Qed
.
(* Since we check the output of the test files, this means
our test suite will fail if we ever accidentally add an axiom
to anything used by this proof. *)
Print
Assumptions
client_adequate
.
theories/algebra/ofe.v
View file @
ec27d524
...
@@ -683,6 +683,7 @@ Bind Scope cFunctor_scope with cFunctor.
...
@@ -683,6 +683,7 @@ Bind Scope cFunctor_scope with cFunctor.
Class
cFunctorContractive
(
F
:
cFunctor
)
:
=
Class
cFunctorContractive
(
F
:
cFunctor
)
:
=
cFunctor_contractive
A1
A2
B1
B2
:
>
Contractive
(@
cFunctor_map
F
A1
A2
B1
B2
).
cFunctor_contractive
A1
A2
B1
B2
:
>
Contractive
(@
cFunctor_map
F
A1
A2
B1
B2
).
Hint
Mode
cFunctorContractive
!
:
typeclass_instances
.
Definition
cFunctor_diag
(
F
:
cFunctor
)
(
A
:
ofeT
)
:
ofeT
:
=
cFunctor_car
F
A
A
.
Definition
cFunctor_diag
(
F
:
cFunctor
)
(
A
:
ofeT
)
:
ofeT
:
=
cFunctor_car
F
A
A
.
Coercion
cFunctor_diag
:
cFunctor
>->
Funclass
.
Coercion
cFunctor_diag
:
cFunctor
>->
Funclass
.
...
...
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