@@ -492,14 +492,9 @@ Definition pred_of_finset (K : choiceType)
492492 (f : finSet K) : pred K := fun k => k \in (enum_fset f).
493493Canonical finSetPredType (K : choiceType) := PredType (@pred_of_finset K).
494494
495- Section FinSetCanonicals.
496-
497- Variable (K : choiceType).
498-
499- HB.instance Definition _ := [isSub for (@enum_fset K)].
500- HB.instance Definition _ := [Choice of {fset K} by <:].
501-
502- End FinSetCanonicals.
495+ HB.instance Definition _ (K : choiceType) := [isSub for @enum_fset K].
496+ HB.instance Definition _ (K : choiceType) := [Choice of {fset K} by <:].
497+ HB.instance Definition _ (T : countType) := [Countable of {fset T} by <:].
503498
504499Section FinTypeSet.
505500
@@ -516,7 +511,6 @@ Record fset_sub_type : predArgType :=
516511
517512HB.instance Definition _ := [isSub for fsval].
518513HB.instance Definition _ := [Choice of fset_sub_type by <:].
519- HB.instance Definition _ (T : countType) := [Countable of {fset T} by <:].
520514
521515Definition fset_sub_enum : seq fset_sub_type :=
522516 undup (pmap insub (enum_fset A)).
@@ -575,33 +569,32 @@ Definition set_of_fset (K : choiceType) (A : {fset K}) : {set A} :=
575569
576570Arguments pred_of_finset : simpl never.
577571
572+ HB.lock Definition seq_fset
573+ (finset_key : unit) (K : choiceType) (s : seq K) : {fset K} :=
574+ mkFinSet (@canonical_sort_keys K s).
575+
578576Section SeqFset.
579577
580- Variable finset_key : unit.
581- Definition seq_fset : forall K : choiceType, seq K -> {fset K} :=
582- locked_with finset_key (fun K s => mkFinSet (@canonical_sort_keys K s)).
578+ Variable (finset_key : unit) (K : choiceType) (s : seq K).
583579
584- Variable (K : choiceType) (s : seq K ).
580+ Notation seq_fset := (seq_fset finset_key ).
585581
586582Lemma seq_fsetE : seq_fset s =i s.
587- Proof . by move=> a; rewrite [seq_fset] unlock sort_keysE. Qed .
583+ Proof . by move=> a; rewrite unlock sort_keysE. Qed .
588584
589585Lemma size_seq_fset : size (seq_fset s) = size (undup s).
590- Proof . by rewrite [seq_fset] unlock /= size_sort_keys. Qed .
586+ Proof . by rewrite unlock /= size_sort_keys. Qed .
591587
592588Lemma seq_fset_uniq : uniq (seq_fset s).
593- Proof . by rewrite [seq_fset] unlock /= sort_keys_uniq. Qed .
589+ Proof . by rewrite unlock /= sort_keys_uniq. Qed .
594590
595591Lemma seq_fset_perm : perm_eq (seq_fset s) (undup s).
596- Proof . by rewrite [seq_fset] unlock //= sort_keys_perm. Qed .
592+ Proof . by rewrite unlock //= sort_keys_perm. Qed .
597593
598594End SeqFset.
599595
600596#[global] Hint Resolve keys_canonical sort_keys_uniq : core.
601597
602- HB.instance Definition _ K := [isSub for (@enum_fset K)].
603- HB.instance Definition _ (K : choiceType) := [Choice of {fset K} by <:].
604-
605598Section FinPredStruct.
606599
607600Structure finpredType (T : eqType) := FinPredType {
@@ -739,46 +732,25 @@ Canonical seq_finpredType (T : eqType) :=
739732
740733End CanonicalFinPred.
741734
742- Local Notation imfset_def key :=
743- (fun (T K : choiceType) (f : T -> K) (p : finmempred T)
744- => seq_fset key [seq f x | x <- enum_finmem p]).
745- Local Notation imfset2_def key :=
746- (fun (K T1 : choiceType) (T2 : T1 -> choiceType)
747- (f : forall x : T1, T2 x -> K)
748- (p1 : finmempred T1) (p2 : forall x : T1, finmempred (T2 x)) =>
749- seq_fset key [seq f x y | x <- enum_finmem p1, y <- enum_finmem (p2 x)]).
750-
751- Module Type ImfsetSig.
752- Parameter imfset : forall (key : unit) (T K : choiceType)
753- (f : T -> K) (p : finmempred T), {fset K}.
754- Parameter imfset2 : forall (key : unit) (K T1 : choiceType)
755- (T2 : T1 -> choiceType)(f : forall x : T1, T2 x -> K)
756- (p1 : finmempred T1) (p2 : forall x : T1, finmempred (T2 x)), {fset K}.
757- Axiom imfsetE : forall key, imfset key = imfset_def key.
758- Axiom imfset2E : forall key, imfset2 key = imfset2_def key.
759- End ImfsetSig.
760-
761- Module Imfset : ImfsetSig.
762- Definition imfset key := imfset_def key.
763- Definition imfset2 key := imfset2_def key.
764- Lemma imfsetE key : imfset key = imfset_def key. Proof . by []. Qed .
765- Lemma imfset2E key : imfset2 key = imfset2_def key. Proof . by []. Qed .
766- End Imfset.
767-
768- Notation imfset := Imfset.imfset.
769- Notation imfset2 := Imfset.imfset2.
770- Canonical imfset_unlock k := Unlockable (Imfset.imfsetE k).
771- Canonical imfset2_unlock k := Unlockable (Imfset.imfset2E k).
735+ HB.lock Definition imfset
736+ (key : unit) (T K : choiceType) (f : T -> K) (p : finmempred T) : {fset K} :=
737+ seq_fset key [seq f x | x <- enum_finmem p].
738+
739+ HB.lock Definition imfset2
740+ (key : unit) (K T1 : choiceType) (T2 : T1 -> choiceType)
741+ (f : forall x : T1, T2 x -> K)
742+ (p1 : finmempred T1) (p2 : forall x : T1, finmempred (T2 x)) : {fset K} :=
743+ seq_fset key [seq f x y | x <- enum_finmem p1, y <- enum_finmem (p2 x)].
772744
773745Notation "A `=` B" := (A = B :> {fset _}) (only parsing) : fset_scope.
774746Notation "A `<>` B" := (A <> B :> {fset _}) (only parsing) : fset_scope.
775747Notation "A `==` B" := (A == B :> {fset _}) (only parsing) : fset_scope.
776748Notation "A `!=` B" := (A != B :> {fset _}) (only parsing) : fset_scope.
777749Notation "A `=P` B" := (A =P B :> {fset _}) (only parsing) : fset_scope.
778750
779- Notation "f @`[ key ] A" := (Imfset. imfset key f (mem A)) : fset_scope.
751+ Notation "f @`[ key ] A" := (imfset key f (mem A)) : fset_scope.
780752Notation "f @2`[ key ] ( A , B )" :=
781- (Imfset. imfset2 key f (mem A) (fun x => (mem (B x)))) : fset_scope.
753+ (imfset2 key f (mem A) (fun x => (mem (B x)))) : fset_scope.
782754
783755Fact imfset_key : unit. Proof . exact: tt. Qed .
784756
@@ -3757,17 +3729,9 @@ End FsfunDef.
37573729
37583730Coercion fun_of_fsfun : fsfun >-> Funclass.
37593731
3760- Module Type FinSuppSig.
3761- Axiom fs : forall (K : choiceType) (V : eqType) (default : K -> V),
3762- fsfun default -> {fset K}.
3763- Axiom E : fs = (fun K V d f => domf (@fmap_of_fsfun K V d f)).
3764- End FinSuppSig.
3765- Module FinSupp : FinSuppSig.
3766- Definition fs := (fun K V d f => domf (@fmap_of_fsfun K V d f)).
3767- Definition E := erefl fs.
3768- End FinSupp.
3769- Notation finsupp := FinSupp.fs.
3770- Canonical unlockable_finsupp := Unlockable FinSupp.E.
3732+ HB.lock Definition finsupp
3733+ (K : choiceType) (V : eqType) (d : K -> V) (f : fsfun d) : {fset K} :=
3734+ domf (@fmap_of_fsfun K V d f).
37713735
37723736Section FSfunBasics.
37733737
@@ -3844,38 +3808,34 @@ Notation "{ 'fsfun' 'of' x => dflt }" := {fsfun of x : _ => dflt}
38443808Notation "{ 'fsfun' 'with ' dflt }" := {fsfun of _ => dflt}
38453809 (at level 0, only parsing) : type_scope.
38463810
3847- Module Type FsfunSig.
3848- Section FsfunSig.
3849- Variables (K : choiceType) (V : eqType) (default : K -> V).
3850-
3851- Parameter of_ffun : forall (S : {fset K}), (S -> V) -> unit -> fsfun default.
3852- Variables (S : {fset K}) (h : S -> V).
3853- Axiom of_ffunE :forall key x, of_ffun h key x = odflt (default x) (omap h (insub x)).
3854-
3855- End FsfunSig.
3856- End FsfunSig.
3857-
3858- Module Fsfun : FsfunSig.
38593811Section FsfunOfFinfun.
38603812
3861- Variables (K : choiceType) (V : eqType) (default : K -> V)
3862- (S : {fset K}) (h : S -> V).
3813+ Variables (K : choiceType) (V : eqType).
3814+ Variables (S : {fset K}) (h : S -> V) (default : K -> V).
38633815
3864- Let fmap :=
3816+ Let fmap : {fmap K -> V} : =
38653817 [fmap a : [fsetval a in {: S} | h a != default (val a)]
38663818 => h (fincl (fset_sub_val _ _) a)].
38673819
3868- Fact fmapP a : fmap a != default (val a).
3820+ Fact fsfun_of_ffun_subproof (a : domf fmap) : fmap a != default (val a).
38693821Proof .
38703822rewrite ffunE; have /in_fset_valP [a_in_S] := valP a.
38713823by have -> : [` a_in_S] = fincl (fset_sub_val _ _) a by exact/val_inj.
38723824Qed .
38733825
3874- Definition of_ffun (k : unit) := fsfun_of_can_ffun fmapP.
3826+ End FsfunOfFinfun.
3827+
3828+ HB.lock Definition fsfun_of_ffun
3829+ (k : unit) (K : choiceType) (V : eqType)
3830+ (S : {fset K}) (h : S -> V) (default : K -> V) :=
3831+ fsfun_of_can_ffun (@fsfun_of_ffun_subproof K V S h default).
38753832
3876- Lemma of_ffunE key x : of_ffun key x = odflt (default x) (omap h (insub x)).
3833+ Lemma fsfun_ffun
3834+ (k : unit) (K : choiceType) (V : eqType)
3835+ (default : K -> V) (S : {fset K}) (h : S -> V) (x : K) :
3836+ fsfun_of_ffun k h default x = odflt (default x) (omap h (insub x)).
38773837Proof .
3878- rewrite /fun_of_fsfun /=.
3838+ rewrite unlock /fun_of_fsfun /=.
38793839case: insubP => /= [u _ <-|xNS]; last first.
38803840 case: fndP => //= kf; rewrite !ffunE /=.
38813841 by set y := (X in h X); rewrite (valP y) in xNS.
@@ -3884,17 +3844,8 @@ case: fndP => /= [kf|].
38843844by rewrite inE /= -topredE /= negbK => /eqP ->.
38853845Qed .
38863846
3887- End FsfunOfFinfun.
3888- End Fsfun.
3889- Canonical fsfun_of_funE K V default S h key x :=
3890- Unlockable (@Fsfun.of_ffunE K V default S h key x).
3891-
38923847Fact fsfun_key : unit. Proof . exact: tt. Qed .
38933848
3894- Definition fsfun_of_ffun key (K : choiceType) (V : eqType)
3895- (S : {fset K}) (h : S -> V) (default : K -> V) :=
3896- (Fsfun.of_ffun default h key).
3897-
38983849HB.instance Definition _ (K V : choiceType) (d : K -> V) :=
38993850 [Choice of fsfun d by <:].
39003851
@@ -4008,18 +3959,13 @@ Notation "[ 'fsfun' ]" := [fsfun for _]
40083959Section FsfunTheory.
40093960Variables (key : unit) (K : choiceType) (V : eqType) (default : K -> V).
40103961
4011- Lemma fsfun_ffun (S : {fset K}) (h : S -> V) (x : K) :
4012- [fsfun[key] a : S => h a | default a] x =
4013- odflt (default x) (omap h (insub x)).
4014- Proof . by rewrite unlock. Qed .
4015-
40163962Lemma fsfun_fun (S : {fset K}) (h : K -> V) (x : K) :
40173963 [fsfun[key] a in S => h a | default a] x =
40183964 if x \in S then h x else (default x).
40193965Proof . by rewrite fsfun_ffun; case: insubP => //= [u -> ->|/negPf ->]. Qed .
40203966
40213967Lemma fsfun0E : [fsfun for default] =1 default.
4022- Proof . by move=> x; rewrite unlock insubF ?inE. Qed .
3968+ Proof . by move=> x; rewrite fsfun_ffun insubF ?inE. Qed .
40233969
40243970Definition fsfunE := (fsfun_fun, fsfun0E).
40253971
@@ -4034,7 +3980,7 @@ Qed.
40343980Lemma finsupp_sub (S : {fset K}) (h : S -> V) :
40353981 finsupp [fsfun[key] a : S => h a | default a] `<=` S.
40363982Proof .
4037- apply/fsubsetP => a; rewrite mem_finsupp unlock /= .
3983+ apply/fsubsetP => a; rewrite mem_finsupp fsfun_ffun .
40383984by case: insubP => //=; rewrite eqxx.
40393985Qed .
40403986
0 commit comments