Skip to content
GitLab
Menu
Projects
Groups
Snippets
Help
Help
Support
Community forum
Keyboard shortcuts
?
Submit feedback
Contribute to GitLab
Sign in / Register
Toggle navigation
Menu
Open sidebar
Marianna Rapoport
iris-coq
Commits
eed2dc1d
Commit
eed2dc1d
authored
Feb 22, 2016
by
Ralf Jung
Browse files
hide which exact functors the barrier needs
parent
1c4b9888
Changes
3
Hide whitespace changes
Inline
Side-by-side
barrier/barrier.v
View file @
eed2dc1d
...
...
@@ -126,6 +126,9 @@ End barrier_proto.
the module into our namespaces. But Coq doesn't seem to support that...?? *)
Import
barrier_proto
.
(* The functors we need. *)
Definition
barrierFs
:
=
stsF
sts
`
::
`
agreeF
`
::
`
pnil
.
(** Now we come to the Iris part of the proof. *)
Section
proof
.
Context
{
Σ
:
iFunctorG
}
(
N
:
namespace
).
...
...
barrier/client.v
View file @
eed2dc1d
...
...
@@ -28,8 +28,7 @@ Section client.
End
client
.
Section
ClosedProofs
.
Definition
Σ
:
iFunctorG
:
=
agreeF
.
::
constF
(
stsRA
barrier_proto
.
sts
)
.
::
heapF
.
::
(
λ
_
,
constF
unitRA
)
:
iFunctorG
.
Definition
Σ
:
iFunctorG
:
=
heapF
.
::
barrierFs
.++
endF
.
Notation
iProp
:
=
(
iPropG
heap_lang
Σ
).
Lemma
client_safe_closed
σ
:
{{
ownP
σ
:
iProp
}}
client
{{
λ
v
,
True
}}.
...
...
prelude/functions.v
View file @
eed2dc1d
...
...
@@ -31,20 +31,41 @@ Section functions.
End
functions
.
(** "Cons-ing" of functions from nat to T *)
(* Coq's standard lists are not universe polymorphic. Hence we have to re-define them. Ouch.
TODO: If we decide to end up going with this, we should move this elsewhere. *)
Polymorphic
Inductive
plist
{
A
:
Type
}
:
Type
:
=
|
pnil
:
plist
|
pcons
:
A
→
plist
→
plist
.
Arguments
plist
:
clear
implicits
.
Polymorphic
Fixpoint
papp
{
A
:
Type
}
(
l1
l2
:
plist
A
)
:
plist
A
:
=
match
l1
with
|
pnil
=>
l2
|
pcons
a
l
=>
pcons
a
(
papp
l
l2
)
end
.
(* TODO: Notation is totally up for debate. *)
Infix
"`::`"
:
=
pcons
(
at
level
60
,
right
associativity
)
:
C_scope
.
Infix
"`++`"
:
=
papp
(
at
level
60
,
right
associativity
)
:
C_scope
.
Polymorphic
Definition
fn_cons
{
T
:
Type
}
(
t
:
T
)
(
f
:
nat
→
T
)
:
nat
→
T
:
=
λ
n
,
match
n
with
|
O
=>
t
|
S
n
=>
f
n
end
.
Polymorphic
Definition
fn_mcons
{
T
:
Type
}
(
ts
:
list
T
)
(
f
:
nat
→
T
)
:
nat
→
T
:
=
fold_right
fn_cons
f
ts
.
Polymorphic
Fixpoint
fn_mcons
{
T
:
Type
}
(
ts
:
plist
T
)
(
f
:
nat
→
T
)
:
nat
→
T
:
=
match
ts
with
|
pnil
=>
f
|
pcons
t
ts
=>
fn_cons
t
(
fn_mcons
ts
f
)
end
.
(* TODO: Notation is totally up for debate. *)
Infix
".::"
:
=
fn_cons
(
at
level
60
,
right
associativity
)
:
C_scope
.
Infix
".++"
:
=
fn_mcons
(
at
level
60
,
right
associativity
)
:
C_scope
.
Polymorphic
Lemma
fn_mcons_app
{
T
:
Type
}
(
ts1
ts2
:
list
T
)
f
:
(
ts1
++
ts2
)
.++
f
=
ts1
.++
(
ts2
.++
f
).
Polymorphic
Lemma
fn_mcons_app
{
T
:
Type
}
(
ts1
ts2
:
p
list
T
)
f
:
(
ts1
`
++
`
ts2
)
.++
f
=
ts1
.++
(
ts2
.++
f
).
Proof
.
unfold
fn_mcons
.
rewrite
fold_right_app
.
d
one
.
induction
ts1
;
simpl
;
eauto
.
c
on
gruenc
e
.
Qed
.
Write
Preview
Supports
Markdown
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