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
Joshua Yanovski
iris-coq
Commits
05817be7
Commit
05817be7
authored
May 19, 2016
by
Ralf Jung
Browse files
show that we can rewrite below a contractive function even if we have an equality only later
parent
b2a210fe
Changes
1
Hide whitespace changes
Inline
Side-by-side
algebra/upred.v
View file @
05817be7
...
...
@@ -497,10 +497,11 @@ Proof. unseal; intros HΦΨ; split=> n x ? [a ?]; by apply HΦΨ with a. Qed.
Lemma
eq_refl
{
A
:
cofeT
}
(
a
:
A
)
:
True
⊢
(
a
≡
a
).
Proof
.
unseal
;
by
split
=>
n
x
??
;
simpl
.
Qed
.
Lemma
eq_rewrite
{
A
:
cofeT
}
a
b
(
Ψ
:
A
→
uPred
M
)
P
`
{
H
Ψ
:
∀
n
,
Proper
(
dist
n
==>
dist
n
)
Ψ
}
:
P
⊢
(
a
≡
b
)
→
P
⊢
Ψ
a
→
P
⊢
Ψ
b
.
{
H
Ψ
:
∀
n
,
Proper
(
dist
n
==>
dist
n
)
Ψ
}
:
P
⊢
(
a
≡
b
)
→
P
⊢
Ψ
a
→
P
⊢
Ψ
b
.
Proof
.
unseal
;
intros
Hab
Ha
;
split
=>
n
x
??
.
apply
H
Ψ
with
n
a
;
auto
.
by
symmetry
;
apply
Hab
with
x
.
by
apply
Ha
.
unseal
;
intros
Hab
Ha
;
split
=>
n
x
??
.
apply
H
Ψ
with
n
a
;
auto
.
-
by
symmetry
;
apply
Hab
with
x
.
-
by
apply
Ha
.
Qed
.
Lemma
eq_equiv
`
{
Empty
M
,
!
CMRAUnit
M
}
{
A
:
cofeT
}
(
a
b
:
A
)
:
True
⊢
(
a
≡
b
)
→
a
≡
b
.
...
...
@@ -508,6 +509,14 @@ Proof.
unseal
=>
Hab
;
apply
equiv_dist
;
intros
n
;
apply
Hab
with
∅
;
last
done
.
apply
cmra_valid_validN
,
cmra_unit_valid
.
Qed
.
Lemma
eq_rewrite_contractive
{
A
:
cofeT
}
a
b
(
Ψ
:
A
→
uPred
M
)
P
{
H
Ψ
:
Contractive
Ψ
}
:
P
⊢
▷
(
a
≡
b
)
→
P
⊢
Ψ
a
→
P
⊢
Ψ
b
.
Proof
.
unseal
;
intros
Hab
Ha
;
split
=>
n
x
??
.
apply
H
Ψ
with
n
a
;
auto
.
-
destruct
n
;
intros
m
?
;
first
omega
.
apply
(
dist_le
n
);
last
omega
.
symmetry
.
by
destruct
Hab
as
[
Hab
];
eapply
(
Hab
(
S
n
)).
-
by
apply
Ha
.
Qed
.
(
*
Derived
logical
stuff
*
)
Lemma
False_elim
P
:
False
⊢
P
.
...
...
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