aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
authorGravatar Pierre Boutillier <pierre.boutillier@ens-lyon.org>2014-02-20 17:39:15 +0100
committerGravatar Pierre Boutillier <pierre.boutillier@ens-lyon.org>2014-02-24 14:07:07 +0100
commit26c390aae1275f02892412f121360668ad98a660 (patch)
tree4fd13daac049adc7b618d115766fb025f9637ec0
parent80410d825befa1890c872596cf77378a437cee73 (diff)
fixup complement Fin
-rw-r--r--doc/stdlib/index-list.html.template2
-rw-r--r--theories/Logic/FinFun.v2
-rw-r--r--theories/Vectors/Fin.v15
3 files changed, 9 insertions, 10 deletions
diff --git a/doc/stdlib/index-list.html.template b/doc/stdlib/index-list.html.template
index 9382a0890..1f8ff94e1 100644
--- a/doc/stdlib/index-list.html.template
+++ b/doc/stdlib/index-list.html.template
@@ -70,6 +70,7 @@ through the <tt>Require Import</tt> command.</p>
theories/Logic/FunctionalExtensionality.v
theories/Logic/ExtensionalityFacts.v
theories/Logic/Fan.v
+ theories/Logic/FinFun.v
</dd>
<dt> <b>Structures</b>:
@@ -410,6 +411,7 @@ through the <tt>Require Import</tt> command.</p>
</dt>
<dd>
theories/Lists/List.v
+ theories/Lists/ListDec.v
theories/Lists/ListSet.v
theories/Lists/SetoidList.v
theories/Lists/SetoidPermutation.v
diff --git a/theories/Logic/FinFun.v b/theories/Logic/FinFun.v
index f070a6889..2f72f16de 100644
--- a/theories/Logic/FinFun.v
+++ b/theories/Logic/FinFun.v
@@ -245,7 +245,7 @@ Notation n2f := Fin.of_nat_lt.
Definition f2n {n} (x:Fin.t n) := proj1_sig (Fin.to_nat x).
Definition f2n_ok n (x:Fin.t n) : f2n x < n := proj2_sig (Fin.to_nat x).
Definition n2f_f2n : forall n x, n2f (f2n_ok x) = x := @Fin.of_nat_to_nat_inv.
-Definition f2n_n2f : forall x n h, f2n (n2f h) = x := @Fin.to_nat_of_nat.
+Definition f2n_n2f x n h : f2n (n2f h) = x := f_equal (@proj1_sig _ _) (@Fin.to_nat_of_nat x n h).
Definition n2f_ext : forall x n h h', n2f h = n2f h' := @Fin.of_nat_ext.
Definition f2n_inj : forall n x y, f2n x = f2n y -> x = y := @Fin.to_nat_inj.
diff --git a/theories/Vectors/Fin.v b/theories/Vectors/Fin.v
index a1e6b14d6..f57726bea 100644
--- a/theories/Vectors/Fin.v
+++ b/theories/Vectors/Fin.v
@@ -10,8 +10,9 @@ Require Arith_base.
(** [fin n] is a convenient way to represent \[1 .. n\]
-[fin n] can be seen as a n-uplet of unit where [F1] is the first element of
-the n-uplet and [FS] set (n-1)-uplet of all the elements but the first.
+[fin n] can be seen as a n-uplet of unit. [F1] is the first element of
+the n-uplet. If [f] is the k-th element of the (n-1)-uplet, [FS f] is the
+(k+1)-th element of the n-uplet.
Author: Pierre Boutillier
Institution: PPS, INRIA 12/2010-01/2012-07/2012
@@ -115,15 +116,11 @@ induction p; simpl.
- destruct (to_nat p); simpl in *. f_equal. subst p. apply of_nat_ext.
Qed.
-Lemma to_nat_of_nat {p}{n} (h : p < n) : proj1_sig (to_nat (of_nat_lt h)) = p.
+Lemma to_nat_of_nat {p}{n} (h : p < n) : to_nat (of_nat_lt h) = exist _ p h.
Proof.
revert n h.
- induction p; intros; destruct n; simpl; trivial.
- - inversion h.
- - inversion h.
- - set (h' := Lt.lt_S_n p n h). clearbody h'.
- specialize (IHp n h').
- destruct (to_nat (of_nat_lt h')); simpl in *. now f_equal.
+ induction p; (destruct n ; intros h; [ destruct (Lt.lt_n_O _ h) | cbn]);
+ [ | rewrite (IHp _ (Lt.lt_S_n p n h))]; f_equal; apply Peano_dec.le_unique.
Qed.
Lemma to_nat_inj {n} (p q : t n) :