Skip to content

Commit e2c2589

Browse files
committed
Use HB.lock
1 parent a212a4f commit e2c2589

File tree

1 file changed

+44
-98
lines changed

1 file changed

+44
-98
lines changed

finmap.v

Lines changed: 44 additions & 98 deletions
Original file line numberDiff line numberDiff line change
@@ -494,14 +494,9 @@ Definition pred_of_finset (K : choiceType)
494494
(f : finSet K) : pred K := fun k => k \in (enum_fset f).
495495
Canonical finSetPredType (K : choiceType) := PredType (@pred_of_finset K).
496496

497-
Section FinSetCanonicals.
498-
499-
Variable (K : choiceType).
500-
501-
HB.instance Definition _ := [isSub for (@enum_fset K)].
502-
HB.instance Definition _ := [Choice of {fset K} by <:].
503-
504-
End FinSetCanonicals.
497+
HB.instance Definition _ (K : choiceType) := [isSub for @enum_fset K].
498+
HB.instance Definition _ (K : choiceType) := [Choice of {fset K} by <:].
499+
HB.instance Definition _ (T : countType) := [Countable of {fset T} by <:].
505500

506501
Section FinTypeSet.
507502

@@ -518,7 +513,6 @@ Record fset_sub_type : predArgType :=
518513

519514
HB.instance Definition _ := [isSub for fsval].
520515
HB.instance Definition _ := [Choice of fset_sub_type by <:].
521-
HB.instance Definition _ (T : countType) := [Countable of {fset T} by <:].
522516

523517
Definition fset_sub_enum : seq fset_sub_type :=
524518
undup (pmap insub (enum_fset A)).
@@ -577,33 +571,32 @@ Definition set_of_fset (K : choiceType) (A : {fset K}) : {set A} :=
577571

578572
Arguments pred_of_finset : simpl never.
579573

574+
HB.lock Definition seq_fset
575+
(finset_key : unit) (K : choiceType) (s : seq K) : {fset K} :=
576+
mkFinSet (@canonical_sort_keys K s).
577+
580578
Section SeqFset.
581579

582-
Variable finset_key : unit.
583-
Definition seq_fset : forall K : choiceType, seq K -> {fset K} :=
584-
locked_with finset_key (fun K s => mkFinSet (@canonical_sort_keys K s)).
580+
Variable (finset_key : unit) (K : choiceType) (s : seq K).
585581

586-
Variable (K : choiceType) (s : seq K).
582+
Notation seq_fset := (seq_fset finset_key).
587583

588584
Lemma seq_fsetE : seq_fset s =i s.
589-
Proof. by move=> a; rewrite [seq_fset]unlock sort_keysE. Qed.
585+
Proof. by move=> a; rewrite unlock sort_keysE. Qed.
590586

591587
Lemma size_seq_fset : size (seq_fset s) = size (undup s).
592-
Proof. by rewrite [seq_fset]unlock /= size_sort_keys. Qed.
588+
Proof. by rewrite unlock /= size_sort_keys. Qed.
593589

594590
Lemma seq_fset_uniq : uniq (seq_fset s).
595-
Proof. by rewrite [seq_fset]unlock /= sort_keys_uniq. Qed.
591+
Proof. by rewrite unlock /= sort_keys_uniq. Qed.
596592

597593
Lemma seq_fset_perm : perm_eq (seq_fset s) (undup s).
598-
Proof. by rewrite [seq_fset]unlock //= sort_keys_perm. Qed.
594+
Proof. by rewrite unlock //= sort_keys_perm. Qed.
599595

600596
End SeqFset.
601597

602598
#[global] Hint Resolve keys_canonical sort_keys_uniq : core.
603599

604-
HB.instance Definition _ K := [isSub for (@enum_fset K)].
605-
HB.instance Definition _ (K : choiceType) := [Choice of {fset K} by <:].
606-
607600
Section FinPredStruct.
608601

