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
Rice Wine
Iris
Commits
be81fb92
Commit
be81fb92
authored
Jan 09, 2017
by
Ralf Jung
Browse files
add a tactic to automatically solve goals that show inG from subG
parent
673f0a14
Changes
17
Hide whitespace changes
Inline
Side-by-side
theories/base_logic/lib/auth.v
View file @
be81fb92
...
...
@@ -14,7 +14,7 @@ Class authG Σ (A : ucmraT) := AuthG {
Definition
auth
Σ
(
A
:
ucmraT
)
:
gFunctors
:
=
#[
GFunctor
(
authR
A
)
].
Instance
subG_auth
Σ
Σ
A
:
subG
(
auth
Σ
A
)
Σ
→
CMRADiscrete
A
→
authG
Σ
A
.
Proof
.
intros
?%
subG_inG
?.
by
split
.
Qed
.
Proof
.
solve_inG
.
Qed
.
Section
definitions
.
Context
`
{
invG
Σ
,
authG
Σ
A
}
{
T
:
Type
}
(
γ
:
gname
).
...
...
theories/base_logic/lib/boxes.v
View file @
be81fb92
...
...
@@ -16,7 +16,7 @@ Definition boxΣ : gFunctors := #[ GFunctor (authR (optionUR (exclR boolC)) *
Instance
subG_sts
Σ
Σ
:
subG
box
Σ
Σ
→
boxG
Σ
.
Proof
.
apply
subG
_inG
.
Qed
.
Proof
.
solve
_inG
.
Qed
.
Section
box_defs
.
Context
`
{
invG
Σ
,
boxG
Σ
}
(
N
:
namespace
).
...
...
theories/base_logic/lib/gen_heap.v
View file @
be81fb92
...
...
@@ -24,7 +24,7 @@ Definition gen_heapΣ (L V : Type) `{Countable L} : gFunctors :=
Instance
subG_gen_heapPreG
{
Σ
L
V
}
`
{
Countable
L
}
:
subG
(
gen_heap
Σ
L
V
)
Σ
→
gen_heapPreG
L
V
Σ
.
Proof
.
intros
?%
subG_inG
;
split
;
apply
_
.
Qed
.
Proof
.
solve_inG
.
Qed
.
Section
definitions
.
Context
`
{
gen_heapG
L
V
Σ
}.
...
...
theories/base_logic/lib/own.v
View file @
be81fb92
...
...
@@ -17,6 +17,33 @@ Arguments inG_id {_ _} _.
Lemma
subG_inG
Σ
(
F
:
gFunctor
)
:
subG
F
Σ
→
inG
Σ
(
F
(
iPreProp
Σ
)).
Proof
.
move
=>
/(
_
0
%
fin
)
/=
[
j
->].
by
exists
j
.
Qed
.
(** This tactic solves the usual obligations "subG ? Σ → {in,?}G ? Σ" *)
Ltac
solve_inG
:
=
(* Get all assumptions *)
intros
;
(* Unfold the top-level xΣ *)
lazymatch
goal
with
|
H
:
subG
(
?x
Σ
_
)
_
|-
_
=>
try
unfold
x
Σ
in
H
|
H
:
subG
?x
Σ
_
|-
_
=>
try
unfold
x
Σ
in
H
end
;
(* Take apart subG for non-"atomic" lists *)
repeat
match
goal
with
|
H
:
subG
(
gFunctors
.
app
_
_
)
_
|-
_
=>
apply
subG_inv
in
H
;
destruct
H
end
;
(* Try to turn singleton subG into inG; but also keep the subG for typeclass
resolution -- to keep them, we put them onto the goal. *)
repeat
match
goal
with
|
H
:
subG
_
_
|-
_
=>
move
:
(
H
)
;
(
apply
subG_inG
in
H
||
clear
H
)
end
;
(* Again get all assumptions *)
intros
;
(* We support two kinds of goals: Things convertible to inG;
and records with inG and typeclass fields. Try to solve the
first case. *)
try
done
;
(* That didn't work, now we're in for the second case. *)
split
;
(
assumption
||
by
apply
_
).
(** * Definition of the connective [own] *)
Definition
iRes_singleton
`
{
i
:
inG
Σ
A
}
(
γ
:
gname
)
(
a
:
A
)
:
iResUR
Σ
:
=
iprod_singleton
(
inG_id
i
)
{[
γ
:
=
cmra_transport
inG_prf
a
]}.
...
...
theories/base_logic/lib/saved_prop.v
View file @
be81fb92
...
...
@@ -10,7 +10,7 @@ Definition savedPropΣ (F : cFunctor) : gFunctors :=
#[
GFunctor
(
agreeRF
(
▶
F
))
].
Instance
subG_savedProp
Σ
{
Σ
F
}
:
subG
(
savedProp
Σ
F
)
Σ
→
savedPropG
Σ
F
.
Proof
.
apply
subG
_inG
.
Qed
.
Proof
.
solve
_inG
.
Qed
.
Definition
saved_prop_own
`
{
savedPropG
Σ
F
}
(
γ
:
gname
)
(
x
:
F
(
iProp
Σ
))
:
iProp
Σ
:
=
...
...
theories/base_logic/lib/sts.v
View file @
be81fb92
...
...
@@ -13,7 +13,7 @@ Class stsG Σ (sts : stsT) := StsG {
Definition
sts
Σ
(
sts
:
stsT
)
:
gFunctors
:
=
#[
GFunctor
(
stsR
sts
)
].
Instance
subG_sts
Σ
Σ
sts
:
subG
(
sts
Σ
sts
)
Σ
→
Inhabited
(
sts
.
state
sts
)
→
stsG
Σ
sts
.
Proof
.
intros
?%
subG_inG
?.
by
split
.
Qed
.
Proof
.
solve_inG
.
Qed
.
Section
definitions
.
Context
`
{
stsG
Σ
sts
}
(
γ
:
gname
).
...
...
theories/heap_lang/adequacy.v
View file @
be81fb92
...
...
@@ -12,7 +12,7 @@ Class heapPreG Σ := HeapPreG {
Definition
heap
Σ
:
gFunctors
:
=
#[
inv
Σ
;
gen_heap
Σ
loc
val
].
Instance
subG_heapPreG
{
Σ
}
:
subG
heap
Σ
Σ
→
heapPreG
Σ
.
Proof
.
intros
[?
?]%
subG_inv
;
split
;
apply
_
.
Qed
.
Proof
.
solve_inG
.
Qed
.
Definition
heap_adequacy
Σ
`
{
heapPreG
Σ
}
e
σ
φ
:
(
∀
`
{
heapG
Σ
},
True
⊢
WP
e
{{
v
,
⌜φ
v
⌝
}})
→
...
...
theories/heap_lang/lib/barrier/proof.v
View file @
be81fb92
...
...
@@ -8,7 +8,6 @@ From iris.heap_lang.lib.barrier Require Import protocol.
Set
Default
Proof
Using
"Type"
.
(** The CMRAs/functors we need. *)
(* Not bundling heapG, as it may be shared with other users. *)
Class
barrierG
Σ
:
=
BarrierG
{
barrier_stsG
:
>
stsG
Σ
sts
;
barrier_savedPropG
:
>
savedPropG
Σ
idCF
;
...
...
@@ -16,7 +15,7 @@ Class barrierG Σ := BarrierG {
Definition
barrier
Σ
:
gFunctors
:
=
#[
sts
Σ
sts
;
savedProp
Σ
idCF
].
Instance
subG_barrier
Σ
{
Σ
}
:
subG
barrier
Σ
Σ
→
barrierG
Σ
.
Proof
.
intros
[?
[?
_
]%
subG_inv
]%
subG_inv
.
split
;
apply
_
.
Qed
.
Proof
.
solve_inG
.
Qed
.
(** Now we come to the Iris part of the proof. *)
Section
proof
.
...
...
theories/heap_lang/lib/counter.v
View file @
be81fb92
...
...
@@ -17,7 +17,7 @@ Class mcounterG Σ := MCounterG { mcounter_inG :> inG Σ (authR mnatUR) }.
Definition
mcounter
Σ
:
gFunctors
:
=
#[
GFunctor
(
authR
mnatUR
)].
Instance
subG_mcounter
Σ
{
Σ
}
:
subG
mcounter
Σ
Σ
→
mcounterG
Σ
.
Proof
.
intros
[?%
subG_inG
_
]%
subG_inv
.
split
;
apply
_
.
Qed
.
Proof
.
solve_inG
.
Qed
.
Section
mono_proof
.
Context
`
{!
heapG
Σ
,
!
mcounterG
Σ
}
(
N
:
namespace
).
...
...
theories/heap_lang/lib/spawn.v
View file @
be81fb92
...
...
@@ -23,7 +23,7 @@ Class spawnG Σ := SpawnG { spawn_tokG :> inG Σ (exclR unitC) }.
Definition
spawn
Σ
:
gFunctors
:
=
#[
GFunctor
(
exclR
unitC
)].
Instance
subG_spawn
Σ
{
Σ
}
:
subG
spawn
Σ
Σ
→
spawnG
Σ
.
Proof
.
intros
[?%
subG_inG
_
]%
subG_inv
.
split
;
apply
_
.
Qed
.
Proof
.
solve_inG
.
Qed
.
(** Now we come to the Iris part of the proof. *)
Section
proof
.
...
...
theories/heap_lang/lib/spin_lock.v
View file @
be81fb92
...
...
@@ -18,7 +18,7 @@ Class lockG Σ := LockG { lock_tokG :> inG Σ (exclR unitC) }.
Definition
lock
Σ
:
gFunctors
:
=
#[
GFunctor
(
exclR
unitC
)].
Instance
subG_lock
Σ
{
Σ
}
:
subG
lock
Σ
Σ
→
lockG
Σ
.
Proof
.
intros
?%
subG_inG
.
split
;
apply
_
.
Qed
.
Proof
.
solve_inG
.
Qed
.
Section
proof
.
Context
`
{!
heapG
Σ
,
!
lockG
Σ
}
(
N
:
namespace
).
...
...
theories/heap_lang/lib/ticket_lock.v
View file @
be81fb92
...
...
@@ -34,7 +34,7 @@ Definition tlockΣ : gFunctors :=
#[
GFunctor
(
authR
(
prodUR
(
optionUR
(
exclR
natC
))
(
gset_disjUR
nat
)))
].
Instance
subG_tlock
Σ
{
Σ
}
:
subG
tlock
Σ
Σ
→
tlockG
Σ
.
Proof
.
by
intros
?%
subG
_inG
.
Qed
.
Proof
.
solve
_inG
.
Qed
.
Section
proof
.
Context
`
{!
heapG
Σ
,
!
tlockG
Σ
}
(
N
:
namespace
).
...
...
theories/program_logic/adequacy.v
View file @
be81fb92
...
...
@@ -19,9 +19,7 @@ Class invPreG (Σ : gFunctors) : Set := WsatPreG {
}.
Instance
subG_inv
Σ
{
Σ
}
:
subG
inv
Σ
Σ
→
invPreG
Σ
.
Proof
.
intros
[?%
subG_inG
[?%
subG_inG
?%
subG_inG
]%
subG_inv
]%
subG_inv
;
by
constructor
.
Qed
.
Proof
.
solve_inG
.
Qed
.
(* Allocation *)
Lemma
wsat_alloc
`
{
invPreG
Σ
}
:
(|==>
∃
_
:
invG
Σ
,
wsat
∗
ownE
⊤
)%
I
.
...
...
theories/program_logic/ownp.v
View file @
be81fb92
...
...
@@ -29,7 +29,7 @@ Class ownPPreG' (Λstate : Type) (Σ : gFunctors) : Set := IrisPreG {
Notation
ownPPreG
Λ
Σ
:
=
(
ownPPreG'
(
state
Λ
)
Σ
).
Instance
subG_ownP
Σ
{
Λ
state
Σ
}
:
subG
(
ownP
Σ
Λ
state
)
Σ
→
ownPPreG'
Λ
state
Σ
.
Proof
.
intros
[??%
subG_inG
]%
subG_inv
;
constructor
;
apply
_
.
Qed
.
Proof
.
solve_inG
.
Qed
.
(** Ownership *)
Definition
ownP
`
{
ownPG'
Λ
state
Σ
}
(
σ
:
Λ
state
)
:
iProp
Σ
:
=
...
...
theories/tests/counter.v
View file @
be81fb92
...
...
@@ -75,7 +75,7 @@ End M.
Class
counterG
Σ
:
=
CounterG
{
counter_tokG
:
>
inG
Σ
M_UR
}.
Definition
counter
Σ
:
gFunctors
:
=
#[
GFunctor
M_UR
].
Instance
subG_counter
Σ
{
Σ
}
:
subG
counter
Σ
Σ
→
counterG
Σ
.
Proof
.
intros
[?%
subG_inG
_
]%
subG_inv
.
split
;
apply
_
.
Qed
.
Proof
.
solve_inG
.
Qed
.
Section
proof
.
Context
`
{!
heapG
Σ
,
!
counterG
Σ
}.
...
...
theories/tests/joining_existentials.v
View file @
be81fb92
...
...
@@ -17,7 +17,7 @@ Class oneShotG (Σ : gFunctors) (F : cFunctor) :=
Definition
oneShot
Σ
(
F
:
cFunctor
)
:
gFunctors
:
=
#[
GFunctor
(
csumRF
(
exclRF
unitC
)
(
agreeRF
(
▶
F
)))
].
Instance
subG_oneShot
Σ
{
Σ
F
}
:
subG
(
oneShot
Σ
F
)
Σ
→
oneShotG
Σ
F
.
Proof
.
apply
subG
_inG
.
Qed
.
Proof
.
solve
_inG
.
Qed
.
Definition
client
eM
eW1
eW2
:
expr
:
=
let
:
"b"
:
=
newbarrier
#()
in
...
...
theories/tests/one_shot.v
View file @
be81fb92
...
...
@@ -27,7 +27,7 @@ Definition Shot (n : Z) : one_shotR := (Cinr (to_agree n) : one_shotR).
Class
one_shotG
Σ
:
=
{
one_shot_inG
:
>
inG
Σ
one_shotR
}.
Definition
one_shot
Σ
:
gFunctors
:
=
#[
GFunctor
one_shotR
].
Instance
subG_one_shot
Σ
{
Σ
}
:
subG
one_shot
Σ
Σ
→
one_shotG
Σ
.
Proof
.
intros
[?%
subG_inG
_
]%
subG_inv
.
split
;
apply
_
.
Qed
.
Proof
.
solve_inG
.
Qed
.
Section
proof
.
Set
Default
Proof
Using
"Type*"
.
...
...
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