Library compcert.Coqlib


This file collects a number of definitions and theorems that are used throughout the development. It complements the Coq standard library.

Require Export ZArith.
Require Export Znumtheory.
Require Export List.
Require Export Bool.

Useful tactics


Ltac inv H := inversion H; clear H; subst.

Ltac predSpec pred predspec x y :=
  generalize (predspec x y); case (pred x y); intro.

Ltac caseEq name :=
  generalize (refl_equal name); pattern name at -1 in |- *; case name.

Ltac destructEq name :=
  destruct name eqn:?.

Ltac decEq :=
  match goal with
  | [ |- _ = _ ] ⇒ f_equal
  | [ |- (?X ?A ?X ?B) ] ⇒
      cut (A B); [intro; congruence | try discriminate]
  end.

Ltac byContradiction :=
  cut False; [contradiction|idtac].

Ltac omegaContradiction :=
  cut False; [contradiction|omega].

Lemma modusponens: (P Q: Prop), P → (PQ) → Q.

Ltac exploit x :=
    refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _)
 || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _)
 || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _)
 || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _)
 || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _)
 || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _)
 || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _)
 || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _)
 || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _)
 || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _)
 || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _)
 || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _)
 || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _)
 || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _)
 || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _)
 || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _)
 || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _)
 || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _)
 || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _)
 || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _)
 || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _)
 || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _) _)
 || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _) _)
 || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _) _)
 || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _) _)
 || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _) _)
 || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _) _)
 || refine (modusponens _ _ (x _ _ _ _ _ _ _ _) _)
 || refine (modusponens _ _ (x _ _ _ _ _ _ _) _)
 || refine (modusponens _ _ (x _ _ _ _ _ _) _)
 || refine (modusponens _ _ (x _ _ _ _ _) _)
 || refine (modusponens _ _ (x _ _ _ _) _)
 || refine (modusponens _ _ (x _ _ _) _)
 || refine (modusponens _ _ (x _ _) _)
 || refine (modusponens _ _ (x _) _).

Definitions and theorems over the type positive


Definition peq: (x y: positive), {x = y} + {x y} := Pos.eq_dec.
Global Opaque peq.

Lemma peq_true:
   (A: Type) (x: positive) (a b: A), (if peq x x then a else b) = a.

Lemma peq_false:
   (A: Type) (x y: positive) (a b: A), x y(if peq x y then a else b) = b.

Definition Plt: positivepositiveProp := Pos.lt.

Lemma Plt_ne:
   (x y: positive), Plt x yx y.
Hint Resolve Plt_ne: coqlib.

Lemma Plt_trans:
   (x y z: positive), Plt x yPlt y zPlt x z.

Lemma Plt_succ:
   (x: positive), Plt x (Psucc x).
Hint Resolve Plt_succ: coqlib.

Lemma Plt_trans_succ:
   (x y: positive), Plt x yPlt x (Psucc y).
Hint Resolve Plt_succ: coqlib.

Lemma Plt_succ_inv:
   (x y: positive), Plt x (Psucc y) → Plt x y x = y.

Definition plt (x y: positive) : {Plt x y} + {¬ Plt x y}.
Global Opaque plt.

Definition Ple: positivepositiveProp := Pos.le.

Lemma Ple_refl: (p: positive), Ple p p.

Lemma Ple_trans: (p q r: positive), Ple p qPle q rPle p r.

Lemma Plt_Ple: (p q: positive), Plt p qPle p q.

Lemma Ple_succ: (p: positive), Ple p (Psucc p).

Lemma Plt_Ple_trans:
   (p q r: positive), Plt p qPle q rPlt p r.

Lemma Plt_strict: p, ¬ Plt p p.

Hint Resolve Ple_refl Plt_Ple Ple_succ Plt_strict: coqlib.

Ltac xomega := unfold Plt, Ple in *; zify; omega.
Ltac xomegaContradiction := exfalso; xomega.

Peano recursion over positive numbers.

Section POSITIVE_ITERATION.

Lemma Plt_wf: well_founded Plt.

Variable A: Type.
Variable v1: A.
Variable f: positiveAA.

