Skip to content

Commit 02b0059

Browse files
committed
Remove the use of the phantom type
1 parent c0cf3df commit 02b0059

File tree

1 file changed

+22
-22
lines changed

1 file changed

+22
-22
lines changed

finmap.v

Lines changed: 22 additions & 22 deletions
Original file line numberDiff line numberDiff line change
@@ -338,7 +338,10 @@ Reserved Notation "[ 'f' 'setval' x : A | P ]"
338338
Reserved Notation "[ 'f' 'setval' x : A | P & Q ]"
339339
(at level 0, x at level 99, format "[ 'f' 'setval' x : A | P & Q ]").
340340

341-
Reserved Notation "{fmap T }" (at level 0, format "{fmap T }").
341+
Reserved Notation "{ 'fmap' K -> T }"
342+
(at level 0, K at level 98, T at level 99, format "{ 'fmap' K -> T }").
343+
Reserved Notation "[ 'fmap' 'of' K -> T ]"
344+
(at level 0, K at level 98, T at level 99, format "[ 'fmap' 'of' K -> T ]").
342345
Reserved Notation "x .[ k <- v ]"
343346
(left associativity, format "x .[ k <- v ]").
344347
Reserved Notation "x .[~ k ]" (format "x .[~ k ]").
@@ -476,12 +479,14 @@ Structure finSet : Type := mkFinSet {
476479
_ : canonical_keys enum_fset
477480
}.
478481

482+
#[deprecated(note="Use finSet instead")]
479483
Definition finset_of (_ : phant K) := finSet.
480484

481485
End Def.
482486

487+
#[warning="-deprecated-reference"]
483488
Identity Coercion type_of_finset : finset_of >-> finSet.
484-
Notation "{fset T }" := (@finset_of _ (Phant T)) : type_scope.
489+
Notation "{fset T }" := (finSet T) : type_scope.
485490

486491
Definition pred_of_finset (K : choiceType)
487492
(f : finSet K) : pred K := fun k => k \in (enum_fset f).
@@ -547,6 +552,7 @@ Proof. by rewrite cardE enum_fsetE size_map. Qed.
547552

548553
End FinTypeSet.
549554

555+
#[warning="-deprecated-reference"]
550556
Identity Coercion finSet_sub_type : finset_of >-> finSet.
551557
Coercion fset_sub_type : finSet >-> predArgType.
552558
#[global] Hint Resolve fsvalP fset_uniq mem_fset_sub_enum : core.
@@ -673,7 +679,7 @@ Canonical fin_finpred (T : eqType) (pT : finpredType T) (p : pT) :=
673679
(enum_finpred_uniq p) (enum_finpredE p)).
674680

675681
Definition finpred_of (T : eqType) (pT : predType T) (p : pT)
676-
(fp : finpred pT) & phantom pT fp : finpred pT := fp.
682+
(fp : finpred pT) : finpred pT := fp.
677683

678684
Structure finmempred (T : eqType) := FinMemPred {
679685
pred_of_finmempred :> mem_pred T;
@@ -692,7 +698,7 @@ Lemma enum_finmemE (T : eqType) (p : finmempred T) :
692698
Proof. by case: p => ?[]. Qed.
693699

694700
Definition finmempred_of (T : eqType) (P : pred T)
695-
(mP : finmempred T) & phantom (mem_pred T) mP : finmempred T := mP.
701+
(mP : finmempred T) : finmempred T := mP.
696702

697703
Canonical mem_fin (T : eqType) (pT : predType T) (p : finpred pT) :=
698704
@FinMemPred _ (mem p)
@@ -735,22 +741,19 @@ End CanonicalFinPred.
735741

736742
Local Notation imfset_def key :=
737743
(fun (T K : choiceType) (f : T -> K) (p : finmempred T)
738-
of phantom (mem_pred T) p => seq_fset key [seq f x | x <- enum_finmem p]).
744+
=> seq_fset key [seq f x | x <- enum_finmem p]).
739745
Local Notation imfset2_def key :=
740746
(fun (K T1 : choiceType) (T2 : T1 -> choiceType)
741747
(f : forall x : T1, T2 x -> K)
742-
(p1 : finmempred T1) (p2 : forall x : T1, finmempred (T2 x))
743-
of phantom (mem_pred T1) p1 & phantom (forall x, mem_pred (T2 x)) p2 =>
748+
(p1 : finmempred T1) (p2 : forall x : T1, finmempred (T2 x)) =>
744749
seq_fset key [seq f x y | x <- enum_finmem p1, y <- enum_finmem (p2 x)]).
745750

