aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
-rw-r--r--Makefile3
-rw-r--r--checker/term.ml4
-rw-r--r--dev/doc/cic.dtd2
-rw-r--r--doc/RecTutorial/coqartmacros.tex2
-rw-r--r--doc/refman/RefMan-tac.tex6
-rw-r--r--doc/refman/RefMan-tus.tex18
-rw-r--r--engine/evarutil.ml2
-rw-r--r--ide/texmacspp.ml2
-rw-r--r--interp/topconstr.ml7
-rw-r--r--intf/constrexpr.mli2
-rw-r--r--intf/vernacexpr.mli2
-rw-r--r--kernel/constr.ml11
-rw-r--r--kernel/constr.mli2
-rw-r--r--kernel/nativeconv.ml2
-rw-r--r--kernel/term.ml2
-rw-r--r--kernel/term.mli2
-rw-r--r--kernel/vars.ml2
-rw-r--r--kernel/vars.mli4
-rw-r--r--parsing/g_vernac.ml44
-rw-r--r--printing/ppvernac.ml3
-rw-r--r--tactics/dnet.mli2
-rw-r--r--tactics/tactics.ml3
-rw-r--r--test-suite/success/ImplicitArguments.v6
-rw-r--r--theories/Reals/Rfunctions.v98
-rw-r--r--theories/Reals/Rpower.v14
-rw-r--r--vernac/auto_ind_decl.ml4
-rw-r--r--vernac/command.ml24
-rw-r--r--vernac/command.mli8
-rw-r--r--vernac/lemmas.ml29
-rw-r--r--vernac/lemmas.mli4
-rw-r--r--vernac/obligations.ml6
-rw-r--r--vernac/vernacentries.ml4
32 files changed, 202 insertions, 82 deletions
diff --git a/Makefile b/Makefile
index e1d6e8e1d..826ed17b0 100644
--- a/Makefile
+++ b/Makefile
@@ -52,7 +52,8 @@ FIND_VCS_CLAUSE:='(' \
-name '.bzr' -o \
-name 'debian' -o \
-name "$${GIT_DIR}" -o \
- -name '_build' \
+ -name '_build' -o \
+ -name '_build_ci' \
')' -prune -o
define find
diff --git a/checker/term.ml b/checker/term.ml
index 591348cb6..24e6008d3 100644
--- a/checker/term.ml
+++ b/checker/term.ml
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* This module instantiates the structure of generic deBruijn terms to Coq *)
+(* This module instantiates the structure of generic de Bruijn terms to Coq *)
open CErrors
open Util
@@ -94,7 +94,7 @@ let closedn n c =
in
try closed_rec n c; true with LocalOccur -> false
-(* [closed0 M] is true iff [M] is a (deBruijn) closed term *)
+(* [closed0 M] is true iff [M] is a (de Bruijn) closed term *)
let closed0 = closedn 0
diff --git a/dev/doc/cic.dtd b/dev/doc/cic.dtd
index f2314e224..cc33efd48 100644
--- a/dev/doc/cic.dtd
+++ b/dev/doc/cic.dtd
@@ -125,7 +125,7 @@
id ID #REQUIRED
sort %sort; #REQUIRED>
-<!-- The substitutions are ordered by increasing DeBrujin -->
+<!-- The substitutions are ordered by increasing de Bruijn -->
<!-- index. An empty substitution means that that index is -->
<!-- not accessible. -->
<!ELEMENT META (substitution*)>
diff --git a/doc/RecTutorial/coqartmacros.tex b/doc/RecTutorial/coqartmacros.tex
index 2a2c21196..72d749269 100644
--- a/doc/RecTutorial/coqartmacros.tex
+++ b/doc/RecTutorial/coqartmacros.tex
@@ -149,7 +149,7 @@
\newcommand{\PicAbst}[3]{\begin{bundle}{\bf abst}\chunk{#1}\chunk{#2}\chunk{#3}%
\end{bundle}}
-% the same in DeBruijn form
+% the same in de Bruijn form
\newcommand{\PicDbj}[2]{\begin{bundle}{\bf abst}\chunk{#1}\chunk{#2}
\end{bundle}}
diff --git a/doc/refman/RefMan-tac.tex b/doc/refman/RefMan-tac.tex
index 3f1241186..0edc66f83 100644
--- a/doc/refman/RefMan-tac.tex
+++ b/doc/refman/RefMan-tac.tex
@@ -2618,9 +2618,9 @@ as the ones described in Section~\ref{Tac-induction}.
In the syntax of the tactic, the identifier {\ident} is the name given
to the induction hypothesis. The natural number {\num} tells on which
premise of the current goal the induction acts, starting
-from 1 and counting both dependent and non dependent
-products. Especially, the current lemma must be composed of at least
-{\num} products.
+from 1, counting both dependent and non dependent
+products, but skipping local definitions. Especially, the current
+lemma must be composed of at least {\num} products.
Like in a {\tt fix} expression, the induction
hypotheses have to be used on structurally smaller arguments.
diff --git a/doc/refman/RefMan-tus.tex b/doc/refman/RefMan-tus.tex
index 797b0aded..017de6d48 100644
--- a/doc/refman/RefMan-tus.tex
+++ b/doc/refman/RefMan-tus.tex
@@ -288,8 +288,8 @@ constructors:
\item $(\texttt{VAR}\;id)$, a reference to a global identifier called $id$;
\item $(\texttt{Rel}\;n)$, a bound variable, whose binder is the $nth$
binder up in the term;
-\item $\texttt{DLAM}\;(x,t)$, a deBruijn's binder on the term $t$;
-\item $\texttt{DLAMV}\;(x,vt)$, a deBruijn's binder on all the terms of
+\item $\texttt{DLAM}\;(x,t)$, a de Bruijn's binder on the term $t$;
+\item $\texttt{DLAMV}\;(x,vt)$, a de Bruijn's binder on all the terms of
the vector $vt$;
\item $(\texttt{DOP0}\;op)$, a unary operator $op$;
\item $\texttt{DOP2}\;(op,t_1,t_2)$, the application of a binary
@@ -299,7 +299,7 @@ vector of terms $vt$.
\end{itemize}
In this meta-language, bound variables are represented using the
-so-called deBrujin's indexes. In this representation, an occurrence of
+so-called de Bruijn's indexes. In this representation, an occurrence of
a bound variable is denoted by an integer, meaning the number of
binders that must be traversed to reach its own
binder\footnote{Actually, $(\texttt{Rel}\;n)$ means that $(n-1)$ binders
@@ -339,7 +339,7 @@ on the terms of the meta-language:
\fun{val Generic.dependent : 'op term -> 'op term -> bool}
{Returns true if the first term is a sub-term of the second.}
%\fun{val Generic.subst\_var : identifier -> 'op term -> 'op term}
-% { $(\texttt{subst\_var}\;id\;t)$ substitutes the deBruijn's index
+% { $(\texttt{subst\_var}\;id\;t)$ substitutes the de Bruijn's index
% associated to $id$ to every occurrence of the term
% $(\texttt{VAR}\;id)$ in $t$.}
\end{description}
@@ -482,7 +482,7 @@ following constructor functions:
\begin{description}
\fun{val Term.mkRel : int -> constr}
- {$(\texttt{mkRel}\;n)$ represents deBrujin's index $n$.}
+ {$(\texttt{mkRel}\;n)$ represents de Bruijn's index $n$.}
\fun{val Term.mkVar : identifier -> constr}
{$(\texttt{mkVar}\;id)$
@@ -545,7 +545,7 @@ following constructor functions:
\fun{val Term.mkProd : name ->constr ->constr -> constr}
{$(\texttt{mkProd}\;x\;A\;B)$ represents the product $(x:A)B$.
- The free ocurrences of $x$ in $B$ are represented by deBrujin's
+ The free ocurrences of $x$ in $B$ are represented by de Bruijn's
indexes.}
\fun{val Term.mkNamedProd : identifier -> constr -> constr -> constr}
@@ -553,14 +553,14 @@ following constructor functions:
but the bound occurrences of $x$ in $B$ are denoted by
the identifier $(\texttt{mkVar}\;x)$. The function automatically
changes each occurrences of this identifier into the corresponding
- deBrujin's index.}
+ de Bruijn's index.}
\fun{val Term.mkArrow : constr -> constr -> constr}
{$(\texttt{arrow}\;A\;B)$ represents the type $(A\rightarrow B)$.}
\fun{val Term.mkLambda : name -> constr -> constr -> constr}
{$(\texttt{mkLambda}\;x\;A\;b)$ represents the lambda abstraction
- $[x:A]b$. The free ocurrences of $x$ in $B$ are represented by deBrujin's
+ $[x:A]b$. The free ocurrences of $x$ in $B$ are represented by de Bruijn's
indexes.}
\fun{val Term.mkNamedLambda : identifier -> constr -> constr -> constr}
@@ -666,7 +666,7 @@ use the primitive \textsl{Case} described in Chapter \ref{Cic}
\item Restoring type coercions and synthesizing the implicit arguments
(the one denoted by question marks in
{\Coq} syntax: see Section~\ref{Coercions}).
-\item Transforming the named bound variables into deBrujin's indexes.
+\item Transforming the named bound variables into de Bruijn's indexes.
\item Classifying the global names into the different classes of
constants (defined constants, constructors, inductive types, etc).
\end{enumerate}
diff --git a/engine/evarutil.ml b/engine/evarutil.ml
index fba466107..e85c1f6fd 100644
--- a/engine/evarutil.ml
+++ b/engine/evarutil.ml
@@ -461,7 +461,7 @@ let e_new_evar env evdref ?(src=default_source) ?filter ?candidates ?store ?nami
ev
(* This assumes an evar with identity instance and generalizes it over only
- the De Bruijn part of the context *)
+ the de Bruijn part of the context *)
let generalize_evar_over_rels sigma (ev,args) =
let open EConstr in
let evi = Evd.find sigma ev in
diff --git a/ide/texmacspp.ml b/ide/texmacspp.ml
index e20704b7f..05f1820cf 100644
--- a/ide/texmacspp.ml
+++ b/ide/texmacspp.ml
@@ -552,7 +552,7 @@ let rec tmpp v loc =
let str_dk = Kindops.string_of_definition_kind (l, false, dk) in
let str_id = Id.to_string id in
(xmlDef str_dk str_id loc [pp_expr e])
- | VernacStartTheoremProof (tk, [ Some ((_,id),_), ([], statement, None) ], b) ->
+ | VernacStartTheoremProof (tk, [ Some ((_,id),_), ([], statement) ], b) ->
let str_tk = Kindops.string_of_theorem_kind tk in
let str_id = Id.to_string id in
(xmlThm str_tk str_id loc [pp_expr statement])
diff --git a/interp/topconstr.ml b/interp/topconstr.ml
index d3142e7f0..e05be65fb 100644
--- a/interp/topconstr.ml
+++ b/interp/topconstr.ml
@@ -178,7 +178,12 @@ let split_at_annot bl na =
in
(List.rev ans, CLocalAssum (r, k, t) :: rest)
end
- | CLocalDef _ as x :: rest -> aux (x :: acc) rest
+ | CLocalDef ((_,na),_,_) as x :: rest ->
+ if Name.equal (Name id) na then
+ user_err ~loc
+ (Nameops.pr_id id ++ str" must be a proper parameter and not a local definition.")
+ else
+ aux (x :: acc) rest
| CLocalPattern (loc,_,_) :: rest ->
Loc.raise ~loc (Stream.Error "pattern with quote not allowed after fix")
| [] ->
diff --git a/intf/constrexpr.mli b/intf/constrexpr.mli
index 49bafadc8..a4a6eb909 100644
--- a/intf/constrexpr.mli
+++ b/intf/constrexpr.mli
@@ -19,7 +19,7 @@ open Decl_kinds
type notation = string
type explicitation =
- | ExplByPos of int * Id.t option
+ | ExplByPos of int * Id.t option (* a reference to the n-th product starting from left *)
| ExplByName of Id.t
type binder_kind =
diff --git a/intf/vernacexpr.mli b/intf/vernacexpr.mli
index f018d59e6..cb093d85d 100644
--- a/intf/vernacexpr.mli
+++ b/intf/vernacexpr.mli
@@ -209,7 +209,7 @@ type one_inductive_expr =
plident * local_binder_expr list * constr_expr option * constructor_expr list
type proof_expr =
- plident option * (local_binder_expr list * constr_expr * (lident option * recursion_order_expr) option)
+ plident option * (local_binder_expr list * constr_expr)
type syntax_modifier =
| SetItemLevel of string list * Extend.production_level
diff --git a/kernel/constr.ml b/kernel/constr.ml
index 71efe9032..eecceb32a 100644
--- a/kernel/constr.ml
+++ b/kernel/constr.ml
@@ -107,7 +107,16 @@ type constr = t
type existential = existential_key * constr array
type rec_declaration = Name.t array * constr array * constr array
type fixpoint = (int array * int) * rec_declaration
+ (* The array of [int]'s tells for each component of the array of
+ mutual fixpoints the number of lambdas to skip before finding the
+ recursive argument (e.g., value is 2 in "fix f (x:A) (y:=t) (z:B)
+ (v:=u) (w:I) {struct w}"), telling to skip x and z and that w is
+ the recursive argument);
+ The second component [int] tells which component of the block is
+ returned *)
type cofixpoint = int * rec_declaration
+ (* The component [int] tells which component of the block of
+ cofixpoint is returned *)
type types = constr
@@ -115,7 +124,7 @@ type types = constr
(* Term constructors *)
(*********************)
-(* Constructs a DeBrujin index with number n *)
+(* Constructs a de Bruijn index with number n *)
let rels =
[|Rel 1;Rel 2;Rel 3;Rel 4;Rel 5;Rel 6;Rel 7; Rel 8;
Rel 9;Rel 10;Rel 11;Rel 12;Rel 13;Rel 14;Rel 15; Rel 16|]
diff --git a/kernel/constr.mli b/kernel/constr.mli
index 700c235e6..e0954160f 100644
--- a/kernel/constr.mli
+++ b/kernel/constr.mli
@@ -70,7 +70,7 @@ type types = constr
(** {6 Term constructors. } *)
-(** Constructs a DeBrujin index (DB indices begin at 1) *)
+(** Constructs a de Bruijn index (DB indices begin at 1) *)
val mkRel : int -> constr
(** Constructs a Variable *)
diff --git a/kernel/nativeconv.ml b/kernel/nativeconv.ml
index 3c0afe380..3593d94c2 100644
--- a/kernel/nativeconv.ml
+++ b/kernel/nativeconv.ml
@@ -141,7 +141,7 @@ let native_conv_gen pb sigma env univs t1 t2 =
let t1 = Sys.time () in
let time_info = Format.sprintf "Evaluation done in %.5f@." (t1 -. t0) in
if !Flags.debug then Feedback.msg_debug (Pp.str time_info);
- (* TODO change 0 when we can have deBruijn *)
+ (* TODO change 0 when we can have de Bruijn *)
fst (conv_val env pb 0 !rt1 !rt2 univs)
end
| _ -> anomaly (Pp.str "Compilation failure")
diff --git a/kernel/term.ml b/kernel/term.ml
index e5a681375..03562d9f3 100644
--- a/kernel/term.ml
+++ b/kernel/term.ml
@@ -169,7 +169,7 @@ let hcons_types = Constr.hcons
exception DestKO
-(* Destructs a DeBrujin index *)
+(* Destructs a de Bruijn index *)
let destRel c = match kind_of_term c with
| Rel n -> n
| _ -> raise DestKO
diff --git a/kernel/term.mli b/kernel/term.mli
index a9bb67705..241ef322f 100644
--- a/kernel/term.mli
+++ b/kernel/term.mli
@@ -127,7 +127,7 @@ val is_small : sorts -> bool
exception DestKO
-(** Destructs a DeBrujin index *)
+(** Destructs a de Bruijn index *)
val destRel : constr -> int
(** Destructs an existential variable *)
diff --git a/kernel/vars.ml b/kernel/vars.ml
index 4affb5f9f..f1c0a4f08 100644
--- a/kernel/vars.ml
+++ b/kernel/vars.ml
@@ -27,7 +27,7 @@ let closedn n c =
in
try closed_rec n c; true with LocalOccur -> false
-(* [closed0 M] is true iff [M] is a (deBruijn) closed term *)
+(* [closed0 M] is true iff [M] is a (de Bruijn) closed term *)
let closed0 c = closedn 0 c
diff --git a/kernel/vars.mli b/kernel/vars.mli
index adeac422e..df5c55118 100644
--- a/kernel/vars.mli
+++ b/kernel/vars.mli
@@ -11,10 +11,10 @@ open Constr
(** {6 Occur checks } *)
-(** [closedn n M] is true iff [M] is a (deBruijn) closed term under n binders *)
+(** [closedn n M] is true iff [M] is a (de Bruijn) closed term under n binders *)
val closedn : int -> constr -> bool
-(** [closed0 M] is true iff [M] is a (deBruijn) closed term *)
+(** [closed0 M] is true iff [M] is a (de Bruijn) closed term *)
val closed0 : constr -> bool
(** [noccurn n M] returns true iff [Rel n] does NOT occur in term [M] *)
diff --git a/parsing/g_vernac.ml4 b/parsing/g_vernac.ml4
index 3438635cf..085c98e37 100644
--- a/parsing/g_vernac.ml4
+++ b/parsing/g_vernac.ml4
@@ -136,8 +136,8 @@ GEXTEND Gram
[ [ thm = thm_token; id = pidentref; bl = binders; ":"; c = lconstr;
l = LIST0
[ "with"; id = pidentref; bl = binders; ":"; c = lconstr ->
- (Some id,(bl,c,None)) ] ->
- VernacStartTheoremProof (thm, (Some id,(bl,c,None))::l, false)
+ (Some id,(bl,c)) ] ->
+ VernacStartTheoremProof (thm, (Some id,(bl,c))::l, false)
| stre = assumption_token; nl = inline; bl = assum_list ->
VernacAssumption (stre, nl, bl)
| (kwd,stre) = assumptions_token; nl = inline; bl = assum_list ->
diff --git a/printing/ppvernac.ml b/printing/ppvernac.ml
index 58475630e..3e41439c8 100644
--- a/printing/ppvernac.ml
+++ b/printing/ppvernac.ml
@@ -386,13 +386,12 @@ open Decl_kinds
++ pr_opt (fun def -> str":=" ++ brk(1,2) ++ pr_pure_lconstr def) def
++ prlist (pr_decl_notation pr_constr) ntn
- let pr_statement head (idpl,(bl,c,guard)) =
+ let pr_statement head (idpl,(bl,c)) =
assert (not (Option.is_empty idpl));
let id, pl = Option.get idpl in
hov 2
(head ++ spc() ++ pr_lident id ++ pr_univs pl ++ spc() ++
(match bl with [] -> mt() | _ -> pr_binders bl ++ spc()) ++
- pr_opt (pr_guard_annot pr_lconstr_expr bl) guard ++
str":" ++ pr_spc_lconstr c)
(**************************************)
diff --git a/tactics/dnet.mli b/tactics/dnet.mli
index 9f29c60b6..565a916f8 100644
--- a/tactics/dnet.mli
+++ b/tactics/dnet.mli
@@ -26,7 +26,7 @@
distincts, or you'll get unexpected behaviours.
Warning 2: This structure is perfect, i.e. the set of candidates
- returned is equal to the set of solutions. Beware of DeBruijn
+ returned is equal to the set of solutions. Beware of de Bruijn
shifts and sorts subtyping though (which makes the comparison not
symmetric, see term_dnet.ml).
diff --git a/tactics/tactics.ml b/tactics/tactics.ml
index f3f7d16b7..211a7338b 100644
--- a/tactics/tactics.ml
+++ b/tactics/tactics.ml
@@ -512,6 +512,9 @@ let rec check_mutind env sigma k cl = match EConstr.kind sigma (strip_outer_cast
else
let open Context.Rel.Declaration in
check_mutind (push_rel (LocalAssum (na, c1)) env) sigma (pred k) b
+| LetIn (na, c1, t, b) ->
+ let open Context.Rel.Declaration in
+ check_mutind (push_rel (LocalDef (na, c1, t)) env) sigma k b
| _ -> error "Not enough products."
(* Refine as a fixpoint *)
diff --git a/test-suite/success/ImplicitArguments.v b/test-suite/success/ImplicitArguments.v
index f702aa62f..f07773f8b 100644
--- a/test-suite/success/ImplicitArguments.v
+++ b/test-suite/success/ImplicitArguments.v
@@ -21,3 +21,9 @@ Fixpoint app {A : Type} {n m : nat} (v : vector A n) (w : vector A m) : vector A
(* Test sharing information between different hypotheses *)
Parameters (a:_) (b:a=0).
+
+(* These examples were failing due to a lifting wrongly taking let-in into account *)
+
+Definition foo6 (x:=1) : forall {n:nat}, n=n := fun n => eq_refl.
+
+Fixpoint foo7 (x:=1) (n:nat) {p:nat} {struct n} : nat.
diff --git a/theories/Reals/Rfunctions.v b/theories/Reals/Rfunctions.v
index 99acdd0a1..606136333 100644
--- a/theories/Reals/Rfunctions.v
+++ b/theories/Reals/Rfunctions.v
@@ -532,6 +532,36 @@ Qed.
(*******************************)
(*i Due to L.Thery i*)
+Section PowerRZ.
+
+Local Coercion Z_of_nat : nat >-> Z.
+
+(* the following section should probably be somewhere else, but not sure where *)
+Section Z_compl.
+
+Local Open Scope Z_scope.
+
+(* Provides a way to reason directly on Z in terms of nats instead of positive *)
+Inductive Z_spec (x : Z) : Z -> Type :=
+| ZintNull : x = 0 -> Z_spec x 0
+| ZintPos (n : nat) : x = n -> Z_spec x n
+| ZintNeg (n : nat) : x = - n -> Z_spec x (- n).
+
+Lemma intP (x : Z) : Z_spec x x.
+Proof.
+ destruct x as [|p|p].
+ - now apply ZintNull.
+ - rewrite <-positive_nat_Z at 2.
+ apply ZintPos.
+ now rewrite positive_nat_Z.
+ - rewrite <-Pos2Z.opp_pos.
+ rewrite <-positive_nat_Z at 2.
+ apply ZintNeg.
+ now rewrite positive_nat_Z.
+Qed.
+
+End Z_compl.
+
Definition powerRZ (x:R) (n:Z) :=
match n with
| Z0 => 1
@@ -658,6 +688,74 @@ Proof.
auto with real.
Qed.
+Local Open Scope Z_scope.
+
+Lemma pow_powerRZ (r : R) (n : nat) :
+ (r ^ n)%R = powerRZ r (Z_of_nat n).
+Proof.
+ induction n; [easy|simpl].
+ now rewrite SuccNat2Pos.id_succ.
+Qed.
+
+Lemma powerRZ_ind (P : Z -> R -> R -> Prop) :
+ (forall x, P 0 x 1%R) ->
+ (forall x n, P (Z.of_nat n) x (x ^ n)%R) ->
+ (forall x n, P ((-(Z.of_nat n))%Z) x (Rinv (x ^ n))) ->
+ forall x (m : Z), P m x (powerRZ x m)%R.
+Proof.
+ intros ? ? ? x m.
+ destruct (intP m) as [Hm|n Hm|n Hm].
+ - easy.
+ - now rewrite <- pow_powerRZ.
+ - unfold powerRZ.
+ destruct n as [|n]; [ easy |].
+ rewrite Nat2Z.inj_succ, <- Zpos_P_of_succ_nat, Pos2Z.opp_pos.
+ now rewrite <- Pos2Z.opp_pos, <- positive_nat_Z.
+Qed.
+
+Lemma powerRZ_inv x alpha : (x <> 0)%R -> powerRZ (/ x) alpha = Rinv (powerRZ x alpha).
+Proof.
+ intros; destruct (intP alpha).
+ - now simpl; rewrite Rinv_1.
+ - now rewrite <-!pow_powerRZ, ?Rinv_pow, ?pow_powerRZ.
+ - unfold powerRZ.
+ destruct (- n).
+ + now rewrite Rinv_1.
+ + now rewrite Rinv_pow.
+ + now rewrite <-Rinv_pow.
+Qed.
+
+Lemma powerRZ_neg x : forall alpha, x <> R0 -> powerRZ x (- alpha) = powerRZ (/ x) alpha.
+Proof.
+ intros [|n|n] H ; simpl.
+ - easy.
+ - now rewrite Rinv_pow.
+ - rewrite Rinv_pow by now apply Rinv_neq_0_compat.
+ now rewrite Rinv_involutive.
+Qed.
+
+Lemma powerRZ_mult_distr :
+ forall m x y, ((0 <= m)%Z \/ (x * y <> 0)%R) ->
+ (powerRZ (x*y) m = powerRZ x m * powerRZ y m)%R.
+Proof.
+ intros m x0 y0 Hmxy.
+ destruct (intP m) as [ | | n Hm ].
+ - now simpl; rewrite Rmult_1_l.
+ - now rewrite <- !pow_powerRZ, Rpow_mult_distr.
+ - destruct Hmxy as [H|H].
+ + assert(m = 0) as -> by now omega.
+ now rewrite <- Hm, Rmult_1_l.
+ + assert(x0 <> 0)%R by now intros ->; apply H; rewrite Rmult_0_l.
+ assert(y0 <> 0)%R by now intros ->; apply H; rewrite Rmult_0_r.
+ rewrite !powerRZ_neg by assumption.
+ rewrite Rinv_mult_distr by assumption.
+ now rewrite <- !pow_powerRZ, Rpow_mult_distr.
+Qed.
+
+End PowerRZ.
+
+Local Infix "^Z" := powerRZ (at level 30, right associativity) : R_scope.
+
(*******************************)
(* For easy interface *)
(*******************************)
diff --git a/theories/Reals/Rpower.v b/theories/Reals/Rpower.v
index b8040bb4f..0e0246cbf 100644
--- a/theories/Reals/Rpower.v
+++ b/theories/Reals/Rpower.v
@@ -473,6 +473,20 @@ Proof.
apply exp_Ropp.
Qed.
+Lemma powerRZ_Rpower x z : (0 < x)%R -> powerRZ x z = Rpower x (IZR z).
+Proof.
+ intros Hx.
+ assert (x <> 0)%R
+ by now intros Habs; rewrite Habs in Hx; apply (Rlt_irrefl 0).
+ destruct (intP z).
+ - now rewrite Rpower_O.
+ - rewrite <- pow_powerRZ, <- Rpower_pow by assumption.
+ now rewrite INR_IZR_INZ.
+ - rewrite opp_IZR, Rpower_Ropp.
+ rewrite powerRZ_neg, powerRZ_inv by assumption.
+ now rewrite <- pow_powerRZ, <- INR_IZR_INZ, Rpower_pow.
+Qed.
+
Theorem Rle_Rpower :
forall e n m:R, 1 < e -> 0 <= n -> n <= m -> e ^R n <= e ^R m.
Proof.
diff --git a/vernac/auto_ind_decl.ml b/vernac/auto_ind_decl.ml
index 25091f783..f363deac6 100644
--- a/vernac/auto_ind_decl.ml
+++ b/vernac/auto_ind_decl.ml
@@ -92,7 +92,7 @@ let destruct_on_using c id =
let destruct_on_as c l =
destruct false None c (Some (dl,l)) None
-(* reconstruct the inductive with the correct deBruijn indexes *)
+(* reconstruct the inductive with the correct de Bruijn indexes *)
let mkFullInd (ind,u) n =
let mib = Global.lookup_mind (fst ind) in
let nparams = mib.mind_nparams in
@@ -171,7 +171,7 @@ let build_beq_scheme mode kn =
(* give a type A, this function tries to find the equality on A declared
previously *)
(* nlist = the number of args (A , B , ... )
- eqA = the deBruijn index of the first eq param
+ eqA = the de Bruijn index of the first eq param
ndx = how much to translate due to the 2nd Case
*)
let compute_A_equality rel_list nlist eqA ndx t =
diff --git a/vernac/command.ml b/vernac/command.ml
index b2f5755ce..2fa2aa4e3 100644
--- a/vernac/command.ml
+++ b/vernac/command.ml
@@ -96,7 +96,7 @@ let interp_definition pl bl p red_option c ctypopt =
let evdref = ref (Evd.from_ctx ctx) in
let impls, ((env_bl, ctx), imps1) = interp_context_evars env evdref bl in
let ctx = List.map (fun d -> map_rel_decl EConstr.Unsafe.to_constr d) ctx in
- let nb_args = List.length ctx in
+ let nb_args = Context.Rel.nhyps ctx in
let imps,pl,ce =
match ctypopt with
None ->
@@ -849,7 +849,7 @@ type structured_fixpoint_expr = {
let interp_fix_context env evdref isfix fix =
let before, after = if isfix then split_at_annot fix.fix_binders fix.fix_annot else [], fix.fix_binders in
let impl_env, ((env', ctx), imps) = interp_context_evars env evdref before in
- let impl_env', ((env'', ctx'), imps') = interp_context_evars ~impl_env ~shift:(List.length before) env' evdref after in
+ let impl_env', ((env'', ctx'), imps') = interp_context_evars ~impl_env ~shift:(Context.Rel.nhyps ctx) env' evdref after in
let annot = Option.map (fun _ -> List.length (assums_of_rel_context ctx)) fix.fix_annot in
((env'', ctx' @ ctx), (impl_env',imps @ imps'), annot)
@@ -880,8 +880,10 @@ let prepare_recursive_declaration fixnames fixtypes fixdefs =
(* Jump over let-bindings. *)
-let compute_possible_guardness_evidences (ids,_,na) =
- match na with
+let compute_possible_guardness_evidences (ctx,_,recindex) =
+ (* A recursive index is characterized by the number of lambdas to
+ skip before finding the relevant inductive argument *)
+ match recindex with
| Some i -> [i]
| None ->
(* If recursive argument was not given by user, we try all args.
@@ -889,7 +891,7 @@ let compute_possible_guardness_evidences (ids,_,na) =
but doing it properly involves delta-reduction, and it finally
doesn't seem to worth the effort (except for huge mutual
fixpoints ?) *)
- List.interval 0 (List.length ids - 1)
+ List.interval 0 (Context.Rel.nhyps ctx - 1)
type recursive_preentry =
Id.t list * constr option list * types list
@@ -1130,7 +1132,7 @@ let interp_recursive isfix fixl notations =
let fixtypes = List.map2 build_fix_type fixctxs fixccls in
let fixtypes = List.map (fun c -> nf_evar !evdref c) fixtypes in
let fiximps = List.map3
- (fun ctximps cclimps (_,ctx) -> ctximps@(Impargs.lift_implicits (List.length ctx) cclimps))
+ (fun ctximps cclimps (_,ctx) -> ctximps@(Impargs.lift_implicits (Context.Rel.nhyps ctx) cclimps))
fixctximps fixcclimps fixctxs in
let rec_sign =
List.fold_left2
@@ -1169,10 +1171,10 @@ let interp_recursive isfix fixl notations =
let fixdefs = List.map (fun c -> Option.map EConstr.Unsafe.to_constr c) fixdefs in
let fixdefs = List.map (Option.map nf) fixdefs in
let fixtypes = List.map nf fixtypes in
- let fixctxnames = List.map (fun (_,ctx) -> List.map RelDecl.get_name ctx) fixctxs in
+ let fixctxs = List.map (fun (_,ctx) -> ctx) fixctxs in
(* Build the fix declaration block *)
- (env,rec_sign,all_universes,evd), (fixnames,fixdefs,fixtypes), List.combine3 fixctxnames fiximps fixannots
+ (env,rec_sign,all_universes,evd), (fixnames,fixdefs,fixtypes), List.combine3 fixctxs fiximps fixannots
let check_recursive isfix env evd (fixnames,fixdefs,_) =
check_evars_are_solved env evd Evd.empty;
@@ -1188,14 +1190,14 @@ let interp_fixpoint l ntns =
let interp_cofixpoint l ntns =
let (env,_,pl,evd),fix,info = interp_recursive false l ntns in
- check_recursive false env evd fix;
+ check_recursive false env evd fix;
(fix,pl,Evd.evar_universe_context evd,info)
let declare_fixpoint local poly ((fixnames,fixdefs,fixtypes),pl,ctx,fiximps) indexes ntns =
if List.exists Option.is_empty fixdefs then
(* Some bodies to define by proof *)
let thms =
- List.map3 (fun id t (len,imps,_) -> ((id,pl),(t,(len,imps))))
+ List.map3 (fun id t (ctx,imps,_) -> ((id,pl),(t,(List.map RelDecl.get_name ctx,imps))))
fixnames fixtypes fiximps in
let init_tac =
Some (List.map (Option.cata (EConstr.of_constr %> Tactics.exact_no_check) Tacticals.New.tclIDTAC)
@@ -1229,7 +1231,7 @@ let declare_cofixpoint local poly ((fixnames,fixdefs,fixtypes),pl,ctx,fiximps) n
if List.exists Option.is_empty fixdefs then
(* Some bodies to define by proof *)
let thms =
- List.map3 (fun id t (len,imps,_) -> ((id,pl),(t,(len,imps))))
+ List.map3 (fun id t (ctx,imps,_) -> ((id,pl),(t,(List.map RelDecl.get_name ctx,imps))))
fixnames fixtypes fiximps in
let init_tac =
Some (List.map (Option.cata (EConstr.of_constr %> Tactics.exact_no_check) Tacticals.New.tclIDTAC)
diff --git a/vernac/command.mli b/vernac/command.mli
index 7cd0afeec..9bbc2fdac 100644
--- a/vernac/command.mli
+++ b/vernac/command.mli
@@ -138,24 +138,24 @@ type recursive_preentry =
val interp_fixpoint :
structured_fixpoint_expr list -> decl_notation list ->
recursive_preentry * lident list option * Evd.evar_universe_context *
- (Name.t list * Impargs.manual_implicits * int option) list
+ (EConstr.rel_context * Impargs.manual_implicits * int option) list
val interp_cofixpoint :
structured_fixpoint_expr list -> decl_notation list ->
recursive_preentry * lident list option * Evd.evar_universe_context *
- (Name.t list * Impargs.manual_implicits * int option) list
+ (EConstr.rel_context * Impargs.manual_implicits * int option) list
(** Registering fixpoints and cofixpoints in the environment *)
val declare_fixpoint :
locality -> polymorphic ->
recursive_preentry * lident list option * Evd.evar_universe_context *
- (Name.t list * Impargs.manual_implicits * int option) list ->
+ (Context.Rel.t * Impargs.manual_implicits * int option) list ->
lemma_possible_guards -> decl_notation list -> unit
val declare_cofixpoint : locality -> polymorphic ->
recursive_preentry * lident list option * Evd.evar_universe_context *
- (Name.t list * Impargs.manual_implicits * int option) list ->
+ (Context.Rel.t * Impargs.manual_implicits * int option) list ->
decl_notation list -> unit
(** Entry points for the vernacular commands Fixpoint and CoFixpoint *)
diff --git a/vernac/lemmas.ml b/vernac/lemmas.ml
index 993a2c260..1344701ff 100644
--- a/vernac/lemmas.ml
+++ b/vernac/lemmas.ml
@@ -88,25 +88,9 @@ let adjust_guardness_conditions const = function
let find_mutually_recursive_statements thms =
let n = List.length thms in
- let inds = List.map (fun (id,(t,impls,annot)) ->
+ let inds = List.map (fun (id,(t,impls)) ->
let (hyps,ccl) = decompose_prod_assum t in
let x = (id,(t,impls)) in
- match annot with
- (* Explicit fixpoint decreasing argument is given *)
- | Some (Some (_,id),CStructRec) ->
- let i,b,typ = lookup_rel_id id hyps in
- (match kind_of_term t with
- | Ind ((kn,_ as ind), u) when
- let mind = Global.lookup_mind kn in
- mind.mind_finite == Decl_kinds.Finite && Option.is_empty b ->
- [ind,x,i],[]
- | _ ->
- error "Decreasing argument is not an inductive assumption.")
- (* Unsupported cases *)
- | Some (_,(CWfRec _|CMeasureRec _)) ->
- error "Only structural decreasing is supported for mutual statements."
- (* Cofixpoint or fixpoint w/o explicit decreasing argument *)
- | None | Some (None, CStructRec) ->
let whnf_hyp_hds = map_rel_context_in_env
(fun env c -> EConstr.Unsafe.to_constr (fst (whd_all_stack env Evd.empty (EConstr.of_constr c))))
(Global.env()) hyps in
@@ -116,10 +100,10 @@ let find_mutually_recursive_statements thms =
match kind_of_term t with
| Ind ((kn,_ as ind),u) when
let mind = Global.lookup_mind kn in
- mind.mind_finite <> Decl_kinds.CoFinite && is_local_assum decl ->
+ mind.mind_finite <> Decl_kinds.CoFinite ->
[ind,x,i]
| _ ->
- []) 0 (List.rev whnf_hyp_hds)) in
+ []) 0 (List.rev (List.filter RelDecl.is_local_assum whnf_hyp_hds))) in
let ind_ccl =
let cclenv = push_rel_context hyps (Global.env()) in
let whnf_ccl,_ = whd_all_stack cclenv Evd.empty (EConstr.of_constr ccl) in
@@ -178,7 +162,7 @@ let find_mutually_recursive_statements thms =
(finite,guard,None), ordered_inds
let look_for_possibly_mutual_statements = function
- | [id,(t,impls,None)] ->
+ | [id,(t,impls)] ->
(* One non recursively proved theorem *)
None,[id,(t,impls)],None
| _::_ as thms ->
@@ -458,7 +442,7 @@ let start_proof_com ?inference_hook kind thms hook =
| None -> Evd.from_env env0
| Some l -> Evd.from_ctx (Evd.make_evar_universe_context env0 l))
in
- let thms = List.map (fun (sopt,(bl,t,guard)) ->
+ let thms = List.map (fun (sopt,(bl,t)) ->
let impls, ((env, ctx), imps) = interp_context_evars env0 evdref bl in
let t', imps' = interp_type_evars_impls ~impls env evdref t in
let flags = all_and_fail_flags in
@@ -467,8 +451,7 @@ let start_proof_com ?inference_hook kind thms hook =
let ids = List.map RelDecl.get_name ctx in
(compute_proof_name (pi1 kind) sopt,
(EConstr.Unsafe.to_constr (nf_evar !evdref (EConstr.it_mkProd_or_LetIn t' ctx)),
- (ids, imps @ lift_implicits (List.length ids) imps'),
- guard)))
+ (ids, imps @ lift_implicits (Context.Rel.nhyps ctx) imps'))))
thms in
let recguard,thms,snl = look_for_possibly_mutual_statements thms in
let evd, nf = Evarutil.nf_evars_and_universes !evdref in
diff --git a/vernac/lemmas.mli b/vernac/lemmas.mli
index 681413a29..d06b8fd14 100644
--- a/vernac/lemmas.mli
+++ b/vernac/lemmas.mli
@@ -41,8 +41,8 @@ val start_proof_com :
val start_proof_with_initialization :
goal_kind -> Evd.evar_map ->
(bool * lemma_possible_guards * unit Proofview.tactic list option) option ->
- ((Id.t * universe_binders option) *
- (types * (Name.t list * Impargs.manual_explicitation list))) list
+ ((Id.t (* name of thm *) * universe_binders option) *
+ (types (* type of thm *) * (Name.t list (* names to pre-introduce *) * Impargs.manual_explicitation list))) list
-> int list option -> unit declaration_hook -> unit
val universe_proof_terminator :
diff --git a/vernac/obligations.ml b/vernac/obligations.ml
index ea2401b5c..e0520216b 100644
--- a/vernac/obligations.ml
+++ b/vernac/obligations.ml
@@ -8,7 +8,7 @@ open Declare
(**
- Get types of existentials ;
- Flatten dependency tree (prefix order) ;
- - Replace existentials by De Bruijn indices in term, applied to the right arguments ;
+ - Replace existentials by de Bruijn indices in term, applied to the right arguments ;
- Apply term prefixed by quantification on "existentials".
*)
@@ -51,7 +51,7 @@ type oblinfo =
ev_tac: unit Proofview.tactic option;
ev_deps: Int.Set.t }
-(** Substitute evar references in t using De Bruijn indices,
+(** Substitute evar references in t using de Bruijn indices,
where n binders were passed through. *)
let subst_evar_constr evs n idf t =
@@ -102,7 +102,7 @@ let subst_evar_constr evs n idf t =
t', !seen, !transparent
-(** Substitute variable references in t using De Bruijn indices,
+(** Substitute variable references in t using de Bruijn indices,
where n binders were passed through. *)
let subst_vars acc n t =
let var_index id = Util.List.index Id.equal id acc in
diff --git a/vernac/vernacentries.ml b/vernac/vernacentries.ml
index 0a5a000fe..2cb6f3918 100644
--- a/vernac/vernacentries.ml
+++ b/vernac/vernacentries.ml
@@ -477,7 +477,7 @@ let vernac_definition locality p (local,k) ((loc,id as lid),pl) def =
(match def with
| ProveBody (bl,t) -> (* local binders, typ *)
start_proof_and_print (local,p,DefinitionBody k)
- [Some (lid,pl), (bl,t,None)] hook
+ [Some (lid,pl), (bl,t)] hook
| DefineBody (bl,red_option,c,typ_opt) ->
let red_option = match red_option with
| None -> None
@@ -2055,7 +2055,7 @@ let interp ?proof ~loc locality poly c =
| VernacComments l -> if_verbose Feedback.msg_info (str "Comments ok\n")
(* Proof management *)
- | VernacGoal t -> vernac_start_proof locality poly Theorem [None,([],t,None)] false
+ | VernacGoal t -> vernac_start_proof locality poly Theorem [None,([],t)] false
| VernacFocus n -> vernac_focus n
| VernacUnfocus -> vernac_unfocus ()
| VernacUnfocused -> vernac_unfocused ()