aboutsummaryrefslogtreecommitdiff
diff options
context:
space:
mode:
-rw-r--r--_CoqProject3
-rw-r--r--src/Util/Equality.v25
-rw-r--r--src/Util/FixCoqMistakes.v1
-rw-r--r--src/Util/GlobalSettings.v24
-rw-r--r--src/Util/Notations.v3
-rw-r--r--src/Util/Sigma.v116
6 files changed, 171 insertions, 1 deletions
diff --git a/_CoqProject b/_CoqProject
index f4a6f121b..ea1b1b4bf 100644
--- a/_CoqProject
+++ b/_CoqProject
@@ -64,13 +64,16 @@ src/Tactics/Algebra_syntax/Nsatz.v
src/Util/AdditionChainExponentiation.v
src/Util/CaseUtil.v
src/Util/Decidable.v
+src/Util/Equality.v
src/Util/FixCoqMistakes.v
+src/Util/GlobalSettings.v
src/Util/IterAssocOp.v
src/Util/ListUtil.v
src/Util/NatUtil.v
src/Util/Notations.v
src/Util/NumTheoryUtil.v
src/Util/Option.v
+src/Util/Sigma.v
src/Util/Sum.v
src/Util/Tactics.v
src/Util/Tuple.v
diff --git a/src/Util/Equality.v b/src/Util/Equality.v
new file mode 100644
index 000000000..1ff76a009
--- /dev/null
+++ b/src/Util/Equality.v
@@ -0,0 +1,25 @@
+(** * Lemmas about [eq] *)
+
+Definition concat_1p {A x y} (p : x = y :> A) : eq_trans eq_refl p = p.
+Proof. case p; reflexivity. Defined.
+Definition concat_p1 {A x y} (p : x = y :> A) : eq_trans p eq_refl = p.
+Proof. case p; reflexivity. Defined.
+Definition concat_pV {A x y} (p : x = y :> A) : eq_trans p (eq_sym p) = eq_refl.
+Proof. case p; reflexivity. Defined.
+Definition concat_Vp {A x y} (p : x = y :> A) : eq_trans (eq_sym p) p = eq_refl.
+Proof. case p; reflexivity. Defined.
+Definition transport_pp {A} {P : A -> Type} {x y z} (p : x = y) (q : y = z) (k : P x)
+: eq_rect _ P k _ (eq_trans p q) = eq_rect _ P (eq_rect _ P k _ p) _ q.
+Proof. case q; simpl; reflexivity. Defined.
+Lemma transport_const {A P x y} (p : x = y :> A) k
+: eq_rect _ (fun _ : A => P) k _ p = k.
+Proof. case p; reflexivity. Defined.
+Lemma ap_const {A B x y} (b : B) (p : x = y :> A)
+: f_equal (fun _ => b) p = eq_refl.
+Proof. case p; reflexivity. Defined.
+Lemma inv_pp {A x y z} (p : x = y :> A) (q : y = z :> A)
+: eq_sym (eq_trans p q) = eq_trans (eq_sym q) (eq_sym p).
+Proof. case q; case p; reflexivity. Defined.
+Lemma inv_V {A x y} (p : x = y :> A)
+: eq_sym (eq_sym p) = p.
+Proof. case p; reflexivity. Defined.
diff --git a/src/Util/FixCoqMistakes.v b/src/Util/FixCoqMistakes.v
index abacfa580..129344f83 100644
--- a/src/Util/FixCoqMistakes.v
+++ b/src/Util/FixCoqMistakes.v
@@ -1,4 +1,5 @@
(** * Fixes *)
+Require Export Crypto.Util.GlobalSettings.
(** Coq is poorly designed in some ways. We fix some of these issues
in this file. *)
diff --git a/src/Util/GlobalSettings.v b/src/Util/GlobalSettings.v
new file mode 100644
index 000000000..001799aa3
--- /dev/null
+++ b/src/Util/GlobalSettings.v
@@ -0,0 +1,24 @@
+(** * Global Settings across the project *)
+
+(** Compatibility with 8.4 so we can write, e.g., [match p with
+ ex_intro x y => _ end], rather than [match p with ex_intro _ x y
+ => _ end]. *)
+Global Set Asymmetric Patterns.
+
+(** Consider also: *)
+(** Judgmental η for records, faster projections *)
+(** Set Primitive Projections. *)
+
+(** Don't use non-imported hints *)
+(** Set Loose Hint Behavior "Strict". *)
+
+(** Universes *)
+(** Set Universe Polymorphism. *)
+(** Set Strict Universe Declaration. *)
+(** Unset Universe Minimization ToSet. *)
+
+(** Better control of unfolding in [rewrite] and [setoid_rewrite] *)
+(** Set Keyed Unification. *)
+
+(** Better-behaved [simpl] *)
+(** Set SimplIsCbn. *)
diff --git a/src/Util/Notations.v b/src/Util/Notations.v
index c43a58dc3..8595b0189 100644
--- a/src/Util/Notations.v
+++ b/src/Util/Notations.v
@@ -1,5 +1,6 @@
(** * Reserved Notations *)
Require Export Crypto.Util.FixCoqMistakes.
+Require Export Crypto.Util.GlobalSettings.
(** Putting them all together in one file prevents conflicts. Coq's
parser (camlpX) is really bad at conflicting notation levels and
@@ -20,7 +21,7 @@ Reserved Notation "'canonical' 'encoding' 'of' T 'as' B" (at level 50).
Reserved Notation "@ 'is_eq_dec' T R" (at level 10, T at level 8, R at level 8).
Reserved Infix "<<" (at level 30, no associativity).
Reserved Infix ">>" (at level 30, no associativity).
-Reserved Infix "&" (at level 50).
+Reserved Infix "&" (at level 50). (* N.B. This conflicts with [{ a : T & P}] for [sigT] *)
Reserved Infix "∣" (at level 50).
Reserved Infix "~=" (at level 70).
Reserved Infix "==" (at level 70, no associativity).
diff --git a/src/Util/Sigma.v b/src/Util/Sigma.v
new file mode 100644
index 000000000..9b1ff2c96
--- /dev/null
+++ b/src/Util/Sigma.v
@@ -0,0 +1,116 @@
+Require Import Crypto.Util.Equality.
+Require Import Crypto.Util.GlobalSettings.
+
+Section sigT.
+ Definition pr1_path {A} {P : A -> Type} {u v : sigT P} (p : u = v)
+ : projT1 u = projT1 v
+ := f_equal (@projT1 _ _) p.
+
+ Definition pr2_path {A} {P : A -> Type} {u v : sigT P} (p : u = v)
+ : eq_rect _ _ (projT2 u) _ (pr1_path p) = projT2 v.
+ Proof.
+ destruct p; reflexivity.
+ Defined.
+
+ Definition path_sigT_uncurried {A : Type} {P : A -> Type} (u v : sigT P)
+ (pq : sigT (fun p : projT1 u = projT1 v => eq_rect _ _ (projT2 u) _ p = projT2 v))
+ : u = v.
+ Proof.
+ destruct u as [u1 u2], v as [v1 v2]; simpl in *.
+ destruct pq as [p q].
+ destruct q; simpl in *.
+ destruct p; reflexivity.
+ Defined.
+
+ Definition path_sigT {A : Type} {P : A -> Type} (u v : sigT P)
+ (p : projT1 u = projT1 v) (q : eq_rect _ _ (projT2 u) _ p = projT2 v)
+ : u = v
+ := path_sigT_uncurried u v (existT _ p q).
+
+ Definition path_sigT_nondep {A B : Type} (u v : @sigT A (fun _ => B))
+ (p : projT1 u = projT1 v) (q : projT2 u = projT2 v)
+ : u = v
+ := @path_sigT _ _ u v p (eq_trans (transport_const _ _) q).
+
+ Lemma eq_rect_sigT {A x} {P : A -> Type} (Q : forall a, P a -> Prop) (u : sigT (Q x)) {y} (H : x = y)
+ : eq_rect x (fun a => sigT (Q a)) u y H
+ = existT
+ (Q y)
+ (eq_rect x P (projT1 u) y H)
+ match H in (_ = y) return Q y (eq_rect x P (projT1 u) y H) with
+ | eq_refl => projT2 u
+ end.
+ Proof.
+ destruct H, u; reflexivity.
+ Defined.
+End sigT.
+
+Section sig.
+ Definition proj1_sig_path {A} {P : A -> Prop} {u v : sig P} (p : u = v)
+ : proj1_sig u = proj1_sig v
+ := f_equal (@proj1_sig _ _) p.
+
+ Definition proj2_sig_path {A} {P : A -> Prop} {u v : sig P} (p : u = v)
+ : eq_rect _ _ (proj2_sig u) _ (proj1_sig_path p) = proj2_sig v.
+ Proof.
+ destruct p; reflexivity.
+ Defined.
+
+ Definition path_sig_uncurried {A : Type} {P : A -> Prop} (u v : sig P)
+ (pq : {p : proj1_sig u = proj1_sig v | eq_rect _ _ (proj2_sig u) _ p = proj2_sig v})
+ : u = v.
+ Proof.
+ destruct u as [u1 u2], v as [v1 v2]; simpl in *.
+ destruct pq as [p q].
+ destruct q; simpl in *.
+ destruct p; reflexivity.
+ Defined.
+
+ Definition path_sig {A : Type} (P : A -> Prop) (u v : sig P)
+ (p : proj1_sig u = proj1_sig v) (q : eq_rect _ _ (proj2_sig u) _ p = proj2_sig v)
+ : u = v
+ := path_sig_uncurried u v (exist _ p q).
+
+ Lemma eq_rect_sig {A x} {P : A -> Type} (Q : forall a, P a -> Prop) (u : sig (Q x)) {y} (H : x = y)
+ : eq_rect x (fun a => sig (Q a)) u y H
+ = exist
+ (Q y)
+ (eq_rect x P (proj1_sig u) y H)
+ match H in (_ = y) return Q y (eq_rect x P (proj1_sig u) y H) with
+ | eq_refl => proj2_sig u
+ end.
+ Proof.
+ destruct H, u; reflexivity.
+ Defined.
+End sig.
+
+Section ex.
+ Definition path_ex_uncurried' {A : Type} (P : A -> Prop) {u1 v1 : A} {u2 : P u1} {v2 : P v1}
+ (pq : exists p : u1 = v1, eq_rect _ _ u2 _ p = v2)
+ : ex_intro P u1 u2 = ex_intro P v1 v2.
+ Proof.
+ destruct pq as [p q].
+ destruct q; simpl in *.
+ destruct p; reflexivity.
+ Defined.
+
+ Definition path_ex' {A : Type} (P : A -> Prop) (u1 v1 : A) (u2 : P u1) (v2 : P v1)
+ (p : u1 = v1) (q : eq_rect _ _ u2 _ p = v2)
+ : ex_intro P u1 u2 = ex_intro P v1 v2
+ := path_ex_uncurried' P (ex_intro _ p q).
+
+ Lemma eq_rect_ex {A x} {P : A -> Type} (Q : forall a, P a -> Prop) (u : ex (Q x)) {y} (H : x = y)
+ : eq_rect x (fun a => ex (Q a)) u y H
+ = match u with
+ | ex_intro u1 u2
+ => ex_intro
+ (Q y)
+ (eq_rect x P u1 y H)
+ match H in (_ = y) return Q y (eq_rect x P u1 y H) with
+ | eq_refl => u2
+ end
+ end.
+ Proof.
+ destruct H, u; reflexivity.
+ Defined.
+End ex.