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
George Pirlea
Iris
Commits
d4a687a8
Commit
d4a687a8
authored
Nov 20, 2016
by
Robbert Krebbers
Browse files
Prove (■ φ → P) ⊣⊢ (∀ _ : φ, P).
parent
588b9559
Changes
1
Show whitespace changes
Inline
Side-by-side
base_logic/derived.v
View file @
d4a687a8
...
...
@@ -265,6 +265,12 @@ Proof. by intros ->. Qed.
Lemma
internal_eq_sym
{
A
:
cofeT
}
(
a
b
:
A
)
:
a
≡
b
⊢
b
≡
a
.
Proof
.
apply
(
internal_eq_rewrite
a
b
(
λ
b
,
b
≡
a
)%
I
)
;
auto
.
solve_proper
.
Qed
.
Lemma
pure_impl_forall
φ
P
:
(
■
φ
→
P
)
⊣
⊢
(
∀
_
:
φ
,
P
).
Proof
.
apply
(
anti_symm
_
).
-
apply
forall_intro
=>
?.
by
rewrite
pure_equiv
//
left_id
.
-
apply
impl_intro_l
,
pure_elim_l
=>
H
φ
.
by
rewrite
(
forall_elim
H
φ
).
Qed
.
Lemma
pure_alt
φ
:
■
φ
⊣
⊢
∃
_
:
φ
,
True
.
Proof
.
apply
(
anti_symm
_
).
...
...
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