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
Rodolphe Lepigre
Iris
Commits
0f11453a
Commit
0f11453a
authored
Nov 29, 2017
by
Robbert Krebbers
Browse files
More consistent variable names and `Implicit Types`.
parent
93315c9e
Changes
2
Hide whitespace changes
Inline
Side-by-side
theories/algebra/cmra.v
View file @
0f11453a
...
...
@@ -1235,92 +1235,93 @@ Qed.
(** ** CMRA for the option type *)
Section
option
.
Context
{
A
:
cmraT
}.
Implicit
Types
a
:
A
.
Implicit
Types
a
b
:
A
.
Implicit
Types
ma
mb
:
option
A
.
Local
Arguments
core
_
_
!
_
/.
Local
Arguments
pcore
_
_
!
_
/.
Instance
option_valid
:
Valid
(
option
A
)
:
=
λ
m
x
,
match
m
x
with
Some
x
=>
✓
x
|
None
=>
True
end
.
Instance
option_validN
:
ValidN
(
option
A
)
:
=
λ
n
m
x
,
match
m
x
with
Some
x
=>
✓
{
n
}
x
|
None
=>
True
end
.
Instance
option_pcore
:
PCore
(
option
A
)
:
=
λ
m
x
,
Some
(
m
x
≫
=
pcore
).
Instance
option_valid
:
Valid
(
option
A
)
:
=
λ
m
a
,
match
m
a
with
Some
a
=>
✓
a
|
None
=>
True
end
.
Instance
option_validN
:
ValidN
(
option
A
)
:
=
λ
n
m
a
,
match
m
a
with
Some
a
=>
✓
{
n
}
a
|
None
=>
True
end
.
Instance
option_pcore
:
PCore
(
option
A
)
:
=
λ
m
a
,
Some
(
m
a
≫
=
pcore
).
Arguments
option_pcore
!
_
/.
Instance
option_op
:
Op
(
option
A
)
:
=
union_with
(
λ
x
y
,
Some
(
x
⋅
y
)).
Instance
option_op
:
Op
(
option
A
)
:
=
union_with
(
λ
a
b
,
Some
(
a
⋅
b
)).
Definition
Some_valid
a
:
✓
Some
a
↔
✓
a
:
=
reflexivity
_
.
Definition
Some_validN
a
n
:
✓
{
n
}
Some
a
↔
✓
{
n
}
a
:
=
reflexivity
_
.
Definition
Some_op
a
b
:
Some
(
a
⋅
b
)
=
Some
a
⋅
Some
b
:
=
eq_refl
.
Lemma
Some_core
`
{
CmraTotal
A
}
a
:
Some
(
core
a
)
=
core
(
Some
a
).
Proof
.
rewrite
/
core
/=.
by
destruct
(
cmra_total
a
)
as
[?
->].
Qed
.
Lemma
Some_op_opM
x
m
y
:
Some
x
⋅
m
y
=
Some
(
x
⋅
?
m
y
).
Proof
.
by
destruct
m
y
.
Qed
.
Lemma
Some_op_opM
a
m
a
:
Some
a
⋅
m
a
=
Some
(
a
⋅
?
m
a
).
Proof
.
by
destruct
m
a
.
Qed
.
Lemma
option_included
(
mx
my
:
option
A
)
:
m
x
≼
m
y
↔
m
x
=
None
∨
∃
x
y
,
m
x
=
Some
x
∧
m
y
=
Some
y
∧
(
x
≡
y
∨
x
≼
y
).
Lemma
option_included
ma
mb
:
m
a
≼
m
b
↔
m
a
=
None
∨
∃
a
b
,
m
a
=
Some
a
∧
m
b
=
Some
b
∧
(
a
≡
b
∨
a
≼
b
).
Proof
.
split
.
-
intros
[
m
z
Hm
z
].
destruct
m
x
as
[
x
|]
;
[
right
|
by
left
].
destruct
m
y
as
[
y
|]
;
[
exists
x
,
y
|
destruct
m
z
;
inversion_clear
Hm
z
].
destruct
m
z
as
[
z
|]
;
inversion_clear
Hm
z
;
split_and
?
;
auto
;
-
intros
[
m
c
Hm
c
].
destruct
m
a
as
[
a
|]
;
[
right
|
by
left
].
destruct
m
b
as
[
b
|]
;
[
exists
a
,
b
|
destruct
m
c
;
inversion_clear
Hm
c
].
destruct
m
c
as
[
c
|]
;
inversion_clear
Hm
c
;
split_and
?
;
auto
;
setoid_subst
;
eauto
using
cmra_included_l
.
-
intros
[->|(
x
&
y
&->&->&[
H
z
|[
z
H
z
]])].
+
exists
m
y
.
by
destruct
m
y
.
-
intros
[->|(
a
&
b
&->&->&[
H
c
|[
c
H
c
]])].
+
exists
m
b
.
by
destruct
m
b
.
+
exists
None
;
by
constructor
.
+
exists
(
Some
z
)
;
by
constructor
.
+
exists
(
Some
c
)
;
by
constructor
.
Qed
.
Lemma
option_includedN
n
(
mx
my
:
option
A
)
:
m
x
≼
{
n
}
m
y
↔
m
x
=
None
∨
∃
x
y
,
m
x
=
Some
x
∧
m
y
=
Some
y
∧
(
x
≡
{
n
}
≡
y
∨
x
≼
{
n
}
y
).
Lemma
option_includedN
n
ma
mb
:
m
a
≼
{
n
}
m
b
↔
m
a
=
None
∨
∃
x
y
,
m
a
=
Some
x
∧
m
b
=
Some
y
∧
(
x
≡
{
n
}
≡
y
∨
x
≼
{
n
}
y
).
Proof
.
split
.
-
intros
[
m
z
Hm
z
].
destruct
m
x
as
[
x
|]
;
[
right
|
by
left
].
destruct
m
y
as
[
y
|]
;
[
exists
x
,
y
|
destruct
m
z
;
inversion_clear
Hm
z
].
destruct
m
z
as
[
z
|]
;
inversion_clear
Hm
z
;
split_and
?
;
auto
;
-
intros
[
m
c
Hm
c
].
destruct
m
a
as
[
a
|]
;
[
right
|
by
left
].
destruct
m
b
as
[
b
|]
;
[
exists
a
,
b
|
destruct
m
c
;
inversion_clear
Hm
c
].
destruct
m
c
as
[
c
|]
;
inversion_clear
Hm
c
;
split_and
?
;
auto
;
ofe_subst
;
eauto
using
cmra_includedN_l
.
-
intros
[->|(
x
&
y
&->&->&[
H
z
|[
z
H
z
]])].
+
exists
m
y
.
by
destruct
m
y
.
-
intros
[->|(
a
&
y
&->&->&[
H
c
|[
c
H
c
]])].
+
exists
m
b
.
by
destruct
m
b
.
+
exists
None
;
by
constructor
.
+
exists
(
Some
z
)
;
by
constructor
.
+
exists
(
Some
c
)
;
by
constructor
.
Qed
.
Lemma
option_cmra_mixin
:
CmraMixin
(
option
A
).
Proof
.
apply
cmra_total_mixin
.
-
eauto
.
-
by
intros
[
x
|]
n
;
destruct
1
;
constructor
;
ofe_subst
.
-
by
intros
[
a
|]
n
;
destruct
1
;
constructor
;
ofe_subst
.
-
destruct
1
;
by
ofe_subst
.
-
by
destruct
1
;
rewrite
/
validN
/
option_validN
//=
;
ofe_subst
.
-
intros
[
x
|]
;
[
apply
cmra_valid_validN
|
done
].
-
intros
n
[
x
|]
;
unfold
validN
,
option_validN
;
eauto
using
cmra_validN_S
.
-
intros
[
x
|]
[
y
|]
[
z
|]
;
constructor
;
rewrite
?assoc
;
auto
.
-
intros
[
x
|]
[
y
|]
;
constructor
;
rewrite
1
?comm
;
auto
.
-
intros
[
x
|]
;
simpl
;
auto
.
destruct
(
pcore
x
)
as
[
c
x
|]
eqn
:
?
;
constructor
;
eauto
using
cmra_pcore_l
.
-
intros
[
x
|]
;
simpl
;
auto
.
destruct
(
pcore
x
)
as
[
c
x
|]
eqn
:
?
;
simpl
;
eauto
using
cmra_pcore_idemp
.
-
intros
m
x
m
y
;
setoid_rewrite
option_included
.
intros
[->|(
x
&
y
&->&->&[?|?])]
;
simpl
;
eauto
.
+
destruct
(
pcore
x
)
as
[
c
x
|]
eqn
:
?
;
eauto
.
destruct
(
cmra_pcore_proper
x
y
c
x
)
as
(?&?&?)
;
eauto
10
.
+
destruct
(
pcore
x
)
as
[
c
x
|]
eqn
:
?
;
eauto
.
destruct
(
cmra_pcore_mono
x
y
c
x
)
as
(?&?&?)
;
eauto
10
.
-
intros
n
[
x
|]
[
y
|]
;
rewrite
/
validN
/
option_validN
/=
;
-
intros
[
a
|]
;
[
apply
cmra_valid_validN
|
done
].
-
intros
n
[
a
|]
;
unfold
validN
,
option_validN
;
eauto
using
cmra_validN_S
.
-
intros
[
a
|]
[
b
|]
[
c
|]
;
constructor
;
rewrite
?assoc
;
auto
.
-
intros
[
a
|]
[
b
|]
;
constructor
;
rewrite
1
?comm
;
auto
.
-
intros
[
a
|]
;
simpl
;
auto
.
destruct
(
pcore
a
)
as
[
c
a
|]
eqn
:
?
;
constructor
;
eauto
using
cmra_pcore_l
.
-
intros
[
a
|]
;
simpl
;
auto
.
destruct
(
pcore
a
)
as
[
c
a
|]
eqn
:
?
;
simpl
;
eauto
using
cmra_pcore_idemp
.
-
intros
m
a
m
b
;
setoid_rewrite
option_included
.
intros
[->|(
a
&
b
&->&->&[?|?])]
;
simpl
;
eauto
.
+
destruct
(
pcore
a
)
as
[
c
a
|]
eqn
:
?
;
eauto
.
destruct
(
cmra_pcore_proper
a
b
c
a
)
as
(?&?&?)
;
eauto
10
.
+
destruct
(
pcore
a
)
as
[
c
a
|]
eqn
:
?
;
eauto
.
destruct
(
cmra_pcore_mono
a
b
c
a
)
as
(?&?&?)
;
eauto
10
.
-
intros
n
[
a
|]
[
b
|]
;
rewrite
/
validN
/
option_validN
/=
;
eauto
using
cmra_validN_op_l
.
-
intros
n
m
x
m
y
1
m
y
2
.
destruct
m
x
as
[
x
|],
m
y
1
as
[
y
1
|],
m
y
2
as
[
y
2
|]
;
intros
Hx
Hx'
;
-
intros
n
m
a
m
b
1
m
b
2
.
destruct
m
a
as
[
a
|],
m
b
1
as
[
b
1
|],
m
b
2
as
[
b
2
|]
;
intros
Hx
Hx'
;
inversion_clear
Hx'
;
auto
.
+
destruct
(
cmra_extend
n
x
y
1
y
2
)
as
(
z
1
&
z
2
&?&?&?)
;
auto
.
by
exists
(
Some
z
1
),
(
Some
z
2
)
;
repeat
constructor
.
+
by
exists
(
Some
x
),
None
;
repeat
constructor
.
+
by
exists
None
,
(
Some
x
)
;
repeat
constructor
.
+
destruct
(
cmra_extend
n
a
b
1
b
2
)
as
(
c
1
&
c
2
&?&?&?)
;
auto
.
by
exists
(
Some
c
1
),
(
Some
c
2
)
;
repeat
constructor
.
+
by
exists
(
Some
a
),
None
;
repeat
constructor
.
+
by
exists
None
,
(
Some
a
)
;
repeat
constructor
.
+
exists
None
,
None
;
repeat
constructor
.
Qed
.
Canonical
Structure
optionR
:
=
CmraT
(
option
A
)
option_cmra_mixin
.
Global
Instance
option_cmra_discrete
:
CmraDiscrete
A
→
CmraDiscrete
optionR
.
Proof
.
split
;
[
apply
_
|].
by
intros
[
x
|]
;
[
apply
(
cmra_discrete_valid
x
)|].
Qed
.
Proof
.
split
;
[
apply
_
|].
by
intros
[
a
|]
;
[
apply
(
cmra_discrete_valid
a
)|].
Qed
.
Instance
option_unit
:
Unit
(
option
A
)
:
=
None
.
Lemma
option_ucmra_mixin
:
UcmraMixin
optionR
.
...
...
@@ -1328,64 +1329,63 @@ Section option.
Canonical
Structure
optionUR
:
=
UcmraT
(
option
A
)
option_ucmra_mixin
.
(** Misc *)
Lemma
op_None
m
x
m
y
:
m
x
⋅
m
y
=
None
↔
m
x
=
None
∧
m
y
=
None
.
Proof
.
destruct
m
x
,
m
y
;
naive_solver
.
Qed
.
Lemma
op_is_Some
m
x
m
y
:
is_Some
(
m
x
⋅
m
y
)
↔
is_Some
m
x
∨
is_Some
m
y
.
Proof
.
rewrite
-!
not_eq_None_Some
op_None
.
destruct
m
x
,
m
y
;
naive_solver
.
Qed
.
Lemma
op_None
m
a
m
b
:
m
a
⋅
m
b
=
None
↔
m
a
=
None
∧
m
b
=
None
.
Proof
.
destruct
m
a
,
m
b
;
naive_solver
.
Qed
.
Lemma
op_is_Some
m
a
m
b
:
is_Some
(
m
a
⋅
m
b
)
↔
is_Some
m
a
∨
is_Some
m
b
.
Proof
.
rewrite
-!
not_eq_None_Some
op_None
.
destruct
m
a
,
m
b
;
naive_solver
.
Qed
.
Lemma
cmra_opM_assoc'
x
m
y
m
z
:
x
⋅
?
m
y
⋅
?
m
z
≡
x
⋅
?
(
m
y
⋅
m
z
).
Proof
.
destruct
m
y
,
m
z
;
by
rewrite
/=
-
?assoc
.
Qed
.
Lemma
cmra_opM_assoc'
a
m
b
m
c
:
a
⋅
?
m
b
⋅
?
m
c
≡
a
⋅
?
(
m
b
⋅
m
c
).
Proof
.
destruct
m
b
,
m
c
;
by
rewrite
/=
-
?assoc
.
Qed
.
Global
Instance
Some_core_id
(
x
:
A
)
:
CoreId
x
→
CoreId
(
Some
x
).
Global
Instance
Some_core_id
a
:
CoreId
a
→
CoreId
(
Some
a
).
Proof
.
by
constructor
.
Qed
.
Global
Instance
option_core_id
(
mx
:
option
A
)
:
(
∀
x
:
A
,
CoreId
x
)
→
CoreId
mx
.
Proof
.
intros
.
destruct
mx
;
apply
_
.
Qed
.
Lemma
exclusiveN_Some_l
n
x
`
{!
Exclusive
x
}
my
:
✓
{
n
}
(
Some
x
⋅
my
)
→
my
=
None
.
Proof
.
destruct
my
.
move
=>
/(
exclusiveN_l
_
x
)
[].
done
.
Qed
.
Lemma
exclusiveN_Some_r
n
x
`
{!
Exclusive
x
}
my
:
✓
{
n
}
(
my
⋅
Some
x
)
→
my
=
None
.
Global
Instance
option_core_id
ma
:
(
∀
x
:
A
,
CoreId
x
)
→
CoreId
ma
.
Proof
.
intros
.
destruct
ma
;
apply
_
.
Qed
.
Lemma
exclusiveN_Some_l
n
a
`
{!
Exclusive
a
}
mb
:
✓
{
n
}
(
Some
a
⋅
mb
)
→
mb
=
None
.
Proof
.
destruct
mb
.
move
=>
/(
exclusiveN_l
_
a
)
[].
done
.
Qed
.
Lemma
exclusiveN_Some_r
n
a
`
{!
Exclusive
a
}
mb
:
✓
{
n
}
(
mb
⋅
Some
a
)
→
mb
=
None
.
Proof
.
rewrite
comm
.
by
apply
exclusiveN_Some_l
.
Qed
.
Lemma
exclusive_Some_l
x
`
{!
Exclusive
x
}
m
y
:
✓
(
Some
x
⋅
m
y
)
→
m
y
=
None
.
Proof
.
destruct
m
y
.
move
=>
/(
exclusive_l
x
)
[].
done
.
Qed
.
Lemma
exclusive_Some_r
x
`
{!
Exclusive
x
}
m
y
:
✓
(
m
y
⋅
Some
x
)
→
m
y
=
None
.
Lemma
exclusive_Some_l
a
`
{!
Exclusive
a
}
m
b
:
✓
(
Some
a
⋅
m
b
)
→
m
b
=
None
.
Proof
.
destruct
m
b
.
move
=>
/(
exclusive_l
a
)
[].
done
.
Qed
.
Lemma
exclusive_Some_r
a
`
{!
Exclusive
a
}
m
b
:
✓
(
m
b
⋅
Some
a
)
→
m
b
=
None
.
Proof
.
rewrite
comm
.
by
apply
exclusive_Some_l
.
Qed
.
Lemma
Some_includedN
n
x
y
:
Some
x
≼
{
n
}
Some
y
↔
x
≡
{
n
}
≡
y
∨
x
≼
{
n
}
y
.
Lemma
Some_includedN
n
a
b
:
Some
a
≼
{
n
}
Some
b
↔
a
≡
{
n
}
≡
b
∨
a
≼
{
n
}
b
.
Proof
.
rewrite
option_includedN
;
naive_solver
.
Qed
.
Lemma
Some_included
x
y
:
Some
x
≼
Some
y
↔
x
≡
y
∨
x
≼
y
.
Lemma
Some_included
a
b
:
Some
a
≼
Some
b
↔
a
≡
b
∨
a
≼
b
.
Proof
.
rewrite
option_included
;
naive_solver
.
Qed
.
Lemma
Some_included_2
x
y
:
x
≼
y
→
Some
x
≼
Some
y
.
Lemma
Some_included_2
a
b
:
a
≼
b
→
Some
a
≼
Some
b
.
Proof
.
rewrite
Some_included
;
eauto
.
Qed
.
Lemma
Some_includedN_total
`
{
CmraTotal
A
}
n
x
y
:
Some
x
≼
{
n
}
Some
y
↔
x
≼
{
n
}
y
.
Lemma
Some_includedN_total
`
{
CmraTotal
A
}
n
a
b
:
Some
a
≼
{
n
}
Some
b
↔
a
≼
{
n
}
b
.
Proof
.
rewrite
Some_includedN
.
split
.
by
intros
[->|?].
eauto
.
Qed
.
Lemma
Some_included_total
`
{
CmraTotal
A
}
x
y
:
Some
x
≼
Some
y
↔
x
≼
y
.
Lemma
Some_included_total
`
{
CmraTotal
A
}
a
b
:
Some
a
≼
Some
b
↔
a
≼
b
.
Proof
.
rewrite
Some_included
.
split
.
by
intros
[->|?].
eauto
.
Qed
.
Lemma
Some_includedN_exclusive
n
x
`
{!
Exclusive
x
}
y
:
Some
x
≼
{
n
}
Some
y
→
✓
{
n
}
y
→
x
≡
{
n
}
≡
y
.
Lemma
Some_includedN_exclusive
n
a
`
{!
Exclusive
a
}
b
:
Some
a
≼
{
n
}
Some
b
→
✓
{
n
}
b
→
a
≡
{
n
}
≡
b
.
Proof
.
move
=>
/
Some_includedN
[//|/
exclusive_includedN
]
;
tauto
.
Qed
.
Lemma
Some_included_exclusive
x
`
{!
Exclusive
x
}
y
:
Some
x
≼
Some
y
→
✓
y
→
x
≡
y
.
Lemma
Some_included_exclusive
a
`
{!
Exclusive
a
}
b
:
Some
a
≼
Some
b
→
✓
b
→
a
≡
b
.
Proof
.
move
=>
/
Some_included
[//|/
exclusive_included
]
;
tauto
.
Qed
.
Lemma
is_Some_includedN
n
m
x
m
y
:
m
x
≼
{
n
}
m
y
→
is_Some
m
x
→
is_Some
m
y
.
Lemma
is_Some_includedN
n
m
a
m
b
:
m
a
≼
{
n
}
m
b
→
is_Some
m
a
→
is_Some
m
b
.
Proof
.
rewrite
-!
not_eq_None_Some
option_includedN
.
naive_solver
.
Qed
.
Lemma
is_Some_included
m
x
m
y
:
m
x
≼
m
y
→
is_Some
m
x
→
is_Some
m
y
.
Lemma
is_Some_included
m
a
m
b
:
m
a
≼
m
b
→
is_Some
m
a
→
is_Some
m
b
.
Proof
.
rewrite
-!
not_eq_None_Some
option_included
.
naive_solver
.
Qed
.
Global
Instance
cancelable_Some
x
:
IdFree
x
→
Cancelable
x
→
Cancelable
(
Some
x
).
Global
Instance
cancelable_Some
a
:
IdFree
a
→
Cancelable
a
→
Cancelable
(
Some
a
).
Proof
.
intros
Hirr
??
[
y
|]
[
z
|]
?
EQ
;
inversion_clear
EQ
.
-
constructor
.
by
apply
(
cancelableN
x
).
-
destruct
(
Hirr
y
)
;
[|
eauto
using
dist_le
with
lia
].
by
eapply
(
cmra_validN_op_l
0
x
y
),
(
cmra_validN_le
n
)
;
last
lia
.
-
destruct
(
Hirr
z
)
;
[|
symmetry
;
eauto
using
dist_le
with
lia
].
intros
Hirr
??
[
b
|]
[
c
|]
?
EQ
;
inversion_clear
EQ
.
-
constructor
.
by
apply
(
cancelableN
a
).
-
destruct
(
Hirr
b
)
;
[|
eauto
using
dist_le
with
lia
].
by
eapply
(
cmra_validN_op_l
0
a
b
),
(
cmra_validN_le
n
)
;
last
lia
.
-
destruct
(
Hirr
c
)
;
[|
symmetry
;
eauto
using
dist_le
with
lia
].
by
eapply
(
cmra_validN_le
n
)
;
last
lia
.
-
done
.
Qed
.
...
...
@@ -1396,35 +1396,37 @@ Arguments optionUR : clear implicits.
Section
option_prod
.
Context
{
A
B
:
cmraT
}.
Implicit
Types
a
:
A
.
Implicit
Types
b
:
B
.
Lemma
Some_pair_includedN
n
(
x
1
x
2
:
A
)
(
y
1
y
2
:
B
)
:
Some
(
x
1
,
y
1
)
≼
{
n
}
Some
(
x
2
,
y
2
)
→
Some
x
1
≼
{
n
}
Some
x
2
∧
Some
y
1
≼
{
n
}
Some
y
2
.
Lemma
Some_pair_includedN
n
a
1
a
2
b
1
b
2
:
Some
(
a
1
,
b
1
)
≼
{
n
}
Some
(
a
2
,
b
2
)
→
Some
a
1
≼
{
n
}
Some
a
2
∧
Some
b
1
≼
{
n
}
Some
b
2
.
Proof
.
rewrite
!
Some_includedN
.
intros
[[??]|[??]%
prod_includedN
]
;
eauto
.
Qed
.
Lemma
Some_pair_includedN_total_1
`
{
CmraTotal
A
}
n
(
x
1
x
2
:
A
)
(
y
1
y
2
:
B
)
:
Some
(
x
1
,
y
1
)
≼
{
n
}
Some
(
x
2
,
y
2
)
→
x
1
≼
{
n
}
x
2
∧
Some
y
1
≼
{
n
}
Some
y
2
.
Proof
.
intros
?%
Some_pair_includedN
.
by
rewrite
-(
Some_includedN_total
_
x
1
).
Qed
.
Lemma
Some_pair_includedN_total_2
`
{
CmraTotal
B
}
n
(
x
1
x
2
:
A
)
(
y
1
y
2
:
B
)
:
Some
(
x
1
,
y
1
)
≼
{
n
}
Some
(
x
2
,
y
2
)
→
Some
x
1
≼
{
n
}
Some
x
2
∧
y
1
≼
{
n
}
y
2
.
Proof
.
intros
?%
Some_pair_includedN
.
by
rewrite
-(
Some_includedN_total
_
y
1
).
Qed
.
Lemma
Some_pair_included
(
x
1
x
2
:
A
)
(
y
1
y
2
:
B
)
:
Some
(
x
1
,
y
1
)
≼
Some
(
x
2
,
y
2
)
→
Some
x
1
≼
Some
x
2
∧
Some
y
1
≼
Some
y
2
.
Lemma
Some_pair_includedN_total_1
`
{
CmraTotal
A
}
n
a
1
a
2
b
1
b
2
:
Some
(
a
1
,
b
1
)
≼
{
n
}
Some
(
a
2
,
b
2
)
→
a
1
≼
{
n
}
a
2
∧
Some
b
1
≼
{
n
}
Some
b
2
.
Proof
.
intros
?%
Some_pair_includedN
.
by
rewrite
-(
Some_includedN_total
_
a
1
).
Qed
.
Lemma
Some_pair_includedN_total_2
`
{
CmraTotal
B
}
n
a
1
a
2
b
1
b
2
:
Some
(
a
1
,
b
1
)
≼
{
n
}
Some
(
a
2
,
b
2
)
→
Some
a
1
≼
{
n
}
Some
a
2
∧
b
1
≼
{
n
}
b
2
.
Proof
.
intros
?%
Some_pair_includedN
.
by
rewrite
-(
Some_includedN_total
_
b
1
).
Qed
.
Lemma
Some_pair_included
a
1
a
2
b
1
b
2
:
Some
(
a
1
,
b
1
)
≼
Some
(
a
2
,
b
2
)
→
Some
a
1
≼
Some
a
2
∧
Some
b
1
≼
Some
b
2
.
Proof
.
rewrite
!
Some_included
.
intros
[[??]|[??]%
prod_included
]
;
eauto
.
Qed
.
Lemma
Some_pair_included_total_1
`
{
CmraTotal
A
}
(
x
1
x
2
:
A
)
(
y
1
y
2
:
B
)
:
Some
(
x
1
,
y
1
)
≼
Some
(
x
2
,
y
2
)
→
x
1
≼
x
2
∧
Some
y
1
≼
Some
y
2
.
Proof
.
intros
?%
Some_pair_included
.
by
rewrite
-(
Some_included_total
x
1
).
Qed
.
Lemma
Some_pair_included_total_2
`
{
CmraTotal
B
}
(
x
1
x
2
:
A
)
(
y
1
y
2
:
B
)
:
Some
(
x
1
,
y
1
)
≼
Some
(
x
2
,
y
2
)
→
Some
x
1
≼
Some
x
2
∧
y
1
≼
y
2
.
Proof
.
intros
?%
Some_pair_included
.
by
rewrite
-(
Some_included_total
y
1
).
Qed
.
Lemma
Some_pair_included_total_1
`
{
CmraTotal
A
}
a
1
a
2
b
1
b
2
:
Some
(
a
1
,
b
1
)
≼
Some
(
a
2
,
b
2
)
→
a
1
≼
a
2
∧
Some
b
1
≼
Some
b
2
.
Proof
.
intros
?%
Some_pair_included
.
by
rewrite
-(
Some_included_total
a
1
).
Qed
.
Lemma
Some_pair_included_total_2
`
{
CmraTotal
B
}
a
1
a
2
b
1
b
2
:
Some
(
a
1
,
b
1
)
≼
Some
(
a
2
,
b
2
)
→
Some
a
1
≼
Some
a
2
∧
b
1
≼
b
2
.
Proof
.
intros
?%
Some_pair_included
.
by
rewrite
-(
Some_included_total
b
1
).
Qed
.
End
option_prod
.
Instance
option_fmap_cmra_morphism
{
A
B
:
cmraT
}
(
f
:
A
→
B
)
`
{!
CmraMorphism
f
}
:
CmraMorphism
(
fmap
f
:
option
A
→
option
B
).
Proof
.
split
;
first
apply
_
.
-
intros
n
[
x
|]
?
;
rewrite
/
cmra_validN
//=.
by
apply
(
cmra_morphism_validN
f
).
-
move
=>
[
x
|]
//.
by
apply
Some_proper
,
cmra_morphism_pcore
.
-
move
=>
[
x
|]
[
y
|]
//=.
by
rewrite
-(
cmra_morphism_op
f
).
-
intros
n
[
a
|]
?
;
rewrite
/
cmra_validN
//=.
by
apply
(
cmra_morphism_validN
f
).
-
move
=>
[
a
|]
//.
by
apply
Some_proper
,
cmra_morphism_pcore
.
-
move
=>
[
a
|]
[
b
|]
//=.
by
rewrite
-(
cmra_morphism_op
f
).
Qed
.
Program
Definition
optionRF
(
F
:
rFunctor
)
:
rFunctor
:
=
{|
...
...
theories/algebra/list.v
View file @
0f11453a
...
...
@@ -6,6 +6,7 @@ Set Default Proof Using "Type".
Section
cofe
.
Context
{
A
:
ofeT
}.
Implicit
Types
l
:
list
A
.
Instance
list_dist
:
Dist
(
list
A
)
:
=
λ
n
,
Forall2
(
dist
n
).
...
...
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