11(* (c) Copyright 2006-2019 Microsoft Corporation and Inria. *)
22(* Distributed under the terms of CeCILL-B. *)
33
4+ From HB Require Import structures.
45Set Warnings "-notation-incompatible-format".
56From mathcomp Require Import ssreflect ssrbool ssrnat eqtype ssrfun seq.
67Set Warnings "notation-incompatible-format".
@@ -505,11 +506,8 @@ Section FinSetCanonicals.
505506
506507Variable (K : choiceType).
507508
508- Canonical fsetType := Eval hnf in [subType for (@enum_fset K)].
509- Definition fset_eqMixin := Eval hnf in [eqMixin of {fset K} by <:].
510- Canonical fset_eqType := Eval hnf in EqType {fset K} fset_eqMixin.
511- Definition fset_choiceMixin := Eval hnf in [choiceMixin of {fset K} by <:].
512- Canonical fset_choiceType := Eval hnf in ChoiceType {fset K} fset_choiceMixin.
509+ HB.instance Definition _ := [isSub for (@enum_fset K)].
510+ HB.instance Definition _ := [Choice of {fset K} by <:].
513511
514512End FinSetCanonicals.
515513
@@ -526,14 +524,9 @@ Proof. by rewrite canonical_uniq // keys_canonical. Qed.
526524Record fset_sub_type : predArgType :=
527525 FSetSub {fsval : K; fsvalP : in_mem fsval (@mem K _ A)}.
528526
529- Canonical fset_sub_subType := Eval hnf in [subType for fsval].
530- Definition fset_sub_eqMixin := Eval hnf in [eqMixin of fset_sub_type by <:].
531- Canonical fset_sub_eqType := Eval hnf in EqType fset_sub_type fset_sub_eqMixin.
532- Definition fset_sub_choiceMixin := Eval hnf in [choiceMixin of fset_sub_type by <:].
533- Canonical fset_sub_choiceType := Eval hnf in ChoiceType fset_sub_type fset_sub_choiceMixin.
534- Definition fset_countMixin (T : countType) := Eval hnf in [countMixin of {fset T} by <:].
535- Canonical fset_countType (T : countType) := Eval hnf in CountType {fset T} (fset_countMixin T).
536-
527+ HB.instance Definition _ := [isSub for fsval].
528+ HB.instance Definition _ := [Choice of fset_sub_type by <:].
529+ HB.instance Definition _ (T : countType) := [Countable of {fset T} by <:].
537530
538531Definition fset_sub_enum : seq fset_sub_type :=
539532 undup (pmap insub (enum_fset A)).
@@ -556,13 +549,10 @@ rewrite /fset_sub_unpickle => x.
556549by rewrite (nth_map x) ?nth_index ?index_mem ?mem_fset_sub_enum.
557550Qed .
558551
559- Definition fset_sub_countMixin := CountMixin fset_sub_pickleK.
560- Canonical fset_sub_countType := Eval hnf in CountType fset_sub_type fset_sub_countMixin.
561-
562- Definition fset_sub_finMixin :=
563- Eval hnf in UniqFinMixin (undup_uniq _) mem_fset_sub_enum.
564- Canonical fset_sub_finType := Eval hnf in FinType fset_sub_type fset_sub_finMixin.
565- Canonical fset_sub_subfinType := [subFinType of fset_sub_type].
552+ HB.instance Definition _ :=
553+ Countable.copy fset_sub_type (pcan_type fset_sub_pickleK).
554+ HB.instance Definition _ := isFinite.Build fset_sub_type
555+ (Finite.uniq_enumP (undup_uniq _) mem_fset_sub_enum).
566556
567557Lemma enum_fsetE : enum_fset A = [seq val i | i <- enum fset_sub_type].
568558Proof . by rewrite enumT unlock val_fset_sub_enum. Qed .
@@ -618,11 +608,8 @@ End SeqFset.
618608
619609#[global] Hint Resolve keys_canonical sort_keys_uniq : core.
620610
621- Canonical finSetSubType K := [subType for (@enum_fset K)].
622- Definition finSetEqMixin (K : choiceType) := [eqMixin of {fset K} by <:].
623- Canonical finSetEqType (K : choiceType) := EqType {fset K} (finSetEqMixin K).
624- Definition finSetChoiceMixin (K : choiceType) := [choiceMixin of {fset K} by <:].
625- Canonical finSetChoiceType (K : choiceType) := ChoiceType {fset K} (finSetChoiceMixin K).
611+ HB.instance Definition _ K := [isSub for (@enum_fset K)].
612+ HB.instance Definition _ (K : choiceType) := [Choice of {fset K} by <:].
626613
627614Section FinPredStruct.
628615
@@ -1057,7 +1044,7 @@ Lemma imfset_rec (T : choiceType) (f : T -> K) (p : finmempred T)
10571044Proof .
10581045move=> PP v; have /imfsetP [k pk vv_eq] := valP v.
10591046pose vP := in_imfset f pk; suff -> : P v = P [` vP] by apply: PP.
1060- by congr P; apply/val_inj => /=; rewrite vv_eq.
1047+ by congr P; apply/val_inj => /=; rewrite - vv_eq (*FIXME: was rewrite vv_eq*) .
10611048Qed .
10621049
10631050Lemma mem_imfset (T : choiceType) (f : T -> K) (p : finmempred T) :
@@ -1138,7 +1125,7 @@ Lemma val_in_fset A (p : finmempred _) (k : A) :
11381125 (val k \in imfset key val p) = (in_mem k p).
11391126Proof . by rewrite mem_imfset ?in_finmempred //; exact: val_inj. Qed .
11401127
1141- Lemma in_fset_val A (p : finmempred [eqType of A] ) (k : K) :
1128+ Lemma in_fset_val A (p : finmempred A ) (k : K) :
11421129 (k \in imfset key val p) = if insub k is Some a then in_mem a p else false.
11431130Proof .
11441131have [a _ <- /=|kNA] := insubP; first by rewrite val_in_fset.
@@ -1156,7 +1143,7 @@ apply: (iffP (imfsetP _ _ _ _)) => [|[kA xkA]]; last by exists [`kA].
11561143by move=> /sig2_eqW[/= x Xx ->]; exists (valP x); rewrite fsetsubE.
11571144Qed .
11581145
1159- Lemma in_fset_valF A (p : finmempred [eqType of A] ) (k : K) : k \notin A ->
1146+ Lemma in_fset_valF A (p : finmempred A ) (k : K) : k \notin A ->
11601147 (k \in imfset key val p) = false.
11611148Proof . by apply: contraNF => /imfsetP[/= a Xa->]. Qed .
11621149
@@ -1370,7 +1357,7 @@ apply: (iffP fset_eqP) => AsubB a; first by rewrite -AsubB inE => /andP[].
13701357by rewrite inE; have [/AsubB|] := boolP (a \in A).
13711358Qed .
13721359
1373- Lemma fset_sub_val A (p : finmempred [eqType of A] ) :
1360+ Lemma fset_sub_val A (p : finmempred A ) :
13741361 (imfset key val p) `<=` A.
13751362Proof . by apply/fsubsetP => k /in_fset_valP []. Qed .
13761363
@@ -2128,11 +2115,11 @@ End FSetInE.
21282115Section Enum.
21292116
21302117Lemma enum_fset0 (T : choiceType) :
2131- enum [finType of fset0] = [::] :> seq (@fset0 T).
2118+ enum (fset0 : finType) = [::] :> seq (@fset0 T).
21322119Proof . by rewrite enumT unlock. Qed .
21332120
21342121Lemma enum_fset1 (T : choiceType) (x : T) :
2135- enum [finType of [ fset x]] = [:: [`fset11 x]].
2122+ enum ([ fset x] : finType) = [:: [`fset11 x]].
21362123Proof .
21372124apply/perm_small_eq=> //; apply/uniq_perm => //.
21382125 by apply/enum_uniq.
@@ -2231,8 +2218,7 @@ Proof. by move=> s; apply/fsetP=> x; rewrite !inE. Qed.
22312218Lemma unpickleK : cancel unpickle pickle.
22322219Proof . by move=> s; apply/setP=> x; rewrite !inE. Qed .
22332220
2234- Definition fset_finMixin := CanFinMixin pickleK.
2235- Canonical fset_finType := Eval hnf in FinType {fset T} fset_finMixin.
2221+ HB.instance Definition _ : fintype.isFinite {fset T} := CanIsFinite pickleK.
22362222
22372223Lemma card_fsets : #|{:{fset T}}| = 2^#|T|.
22382224Proof .
@@ -2449,8 +2435,8 @@ Section FSetMonoids.
24492435Import Monoid.
24502436Variable (T : choiceType).
24512437
2452- Canonical fsetU_monoid := Law (@fsetUA T) (@fset0U T) (@fsetU0 T).
2453- Canonical fsetU_comoid := ComLaw (@fsetUC T).
2438+ HB.instance Definition _ := isComLaw.Build {fset T} fset0 fsetU
2439+ (@fsetUA T) (@fsetUC T) (@fset0U T).
24542440
24552441End FSetMonoids.
24562442Section BigFOpsSeq.
@@ -2969,26 +2955,26 @@ Proof. by case. Qed.
29692955End FinMapEncode.
29702956
29712957Section FinMapEqType.
2972- Variable (K : choiceType) (V : eqType).
2958+ Variables (K : choiceType) (V : eqType).
29732959
2974- Definition finMap_eqMixin := CanEqMixin (@finMap_codeK K V).
2975- Canonical finMap_eqType := EqType {fmap K -> V} finMap_eqMixin .
2960+ HB.instance Definition _ := Equality.copy {fmap K -> V}
2961+ (can_type (@finMap_codeK K V)) .
29762962
29772963End FinMapEqType.
29782964
29792965Section FinMapChoiceType.
2980- Variable ( K V : choiceType) .
2966+ Variables K V : choiceType.
29812967
2982- Definition finMap_choiceMixin := CanChoiceMixin (@finMap_codeK K V).
2983- Canonical finMap_choiceType := ChoiceType {fmap K -> V} finMap_choiceMixin .
2968+ HB.instance Definition _ := Choice.copy {fmap K -> V}
2969+ (can_type (@finMap_codeK K V)) .
29842970
29852971End FinMapChoiceType.
29862972
29872973Section FinMapCountType.
2988- Variable ( K V : countType) .
2974+ Variables K V : countType.
29892975
2990- Definition finMap_countMixin := CanCountMixin (@finMap_codeK K V).
2991- Canonical finMap_countType := CountType {fmap K -> V} finMap_countMixin .
2976+ HB.instance Definition _ := Countable.copy {fmap K -> V}
2977+ (can_type (@finMap_codeK K V)) .
29922978
29932979End FinMapCountType.
29942980
@@ -3617,9 +3603,8 @@ Record fsfun := Fsfun {
36173603 fmap_of_fsfun k != default (val k)]
36183604}.
36193605
3620- Canonical fsfun_subType := Eval hnf in [subType for fmap_of_fsfun].
3621- Definition fsfun_eqMixin := [eqMixin of fsfun by <:].
3622- Canonical fsfun_eqType := EqType fsfun fsfun_eqMixin.
3606+ HB.instance Definition _ := [isSub for fmap_of_fsfun].
3607+ HB.instance Definition _ := [Equality of fsfun by <:].
36233608
36243609Fact fsfun_subproof (f : fsfun) :
36253610 forall (k : K) (kf : k \in fmap_of_fsfun f),
@@ -3773,15 +3758,10 @@ Definition fsfun_of_ffun key (K : choiceType) (V : eqType)
37733758 (S : {fset K}) (h : S -> V) (default : K -> V) :=
37743759 (Fsfun.of_ffun default h key).
37753760
3776- Definition fsfun_choiceMixin (K V : choiceType) (d : K -> V) :=
3777- [choiceMixin of fsfun d by <:].
3778- Canonical fsfun_choiceType (K V : choiceType) (d : K -> V) :=
3779- ChoiceType (fsfun d) (fsfun_choiceMixin d).
3780-
3781- Definition fsfun_countMixin (K V : countType) (d : K -> V) :=
3782- [countMixin of fsfun d by <:].
3783- Canonical fsfun_countType (K V : countType) (d : K -> V) :=
3784- CountType (fsfun d) (fsfun_countMixin d).
3761+ HB.instance Definition _ (K V : choiceType) (d : K -> V) :=
3762+ [Choice of fsfun d by <:].
3763+ HB.instance Definition _ (K V : countType) (d : K -> V) :=
3764+ [Countable of fsfun d by <:].
37853765
37863766Declare Scope fsfun_scope.
37873767Delimit Scope fsfun_scope with fsfun.
0 commit comments