Skip to content
GitLab
Explore
Sign in
Primary navigation
Search or go to…
Project
A
Actris
Manage
Activity
Members
Labels
Plan
Issues
Issue boards
Milestones
Code
Merge requests
Repository
Branches
Commits
Tags
Repository graph
Compare revisions
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
Actris
Commits
a2df6db0
Commit
a2df6db0
authored
4 years ago
by
Jonas Kastberg
Browse files
Options
Downloads
Patches
Plain Diff
Revert "Bounded subtyping."
This reverts commit
582ce16e
.
parent
d20b3324
No related branches found
No related tags found
No related merge requests found
Pipeline
#27519
passed
4 years ago
Stage: build
Changes
3
Pipelines
1
Hide whitespace changes
Inline
Side-by-side
Showing
3 changed files
theories/logrel/subtyping_rules.v
+21
-34
21 additions, 34 deletions
theories/logrel/subtyping_rules.v
theories/logrel/term_types.v
+14
-21
14 additions, 21 deletions
theories/logrel/term_types.v
theories/logrel/term_typing_rules.v
+20
-16
20 additions, 16 deletions
theories/logrel/term_typing_rules.v
with
55 additions
and
71 deletions
theories/logrel/subtyping_rules.v
+
21
−
34
View file @
a2df6db0
...
@@ -49,13 +49,9 @@ Section subtyping_rules.
...
@@ -49,13 +49,9 @@ Section subtyping_rules.
Qed
.
Qed
.
(** Term subtyping *)
(** Term subtyping *)
Lemma
lty_le_
bot
A
:
⊢
⊥
<:
A
.
Lemma
lty_le_
any
A
:
⊢
A
<:
any
.
Proof
.
by
iIntros
(
v
)
"!> H"
.
Qed
.
Proof
.
by
iIntros
(
v
)
"!> H"
.
Qed
.
Lemma
lty_copyable_bot
:
⊢@
{
iPropI
Σ
}
lty_copyable
⊥
.
Lemma
lty_copyable_any
:
⊢@
{
iPropI
Σ
}
lty_copyable
any
.
Proof
.
iApply
lty_le_bot
.
Qed
.
Lemma
lty_le_top
A
:
⊢
A
<:
⊤
.
Proof
.
by
iIntros
(
v
)
"!> H"
.
Qed
.
Lemma
lty_copyable_top
:
⊢@
{
iPropI
Σ
}
lty_copyable
⊤
.
Proof
.
iIntros
(
v
)
"!> #Hv !>"
.
iFrame
"Hv"
.
Qed
.
Proof
.
iIntros
(
v
)
"!> #Hv !>"
.
iFrame
"Hv"
.
Qed
.
Lemma
lty_le_copy
A
B
:
A
<:
B
-∗
copy
A
<:
copy
B
.
Lemma
lty_le_copy
A
B
:
A
<:
B
-∗
copy
A
<:
copy
B
.
...
@@ -133,41 +129,32 @@ Section subtyping_rules.
...
@@ -133,41 +129,32 @@ Section subtyping_rules.
try
iModIntro
;
first
[
iLeft
;
by
auto
|
iRight
;
by
auto
]
.
try
iModIntro
;
first
[
iLeft
;
by
auto
|
iRight
;
by
auto
]
.
Qed
.
Qed
.
Lemma
lty_le_forall
{
k
}
(
Mlow1
Mup1
Mlow2
Mup2
:
lty
Σ
k
)
C1
C2
:
Lemma
lty_le_forall
C1
C2
:
Mlow1
<:
Mlow2
-∗
Mup2
<:
Mup1
-∗
▷
(
∀
A
,
C1
A
<:
C2
A
)
-∗
▷
(
∀
M
,
Mlow2
<:
M
-∗
M
<:
Mup2
-∗
C1
M
<:
C2
M
)
-∗
(
∀
A
,
C1
A
)
<:
(
∀
A
,
C2
A
)
.
lty_forall
Mlow1
Mup1
C1
<:
lty_forall
Mlow2
Mup2
C2
.
Proof
.
Proof
.
iIntros
"#Hlow #Hup #Hle"
(
v
)
"!> H"
.
iIntros
(
M
)
"#Hlow' #Hup'"
.
iIntros
"#Hle"
(
v
)
"!> H"
.
iIntros
(
w
)
.
iApply
wp_step_fupd
;
first
done
.
iApply
(
wp_step_fupd
);
first
done
.
{
iIntros
"!> !> !>"
.
iExact
"Hle"
.
}
{
iIntros
"!>!>!>"
.
iExact
"Hle"
.
}
iApply
(
wp_wand
with
"(H [] [])"
)
.
iApply
(
wp_wand
with
"H"
)
.
iIntros
(
v'
)
"H Hle' !>"
.
{
iApply
(
lty_le_trans
with
"Hlow Hlow'"
)
.
}
by
iApply
"Hle'"
.
{
iApply
(
lty_le_trans
with
"Hup' Hup"
)
.
}
iIntros
(
v'
)
"H Hle' !>"
.
by
iApply
"Hle'"
.
Qed
.
Qed
.
(* TODO(COPY) TODO(VALUERES): Do the forall type former, once we have the value restriction *)
(* TODO(COPY) TODO(VALUERES): Do the forall type former, once we have the value restriction *)
Lemma
lty_le_exist
{
k
}
(
Mlow1
Mup1
Mlow2
Mup2
:
lty
Σ
k
)
C1
C2
:
Lemma
lty_le_exist
C1
C2
:
Mlow2
<:
Mlow1
-∗
Mup1
<:
Mup2
-∗
▷
(
∀
A
,
C1
A
<:
C2
A
)
-∗
▷
(
∀
M
,
Mlow1
<:
M
-∗
M
<:
Mup1
-∗
C1
M
<:
C2
M
)
-∗
(
∃
A
,
C1
A
)
<:
(
∃
A
,
C2
A
)
.
lty_exist
Mlow1
Mup1
C1
<:
lty_exist
Mlow2
Mup2
C2
.
Proof
.
Proof
.
iIntros
"#Hlow #Hup #Hle"
(
v
)
"!>"
.
iDestruct
1
as
(
M
)
"(#Hlow' & #Hup' & H)"
.
iIntros
"#Hle"
(
v
)
"!>"
.
iDestruct
1
as
(
A
)
"H"
.
iExists
A
.
by
iApply
"Hle"
.
iExists
M
.
iSplit
;
[|
iSplit
]
.
{
iApply
(
lty_le_trans
with
"Hlow Hlow'"
)
.
}
{
iApply
(
lty_le_trans
with
"Hup' Hup"
)
.
}
by
iApply
"Hle"
.
Qed
.
Qed
.
Lemma
lty_le_exist_intro
{
k
}
(
Mlow
Mup
:
lty
Σ
k
)
C
M
:
Lemma
lty_le_exist_elim
C
B
:
Mlow
<:
M
-∗
M
<:
Mup
-∗
⊢
C
B
<:
∃
A
,
C
A
.
C
M
<:
lty_exist
Mlow
Mup
C
.
Proof
.
iIntros
"!>"
(
v
)
"Hle"
.
by
iExists
B
.
Qed
.
Proof
.
iIntros
"#Hlow #Hup !>"
(
v
)
"Hle"
.
iExists
M
.
auto
.
Qed
.
Lemma
lty_le_exist_copy
F
:
Lemma
lty_le_exist_copy
{
k
}
(
Mlow
Mup
:
lty
Σ
k
)
C
:
⊢
(
∃
A
,
copy
(
F
A
))
<:>
copy
(
∃
A
,
F
A
)
.
⊢
lty_exist
Mlow
Mup
(
λ
M
,
copy
(
C
M
))
%
lty
<:>
copy
(
lty_exist
Mlow
Mup
C
)
.
Proof
.
Proof
.
iSplit
;
iIntros
"!>"
(
v
);
iSplit
;
iIntros
"!>"
(
v
);
iDestruct
1
as
(
A
)
"#Hv"
;
i
Destruct
1
as
(
A
)
"#(Hlow & Hup & Hv)"
;
iExists
A
;
auto
.
i
Exists
A
;
repeat
iModIntro
;
iApply
"Hv"
.
Qed
.
Qed
.
(* TODO(COPY): Commuting rule for μ, allowing `copy` to move outside the μ *)
(* TODO(COPY): Commuting rule for μ, allowing `copy` to move outside the μ *)
...
...
This diff is collapsed.
Click to expand it.
theories/logrel/term_types.v
+
14
−
21
View file @
a2df6db0
...
@@ -4,8 +4,7 @@ From iris.heap_lang Require Export spin_lock.
...
@@ -4,8 +4,7 @@ From iris.heap_lang Require Export spin_lock.
From
actris
.
logrel
Require
Export
subtyping
.
From
actris
.
logrel
Require
Export
subtyping
.
From
actris
.
channel
Require
Export
channel
.
From
actris
.
channel
Require
Export
channel
.
Instance
lty_bot
{
Σ
}
:
Bottom
(
ltty
Σ
)
:=
Ltty
(
λ
w
,
False
%
I
)
.
Definition
lty_any
{
Σ
}
:
ltty
Σ
:=
Ltty
(
λ
w
,
True
%
I
)
.
Instance
lty_top
{
Σ
}
:
Top
(
ltty
Σ
)
:=
Ltty
(
λ
w
,
True
%
I
)
.
Definition
lty_copy
{
Σ
}
(
A
:
ltty
Σ
)
:
ltty
Σ
:=
Ltty
(
λ
w
,
□
ltty_car
A
w
)
%
I
.
Definition
lty_copy
{
Σ
}
(
A
:
ltty
Σ
)
:
ltty
Σ
:=
Ltty
(
λ
w
,
□
ltty_car
A
w
)
%
I
.
Definition
lty_copy_inv
{
Σ
}
(
A
:
ltty
Σ
)
:
ltty
Σ
:=
Ltty
(
λ
w
,
coreP
(
ltty_car
A
w
))
.
Definition
lty_copy_inv
{
Σ
}
(
A
:
ltty
Σ
)
:
ltty
Σ
:=
Ltty
(
λ
w
,
coreP
(
ltty_car
A
w
))
.
...
@@ -25,11 +24,10 @@ Definition lty_sum {Σ} (A1 A2 : ltty Σ) : ltty Σ := Ltty (λ w,
...
@@ -25,11 +24,10 @@ Definition lty_sum {Σ} (A1 A2 : ltty Σ) : ltty Σ := Ltty (λ w,
(
∃
w1
,
⌜
w
=
InjLV
w1
⌝
∗
▷
ltty_car
A1
w1
)
∨
(
∃
w1
,
⌜
w
=
InjLV
w1
⌝
∗
▷
ltty_car
A1
w1
)
∨
(
∃
w2
,
⌜
w
=
InjRV
w2
⌝
∗
▷
ltty_car
A2
w2
))
%
I
.
(
∃
w2
,
⌜
w
=
InjRV
w2
⌝
∗
▷
ltty_car
A2
w2
))
%
I
.
Definition
lty_forall
`{
heapG
Σ
}
{
k
}
(
Mlow
Mup
:
lty
Σ
k
)
Definition
lty_forall
`{
heapG
Σ
}
{
k
}
(
C
:
lty
Σ
k
→
ltty
Σ
)
:
ltty
Σ
:=
(
C
:
lty
Σ
k
→
ltty
Σ
)
:
ltty
Σ
:=
Ltty
(
λ
w
,
∀
A
,
WP
w
#
()
{{
ltty_car
(
C
A
)
}})
%
I
.
Ltty
(
λ
w
,
∀
M
,
Mlow
<:
M
-∗
M
<:
Mup
-∗
WP
w
#
()
{{
ltty_car
(
C
M
)
}})
%
I
.
Definition
lty_exist
{
Σ
k
}
(
C
:
lty
Σ
k
→
ltty
Σ
)
:
ltty
Σ
:=
Definition
lty_exist
{
Σ
k
}
(
Mlow
Mup
:
lty
Σ
k
)
(
C
:
lty
Σ
k
→
ltty
Σ
)
:
ltty
Σ
:=
Ltty
(
λ
w
,
∃
A
,
▷
ltty_car
(
C
A
)
w
)
%
I
.
Ltty
(
λ
w
,
∃
M
,
Mlow
<:
M
∗
M
<:
Mup
∗
▷
ltty_car
(
C
M
)
w
)
%
I
.
Definition
lty_ref_mut
`{
heapG
Σ
}
(
A
:
ltty
Σ
)
:
ltty
Σ
:=
Ltty
(
λ
w
,
Definition
lty_ref_mut
`{
heapG
Σ
}
(
A
:
ltty
Σ
)
:
ltty
Σ
:=
Ltty
(
λ
w
,
∃
(
l
:
loc
)
(
v
:
val
),
⌜
w
=
#
l
⌝
∗
l
↦
v
∗
▷
ltty_car
A
v
)
%
I
.
∃
(
l
:
loc
)
(
v
:
val
),
⌜
w
=
#
l
⌝
∗
l
↦
v
∗
▷
ltty_car
A
v
)
%
I
.
...
@@ -64,6 +62,7 @@ Instance: Params (@lty_mutex) 3 := {}.
...
@@ -64,6 +62,7 @@ Instance: Params (@lty_mutex) 3 := {}.
Instance
:
Params
(
@
lty_mutexguard
)
3
:=
{}
.
Instance
:
Params
(
@
lty_mutexguard
)
3
:=
{}
.
Instance
:
Params
(
@
lty_chan
)
3
:=
{}
.
Instance
:
Params
(
@
lty_chan
)
3
:=
{}
.
Notation
any
:=
lty_any
.
Notation
"()"
:=
lty_unit
:
lty_scope
.
Notation
"()"
:=
lty_unit
:
lty_scope
.
Notation
"'copy' A"
:=
(
lty_copy
A
)
(
at
level
10
)
:
lty_scope
.
Notation
"'copy' A"
:=
(
lty_copy
A
)
(
at
level
10
)
:
lty_scope
.
Notation
"'copy-' A"
:=
(
lty_copy_inv
A
)
(
at
level
10
)
:
lty_scope
.
Notation
"'copy-' A"
:=
(
lty_copy_inv
A
)
(
at
level
10
)
:
lty_scope
.
...
@@ -74,11 +73,10 @@ Notation "A → B" := (lty_copy (lty_arr A B)) : lty_scope.
...
@@ -74,11 +73,10 @@ Notation "A → B" := (lty_copy (lty_arr A B)) : lty_scope.
Infix
"*"
:=
lty_prod
:
lty_scope
.
Infix
"*"
:=
lty_prod
:
lty_scope
.
Infix
"+"
:=
lty_sum
:
lty_scope
.
Infix
"+"
:=
lty_sum
:
lty_scope
.
(* TODO: Have nice notations for bounded quantifiers *)
Notation
"∀ A1 .. An , C"
:=
Notation
"∀ A1 .. An , C"
:=
(
lty_forall
⊥
⊤
(
λ
A1
,
.
.
(
lty_forall
⊥
⊤
(
λ
An
,
C
%
lty
))
..))
:
lty_scope
.
(
lty_forall
(
λ
A1
,
.
.
(
lty_forall
(
λ
An
,
C
%
lty
))
..))
:
lty_scope
.
Notation
"∃ A1 .. An , C"
:=
Notation
"∃ A1 .. An , C"
:=
(
lty_exist
⊥
⊤
(
λ
A1
,
.
.
(
lty_exist
⊥
⊤
(
λ
An
,
C
%
lty
))
..))
:
lty_scope
.
(
lty_exist
(
λ
A1
,
.
.
(
lty_exist
(
λ
An
,
C
%
lty
))
..))
:
lty_scope
.
Notation
"'ref_mut' A"
:=
(
lty_ref_mut
A
)
(
at
level
10
)
:
lty_scope
.
Notation
"'ref_mut' A"
:=
(
lty_ref_mut
A
)
(
at
level
10
)
:
lty_scope
.
Notation
"'ref_shr' A"
:=
(
lty_ref_shr
A
)
(
at
level
10
)
:
lty_scope
.
Notation
"'ref_shr' A"
:=
(
lty_ref_shr
A
)
(
at
level
10
)
:
lty_scope
.
...
@@ -123,25 +121,20 @@ Section term_types.
...
@@ -123,25 +121,20 @@ Section term_types.
Proof
.
solve_proper
.
Qed
.
Proof
.
solve_proper
.
Qed
.
Global
Instance
lty_forall_contractive
`{
heapG
Σ
}
k
n
:
Global
Instance
lty_forall_contractive
`{
heapG
Σ
}
k
n
:
Proper
(
dist
n
==>
dist
n
==>
pointwise_relation
_
(
dist_later
n
)
==>
dist
n
)
Proper
(
pointwise_relation
_
(
dist_later
n
)
==>
dist
n
)
(
@
lty_forall
Σ
_
k
)
.
(
@
lty_forall
Σ
_
k
)
.
Proof
.
Proof
.
intros
Ml
Ml'
Hl
Mu
Mu'
Hu
C
C'
HC
w
.
apply
Ltty_ne
=>
v
.
f_equiv
=>
M
.
intros
F
F'
A
.
apply
Ltty_ne
=>
w
.
f_equiv
=>
B
.
do
2
(
f_equiv
;
[
by
f_equiv
|])
.
apply
(
wp_contractive
_
_
_
_
_)=>
u
.
specialize
(
A
B
)
.
apply
(
wp_contractive
_
_
_
_
_)=>
u
.
specialize
(
HC
M
)
.
by
destruct
n
as
[|
n
];
simpl
.
by
destruct
n
as
[|
n
];
simpl
.
Qed
.
Qed
.
Global
Instance
lty_forall_ne
`{
heapG
Σ
}
k
n
:
Global
Instance
lty_forall_ne
`{
heapG
Σ
}
k
n
:
Proper
(
dist
n
==>
dist
n
==>
pointwise_relation
_
(
dist
n
)
==>
dist
n
)
Proper
(
pointwise_relation
_
(
dist
n
)
==>
dist
n
)
(
@
lty_forall
Σ
_
k
)
.
(
@
lty_forall
Σ
_
k
)
.
Proof
.
solve_proper
.
Qed
.
Proof
.
solve_proper
.
Qed
.
Global
Instance
lty_exist_contractive
k
n
:
Global
Instance
lty_exist_contractive
k
n
:
Proper
(
dist
n
==>
dist
n
==>
pointwise_relation
_
(
dist_later
n
)
==>
dist
n
)
Proper
(
pointwise_relation
_
(
dist_later
n
)
==>
dist
n
)
(
@
lty_exist
Σ
k
)
.
(
@
lty_exist
Σ
k
)
.
Proof
.
solve_contractive
.
Qed
.
Proof
.
solve_contractive
.
Qed
.
Global
Instance
lty_exist_ne
k
n
:
Global
Instance
lty_exist_ne
k
n
:
Proper
(
dist
n
==>
dist
n
==>
pointwise_relation
_
(
dist
n
)
==>
dist
n
)
Proper
(
pointwise_relation
_
(
dist
n
)
==>
dist
n
)
(
@
lty_exist
Σ
k
)
.
(
@
lty_exist
Σ
k
)
.
Proof
.
solve_proper
.
Qed
.
Proof
.
solve_proper
.
Qed
.
Global
Instance
lty_ref_mut_contractive
`{
heapG
Σ
}
:
Contractive
lty_ref_mut
.
Global
Instance
lty_ref_mut_contractive
`{
heapG
Σ
}
:
Contractive
lty_ref_mut
.
...
...
This diff is collapsed.
Click to expand it.
theories/logrel/term_typing_rules.v
+
20
−
16
View file @
a2df6db0
...
@@ -202,39 +202,43 @@ Section properties.
...
@@ -202,39 +202,43 @@ Section properties.
Qed
.
Qed
.
(** Universal Properties *)
(** Universal Properties *)
Lemma
ltyped_tlam
Γ
e
k
Mlow
Mup
(
C
:
lty
Σ
k
→
ltty
Σ
)
:
Lemma
ltyped_tlam
Γ
e
k
(
C
:
lty
Σ
k
→
ltty
Σ
)
:
(
∀
M
,
Mlow
<:
M
-∗
M
<:
Mup
-∗
Γ
⊨
e
:
C
M
⫤
∅
)
-∗
(
∀
M
,
Γ
⊨
e
:
C
M
⫤
∅
)
-∗
Γ
⊨
(
λ
:
<>
,
e
)
:
∀
M
,
C
M
⫤
∅.
Γ
⊨
(
λ
:
<>
,
e
)
:
lty_forall
Mlow
Mup
C
⫤
∅.
Proof
.
Proof
.
iIntros
"#He"
(
vs
)
"!> HΓ /="
.
wp_pures
.
iIntros
"#He"
(
vs
)
"!> HΓ /="
.
wp_pures
.
iSplitL
;
last
by
iApply
env_ltyped_empty
.
iSplitL
;
last
by
iApply
env_ltyped_empty
.
iIntros
(
M
)
"/=
Hlow Hup
"
.
wp_pures
.
iIntros
(
M
)
"/="
.
wp_pures
.
iApply
(
wp_wand
with
"(He
Hlow Hup
HΓ)"
)
.
iIntros
(
v
)
"[$ _]"
.
iApply
(
wp_wand
with
"(He HΓ)"
)
.
iIntros
(
v
)
"[$ _]"
.
Qed
.
Qed
.
Lemma
ltyped_tapp
Γ
Γ2
e
k
(
C
:
lty
Σ
k
→
ltty
Σ
)
Mlow
Mup
M
:
Lemma
ltyped_tapp
Γ
Γ2
e
k
(
C
:
lty
Σ
k
→
ltty
Σ
)
M
:
Mlow
<:
M
-∗
M
<:
Mup
-∗
(
Γ
⊨
e
:
∀
M
,
C
M
⫤
Γ2
)
-∗
Γ
⊨
e
#
()
:
C
M
⫤
Γ2
.
(
Γ
⊨
e
:
lty_forall
Mlow
Mup
C
⫤
Γ2
)
-∗
Γ
⊨
e
#
()
:
C
M
⫤
Γ2
.
Proof
.
Proof
.
iIntros
"
#Hlow #Hup
#He"
(
vs
)
"!> HΓ /="
.
iIntros
"#He"
(
vs
)
"!> HΓ /="
.
wp_apply
(
wp_wand
with
"(He [HΓ //])"
);
iIntros
(
w
)
"[HB HΓ] /="
.
wp_apply
(
wp_wand
with
"(He [HΓ //])"
);
iIntros
(
w
)
"[HB HΓ] /="
.
iApply
(
wp_wand
with
"
(
HB
Hlow Hup)
[HΓ]"
)
.
by
iIntros
(
v
)
"$"
.
iApply
(
wp_wand
with
"HB [HΓ]"
)
.
by
iIntros
(
v
)
"$"
.
Qed
.
Qed
.
(** Existential properties *)
(** Existential properties *)
Lemma
ltyped_pack
Γ
e
k
(
C
:
lty
Σ
k
→
ltty
Σ
)
M
:
(
Γ
⊨
e
:
C
M
)
-∗
Γ
⊨
e
:
∃
M
,
C
M
.
Proof
.
iIntros
"#He"
(
vs
)
"!> HΓ /="
.
wp_apply
(
wp_wand
with
"(He [HΓ //])"
);
iIntros
(
w
)
"[HB $]"
.
by
iExists
M
.
Qed
.
Definition
unpack
:
val
:=
λ
:
"exist"
"f"
,
"f"
#
()
"exist"
.
Definition
unpack
:
val
:=
λ
:
"exist"
"f"
,
"f"
#
()
"exist"
.
Lemma
ltyped_unpack
k
(
C
:
lty
Σ
k
→
ltty
Σ
)
Mlow
Mup
B
:
Lemma
ltyped_unpack
k
(
C
:
lty
Σ
k
→
ltty
Σ
)
B
:
⊢
∅
⊨
unpack
:
lty_exist
Mlow
Mup
C
→
lty_forall
Mlow
Mup
(
λ
M
,
C
M
⊸
B
)
%
lty
⊸
B
.
⊢
∅
⊨
unpack
:
(
∃
M
,
C
M
)
→
(
∀
M
,
C
M
⊸
B
)
⊸
B
.
Proof
.
Proof
.
iIntros
(
vs
)
"!> HΓ /="
.
iApply
wp_value
.
iIntros
(
vs
)
"!> HΓ /="
.
iApply
wp_value
.
iSplitL
;
last
by
iApply
env_ltyped_empty
.
iSplitL
;
last
by
iApply
env_ltyped_empty
.
iIntros
"!>"
(
v
)
.
iDestruct
1
as
(
M
)
"
(Hlow & Hup &
Hv
)
"
.
iIntros
"!>"
(
v
)
.
iDestruct
1
as
(
M
)
"Hv"
.
rewrite
/
unpack
.
wp_pures
.
rewrite
/
unpack
.
wp_pures
.
iIntros
(
fty
)
"Hfty"
.
wp_pures
.
iIntros
(
fty
)
"Hfty"
.
wp_pures
.
iSpecialize
(
"Hfty"
$!
M
)
.
iSpecialize
(
"Hfty"
$!
M
)
.
wp_bind
(
fty
_)
.
wp_apply
(
wp_wand
with
"
(
Hfty
Hlow Hup)
"
)
.
wp_bind
(
fty
_)
.
wp_apply
(
wp_wand
with
"Hfty"
)
.
iIntros
(
f
)
"Hf"
.
iIntros
(
f
)
"Hf"
.
wp_apply
(
wp_wand
with
"(Hf [Hv //])"
)
.
wp_apply
(
wp_wand
with
"(Hf [Hv //])"
)
.
iIntros
(
w
)
"HB"
.
iApply
"HB"
.
iIntros
(
w
)
"HB"
.
iApply
"HB"
.
...
@@ -260,7 +264,7 @@ Section properties.
...
@@ -260,7 +264,7 @@ Section properties.
longer use the memory at the old location. *)
longer use the memory at the old location. *)
Definition
load
:
val
:=
λ
:
"r"
,
(
!
"r"
,
"r"
)
.
Definition
load
:
val
:=
λ
:
"r"
,
(
!
"r"
,
"r"
)
.
Lemma
ltyped_load
A
:
Lemma
ltyped_load
A
:
⊢
∅
⊨
load
:
ref_mut
A
→
A
*
ref_mut
⊤
.
⊢
∅
⊨
load
:
ref_mut
A
→
A
*
ref_mut
any
.
Proof
.
Proof
.
iIntros
(
vs
)
"!> HΓ /="
.
iApply
wp_value
.
iIntros
(
vs
)
"!> HΓ /="
.
iApply
wp_value
.
iSplitL
;
last
by
iApply
env_ltyped_empty
.
iSplitL
;
last
by
iApply
env_ltyped_empty
.
...
...
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