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
I
Iris
Project overview
Project overview
Details
Activity
Releases
Repository
Repository
Files
Commits
Branches
Tags
Contributors
Graph
Compare
Issues
116
Issues
116
List
Boards
Labels
Service Desk
Milestones
Merge Requests
21
Merge Requests
21
CI / CD
CI / CD
Pipelines
Jobs
Schedules
Operations
Operations
Incidents
Environments
Analytics
Analytics
CI / CD
Repository
Value Stream
Wiki
Wiki
Members
Members
Collapse sidebar
Close sidebar
Activity
Graph
Create a new issue
Jobs
Commits
Issue Boards
Open sidebar
Iris
Iris
Commits
8d8d9eef
Commit
8d8d9eef
authored
May 26, 2016
by
Robbert Krebbers
Browse files
Options
Browse Files
Download
Email Patches
Plain Diff
Lift a relation to the sum type.
parent
b3ec9bd3
Changes
1
Hide whitespace changes
Inline
Side-by-side
Showing
1 changed file
with
38 additions
and
0 deletions
+38
-0
prelude/base.v
prelude/base.v
+38
-0
No files found.
prelude/base.v
View file @
8d8d9eef
...
...
@@ -490,6 +490,10 @@ Instance snd_proper `{Equiv A, Equiv B} : Proper ((≡) ==> (≡)) (@snd A B) :=
Typeclasses
Opaque
prod_equiv
.
(** ** Sums *)
Definition
sum_map
{
A
A'
B
B'
}
(
f
:
A
→
A'
)
(
g
:
B
→
B'
)
(
xy
:
A
+
B
)
:
A'
+
B'
:
=
match
xy
with
inl
x
=>
inl
(
f
x
)
|
inr
y
=>
inr
(
g
y
)
end
.
Arguments
sum_map
{
_
_
_
_
}
_
_
!
_
/.
Instance
sum_inhabited_l
{
A
B
}
(
iA
:
Inhabited
A
)
:
Inhabited
(
A
+
B
)
:
=
match
iA
with
populate
x
=>
populate
(
inl
x
)
end
.
Instance
sum_inhabited_r
{
A
B
}
(
iB
:
Inhabited
A
)
:
Inhabited
(
A
+
B
)
:
=
...
...
@@ -500,6 +504,40 @@ Proof. injection 1; auto. Qed.
Instance
inr_inj
:
Inj
(=)
(=)
(@
inr
A
B
).
Proof
.
injection
1
;
auto
.
Qed
.
Instance
sum_map_inj
{
A
A'
B
B'
}
(
f
:
A
→
A'
)
(
g
:
B
→
B'
)
:
Inj
(=)
(=)
f
→
Inj
(=)
(=)
g
→
Inj
(=)
(=)
(
sum_map
f
g
).
Proof
.
intros
??
[?|?]
[?|?]
[=]
;
f_equal
;
apply
(
inj
_
)
;
auto
.
Qed
.
Inductive
sum_relation
{
A
B
}
(
R1
:
relation
A
)
(
R2
:
relation
B
)
:
relation
(
A
+
B
)
:
=
|
inl_related
x1
x2
:
R1
x1
x2
→
sum_relation
R1
R2
(
inl
x1
)
(
inl
x2
)
|
inr_related
y1
y2
:
R2
y1
y2
→
sum_relation
R1
R2
(
inr
y1
)
(
inr
y2
).
Section
sum_relation
.
Context
`
{
R1
:
relation
A
,
R2
:
relation
B
}.
Global
Instance
sum_relation_refl
:
Reflexive
R1
→
Reflexive
R2
→
Reflexive
(
sum_relation
R1
R2
).
Proof
.
intros
??
[?|?]
;
constructor
;
reflexivity
.
Qed
.
Global
Instance
sum_relation_sym
:
Symmetric
R1
→
Symmetric
R2
→
Symmetric
(
sum_relation
R1
R2
).
Proof
.
destruct
3
;
constructor
;
eauto
.
Qed
.
Global
Instance
sum_relation_trans
:
Transitive
R1
→
Transitive
R2
→
Transitive
(
sum_relation
R1
R2
).
Proof
.
destruct
3
;
inversion_clear
1
;
constructor
;
eauto
.
Qed
.
Global
Instance
sum_relation_equiv
:
Equivalence
R1
→
Equivalence
R2
→
Equivalence
(
sum_relation
R1
R2
).
Proof
.
split
;
apply
_
.
Qed
.
Global
Instance
inl_proper'
:
Proper
(
R1
==>
sum_relation
R1
R2
)
inl
.
Proof
.
constructor
;
auto
.
Qed
.
Global
Instance
inr_proper'
:
Proper
(
R2
==>
sum_relation
R1
R2
)
inr
.
Proof
.
constructor
;
auto
.
Qed
.
End
sum_relation
.
Instance
sum_equiv
`
{
Equiv
A
,
Equiv
B
}
:
Equiv
(
A
+
B
)
:
=
sum_relation
(
≡
)
(
≡
).
Instance
inl_proper
`
{
Equiv
A
,
Equiv
B
}
:
Proper
((
≡
)
==>
(
≡
))
(@
inl
A
B
)
:
=
_
.
Instance
inr_proper
`
{
Equiv
A
,
Equiv
B
}
:
Proper
((
≡
)
==>
(
≡
))
(@
inr
A
B
)
:
=
_
.
Typeclasses
Opaque
sum_equiv
.
(** ** Option *)
Instance
option_inhabited
{
A
}
:
Inhabited
(
option
A
)
:
=
populate
None
.
...
...
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