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
Iris
Iris
Commits
c02ea520
Commit
c02ea520
authored
Jan 27, 2016
by
Ralf Jung
Browse files
prove that we can pick a fresh location
parent
f71e526a
Changes
3
Hide whitespace changes
Inline
Side-by-side
barrier/heap_lang.v
View file @
c02ea520
Require
Import
Autosubst
.
Autosubst
.
Require
Import
prelude
.
option
prelude
.
gmap
iris
.
language
.
(** Some tactics useful when dealing with equality of sigma-like types: existT T0 t0 = existT T1 t1.
(** Some tactics useful when dealing with equality of sigma-like types:
existT T0 t0 = existT T1 t1.
They all assume such an equality is the first thing on the "stack" (goal). *)
Ltac
case_depeq1
:
=
let
Heq
:
=
fresh
"Heq"
in
case
=>
_
/
EqdepFacts
.
eq_sigT_sig_eq
=>
Heq
;
...
...
@@ -20,7 +21,7 @@ Ltac case_depeq3 := let Heq := fresh "Heq" in
destruct
Heq
as
(->,<-).
(** Expressions and values. *)
Definition
loc
:
=
nat
.
(* A
ny countable type. *)
Definition
loc
:
=
positive
.
(* Really, a
ny countable type. *)
Inductive
expr
:
=
(* Base lambda calculus *)
...
...
barrier/lifting.v
View file @
c02ea520
...
...
@@ -56,7 +56,9 @@ Proof.
-
intros
e2
σ
2
ef
Hstep
.
inversion_clear
Hstep
.
split
;
first
done
.
rewrite
v2v
in
Hv
.
inversion_clear
Hv
.
eexists
;
split_ands
;
done
.
-
(* RJ FIXME: Need to find a fresh location. *)
admit
.
-
set
(
l
:
=
fresh
$
dom
(
gset
loc
)
σ
).
exists
(
Loc
l
),
((<[
l
:
=
v
]>)
σ
),
None
.
eapply
AllocS
;
first
by
rewrite
v2v
.
apply
(
not_elem_of_dom
(
D
:
=
gset
loc
)).
apply
is_fresh
.
-
reflexivity
.
-
reflexivity
.
-
(* RJ FIXME I am sure there is a better way to invoke right_id, but I could not find it. *)
...
...
@@ -69,7 +71,7 @@ Proof.
erewrite
<-
exist_intro
with
(
a
:
=
l
).
apply
and_intro
.
+
by
apply
const_intro
.
+
done
.
Abort
.
Qed
.
Lemma
wp_load
E
σ
l
v
:
σ
!!
l
=
Some
v
→
...
...
prelude/gmap.v
View file @
c02ea520
...
...
@@ -117,3 +117,20 @@ Notation gset K := (mapset (gmap K)).
Instance
gset_dom
`
{
Countable
K
}
{
A
}
:
Dom
(
gmap
K
A
)
(
gset
K
)
:
=
mapset_dom
.
Instance
gset_dom_spec
`
{
Countable
K
}
:
FinMapDom
K
(
gmap
K
)
(
gset
K
)
:
=
mapset_dom_spec
.
(** * Fresh positive *)
Definition
Gfresh
{
A
}
(
m
:
gmap
positive
A
)
:
positive
:
=
Pfresh
$
gmap_car
m
.
Lemma
Gfresh_fresh
{
A
}
(
m
:
gmap
positive
A
)
:
m
!!
Gfresh
m
=
None
.
Proof
.
destruct
m
as
[[]].
apply
Pfresh_fresh
;
done
.
Qed
.
Instance
Gset_fresh
:
Fresh
positive
(
gset
positive
)
:
=
λ
X
,
let
(
m
)
:
=
X
in
Gfresh
m
.
Instance
Gset_fresh_spec
:
FreshSpec
positive
(
gset
positive
).
Proof
.
split
.
*
apply
_
.
*
intros
X
Y
;
rewrite
<-
elem_of_equiv_L
.
by
intros
->.
*
unfold
elem_of
,
mapset_elem_of
,
fresh
;
intros
[
m
]
;
simpl
.
by
rewrite
Gfresh_fresh
.
Qed
.
Robbert Krebbers
@robbertkrebbers
mentioned in commit
9c4b7e80
·
Feb 01, 2016
mentioned in commit
9c4b7e80
mentioned in commit 9c4b7e80165a3a61c4541d5aafb8187462c8b81a
Toggle commit list
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