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
Janno
iris-coq
Commits
9379001a
Commit
9379001a
authored
Oct 22, 2016
by
Robbert Krebbers
Browse files
Avoid naming a type class constraint.
parent
58bd4b02
Changes
1
Show whitespace changes
Inline
Side-by-side
program_logic/ghost_ownership.v
View file @
9379001a
...
@@ -31,7 +31,7 @@ Typeclasses Opaque own.
...
@@ -31,7 +31,7 @@ Typeclasses Opaque own.
(** * Properties about ghost ownership *)
(** * Properties about ghost ownership *)
Section
global
.
Section
global
.
Context
`
{
i
:
inG
Σ
A
}.
Context
`
{
inG
Σ
A
}.
Implicit
Types
a
:
A
.
Implicit
Types
a
:
A
.
(** ** Properties of [iRes_singleton] *)
(** ** Properties of [iRes_singleton] *)
...
@@ -63,7 +63,7 @@ Proof. intros a1 a2. apply own_mono. Qed.
...
@@ -63,7 +63,7 @@ Proof. intros a1 a2. apply own_mono. Qed.
Lemma
own_valid
γ
a
:
own
γ
a
⊢
✓
a
.
Lemma
own_valid
γ
a
:
own
γ
a
⊢
✓
a
.
Proof
.
Proof
.
rewrite
!
own_eq
/
own_def
ownM_valid
/
iRes_singleton
.
rewrite
!
own_eq
/
own_def
ownM_valid
/
iRes_singleton
.
rewrite
iprod_validI
(
forall_elim
(
inG_id
i
))
iprod_lookup_singleton
.
rewrite
iprod_validI
(
forall_elim
(
inG_id
_
))
iprod_lookup_singleton
.
rewrite
gmap_validI
(
forall_elim
γ
)
lookup_singleton
option_validI
.
rewrite
gmap_validI
(
forall_elim
γ
)
lookup_singleton
option_validI
.
(* implicit arguments differ a bit *)
(* implicit arguments differ a bit *)
by
trans
(
✓
cmra_transport
inG_prf
a
:
iProp
Σ
)%
I
;
last
destruct
inG_prf
.
by
trans
(
✓
cmra_transport
inG_prf
a
:
iProp
Σ
)%
I
;
last
destruct
inG_prf
.
...
@@ -92,7 +92,7 @@ Proof.
...
@@ -92,7 +92,7 @@ Proof.
intros
Ha
.
intros
Ha
.
rewrite
-(
bupd_mono
(
∃
m
,
■
(
∃
γ
,
γ
∉
G
∧
m
=
iRes_singleton
γ
a
)
∧
uPred_ownM
m
)%
I
).
rewrite
-(
bupd_mono
(
∃
m
,
■
(
∃
γ
,
γ
∉
G
∧
m
=
iRes_singleton
γ
a
)
∧
uPred_ownM
m
)%
I
).
-
rewrite
ownM_empty
.
-
rewrite
ownM_empty
.
eapply
bupd_ownM_updateP
,
(
iprod_singleton_updateP_empty
(
inG_id
i
))
;
eapply
bupd_ownM_updateP
,
(
iprod_singleton_updateP_empty
(
inG_id
_
))
;
first
(
eapply
alloc_updateP_strong'
,
cmra_transport_valid
,
Ha
)
;
first
(
eapply
alloc_updateP_strong'
,
cmra_transport_valid
,
Ha
)
;
naive_solver
.
naive_solver
.
-
apply
exist_elim
=>
m
;
apply
pure_elim_l
=>-[
γ
[
Hfresh
->]].
-
apply
exist_elim
=>
m
;
apply
pure_elim_l
=>-[
γ
[
Hfresh
->]].
...
...
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