Lemma Ppred_Plt:
   x, x xHPlt (Ppred x) x.

Let iter (x: positive) (P: y, Plt y xA) : A :=
  match peq x xH with
  | left EQv1
  | right NOTEQf (Ppred x) (P (Ppred x) (Ppred_Plt x NOTEQ))
  end.

Definition positive_rec : positiveA :=
  Fix Plt_wf (fun _A) iter.

Lemma unroll_positive_rec:
   x,
  positive_rec x = iter x (fun y _positive_rec y).

Lemma positive_rec_base:
  positive_rec 1%positive = v1.

Lemma positive_rec_succ:
   x, positive_rec (Psucc x) = f x (positive_rec x).

Lemma positive_Peano_ind:
   (P: positiveProp),
  P xH
  ( x, P xP (Psucc x)) →
   x, P x.

End POSITIVE_ITERATION.

Definitions and theorems over the type Z


Definition zeq: (x y: Z), {x = y} + {x y} := Z.eq_dec.

Lemma zeq_true:
   (A: Type) (x: Z) (a b: A), (if zeq x x then a else b) = a.

Lemma zeq_false:
   (A: Type) (x y: Z) (a b: A), x y(if zeq x y then a else b) = b.

Open Scope Z_scope.

Definition zlt: (x y: Z), {x < y} + {x y} := Z_lt_dec.

Lemma zlt_true:
   (A: Type) (x y: Z) (a b: A),
  x < y(if zlt x y then a else b) = a.

Lemma zlt_false:
   (A: Type) (x y: Z) (a b: A),
  x y(if zlt x y then a else b) = b.

Definition zle: (x y: Z), {x y} + {x > y} := Z_le_gt_dec.

Lemma zle_true:
   (A: Type) (x y: Z) (a b: A),
  x y(if zle x y then a else b) = a.

Lemma zle_false:
   (A: Type) (x y: Z) (a b: A),
  x > y(if zle x y then a else b) = b.

Properties of powers of two.

Lemma two_power_nat_O : two_power_nat O = 1.

Lemma two_power_nat_pos : n : nat, two_power_nat n > 0.

Lemma two_power_nat_two_p:
   x, two_power_nat x = two_p (Z_of_nat x).

Lemma two_p_monotone:
   x y, 0 x ytwo_p x two_p y.

Lemma two_p_monotone_strict:
   x y, 0 x < ytwo_p x < two_p y.

Lemma two_p_strict:
   x, x 0 → x < two_p x.

Lemma two_p_strict_2:
   x, x 0 → 2 × x - 1 < two_p x.

Properties of Zmin and Zmax

Lemma Zmin_spec:
   x y, Zmin x y = if zlt x y then x else y.

Lemma Zmax_spec:
   x y, Zmax x y = if zlt y x then x else y.

Lemma Zmax_bound_l:
   x y z, x yx Zmax y z.
Lemma Zmax_bound_r:
   x y z, x zx Zmax y z.

Properties of Euclidean division and modulus.

Lemma Zdiv_small:
   x y, 0 x < yx / y = 0.

Lemma Zmod_small:
   x y, 0 x < yx mod y = x.

Lemma Zmod_unique:
   x y a b,
  x = a × y + b → 0 b < yx mod y = b.

Lemma Zdiv_unique:
   x y a b,
  x = a × y + b → 0 b < yx / y = a.

Lemma Zdiv_Zdiv:
   a b c,
  b > 0 → c > 0 → (a / b) / c = a / (b × c).

Lemma Zmult_le_compat_l_neg :
   n m p:Z, n mp 0 → p × n p × m.

Lemma Zdiv_interval_1:
   lo hi a b,
  lo 0 → hi > 0 → b > 0 →
  lo × b a < hi × b
  lo a/b < hi.

Lemma Zdiv_interval_2:
   lo hi a b,
  lo a hilo 0 → hi 0 → b > 0 →
  lo a/b hi.

Lemma Zmod_recombine:
   x a b,
  a > 0 → b > 0 →
  x mod (a × b) = ((x/b) mod a) × b + (x mod b).

Properties of divisibility.

Lemma Zdivides_trans:
   x y z, (x | y)(y | z)(x | z).

