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
Tej Chajed
iris
Commits
e2bb6d8d
Commit
e2bb6d8d
authored
May 04, 2018
by
Ralf Jung
Browse files
all SBI must be a COFE
parent
19f3e261
Changes
5
Hide whitespace changes
Inline
Side-by-side
theories/bi/derived_laws_sbi.v
View file @
e2bb6d8d
...
...
@@ -210,8 +210,6 @@ Section löb.
Local
Instance
fl
ö
b_pre_contractive
P
:
Contractive
(
fl
ö
b_pre
P
).
Proof
.
solve_contractive
.
Qed
.
Context
`
{
Cofe
PROP
}.
Definition
fl
ö
b
(
P
:
PROP
)
:
=
fixpoint
(
fl
ö
b_pre
P
).
Lemma
weak_l
ö
b
P
:
(
▷
P
⊢
P
)
→
(
True
⊢
P
).
...
...
@@ -412,7 +410,7 @@ Proof. intros; rewrite /Timeless except_0_and later_and; auto. Qed.
Global
Instance
or_timeless
P
Q
:
Timeless
P
→
Timeless
Q
→
Timeless
(
P
∨
Q
).
Proof
.
intros
;
rewrite
/
Timeless
except_0_or
later_or
;
auto
.
Qed
.
Global
Instance
impl_timeless
`
{
Cofe
PROP
}
P
Q
:
Timeless
Q
→
Timeless
(
P
→
Q
).
Global
Instance
impl_timeless
P
Q
:
Timeless
Q
→
Timeless
(
P
→
Q
).
Proof
.
rewrite
/
Timeless
=>
HQ
.
rewrite
later_false_em
.
apply
or_mono
,
impl_intro_l
;
first
done
.
...
...
@@ -425,7 +423,7 @@ Proof.
intros
;
rewrite
/
Timeless
except_0_sep
later_sep
;
auto
using
sep_mono
.
Qed
.
Global
Instance
wand_timeless
`
{
Cofe
PROP
}
P
Q
:
Timeless
Q
→
Timeless
(
P
-
∗
Q
).
Global
Instance
wand_timeless
P
Q
:
Timeless
Q
→
Timeless
(
P
-
∗
Q
).
Proof
.
rewrite
/
Timeless
=>
HQ
.
rewrite
later_false_em
.
apply
or_mono
,
wand_intro_l
;
first
done
.
...
...
theories/bi/interface.v
View file @
e2bb6d8d
...
...
@@ -229,6 +229,7 @@ Structure sbi := Sbi {
sbi_internal_eq
:
∀
A
:
ofeT
,
A
→
A
→
sbi_car
;
sbi_later
:
sbi_car
→
sbi_car
;
sbi_ofe_mixin
:
OfeMixin
sbi_car
;
sbi_cofe
:
Cofe
(
OfeT
sbi_car
sbi_ofe_mixin
)
;
sbi_bi_mixin
:
BiMixin
sbi_entails
sbi_emp
sbi_pure
sbi_and
sbi_or
sbi_impl
sbi_forall
sbi_exist
sbi_sep
sbi_wand
sbi_persistently
;
sbi_sbi_mixin
:
SbiMixin
sbi_entails
sbi_pure
sbi_or
sbi_impl
...
...
@@ -247,6 +248,8 @@ Canonical Structure sbi_ofeC.
Coercion
sbi_bi
(
PROP
:
sbi
)
:
bi
:
=
{|
bi_ofe_mixin
:
=
sbi_ofe_mixin
PROP
;
bi_bi_mixin
:
=
sbi_bi_mixin
PROP
|}.
Canonical
Structure
sbi_bi
.
Global
Instance
sbi_cofe'
(
PROP
:
sbi
)
:
Cofe
PROP
.
Proof
.
apply
sbi_cofe
.
Qed
.
Arguments
sbi_car
:
simpl
never
.
Arguments
sbi_dist
:
simpl
never
.
...
...
theories/proofmode/coq_tactics.v
View file @
e2bb6d8d
...
...
@@ -1396,7 +1396,7 @@ Proof.
-
by
rewrite
Hs
//=
right_id
.
Qed
.
Lemma
tac_l
ö
b
`
{
Cofe
PROP
}
Δ
Δ
'
i
Q
:
Lemma
tac_l
ö
b
Δ
Δ
'
i
Q
:
env_spatial_is_nil
Δ
=
true
→
envs_app
true
(
Esnoc
Enil
i
(
▷
Q
)%
I
)
Δ
=
Some
Δ
'
→
envs_entails
Δ
'
Q
→
envs_entails
Δ
Q
.
...
...
theories/proofmode/ltac_tactics.v
View file @
e2bb6d8d
...
...
@@ -1659,8 +1659,7 @@ Tactic Notation "iLöbCore" "as" constr (IH) :=
refine should use the other unification algorithm, which should
not have this issue. *)
notypeclasses
refine
(
tac_l
ö
b
_
_
IH
_
_
_
_
)
;
[
iSolveTC
||
fail
"iLöb: PROP is not a Cofe"
|
reflexivity
||
fail
"iLöb: spatial context not empty, this should not happen"
[
reflexivity
||
fail
"iLöb: spatial context not empty, this should not happen"
|
env_reflexivity
||
fail
"iLöb:"
IH
"not fresh"
|].
Tactic
Notation
"iLöbRevert"
constr
(
Hs
)
"with"
tactic
(
tac
)
:
=
...
...
theories/tests/proofmode.v
View file @
e2bb6d8d
From
iris
.
proofmode
Require
Import
tactics
intro_patterns
.
From
stdpp
Require
Import
gmap
hlist
.
Set
Default
Proof
Using
"Type
*
"
.
Set
Default
Proof
Using
"Type"
.
Section
tests
.
Context
{
PROP
:
sbi
}
`
{
Cofe
PROP
}
.
Context
{
PROP
:
sbi
}.
Implicit
Types
P
Q
R
:
PROP
.
Lemma
demo_0
P
Q
:
□
(
P
∨
Q
)
-
∗
(
∀
x
,
⌜
x
=
0
⌝
∨
⌜
x
=
1
⌝
)
→
(
Q
∨
P
).
...
...
@@ -381,8 +381,8 @@ Proof. iIntros (Hφ). iAssert ⌜φ⌝%I with "[%]" as "$"; auto with iFrame. Qe
Lemma
test_iEval
x
y
:
⌜
(
y
+
x
)%
nat
=
1
⌝
-
∗
⌜
S
(
x
+
y
)
=
2
%
nat
⌝
:
PROP
.
Proof
.
iIntros
(
H
n
).
iEval
(
rewrite
(
Nat
.
add_comm
x
y
)
//
H
n
).
iIntros
(
H
).
iEval
(
rewrite
(
Nat
.
add_comm
x
y
)
//
H
).
done
.
Qed
.
...
...
Write
Preview
Supports
Markdown
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