Skip to content
GitLab
Explore
Sign in
Primary navigation
Search or go to…
Project
ReLoC
Manage
Activity
Members
Labels
Plan
Issues
Issue boards
Milestones
Code
Merge requests
Repository
Branches
Commits
Tags
Repository graph
Compare revisions
Snippets
Build
Pipelines
Jobs
Pipeline schedules
Artifacts
Deploy
Releases
Model registry
Operate
Environments
Monitor
Incidents
Service Desk
Analyze
Value stream analytics
Contributor analytics
CI/CD analytics
Repository analytics
Model experiments
Help
Help
Support
GitLab documentation
Compare GitLab plans
Community forum
Contribute to GitLab
Provide feedback
Terms and privacy
Keyboard shortcuts
?
Snippets
Groups
Projects
Show more breadcrumbs
Iris
ReLoC
Commits
aa5cdd7c
Commit
aa5cdd7c
authored
4 years ago
by
Dan Frumin
Browse files
Options
Downloads
Patches
Plain Diff
Finish the closedness theorem.
parent
adb7f80c
Branches
Branches containing commit
No related tags found
1 merge request
!3
CKA stuff
Changes
1
Hide whitespace changes
Inline
Side-by-side
Showing
1 changed file
theories/typing/types.v
+82
-39
82 additions, 39 deletions
theories/typing/types.v
with
82 additions
and
39 deletions
theories/typing/types.v
+
82
−
39
View file @
aa5cdd7c
...
@@ -251,7 +251,7 @@ Proof. destruct op; naive_solver. Qed.
...
@@ -251,7 +251,7 @@ Proof. destruct op; naive_solver. Qed.
Lemma elements_insert Γ x τ :
Lemma elements_insert Γ x τ :
elements (dom stringset (<[x:=τ]> Γ)) = x :: elements (dom stringset Γ).
elements (dom stringset (<[x:=τ]> Γ)) = x :: elements (dom stringset Γ).
But this does not hold (only up to multiset equality).
But this does not hold (
it holds
only up to multiset equality).
So we use an auxiliary definition with sets *)
So we use an auxiliary definition with sets *)
Definition
maybe_insert_binder
(
x
:
binder
)
(
X
:
stringset
)
:
stringset
:=
Definition
maybe_insert_binder
(
x
:
binder
)
(
X
:
stringset
)
:
stringset
:=
...
@@ -260,7 +260,7 @@ Definition maybe_insert_binder (x : binder) (X : stringset) : stringset :=
...
@@ -260,7 +260,7 @@ Definition maybe_insert_binder (x : binder) (X : stringset) : stringset :=
|
BNamed
f
=>
{[
f
]}
∪
X
|
BNamed
f
=>
{[
f
]}
∪
X
end
.
end
.
Fixpoint
is_closed_expr_set
(
X
:
stringset
)
(
e
:
expr
)
:
bool
:=
Local
Fixpoint
is_closed_expr_set
(
X
:
stringset
)
(
e
:
expr
)
:
bool
:=
match
e
with
match
e
with
|
Val
v
=>
is_closed_val_set
v
|
Val
v
=>
is_closed_val_set
v
|
Var
x
=>
bool_decide
(
x
∈
X
)
|
Var
x
=>
bool_decide
(
x
∈
X
)
...
@@ -281,62 +281,105 @@ with is_closed_val_set (v : val) : bool :=
...
@@ -281,62 +281,105 @@ with is_closed_val_set (v : val) : bool :=
|
InjLV
v
|
InjRV
v
=>
is_closed_val_set
v
|
InjLV
v
|
InjRV
v
=>
is_closed_val_set
v
end
.
end
.
(* Lemma is_closed_expr_set_sound (X : stringset) (e : expr) : *)
Lemma
is_closed_expr_permute
(
e
:
expr
)
(
xs
ys
:
list
string
)
:
(* is_closed_expr_set X e → is_closed_expr (elements X) e *)
xs
≡
ₚ
ys
→
(* with is_closed_val_set_sound (v : val) : *)
is_closed_expr
xs
e
=
is_closed_expr
ys
e
.
(* is_closed_val_set v → is_closed_val v. *)
Proof
.
(* Proof. *)
revert
xs
ys
.
induction
e
=>
xs
ys
Hxsys
/=
;
(* - induction e; simplify_eq/=. *)
repeat
match
goal
with
(* + apply is_closed_val_set_sound. *)
|
[
|
-
_
&&
_
=
_
&&
_
]
=>
f_equal
(* + intros. case_bool_decide; last done. *)
|
[
H
:
∀
xs
ys
,
xs
≡
ₚ
ys
→
is_closed_expr
xs
_
=
is_closed_expr
ys
_
(* apply bool_decide_pack. by apply (elem_of_elements X x). *)
|
-
is_closed_expr
_
_
=
is_closed_expr
_
_
]
=>
eapply
H
;
eauto
(* + *)
end
;
try
done
.
-
apply
bool_decide_iff
.
by
rewrite
Hxsys
.
-
by
rewrite
Hxsys
.
Qed
.
Global
Instance
is_closed_expr_proper
:
Proper
(
Permutation
==>
eq
==>
eq
)
is_closed_expr
.
Proof
.
intros
X1
X2
HX
x
y
->
.
by
eapply
is_closed_expr_permute
.
Qed
.
Lemma
typed_is_closed
Γ
e
τ
:
Lemma
is_closed_expr_set_sound
(
X
:
stringset
)
(
e
:
expr
)
:
is_closed_expr_set
X
e
→
is_closed_expr
(
elements
X
)
e
with
is_closed_val_set_sound
(
v
:
val
)
:
is_closed_val_set
v
→
is_closed_val
v
.
Proof
.
-
induction
e
;
simplify_eq
/=
;
try
by
(
intros
;
destruct_and
?;
split_and
?;
eauto
)
.
+
intros
.
case_bool_decide
;
last
done
.
by
apply
bool_decide_pack
,
elem_of_elements
.
+
destruct
f
as
[|
f
],
x
as
[|
x
];
simplify_eq
/=.
*
eapply
IHe
.
*
intros
H
%
is_closed_expr_set_sound
.
eapply
is_closed_weaken
;
eauto
.
by
set_solver
.
*
intros
H
%
is_closed_expr_set_sound
.
eapply
is_closed_weaken
;
eauto
.
by
set_solver
.
*
intros
H
%
is_closed_expr_set_sound
.
eapply
is_closed_weaken
;
eauto
.
by
set_solver
.
-
induction
v
;
simplify_eq
/=
;
try
naive_solver
.
destruct
f
as
[|
f
],
x
as
[|
x
];
simplify_eq
/=
;
intros
H
%
is_closed_expr_set_sound
;
revert
H
.
+
set_solver
.
+
by
rewrite
?right_id_L
elements_singleton
.
+
by
rewrite
?right_id_L
elements_singleton
.
+
rewrite
?right_id_L
.
intros
.
eapply
is_closed_weaken
;
eauto
.
destruct
(
decide
(
f
=
x
))
as
[
->
|?]
.
*
rewrite
union_idemp_L
elements_singleton
.
set_solver
.
*
rewrite
elements_disj_union
;
last
set_solver
.
rewrite
!
elements_singleton
.
set_solver
.
Qed
.
Local
Lemma
typed_is_closed_set
Γ
e
τ
:
Γ
⊢
ₜ
e
:
τ
→
is_closed_expr_set
(
dom
stringset
Γ
)
e
Γ
⊢
ₜ
e
:
τ
→
is_closed_expr_set
(
dom
stringset
Γ
)
e
with
typed_is_closed_val
v
τ
:
with
typed_is_closed_val
_set
v
τ
:
⊢
ᵥ
v
:
τ
→
is_closed_val_set
v
.
⊢
ᵥ
v
:
τ
→
is_closed_val_set
v
.
Proof
.
Proof
.
-
induction
1
;
simplify_eq
/=
;
try
(
split_and
?;
by
eapply
typed_is_closed
)
.
-
induction
1
;
simplify_eq
/=
;
try
(
split_and
?;
by
eapply
typed_is_closed
_set
)
.
+
apply
bool_decide_pack
.
apply
elem_of_dom
.
eauto
.
+
apply
bool_decide_pack
.
apply
elem_of_dom
.
eauto
.
+
by
eapply
typed_is_closed_val
.
+
by
eapply
typed_is_closed_val
_set
.
+
destruct
f
as
[|
f
],
x
as
[|
x
];
simplify_eq
/=.
+
destruct
f
as
[|
f
],
x
as
[|
x
];
simplify_eq
/=.
*
eapply
typed_is_closed
.
eauto
.
*
eapply
typed_is_closed
_set
.
eauto
.
*
specialize
(
typed_is_closed
(
<
[
x
:=
τ1
]
>
Γ
)
e
τ2
H
)
.
*
specialize
(
typed_is_closed
_set
(
<
[
x
:=
τ1
]
>
Γ
)
e
τ2
H
)
.
revert
typed_is_closed
.
by
rewrite
dom_insert_L
.
revert
typed_is_closed
_set
.
by
rewrite
dom_insert_L
.
*
specialize
(
typed_is_closed
(
<
[
f
:=(
τ1
→
τ2
)
%
ty
]
>
Γ
)
e
τ2
H
)
.
*
specialize
(
typed_is_closed
_set
(
<
[
f
:=(
τ1
→
τ2
)
%
ty
]
>
Γ
)
e
τ2
H
)
.
revert
typed_is_closed
.
by
rewrite
dom_insert_L
.
revert
typed_is_closed
_set
.
by
rewrite
dom_insert_L
.
*
specialize
(
typed_is_closed
_
e
τ2
H
)
.
*
specialize
(
typed_is_closed
_set
_
e
τ2
H
)
.
revert
typed_is_closed
.
revert
typed_is_closed
_set
.
rewrite
(
dom_insert_L
(
D
:=
stringset
)
(
<
[
x
:=
τ1
]
>
Γ
)
f
(
τ1
→
τ2
)
%
ty
)
.
rewrite
(
dom_insert_L
(
D
:=
stringset
)
(
<
[
x
:=
τ1
]
>
Γ
)
f
(
τ1
→
τ2
)
%
ty
)
.
by
rewrite
dom_insert_L
.
by
rewrite
dom_insert_L
.
+
specialize
(
typed_is_closed
(
⤉
Γ
)
e
τ
H
)
.
+
specialize
(
typed_is_closed
_set
(
⤉
Γ
)
e
τ
H
)
.
revert
typed_is_closed
.
by
rewrite
dom_fmap_L
.
revert
typed_is_closed
_set
.
by
rewrite
dom_fmap_L
.
+
by
split_and
?
.
+
by
split_and
?
.
+
by
split_and
?
.
+
by
split_and
?
.
+
split_and
?;
eauto
.
try
(
apply
bool_decide_pack
;
by
set_solver
)
.
+
split_and
?;
eauto
.
try
(
apply
bool_decide_pack
;
by
set_solver
)
.
destruct
x
as
[|
x
];
simplify_eq
/=.
destruct
x
as
[|
x
];
simplify_eq
/=.
++
specialize
(
typed_is_closed
_
_
_
H0
)
.
++
specialize
(
typed_is_closed
_set
_
_
_
H0
)
.
revert
typed_is_closed
.
by
rewrite
dom_fmap_L
.
revert
typed_is_closed
_set
.
by
rewrite
dom_fmap_L
.
++
specialize
(
typed_is_closed
_
_
_
H0
)
.
++
specialize
(
typed_is_closed
_set
_
_
_
H0
)
.
revert
typed_is_closed
.
rewrite
(
dom_insert_L
(
D
:=
stringset
)
(
⤉
Γ
)
x
)
.
revert
typed_is_closed
_set
.
rewrite
(
dom_insert_L
(
D
:=
stringset
)
(
⤉
Γ
)
x
)
.
by
rewrite
dom_fmap_L
.
by
rewrite
dom_fmap_L
.
-
induction
1
;
simplify_eq
/=
;
try
done
.
-
induction
1
;
simplify_eq
/=
;
try
done
.
+
by
split_and
?
.
+
by
split_and
?
.
+
destruct
f
as
[|
f
],
x
as
[|
x
];
simplify_eq
/=.
+
destruct
f
as
[|
f
],
x
as
[|
x
];
simplify_eq
/=.
*
specialize
(
typed_is_closed
_
_
_
H
)
.
*
specialize
(
typed_is_closed
_set
_
_
_
H
)
.
revert
typed_is_closed
.
by
rewrite
dom_empty_L
.
revert
typed_is_closed
_set
.
by
rewrite
dom_empty_L
.
*
specialize
(
typed_is_closed
(
<
[
x
:=
τ1
]
>∅
)
_
_
H
)
.
*
specialize
(
typed_is_closed
_set
(
<
[
x
:=
τ1
]
>∅
)
_
_
H
)
.
revert
typed_is_closed
.
by
rewrite
dom_insert_L
dom_empty_L
.
revert
typed_is_closed
_set
.
by
rewrite
dom_insert_L
dom_empty_L
.
*
specialize
(
typed_is_closed
_
_
_
H
)
.
*
specialize
(
typed_is_closed
_set
_
_
_
H
)
.
revert
typed_is_closed
.
revert
typed_is_closed
_set
.
rewrite
(
dom_insert_L
(
D
:=
stringset
)
_
f
(
τ1
→
τ2
)
%
ty
)
.
rewrite
(
dom_insert_L
(
D
:=
stringset
)
_
f
(
τ1
→
τ2
)
%
ty
)
.
by
rewrite
dom_empty_L
.
by
rewrite
dom_empty_L
.
*
specialize
(
typed_is_closed
_
_
_
H
)
.
*
specialize
(
typed_is_closed
_set
_
_
_
H
)
.
revert
typed_is_closed
.
revert
typed_is_closed
_set
.
rewrite
(
dom_insert_L
(
D
:=
stringset
)
(
<
[
x
:=
τ1
]
>
∅
)
f
(
τ1
→
τ2
)
%
ty
)
.
rewrite
(
dom_insert_L
(
D
:=
stringset
)
(
<
[
x
:=
τ1
]
>
∅
)
f
(
τ1
→
τ2
)
%
ty
)
.
by
rewrite
dom_insert_L
dom_empty_L
.
by
rewrite
dom_insert_L
dom_empty_L
.
+
specialize
(
typed_is_closed
_
_
_
H
)
.
+
specialize
(
typed_is_closed
_set
_
_
_
H
)
.
revert
typed_is_closed
.
by
rewrite
dom_empty_L
.
revert
typed_is_closed
_set
.
by
rewrite
dom_empty_L
.
Qed
.
Qed
.
Theorem
typed_is_closed
Γ
e
τ
:
Γ
⊢
ₜ
e
:
τ
→
is_closed_expr
(
elements
(
dom
stringset
Γ
))
e
.
Proof
.
intros
.
eapply
is_closed_expr_set_sound
,
typed_is_closed_set
;
eauto
.
Qed
.
This diff is collapsed.
Click to expand it.
Preview
0%
Loading
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!
Save comment
Cancel
Please
register
or
sign in
to comment