609602
Structure finpredType (T : eqType) := FinPredType {
@@ -746,46 +739,25 @@ Canonical seq_finpredType (T : eqType) :=
746739

747740
End CanonicalFinPred.
748741

749-
Local Notation imfset_def key :=
750-
(fun (T K : choiceType) (f : T -> K) (p : finmempred T)
751-
=> seq_fset key [seq f x | x <- enum_finmem p]).
752-
Local Notation imfset2_def key :=
753-
(fun (K T1 : choiceType) (T2 : T1 -> choiceType)
754-
(f : forall x : T1, T2 x -> K)
755-
(p1 : finmempred T1) (p2 : forall x : T1, finmempred (T2 x)) =>
756-
seq_fset key [seq f x y | x <- enum_finmem p1, y <- enum_finmem (p2 x)]).
757-
758-
Module Type ImfsetSig.
759-
Parameter imfset : forall (key : unit) (T K : choiceType)
760-
(f : T -> K) (p : finmempred T), {fset K}.
761-
Parameter imfset2 : forall (key : unit) (K T1 : choiceType)
762-
(T2 : T1 -> choiceType)(f : forall x : T1, T2 x -> K)
763-
(p1 : finmempred T1) (p2 : forall x : T1, finmempred (T2 x)), {fset K}.
764-
Axiom imfsetE : forall key, imfset key = imfset_def key.
765-
Axiom imfset2E : forall key, imfset2 key = imfset2_def key.
766-
End ImfsetSig.
767-
768-
Module Imfset : ImfsetSig.
769-
Definition imfset key := imfset_def key.
770-
Definition imfset2 key := imfset2_def key.
771-
Lemma imfsetE key : imfset key = imfset_def key. Proof. by []. Qed.
772-
Lemma imfset2E key : imfset2 key = imfset2_def key. Proof. by []. Qed.
773-
End Imfset.
774-
775-
Notation imfset := Imfset.imfset.
776-
Notation imfset2 := Imfset.imfset2.
777-
Canonical imfset_unlock k := Unlockable (Imfset.imfsetE k).
778-
Canonical imfset2_unlock k := Unlockable (Imfset.imfset2E k).
742+
HB.lock Definition imfset
743+
(key : unit) (T K : choiceType) (f : T -> K) (p : finmempred T) : {fset K} :=
744+
seq_fset key [seq f x | x <- enum_finmem p].
745+
746+
HB.lock Definition imfset2
747+
(key : unit) (K T1 : choiceType) (T2 : T1 -> choiceType)
748+
(f : forall x : T1, T2 x -> K)
749+
(p1 : finmempred T1) (p2 : forall x : T1, finmempred (T2 x)) : {fset K} :=
750+
seq_fset key [seq f x y | x <- enum_finmem p1, y <- enum_finmem (p2 x)].
779751

780752
Notation "A `=` B" := (A = B :> {fset _}) (only parsing) : fset_scope.
781753
Notation "A `<>` B" := (A <> B :> {fset _}) (only parsing) : fset_scope.
782754
Notation "A `==` B" := (A == B :> {fset _}) (only parsing) : fset_scope.
783755
Notation "A `!=` B" := (A != B :> {fset _}) (only parsing) : fset_scope.
784756
Notation "A `=P` B" := (A =P B :> {fset _}) (only parsing) : fset_scope.
785757

786-
Notation "f @`[ key ] A" := (Imfset.imfset key f (mem A)) : fset_scope.
758+
Notation "f @`[ key ] A" := (imfset key f (mem A)) : fset_scope.
787759
Notation "f @2`[ key ] ( A , B )" :=
788-
(Imfset.imfset2 key f (mem A) (fun x => (mem (B x)))) : fset_scope.
760+
(imfset2 key f (mem A) (fun x => (mem (B x)))) : fset_scope.
789761

790762
Fact imfset_key : unit. Proof. exact: tt. Qed.
791763

@@ -3766,17 +3738,9 @@ End FsfunDef.
37663738

37673739
Coercion fun_of_fsfun : fsfun >-> Funclass.
37683740

3769-
Module Type FinSuppSig.
3770-
Axiom fs : forall (K : choiceType) (V : eqType) (default : K -> V),
3771-
fsfun default -> {fset K}.
3772-
Axiom E : fs = (fun K V d f => domf (@fmap_of_fsfun K V d f)).
3773-
End FinSuppSig.
3774-
Module FinSupp : FinSuppSig.
3775-
Definition fs := (fun K V d f => domf (@fmap_of_fsfun K V d f)).
3776-
Definition E := erefl fs.
3777-
End FinSupp.
3778-
Notation finsupp := FinSupp.fs.
3779-
Canonical unlockable_finsupp := Unlockable FinSupp.E.
3741+
HB.lock Definition finsupp
3742+
(K : choiceType) (V : eqType) (d : K -> V) (f : fsfun d) : {fset K} :=
3743+
domf (@fmap_of_fsfun K V d f).
37803744

37813745
Section FSfunBasics.
37823746

@@ -3853,38 +3817,34 @@ Notation "{ 'fsfun' 'of' x => dflt }" := {fsfun of x : _ => dflt}
38533817
Notation "{ 'fsfun' 'with' dflt }" := {fsfun of _ => dflt}
38543818
(at level 0, only parsing) : type_scope.
38553819

3856-
Module Type FsfunSig.
3857-
Section FsfunSig.
3858-
Variables (K : choiceType) (V : eqType) (default : K -> V).
3859-
3860-
Parameter of_ffun : forall (S : {fset K}), (S -> V) -> unit -> fsfun default.
3861-
Variables (S : {fset K}) (h : S -> V).
3862-
Axiom of_ffunE :forall key x, of_ffun h key x = odflt (default x) (omap h (insub x)).
3863-
3864-
End FsfunSig.
3865-
End FsfunSig.
3866-
3867-
Module Fsfun : FsfunSig.
38683820
Section FsfunOfFinfun.
38693821

3870-
Variables (K : choiceType) (V : eqType) (default : K -> V)
3871-
(S : {fset K}) (h : S -> V).
3822+
Variables (K : choiceType) (V : eqType).
3823+
Variables (S : {fset K}) (h : S -> V) (default : K -> V).
38723824

3873-
Let fmap :=
3825+
Let fmap : {fmap K -> V} :=
38743826
[fmap a : [fsetval a in {: S} | h a != default (val a)]
38753827
=> h (fincl (fset_sub_val _ _) a)].
38763828

3877-
Fact fmapP a : fmap a != default (val a).
3829+
Fact fsfun_of_ffun_subproof (a : domf fmap) : fmap a != default (val a).
38783830
Proof.
38793831
rewrite ffunE; have /in_fset_valP [a_in_S] := valP a.
38803832
by have -> : [` a_in_S] = fincl (fset_sub_val _ _) a by exact/val_inj.
38813833
Qed.
38823834

3883-
Definition of_ffun (k : unit) := fsfun_of_can_ffun fmapP.
3835+
End FsfunOfFinfun.
3836+
3837+
HB.lock Definition fsfun_of_ffun
3838+
(k : unit) (K : choiceType) (V : eqType)
3839+
(S : {fset K}) (h : S -> V) (default : K -> V) :=
3840+
fsfun_of_can_ffun (@fsfun_of_ffun_subproof K V S h default).
38843841

3885-
Lemma of_ffunE key x : of_ffun key x = odflt (default x) (omap h (insub x)).
3842+
Lemma fsfun_ffun
3843+
(k : unit) (K : choiceType) (V : eqType)
3844+
(default : K -> V) (S : {fset K}) (h : S -> V) (x : K) :
3845+
fsfun_of_ffun k h default x = odflt (default x) (omap h (insub x)).
38863846
Proof.
3887-
rewrite /fun_of_fsfun /=.
3847+
rewrite unlock /fun_of_fsfun /=.
38883848
case: insubP => /= [u _ <-|xNS]; last first.
38893849
case: fndP => //= kf; rewrite !ffunE /=.
38903850
by set y := (X in h X); rewrite (valP y) in xNS.
@@ -3893,17 +3853,8 @@ case: fndP => /= [kf|].
38933853
by rewrite inE /= -topredE /= negbK => /eqP ->.
38943854
Qed.
38953855

3896-
End FsfunOfFinfun.
3897-
End Fsfun.
3898-
Canonical fsfun_of_funE K V default S h key x :=
3899-
Unlockable (@Fsfun.of_ffunE K V default S h key x).
3900-
39013856
Fact fsfun_key : unit. Proof. exact: tt. Qed.
39023857

3903-
Definition fsfun_of_ffun key (K : choiceType) (V : eqType)
3904-
(S : {fset K}) (h : S -> V) (default : K -> V) :=
3905-
(Fsfun.of_ffun default h key).
3906-
39073858
HB.instance Definition _ (K V : choiceType) (d : K -> V) :=
39083859
[Choice of fsfun d by <:].
39093860

@@ -4017,18 +3968,13 @@ Notation "[ 'fsfun' ]" := [fsfun for _]
40173968
Section FsfunTheory.
40183969
Variables (key : unit) (K : choiceType) (V : eqType) (default : K -> V).
40193970

4020-
Lemma fsfun_ffun (S : {fset K}) (h : S -> V) (x : K) :
4021-
[fsfun[key] a : S => h a | default a] x =
4022-
odflt (default x) (omap h (insub x)).
4023-
Proof. by rewrite unlock. Qed.
4024-
40253971
Lemma fsfun_fun (S : {fset K}) (h : K -> V) (x : K) :
40263972
[fsfun[key] a in S => h a | default a] x =
40273973
if x \in S then h x else (default x).
40283974
Proof. by rewrite fsfun_ffun; case: insubP => //= [u -> ->|/negPf ->]. Qed.
40293975

40303976
Lemma fsfun0E : [fsfun for default] =1 default.
4031-
Proof. by move=> x; rewrite unlock insubF ?inE. Qed.
3977+
Proof. by move=> x; rewrite fsfun_ffun insubF ?inE. Qed.
40323978

40333979
Definition fsfunE := (fsfun_fun, fsfun0E).
40343980

@@ -4043,7 +3989,7 @@ Qed.
40433989
Lemma finsupp_sub (S : {fset K}) (h : S -> V) :
40443990
finsupp [fsfun[key] a : S => h a | default a] `<=` S.
40453991
Proof.
4046-
apply/fsubsetP => a; rewrite mem_finsupp unlock /=.
3992+
apply/fsubsetP => a; rewrite mem_finsupp fsfun_ffun.
40473993
by case: insubP => //=; rewrite eqxx.
40483994
Qed.
40493995

0 commit comments

Comments
 (0)