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
fbf07f03
Commit
fbf07f03
authored
Jan 05, 2017
by
Ralf Jung
Browse files
program_logic: improve 'proof using' hint to be more minimal, more future-proof
parent
599c70e1
Changes
9
Hide whitespace changes
Inline
Side-by-side
theories/program_logic/adequacy.v
View file @
fbf07f03
...
...
@@ -3,7 +3,7 @@ From iris.algebra Require Import gmap auth agree gset coPset.
From
iris
.
base_logic
Require
Import
big_op
soundness
.
From
iris
.
base_logic
.
lib
Require
Import
wsat
.
From
iris
.
proofmode
Require
Import
tactics
.
Set
Default
Proof
Using
"Type
*
"
.
Set
Default
Proof
Using
"Type"
.
Import
uPred
.
(* Global functor setup *)
...
...
theories/program_logic/ectx_language.v
View file @
fbf07f03
...
...
@@ -2,7 +2,7 @@
that this gives rise to a "language" in the Iris sense. *)
From
iris
.
algebra
Require
Export
base
.
From
iris
.
program_logic
Require
Import
language
.
Set
Default
Proof
Using
"Type
*
"
.
Set
Default
Proof
Using
"Type"
.
(* We need to make thos arguments indices that we want canonical structure
inference to use a keys. *)
...
...
theories/program_logic/ectx_lifting.v
View file @
fbf07f03
(** Some derived lemmas for ectx-based languages *)
From
iris
.
program_logic
Require
Export
ectx_language
weakestpre
lifting
.
From
iris
.
proofmode
Require
Import
tactics
.
(* FIXME: This file needs a 'Proof Using' hint, but the default we use
everywhere makes for lots of extra ssumptions. *)
Set
Default
Proof
Using
"Type"
.
Section
wp
.
Context
{
expr
val
ectx
state
}
{
Λ
:
EctxLanguage
expr
val
ectx
state
}.
Context
`
{
irisG
(
ectx_lang
expr
)
Σ
}
`
{
Inhabited
state
}.
Context
`
{
irisG
(
ectx_lang
expr
)
Σ
}
{
Hinh
:
Inhabited
state
}.
Implicit
Types
P
:
iProp
Σ
.
Implicit
Types
Φ
:
val
→
iProp
Σ
.
Implicit
Types
v
:
val
.
...
...
@@ -37,7 +36,7 @@ Lemma wp_lift_pure_head_step E Φ e1 :
(
▷
∀
e2
efs
σ
,
⌜
head_step
e1
σ
e2
σ
efs
⌝
→
WP
e2
@
E
{{
Φ
}}
∗
[
∗
list
]
ef
∈
efs
,
WP
ef
{{
_
,
True
}})
⊢
WP
e1
@
E
{{
Φ
}}.
Proof
.
Proof
using
Hinh
.
iIntros
(??)
"H"
.
iApply
wp_lift_pure_step
;
eauto
.
iNext
.
iIntros
(????).
iApply
"H"
.
eauto
.
Qed
.
...
...
@@ -77,7 +76,7 @@ Lemma wp_lift_pure_det_head_step {E Φ} e1 e2 efs :
head_step
e1
σ
1 e2
'
σ
2
efs'
→
σ
1
=
σ
2
∧
e2
=
e2'
∧
efs
=
efs'
)
→
▷
(
WP
e2
@
E
{{
Φ
}}
∗
[
∗
list
]
ef
∈
efs
,
WP
ef
{{
_
,
True
}})
⊢
WP
e1
@
E
{{
Φ
}}.
Proof
.
eauto
using
wp_lift_pure_det_step
.
Qed
.
Proof
using
Hinh
.
eauto
using
wp_lift_pure_det_step
.
Qed
.
Lemma
wp_lift_pure_det_head_step_no_fork
{
E
Φ
}
e1
e2
:
to_val
e1
=
None
→
...
...
@@ -85,7 +84,7 @@ Lemma wp_lift_pure_det_head_step_no_fork {E Φ} e1 e2 :
(
∀
σ
1 e2
'
σ
2
efs'
,
head_step
e1
σ
1 e2
'
σ
2
efs'
→
σ
1
=
σ
2
∧
e2
=
e2'
∧
[]
=
efs'
)
→
▷
WP
e2
@
E
{{
Φ
}}
⊢
WP
e1
@
E
{{
Φ
}}.
Proof
.
Proof
using
Hinh
.
intros
.
rewrite
-(
wp_lift_pure_det_step
e1
e2
[])
?big_sepL_nil
?right_id
;
eauto
.
Qed
.
End
wp
.
theories/program_logic/ectxi_language.v
View file @
fbf07f03
...
...
@@ -2,7 +2,7 @@
a proof that these are instances of general ectx-based languages. *)
From
iris
.
algebra
Require
Export
base
.
From
iris
.
program_logic
Require
Import
language
ectx_language
.
Set
Default
Proof
Using
"Type
*
"
.
Set
Default
Proof
Using
"Type"
.
(* We need to make thos arguments indices that we want canonical structure
inference to use a keys. *)
...
...
theories/program_logic/hoare.v
View file @
fbf07f03
From
iris
.
program_logic
Require
Export
weakestpre
.
From
iris
.
base_logic
.
lib
Require
Export
viewshifts
.
From
iris
.
proofmode
Require
Import
tactics
.
Set
Default
Proof
Using
"Type
*
"
.
Set
Default
Proof
Using
"Type"
.
Definition
ht
`
{
irisG
Λ
Σ
}
(
E
:
coPset
)
(
P
:
iProp
Σ
)
(
e
:
expr
Λ
)
(
Φ
:
val
Λ
→
iProp
Σ
)
:
iProp
Σ
:
=
...
...
theories/program_logic/language.v
View file @
fbf07f03
From
iris
.
algebra
Require
Export
ofe
.
Set
Default
Proof
Using
"Type
*
"
.
Set
Default
Proof
Using
"Type"
.
Structure
language
:
=
Language
{
expr
:
Type
;
...
...
theories/program_logic/lifting.v
View file @
fbf07f03
From
iris
.
program_logic
Require
Export
weakestpre
.
From
iris
.
base_logic
Require
Export
big_op
.
From
iris
.
proofmode
Require
Import
tactics
.
Set
Default
Proof
Using
"Type
*
"
.
Set
Default
Proof
Using
"Type"
.
Section
lifting
.
Context
`
{
irisG
Λ
Σ
}.
...
...
theories/program_logic/ownp.v
View file @
fbf07f03
...
...
@@ -3,7 +3,7 @@ From iris.program_logic Require Import lifting adequacy.
From
iris
.
program_logic
Require
ectx_language
.
From
iris
.
algebra
Require
Import
auth
.
From
iris
.
proofmode
Require
Import
tactics
classes
.
Set
Default
Proof
Using
"Type
*
"
.
Set
Default
Proof
Using
"Type"
.
Class
ownPG'
(
Λ
state
:
Type
)
(
Σ
:
gFunctors
)
:
=
OwnPG
{
ownP_invG
:
invG
Σ
;
...
...
@@ -158,7 +158,7 @@ End lifting.
Section
ectx_lifting
.
Import
ectx_language
.
Context
{
expr
val
ectx
state
}
{
Λ
:
EctxLanguage
expr
val
ectx
state
}.
Context
`
{
ownPG
(
ectx_lang
expr
)
Σ
}
`
{
Inhabited
state
}.
Context
`
{
ownPG
(
ectx_lang
expr
)
Σ
}
{
Hinh
:
Inhabited
state
}.
Implicit
Types
Φ
:
val
→
iProp
Σ
.
Implicit
Types
e
:
expr
.
Hint
Resolve
head_prim_reducible
head_reducible_prim_step
.
...
...
@@ -181,7 +181,7 @@ Section ectx_lifting.
(
▷
∀
e2
efs
σ
,
⌜
head_step
e1
σ
e2
σ
efs
⌝
→
WP
e2
@
E
{{
Φ
}}
∗
[
∗
list
]
ef
∈
efs
,
WP
ef
{{
_
,
True
}})
⊢
WP
e1
@
E
{{
Φ
}}.
Proof
.
Proof
using
Hinh
.
iIntros
(??)
"H"
.
iApply
ownP_lift_pure_step
;
eauto
.
iNext
.
iIntros
(????).
iApply
"H"
.
eauto
.
Qed
.
...
...
@@ -220,14 +220,14 @@ Section ectx_lifting.
(
∀
σ
1 e2
'
σ
2
efs'
,
head_step
e1
σ
1 e2
'
σ
2
efs'
→
σ
1
=
σ
2
∧
e2
=
e2'
∧
efs
=
efs'
)
→
▷
(
WP
e2
@
E
{{
Φ
}}
∗
[
∗
list
]
ef
∈
efs
,
WP
ef
{{
_
,
True
}})
⊢
WP
e1
@
E
{{
Φ
}}.
Proof
.
eauto
using
wp_lift_pure_det_step
.
Qed
.
Proof
using
Hinh
.
eauto
using
wp_lift_pure_det_step
.
Qed
.
Lemma
ownP_lift_pure_det_head_step_no_fork
{
E
Φ
}
e1
e2
:
to_val
e1
=
None
→
(
∀
σ
1
,
head_reducible
e1
σ
1
)
→
(
∀
σ
1 e2
'
σ
2
efs'
,
head_step
e1
σ
1 e2
'
σ
2
efs'
→
σ
1
=
σ
2
∧
e2
=
e2'
∧
[]
=
efs'
)
→
▷
WP
e2
@
E
{{
Φ
}}
⊢
WP
e1
@
E
{{
Φ
}}.
Proof
.
Proof
using
Hinh
.
intros
.
rewrite
-(
wp_lift_pure_det_step
e1
e2
[])
?big_sepL_nil
?right_id
;
eauto
.
Qed
.
End
ectx_lifting
.
theories/program_logic/weakestpre.v
View file @
fbf07f03
...
...
@@ -2,7 +2,7 @@ From iris.base_logic.lib Require Export fancy_updates.
From
iris
.
program_logic
Require
Export
language
.
From
iris
.
base_logic
Require
Import
big_op
.
From
iris
.
proofmode
Require
Import
tactics
classes
.
Set
Default
Proof
Using
"Type
*
"
.
Set
Default
Proof
Using
"Type"
.
Import
uPred
.
Class
irisG'
(
Λ
state
:
Type
)
(
Σ
:
gFunctors
)
:
=
IrisG
{
...
...
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