746751
Module Type ImfsetSig.
747752
Parameter imfset : forall (key : unit) (T K : choiceType)
748-
(f : T -> K) (p : finmempred T),
749-
phantom (mem_pred T) p -> {fset K}.
753+
(f : T -> K) (p : finmempred T), {fset K}.
750754
Parameter imfset2 : forall (key : unit) (K T1 : choiceType)
751755
(T2 : T1 -> choiceType)(f : forall x : T1, T2 x -> K)
752-
(p1 : finmempred T1) (p2 : forall x : T1, finmempred (T2 x)),
753-
phantom (mem_pred T1) p1 -> phantom (forall x, mem_pred (T2 x)) p2 -> {fset K}.
756+
(p1 : finmempred T1) (p2 : forall x : T1, finmempred (T2 x)), {fset K}.
754757
Axiom imfsetE : forall key, imfset key = imfset_def key.
755758
Axiom imfset2E : forall key, imfset2 key = imfset2_def key.
756759
End ImfsetSig.
@@ -762,11 +765,8 @@ Lemma imfsetE key : imfset key = imfset_def key. Proof. by []. Qed.
762765
Lemma imfset2E key : imfset2 key = imfset2_def key. Proof. by []. Qed.
763766
End Imfset.
764767

765-
Notation imfset key f p :=
766-
(Imfset.imfset key f (Phantom _ (pred_of_finmempred p))).
767-
Notation imfset2 key f p q :=
768-
(Imfset.imfset2 key f (Phantom _ (pred_of_finmempred p))
769-
(Phantom _ (fun x => (pred_of_finmempred (q x))))).
768+
Notation imfset := Imfset.imfset.
769+
Notation imfset2 := Imfset.imfset2.
770770
Canonical imfset_unlock k := Unlockable (Imfset.imfsetE k).
771771
Canonical imfset2_unlock k := Unlockable (Imfset.imfset2E k).
772772

@@ -776,11 +776,9 @@ Notation "A `==` B" := (A == B :> {fset _}) (only parsing) : fset_scope.
776776
Notation "A `!=` B" := (A != B :> {fset _}) (only parsing) : fset_scope.
777777
Notation "A `=P` B" := (A =P B :> {fset _}) (only parsing) : fset_scope.
778778

779-
Notation "f @`[ key ] A" :=
780-
(Imfset.imfset key f (Phantom _ (mem A))) : fset_scope.
779+
Notation "f @`[ key ] A" := (Imfset.imfset key f (mem A)) : fset_scope.
781780
Notation "f @2`[ key ] ( A , B )" :=
782-
(Imfset.imfset2 key f (Phantom _ (mem A)) (Phantom _ (fun x => (mem (B x)))))
783-
: fset_scope.
781+
(Imfset.imfset2 key f (mem A) (fun x => (mem (B x)))) : fset_scope.
784782

785783
Fact imfset_key : unit. Proof. exact: tt. Qed.
786784

@@ -2984,14 +2982,15 @@ Record finMap : Type := FinMap {
29842982
ffun_of_fmap :> {ffun domf -> V}
29852983
}.
29862984

2985+
#[deprecated(note="Use finMap instead.")]
29872986
Definition finmap_of (_ : phant (K -> V)) := finMap.
29882987

29892988
Let T_ (domf : {fset K}) := {ffun domf -> V}.
29902989
Local Notation finMap' := {domf : _ & T_ domf}.
29912990

29922991
End DefMap.
29932992

2994-
Notation "{fmap T }" := (@finmap_of _ _ (Phant T)) : type_scope.
2993+
Notation "{ 'fmap' K -> T }" := (finMap K T) : type_scope.
29952994

29962995
Definition pred_of_finmap (K : choiceType) (V : Type)
29972996
(f : {fmap K -> V}) : pred K := mem (domf f).
@@ -3045,7 +3044,8 @@ Arguments fmap0 {K V}.
30453044
Arguments setf : simpl never.
30463045
Arguments fnd : simpl never.
30473046

3048-
Notation "[ 'fmap' 'of' T ]" := (fmap0 : {fmap T}) (only parsing) : fmap_scope.
3047+
Notation "[ 'fmap' 'of' K -> T ]" :=
3048+
(fmap0 : {fmap K -> T}) (only parsing) : fmap_scope.
30493049
Notation "[fmap]" := fmap0 : fmap_scope.
30503050
Notation "x .[ k <- v ]" := (setf x k v) : fmap_scope.
30513051
Notation "f .[? k ]" := (fnd f k) : fmap_scope.

0 commit comments

Comments
 (0)