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
58b8eafa
Commit
58b8eafa
authored
Nov 02, 2017
by
Jacques-Henri Jourdan
Browse files
Renaming: sink->absorbingly bare->affinely.
parent
a38db108
Changes
13
Expand all
Hide whitespace changes
Inline
Side-by-side
theories/base_logic/derived.v
View file @
58b8eafa
...
...
@@ -36,9 +36,9 @@ Global Instance uPred_affine : AffineBI (uPredI M) | 0.
Proof
.
intros
P
.
rewrite
/
Affine
.
by
apply
bi
.
pure_intro
.
Qed
.
(* Own and valid derived *)
Lemma
bare
_persistently_ownM
(
a
:
M
)
:
CoreId
a
→
□
uPred_ownM
a
⊣
⊢
uPred_ownM
a
.
Lemma
affinely
_persistently_ownM
(
a
:
M
)
:
CoreId
a
→
□
uPred_ownM
a
⊣
⊢
uPred_ownM
a
.
Proof
.
rewrite
affine_
bare
=>?
;
apply
(
anti_symm
_
)
;
[
by
rewrite
persistently_elim
|].
rewrite
affine_
affinely
=>?
;
apply
(
anti_symm
_
)
;
[
by
rewrite
persistently_elim
|].
by
rewrite
{
1
}
persistently_ownM_core
core_id_core
.
Qed
.
Lemma
ownM_invalid
(
a
:
M
)
:
¬
✓
{
0
}
a
→
uPred_ownM
a
⊢
False
.
...
...
@@ -47,9 +47,9 @@ Global Instance ownM_mono : Proper (flip (≼) ==> (⊢)) (@uPred_ownM M).
Proof
.
intros
a
b
[
b'
->].
by
rewrite
ownM_op
sep_elim_l
.
Qed
.
Lemma
ownM_unit'
:
uPred_ownM
ε
⊣
⊢
True
.
Proof
.
apply
(
anti_symm
_
)
;
first
by
apply
pure_intro
.
apply
ownM_empty
.
Qed
.
Lemma
bare
_persistently_cmra_valid
{
A
:
cmraT
}
(
a
:
A
)
:
□
✓
a
⊣
⊢
✓
a
.
Lemma
affinely
_persistently_cmra_valid
{
A
:
cmraT
}
(
a
:
A
)
:
□
✓
a
⊣
⊢
✓
a
.
Proof
.
rewrite
affine_
bare
.
intros
;
apply
(
anti_symm
_
)
;
first
by
rewrite
persistently_elim
.
rewrite
affine_
affinely
.
intros
;
apply
(
anti_symm
_
)
;
first
by
rewrite
persistently_elim
.
apply
:
persistently_cmra_valid_1
.
Qed
.
...
...
theories/base_logic/lib/fancy_updates.v
View file @
58b8eafa
...
...
@@ -157,14 +157,14 @@ Section proofmode_classes.
IntoWand
false
false
R
P
Q
→
IntoWand
p
q
(|={
E
}=>
R
)
(|={
E
}=>
P
)
(|={
E
}=>
Q
).
Proof
.
rewrite
/
IntoWand
/=
=>
HR
.
rewrite
!
bare
_persistently_if_elim
HR
.
rewrite
/
IntoWand
/=
=>
HR
.
rewrite
!
affinely
_persistently_if_elim
HR
.
apply
wand_intro_l
.
by
rewrite
fupd_sep
wand_elim_r
.
Qed
.
Global
Instance
into_wand_fupd_persistent
E1
E2
p
q
R
P
Q
:
IntoWand
false
q
R
P
Q
→
IntoWand
p
q
(|={
E1
,
E2
}=>
R
)
P
(|={
E1
,
E2
}=>
Q
).
Proof
.
rewrite
/
IntoWand
/=
=>
HR
.
rewrite
bare
_persistently_if_elim
HR
.
rewrite
/
IntoWand
/=
=>
HR
.
rewrite
affinely
_persistently_if_elim
HR
.
apply
wand_intro_l
.
by
rewrite
fupd_frame_l
wand_elim_r
.
Qed
.
...
...
@@ -172,7 +172,7 @@ Section proofmode_classes.
IntoWand
p
false
R
P
Q
→
IntoWand'
p
q
R
(|={
E1
,
E2
}=>
P
)
(|={
E1
,
E2
}=>
Q
).
Proof
.
rewrite
/
IntoWand'
/
IntoWand
/=
=>
->.
apply
wand_intro_l
.
by
rewrite
bare
_persistently_if_elim
fupd_wand_r
.
apply
wand_intro_l
.
by
rewrite
affinely
_persistently_if_elim
fupd_wand_r
.
Qed
.
Global
Instance
from_sep_fupd
E
P
Q1
Q2
:
...
...
theories/base_logic/lib/viewshifts.v
View file @
58b8eafa
...
...
@@ -82,7 +82,7 @@ Proof. iIntros "!# HP". by iApply inv_alloc. Qed.
Lemma
wand_fupd_alt
E1
E2
P
Q
:
(
P
={
E1
,
E2
}=
∗
Q
)
⊣
⊢
∃
R
,
R
∗
(
P
∗
R
={
E1
,
E2
}=>
Q
).
Proof
.
rewrite
bi
.
wand_alt
.
do
2
f_equiv
.
setoid_rewrite
bi
.
affine_
bare
;
last
apply
_
.
rewrite
bi
.
wand_alt
.
do
2
f_equiv
.
setoid_rewrite
bi
.
affine_
affinely
;
last
apply
_
.
by
rewrite
bi
.
persistently_impl_wand
.
Qed
.
End
vs
.
theories/base_logic/proofmode.v
View file @
58b8eafa
...
...
@@ -31,14 +31,14 @@ Qed.
Global
Instance
into_wand_bupd
p
q
R
P
Q
:
IntoWand
false
false
R
P
Q
→
IntoWand
p
q
(|==>
R
)
(|==>
P
)
(|==>
Q
).
Proof
.
rewrite
/
IntoWand
/=
=>
HR
.
rewrite
!
bare
_persistently_if_elim
HR
.
rewrite
/
IntoWand
/=
=>
HR
.
rewrite
!
affinely
_persistently_if_elim
HR
.
apply
wand_intro_l
.
by
rewrite
bupd_sep
wand_elim_r
.
Qed
.
Global
Instance
into_wand_bupd_persistent
p
q
R
P
Q
:
IntoWand
false
q
R
P
Q
→
IntoWand
p
q
(|==>
R
)
P
(|==>
Q
).
Proof
.
rewrite
/
IntoWand
/=
=>
HR
.
rewrite
bare
_persistently_if_elim
HR
.
rewrite
/
IntoWand
/=
=>
HR
.
rewrite
affinely
_persistently_if_elim
HR
.
apply
wand_intro_l
.
by
rewrite
bupd_frame_l
wand_elim_r
.
Qed
.
...
...
@@ -46,7 +46,7 @@ Global Instance into_wand_bupd_args p q R P Q :
IntoWand
p
false
R
P
Q
→
IntoWand'
p
q
R
(|==>
P
)
(|==>
Q
).
Proof
.
rewrite
/
IntoWand'
/
IntoWand
/=
=>
->.
apply
wand_intro_l
.
by
rewrite
bare
_persistently_if_elim
bupd_wand_r
.
apply
wand_intro_l
.
by
rewrite
affinely
_persistently_if_elim
bupd_wand_r
.
Qed
.
(* FromOp *)
...
...
@@ -85,7 +85,7 @@ Qed.
Global
Instance
into_and_ownM
p
(
a
b1
b2
:
M
)
:
IsOp
a
b1
b2
→
IntoAnd
p
(
uPred_ownM
a
)
(
uPred_ownM
b1
)
(
uPred_ownM
b2
).
Proof
.
intros
.
apply
bare
_persistently_if_mono
.
by
rewrite
(
is_op
a
)
ownM_op
sep_and
.
intros
.
apply
affinely
_persistently_if_mono
.
by
rewrite
(
is_op
a
)
ownM_op
sep_and
.
Qed
.
Global
Instance
into_sep_ownM
(
a
b1
b2
:
M
)
:
...
...
theories/bi/big_op.v
View file @
58b8eafa
...
...
@@ -150,12 +150,12 @@ Section sep_list.
Proof
.
apply
wand_intro_l
.
revert
Φ
Ψ
.
induction
l
as
[|
x
l
IH
]=>
Φ
Ψ
/=.
{
by
rewrite
sep_elim_r
.
}
rewrite
bare
_persistently_sep_dup
-
assoc
[(
□
_
∗
_
)%
I
]
comm
-!
assoc
assoc
.
rewrite
affinely
_persistently_sep_dup
-
assoc
[(
□
_
∗
_
)%
I
]
comm
-!
assoc
assoc
.
apply
sep_mono
.
-
rewrite
(
forall_elim
0
)
(
forall_elim
x
)
pure_True
//
True_impl
.
by
rewrite
bare
_persistently_elim
wand_elim_l
.
by
rewrite
affinely
_persistently_elim
wand_elim_l
.
-
rewrite
comm
-(
IH
(
Φ
∘
S
)
(
Ψ
∘
S
))
/=.
apply
sep_mono_l
,
bare
_mono
,
persistently_mono
.
apply
sep_mono_l
,
affinely
_mono
,
persistently_mono
.
apply
forall_intro
=>
k
.
by
rewrite
(
forall_elim
(
S
k
)).
Qed
.
...
...
@@ -424,12 +424,12 @@ Section gmap.
Proof
.
apply
wand_intro_l
.
induction
m
as
[|
i
x
m
?
IH
]
using
map_ind
.
{
by
rewrite
sep_elim_r
.
}
rewrite
!
big_sepM_insert
//
bare
_persistently_sep_dup
.
rewrite
!
big_sepM_insert
//
affinely
_persistently_sep_dup
.
rewrite
-
assoc
[(
□
_
∗
_
)%
I
]
comm
-!
assoc
assoc
.
apply
sep_mono
.
-
rewrite
(
forall_elim
i
)
(
forall_elim
x
)
pure_True
?lookup_insert
//.
by
rewrite
True_impl
bare
_persistently_elim
wand_elim_l
.
by
rewrite
True_impl
affinely
_persistently_elim
wand_elim_l
.
-
rewrite
comm
-
IH
/=.
apply
sep_mono_l
,
bare
_mono
,
persistently_mono
,
forall_mono
=>
k
.
apply
sep_mono_l
,
affinely
_mono
,
persistently_mono
,
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_True
//
True_impl
.
...
...
@@ -585,11 +585,11 @@ Section gset.
Proof
.
apply
wand_intro_l
.
induction
X
as
[|
x
X
?
IH
]
using
collection_ind_L
.
{
by
rewrite
sep_elim_r
.
}
rewrite
!
big_sepS_insert
//
bare
_persistently_sep_dup
.
rewrite
!
big_sepS_insert
//
affinely
_persistently_sep_dup
.
rewrite
-
assoc
[(
□
_
∗
_
)%
I
]
comm
-!
assoc
assoc
.
apply
sep_mono
.
-
rewrite
(
forall_elim
x
)
pure_True
;
last
set_solver
.
by
rewrite
True_impl
bare
_persistently_elim
wand_elim_l
.
-
rewrite
comm
-
IH
/=.
apply
sep_mono_l
,
bare
_mono
,
persistently_mono
.
by
rewrite
True_impl
affinely
_persistently_elim
wand_elim_l
.
-
rewrite
comm
-
IH
/=.
apply
sep_mono_l
,
affinely
_mono
,
persistently_mono
.
apply
forall_mono
=>
y
.
apply
impl_intro_l
,
pure_elim_l
=>
?.
by
rewrite
pure_True
?True_impl
;
last
set_solver
.
Qed
.
...
...
theories/bi/derived.v
View file @
58b8eafa
This diff is collapsed.
Click to expand it.
theories/bi/fractional.v
View file @
58b8eafa
...
...
@@ -173,6 +173,6 @@ Section fractional.
-
rewrite
fractional
/
Frame
/
MakeSep
=><-<-.
by
rewrite
assoc
.
-
rewrite
fractional
/
Frame
/
MakeSep
=><-<-=>
_
.
by
rewrite
(
comm
_
Q
(
Φ
q0
))
!
assoc
(
comm
_
(
Φ
_
)).
-
move
=>-[->
_
]->.
by
rewrite
bi
.
bare
_persistently_if_elim
-
fractional
Qp_div_2
.
-
move
=>-[->
_
]->.
by
rewrite
bi
.
affinely
_persistently_if_elim
-
fractional
Qp_div_2
.
Qed
.
End
fractional
.
theories/program_logic/hoare.v
View file @
58b8eafa
...
...
@@ -37,7 +37,7 @@ Global Instance ht_proper E :
Proof
.
solve_proper
.
Qed
.
Lemma
ht_mono
E
P
P'
Φ
Φ
'
e
:
(
P
⊢
P'
)
→
(
∀
v
,
Φ
'
v
⊢
Φ
v
)
→
{{
P'
}}
e
@
E
{{
Φ
'
}}
⊢
{{
P
}}
e
@
E
{{
Φ
}}.
Proof
.
intros
.
by
apply
bare
_mono
,
persistently_mono
,
wand_mono
,
wp_mono
.
Qed
.
Proof
.
intros
.
by
apply
affinely
_mono
,
persistently_mono
,
wand_mono
,
wp_mono
.
Qed
.
Global
Instance
ht_mono'
E
:
Proper
(
flip
(
⊢
)
==>
eq
==>
pointwise_relation
_
(
⊢
)
==>
(
⊢
))
(
ht
E
).
Proof
.
solve_proper
.
Qed
.
...
...
theories/proofmode/class_instances.v
View file @
58b8eafa
This diff is collapsed.
Click to expand it.
theories/proofmode/classes.v
View file @
58b8eafa
...
...
@@ -56,23 +56,23 @@ Arguments into_persistent {_} _ _%I _%I {_}.
Hint
Mode
IntoPersistent
+
+
!
-
:
typeclass_instances
.
Class
FromPersistent
{
PROP
:
bi
}
(
a
p
:
bool
)
(
P
Q
:
PROP
)
:
=
from_persistent
:
bi_
bare
_if
a
(
bi_persistently_if
p
Q
)
⊢
P
.
from_persistent
:
bi_
affinely
_if
a
(
bi_persistently_if
p
Q
)
⊢
P
.
Arguments
FromPersistent
{
_
}
_
_
_
%
I
_
%
I
:
simpl
never
.
Arguments
from_persistent
{
_
}
_
_
_
%
I
_
%
I
{
_
}.
Hint
Mode
FromPersistent
+
-
-
!
-
:
typeclass_instances
.
Class
From
Bare
{
PROP
:
bi
}
(
P
Q
:
PROP
)
:
=
from_
bare
:
bi_bare
Q
⊢
P
.
Arguments
From
Bare
{
_
}
_
%
I
_
%
type_scope
:
simpl
never
.
Arguments
from_
bare
{
_
}
_
%
I
_
%
type_scope
{
_
}.
Hint
Mode
From
Bare
+
!
-
:
typeclass_instances
.
Hint
Mode
From
Bare
+
-
!
:
typeclass_instances
.
Class
Into
Sink
{
PROP
:
bi
}
(
P
Q
:
PROP
)
:
=
into_
sink
:
P
⊢
▲
Q
.
Arguments
Into
Sink
{
_
}
_
%
I
_
%
I
.
Arguments
into_
sink
{
_
}
_
%
I
_
%
I
{
_
}.
Hint
Mode
Into
Sink
+
!
-
:
typeclass_instances
.
Hint
Mode
Into
Sink
+
-
!
:
typeclass_instances
.
Class
From
Affinely
{
PROP
:
bi
}
(
P
Q
:
PROP
)
:
=
from_
affinely
:
bi_affinely
Q
⊢
P
.
Arguments
From
Affinely
{
_
}
_
%
I
_
%
type_scope
:
simpl
never
.
Arguments
from_
affinely
{
_
}
_
%
I
_
%
type_scope
{
_
}.
Hint
Mode
From
Affinely
+
!
-
:
typeclass_instances
.
Hint
Mode
From
Affinely
+
-
!
:
typeclass_instances
.
Class
Into
Absorbingly
{
PROP
:
bi
}
(
P
Q
:
PROP
)
:
=
into_
absorbingly
:
P
⊢
▲
Q
.
Arguments
Into
Absorbingly
{
_
}
_
%
I
_
%
I
.
Arguments
into_
absorbingly
{
_
}
_
%
I
_
%
I
{
_
}.
Hint
Mode
Into
Absorbingly
+
!
-
:
typeclass_instances
.
Hint
Mode
Into
Absorbingly
+
-
!
:
typeclass_instances
.
Class
IntoInternalEq
{
PROP
:
bi
}
{
A
:
ofeT
}
(
P
:
PROP
)
(
x
y
:
A
)
:
=
into_internal_eq
:
P
⊢
x
≡
y
.
...
...
theories/proofmode/coq_tactics.v
View file @
58b8eafa
This diff is collapsed.
Click to expand it.
theories/proofmode/tactics.v
View file @
58b8eafa
...
...
@@ -1269,8 +1269,8 @@ Instance copy_destruct_impl {PROP : bi} (P Q : PROP) :
CopyDestruct
Q
→
CopyDestruct
(
P
→
Q
).
Instance
copy_destruct_wand
{
PROP
:
bi
}
(
P
Q
:
PROP
)
:
CopyDestruct
Q
→
CopyDestruct
(
P
-
∗
Q
).
Instance
copy_destruct_
bare
{
PROP
:
bi
}
(
P
:
PROP
)
:
CopyDestruct
P
→
CopyDestruct
(
bi_
bare
P
).
Instance
copy_destruct_
affinely
{
PROP
:
bi
}
(
P
:
PROP
)
:
CopyDestruct
P
→
CopyDestruct
(
bi_
affinely
P
).
Instance
copy_destruct_persistently
{
PROP
:
bi
}
(
P
:
PROP
)
:
CopyDestruct
P
→
CopyDestruct
(
bi_persistently
P
).
...
...
@@ -1771,7 +1771,7 @@ Hint Extern 1 (of_envs _ ⊢ _ ∧ _) => iSplit.
Hint
Extern
1
(
of_envs
_
⊢
_
∗
_
)
=>
iSplit
.
Hint
Extern
1
(
of_envs
_
⊢
▷
_
)
=>
iNext
.
Hint
Extern
1
(
of_envs
_
⊢
bi_persistently
_
)
=>
iAlways
.
Hint
Extern
1
(
of_envs
_
⊢
bi_
bare
_
)
=>
iAlways
.
Hint
Extern
1
(
of_envs
_
⊢
bi_
affinely
_
)
=>
iAlways
.
Hint
Extern
1
(
of_envs
_
⊢
∃
_
,
_
)
=>
iExists
_
.
Hint
Extern
1
(
of_envs
_
⊢
◇
_
)
=>
iModIntro
.
Hint
Extern
1
(
of_envs
_
⊢
_
∨
_
)
=>
iLeft
.
...
...
theories/tests/proofmode.v
View file @
58b8eafa
...
...
@@ -44,7 +44,7 @@ Lemma test_unfold_constants : bar.
Proof
.
iIntros
(
P
)
"HP //"
.
Qed
.
Lemma
test_iRewrite
{
A
:
ofeT
}
(
x
y
:
A
)
P
:
□
(
∀
z
,
P
-
∗
bi_
bare
(
z
≡
y
))
-
∗
(
P
-
∗
P
∧
(
x
,
x
)
≡
(
y
,
x
)).
□
(
∀
z
,
P
-
∗
bi_
affinely
(
z
≡
y
))
-
∗
(
P
-
∗
P
∧
(
x
,
x
)
≡
(
y
,
x
)).
Proof
.
iIntros
"#H1 H2"
.
iRewrite
(
bi
.
internal_eq_sym
x
x
with
"[# //]"
).
...
...
@@ -53,7 +53,7 @@ Proof.
Qed
.
Lemma
test_iDestruct_and_emp
P
Q
`
{!
Persistent
P
,
!
Persistent
Q
}
:
P
∧
emp
-
∗
emp
∧
Q
-
∗
bi_
bare
(
P
∗
Q
).
P
∧
emp
-
∗
emp
∧
Q
-
∗
bi_
affinely
(
P
∗
Q
).
Proof
.
iIntros
"[#? _] [_ #?]"
.
auto
.
Qed
.
Lemma
test_iIntros_persistent
P
Q
`
{!
Persistent
Q
}
:
(
P
→
Q
→
P
∧
Q
)%
I
.
...
...
@@ -110,17 +110,18 @@ Lemma test_iSpecialize_tc P : (∀ x y z : gset positive, P) -∗ P.
Proof
.
iIntros
"H"
.
iSpecialize
(
"H"
$!
∅
{[
1
%
positive
]}
∅
).
done
.
Qed
.
Lemma
test_iFrame_pure
{
A
:
ofeT
}
(
φ
:
Prop
)
(
y
z
:
A
)
:
φ
→
bi_
bare
⌜
y
≡
z
⌝
-
∗
(
⌜
φ
⌝
∧
⌜
φ
⌝
∧
y
≡
z
:
PROP
).
φ
→
bi_
affinely
⌜
y
≡
z
⌝
-
∗
(
⌜
φ
⌝
∧
⌜
φ
⌝
∧
y
≡
z
:
PROP
).
Proof
.
iIntros
(
Hv
)
"#Hxy"
.
iFrame
(
Hv
)
"Hxy"
.
Qed
.
Lemma
test_iAssert_modality
P
:
◇
False
-
∗
▷
P
.
Proof
.
iIntros
"HF"
.
iAssert
(
bi_
bare
False
)%
I
with
"[> -]"
as
%[].
iAssert
(
bi_
affinely
False
)%
I
with
"[> -]"
as
%[].
by
iMod
"HF"
.
Qed
.
Lemma
test_iMod_bare_timeless
P
`
{!
Timeless
P
}
:
bi_bare
(
▷
P
)
-
∗
◇
bi_bare
P
.
Lemma
test_iMod_affinely_timeless
P
`
{!
Timeless
P
}
:
bi_affinely
(
▷
P
)
-
∗
◇
bi_affinely
P
.
Proof
.
iIntros
"H"
.
iMod
"H"
.
done
.
Qed
.
Lemma
test_iAssumption_False
P
:
False
-
∗
P
.
...
...
@@ -220,7 +221,8 @@ Lemma test_iIntros_let P :
∀
Q
,
let
R
:
=
emp
%
I
in
P
-
∗
R
-
∗
Q
-
∗
P
∗
Q
.
Proof
.
iIntros
(
Q
R
)
"$ _ $"
.
Qed
.
Lemma
test_foo
P
Q
:
bi_bare
(
▷
(
Q
≡
P
))
-
∗
bi_bare
(
▷
Q
)
-
∗
bi_bare
(
▷
P
).
Lemma
test_foo
P
Q
:
bi_affinely
(
▷
(
Q
≡
P
))
-
∗
bi_affinely
(
▷
Q
)
-
∗
bi_affinely
(
▷
P
).
Proof
.
iIntros
"#HPQ HQ !#"
.
iNext
.
by
iRewrite
"HPQ"
in
"HQ"
.
Qed
.
...
...
@@ -234,10 +236,10 @@ Proof.
Qed
.
Lemma
test_iNext_affine
P
Q
:
bi_
bare
(
▷
(
Q
≡
P
))
-
∗
bi_
bare
(
▷
Q
)
-
∗
bi_
bare
(
▷
P
).
bi_
affinely
(
▷
(
Q
≡
P
))
-
∗
bi_
affinely
(
▷
Q
)
-
∗
bi_
affinely
(
▷
P
).
Proof
.
iIntros
"#HPQ HQ !#"
.
iNext
.
by
iRewrite
"HPQ"
in
"HQ"
.
Qed
.
Lemma
test_iAlways
P
Q
R
:
□
P
-
∗
bi_persistently
Q
→
R
-
∗
bi_persistently
(
bi_
bare
(
bi_bare
P
))
∗
□
Q
.
□
P
-
∗
bi_persistently
Q
→
R
-
∗
bi_persistently
(
bi_
affinely
(
bi_affinely
P
))
∗
□
Q
.
Proof
.
iIntros
"#HP #HQ HR"
.
iSplitL
.
iAlways
.
done
.
iAlways
.
done
.
Qed
.
End
tests
.
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