Skip to content
GitLab
Projects
Groups
Snippets
Help
Loading...
Help
Help
Support
Community forum
Keyboard shortcuts
?
Submit feedback
Contribute to GitLab
Sign in / Register
Toggle navigation
C
coq-stdpp
Project overview
Project overview
Details
Activity
Releases
Repository
Repository
Files
Commits
Branches
Tags
Contributors
Graph
Compare
Issues
0
Issues
0
List
Boards
Labels
Service Desk
Milestones
Merge Requests
0
Merge Requests
0
CI / CD
CI / CD
Pipelines
Jobs
Schedules
Operations
Operations
Incidents
Environments
Analytics
Analytics
CI / CD
Repository
Value Stream
Wiki
Wiki
Snippets
Snippets
Members
Members
Collapse sidebar
Close sidebar
Activity
Graph
Create a new issue
Jobs
Commits
Issue Boards
Open sidebar
David Swasey
coq-stdpp
Commits
9365ea8b
Commit
9365ea8b
authored
May 26, 2016
by
Robbert Krebbers
Browse files
Options
Browse Files
Download
Email Patches
Plain Diff
Generalize from_option and define default using it.
parent
e36e7f99
Changes
4
Hide whitespace changes
Inline
Side-by-side
Showing
4 changed files
with
17 additions
and
22 deletions
+17
-22
theories/co_pset.v
theories/co_pset.v
+1
-1
theories/finite.v
theories/finite.v
+2
-2
theories/list.v
theories/list.v
+2
-2
theories/option.v
theories/option.v
+12
-17
No files found.
theories/co_pset.v
View file @
9365ea8b
...
...
@@ -225,7 +225,7 @@ Fixpoint coPpick_raw (t : coPset_raw) : option positive :=
|
Some
i
=>
Some
(
i
~
0
)
|
None
=>
(~
1
)
<$>
coPpick_raw
r
end
end
.
Definition
coPpick
(
X
:
coPset
)
:
positive
:
=
from_option
1
(
coPpick_raw
(
`
X
)).
Definition
coPpick
(
X
:
coPset
)
:
positive
:
=
from_option
id
1
(
coPpick_raw
(
`
X
)).
Lemma
coPpick_raw_elem_of
t
i
:
coPpick_raw
t
=
Some
i
→
e_of
i
t
.
Proof
.
...
...
theories/finite.v
View file @
9365ea8b
...
...
@@ -12,7 +12,7 @@ Arguments NoDup_enum _ {_ _} : clear implicits.
Definition
card
A
`
{
Finite
A
}
:
=
length
(
enum
A
).
Program
Instance
finite_countable
`
{
Finite
A
}
:
Countable
A
:
=
{|
encode
:
=
λ
x
,
Pos
.
of_nat
$
S
$
from_option
0
$
fst
<$>
list_find
(
x
=)
(
enum
A
)
;
Pos
.
of_nat
$
S
$
from_option
id
0
$
fst
<$>
list_find
(
x
=)
(
enum
A
)
;
decode
:
=
λ
p
,
enum
A
!!
pred
(
Pos
.
to_nat
p
)
|}.
Arguments
Pos
.
of_nat
_
:
simpl
never
.
...
...
@@ -127,7 +127,7 @@ Lemma finite_surj A `{Finite A} B `{Finite B} :
0
<
card
A
≤
card
B
→
∃
g
:
B
→
A
,
Surj
(=)
g
.
Proof
.
intros
[??].
destruct
(
finite_inhabited
A
)
as
[
x'
]
;
auto
with
lia
.
exists
(
λ
y
:
B
,
from_option
x'
(
decode_nat
(
encode_nat
y
))).
exists
(
λ
y
:
B
,
from_option
id
x'
(
decode_nat
(
encode_nat
y
))).
intros
x
.
destruct
(
encode_decode
B
(
encode_nat
x
))
as
(
y
&
Hy1
&
Hy2
).
{
pose
proof
(
encode_lt_card
x
)
;
lia
.
}
exists
y
.
by
rewrite
Hy2
,
decode_encode_nat
.
...
...
theories/list.v
View file @
9365ea8b
...
...
@@ -3393,7 +3393,7 @@ Definition eval {A} (E : env A) : rlist nat → list A :=
fix go t :=
match t with
| rnil => []
| rnode i => from_option [] (E !! i)
| rnode i => from_option
id
[] (E !! i)
| rapp t1 t2 => go t1 ++ go t2
end.
...
...
@@ -3427,7 +3427,7 @@ End quote.
Section eval.
Context {A} (E : env A).
Lemma eval_alt t : eval E t = to_list t ≫= from_option [] ∘ (E !!).
Lemma eval_alt t : eval E t = to_list t ≫= from_option
id
[] ∘ (E !!).
Proof.
induction t; csimpl.
- done.
...
...
theories/option.v
View file @
9365ea8b
...
...
@@ -19,16 +19,15 @@ Proof. congruence. Qed.
Instance
Some_inj
{
A
}
:
Inj
(=)
(=)
(@
Some
A
).
Proof
.
congruence
.
Qed
.
(** The
non dependent elimination principle on the option type
. *)
Definition
default
{
A
B
}
(
y
:
B
)
(
mx
:
option
A
)
(
f
:
A
→
B
)
:
B
:
=
(** The
[from_option] is the eliminator for option
. *)
Definition
from_option
{
A
B
}
(
f
:
A
→
B
)
(
y
:
B
)
(
mx
:
option
A
)
:
B
:
=
match
mx
with
None
=>
y
|
Some
x
=>
f
x
end
.
Instance
:
Params
(@
default
)
2
.
Instance
:
Params
(@
from_option
)
3
.
Arguments
from_option
{
_
_
}
_
_
!
_
/.
(** The [from_option] function allows us to get the value out of the option
type by specifying a default value. *)
Definition
from_option
{
A
}
(
x
:
A
)
(
mx
:
option
A
)
:
A
:
=
match
mx
with
None
=>
x
|
Some
y
=>
y
end
.
Instance
:
Params
(@
from_option
)
1
.
(* The eliminator again, but with the arguments in different order, which is
sometimes more convenient. *)
Notation
default
y
mx
f
:
=
(
from_option
f
y
mx
)
(
only
parsing
).
(** An alternative, but equivalent, definition of equality on the option
data type. This theorem is useful to prove that two options are the same. *)
...
...
@@ -137,9 +136,9 @@ Section setoids.
Global
Instance
is_Some_proper
:
Proper
((
≡
)
==>
iff
)
(@
is_Some
A
).
Proof
.
inversion_clear
1
;
split
;
eauto
.
Qed
.
Global
Instance
from_option_proper
:
Proper
((
≡
)
==>
(
≡
)
==>
(
≡
))
(@
from_option
A
).
Proof
.
by
destruct
2
.
Qed
.
Global
Instance
from_option_proper
{
B
}
(
R
:
relation
B
)
(
f
:
A
→
B
)
:
Proper
((
≡
)
==>
R
)
f
→
Proper
(
R
==>
(
≡
)
==>
R
)
(
from_option
f
).
Proof
.
destruct
3
;
simpl
;
auto
.
Qed
.
End
setoids
.
Typeclasses
Opaque
option_equiv
.
...
...
@@ -323,9 +322,7 @@ Tactic Notation "simpl_option" "by" tactic3(tac) :=
let
Hx
:
=
fresh
in
assert_Some_None
A
mx
Hx
;
rewrite
Hx
in
H
;
clear
Hx
|
H
:
context
[
fmap
(
M
:
=
option
)
(
A
:
=
?A
)
?f
?mx
]
|-
_
=>
let
Hx
:
=
fresh
in
assert_Some_None
A
mx
Hx
;
rewrite
Hx
in
H
;
clear
Hx
|
H
:
context
[
default
(
A
:
=
?A
)
_
?mx
_
]
|-
_
=>
let
Hx
:
=
fresh
in
assert_Some_None
A
mx
Hx
;
rewrite
Hx
in
H
;
clear
Hx
|
H
:
context
[
from_option
(
A
:
=
?A
)
_
?mx
]
|-
_
=>
|
H
:
context
[
from_option
(
A
:
=
?A
)
_
_
?mx
]
|-
_
=>
let
Hx
:
=
fresh
in
assert_Some_None
A
mx
Hx
;
rewrite
Hx
in
H
;
clear
Hx
|
H
:
context
[
match
?mx
with
_
=>
_
end
]
|-
_
=>
match
type
of
mx
with
...
...
@@ -336,9 +333,7 @@ Tactic Notation "simpl_option" "by" tactic3(tac) :=
let
Hx
:
=
fresh
in
assert_Some_None
A
mx
Hx
;
rewrite
Hx
;
clear
Hx
|
|-
context
[
fmap
(
M
:
=
option
)
(
A
:
=
?A
)
?f
?mx
]
=>
let
Hx
:
=
fresh
in
assert_Some_None
A
mx
Hx
;
rewrite
Hx
;
clear
Hx
|
|-
context
[
default
(
A
:
=
?A
)
_
?mx
_
]
=>
let
Hx
:
=
fresh
in
assert_Some_None
A
mx
Hx
;
rewrite
Hx
;
clear
Hx
|
|-
context
[
from_option
(
A
:
=
?A
)
_
?mx
]
=>
|
|-
context
[
from_option
(
A
:
=
?A
)
_
_
?mx
]
=>
let
Hx
:
=
fresh
in
assert_Some_None
A
mx
Hx
;
rewrite
Hx
;
clear
Hx
|
|-
context
[
match
?mx
with
_
=>
_
end
]
=>
match
type
of
mx
with
...
...
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