aboutsummaryrefslogtreecommitdiff
path: root/mathcomp/ssreflect
diff options
context:
space:
mode:
authorAnton Trunov2018-11-16 10:55:30 +0100
committerAnton Trunov2018-12-04 14:25:53 +0100
commit79aa2b1ab5b233f103cd3e402094cd93d9028866 (patch)
tree4d3afa56b14db2de71dbd52c959e71c9766e66fb /mathcomp/ssreflect
parente99b8b9695cdb53492e63077cb0dd551c4a06dc3 (diff)
Remove `_ : Type` from packed classes
This increases performance 10% - 15% for Coq v8.6.1 - v8.9.dev. Tested on a Debain-based 16-core build server and a Macbook Pro laptop with 2,3 GHz Intel Core i5. | | Compilation time, old | Compilation | Speedup | | | (mathcomp commit 967088a6f87) | time, new | | | Coq 8.6.1 | 10min 33s | 9min 10s | 15% | | Coq 8.7.2 | 10min 12s | 8min 50s | 15% | | Coq 8.8.2 | 9min 39s | 8min 32s | 13% | | Coq 8.9.dev(05d827c800544) | 9min 12s | 8min 16s | 11% | | | | | | It seems Coq at some point fixed the problem `_ : Type` was supposed to solve.
Diffstat (limited to 'mathcomp/ssreflect')
-rw-r--r--mathcomp/ssreflect/choice.v28
-rw-r--r--mathcomp/ssreflect/eqtype.v6
-rw-r--r--mathcomp/ssreflect/fintype.v22
-rw-r--r--mathcomp/ssreflect/generic_quotient.v19
4 files changed, 37 insertions, 38 deletions
diff --git a/mathcomp/ssreflect/choice.v b/mathcomp/ssreflect/choice.v
index baa7231..6f46832 100644
--- a/mathcomp/ssreflect/choice.v
+++ b/mathcomp/ssreflect/choice.v
@@ -255,19 +255,19 @@ Record mixin_of T := Mixin {
Record class_of T := Class {base : Equality.class_of T; mixin : mixin_of T}.
Local Coercion base : class_of >-> Equality.class_of.
-Structure type := Pack {sort; _ : class_of sort; _ : Type}.
+Structure type := Pack {sort; _ : class_of sort}.
Local Coercion sort : type >-> Sortclass.
Variables (T : Type) (cT : type).
-Definition class := let: Pack _ c _ as cT' := cT return class_of cT' in c.
-Definition clone c of phant_id class c := @Pack T c T.
-Let xT := let: Pack T _ _ := cT in T.
+Definition class := let: Pack _ c as cT' := cT return class_of cT' in c.
+Definition clone c of phant_id class c := @Pack T c.
+Let xT := let: Pack T _ := cT in T.
Notation xclass := (class : class_of xT).
Definition pack m :=
- fun b bT & phant_id (Equality.class bT) b => Pack (@Class T b m) T.
+ fun b bT & phant_id (Equality.class bT) b => Pack (@Class T b m).
(* Inheritance *)
-Definition eqType := @Equality.Pack cT xclass xT.
+Definition eqType := @Equality.Pack cT xclass.
End ClassDef.
@@ -408,7 +408,7 @@ Variables (P : pred T) (sT : subType P).
Definition sub_choiceMixin := PcanChoiceMixin (@valK T P sT).
Definition sub_choiceClass := @Choice.Class sT (sub_eqMixin sT) sub_choiceMixin.
-Canonical sub_choiceType := Choice.Pack sub_choiceClass sT.
+Canonical sub_choiceType := Choice.Pack sub_choiceClass.
End SubChoice.
@@ -522,19 +522,19 @@ Section ClassDef.
Record class_of T := Class { base : Choice.class_of T; mixin : mixin_of T }.
Local Coercion base : class_of >-> Choice.class_of.
-Structure type : Type := Pack {sort : Type; _ : class_of sort; _ : Type}.
+Structure type : Type := Pack {sort : Type; _ : class_of sort}.
Local Coercion sort : type >-> Sortclass.
Variables (T : Type) (cT : type).
-Definition class := let: Pack _ c _ as cT' := cT return class_of cT' in c.
-Definition clone c of phant_id class c := @Pack T c T.
-Let xT := let: Pack T _ _ := cT in T.
+Definition class := let: Pack _ c as cT' := cT return class_of cT' in c.
+Definition clone c of phant_id class c := @Pack T c.
+Let xT := let: Pack T _ := cT in T.
Notation xclass := (class : class_of xT).
Definition pack m :=
- fun bT b & phant_id (Choice.class bT) b => Pack (@Class T b m) T.
+ fun bT b & phant_id (Choice.class bT) b => Pack (@Class T b m).
-Definition eqType := @Equality.Pack cT xclass xT.
-Definition choiceType := @Choice.Pack cT xclass xT.
+Definition eqType := @Equality.Pack cT xclass.
+Definition choiceType := @Choice.Pack cT xclass.
End ClassDef.
diff --git a/mathcomp/ssreflect/eqtype.v b/mathcomp/ssreflect/eqtype.v
index 7473ed7..a5e8784 100644
--- a/mathcomp/ssreflect/eqtype.v
+++ b/mathcomp/ssreflect/eqtype.v
@@ -126,13 +126,13 @@ Notation class_of := mixin_of (only parsing).
Section ClassDef.
-Structure type := Pack {sort; _ : class_of sort; _ : Type}.
+Structure type := Pack {sort; _ : class_of sort}.
Local Coercion sort : type >-> Sortclass.
Variables (T : Type) (cT : type).
-Definition class := let: Pack _ c _ := cT return class_of cT in c.
+Definition class := let: Pack _ c := cT return class_of cT in c.
-Definition pack c := @Pack T c T.
+Definition pack c := @Pack T c.
Definition clone := fun c & cT -> T & phant_id (pack c) cT => pack c.
End ClassDef.
diff --git a/mathcomp/ssreflect/fintype.v b/mathcomp/ssreflect/fintype.v
index d4f564f..8be2eb0 100644
--- a/mathcomp/ssreflect/fintype.v
+++ b/mathcomp/ssreflect/fintype.v
@@ -172,7 +172,7 @@ Section Mixins.
Variable T : countType.
Definition EnumMixin :=
- let: Countable.Pack _ (Countable.Class _ m) _ as cT := T
+ let: Countable.Pack _ (Countable.Class _ m) as cT := T
return forall e : seq cT, axiom e -> mixin_of cT in
@Mixin (EqType _ _) m.
@@ -198,26 +198,26 @@ Section ClassDef.
Record class_of T := Class {
base : Choice.class_of T;
- mixin : mixin_of (Equality.Pack base T)
+ mixin : mixin_of (Equality.Pack base)
}.
Definition base2 T c := Countable.Class (@base T c) (mixin_base (mixin c)).
Local Coercion base : class_of >-> Choice.class_of.
-Structure type : Type := Pack {sort; _ : class_of sort; _ : Type}.
+Structure type : Type := Pack {sort; _ : class_of sort}.
Local Coercion sort : type >-> Sortclass.
Variables (T : Type) (cT : type).
-Definition class := let: Pack _ c _ as cT' := cT return class_of cT' in c.
-Definition clone c of phant_id class c := @Pack T c T.
-Let xT := let: Pack T _ _ := cT in T.
+Definition class := let: Pack _ c as cT' := cT return class_of cT' in c.
+Definition clone c of phant_id class c := @Pack T c.
+Let xT := let: Pack T _ := cT in T.
Notation xclass := (class : class_of xT).
Definition pack b0 (m0 : mixin_of (EqType T b0)) :=
fun bT b & phant_id (Choice.class bT) b =>
- fun m & phant_id m0 m => Pack (@Class T b m) T.
+ fun m & phant_id m0 m => Pack (@Class T b m).
-Definition eqType := @Equality.Pack cT xclass xT.
-Definition choiceType := @Choice.Pack cT xclass xT.
-Definition countType := @Countable.Pack cT (base2 xclass) xT.
+Definition eqType := @Equality.Pack cT xclass.
+Definition choiceType := @Choice.Pack cT xclass.
+Definition countType := @Countable.Pack cT (base2 xclass).
End ClassDef.
@@ -1379,7 +1379,7 @@ Coercion subFinType_subCountType sT := @SubCountType _ _ sT (subFin_mixin sT).
Canonical subFinType_subCountType.
Coercion subFinType_finType sT :=
- Pack (@Class sT (sub_choiceClass sT) (subFin_mixin sT)) sT.
+ Pack (@Class sT (sub_choiceClass sT) (subFin_mixin sT)).
Canonical subFinType_finType.
Lemma codom_val sT x : (x \in codom (val : sT -> T)) = P x.
diff --git a/mathcomp/ssreflect/generic_quotient.v b/mathcomp/ssreflect/generic_quotient.v
index 1475d55..ea929d7 100644
--- a/mathcomp/ssreflect/generic_quotient.v
+++ b/mathcomp/ssreflect/generic_quotient.v
@@ -128,11 +128,10 @@ Notation quot_class_of := quot_mixin_of.
Record quotType := QuotTypePack {
quot_sort :> Type;
- quot_class : quot_class_of quot_sort;
- _ : Type
+ quot_class : quot_class_of quot_sort
}.
-Definition QuotType_pack qT m := @QuotTypePack qT m qT.
+Definition QuotType_pack qT m := @QuotTypePack qT m.
Variable qT : quotType.
Definition pi_phant of phant qT := quot_pi (quot_class qT).
@@ -143,7 +142,7 @@ Lemma repr_ofK : cancel repr_of \pi.
Proof. by rewrite /pi_phant /repr_of /=; case: qT=> [? []]. Qed.
Definition QuotType_clone (Q : Type) qT cT
- of phant_id (quot_class qT) cT := @QuotTypePack Q cT Q.
+ of phant_id (quot_class qT) cT := @QuotTypePack Q cT.
End QuotientDef.
@@ -329,8 +328,8 @@ Variable eq_quot_op : rel T.
Definition eq_quot_mixin_of (Q : Type) (qc : quot_class_of T Q)
(ec : Equality.class_of Q) :=
- {mono \pi_(QuotTypePack qc Q) : x y /
- eq_quot_op x y >-> @eq_op (Equality.Pack ec Q) x y}.
+ {mono \pi_(QuotTypePack qc) : x y /
+ eq_quot_op x y >-> @eq_op (Equality.Pack ec) x y}.
Record eq_quot_class_of (Q : Type) : Type := EqQuotClass {
eq_quot_quot_class :> quot_class_of T Q;
@@ -341,13 +340,13 @@ Record eq_quot_class_of (Q : Type) : Type := EqQuotClass {
Record eqQuotType : Type := EqQuotTypePack {
eq_quot_sort :> Type;
_ : eq_quot_class_of eq_quot_sort;
- _ : Type
+
}.
Implicit Type eqT : eqQuotType.
Definition eq_quot_class eqT : eq_quot_class_of eqT :=
- let: EqQuotTypePack _ cT _ as qT' := eqT return eq_quot_class_of qT' in cT.
+ let: EqQuotTypePack _ cT as qT' := eqT return eq_quot_class_of qT' in cT.
Canonical eqQuotType_eqType eqT := EqType eqT (eq_quot_class eqT).
Canonical eqQuotType_quotType eqT := QuotType eqT (eq_quot_class eqT).
@@ -358,10 +357,10 @@ Coercion eqQuotType_quotType : eqQuotType >-> quotType.
Definition EqQuotType_pack Q :=
fun (qT : quotType T) (eT : eqType) qc ec
of phant_id (quot_class qT) qc & phant_id (Equality.class eT) ec =>
- fun m => EqQuotTypePack (@EqQuotClass Q qc ec m) Q.
+ fun m => EqQuotTypePack (@EqQuotClass Q qc ec m).
Definition EqQuotType_clone (Q : Type) eqT cT
- of phant_id (eq_quot_class eqT) cT := @EqQuotTypePack Q cT Q.
+ of phant_id (eq_quot_class eqT) cT := @EqQuotTypePack Q cT.
Lemma pi_eq_quot eqT : {mono \pi_eqT : x y / eq_quot_op x y >-> x == y}.
Proof. by case: eqT => [] ? []. Qed.