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
Simon Spies
Iris
Commits
51e2363b
Commit
51e2363b
authored
Jun 04, 2019
by
Robbert Krebbers
Browse files
Move `array` out of section to avoid duplicating the notation.
parent
9f34cfe5
Changes
1
Hide whitespace changes
Inline
Side-by-side
theories/heap_lang/lifting.v
View file @
51e2363b
...
...
@@ -32,6 +32,11 @@ Notation "l ↦{ q } -" := (∃ v, l ↦{q} v)%I
(
at
level
20
,
q
at
level
50
,
format
"l ↦{ q } -"
)
:
bi_scope
.
Notation
"l ↦ -"
:
=
(
l
↦
{
1
}
-)%
I
(
at
level
20
)
:
bi_scope
.
Definition
array
`
{!
heapG
Σ
}
(
l
:
loc
)
(
vs
:
list
val
)
:
iProp
Σ
:
=
([
∗
list
]
i
↦
v
∈
vs
,
loc_add
l
i
↦
v
)%
I
.
Notation
"l ↦∗ vs"
:
=
(
array
l
vs
)
(
at
level
20
,
format
"l ↦∗ vs"
)
:
bi_scope
.
(** The tactic [inv_head_step] performs inversion on hypotheses of the shape
[head_step]. The tactic will discharge head-reductions starting from values, and
simplifies hypothesis related to conversions from and to values, and finite map
...
...
@@ -207,12 +212,6 @@ Proof.
iIntros
(
κ
v2
σ
2
efs
Hstep
)
;
inv_head_step
.
by
iFrame
.
Qed
.
Definition
array
(
l
:
loc
)
(
vs
:
list
val
)
:
iProp
Σ
:
=
([
∗
list
]
i
↦
v
∈
vs
,
loc_add
l
i
↦
v
)%
I
.
Notation
"l ↦∗ vs"
:
=
(
array
l
vs
)
(
at
level
20
,
format
"l ↦∗ vs"
)
:
bi_scope
.
Lemma
array_nil
l
:
l
↦∗
[]
⊣
⊢
emp
.
Proof
.
by
rewrite
/
array
.
Qed
.
...
...
@@ -504,6 +503,3 @@ Proof.
Qed
.
End
lifting
.
Notation
"l ↦∗ vs"
:
=
(
array
l
vs
)
(
at
level
20
,
format
"l ↦∗ vs"
)
:
bi_scope
.
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