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
5b36e201
Commit
5b36e201
authored
Feb 12, 2016
by
Robbert Krebbers
Browse files
More general and Global version of auth_timeless.
parent
56c5b62d
Changes
2
Hide whitespace changes
Inline
Side-by-side
algebra/auth.v
View file @
5b36e201
...
...
@@ -50,8 +50,8 @@ Proof.
apply
(
conv_compl
(
chain_map
own
c
)
n
).
Qed
.
Canonical
Structure
authC
:
=
CofeT
auth_cofe_mixin
.
Instance
A
uth_timeless
(
ea
:
excl
A
)
(
b
:
A
)
:
Timeless
ea
→
Timeless
b
→
Timeless
(
Auth
ea
b
)
.
Global
Instance
a
uth_timeless
(
x
:
auth
A
)
:
Timeless
(
authoritative
x
)
→
Timeless
(
own
x
)
→
Timeless
x
.
Proof
.
by
intros
??
[??]
[??]
;
split
;
simpl
in
*
;
apply
(
timeless
_
).
Qed
.
Global
Instance
auth_leibniz
:
LeibnizEquiv
A
→
LeibnizEquiv
(
auth
A
).
Proof
.
by
intros
?
[??]
[??]
[??]
;
f_equal'
;
apply
leibniz_equiv
.
Qed
.
...
...
@@ -140,7 +140,7 @@ Proof.
split
;
simpl
.
*
by
apply
(@
cmra_empty_valid
A
_
).
*
by
intros
x
;
constructor
;
rewrite
/=
left_id
.
*
apply
Auth_timeless
;
apply
_
.
*
apply
_
.
Qed
.
Lemma
auth_frag_op
a
b
:
◯
(
a
⋅
b
)
≡
◯
a
⋅
◯
b
.
Proof
.
done
.
Qed
.
...
...
program_logic/auth.v
View file @
5b36e201
...
...
@@ -11,14 +11,6 @@ Section auth.
Implicit
Types
a
b
:
A
.
Implicit
Types
γ
:
gname
.
(* Adding this locally only, since it overlaps with Auth_timelss in algebra/auth.v.
TODO: Would moving this to auth.v and making it global break things? *)
Local
Instance
AuthA_timeless
(
x
:
auth
A
)
:
Timeless
x
.
Proof
.
(* FIXME: "destruct x; auto with typeclass_instances" should find this through Auth, right? *)
destruct
x
.
apply
Auth_timeless
;
apply
_
.
Qed
.
(* TODO: Need this to be proven somewhere. *)
Hypothesis
auth_valid
:
forall
a
b
,
(
✓
Auth
(
Excl
a
)
b
:
iPropG
Λ
Σ
)
⊑
(
∃
b'
,
a
≡
b
⋅
b'
).
...
...
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