1 2 3 4 5 6 7 8 9 10 11 12 13 14 15 16 17 18 19 20 21 22 23 24 25 26 27 28 29 30 31 32 33 34 35 36 37 38 39 40 41 42 43 44 45 46 47 48 49 50 51 52 53 54 55 56 57 58 59 60 61 62 63 64 65 66 67 68 69 70 71 72 73 74 75 76 77 78 79 80 81 82 83 84 85 86 87 88 89 90 91 92 93 94 95 96 97 98 99 100 101 102 103 104 105 106 107 108 109 110 111 112 113 114 115 116 117 118 119 120 121 122 123 124 125 126 127 128 129 130 131 132 133 134 135 136 137 138 139 140 141 142 143 144 145 146
|
(************************************************************************)
(* * The Coq Proof Assistant / The Coq Development Team *)
(* v * Copyright INRIA, CNRS and contributors *)
(* <O___,, * (see version control and CREDITS file for authors & dates) *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
(* Pierre Letouzey, Jerome Vouillon, PPS, Paris 7, 2008 *)
(************************************************************************)
Local Open Scope type_scope.
Require Import List.
(** * Generic dependently-typed operators about [n]-ary functions *)
(** The type of [n]-ary function: [nfun A n B] is
[A -> ... -> A -> B] with [n] occurrences of [A] in this type. *)
Fixpoint nfun A n B :=
match n with
| O => B
| S n => A -> (nfun A n B)
end.
Notation " A ^^ n --> B " := (nfun A n B)
(at level 50, n at next level) : type_scope.
(** [napply_cst _ _ a n f] iterates [n] times the application of a
particular constant [a] to the [n]-ary function [f]. *)
Fixpoint napply_cst (A B:Type)(a:A) n : (A^^n-->B) -> B :=
match n return (A^^n-->B) -> B with
| O => fun x => x
| S n => fun x => napply_cst _ _ a n (x a)
end.
(** A generic transformation from an n-ary function to another one.*)
Fixpoint nfun_to_nfun (A B C:Type)(f:B -> C) n :
(A^^n-->B) -> (A^^n-->C) :=
match n return (A^^n-->B) -> (A^^n-->C) with
| O => f
| S n => fun g a => nfun_to_nfun _ _ _ f n (g a)
end.
(** [napply_except_last _ _ n f] expects [S n] arguments of type [A],
applies [n] of them to [f] and discards the last one. *)
Fixpoint napply_except_last (A B:Type) (n : nat) (f : A^^n-->B) {struct n} : A^^S n-->B.
Proof.
destruct n.
- exact (fun _ => f).
- exact (fun arg => napply_except_last A B n (f arg)).
Defined.
(** [napply_then_last _ _ a n f] expects [n] arguments of type [A],
applies them to [f] and then apply [a] to the result. *)
Definition napply_then_last (A B:Type)(a:A) :=
nfun_to_nfun A (A->B) B (fun fab => fab a).
(** [napply_discard _ b n] expects [n] arguments, discards then,
and returns [b]. *)
Fixpoint napply_discard (A B:Type)(b:B) n : A^^n-->B :=
match n return A^^n-->B with
| O => b
| S n => fun _ => napply_discard _ _ b n
end.
(** A fold function *)
Fixpoint nfold A B (f:A->B->B)(b:B) n : (A^^n-->B) :=
match n return (A^^n-->B) with
| O => b
| S n => fun a => (nfold _ _ f (f a b) n)
end.
(** [n]-ary products : [nprod A n] is [A*...*A*unit],
with [n] occurrences of [A] in this type. *)
Fixpoint nprod A n : Type := match n with
| O => unit
| S n => (A * nprod A n)%type
end.
Notation "A ^ n" := (nprod A n) : type_scope.
(** [n]-ary curryfication / uncurryfication *)
Fixpoint ncurry (A B:Type) n : (A^n -> B) -> (A^^n-->B) :=
match n return (A^n -> B) -> (A^^n-->B) with
| O => fun x => x tt
| S n => fun f a => ncurry _ _ n (fun p => f (a,p))
end.
Fixpoint nuncurry (A B:Type) n : (A^^n-->B) -> (A^n -> B) :=
match n return (A^^n-->B) -> (A^n -> B) with
| O => fun x _ => x
| S n => fun f p => let (x,p) := p in nuncurry _ _ n (f x) p
end.
(** Earlier functions can also be defined via [ncurry/nuncurry].
For instance : *)
Definition nfun_to_nfun_bis A B C (f:B->C) n :
(A^^n-->B) -> (A^^n-->C) :=
fun anb => ncurry _ _ n (fun an => f ((nuncurry _ _ n anb) an)).
(** We can also us it to obtain another [fold] function,
equivalent to the previous one, but with a nicer expansion
(see for instance Int31.iszero). *)
Fixpoint nfold_bis A B (f:A->B->B)(b:B) n : (A^^n-->B) :=
match n return (A^^n-->B) with
| O => b
| S n => fun a =>
nfun_to_nfun_bis _ _ _ (f a) n (nfold_bis _ _ f b n)
end.
(** From [nprod] to [list] *)
Fixpoint nprod_to_list (A:Type) n : A^n -> list A :=
match n with
| O => fun _ => nil
| S n => fun p => let (x,p) := p in x::(nprod_to_list _ n p)
end.
(** From [list] to [nprod] *)
Fixpoint nprod_of_list (A:Type)(l:list A) : A^(length l) :=
match l return A^(length l) with
| nil => tt
| x::l => (x, nprod_of_list _ l)
end.
(** This gives an additional way to write the fold *)
Definition nfold_list (A B:Type)(f:A->B->B)(b:B) n : (A^^n-->B) :=
ncurry _ _ n (fun p => fold_right f b (nprod_to_list _ _ p)).
|