Search code examples
coqssreflect

Derive a ssreflect finType from a seq over a finType with uniq


I have a structure consisting of a sequence over a finite type and a proof of uniq of this sequence. This should describe a type that is obviously finite, but I do not see how to show that.

I thought I could use UniqFinMixin, however it requires - if I understand this correctly - to provide an explicit seq of all elements of the type, which I don't know how to compute. I tried using Finite.enum on the finite type, but it only produces a seq with all the elements of the finite type and I didn't find an elegant way of computing all the subsequences / permutations.

From mathcomp
Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq choice fintype.
From mathcomp
Require Import tuple finfun bigop finset.

Variable ft : finType.

Structure dbranch := {branch :> seq ft ; buniq : uniq branch}.

Canonical dbranch_subType := Eval hnf in [subType for branch].
Canonical dbranch_eqType := Eval hnf in EqType _ [eqMixin of dbranch by <:].
Canonical dbranch_choiceType := Eval hnf in ChoiceType _ [choiceMixin of dbranch by <:].
Canonical dbranch_countType := Eval hnf in CountType _ [countMixin of dbranch by <:].
Canonical dbranch_subCountType := Eval hnf in [subCountType of dbranch].


Lemma dbranchFin : Finite.mixin_of [eqType of dbranch].
Admitted. (* :-( *)

Canonical dbranch_finType := Eval hnf in FinType _ dbranchFin.

It seems weird to me that there wouldn't be a simple way to derive a finType here, but I couldn't find it in the finset.v file. Thank you in advance for your help.


Solution

  • You can show that dbranch embeds into another finType -- for example, the type of lists of elements of ft whose size is bounded by #|ft|.

    Lemma size_dbranch d : size (branch d) < #|ft|.+1.
    Proof.
    rewrite ltnS [card]unlock uniq_leq_size ?buniq // => ?.
    by rewrite mem_enum.
    Qed.
    
    Definition tag_of_dbranch d : {k : 'I_#|ft|.+1 & k.-tuple ft} :=
      @Tagged _ (Sub (size (branch d)) (size_dbranch d))
                (fun k : 'I_#|ft|.+1 => k.-tuple ft)
                (in_tuple (branch d)).
    
    Definition dbranch_of_tag (t : {k : 'I_#|ft|.+1 & k.-tuple ft}) : option dbranch :=
      insub (val (tagged t)).
    
    Lemma tag_of_dbranchK : pcancel tag_of_dbranch dbranch_of_tag.
    Proof. by rewrite /tag_of_dbranch /dbranch_of_tag=> x; rewrite valK. Qed.
    
    Definition dbranch_finMixin := PcanFinMixin tag_of_dbranchK.
    Canonical dbranch_finType := Eval hnf in FinType dbranch dbranch_finMixin.