Skip to content
Snippets Groups Projects
Commit f59a2c06 authored by Robbert Krebbers's avatar Robbert Krebbers
Browse files

Merge branch 'injective_finite' into 'master'

surjective_finite

See merge request !238
parents b2c5f1fe 57c98ce1
No related branches found
No related tags found
1 merge request!238surjective_finite
Pipeline #43850 passed
...@@ -220,15 +220,26 @@ Section enc_finite. ...@@ -220,15 +220,26 @@ Section enc_finite.
Proof. unfold card. simpl. by rewrite fmap_length, seq_length. Qed. Proof. unfold card. simpl. by rewrite fmap_length, seq_length. Qed.
End enc_finite. End enc_finite.
Section surjective_finite.
Context `{Finite A, EqDecision B} (f : A B).
Context `{!Surj (=) f}.
Program Instance surjective_finite: Finite B :=
{| enum := remove_dups (f <$> enum A) |}.
Next Obligation. apply NoDup_remove_dups. Qed.
Next Obligation.
intros y. rewrite elem_of_remove_dups, elem_of_list_fmap.
destruct (surj f y). eauto using elem_of_enum.
Qed.
End surjective_finite.
Section bijective_finite. Section bijective_finite.
Context `{Finite A, EqDecision B} (f : A B) (g : B A). Context `{Finite A, EqDecision B} (f : A B) (g : B A).
Context `{!Inj (=) (=) f, !Cancel (=) f g}. Context `{!Inj (=) (=) f, !Cancel (=) f g}.
Program Instance bijective_finite: Finite B := {| enum := f <$> enum A |}. Global Instance bijective_finite: Finite B :=
Next Obligation. apply (NoDup_fmap_2 _), NoDup_enum. Qed. let _ := cancel_surj (f:=f) (g:=g) in
Next Obligation. surjective_finite f.
intros y. rewrite elem_of_list_fmap. eauto using elem_of_enum.
Qed.
End bijective_finite. End bijective_finite.
Program Instance option_finite `{Finite A} : Finite (option A) := Program Instance option_finite `{Finite A} : Finite (option A) :=
......
0% Loading or .
You are about to add 0 people to the discussion. Proceed with caution.
Finish editing this message first!
Please register or to comment