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
David Swasey
coq-stdpp
Commits
596a0a2c
Commit
596a0a2c
authored
Sep 20, 2016
by
Robbert Krebbers
Browse files
Options
Browse Files
Download
Email Patches
Plain Diff
Use ⊆ type class for set-like inclusion of lists.
This also solves a name clash with the extension order of CMRAs.
parent
db08223a
Changes
3
Hide whitespace changes
Inline
Side-by-side
Showing
3 changed files
with
10 additions
and
11 deletions
+10
-11
theories/collections.v
theories/collections.v
+2
-2
theories/fin_maps.v
theories/fin_maps.v
+4
-4
theories/list.v
theories/list.v
+4
-5
No files found.
theories/collections.v
View file @
596a0a2c
...
...
@@ -711,8 +711,8 @@ Section list_unfold.
Qed
.
Global
Instance
set_unfold_included
l
k
(
P
Q
:
A
→
Prop
)
:
(
∀
x
,
SetUnfold
(
x
∈
l
)
(
P
x
))
→
(
∀
x
,
SetUnfold
(
x
∈
k
)
(
Q
x
))
→
SetUnfold
(
l
`
included
`
k
)
(
∀
x
,
P
x
→
Q
x
).
Proof
.
by
constructor
;
unfold
included
;
set_unfold
.
Qed
.
SetUnfold
(
l
⊆
k
)
(
∀
x
,
P
x
→
Q
x
).
Proof
.
by
constructor
;
unfold
subseteq
,
list_subseteq
;
set_unfold
.
Qed
.
End
list_unfold
.
...
...
theories/fin_maps.v
View file @
596a0a2c
...
...
@@ -1225,14 +1225,14 @@ Qed.
Lemma
map_union_cancel_l
{
A
}
(
m1
m2
m3
:
M
A
)
:
m1
⊥
ₘ
m3
→
m2
⊥
ₘ
m3
→
m3
∪
m1
=
m3
∪
m2
→
m1
=
m2
.
Proof
.
intros
.
apply
(
anti_symm
(
⊆
))
;
apply
map_union_reflecting_l
with
m3
;
auto
using
(
reflexive_eq
(
R
:
=
(
⊆
)
)).
intros
.
apply
(
anti_symm
(
⊆
))
;
apply
map_union_reflecting_l
with
m3
;
auto
using
(
reflexive_eq
(
R
:
=
@
subseteq
(
M
A
)
_
)).
Qed
.
Lemma
map_union_cancel_r
{
A
}
(
m1
m2
m3
:
M
A
)
:
m1
⊥
ₘ
m3
→
m2
⊥
ₘ
m3
→
m1
∪
m3
=
m2
∪
m3
→
m1
=
m2
.
Proof
.
intros
.
apply
(
anti_symm
(
⊆
))
;
apply
map_union_reflecting_r
with
m3
;
auto
using
(
reflexive_eq
(
R
:
=
(
⊆
)
)).
intros
.
apply
(
anti_symm
(
⊆
))
;
apply
map_union_reflecting_r
with
m3
;
auto
using
(
reflexive_eq
(
R
:
=
@
subseteq
(
M
A
)
_
)).
Qed
.
Lemma
map_disjoint_union_l
{
A
}
(
m1
m2
m3
:
M
A
)
:
m1
∪
m2
⊥
ₘ
m3
↔
m1
⊥
ₘ
m3
∧
m2
⊥
ₘ
m3
.
...
...
theories/list.v
View file @
596a0a2c
...
...
@@ -303,9 +303,8 @@ Inductive Forall3 {A B C} (P : A → B → C → Prop) :
| Forall3_cons x y z l k k' :
P x y z → Forall3 P l k k' → Forall3 P (x :: l) (y :: k) (z :: k').
(** Set operations Decisionon lists *)
Definition included {A} (l1 l2 : list A) := ∀ x, x ∈ l1 → x ∈ l2.
Infix "`included`" := included (at level 70) : C_scope.
(** Set operations on lists *)
Instance list_subseteq {A} : SubsetEq (list A) := λ l1 l2, ∀ x, x ∈ l1 → x ∈ l2.
Section list_set.
Context `{dec : EqDecision A}.
...
...
@@ -2046,9 +2045,9 @@ Section contains_dec.
End contains_dec.
(** ** Properties of [included] *)
Global Instance
included_preorder : PreOrder (@included A
).
Global Instance
list_subseteq_po : PreOrder (@subseteq (list A) _
).
Proof. split; firstorder. Qed.
Lemma
included
_nil l : []
`included`
l.
Lemma
list_subseteq
_nil l : []
⊆
l.
Proof. intros x. by rewrite elem_of_nil. Qed.
(** ** Properties of the [Forall] and [Exists] predicate *)
...
...
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