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
Dmitry Khalanskiy
Iris
Commits
8fd8da22
Commit
8fd8da22
authored
Nov 21, 2016
by
Robbert Krebbers
Browse files
Rename pure_equiv -> pure_True.
The old name didn't make much sense. Also now we can have pure_False too.
parent
863f30c0
Changes
5
Hide whitespace changes
Inline
Side-by-side
base_logic/big_op.v
View file @
8fd8da22
...
...
@@ -260,7 +260,7 @@ Section list.
revert
Φ
H
Φ
.
induction
l
as
[|
x
l
IH
]=>
Φ
H
Φ
.
{
rewrite
big_sepL_nil
;
auto
with
I
.
}
rewrite
big_sepL_cons
.
rewrite
-
always_and_sep_l
;
apply
and_intro
.
-
by
rewrite
(
forall_elim
0
)
(
forall_elim
x
)
pure_
equiv
//
True_impl
.
-
by
rewrite
(
forall_elim
0
)
(
forall_elim
x
)
pure_
True
//
True_impl
.
-
rewrite
-
IH
.
apply
forall_intro
=>
k
;
by
rewrite
(
forall_elim
(
S
k
)).
Qed
.
...
...
@@ -384,11 +384,11 @@ Section gmap.
induction
m
as
[|
i
x
m
?
IH
]
using
map_ind
;
[
rewrite
?big_sepM_empty
;
auto
|].
rewrite
big_sepM_insert
//
-
always_and_sep_l
.
apply
and_intro
.
-
rewrite
(
forall_elim
i
)
(
forall_elim
x
)
lookup_insert
.
by
rewrite
pure_
equiv
//
True_impl
.
by
rewrite
pure_
True
//
True_impl
.
-
rewrite
-
IH
.
apply
forall_mono
=>
k
;
apply
forall_mono
=>
y
.
apply
impl_intro_l
,
pure_elim_l
=>
?.
rewrite
lookup_insert_ne
;
last
by
intros
?
;
simplify_map_eq
.
by
rewrite
pure_
equiv
//
True_impl
.
by
rewrite
pure_
True
//
True_impl
.
Qed
.
Lemma
big_sepM_impl
Φ
Ψ
m
:
...
...
@@ -492,12 +492,12 @@ Section gset.
intros
.
apply
(
anti_symm
_
).
{
apply
forall_intro
=>
x
.
apply
impl_intro_l
,
pure_elim_l
=>
?
;
by
apply
big_sepS_elem_of
.
}
induction
X
as
[|
x
m
?
IH
]
using
collection_ind_L
.
induction
X
as
[|
x
X
?
IH
]
using
collection_ind_L
.
{
rewrite
big_sepS_empty
;
auto
.
}
rewrite
big_sepS_insert
//
-
always_and_sep_l
.
apply
and_intro
.
-
by
rewrite
(
forall_elim
x
)
pure_
equiv
?True_impl
;
last
set_solver
.
-
by
rewrite
(
forall_elim
x
)
pure_
True
?True_impl
;
last
set_solver
.
-
rewrite
-
IH
.
apply
forall_mono
=>
y
.
apply
impl_intro_l
,
pure_elim_l
=>
?.
by
rewrite
pure_
equiv
?True_impl
;
last
set_solver
.
by
rewrite
pure_
True
?True_impl
;
last
set_solver
.
Qed
.
Lemma
big_sepS_impl
Φ
Ψ
X
:
...
...
base_logic/derived.v
View file @
8fd8da22
...
...
@@ -163,6 +163,11 @@ Proof.
-
by
rewrite
-(
left_id
True
%
I
uPred_and
(
_
→
_
)%
I
)
impl_elim_r
.
-
by
apply
impl_intro_l
;
rewrite
left_id
.
Qed
.
Lemma
False_impl
P
:
(
False
→
P
)
⊣
⊢
True
.
Proof
.
apply
(
anti_symm
(
⊢
))
;
[
by
auto
|].
apply
impl_intro_l
.
rewrite
left_absorb
.
auto
.
Qed
.
Lemma
exists_impl_forall
{
A
}
P
(
Ψ
:
A
→
uPred
M
)
:
((
∃
x
:
A
,
Ψ
x
)
→
P
)
⊣
⊢
∀
x
:
A
,
Ψ
x
→
P
.
...
...
@@ -223,8 +228,11 @@ Lemma pure_elim_l φ Q R : (φ → Q ⊢ R) → ■ φ ∧ Q ⊢ R.
Proof
.
intros
;
apply
pure_elim
with
φ
;
eauto
.
Qed
.
Lemma
pure_elim_r
φ
Q
R
:
(
φ
→
Q
⊢
R
)
→
Q
∧
■
φ
⊢
R
.
Proof
.
intros
;
apply
pure_elim
with
φ
;
eauto
.
Qed
.
Lemma
pure_equiv
(
φ
:
Prop
)
:
φ
→
■
φ
⊣
⊢
True
.
Lemma
pure_True
(
φ
:
Prop
)
:
φ
→
■
φ
⊣
⊢
True
.
Proof
.
intros
;
apply
(
anti_symm
_
)
;
auto
.
Qed
.
Lemma
pure_False
(
φ
:
Prop
)
:
¬
φ
→
■
φ
⊣
⊢
False
.
Proof
.
intros
;
apply
(
anti_symm
_
)
;
eauto
using
pure_elim
.
Qed
.
Lemma
pure_and
φ
1
φ
2
:
■
(
φ
1
∧
φ
2
)
⊣
⊢
■
φ
1
∧
■
φ
2
.
Proof
.
...
...
@@ -243,7 +251,7 @@ Proof.
apply
(
anti_symm
_
).
-
apply
impl_intro_l
.
rewrite
-
pure_and
.
apply
pure_mono
.
naive_solver
.
-
rewrite
-
pure_forall_2
.
apply
forall_intro
=>
?.
by
rewrite
-(
left_id
True
uPred_and
(
_→_
))%
I
(
pure_
equiv
φ
1
)
//
impl_elim_r
.
by
rewrite
-(
left_id
True
uPred_and
(
_→_
))%
I
(
pure_
True
φ
1
)
//
impl_elim_r
.
Qed
.
Lemma
pure_forall
{
A
}
(
φ
:
A
→
Prop
)
:
■
(
∀
x
,
φ
x
)
⊣
⊢
∀
x
,
■
φ
x
.
Proof
.
...
...
@@ -268,7 +276,7 @@ Proof. apply (internal_eq_rewrite a b (λ b, b ≡ a)%I); auto. solve_proper. Qe
Lemma
pure_impl_forall
φ
P
:
(
■
φ
→
P
)
⊣
⊢
(
∀
_
:
φ
,
P
).
Proof
.
apply
(
anti_symm
_
).
-
apply
forall_intro
=>
?.
by
rewrite
pure_
equiv
//
left_id
.
-
apply
forall_intro
=>
?.
by
rewrite
pure_
True
//
left_id
.
-
apply
impl_intro_l
,
pure_elim_l
=>
H
φ
.
by
rewrite
(
forall_elim
H
φ
).
Qed
.
Lemma
pure_alt
φ
:
■
φ
⊣
⊢
∃
_
:
φ
,
True
.
...
...
base_logic/lib/own.v
View file @
8fd8da22
...
...
@@ -95,7 +95,7 @@ Proof.
first
(
eapply
alloc_updateP_strong'
,
cmra_transport_valid
,
Ha
)
;
naive_solver
.
-
apply
exist_elim
=>
m
;
apply
pure_elim_l
=>-[
γ
[
Hfresh
->]].
by
rewrite
!
own_eq
/
own_def
-(
exist_intro
γ
)
pure_
equiv
//
left_id
.
by
rewrite
!
own_eq
/
own_def
-(
exist_intro
γ
)
pure_
True
//
left_id
.
Qed
.
Lemma
own_alloc
a
:
✓
a
→
True
==
∗
∃
γ
,
own
γ
a
.
Proof
.
...
...
heap_lang/heap.v
View file @
8fd8da22
...
...
@@ -88,7 +88,7 @@ Section heap.
l
↦
{
q1
}
v1
∗
l
↦
{
q2
}
v2
⊣
⊢
v1
=
v2
∧
l
↦
{
q1
+
q2
}
v1
.
Proof
.
destruct
(
decide
(
v1
=
v2
))
as
[->|].
{
by
rewrite
heap_mapsto_op_eq
pure_
equiv
//
left_id
.
}
{
by
rewrite
heap_mapsto_op_eq
pure_
True
//
left_id
.
}
apply
(
anti_symm
(
⊢
))
;
last
by
apply
pure_elim_l
.
rewrite
heap_mapsto_eq
-
auth_own_op
auth_own_valid
discrete_valid
.
eapply
pure_elim
;
[
done
|]
=>
/=.
...
...
proofmode/coq_tactics.v
View file @
8fd8da22
...
...
@@ -152,10 +152,10 @@ Proof.
rewrite
/
envs_lookup
/
of_envs
=>?
;
apply
pure_elim_sep_l
=>
Hwf
.
destruct
Δ
as
[
Γ
p
Γ
s
],
(
Γ
p
!!
i
)
eqn
:
?
;
simplify_eq
/=.
-
rewrite
(
env_lookup_perm
Γ
p
)
//=
always_sep
.
rewrite
pure_
equiv
//
left_id
.
rewrite
pure_
True
//
left_id
.
cancel
[
□
P
]%
I
.
apply
wand_intro_l
.
solve_sep_entails
.
-
destruct
(
Γ
s
!!
i
)
eqn
:
?
;
simplify_eq
/=.
rewrite
(
env_lookup_perm
Γ
s
)
//=.
rewrite
pure_
equiv
//
left_id
.
rewrite
(
env_lookup_perm
Γ
s
)
//=.
rewrite
pure_
True
//
left_id
.
cancel
[
P
].
apply
wand_intro_l
.
solve_sep_entails
.
Qed
.
...
...
@@ -398,7 +398,7 @@ Proof.
Qed
.
Lemma
tac_pure_revert
Δ
φ
Q
:
(
Δ
⊢
■
φ
→
Q
)
→
(
φ
→
Δ
⊢
Q
).
Proof
.
intros
H
Δ
?.
by
rewrite
H
Δ
pure_
equiv
//
left_id
.
Qed
.
Proof
.
intros
H
Δ
?.
by
rewrite
H
Δ
pure_
True
//
left_id
.
Qed
.
(** * Later *)
Class
IntoLaterEnv
(
Γ
1
Γ
2
:
env
(
uPred
M
))
:
=
...
...
@@ -548,7 +548,7 @@ Lemma tac_specialize_assert_pure Δ Δ' j q R P1 P2 φ Q :
φ
→
(
Δ
'
⊢
Q
)
→
Δ
⊢
Q
.
Proof
.
intros
.
rewrite
envs_simple_replace_sound
//
;
simpl
.
rewrite
right_id
(
into_wand
R
)
-(
from_pure
P1
)
pure_
equiv
//.
rewrite
right_id
(
into_wand
R
)
-(
from_pure
P1
)
pure_
True
//.
by
rewrite
wand_True
wand_elim_r
.
Qed
.
...
...
@@ -745,7 +745,7 @@ Qed.
(** * Framing *)
Lemma
tac_frame_pure
Δ
(
φ
:
Prop
)
P
Q
:
φ
→
Frame
(
■
φ
)
P
Q
→
(
Δ
⊢
Q
)
→
Δ
⊢
P
.
Proof
.
intros
??
->.
by
rewrite
-(
frame
(
■
φ
)
P
)
pure_
equiv
//
left_id
.
Qed
.
Proof
.
intros
??
->.
by
rewrite
-(
frame
(
■
φ
)
P
)
pure_
True
//
left_id
.
Qed
.
Lemma
tac_frame
Δ
Δ
'
i
p
R
P
Q
:
envs_lookup_delete
i
Δ
=
Some
(
p
,
R
,
Δ
'
)
→
Frame
R
P
Q
→
...
...
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