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
Tej Chajed
iris
Commits
d6b49ab2
Commit
d6b49ab2
authored
Jan 05, 2017
by
Ralf Jung
Browse files
more restrictive Proof Using hints in base_logic, algebra
parent
60d82286
Changes
49
Hide whitespace changes
Inline
Side-by-side
theories/algebra/agree.v
View file @
d6b49ab2
...
...
@@ -208,7 +208,7 @@ Section list_theory.
Lemma
list_agrees_fmap
`
{
Equivalence
_
R'
}
al
:
list_agrees
R
al
→
list_agrees
R'
(
f
<$>
al
).
Proof
using
All
.
Proof
using
Type
*
.
move
=>
/
list_agrees_alt
Hl
.
apply
(
list_agrees_alt
R'
)
=>
a'
b'
.
intros
(
a
&
->
&
Ha
)%
elem_of_list_fmap
(
b
&
->
&
Hb
)%
elem_of_list_fmap
.
apply
Hf
.
exact
:
Hl
.
...
...
theories/algebra/auth.v
View file @
d6b49ab2
From
iris
.
algebra
Require
Export
excl
local_updates
.
From
iris
.
base_logic
Require
Import
base_logic
.
From
iris
.
proofmode
Require
Import
classes
.
Set
Default
Proof
Using
"Type
*
"
.
Set
Default
Proof
Using
"Type"
.
Record
auth
(
A
:
Type
)
:
=
Auth
{
authoritative
:
excl'
A
;
auth_own
:
A
}.
Add
Printing
Constructor
auth
.
...
...
theories/algebra/base.v
View file @
d6b49ab2
From
mathcomp
Require
Export
ssreflect
.
From
iris
.
prelude
Require
Export
prelude
.
Set
Default
Proof
Using
"Type
*
"
.
Set
Default
Proof
Using
"Type"
.
Global
Set
Bullet
Behavior
"Strict Subproofs"
.
Global
Open
Scope
general_if_scope
.
Ltac
done
:
=
prelude
.
tactics
.
done
.
theories/algebra/cmra.v
View file @
d6b49ab2
From
iris
.
algebra
Require
Export
ofe
.
Set
Default
Proof
Using
"Type
*
"
.
Set
Default
Proof
Using
"Type"
.
Class
PCore
(
A
:
Type
)
:
=
pcore
:
A
→
option
A
.
Instance
:
Params
(@
pcore
)
2
.
...
...
@@ -428,6 +428,7 @@ Qed.
(** ** Total core *)
Section
total_core
.
Set
Default
Proof
Using
"Type*"
.
Context
`
{
CMRATotal
A
}.
Lemma
cmra_core_l
x
:
core
x
⋅
x
≡
x
.
...
...
theories/algebra/cmra_big_op.v
View file @
d6b49ab2
From
iris
.
algebra
Require
Export
cmra
list
.
From
iris
.
prelude
Require
Import
functions
gmap
gmultiset
.
Set
Default
Proof
Using
"Type
*
"
.
Set
Default
Proof
Using
"Type"
.
(** The operator [ [⋅] Ps ] folds [⋅] over the list [Ps]. This operator is not a
quantifier, so it binds strongly.
...
...
theories/algebra/cmra_tactics.v
View file @
d6b49ab2
From
iris
.
algebra
Require
Export
cmra
.
From
iris
.
algebra
Require
Import
cmra_big_op
.
Set
Default
Proof
Using
"Type
*
"
.
Set
Default
Proof
Using
"Type"
.
(** * Simple solver for validity and inclusion by reflection *)
Module
ra_reflection
.
Section
ra_reflection
.
...
...
theories/algebra/coPset.v
View file @
d6b49ab2
From
iris
.
algebra
Require
Export
cmra
.
From
iris
.
algebra
Require
Import
updates
local_updates
.
From
iris
.
prelude
Require
Export
collections
coPset
.
Set
Default
Proof
Using
"Type
*
"
.
Set
Default
Proof
Using
"Type"
.
(** This is pretty much the same as algebra/gset, but I was not able to
generalize the construction without breaking canonical structures. *)
...
...
theories/algebra/cofe_solver.v
View file @
d6b49ab2
...
...
@@ -205,7 +205,7 @@ Instance fold_ne : Proper (dist n ==> dist n) fold.
Proof
.
by
intros
n
X
Y
HXY
k
;
rewrite
/
fold
/=
HXY
.
Qed
.
Theorem
result
:
solution
F
.
Proof
using
All
.
Proof
using
Type
*
.
apply
(
Solution
F
T
_
(
CofeMor
unfold
)
(
CofeMor
fold
)).
-
move
=>
X
/=.
rewrite
equiv_dist
=>
n
k
;
rewrite
/
unfold
/
fold
/=.
rewrite
-
g_tower
-(
gg_tower
_
n
)
;
apply
(
_
:
Proper
(
_
==>
_
)
(
g
_
)).
...
...
theories/algebra/csum.v
View file @
d6b49ab2
From
iris
.
algebra
Require
Export
cmra
.
From
iris
.
base_logic
Require
Import
base_logic
.
From
iris
.
algebra
Require
Import
local_updates
.
Set
Default
Proof
Using
"Type
*
"
.
Set
Default
Proof
Using
"Type"
.
Local
Arguments
pcore
_
_
!
_
/.
Local
Arguments
cmra_pcore
_
!
_
/.
Local
Arguments
validN
_
_
_
!
_
/.
...
...
theories/algebra/deprecated.v
View file @
d6b49ab2
From
iris
.
algebra
Require
Import
ofe
cmra
.
Set
Default
Proof
Using
"Type
*
"
.
Set
Default
Proof
Using
"Type"
.
(* Old notation for backwards compatibility. *)
...
...
theories/algebra/dra.v
View file @
d6b49ab2
From
iris
.
algebra
Require
Export
cmra
updates
.
Set
Default
Proof
Using
"Type
*
"
.
Set
Default
Proof
Using
"Type"
.
Record
DRAMixin
A
`
{
Equiv
A
,
Core
A
,
Disjoint
A
,
Op
A
,
Valid
A
}
:
=
{
(* setoids *)
...
...
theories/algebra/excl.v
View file @
d6b49ab2
From
iris
.
algebra
Require
Export
cmra
.
From
iris
.
base_logic
Require
Import
base_logic
.
Set
Default
Proof
Using
"Type
*
"
.
Set
Default
Proof
Using
"Type"
.
Local
Arguments
validN
_
_
_
!
_
/.
Local
Arguments
valid
_
_
!
_
/.
...
...
theories/algebra/frac.v
View file @
d6b49ab2
From
Coq
.
QArith
Require
Import
Qcanon
.
From
iris
.
algebra
Require
Export
cmra
.
Set
Default
Proof
Using
"Type
*
"
.
Set
Default
Proof
Using
"Type"
.
Notation
frac
:
=
Qp
(
only
parsing
).
...
...
theories/algebra/gmap.v
View file @
d6b49ab2
...
...
@@ -2,7 +2,7 @@ From iris.algebra Require Export cmra.
From
iris
.
prelude
Require
Export
gmap
.
From
iris
.
algebra
Require
Import
updates
local_updates
.
From
iris
.
base_logic
Require
Import
base_logic
.
Set
Default
Proof
Using
"Type
*
"
.
Set
Default
Proof
Using
"Type"
.
Section
cofe
.
Context
`
{
Countable
K
}
{
A
:
ofeT
}.
...
...
@@ -334,6 +334,7 @@ Proof.
Qed
.
Section
freshness
.
Set
Default
Proof
Using
"Type*"
.
Context
`
{
Fresh
K
(
gset
K
),
!
FreshSpec
K
(
gset
K
)}.
Lemma
alloc_updateP_strong
(
Q
:
gmap
K
A
→
Prop
)
(
I
:
gset
K
)
m
x
:
✓
x
→
(
∀
i
,
m
!!
i
=
None
→
i
∉
I
→
Q
(<[
i
:
=
x
]>
m
))
→
m
~~>
:
Q
.
...
...
theories/algebra/gset.v
View file @
d6b49ab2
From
iris
.
algebra
Require
Export
cmra
.
From
iris
.
algebra
Require
Import
updates
local_updates
.
From
iris
.
prelude
Require
Export
collections
gmap
mapset
.
Set
Default
Proof
Using
"Type
*
"
.
Set
Default
Proof
Using
"Type"
.
(* The union CMRA *)
Section
gset
.
...
...
@@ -155,6 +155,7 @@ Section gset_disj.
Proof
.
eauto
using
gset_disj_alloc_empty_updateP_strong
.
Qed
.
Section
fresh_updates
.
Set
Default
Proof
Using
"Type*"
.
Context
`
{
Fresh
K
(
gset
K
),
!
FreshSpec
K
(
gset
K
)}.
Lemma
gset_disj_alloc_updateP
(
Q
:
gset_disj
K
→
Prop
)
X
:
...
...
theories/algebra/iprod.v
View file @
d6b49ab2
From
iris
.
algebra
Require
Export
cmra
.
From
iris
.
base_logic
Require
Import
base_logic
.
From
iris
.
prelude
Require
Import
finite
.
Set
Default
Proof
Using
"Type
*
"
.
Set
Default
Proof
Using
"Type"
.
(** * Indexed product *)
(** Need to put this in a definition to make canonical structures to work. *)
...
...
theories/algebra/list.v
View file @
d6b49ab2
...
...
@@ -2,7 +2,7 @@ From iris.algebra Require Export cmra.
From
iris
.
prelude
Require
Export
list
.
From
iris
.
base_logic
Require
Import
base_logic
.
From
iris
.
algebra
Require
Import
updates
local_updates
.
Set
Default
Proof
Using
"Type
*
"
.
Set
Default
Proof
Using
"Type"
.
Section
cofe
.
Context
{
A
:
ofeT
}.
...
...
theories/algebra/local_updates.v
View file @
d6b49ab2
From
iris
.
algebra
Require
Export
cmra
.
Set
Default
Proof
Using
"Type
*
"
.
Set
Default
Proof
Using
"Type"
.
(** * Local updates *)
Definition
local_update
{
A
:
cmraT
}
(
x
y
:
A
*
A
)
:
=
∀
n
mz
,
...
...
theories/algebra/ofe.v
View file @
d6b49ab2
From
iris
.
algebra
Require
Export
base
.
Set
Default
Proof
Using
"Type
*
"
.
Set
Default
Proof
Using
"Type"
.
(** This files defines (a shallow embedding of) the category of OFEs:
Complete ordered families of equivalences. This is a cartesian closed
...
...
@@ -159,6 +159,7 @@ Instance const_contractive {A B : ofeT} (x : A) : Contractive (@const A B x).
Proof
.
by
intros
n
y1
y2
.
Qed
.
Section
contractive
.
Set
Default
Proof
Using
"Type*"
.
Context
{
A
B
:
ofeT
}
(
f
:
A
→
B
)
`
{!
Contractive
f
}.
Implicit
Types
x
y
:
A
.
...
...
theories/algebra/sts.v
View file @
d6b49ab2
From
iris
.
prelude
Require
Export
set
.
From
iris
.
algebra
Require
Export
cmra
.
From
iris
.
algebra
Require
Import
dra
.
Set
Default
Proof
Using
"Type
*
"
.
Set
Default
Proof
Using
"Type"
.
Local
Arguments
valid
_
_
!
_
/.
Local
Arguments
op
_
_
!
_
!
_
/.
Local
Arguments
core
_
_
!
_
/.
...
...
Prev
1
2
3
Next
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