Skip to content
GitLab
Projects
Groups
Snippets
/
Help
Help
Support
Community forum
Keyboard shortcuts
?
Submit feedback
Contribute to GitLab
Sign in / Register
Toggle navigation
Menu
Open sidebar
Tej Chajed
iris
Commits
42cf780a
Commit
42cf780a
authored
Nov 25, 2016
by
Jacques-Henri Jourdan
Browse files
Stronger splitting rule for boxes.
parent
26ae0c67
Changes
1
Hide whitespace changes
Inline
Side-by-side
base_logic/lib/boxes.v
View file @
42cf780a
...
...
@@ -233,22 +233,24 @@ Qed.
Lemma
box_split
E
f
P
Q1
Q2
γ
b
:
↑
N
⊆
E
→
f
!!
γ
=
Some
b
→
slice
N
γ
(
Q1
∗
Q2
)
-
∗
▷
box
N
f
P
={
E
}=
∗
∃
γ
1
γ
2
,
⌜
delete
γ
f
!!
γ
1
=
None
⌝
∗
⌜
delete
γ
f
!!
γ
2
=
None
⌝
∗
⌜
delete
γ
f
!!
γ
1
=
None
⌝
∗
⌜
delete
γ
f
!!
γ
2
=
None
⌝
∗
⌜γ
1
≠
γ
2
⌝
∗
slice
N
γ
1
Q1
∗
slice
N
γ
2
Q2
∗
▷
box
N
(<[
γ
2
:
=
b
]>(<[
γ
1
:
=
b
]>(
delete
γ
f
)))
P
.
Proof
.
iIntros
(??)
"#Hslice Hbox"
.
destruct
b
.
-
iMod
(
box_delete_full
with
"Hslice Hbox"
)
as
(
P'
)
"([HQ1 HQ2] & Heq & Hbox)"
;
try
done
.
iMod
(
box_insert_full
Q1
with
"HQ1 Hbox"
)
as
(
γ
1
)
"(% & #Hslice1 & Hbox)"
.
done
.
iMod
(
box_insert_full
Q2
with
"HQ2 Hbox"
)
as
(
γ
2
)
"(% & #Hslice2 & Hbox)"
.
done
.
iExists
γ
1
,
γ
2
.
iFrame
"%#"
.
iModIntro
.
iSplit
.
{
iPureIntro
.
by
eapply
lookup_insert_None
.
}
iExists
γ
1
,
γ
2
.
iFrame
"%#"
.
iModIntro
.
iSplit
;
last
iSplit
;
try
iPureIntro
.
{
by
eapply
lookup_insert_None
.
}
{
by
apply
(
lookup_insert_None
(
delete
γ
f
)
γ
1
γ
2
true
).
}
iNext
.
eapply
internal_eq_rewrite_contractive
;
[
by
apply
_
|
|
by
eauto
].
iNext
.
iRewrite
"Heq"
.
iPureIntro
.
rewrite
assoc
.
f_equiv
.
by
rewrite
comm
.
done
.
-
iMod
(
box_delete_empty
with
"Hslice Hbox"
)
as
(
P'
)
"[Heq Hbox]"
;
try
done
.
iMod
(
box_insert_empty
Q1
with
"Hbox"
)
as
(
γ
1
)
"(% & #Hslice1 & Hbox)"
.
iMod
(
box_insert_empty
Q2
with
"Hbox"
)
as
(
γ
2
)
"(% & #Hslice2 & Hbox)"
.
iExists
γ
1
,
γ
2
.
iFrame
"%#"
.
iModIntro
.
iSplit
.
{
iPureIntro
.
by
eapply
lookup_insert_None
.
}
iExists
γ
1
,
γ
2
.
iFrame
"%#"
.
iModIntro
.
iSplit
;
last
iSplit
;
try
iPureIntro
.
{
by
eapply
lookup_insert_None
.
}
{
by
apply
(
lookup_insert_None
(
delete
γ
f
)
γ
1
γ
2
false
).
}
iNext
.
eapply
internal_eq_rewrite_contractive
;
[
by
apply
_
|
|
by
eauto
].
iNext
.
iRewrite
"Heq"
.
iPureIntro
.
rewrite
assoc
.
f_equiv
.
by
rewrite
comm
.
done
.
Qed
.
...
...
Write
Preview
Supports
Markdown
0%
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!
Cancel
Please
register
or
sign in
to comment