Definition Zdivide_dec:
   (p q: Z), p > 0 → { (p|q) } + { ¬(p|q) }.
Global Opaque Zdivide_dec.

Lemma Zdivide_interval:
   a b c,
  0 < c → 0 a < b(c | a)(c | b) → 0 a b - c.

Conversion from Z to nat.

Definition nat_of_Z: Znat := Z.to_nat.

Lemma nat_of_Z_of_nat:
   n, nat_of_Z (Z_of_nat n) = n.

Lemma nat_of_Z_max:
   z, Z_of_nat (nat_of_Z z) = Zmax z 0.

Lemma nat_of_Z_eq:
   z, z 0 → Z_of_nat (nat_of_Z z) = z.

Lemma nat_of_Z_neg:
   n, n 0 → nat_of_Z n = O.

Lemma nat_of_Z_plus:
   p q,
  p 0 → q 0 →
  nat_of_Z (p + q) = (nat_of_Z p + nat_of_Z q)%nat.

Alignment: align n amount returns the smallest multiple of amount greater than or equal to n.

Definition align (n: Z) (amount: Z) :=
  ((n + amount - 1) / amount) × amount.

Lemma align_le: x y, y > 0 → x align x y.

Lemma align_divides: x y, y > 0 → (y | align x y).

Definitions and theorems on the data types option, sum and list


Set Implicit Arguments.

Comparing option types.

Definition option_eq (A: Type) (eqA: (x y: A), {x=y} + {xy}):
   (x y: option A), {x=y} + {xy}.
Global Opaque option_eq.

Mapping a function over an option type.

Definition option_map (A B: Type) (f: AB) (x: option A) : option B :=
  match x with
  | NoneNone
  | Some ySome (f y)
  end.

Mapping a function over a sum type.

Definition sum_left_map (A B C: Type) (f: AB) (x: A + C) : B + C :=
  match x with
  | inl yinl C (f y)
  | inr zinr B z
  end.

Properties of List.nth (n-th element of a list).

Hint Resolve in_eq in_cons: coqlib.

Lemma nth_error_in:
   (A: Type) (n: nat) (l: list A) (x: A),
  List.nth_error l n = Some xIn x l.
Hint Resolve nth_error_in: coqlib.

Lemma nth_error_nil:
   (A: Type) (idx: nat), nth_error (@nil A) idx = None.
Hint Resolve nth_error_nil: coqlib.

Compute the length of a list, with result in Z.

Fixpoint list_length_z_aux (A: Type) (l: list A) (acc: Z) {struct l}: Z :=
  match l with
  | nilacc
  | hd :: tllist_length_z_aux tl (Zsucc acc)
  end.

Remark list_length_z_aux_shift:
   (A: Type) (l: list A) n m,
  list_length_z_aux l n = list_length_z_aux l m + (n - m).

Definition list_length_z (A: Type) (l: list A) : Z :=
  list_length_z_aux l 0.

Lemma list_length_z_cons:
   (A: Type) (hd: A) (tl: list A),
  list_length_z (hd :: tl) = list_length_z tl + 1.

Lemma list_length_z_pos:
   (A: Type) (l: list A),
  list_length_z l 0.

Lemma list_length_z_map:
   (A B: Type) (f: AB) (l: list A),
  list_length_z (map f l) = list_length_z l.

Extract the n-th element of a list, as List.nth_error does, but the index n is of type Z.

Fixpoint list_nth_z (A: Type) (l: list A) (n: Z) {struct l}: option A :=
  match l with
  | nilNone
  | hd :: tlif zeq n 0 then Some hd else list_nth_z tl (Zpred n)
  end.

Lemma list_nth_z_in:
   (A: Type) (l: list A) n x,
  list_nth_z l n = Some xIn x l.

Lemma list_nth_z_map:
   (A B: Type) (f: AB) (l: list A) n,
  list_nth_z (List.map f l) n = option_map f (list_nth_z l n).

Lemma list_nth_z_range:
   (A: Type) (l: list A) n x,
  list_nth_z l n = Some x → 0 n < list_length_z l.

Properties of List.incl (list inclusion).

Lemma incl_cons_inv:
   (A: Type) (a: A) (b c: list A),
  incl (a :: b) cincl b c.
