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
Andrew Hirsch
Pirouette Coq Code
Commits
bbd1e08d
Commit
bbd1e08d
authored
Jun 08, 2021
by
Andrew Hirsch
Browse files
Completeness done.
parent
16ce9dad
Changes
7
Expand all
Hide whitespace changes
Inline
Side-by-side
Choreography.v
View file @
bbd1e08d
...
...
@@ -1412,7 +1412,7 @@ Module Choreography (Import E : Expression) (L : Locations).
ChorStep
RAppGlobal
[]
(
AppGlobal
(
RecGlobal
C1
)
C2
)
(
C1
[
c
|
AppGlobalSubst
C1
C2
])
|
CSyncStep
:
forall
(
l1
l2
:
Loc
)
(
d
:
LRChoice
)
(
C
:
Chor
)
(
B
:
list
Loc
),
~
In
l1
B
->
~
In
l2
B
->
~
In
l1
B
->
~
In
l2
B
->
l1
<>
l2
->
ChorStep
(
RSync
l1
d
l2
)
B
(
Sync
l1
d
l2
C
)
C
|
CSyncIStep
:
forall
(
l1
l2
:
Loc
)
(
d
:
LRChoice
)
(
C1
C2
:
Chor
)
(
B
:
list
Loc
)
(
R
:
Redex
),
ChorStep
R
(
l1
::
l2
::
B
)
C1
C2
->
...
...
@@ -1463,6 +1463,151 @@ Module Choreography (Import E : Expression) (L : Locations).
intros
q
H1
;
apply
H0
;
auto
.
Qed
.
Lemma
ChorStepSendVDistinguish
:
forall
p
v
q
B
C1
C2
,
ChorStep
(
RSendV
p
v
q
)
B
C1
C2
->
p
<>
q
.
Proof
using
.
intros
p
v
q
B
C1
C2
step
;
dependent
induction
step
;
repeat
match
goal
with
|
[
IH
:
forall
a
b
c
,
RSendV
?
p
?
v
?
q
=
RSendV
a
b
c
->
_
|-
_
]
=>
specialize
(
IH
p
v
q
eq_refl
)
end
;
auto
.
Qed
.
Lemma
ChorStepSyncDistinguish
:
forall
p
d
q
B
C1
C2
,
ChorStep
(
RSync
p
d
q
)
B
C1
C2
->
p
<>
q
.
Proof
using
.
intros
p
v
q
B
C1
C2
step
;
dependent
induction
step
;
repeat
match
goal
with
|
[
IH
:
forall
a
b
c
,
RSync
?
p
?
v
?
q
=
RSync
a
b
c
->
_
|-
_
]
=>
specialize
(
IH
p
v
q
eq_refl
)
end
;
auto
.
Qed
.
Definition
InvolvedWithRedex
(
R
:
Redex
)
(
l
:
L
.
t
)
:
Prop
:=
match
R
with
|
RDone
l
'
e1
e2
=>
l
=
l
'
|
RIfE
l
'
e1
e2
=>
l
=
l
'
|
RIfTT
l
'
=>
l
=
l
'
|
RIfFF
l
'
=>
l
=
l
'
|
RSendE
l1
e1
e2
l2
=>
l
=
l1
|
RSendV
l1
v
l2
=>
l
=
l1
\
/
l
=
l2
|
RSync
l1
d
l2
=>
l
=
l1
\
/
l
=
l2
|
RDefLocal
l
'
e
=>
l
=
l
'
|
RAppLocalE
l
'
e1
e2
=>
l
=
l
'
|
RAppLocal
l
'
e
=>
l
=
l
'
|
RAppGlobal
=>
False
end
.
Lemma
NoChorStepInList
:
forall
p
B
R
,
In
p
B
->
InvolvedWithRedex
R
p
->
forall
C1
C2
,
~
ChorStep
R
B
C1
C2
.
Proof
using
.
intros
p
B
R
H
H0
C1
C2
step
;
induction
step
;
cbn
in
H0
;
match
goal
with
|
[
i
:
In
?
p
?
B
,
n
:
~
In
?
p
'
?
B
,
e
:
?
p
=
?
p
'
|-
False
]
=>
apply
n
;
rewrite
<-
e
;
exact
i
|
[
i
:
In
?
p
?
B
,
n
:
~
In
?
p
'
?
B
,
e
:
?
p
'
=
?
p
|-
False
]
=>
apply
n
;
rewrite
e
;
exact
i
|
[
H
:
_
\
/
_
|-
_
]
=>
destruct
H
;
subst
|
[
H1
:
?
P
,
H2
:
~
?
P
|-
_
]
=>
destruct
(
H2
H1
)
|
_
=>
idtac
end
.
all:
try
(
apply
IHstep
;
auto
;
right
;
right
;
auto
;
fail
).
all:
try
(
apply
IHstep1
;
auto
;
right
;
auto
;
fail
).
all:
try
match
goal
with
|
[
H1
:
?
P
,
H2
:
~?
P
|-
_
]
=>
destruct
(
H2
H1
)
end
.
all:
inversion
H
.
Qed
.
Lemma
ThreadNamesExprSubst
:
forall
C
σ
,
ThreadNames
C
=
ThreadNames
(
C
[
ce
|
σ
]).
Proof
using
.
intro
C
;
ChorInduction
C
;
intros
σ
;
cbn
;
auto
.
all:
try
(
erewrite
IHC
;
eauto
;
fail
).
all:
rewrite
IHC1
with
(
σ
:=
σ
).
1
,
3
:
rewrite
IHC2
with
(
σ
:=
σ
);
reflexivity
.
rewrite
IHC2
with
(
σ
:=
ChorUpExprSubst
σ
l
);
reflexivity
.
Qed
.
Lemma
ThreadNamesRenaming
:
forall
C
ξ
,
ThreadNames
(
C
⟨
c
|
ξ⟩
)
=
ThreadNames
C
.
Proof
using
.
intro
C
;
ChorInduction
C
;
intros
ξ
;
cbn
;
auto
.
all:
try
(
erewrite
IHC
;
eauto
;
fail
).
all:
rewrite
IHC1
with
(
ξ
:=
ξ
);
rewrite
IHC2
with
(
ξ
:=
ξ
);
auto
.
Qed
.
Lemma
ThreadNamesSubst
:
forall
C
σ
l
,
In
l
(
ThreadNames
(
C
[
c
|
σ
]))
->
In
l
(
ThreadNames
C
)
\
/
exists
n
,
In
l
(
ThreadNames
(
σ
n
)).
Proof
using
.
intro
C
;
ChorInduction
C
;
cbn
;
intros
σ
p
i
;
auto
;
repeat
match
goal
with
|
[
i
:
_
\
/
_
|-
_
]
=>
destruct
i
;
auto
|
[
i
:
In
_
(
_
++
_
)
|-
_
]
=>
apply
in_app_or
in
i
|
[
IH
:
forall
σ
l
,
In
l
(
ThreadNames
(
?
C
[
c
|
σ
]))
->
In
l
(
ThreadNames
?
C
)
\
/
(
exists
n
,
In
l
(
ThreadNames
(
σ
n
))),
H
:
In
?
l
(
ThreadNames
(
?
C
[
c
|?
σ
]))
|-
_
]
=>
lazymatch
goal
with
|
[
_
:
In
l
(
ThreadNames
C
)
|-
_
]
=>
fail
|
[
_
:
In
l
(
ThreadNames
(
σ
_
))
|-
_
]
=>
fail
|
_
=>
let
H
'
:=
fresh
in
destruct
(
IH
σ
l
H
)
as
[
H
'
|
H
'
];
[
|
let
n
:=
fresh
"n"
in
destruct
H
'
as
[
n
H
'
]];
auto
end
|
[
H
:
In
?
p
(
ThreadNames
(
?
σ
?
n
))
|-
_
\
/
exists
m
,
In
?
p
(
ThreadNames
(
?
σ
m
))]
=>
right
;
exists
n
;
auto
end
.
all:
try
(
left
;
right
;
apply
in_or_app
;
auto
;
fail
).
-
right
.
unfold
ChorUpSubstForExpr
in
H0
.
exists
n0
.
rewrite
ChorExprRenameSpec
in
H0
.
rewrite
<-
ThreadNamesExprSubst
in
H0
;
auto
.
-
right
.
unfold
ChorUpSubstForExpr
in
H0
.
exists
n0
.
rewrite
ChorExprRenameSpec
in
H0
.
rewrite
<-
ThreadNamesExprSubst
in
H0
;
auto
.
-
right
.
unfold
ChorUpSubstForExpr
in
H0
.
unfold
ChorUpSubst
in
H0
;
destruct
n0
.
cbn
in
H0
;
destruct
H0
.
rewrite
ThreadNamesRenaming
in
H0
.
rewrite
ChorExprRenameSpec
in
H0
;
rewrite
<-
ThreadNamesExprSubst
in
H0
.
exists
n0
;
auto
.
-
unfold
ChorUpSubst
in
H
.
destruct
n0
;
[
cbn
in
H
;
destruct
H
|
].
destruct
n0
;
[
cbn
in
H
;
destruct
H
|
].
repeat
rewrite
ThreadNamesRenaming
in
H
.
right
;
exists
n0
;
auto
.
-
left
;
apply
in_or_app
;
auto
.
-
left
;
apply
in_or_app
;
auto
.
Qed
.
Lemma
ThreadNamesAfterStep
:
forall
R
B
C1
C2
l
,
ChorStep
R
B
C1
C2
->
In
l
(
ThreadNames
C2
)
->
In
l
(
ThreadNames
C1
).
Proof
using
.
intros
R
B
C1
C2
l
step
;
revert
l
;
induction
step
;
cbn
;
intros
p
i
;
auto
;
repeat
match
goal
with
|
[
i
:
_
\
/
_
|-
_
]
=>
destruct
i
;
auto
|
[
i
:
In
_
(
_
++
_
)
|-
_
]
=>
apply
in_app_or
in
i
|
[
IH
:
forall
l
,
In
l
(
ThreadNames
?
C2
)
->
In
l
(
ThreadNames
?
C1
),
i
:
In
?
l
(
ThreadNames
?
C2
)
|-
_
]
=>
lazymatch
goal
with
|
[
_
:
In
l
(
ThreadNames
C1
)
|-
_
]
=>
fail
|
_
=>
pose
proof
(
IH
l
i
);
auto
end
end
.
rewrite
<-
ThreadNamesExprSubst
in
i
;
auto
.
1
-
6
:
right
;
apply
in_or_app
;
auto
.
3
-
7
:
apply
in_or_app
;
auto
.
-
rewrite
<-
ThreadNamesExprSubst
in
i
;
auto
.
-
apply
ThreadNamesSubst
in
i
;
destruct
i
.
rewrite
<-
ThreadNamesExprSubst
in
H0
;
auto
.
destruct
H0
as
[
n
i
];
unfold
AppLocalSubst
in
i
.
destruct
n
;
[
|
cbn
in
i
;
destruct
i
].
cbn
in
i
;
auto
.
-
apply
ThreadNamesSubst
in
i
;
destruct
i
;
auto
.
destruct
H0
as
[
n
i
];
unfold
AppGlobalSubst
in
i
.
destruct
n
;
auto
.
destruct
n
;
auto
.
cbn
in
i
;
destruct
i
.
Qed
.
End
Choreography
.
ChoreographyCompiler.v
0 → 100644
View file @
bbd1e08d
This diff is collapsed.
Click to expand it.
ConcurrentLambda.v
View file @
bbd1e08d
This diff is collapsed.
Click to expand it.
Locations.v
View file @
bbd1e08d
...
...
@@ -32,6 +32,8 @@ Module Type Locations.
Declare
Instance
eq_equiv
:
@
Equivalence
t
eq
.
Parameter
eq_dec
:
forall
x
y
:
t
,
{
x
=
y
}
+
{
x
<>
y
}
.
Parameter
Nontrivial
:
t
.
(
*
We
don
'
t
want
to
talk
about
a
trivial
theory
*
)
(
*
Declare
Instance
lt_strorder
:
StrictOrder
lt
.
*
)
(
*
Declare
Instance
lt_compat
:
Proper
(
eq
==>
eq
==>
iff
)
lt
.
*
)
...
...
RestrictedSemantics.v
View file @
bbd1e08d
...
...
@@ -76,7 +76,7 @@ Module RestrictedSemantics (Import E : Expression) (Import TE : TypedExpression
(
AppGlobal
(
RecGlobal
C1
)
C2
)
(
C1
[
c
|
AppGlobalSubst
C1
C2
])
|
SyncStep
:
forall
(
l1
l2
:
Loc
)
(
d
:
LRChoice
)
(
C
:
Chor
)
(
B
:
list
Loc
),
~
In
l1
B
->
~
In
l2
B
->
~
In
l1
B
->
~
In
l2
B
->
l1
<>
l2
->
RChorStep
(
RSync
l1
d
l2
)
B
(
Sync
l1
d
l2
C
)
C
|
SyncIStep
:
forall
(
l1
l2
:
Loc
)
(
d
:
LRChoice
)
(
C1
C2
:
Chor
)
(
B
:
list
Loc
)
(
R
:
Redex
),
RChorStep
R
(
l1
::
l2
::
B
)
C1
C2
->
...
...
@@ -401,6 +401,7 @@ Module RestrictedSemantics (Import E : Expression) (Import TE : TypedExpression
|
StdIfFalseStep
:
forall
(
l
:
Loc
)
(
C1
C2
:
Chor
),
StdChorStep
(
If
l
ff
C1
C2
)
C2
|
StdSyncStep
:
forall
(
l1
l2
:
Loc
)
(
d
:
LRChoice
)
(
C
:
Chor
),
l1
<>
l2
->
StdChorStep
(
Sync
l1
d
l2
C
)
C
|
StdDefLocalIStep
:
forall
(
l
:
Loc
)
(
C1
C1
'
C2
:
Chor
),
StdChorStep
C1
C1
'
->
StdChorStep
(
DefLocal
l
C1
C2
)
(
DefLocal
l
C1
'
C2
)
...
...
TypedChoreography.v
View file @
bbd1e08d
...
...
@@ -44,6 +44,7 @@ Module TypedChoreography (L : Locations) (E : Expression) (TE : TypedExpression
|
TSync
:
forall
(
Γ
:
Loc
->
nat
->
ExprTyp
)
(
Δ
:
nat
->
ChorTyp
)
(
p
:
Loc
)
(
d
:
LRChoice
)
(
q
:
Loc
)
(
C
:
Chor
)
(
τ
:
ChorTyp
),
Γ
;;
Δ
⊢
c
C
:::
τ
->
p
<>
q
->
(
*
----------------------------------
*
)
Γ
;;
Δ
⊢
c
Sync
p
d
q
C
:::
τ
|
TDefLocal
:
forall
(
Γ
:
Loc
->
nat
->
ExprTyp
)
(
Δ
:
nat
->
ChorTyp
)
...
...
@@ -404,7 +405,7 @@ Module TypedChoreography (L : Locations) (E : Expression) (TE : TypedExpression
apply
Lt
.
lt_S_n
in
l
;
auto
.
-
apply
TIf
.
eapply
ExprClosedBelowTyping
;
[
|
apply
eq
Γ
|
];
eauto
.
eapply
IHtyp1
;
eauto
.
eapply
IHtyp2
;
eauto
.
-
apply
TSync
;
eapply
IHtyp
;
eauto
.
-
apply
TSync
;
auto
;
eapply
IHtyp
;
eauto
.
-
eapply
TDefLocal
.
eapply
IHtyp1
;
eauto
.
eapply
IHtyp2
;
eauto
.
cbn
in
*
;
intros
p0
k
l
.
destruct
(
L
.
eq_dec
p
p0
);
subst
;
auto
.
destruct
k
;
auto
.
...
...
_CoqProject
View file @
bbd1e08d
...
...
@@ -11,4 +11,6 @@ FunLMap.v
Choreography.v
TypedChoreography.v
SoundlyTypedChoreography.v
RestrictedSemantics.v
\ No newline at end of file
RestrictedSemantics.v
ConcurrentLambda.v
ChoreographyCompiler.v
\ No newline at end of file
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