Skip to content
GitLab
Projects
Groups
Snippets
Help
Loading...
Help
What's new
10
Help
Support
Community forum
Keyboard shortcuts
?
Submit feedback
Contribute to GitLab
Sign in / Register
Toggle navigation
Open sidebar
Iris
Actris
Commits
9d43893c
Commit
9d43893c
authored
May 05, 2020
by
Daniël Louwrink
Browse files
Options
Browse Files
Download
Email Patches
Plain Diff
start updating typing rules
parent
01ec9471
Changes
1
Hide whitespace changes
Inline
Side-by-side
Showing
1 changed file
with
80 additions
and
89 deletions
+80
-89
theories/logrel/term_typing_rules.v
theories/logrel/term_typing_rules.v
+80
-89
No files found.
theories/logrel/term_typing_rules.v
View file @
9d43893c
...
...
@@ -15,16 +15,14 @@ Section properties.
Implicit
Types
Γ
:
gmap
string
(
ltty
Σ
).
(** Variable properties *)
(* TODO(TYRULES) *)
Lemma
ltyped_var
Γ
(
x
:
string
)
A
:
Γ
!!
x
=
Some
A
→
⊢
Γ
⊨
x
:
A
⫤
<[
x
:
=
copy
-
A
]>
%
lty
Γ
.
Γ
!!
x
=
Some
A
→
⊢
Γ
⊨
x
:
A
⫤
<[
x
:
=
(
copy
-
A
)
%
lty
]>
Γ
.
Proof
.
iIntros
(
H
Γ
x
)
"!>"
;
iIntros
(
vs
)
"HΓ /="
.
iDestruct
(
env_ltyped_lookup
with
"HΓ"
)
as
(
v
Hx
)
"[HA HΓ]"
;
first
done
.
rewrite
Hx
.
iApply
wp_value
.
iDestruct
(
coreP_intro
with
"HA"
)
as
"#HAc"
.
iFrame
"HA"
.
iEval
(
rewrite
-
insert_delete
-(
insert_id
vs
x
v
)
//).
by
iApply
(
env_ltyped_insert
_
_
x
).
Qed
.
(* iDestruct (env_ltyped_lookup with "HΓ") as (v ->) "[HA HΓ]"; first done. *)
(* iApply wp_value. eauto with iFrame. *)
Admitted
.
(** Subtyping *)
Theorem
ltyped_subsumption
Γ
Γ
2
e
τ
1
τ
2
:
...
...
@@ -196,25 +194,26 @@ Section properties.
iRight
.
iExists
v
.
auto
.
Qed
.
Definition
paircase
:
val
:
=
λ
:
"pair"
"left"
"right"
,
Case
"pair"
"left"
"right"
.
Lemma
ltyped_paircase
A1
A2
B
:
⊢
∅
⊨
paircase
:
A1
+
A2
→
(
A1
⊸
B
)
⊸
(
A2
⊸
B
)
⊸
B
.
(* TODO(TYRULES) *)
Lemma
ltyped_case
Γ
1
Γ
2
Γ
3 e1
e2
e3
A1
A2
B
:
(
Γ
1
⊨
e1
:
A1
+
A2
⫤
Γ
2
)
-
∗
(
Γ
2
⊨
e2
:
A1
⊸
B
⫤
Γ
3
)
-
∗
(
Γ
2
⊨
e3
:
A2
⊸
B
⫤
Γ
3
)
-
∗
(
Γ
1
⊨
Case
e1
e2
e3
:
B
⫤
Γ
3
).
Proof
.
iIntros
(
vs
)
"!> HΓ /="
.
iApply
wp_value
.
iSplitL
;
last
by
iApply
env_ltyped_empty
.
rewrite
/
paircase
.
iIntros
"!>"
(
p
)
"Hp"
.
wp_pures
.
iIntros
(
f_left
)
"Hleft"
.
wp_pures
.
iIntros
(
f_right
)
"Hright"
.
wp_pures
.
iDestruct
"Hp"
as
"[Hp|Hp]"
.
-
iDestruct
"Hp"
as
(
w1
->)
"Hp"
.
wp_pures
.
wp_apply
(
wp_wand
with
"(Hleft [Hp //])"
).
iIntros
(
v
)
"HB"
.
iApply
"HB"
.
-
iDestruct
"Hp"
as
(
w2
->)
"Hp"
.
wp_pures
.
wp_apply
(
wp_wand
with
"(Hright [Hp //])"
).
iIntros
(
v
)
"HB"
.
iApply
"HB"
.
Q
ed
.
(*
iIntros (vs) "!> HΓ /=". iApply wp_value.
*)
(*
iSplitL; last by iApply env_ltyped_empty.
*)
(*
rewrite /paircase. iIntros "!>" (p) "Hp". wp_pures.
*)
(*
iIntros (f_left) "Hleft". wp_pures.
*)
(*
iIntros (f_right) "Hright". wp_pures.
*)
(*
iDestruct "Hp" as "[Hp|Hp]".
*)
(*
- iDestruct "Hp" as (w1 ->) "Hp". wp_pures.
*)
(*
wp_apply (wp_wand with "(Hleft [Hp //])").
*)
(*
iIntros (v) "HB". iApply "HB".
*)
(*
- iDestruct "Hp" as (w2 ->) "Hp". wp_pures.
*)
(*
wp_apply (wp_wand with "(Hright [Hp //])").
*)
(*
iIntros (v) "HB". iApply "HB".
*)
Admitt
ed
.
(** Universal Properties *)
Lemma
ltyped_tlam
Γ
e
k
(
C
:
lty
Σ
k
→
ltty
Σ
)
:
...
...
@@ -242,37 +241,42 @@ Section properties.
wp_apply
(
wp_wand
with
"(He [HΓ //])"
)
;
iIntros
(
w
)
"[HB $]"
.
by
iExists
M
.
Qed
.
Lemma
ltyped_unpack
{
k
}
Γ
1
Γ
2
Γ
3
x
e1
e2
(
C
:
lty
Σ
k
→
ltty
Σ
)
B
:
(
Γ
1
⊨
e1
:
lty_exist
C
⫤
Γ
2
)
-
∗
(
∀
Y
,
binder_insert
x
(
C
Y
)
Γ
2
⊨
e2
:
B
⫤
Γ
3
)
-
∗
Γ
1
⊨
(
let
:
x
:
=
e1
in
e2
)
:
B
⫤
binder_delete
x
Γ
3
.
(* TODO(TYRULES) *)
(* NOTE: This only works when `x` is a string, not when it is a (more general)
binder. This means that it doesn't work for anonymous binders, but there
really isn't any reason to unpack those anyway. *)
Lemma
ltyped_unpack
{
k
}
Γ
1
Γ
2
Γ
3
(
x
:
string
)
e
(
C
:
lty
Σ
k
→
ltty
Σ
)
A
:
Γ
1
!!
x
=
Some
(
lty_exist
C
)
→
(
∀
X
,
binder_insert
x
(
C
X
)
Γ
1
⊨
e
:
A
⫤
Γ
2
)
-
∗
(
Γ
1
⊨
e
:
A
⫤
Γ
2
).
Proof
.
iIntros
"#He1 #He2 !>"
.
iIntros
(
vs
)
"HΓ1"
=>
/=.
wp_apply
(
wp_wand
with
"(He1 HΓ1)"
).
iIntros
(
v
)
"[HC HΓ2]"
.
iDestruct
"HC"
as
(
X
)
"HX"
.
wp_pures
.
iDestruct
(
env_ltyped_insert
_
_
x
with
"HX HΓ2"
)
as
"HΓ2"
.
iDestruct
(
"He2"
with
"HΓ2"
)
as
"He2'"
.
destruct
x
as
[|
x
]
;
rewrite
/=
-
?subst_map_insert
//.
wp_apply
(
wp_wand
with
"He2'"
).
iIntros
(
w
)
"[HA2 HΓ3]"
.
iFrame
.
iApply
env_ltyped_delete
=>
//.
Q
ed
.
(*
iIntros "#He1 #He2 !>". iIntros (vs) "HΓ1"=> /=.
*)
(*
wp_apply (wp_wand with "(He1 HΓ1)").
*)
(*
iIntros (v) "[HC HΓ2]".
*)
(*
iDestruct "HC" as (X) "HX".
*)
(*
wp_pures.
*)
(*
iDestruct (env_ltyped_insert _ _ x with "HX HΓ2") as "HΓ2".
*)
(*
iDestruct ("He2" with "HΓ2") as "He2'".
*)
(*
destruct x as [|x]; rewrite /= -?subst_map_insert //.
*)
(*
wp_apply (wp_wand with "He2'").
*)
(*
iIntros (w) "[HA2 HΓ3]".
*)
(*
iFrame.
*)
(*
iApply env_ltyped_delete=> //.
*)
Admitt
ed
.
(** Mutable Reference properties *)
Definition
alloc
:
val
:
=
λ
:
"init"
,
ref
"init"
.
Lemma
ltyped_alloc
A
:
⊢
∅
⊨
alloc
:
A
→
ref_mut
A
.
(* TODO(TYRULES) *)
Lemma
ltyped_alloc
Γ
1
Γ
2
e
A
:
(
Γ
1
⊨
e
:
A
⫤
Γ
2
)
-
∗
(
Γ
1
⊨
ref
e
:
ref_mut
A
⫤
Γ
2
).
Proof
.
iIntros
(
vs
)
"!> HΓ /="
.
iApply
wp_value
.
iSplitL
;
last
by
iApply
env_ltyped_empty
.
iIntros
"!>"
(
v
)
"Hv"
.
rewrite
/
alloc
.
wp_pures
.
wp_alloc
l
as
"Hl"
.
iExists
l
,
v
.
iSplit
=>
//.
iFrame
"Hv Hl"
.
Q
ed
.
(*
iIntros (vs) "!> HΓ /=". iApply wp_value.
*)
(*
iSplitL; last by iApply env_ltyped_empty.
*)
(*
iIntros "!>" (v) "Hv". rewrite /alloc. wp_pures.
*)
(*
wp_alloc l as "Hl".
*)
(*
iExists l, v. iSplit=> //.
*)
(*
iFrame "Hv Hl".
*)
Admitt
ed
.
(* The intuition for the any is that the value is still there, but
it no longer holds any Iris resources. Just as in Rust, where a move
...
...
@@ -280,45 +284,32 @@ Section properties.
unmodified, but moves the resources, in the sense that you can no
longer use the memory at the old location. *)
Definition
load
:
val
:
=
λ
:
"r"
,
(!
"r"
,
"r"
).
Lemma
ltyped_load
A
:
⊢
∅
⊨
load
:
ref_mut
A
→
A
*
ref_mut
any
.
Lemma
ltyped_load
Γ
(
x
:
string
)
A
:
Γ
!!
x
=
Some
(
ref_mut
A
)%
lty
→
⊢
Γ
⊨
!
x
:
A
⫤
<[
x
:
=
(
ref_mut
(
copy
-
A
))%
lty
]>
Γ
.
Proof
.
iIntros
(
vs
)
"!> HΓ /="
.
iApply
wp_value
.
iSplitL
;
last
by
iApply
env_ltyped_empty
.
iIntros
"!>"
(
v
)
"Hv"
.
rewrite
/
load
.
wp_pures
.
iDestruct
"Hv"
as
(
l
w
->)
"[Hl Hw]"
.
wp_load
.
wp_pures
.
iExists
w
,
#
l
.
iSplit
=>
//.
iFrame
"Hw"
.
iExists
l
,
w
.
iSplit
=>
//.
iFrame
"Hl"
.
by
iModIntro
.
Qed
.
(* TODO(COPY) *)
(* Lemma ltyped_load_copy A {copyA : LTyCopy A} : *)
(* ⊢ ∅ ⊨ load : ref_mut A → A * ref_mut A. *)
(* Proof. *)
(* iIntros (vs) "!> HΓ /=". *)
(* iApply wp_value. *)
(* iIntros "!>" (v) "Hv". rewrite /load. wp_pures. *)
(* iDestruct "Hv" as (l w ->) "[Hl #Hw]". *)
(* wp_load. wp_pures. *)
(* iExists w, #l. iSplit=> //. iFrame "Hw". *)
(* iExists l, w. iSplit=> //. iFrame "Hl". *)
(* by iModIntro. *)
(* Qed. *)
Definition
store
:
val
:
=
λ
:
"r"
"new"
,
"r"
<-
"new"
;;
"r"
.
Lemma
ltyped_store
A
B
:
⊢
∅
⊨
store
:
ref_mut
A
→
B
⊸
ref_mut
B
.
(* iIntros (vs) "!> HΓ /=". iApply wp_value. *)
(* iSplitL; last by iApply env_ltyped_empty. *)
(* iIntros "!>" (v) "Hv". rewrite /load. wp_pures. *)
(* iDestruct "Hv" as (l w ->) "[Hl Hw]". *)
(* wp_load. wp_pures. *)
(* iExists w, #l. iSplit=> //. iFrame "Hw". *)
(* iExists l, w. iSplit=> //. iFrame "Hl". *)
(* by iModIntro. *)
Admitted
.
Lemma
ltyped_store
Γ
Γ
'
x
e1
e2
A
B
:
Γ
'
!!
x
=
Some
(
ref_mut
A
)%
lty
→
(
Γ
⊨
e2
:
B
⫤
Γ
'
)
-
∗
Γ
⊨
e1
<-
e2
:
()
⫤
<[
x
:
=
(
ref_mut
B
)%
lty
]>
Γ
'
.
Proof
.
iIntros
(
vs
)
"!> HΓ /="
.
iApply
wp_value
.
iSplitL
;
last
by
iApply
env_ltyped_empty
.
iIntros
"!>"
(
v
)
"Hv"
.
rewrite
/
store
.
wp_pures
.
iDestruct
"Hv"
as
(
l
old
->)
"[Hl Hold]"
.
iIntros
(
new
)
"Hnew"
.
wp_store
.
iExists
l
,
new
.
eauto
with
iFrame
.
Q
ed
.
(*
iIntros (vs) "!> HΓ /=". iApply wp_value.
*)
(*
iSplitL; last by iApply env_ltyped_empty.
*)
(*
iIntros "!>" (v) "Hv". rewrite /store. wp_pures.
*)
(*
iDestruct "Hv" as (l old ->) "[Hl Hold]".
*)
(*
iIntros (new) "Hnew". wp_store.
*)
(*
iExists l, new. eauto with iFrame.
*)
Admitt
ed
.
(** Weak Reference properties *)
Definition
fetch_and_add
:
val
:
=
λ
:
"r"
"inc"
,
FAA
"r"
"inc"
.
...
...
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