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
1190a321
Commit
1190a321
authored
Apr 25, 2016
by
Robbert Krebbers
Browse files
Prove False_elim in logic instead of model.
parent
0fccdf33
Changes
1
Hide whitespace changes
Inline
Side-by-side
algebra/upred.v
View file @
1190a321
...
...
@@ -467,8 +467,6 @@ Lemma const_elim φ Q R : Q ⊢ ■ φ → (φ → Q ⊢ R) → Q ⊢ R.
Proof
.
unseal
;
intros
HQP
HQR
;
split
=>
n
x
??
;
apply
HQR
;
first
eapply
HQP
;
eauto
.
Qed
.
Lemma
False_elim
P
:
False
⊢
P
.
Proof
.
by
unseal
;
split
=>
n
x
?.
Qed
.
Lemma
and_elim_l
P
Q
:
(
P
∧
Q
)
⊢
P
.
Proof
.
by
unseal
;
split
=>
n
x
?
[??].
Qed
.
Lemma
and_elim_r
P
Q
:
(
P
∧
Q
)
⊢
Q
.
...
...
@@ -512,6 +510,8 @@ Proof.
Qed
.
(* Derived logical stuff *)
Lemma
False_elim
P
:
False
⊢
P
.
Proof
.
by
apply
(
const_elim
False
).
Qed
.
Lemma
True_intro
P
:
P
⊢
True
.
Proof
.
by
apply
const_intro
.
Qed
.
Lemma
and_elim_l'
P
Q
R
:
P
⊢
R
→
(
P
∧
Q
)
⊢
R
.
...
...
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