Skip to content
GitLab
Projects
Groups
Snippets
Help
Loading...
Help
What's new
7
Help
Support
Community forum
Keyboard shortcuts
?
Submit feedback
Contribute to GitLab
Sign in / Register
Toggle navigation
Open sidebar
Simon Spies
Iris
Commits
7177cd8d
Commit
7177cd8d
authored
Sep 12, 2018
by
Robbert Krebbers
Browse files
Options
Browse Files
Download
Email Patches
Plain Diff
Remove `bi_hforall` and `bi_hexist`, these are superseded by the telescoped versions.
parent
220e7379
Changes
3
Hide whitespace changes
Inline
Side-by-side
Showing
3 changed files
with
0 additions
and
47 deletions
+0
-47
theories/bi/derived_connectives.v
theories/bi/derived_connectives.v
+0
-12
theories/bi/derived_laws_bi.v
theories/bi/derived_laws_bi.v
+0
-27
theories/bi/embedding.v
theories/bi/embedding.v
+0
-8
No files found.
theories/bi/derived_connectives.v
View file @
7177cd8d
From
iris
.
bi
Require
Export
interface
.
From
iris
.
bi
Require
Export
interface
.
From
iris
.
algebra
Require
Import
monoid
.
From
iris
.
algebra
Require
Import
monoid
.
From
stdpp
Require
Import
hlist
.
Definition
bi_iff
{
PROP
:
bi
}
(
P
Q
:
PROP
)
:
PROP
:
=
((
P
→
Q
)
∧
(
Q
→
P
))%
I
.
Definition
bi_iff
{
PROP
:
bi
}
(
P
Q
:
PROP
)
:
PROP
:
=
((
P
→
Q
)
∧
(
Q
→
P
))%
I
.
Arguments
bi_iff
{
_
}
_
%
I
_
%
I
:
simpl
never
.
Arguments
bi_iff
{
_
}
_
%
I
_
%
I
:
simpl
never
.
...
@@ -77,17 +76,6 @@ Instance: Params (@bi_intuitionistically_if) 2.
...
@@ -77,17 +76,6 @@ Instance: Params (@bi_intuitionistically_if) 2.
Typeclasses
Opaque
bi_intuitionistically_if
.
Typeclasses
Opaque
bi_intuitionistically_if
.
Notation
"'□?' p P"
:
=
(
bi_intuitionistically_if
p
P
)
:
bi_scope
.
Notation
"'□?' p P"
:
=
(
bi_intuitionistically_if
p
P
)
:
bi_scope
.
Fixpoint
bi_hexist
{
PROP
:
bi
}
{
As
}
:
himpl
As
PROP
→
PROP
:
=
match
As
return
himpl
As
PROP
→
PROP
with
|
tnil
=>
id
|
tcons
A
As
=>
λ
Φ
,
∃
x
,
bi_hexist
(
Φ
x
)
end
%
I
.
Fixpoint
bi_hforall
{
PROP
:
bi
}
{
As
}
:
himpl
As
PROP
→
PROP
:
=
match
As
return
himpl
As
PROP
→
PROP
with
|
tnil
=>
id
|
tcons
A
As
=>
λ
Φ
,
∀
x
,
bi_hforall
(
Φ
x
)
end
%
I
.
Fixpoint
sbi_laterN
{
PROP
:
sbi
}
(
n
:
nat
)
(
P
:
PROP
)
:
PROP
:
=
Fixpoint
sbi_laterN
{
PROP
:
sbi
}
(
n
:
nat
)
(
P
:
PROP
)
:
PROP
:
=
match
n
with
match
n
with
|
O
=>
P
|
O
=>
P
...
...
theories/bi/derived_laws_bi.v
View file @
7177cd8d
From
iris
.
bi
Require
Export
derived_connectives
.
From
iris
.
bi
Require
Export
derived_connectives
.
From
iris
.
algebra
Require
Import
monoid
.
From
iris
.
algebra
Require
Import
monoid
.
From
stdpp
Require
Import
hlist
.
(** Naming schema for lemmas about modalities:
(** Naming schema for lemmas about modalities:
M1_into_M2: M1 P ⊢ M2 P
M1_into_M2: M1 P ⊢ M2 P
...
@@ -1449,31 +1448,6 @@ Global Instance bi_persistently_sep_entails_homomorphism :
...
@@ -1449,31 +1448,6 @@ Global Instance bi_persistently_sep_entails_homomorphism :
MonoidHomomorphism
bi_sep
bi_sep
(
flip
(
⊢
))
(@
bi_persistently
PROP
).
MonoidHomomorphism
bi_sep
bi_sep
(
flip
(
⊢
))
(@
bi_persistently
PROP
).
Proof
.
split
.
apply
_
.
simpl
.
apply
persistently_emp_intro
.
Qed
.
Proof
.
split
.
apply
_
.
simpl
.
apply
persistently_emp_intro
.
Qed
.
(* Heterogeneous lists *)
Lemma
hexist_exist
{
As
B
}
(
f
:
himpl
As
B
)
(
Φ
:
B
→
PROP
)
:
bi_hexist
(
hcompose
Φ
f
)
⊣
⊢
∃
xs
:
hlist
As
,
Φ
(
f
xs
).
Proof
.
apply
(
anti_symm
_
).
-
induction
As
as
[|
A
As
IH
]
;
simpl
.
+
by
rewrite
-(
exist_intro
hnil
)
.
+
apply
exist_elim
=>
x
;
rewrite
IH
;
apply
exist_elim
=>
xs
.
by
rewrite
-(
exist_intro
(
hcons
x
xs
)).
-
apply
exist_elim
=>
xs
;
induction
xs
as
[|
A
As
x
xs
IH
]
;
simpl
;
auto
.
by
rewrite
-(
exist_intro
x
)
IH
.
Qed
.
Lemma
hforall_forall
{
As
B
}
(
f
:
himpl
As
B
)
(
Φ
:
B
→
PROP
)
:
bi_hforall
(
hcompose
Φ
f
)
⊣
⊢
∀
xs
:
hlist
As
,
Φ
(
f
xs
).
Proof
.
apply
(
anti_symm
_
).
-
apply
forall_intro
=>
xs
;
induction
xs
as
[|
A
As
x
xs
IH
]
;
simpl
;
auto
.
by
rewrite
(
forall_elim
x
)
IH
.
-
induction
As
as
[|
A
As
IH
]
;
simpl
.
+
by
rewrite
(
forall_elim
hnil
)
.
+
apply
forall_intro
=>
x
;
rewrite
-
IH
;
apply
forall_intro
=>
xs
.
by
rewrite
(
forall_elim
(
hcons
x
xs
)).
Qed
.
(* Limits *)
(* Limits *)
Lemma
limit_preserving_entails
{
A
:
ofeT
}
`
{
Cofe
A
}
(
Φ
Ψ
:
A
→
PROP
)
:
Lemma
limit_preserving_entails
{
A
:
ofeT
}
`
{
Cofe
A
}
(
Φ
Ψ
:
A
→
PROP
)
:
NonExpansive
Φ
→
NonExpansive
Ψ
→
LimitPreserving
(
λ
x
,
Φ
x
⊢
Ψ
x
).
NonExpansive
Φ
→
NonExpansive
Ψ
→
LimitPreserving
(
λ
x
,
Φ
x
⊢
Ψ
x
).
...
@@ -1492,5 +1466,4 @@ Global Instance limit_preserving_Persistent {A:ofeT} `{Cofe A} (Φ : A → PROP)
...
@@ -1492,5 +1466,4 @@ Global Instance limit_preserving_Persistent {A:ofeT} `{Cofe A} (Φ : A → PROP)
NonExpansive
Φ
→
LimitPreserving
(
λ
x
,
Persistent
(
Φ
x
)).
NonExpansive
Φ
→
LimitPreserving
(
λ
x
,
Persistent
(
Φ
x
)).
Proof
.
intros
.
apply
limit_preserving_entails
;
solve_proper
.
Qed
.
Proof
.
intros
.
apply
limit_preserving_entails
;
solve_proper
.
Qed
.
End
bi_derived
.
End
bi_derived
.
End
bi
.
End
bi
.
theories/bi/embedding.v
View file @
7177cd8d
From
iris
.
algebra
Require
Import
monoid
.
From
iris
.
algebra
Require
Import
monoid
.
From
iris
.
bi
Require
Import
interface
derived_laws_sbi
big_op
plainly
updates
.
From
iris
.
bi
Require
Import
interface
derived_laws_sbi
big_op
plainly
updates
.
From
stdpp
Require
Import
hlist
.
Class
Embed
(
A
B
:
Type
)
:
=
embed
:
A
→
B
.
Class
Embed
(
A
B
:
Type
)
:
=
embed
:
A
→
B
.
Arguments
embed
{
_
_
_
}
_
%
I
:
simpl
never
.
Arguments
embed
{
_
_
_
}
_
%
I
:
simpl
never
.
...
@@ -189,13 +188,6 @@ Section embed.
...
@@ -189,13 +188,6 @@ Section embed.
⎡□
?b
P
⎤
⊣
⊢
□
?b
⎡
P
⎤
.
⎡□
?b
P
⎤
⊣
⊢
□
?b
⎡
P
⎤
.
Proof
.
destruct
b
;
simpl
;
auto
using
embed_intuitionistically
.
Qed
.
Proof
.
destruct
b
;
simpl
;
auto
using
embed_intuitionistically
.
Qed
.
Lemma
embed_hforall
{
As
}
(
Φ
:
himpl
As
PROP1
)
:
⎡
bi_hforall
Φ⎤
⊣
⊢
bi_hforall
(
hcompose
embed
Φ
).
Proof
.
induction
As
=>//.
rewrite
/=
embed_forall
.
by
do
2
f_equiv
.
Qed
.
Lemma
embed_hexist
{
As
}
(
Φ
:
himpl
As
PROP1
)
:
⎡
bi_hexist
Φ⎤
⊣
⊢
bi_hexist
(
hcompose
embed
Φ
).
Proof
.
induction
As
=>//.
rewrite
/=
embed_exist
.
by
do
2
f_equiv
.
Qed
.
Global
Instance
embed_persistent
P
:
Persistent
P
→
Persistent
⎡
P
⎤
.
Global
Instance
embed_persistent
P
:
Persistent
P
→
Persistent
⎡
P
⎤
.
Proof
.
intros
?.
by
rewrite
/
Persistent
-
embed_persistently
-
persistent
.
Qed
.
Proof
.
intros
?.
by
rewrite
/
Persistent
-
embed_persistently
-
persistent
.
Qed
.
Global
Instance
embed_affine
`
{!
BiEmbedEmp
PROP1
PROP2
}
P
:
Affine
P
→
Affine
⎡
P
⎤
.
Global
Instance
embed_affine
`
{!
BiEmbedEmp
PROP1
PROP2
}
P
:
Affine
P
→
Affine
⎡
P
⎤
.
...
...
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