Hint Resolve incl_cons_inv: coqlib.

Lemma incl_app_inv_l:
   (A: Type) (l1 l2 m: list A),
  incl (l1 ++ l2) mincl l1 m.

Lemma incl_app_inv_r:
   (A: Type) (l1 l2 m: list A),
  incl (l1 ++ l2) mincl l2 m.

Hint Resolve incl_tl incl_refl incl_app_inv_l incl_app_inv_r: coqlib.

Lemma incl_same_head:
   (A: Type) (x: A) (l1 l2: list A),
  incl l1 l2incl (x::l1) (x::l2).

Properties of List.map (mapping a function over a list).

Lemma list_map_exten:
   (A B: Type) (f f': AB) (l: list A),
  ( x, In x lf x = f' x) →
  List.map f' l = List.map f l.

Lemma list_map_compose:
   (A B C: Type) (f: AB) (g: BC) (l: list A),
  List.map g (List.map f l) = List.map (fun xg(f x)) l.

Lemma list_map_identity:
   (A: Type) (l: list A),
  List.map (fun (x:A) ⇒ x) l = l.

Lemma list_map_nth:
   (A B: Type) (f: AB) (l: list A) (n: nat),
  nth_error (List.map f l) n = option_map f (nth_error l n).

Lemma list_length_map:
   (A B: Type) (f: AB) (l: list A),
  List.length (List.map f l) = List.length l.

Lemma list_in_map_inv:
   (A B: Type) (f: AB) (l: list A) (y: B),
  In y (List.map f l) → x:A, y = f x In x l.

Lemma list_append_map:
   (A B: Type) (f: AB) (l1 l2: list A),
  List.map f (l1 ++ l2) = List.map f l1 ++ List.map f l2.

Lemma list_append_map_inv:
   (A B: Type) (f: AB) (m1 m2: list B) (l: list A),
  List.map f l = m1 ++ m2
   l1, l2, List.map f l1 = m1 List.map f l2 = m2 l = l1 ++ l2.

Folding a function over a list

Section LIST_FOLD.

Variables A B: Type.
Variable f: ABB.

This is exactly List.fold_left from Coq's standard library, with f taking arguments in a different order.

Fixpoint list_fold_left (accu: B) (l: list A) : B :=
  match l with nilaccu | x :: l'list_fold_left (f x accu) l' end.

This is exactly List.fold_right from Coq's standard library, except that it runs in constant stack space.

Definition list_fold_right (l: list A) (base: B) : B :=
  list_fold_left base (List.rev' l).

Remark list_fold_left_app:
   l1 l2 accu,
  list_fold_left accu (l1 ++ l2) = list_fold_left (list_fold_left accu l1) l2.

Lemma list_fold_right_eq:
   l base,
  list_fold_right l base =
  match l with nilbase | x :: l'f x (list_fold_right l' base) end.

Lemma list_fold_right_spec:
   l base, list_fold_right l base = List.fold_right f base l.

End LIST_FOLD.

Properties of list membership.

Lemma in_cns:
   (A: Type) (x y: A) (l: list A), In x (y :: l) y = x In x l.

Lemma in_app:
   (A: Type) (x: A) (l1 l2: list A), In x (l1 ++ l2) In x l1 In x l2.

Lemma list_in_insert:
   (A: Type) (x: A) (l1 l2: list A) (y: A),
  In x (l1 ++ l2) → In x (l1 ++ y :: l2).

list_disjoint l1 l2 holds iff l1 and l2 have no elements in common.

Definition list_disjoint (A: Type) (l1 l2: list A) : Prop :=
   (x y: A), In x l1In y l2x y.

Lemma list_disjoint_cons_l:
   (A: Type) (a: A) (l1 l2: list A),
  list_disjoint l1 l2¬In a l2list_disjoint (a :: l1) l2.

Lemma list_disjoint_cons_r:
   (A: Type) (a: A) (l1 l2: list A),
  list_disjoint l1 l2¬In a l1list_disjoint l1 (a :: l2).

Lemma list_disjoint_cons_left:
   (A: Type) (a: A) (l1 l2: list A),
  list_disjoint (a :: l1) l2list_disjoint l1 l2.

Lemma list_disjoint_cons_right:
   (A: Type) (a: A) (l1 l2: list A),
  list_disjoint l1 (a :: l2) → list_disjoint l1 l2.

Lemma list_disjoint_notin:
   (A: Type) (l1 l2: list A) (a: A),
  list_disjoint l1 l2In a l1~(In a l2).

Lemma list_disjoint_sym:
   (A: Type) (l1 l2: list A),
  list_disjoint l1 l2list_disjoint l2 l1.

Lemma list_disjoint_dec:
   (A: Type) (eqA_dec: (x y: A), {x=y} + {xy}) (l1 l2: list A),
  {list_disjoint l1 l2} + {¬list_disjoint l1 l2}.

list_equiv l1 l2 holds iff the lists l1 and l2 contain the same elements.

Definition list_equiv (A : Type) (l1 l2: list A) : Prop :=
   x, In x l1 In x l2.

list_norepet l holds iff the list l contains no repetitions, i.e. no element occurs twice.

Inductive list_norepet (A: Type) : list AProp :=
  | list_norepet_nil:
      list_norepet nil
  | list_norepet_cons:
       hd tl,
      ~(In hd tl)list_norepet tllist_norepet (hd :: tl).

Lemma list_norepet_dec:
   (A: Type) (eqA_dec: (x y: A), {x=y} + {xy}) (l: list A),
  {list_norepet l} + {¬list_norepet l}.

Lemma list_map_norepet:
   (A B: Type) (f: AB) (l: list A),
  list_norepet l
  ( x y, In x lIn y lx yf x f y) →
  list_norepet (List.map f l).

Remark list_norepet_append_commut:
   (A: Type) (a b: list A),
  list_norepet (a ++ b) → list_norepet (b ++ a).

Lemma list_norepet_app:
   (A: Type) (l1 l2: list A),
  list_norepet (l1 ++ l2)
  list_norepet l1 list_norepet l2 list_disjoint l1 l2.

Lemma list_norepet_append:
   (A: Type) (l1 l2: list A),
  list_norepet l1list_norepet l2list_disjoint l1 l2
  list_norepet (l1 ++ l2).

Lemma list_norepet_append_right:
   (A: Type) (l1 l2: list A),
  list_norepet (l1 ++ l2) → list_norepet l2.

Lemma list_norepet_append_left:
   (A: Type) (l1 l2: list A),
  list_norepet (l1 ++ l2) → list_norepet l1.

is_tail l1 l2 holds iff l2 is of the form l ++ l1 for some l.

Inductive is_tail (A: Type): list Alist AProp :=
  | is_tail_refl:
       c, is_tail c c
  | is_tail_cons:
       i c1 c2, is_tail c1 c2is_tail c1 (i :: c2).

Lemma is_tail_in:
   (A: Type) (i: A) c1 c2, is_tail (i :: c1) c2In i c2.

Lemma is_tail_cons_left:
   (A: Type) (i: A) c1 c2, is_tail (i :: c1) c2is_tail c1 c2.

Hint Resolve is_tail_refl is_tail_cons is_tail_in is_tail_cons_left: coqlib.

Lemma is_tail_incl:
   (A: Type) (l1 l2: list A), is_tail l1 l2incl l1 l2.

Lemma is_tail_trans:
   (A: Type) (l1 l2: list A),
  is_tail l1 l2 (l3: list A), is_tail l2 l3is_tail l1 l3.

list_forall2 P [x1 ... xN] [y1 ... yM] holds iff N = M and P xi yi holds for all i.

Section FORALL2.

Variable A: Type.
Variable B: Type.
Variable P: ABProp.

Inductive list_forall2: list Alist BProp :=
  | list_forall2_nil:
      list_forall2 nil nil
  | list_forall2_cons:
       a1 al b1 bl,
      P a1 b1
      list_forall2 al bl
      list_forall2 (a1 :: al) (b1 :: bl).

Lemma list_forall2_app:
   a2 b2 a1 b1,
  list_forall2 a1 b1list_forall2 a2 b2
  list_forall2 (a1 ++ a2) (b1 ++ b2).

Lemma list_forall2_length:
   l1 l2,
  list_forall2 l1 l2length l1 = length l2.

End FORALL2.

Lemma list_forall2_imply:
   (A B: Type) (P1: ABProp) (l1: list A) (l2: list B),
  list_forall2 P1 l1 l2
   (P2: ABProp),
  ( v1 v2, In v1 l1In v2 l2P1 v1 v2P2 v1 v2) →
  list_forall2 P2 l1 l2.

Dropping the first N elements of a list.

Fixpoint list_drop (A: Type) (n: nat) (x: list A) {struct n} : list A :=
  match n with
  | Ox
  | S n'match x with nilnil | hd :: tllist_drop n' tl end
  end.

Lemma list_drop_incl:
   (A: Type) (x: A) n (l: list A), In x (list_drop n l) → In x l.

Lemma list_drop_norepet:
   (A: Type) n (l: list A), list_norepet llist_norepet (list_drop n l).

Lemma list_map_drop:
   (A B: Type) (f: AB) n (l: list A),
  list_drop n (map f l) = map f (list_drop n l).

A list of n elements, all equal to x.

Fixpoint list_repeat {A: Type} (n: nat) (x: A) {struct n} :=
  match n with
  | Onil
  | S mx :: list_repeat m x
  end.

Lemma length_list_repeat:
   (A: Type) n (x: A), length (list_repeat n x) = n.

Lemma in_list_repeat:
   (A: Type) n (x: A) y, In y (list_repeat n x) → y = x.

Definitions and theorems over boolean types


Definition proj_sumbool (P Q: Prop) (a: {P} + {Q}) : bool :=
  if a then true else false.

Implicit Arguments proj_sumbool [P Q].

Coercion proj_sumbool: sumbool >-> bool.

Lemma proj_sumbool_true:
   (P Q: Prop) (a: {P}+{Q}), proj_sumbool a = trueP.

Lemma proj_sumbool_is_true:
   (P: Prop) (a: {P}+{¬P}), Pproj_sumbool a = true.

Ltac InvBooleans :=
  match goal with
  | [ H: _ && _ = true |- _ ] ⇒
      destruct (andb_prop _ _ H); clear H; InvBooleans
  | [ H: _ || _ = false |- _ ] ⇒
      destruct (orb_false_elim _ _ H); clear H; InvBooleans
  | [ H: proj_sumbool ?x = true |- _ ] ⇒
      generalize (proj_sumbool_true _ H); clear H; intro; InvBooleans
  | _idtac
  end.

Section DECIDABLE_EQUALITY.

Variable A: Type.
Variable dec_eq: (x y: A), {x=y} + {xy}.
Variable B: Type.

Lemma dec_eq_true:
   (x: A) (ifso ifnot: B),
  (if dec_eq x x then ifso else ifnot) = ifso.

Lemma dec_eq_false:
   (x y: A) (ifso ifnot: B),
  x y(if dec_eq x y then ifso else ifnot) = ifnot.

Lemma dec_eq_sym:
   (x y: A) (ifso ifnot: B),
  (if dec_eq x y then ifso else ifnot) =
  (if dec_eq y x then ifso else ifnot).

End DECIDABLE_EQUALITY.

Section DECIDABLE_PREDICATE.

Variable P: Prop.
Variable dec: {P} + {¬P}.
Variable A: Type.

Lemma pred_dec_true:
   (a b: A), P(if dec then a else b) = a.

Lemma pred_dec_false:
   (a b: A), ¬P(if dec then a else b) = b.

End DECIDABLE_PREDICATE.

Well-founded orderings


Require Import Relations.

A non-dependent version of lexicographic ordering.

Section LEX_ORDER.

Variable A: Type.
Variable B: Type.
Variable ordA: AAProp.
Variable ordB: BBProp.

Inductive lex_ord: A×BA×BProp :=
  | lex_ord_left: a1 b1 a2 b2,
      ordA a1 a2lex_ord (a1,b1) (a2,b2)
  | lex_ord_right: a b1 b2,
      ordB b1 b2lex_ord (a,b1) (a,b2).

Lemma wf_lex_ord:
  well_founded ordAwell_founded ordBwell_founded lex_ord.

Lemma transitive_lex_ord:
  transitive _ ordAtransitive _ ordBtransitive _ lex_ord.

End LEX_ORDER.