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
Rice Wine
Iris
Commits
a953a68d
Commit
a953a68d
authored
Dec 13, 2016
by
Ralf Jung
Browse files
agree: prove non-step-indexed uninjection
parent
766dbcd2
Changes
1
Hide whitespace changes
Inline
Side-by-side
theories/algebra/agree.v
View file @
a953a68d
...
...
@@ -354,7 +354,7 @@ Proof.
intros
a
b
?.
apply
equiv_dist
=>
n
.
apply
to_agree_injN
.
by
apply
equiv_dist
.
Qed
.
Lemma
to_agree_uninj
n
(
x
:
agree
A
)
:
✓
{
n
}
x
→
∃
y
:
A
,
to_agree
y
≡
{
n
}
≡
x
.
Lemma
to_agree_uninj
N
n
(
x
:
agree
A
)
:
✓
{
n
}
x
→
∃
y
:
A
,
to_agree
y
≡
{
n
}
≡
x
.
Proof
.
intros
Hl
.
exists
(
agree_car
x
).
rewrite
/
dist
/
agree_dist
/=.
split
.
...
...
@@ -362,6 +362,16 @@ Proof.
-
apply
(
list_agrees_iff_setincl
_
)
;
first
set_solver
+.
done
.
Qed
.
Lemma
to_agree_uninj
(
x
:
agree
A
)
:
✓
x
→
∃
y
:
A
,
to_agree
y
≡
x
.
Proof
.
intros
Hl
.
exists
(
agree_car
x
).
rewrite
/
dist
/
agree_dist
/=.
split
.
-
apply
:
list_setincl_singleton_in
.
set_solver
+.
-
apply
(
list_agrees_iff_setincl
_
)
;
first
set_solver
+.
eapply
list_agrees_subrel
;
last
exact
:
Hl
;
[
apply
_
..|].
intros
???.
by
apply
equiv_dist
.
Qed
.
Lemma
to_agree_included
(
a
b
:
A
)
:
to_agree
a
≼
to_agree
b
↔
a
≡
b
.
Proof
.
split
.
...
...
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