Codebase list coq / 420f78b
Imported Upstream version 8.4pl4dfsg Stephane Glondu 9 years ago
1014 changed file(s) with 1472 addition(s) and 1269 deletion(s). Raw diff Collapse all Expand all
0 Changes from V8.4pl3 to V8.4pl4
1 ===============================
2
3 WARNING:
4 The current logic of Coq is now known to be inconsistent with
5 Axiom prop_extensionality : forall A B:Prop, (A <-> B) -> A = B.
6 For more details, see:
7 https://gforge.inria.fr/plugins/scmgit/cgi-bin/gitweb.cgi?p=coq/coq.git;a=blob_plain;f=test-suite/failure/subterm2.v;hb=HEAD
8 or
9 https://gforge.inria.fr/plugins/scmgit/cgi-bin/gitweb.cgi?p=coq/coq.git;a=blob_plain;f=test-suite/failure/subterm3.v;hb=HEAD
10
11 Kernel
12
13 - Bug #3211: unsound check of elimination sort.
14 - Fix guard condition for nested cofixpoints.
15 - Bug #3243: Univ constraints of module subtyping were not propagated.
16
17 Tactics
18
19 - A new option "Set Stable Omega" ensures that repeated identical calls
20 to omega will produce identical proof terms. This option is off by default
21 for maximal compatibility, but should be pretty safe to activate.
22 - The interpretation of the open_constr tactic argument was erroneously
23 firing type classes resolution in some corner cases. This has been
24 fixed. The tactic argument type open_constr_wTC is provided for retro
25 compatibility purposes.
26 - Fixing bug #3228 (fixing precedence of ltac variables over variables in
27 env) introduces rare and justified tactic failure.
28
29 Bug fixes
30
31 - Solved bugs:
32 #3260, #2697, #3037, #3262, #2900, #3131, #3238, #3204, #1758, #1039,
33 #3144
34 - micromega: solved an ambiguous symbol resolution.
35 - Coq always uses / as separator between directories on all platforms.
36 - remove trailing '\r' from file names returned by coqtop.
37 - bug correction in proving inversion principles for Function.
38 - ocamlbuild: minor fixes related to camlp4 and cross-compilation.
39
040 Changes from V8.4pl2 to V8.4pl3
141 ===============================
242
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
276276 (try conv env a1 a1'
277277 with NotConvertible -> raise (LocalArity None));
278278 srec (push_rel (na1,None,a1) env) t ar'
279 | Prod (_,a1,a2), [] -> (* whnf of t was not needed here! *)
280 let ksort = match (whd_betadeltaiota env a2) with
281 | Sort s -> family_of_sort s
279 | Prod (na1,a1,a2), [] -> (* whnf of t was not needed here! *)
280 let env' = push_rel (na1,None,a1) env in
281 let ksort = match (whd_betadeltaiota env' a2) with
282 | Sort s -> family_of_sort s
282283 | _ -> raise (LocalArity None) in
283284 let dep_ind = build_dependent_inductive ind specif params in
284285 (try conv env a1 dep_ind
288289 | Sort s', [] ->
289290 check_allowed_sort (family_of_sort s') specif;
290291 false
292 | _, (_,Some _,_ as d)::ar' ->
293 srec (push_rel d env) (lift 1 pt') ar'
291294 | _ ->
292295 raise (LocalArity None)
293296 in
894897 raise (CoFixGuardError (env,RecCallInTypeOfAbstraction a))
895898
896899 | CoFix (j,(_,varit,vdefs as recdef)) ->
897 if (List.for_all (noccur_with_meta n nbfix) args)
900 if List.for_all (noccur_with_meta n nbfix) args
898901 then
899 let nbfix = Array.length vdefs in
900 if (array_for_all (noccur_with_meta n nbfix) varit) then
902 if array_for_all (noccur_with_meta n nbfix) varit then
903 let nbfix = Array.length vdefs in
901904 let env' = push_rec_types recdef env in
902 (Array.iter (check_rec_call env' alreadygrd (n+1) vlra) vdefs;
905 (Array.iter (check_rec_call env' alreadygrd (n+nbfix) vlra) vdefs;
903906 List.iter (check_rec_call env alreadygrd n vlra) args)
904907 else
905908 raise (CoFixGuardError (env,RecCallInTypeOfDef c))
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
55 #
66 ##################################
77
8 VERSION=8.4pl3
8 VERSION=8.4pl4
99 VOMAGIC=08400
1010 STATEMAGIC=58400
1111 DATE=`LC_ALL=C LANG=C date +"%B %Y"`
110110 coq_profile_flag=
111111 coq_annotate_flag=
112112 best_compiler=opt
113 cflags="-fno-defer-pop -Wall -Wno-unused"
113 cflags="-Wall -Wno-unused"
114114 natdynlink=yes
115115
116116 local=false
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
5050 let rec read_all_lines in_chan =
5151 try
5252 let arg = input_line in_chan in
53 let len = String.length arg in
54 let arg =
55 if arg.[len - 1] = '\r' then
56 String.sub arg 0 (len - 1)
57 else arg
58 in
5359 arg::(read_all_lines in_chan)
5460 with End_of_file -> []
5561
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
11151115 (try
11161116 let n = Recordops.find_projection_nparams ref + 1 in
11171117 if nargs <> n then
1118 user_err_loc (loc,"",str "Projection has not the right number of explicit parameters.");
1118 user_err_loc (loc,"",str "Projection does not have the right number of explicit parameters.");
11191119 with Not_found ->
11201120 user_err_loc
11211121 (loc,"",pr_global_env Idset.empty ref ++ str " is not a registered projection."))
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
3131 | ConstrArgType
3232 | ConstrMayEvalArgType
3333 | QuantHypArgType
34 | OpenConstrArgType of bool
34 | OpenConstrArgType of bool * bool (* casted, TC resolution *)
3535 | ConstrWithBindingsArgType
3636 | BindingsArgType
3737 | RedExprArgType
155155 let globwit_constr_may_eval = ConstrMayEvalArgType
156156 let wit_constr_may_eval = ConstrMayEvalArgType
157157
158 let rawwit_open_constr_gen b = OpenConstrArgType b
159 let globwit_open_constr_gen b = OpenConstrArgType b
160 let wit_open_constr_gen b = OpenConstrArgType b
161
162 let rawwit_open_constr = rawwit_open_constr_gen false
163 let globwit_open_constr = globwit_open_constr_gen false
164 let wit_open_constr = wit_open_constr_gen false
165
166 let rawwit_casted_open_constr = rawwit_open_constr_gen true
167 let globwit_casted_open_constr = globwit_open_constr_gen true
168 let wit_casted_open_constr = wit_open_constr_gen true
158 let rawwit_open_constr_gen (b1,b2) = OpenConstrArgType (b1,b2)
159 let globwit_open_constr_gen (b1,b2) = OpenConstrArgType (b1,b2)
160 let wit_open_constr_gen (b1,b2) = OpenConstrArgType (b1,b2)
161
162 let rawwit_open_constr = rawwit_open_constr_gen (false,false)
163 let globwit_open_constr = globwit_open_constr_gen (false,false)
164 let wit_open_constr = wit_open_constr_gen (false,false)
165
166 let rawwit_casted_open_constr = rawwit_open_constr_gen (true,false)
167 let globwit_casted_open_constr = globwit_open_constr_gen (true,false)
168 let wit_casted_open_constr = wit_open_constr_gen (true,false)
169
170 let rawwit_open_constr_wTC = rawwit_open_constr_gen (false,true)
171 let globwit_open_constr_wTC = globwit_open_constr_gen (false,true)
172 let wit_open_constr_wTC = wit_open_constr_gen (false,true)
169173
170174 let rawwit_constr_with_bindings = ConstrWithBindingsArgType
171175 let globwit_constr_with_bindings = ConstrWithBindingsArgType
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
182182 val globwit_constr_may_eval : ((glob_constr_and_expr,evaluable_global_reference and_short_name or_var,glob_constr_pattern_and_expr) may_eval,glevel) abstract_argument_type
183183 val wit_constr_may_eval : (constr,tlevel) abstract_argument_type
184184
185 val rawwit_open_constr_gen : bool -> (open_constr_expr,rlevel) abstract_argument_type
186 val globwit_open_constr_gen : bool -> (open_glob_constr,glevel) abstract_argument_type
187 val wit_open_constr_gen : bool -> (open_constr,tlevel) abstract_argument_type
185 val rawwit_open_constr_gen : bool * bool -> (open_constr_expr,rlevel) abstract_argument_type
186 val globwit_open_constr_gen : bool * bool -> (open_glob_constr,glevel) abstract_argument_type
187 val wit_open_constr_gen : bool * bool -> (open_constr,tlevel) abstract_argument_type
188188
189189 val rawwit_open_constr : (open_constr_expr,rlevel) abstract_argument_type
190190 val globwit_open_constr : (open_glob_constr,glevel) abstract_argument_type
193193 val rawwit_casted_open_constr : (open_constr_expr,rlevel) abstract_argument_type
194194 val globwit_casted_open_constr : (open_glob_constr,glevel) abstract_argument_type
195195 val wit_casted_open_constr : (open_constr,tlevel) abstract_argument_type
196
197 val rawwit_open_constr_wTC : (open_constr_expr,rlevel) abstract_argument_type
198 val globwit_open_constr_wTC : (open_glob_constr,glevel) abstract_argument_type
199 val wit_open_constr_wTC : (open_constr,tlevel) abstract_argument_type
196200
197201 val rawwit_constr_with_bindings : (constr_expr with_bindings,rlevel) abstract_argument_type
198202 val globwit_constr_with_bindings : (glob_constr_and_expr with_bindings,glevel) abstract_argument_type
279283 | ConstrArgType
280284 | ConstrMayEvalArgType
281285 | QuantHypArgType
282 | OpenConstrArgType of bool
286 | OpenConstrArgType of bool * bool
283287 | ConstrWithBindingsArgType
284288 | BindingsArgType
285289 | RedExprArgType
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
295295 try conv env a1 a1'
296296 with NotConvertible -> raise (LocalArity None) in
297297 srec (push_rel (na1,None,a1) env) t ar' (union_constraints u univ)
298 | Prod (_,a1,a2), [] -> (* whnf of t was not needed here! *)
299 let ksort = match kind_of_term (whd_betadeltaiota env a2) with
300 | Sort s -> family_of_sort s
301 | _ -> raise (LocalArity None) in
302 let dep_ind = build_dependent_inductive ind specif params in
303 let univ =
304 try conv env a1 dep_ind
305 with NotConvertible -> raise (LocalArity None) in
306 check_allowed_sort ksort specif;
307 union_constraints u univ
298 (* The last Prod domain is the type of the scrutinee *)
299 | Prod (na1,a1,a2), [] -> (* whnf of t was not needed here! *)
300 let env' = push_rel (na1,None,a1) env in
301 let ksort = match kind_of_term (whd_betadeltaiota env' a2) with
302 | Sort s -> family_of_sort s
303 | _ -> raise (LocalArity None) in
304 let dep_ind = build_dependent_inductive ind specif params in
305 let univ =
306 try conv env a1 dep_ind
307 with NotConvertible -> raise (LocalArity None) in
308 check_allowed_sort ksort specif;
309 union_constraints u univ
308310 | _, (_,Some _,_ as d)::ar' ->
309 srec (push_rel d env) (lift 1 pt') ar' u
311 srec (push_rel d env) (lift 1 pt') ar' u
310312 | _ ->
311313 raise (LocalArity None)
312314 in
894896 raise (CoFixGuardError (env,RecCallInTypeOfAbstraction a))
895897
896898 | CoFix (j,(_,varit,vdefs as recdef)) ->
897 if (List.for_all (noccur_with_meta n nbfix) args)
899 if List.for_all (noccur_with_meta n nbfix) args
898900 then
899 let nbfix = Array.length vdefs in
900 if (array_for_all (noccur_with_meta n nbfix) varit) then
901 if array_for_all (noccur_with_meta n nbfix) varit then
902 let nbfix = Array.length vdefs in
901903 let env' = push_rec_types recdef env in
902 (Array.iter (check_rec_call env' alreadygrd (n+1) vlra) vdefs;
904 (Array.iter (check_rec_call env' alreadygrd (n+nbfix) vlra) vdefs;
903905 List.iter (check_rec_call env alreadygrd n vlra) args)
904906 else
905907 raise (CoFixGuardError (env,RecCallInTypeOfDef c))
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
405405 in
406406 let newenv = oldsenv.env in
407407 let newenv = set_engagement_opt senv.engagement newenv in
408 let senv'= {senv with env=newenv} in
408 let senv'= {senv with env = newenv; univ = cst} in
409409 let senv' =
410410 List.fold_left
411411 (fun env (_,mb) -> full_add_module mb env)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
77
88 (* This file gathers environment variables needed by Coq to run (such
99 as coqlib) *)
10
11 let (//) s1 s2 = s1 ^ "/" ^ s2
1012
1113 let coqbin =
1214 System.canonical_path_name (Filename.dirname Sys.executable_name)
2628
2729 let reldir instdir testfile oth =
2830 let rpath = if Coq_config.local then [] else instdir in
29 let out = List.fold_left Filename.concat coqroot rpath in
30 if Sys.file_exists (Filename.concat out testfile) then out else oth ()
31 let out = List.fold_left (//) coqroot rpath in
32 if Sys.file_exists (out//testfile) then out else oth ()
3133
3234 let guess_coqlib () =
3335 let file = "states/initial.coq" in
3739 | Some coqlib -> coqlib
3840 | None -> coqroot
3941 in
40 if Sys.file_exists (Filename.concat coqlib file)
42 if Sys.file_exists (coqlib//file)
4143 then coqlib
4244 else Util.error "cannot guess a path for Coq libraries; please use -coqlib option")
4345
5355 Util.split_string_at sep p
5456
5557 let xdg_data_home =
56 Filename.concat
57 (System.getenv_else "XDG_DATA_HOME" (Filename.concat System.home ".local/share"))
58 "coq"
58 (System.getenv_else "XDG_DATA_HOME" (System.home//".local/share"))//"coq"
5959
6060 let xdg_config_home =
61 Filename.concat
62 (System.getenv_else "XDG_CONFIG_HOME" (Filename.concat System.home ".config"))
63 "coq"
61 (System.getenv_else "XDG_CONFIG_HOME" (System.home//".config"))//"coq"
6462
6563 let xdg_data_dirs =
6664 (try
67 List.map (fun dir -> Filename.concat dir "coq") (path_to_list (Sys.getenv "XDG_DATA_DIRS"))
65 List.map (fun dir -> dir//"coq") (path_to_list (Sys.getenv "XDG_DATA_DIRS"))
6866 with Not_found -> ["/usr/local/share/coq";"/usr/share/coq"])
6967 @ (match Coq_config.datadir with |None -> [] |Some datadir -> [datadir])
7068
8381 match l with
8482 | [] -> raise Not_found
8583 | p :: tl ->
86 if Sys.file_exists (Filename.concat p f)
84 if Sys.file_exists (p//f)
8785 then p
8886 else which tl f
8987
107105 then Coq_config.camllib
108106 else
109107 let camlbin = camlbin () in
110 let com = (Filename.concat camlbin "ocamlc") ^ " -where" in
108 let com = (camlbin//"ocamlc") ^ " -where" in
111109 let _,res = System.run_command (fun x -> x) (fun _ -> ()) com in
112110 Util.strip res
113111
116114 if !Flags.boot then Coq_config.camlp4bin else
117115 try guess_camlp4bin () with e when e <> Sys.Break ->
118116 let cb = camlbin () in
119 if Sys.file_exists (Filename.concat cb (exe Coq_config.camlp4)) then cb
117 if Sys.file_exists (cb//(exe Coq_config.camlp4)) then cb
120118 else Coq_config.camlp4bin
121119
122120 let camlp4lib () =
124122 then Coq_config.camlp4lib
125123 else
126124 let camlp4bin = camlp4bin () in
127 let com = (Filename.concat camlp4bin Coq_config.camlp4) ^ " -where" in
125 let com = (camlp4bin//Coq_config.camlp4) ^ " -where" in
128126 let ex,res = System.run_command (fun x -> x) (fun _ -> ()) com in
129127 match ex with
130128 |Unix.WEXITED 0 -> Util.strip res
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
3131 | ConstrMayEvalArgType -> <:expr< Genarg.rawwit_constr_may_eval >>
3232 | QuantHypArgType -> <:expr< Genarg.rawwit_quant_hyp >>
3333 | RedExprArgType -> <:expr< Genarg.rawwit_red_expr >>
34 | OpenConstrArgType b -> <:expr< Genarg.rawwit_open_constr_gen $mlexpr_of_bool b$ >>
34 | OpenConstrArgType (b1,b2) -> <:expr< Genarg.rawwit_open_constr_gen ($mlexpr_of_bool b1$,$mlexpr_of_bool b2$) >>
3535 | ConstrWithBindingsArgType -> <:expr< Genarg.rawwit_constr_with_bindings >>
3636 | BindingsArgType -> <:expr< Genarg.rawwit_bindings >>
3737 | List0ArgType t -> <:expr< Genarg.wit_list0 $make_rawwit loc t$ >>
6161 | ConstrArgType -> <:expr< Genarg.globwit_constr >>
6262 | ConstrMayEvalArgType -> <:expr< Genarg.globwit_constr_may_eval >>
6363 | RedExprArgType -> <:expr< Genarg.globwit_red_expr >>
64 | OpenConstrArgType b -> <:expr< Genarg.globwit_open_constr_gen $mlexpr_of_bool b$ >>
64 | OpenConstrArgType (b1,b2) -> <:expr< Genarg.globwit_open_constr_gen ($mlexpr_of_bool b1$,$mlexpr_of_bool b2$) >>
6565 | ConstrWithBindingsArgType -> <:expr< Genarg.globwit_constr_with_bindings >>
6666 | BindingsArgType -> <:expr< Genarg.globwit_bindings >>
6767 | List0ArgType t -> <:expr< Genarg.wit_list0 $make_globwit loc t$ >>
9191 | ConstrArgType -> <:expr< Genarg.wit_constr >>
9292 | ConstrMayEvalArgType -> <:expr< Genarg.wit_constr_may_eval >>
9393 | RedExprArgType -> <:expr< Genarg.wit_red_expr >>
94 | OpenConstrArgType b -> <:expr< Genarg.wit_open_constr_gen $mlexpr_of_bool b$ >>
94 | OpenConstrArgType (b1,b2) -> <:expr< Genarg.wit_open_constr_gen ($mlexpr_of_bool b1$,$mlexpr_of_bool b2$) >>
9595 | ConstrWithBindingsArgType -> <:expr< Genarg.wit_constr_with_bindings >>
9696 | BindingsArgType -> <:expr< Genarg.wit_bindings >>
9797 | List0ArgType t -> <:expr< Genarg.wit_list0 $make_wit loc t$ >>
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
197197
198198 GEXTEND Gram
199199 GLOBAL: simple_tactic constr_with_bindings quantified_hypothesis
200 bindings red_expr int_or_var open_constr casted_open_constr
200 bindings red_expr int_or_var open_constr casted_open_constr open_constr_wTC
201201 simple_intropattern;
202202
203203 int_or_var:
216216 | id = METAIDENT -> MetaId (loc,id) ] ]
217217 ;
218218 open_constr:
219 [ [ c = constr -> ((),c) ] ]
220 ;
221 open_constr_wTC:
219222 [ [ c = constr -> ((),c) ] ]
220223 ;
221224 casted_open_constr:
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
413413 VernacRequireFrom (export, None, filename)
414414 | IDENT "Import"; qidl = LIST1 global -> VernacImport (false,qidl)
415415 | IDENT "Export"; qidl = LIST1 global -> VernacImport (true,qidl)
416 | IDENT "Include"; e = module_expr_inl; l = LIST0 ext_module_expr ->
416 | IDENT "Include"; e = module_type_inl; l = LIST0 ext_module_expr ->
417417 VernacInclude(e::l)
418418 | IDENT "Include"; "Type"; e = module_type_inl; l = LIST0 ext_module_type ->
419419 Flags.if_verbose
841841 -> PrintCoercionPaths (s,t)
842842 | IDENT "Canonical"; IDENT "Projections" -> PrintCanonicalConversions
843843 | IDENT "Tables" -> PrintTables
844 | IDENT "Options" -> PrintTables (* A Synonymous to Tables *)
844845 | IDENT "Hint" -> PrintHintGoal
845846 | IDENT "Hint"; qid = smart_global -> PrintHint qid
846847 | IDENT "Hint"; "*" -> PrintHintDb
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
379379 (* Entries that can be refered via the string -> Gram.entry table *)
380380 (* Typically for tactic user extensions *)
381381 let open_constr =
382 make_gen_entry utactic (rawwit_open_constr_gen false) "open_constr"
382 make_gen_entry utactic (rawwit_open_constr_gen (false,false)) "open_constr"
383383 let casted_open_constr =
384 make_gen_entry utactic (rawwit_open_constr_gen true) "casted_open_constr"
384 make_gen_entry utactic (rawwit_open_constr_gen (true,false)) "casted_open_constr"
385 let open_constr_wTC =
386 make_gen_entry utactic (rawwit_open_constr_gen (false,true)) "open_constr_wTC"
385387 let constr_with_bindings =
386388 make_gen_entry utactic rawwit_constr_with_bindings "constr_with_bindings"
387389 let bindings =
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
210210 sig
211211 open Glob_term
212212 val open_constr : open_constr_expr Gram.entry
213 val open_constr_wTC : open_constr_expr Gram.entry
213214 val casted_open_constr : open_constr_expr Gram.entry
214215 val constr_with_bindings : constr_expr with_bindings Gram.entry
215216 val bindings : constr_expr bindings Gram.entry
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
148148 | RedExprArgType ->
149149 pr_red_expr (prc,prlc,pr_or_by_notation prref,prpat)
150150 (out_gen rawwit_red_expr x)
151 | OpenConstrArgType b -> prc (snd (out_gen (rawwit_open_constr_gen b) x))
151 | OpenConstrArgType (b1,b2) -> prc (snd (out_gen (rawwit_open_constr_gen (b1,b2)) x))
152152 | ConstrWithBindingsArgType ->
153153 pr_with_bindings prc prlc (out_gen rawwit_constr_with_bindings x)
154154 | BindingsArgType ->
194194 pr_red_expr
195195 (prc,prlc,pr_or_var (pr_and_short_name pr_evaluable_reference),prpat)
196196 (out_gen globwit_red_expr x)
197 | OpenConstrArgType b -> prc (snd (out_gen (globwit_open_constr_gen b) x))
197 | OpenConstrArgType (b1,b2) -> prc (snd (out_gen (globwit_open_constr_gen (b1,b2)) x))
198198 | ConstrWithBindingsArgType ->
199199 pr_with_bindings prc prlc (out_gen globwit_constr_with_bindings x)
200200 | BindingsArgType ->
235235 | RedExprArgType ->
236236 pr_red_expr (prc,prlc,pr_evaluable_reference,prpat)
237237 (out_gen wit_red_expr x)
238 | OpenConstrArgType b -> prc (snd (out_gen (wit_open_constr_gen b) x))
238 | OpenConstrArgType (b1,b2) -> prc (snd (out_gen (wit_open_constr_gen (b1,b2)) x))
239239 | ConstrWithBindingsArgType ->
240240 let (c,b) = (out_gen wit_constr_with_bindings x).Evd.it in
241241 pr_with_bindings prc prlc (c,b)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
204204 | None -> (mt ())
205205 | Some c ->
206206 (* Force evaluation *)
207 let pb = pr_lconstr_env env c in
207 let pb = pr_lconstr_core true env c in
208208 let pb = if isCast c then surround pb else pb in
209209 (str" := " ++ pb ++ cut () ) in
210 let pt = pr_ltype_env env typ in
210 let pt = pr_ltype_core true env typ in
211211 let ptyp = (str" : " ++ pt) in
212212 (pr_id id ++ hov 0 (pbody ++ ptyp))
213213
216216 | None -> mt ()
217217 | Some c ->
218218 (* Force evaluation *)
219 let pb = pr_lconstr_env env c in
219 let pb = pr_lconstr_core true env c in
220220 let pb = if isCast c then surround pb else pb in
221221 (str":=" ++ spc () ++ pb ++ spc ()) in
222 let ptyp = pr_ltype_env env typ in
222 let ptyp = pr_ltype_core true env typ in
223223 match na with
224224 | Anonymous -> hov 0 (str"<>" ++ spc () ++ pbody ++ str":" ++ spc () ++ ptyp)
225225 | Name id -> hov 0 (pr_id id ++ spc () ++ pbody ++ str":" ++ spc () ++ ptyp)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
188188 | Genarg.VarArgType -> <:expr< Genarg.VarArgType >>
189189 | Genarg.StringArgType -> <:expr< Genarg.StringArgType >>
190190 | Genarg.QuantHypArgType -> <:expr< Genarg.QuantHypArgType >>
191 | Genarg.OpenConstrArgType b -> <:expr< Genarg.OpenConstrArgType $mlexpr_of_bool b$ >>
191 | Genarg.OpenConstrArgType (b1,b2) -> <:expr< Genarg.OpenConstrArgType ($mlexpr_of_bool b1$, $mlexpr_of_bool b2$) >>
192192 | Genarg.ConstrWithBindingsArgType -> <:expr< Genarg.ConstrWithBindingsArgType >>
193193 | Genarg.BindingsArgType -> <:expr< Genarg.BindingsArgType >>
194194 | Genarg.RedExprArgType -> <:expr< Genarg.RedExprArgType >>
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
187187
188188
189189 let find_rectype env c =
190 let (t, l) = decompose_app (Reduction.whd_betadeltaiota env c) in
190 let (t, l) = decompose_app (Reduction.whd_betaiotazeta c) in
191191 match kind_of_term t with
192192 | Ind ind -> (t, l)
193193 | Construct _ -> (t,l)
575575 onLastHypId (fun heq_id -> tclTHENLIST [
576576 (* Then the new hypothesis *)
577577 tclMAP introduction_no_check dyn_infos.rec_hyps;
578 (* observe_tac "after_introduction" *)(fun g' ->
578 observe_tac "after_introduction" (fun g' ->
579579 (* We get infos on the equations introduced*)
580580 let new_term_value_eq = pf_type_of g' (mkVar heq_id) in
581581 (* compute the new value of the body *)
602602 }
603603 in
604604 clean_goal_with_heq ptes_infos continue_tac new_infos g'
605 )])
605 )])
606606 ]
607607 g
608608
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
7171 raise reraise;;
7272
7373
74 let observe_tac s tac g =
75 if do_observe ()
76 then do_observe_tac (str s) tac g
74
75 let observe_tac_msg s tac g =
76 if do_observe ()
77 then do_observe_tac s tac g
7778 else tac g
79
80 let observe_tac s tac g =
81 observe_tac_msg (str s) tac g
7882
7983 (* [nf_zeta] $\zeta$-normalization of a term *)
8084 let nf_zeta =
286290 branches
287291 in
288292 (* before building the full intro pattern for the principle *)
289 let pat = Some (dummy_loc,Genarg.IntroOrAndPattern intro_pats) in
290293 let eq_ind = Coqlib.build_coq_eq () in
291294 let eq_construct = mkConstruct((destInd eq_ind),1) in
292295 (* The next to referencies will be used to find out which constructor to apply in each branch *)
293296 let ind_number = ref 0
294297 and min_constr_number = ref 0 in
295298 (* The tactic to prove the ith branch of the principle *)
299 let this_branche_ids empty add i =
300 List.fold_right
301 (fun (_,pat) acc ->
302 match pat with
303 | Genarg.IntroIdentifier id -> add id acc
304 | _ -> anomaly "Not an identifier"
305 )
306 (List.nth intro_pats (pred i))
307 empty
308 in
296309 let prove_branche i g =
297310 (* We get the identifiers of this branch *)
298 let this_branche_ids =
299 List.fold_right
300 (fun (_,pat) acc ->
301 match pat with
302 | Genarg.IntroIdentifier id -> Idset.add id acc
303 | _ -> anomaly "Not an identifier"
304 )
305 (List.nth intro_pats (pred i))
306 Idset.empty
307 in
308311 (* and get the real args of the branch by unfolding the defined constant *)
309312 let pre_args,pre_tac =
310313 List.fold_right
311314 (fun (id,b,t) (pre_args,pre_tac) ->
312 if Idset.mem id this_branche_ids
315 if Idset.mem id (this_branche_ids Idset.empty Idset.add i)
313316 then
314317 match b with
315 | None -> (id::pre_args,pre_tac)
318 | None ->
319 (id::pre_args,pre_tac)
316320 | Some b ->
317321 (pre_args,
318322 tclTHEN (h_reduce (Glob_term.Unfold([Glob_term.all_occurrences_expr,EvalVarRef id])) allHyps) pre_tac
346350 when
347351 (eq_constr eq eq_ind) &&
348352 array_exists (eq_constr graph') graphs_constr ->
349 ((mkApp(mkVar hid,[|args.(2);(mkApp(eq_construct,[|args.(0);args.(2)|]))|]))
350 ::args.(2)::acc)
353 ((mkApp(mkVar hid,[|args.(2);(mkApp(eq_construct,[|args.(0);args.(2)|]))|]))
354 ::args.(2)::acc)
351355 | _ -> mkVar hid :: acc
352356 end
353357 | _ -> mkVar hid :: acc
389393 | [res;hres] -> res,hres
390394 | _ -> assert false
391395 in
392 observe (str "constructor := " ++ Printer.pr_lconstr_env (pf_env g) app_constructor);
396 observe_tac_msg (str "constructor := " ++ Printer.pr_lconstr_env (pf_env g) app_constructor)
393397 (
394398 tclTHENSEQ
395399 [
454458 fun g ->
455459 observe
456460 (pr_constr_with_binding (Printer.pr_lconstr_env (pf_env g)) (mkVar principle_id,bindings));
457 functional_induction false (applist(funs_constr.(i),List.map mkVar args_names))
458 (Some (mkVar principle_id,bindings))
459 pat g
461 h_apply false false [dummy_loc,(mkVar principle_id,bindings)] g
460462 ))
461 (fun i g -> observe_tac ("proving branche "^string_of_int i) (prove_branche i) g )
463 (fun i g -> observe_tac ("proving branche "^string_of_int i)
464 (tclTHEN (tclMAP h_intro (this_branche_ids [] (fun a b -> a::b) i)) (prove_branche i)) g )
462465 ]
463466 g
464467
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
912912 | App(_,_) ->
913913 let (f,_) = decompose_app t in
914914 eq_constr f (well_founded ())
915 | _ -> assert false
915 | _ -> false
916916 in
917917 let compare t1 t2 =
918918 let b1,b2= is_well_founded t1,is_well_founded t2 in
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (* -*- coding: utf-8 -*- *)
11 (************************************************************************)
22 (* v * The Coq Proof Assistant / The Coq Development Team *)
3 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
3 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
44 (* \VV/ **************************************************************)
55 (* // * This file is distributed under the terms of the *)
66 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (* -*- coding: utf-8 -*- *)
11 (************************************************************************)
22 (* v * The Coq Proof Assistant / The Coq Development Team *)
3 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
3 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
44 (* \VV/ **************************************************************)
55 (* // * This file is distributed under the terms of the *)
66 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
274274 *)
275275
276276 let logic_dir = ["Coq";"Logic";"Decidable"]
277 let coq_modules =
278 init_modules @
279 [logic_dir] @ arith_modules @ zarith_base_modules @
280 [ ["Coq";"Lists";"List"];
277
278 let mic_modules =
279 [
280 ["Coq";"Lists";"List"];
281281 ["ZMicromega"];
282282 ["Tauto"];
283283 ["RingMicromega"];
292292 ["Coq";"Reals" ; "Rpow_def"];
293293 ["LRing_normalise"]]
294294
295 let coq_modules =
296 init_modules @
297 [logic_dir] @ arith_modules @ zarith_base_modules @ mic_modules
298
295299 let bin_module = [["Coq";"Numbers";"BinNums"]]
296300
297301 let r_modules =
311315 let bin_constant = gen_constant_in_modules "ZMicromega" bin_module
312316 let r_constant = gen_constant_in_modules "ZMicromega" r_modules
313317 let z_constant = gen_constant_in_modules "ZMicromega" z_modules
314 (* let constant = gen_constant_in_modules "Omicron" coq_modules *)
318 let m_constant = gen_constant_in_modules "ZMicromega" mic_modules
315319
316320 let coq_and = lazy (init_constant "and")
317321 let coq_or = lazy (init_constant "or")
353357 let coq_Qmake = lazy (constant "Qmake")
354358
355359 let coq_Rcst = lazy (constant "Rcst")
356 let coq_C0 = lazy (constant "C0")
357 let coq_C1 = lazy (constant "C1")
358 let coq_CQ = lazy (constant "CQ")
359 let coq_CZ = lazy (constant "CZ")
360 let coq_CPlus = lazy (constant "CPlus")
361 let coq_CMinus = lazy (constant "CMinus")
362 let coq_CMult = lazy (constant "CMult")
363 let coq_CInv = lazy (constant "CInv")
364 let coq_COpp = lazy (constant "COpp")
360 let coq_C0 = lazy (m_constant "C0")
361 let coq_C1 = lazy (m_constant "C1")
362 let coq_CQ = lazy (m_constant "CQ")
363 let coq_CZ = lazy (m_constant "CZ")
364 let coq_CPlus = lazy (m_constant "CPlus")
365 let coq_CMinus = lazy (m_constant "CMinus")
366 let coq_CMult = lazy (m_constant "CMult")
367 let coq_CInv = lazy (m_constant "CInv")
368 let coq_COpp = lazy (m_constant "COpp")
365369
366370
367371 let coq_R0 = lazy (constant "R0")
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
4848 let display_action_flag = ref false
4949 let old_style_flag = ref false
5050
51 (* Should we reset all variable labels between two runs of omega ? *)
52
53 let reset_flag = ref false
54
55 (* Historical version of Coq do not perform such resets, and this
56 implies that omega is slightly non-deterministic: successive runs of
57 omega on the same problem may lead to distinct proof-terms.
58 At the very least, these terms will differ on the inner
59 variable names, but they could even be non-convertible :
60 the OmegaSolver relies on Hashtbl.iter, it can hence find a different
61 solution when variable indices differ.
62
63 Starting from Coq 8.4pl4, omega may be made stable via the option
64 [Set Stable Omega]. In the 8.4 branch, this option is unset by default
65 for compatibility. In Coq >= 8.5, this option is set by default.
66 *)
67
5168 let read f () = !f
5269 let write f x = f:=x
5370
8097 optread = read old_style_flag;
8198 optwrite = write old_style_flag }
8299
100 let _ =
101 declare_bool_option
102 { optsync = true;
103 optdepr = false;
104 optname = "Omega automatic reset of generated names";
105 optkey = ["Stable";"Omega"];
106 optread = read reset_flag;
107 optwrite = write reset_flag }
83108
84109 let all_time = timing "Omega "
85110 let solver_time = timing "Solver "
88113 let simpl_time = timing "Simpl "
89114 let generalize_time = timing "Generalize"
90115
116 let intref, reset_all_references =
117 let refs = ref [] in
118 (fun n -> let r = ref n in refs := (r,n) :: !refs; r),
119 (fun () -> List.iter (fun (r,n) -> r:=n) !refs)
120
91121 let new_identifier =
92 let cpt = ref 0 in
122 let cpt = intref 0 in
93123 (fun () -> let s = "Omega" ^ string_of_int !cpt in incr cpt; id_of_string s)
94124
95125 let new_identifier_state =
96 let cpt = ref 0 in
126 let cpt = intref 0 in
97127 (fun () -> let s = make_ident "State" (Some !cpt) in incr cpt; s)
98128
99129 let new_identifier_var =
100 let cpt = ref 0 in
130 let cpt = intref 0 in
101131 (fun () -> let s = "Zvar" ^ string_of_int !cpt in incr cpt; id_of_string s)
102132
103133 let new_id =
104 let cpt = ref 0 in fun () -> incr cpt; !cpt
134 let cpt = intref 0 in fun () -> incr cpt; !cpt
105135
106136 let new_var_num =
107 let cpt = ref 1000 in (fun () -> incr cpt; !cpt)
137 let cpt = intref 1000 in (fun () -> incr cpt; !cpt)
108138
109139 let new_var =
110 let cpt = ref 0 in fun () -> incr cpt; Nameops.make_ident "WW" (Some !cpt)
140 let cpt = intref 0 in fun () -> incr cpt; Nameops.make_ident "WW" (Some !cpt)
111141
112142 let display_var i = Printf.sprintf "X%d" i
113143
114 let intern_id,unintern_id =
144 let intern_id,unintern_id,reset_intern_tables =
115145 let cpt = ref 0 in
116146 let table = Hashtbl.create 7 and co_table = Hashtbl.create 7 in
117147 (fun (name : identifier) ->
123153 (fun idx ->
124154 try Hashtbl.find co_table idx with Not_found ->
125155 let v = new_var () in
126 Hashtbl.add table v idx; Hashtbl.add co_table idx v; v)
156 Hashtbl.add table v idx; Hashtbl.add co_table idx v; v),
157 (fun () -> cpt := 0; Hashtbl.clear table)
127158
128159 let mk_then = tclTHENLIST
129160
140171 in
141172 loop
142173
143 let tag_hypothesis,tag_of_hyp, hyp_of_tag =
174 let tag_hypothesis,tag_of_hyp, hyp_of_tag, clear_tags =
144175 let l = ref ([]:(identifier * int) list) in
145176 (fun h id -> l := (h,id):: !l),
146177 (fun h -> try List.assoc h !l with Not_found -> failwith "tag_hypothesis"),
147 (fun h -> try rev_assoc h !l with Not_found -> failwith "tag_hypothesis")
148
149 let hide_constr,find_constr,clear_tables,dump_tables =
178 (fun h -> try rev_assoc h !l with Not_found -> failwith "tag_hypothesis"),
179 (fun () -> l := [])
180
181 let hide_constr,find_constr,clear_constr_tables,dump_tables =
150182 let l = ref ([]:(constr * (identifier * identifier * bool)) list) in
151183 (fun h id eg b -> l := (h,(id,eg,b)):: !l),
152184 (fun h -> try list_assoc_f eq_constr h !l with Not_found -> failwith "find_contr"),
153185 (fun () -> l := []),
154186 (fun () -> !l)
187
188 let reset_all () =
189 if !reset_flag then begin
190 reset_all_references ();
191 reset_intern_tables ();
192 clear_tags ();
193 clear_constr_tables ()
194 end
155195
156196 (* Lazy evaluation is used for Coq constants, because this code
157197 is evaluated before the compiled modules are loaded.
13871427 tclTHEN (tclTRY (clear [id])) (intro_using id)
13881428
13891429 let coq_omega gl =
1390 clear_tables ();
1430 clear_constr_tables ();
13911431 let tactic_normalisation, system =
13921432 List.fold_left (destructure_omega gl) ([],[]) (pf_hyps_types gl) in
13931433 let prelude,sys =
18131853
18141854 let omega_solver gl =
18151855 Coqlib.check_required_library ["Coq";"omega";"Omega"];
1856 reset_all ();
18161857 let result = destructure_goal gl in
18171858 (* if !display_time_flag then begin text_time ();
18181859 flush Pervasives.stdout end; *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
410410 let _ = e_cumul pb.env pb.isevars indt typ in
411411 current
412412 else
413 (evd_comb2 (Coercion.inh_conv_coerce_to dummy_loc pb.env)
413 (evd_comb2 (Coercion.inh_conv_coerce_to true dummy_loc pb.env)
414414 pb.isevars (make_judge current typ) (mk_tycon_type indt)).uj_val in
415415 let sigma = !(pb.isevars) in
416416 let typ = IsInd (indt,find_rectype pb.env sigma indt) in
14771477 let inh_conv_coerce_to_tycon loc env isevars j tycon =
14781478 match tycon with
14791479 | Some p ->
1480 let (evd',j) = Coercion.inh_conv_coerce_to loc env !isevars j p in
1480 let (evd',j) = Coercion.inh_conv_coerce_to true loc env !isevars j p in
14811481 isevars := evd';
14821482 j
14831483 | None -> j
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
357357 (hj,typ_cl) p)
358358 with e when Errors.noncritical e -> anomaly "apply_coercion"
359359
360 let inh_app_fun env isevars j =
360 let inh_app_fun _ env isevars j =
361361 let isevars = ref isevars in
362362 let t = hnf env !isevars j.uj_type in
363363 match kind_of_term t with
480480 | Some (init, cur) ->
481481 (evd, cj)
482482
483 let inh_conv_coerce_to = inh_conv_coerce_to_gen false
484 let inh_conv_coerce_rigid_to = inh_conv_coerce_to_gen true
483 let inh_conv_coerce_to _ = inh_conv_coerce_to_gen false
484 let inh_conv_coerce_rigid_to _ = inh_conv_coerce_to_gen true
485485
486486 let inh_conv_coerces_to loc env isevars t ((abs, t') as _tycon) =
487487 let nabsinit, nabs =
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
6262 (* str " in environment: " ++ my_print_env env); *)
6363
6464 let interp env isevars c tycon =
65 let j = pretype tycon env isevars ([],[]) c in
65 let j = pretype true tycon env isevars ([],[]) c in
6666 let _ = isevars := Evarutil.nf_evar_map !isevars in
6767 let evd = consider_remaining_unif_problems env !isevars in
6868 (* let unevd = undefined_evars evd in *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
8686 done
8787
8888 (* coerce to tycon if any *)
89 let inh_conv_coerce_to_tycon loc env evdref j = function
89 let inh_conv_coerce_to_tycon resolve_tc loc env evdref j = function
9090 | None -> j
91 | Some t -> evd_comb2 (Coercion.inh_conv_coerce_to loc env) evdref j t
91 | Some t -> evd_comb2 (Coercion.inh_conv_coerce_to resolve_tc loc env) evdref j t
9292
9393 let push_rels vars env = List.fold_right push_rel vars env
9494
187187 (* [pretype tycon env evdref lvar lmeta cstr] attempts to type [cstr] *)
188188 (* in environment [env], with existential variables [( evdref)] and *)
189189 (* the type constraint tycon *)
190 let rec pretype (tycon : type_constraint) env evdref lvar c =
190 let rec pretype resolve_tc (tycon : type_constraint) env evdref lvar c =
191191 (* let _ = try Subtac_utils.trace (str "pretype " ++ Subtac_utils.my_print_glob_constr env c ++ *)
192192 (* str " with tycon " ++ Evarutil.pr_tycon env tycon) *)
193193 (* with _ -> () *)
194194 (* in *)
195 let pretype = pretype resolve_tc in
196 let pretype_type = pretype_type resolve_tc in
197 let inh_conv_coerce_to_tycon = inh_conv_coerce_to_tycon resolve_tc in
195198 match c with
196199 | GRef (loc,ref) ->
197200 inh_conv_coerce_to_tycon loc env evdref
334337 | [] -> resj
335338 | c::rest ->
336339 let argloc = loc_of_glob_constr c in
337 let resj = evd_comb1 (Coercion.inh_app_fun env) evdref resj in
340 let resj = evd_comb1 (Coercion.inh_app_fun resolve_tc env) evdref resj in
338341 let resty = whd_betadeltaiota env !evdref resj.uj_type in
339342 match kind_of_term resty with
340343 | Prod (na,c1,c2) ->
556559 inh_conv_coerce_to_tycon loc env evdref cj tycon
557560
558561 (* [pretype_type valcon env evdref lvar c] coerces [c] into a type *)
559 and pretype_type valcon env evdref lvar = function
562 and pretype_type resolve_tc valcon env evdref lvar = function
560563 | GHole loc ->
561564 (match valcon with
562565 | Some v ->
576579 { utj_val = e_new_evar evdref env ~src:loc (mkSort s);
577580 utj_type = s})
578581 | c ->
579 let j = pretype empty_tycon env evdref lvar c in
582 let j = pretype resolve_tc empty_tycon env evdref lvar c in
580583 let loc = loc_of_glob_constr c in
581584 let tj = evd_comb1 (Coercion.inh_coerce_to_sort loc env) evdref j in
582585 match valcon with
591594 let c' = match kind with
592595 | OfType exptyp ->
593596 let tycon = match exptyp with None -> empty_tycon | Some t -> mk_tycon t in
594 (pretype tycon env evdref lvar c).uj_val
597 (pretype resolve_classes tycon env evdref lvar c).uj_val
595598 | IsType ->
596 (pretype_type empty_valcon env evdref lvar c).utj_val
599 (pretype_type resolve_classes empty_valcon env evdref lvar c).utj_val
597600 in
598601 if resolve_classes then
599602 (try
615618
616619 let understand_judgment sigma env c =
617620 let evdref = ref (create_evar_defs sigma) in
618 let j = pretype empty_tycon env evdref ([],[]) c in
621 let j = pretype true empty_tycon env evdref ([],[]) c in
619622 let evd = consider_remaining_unif_problems env !evdref in
620623 let j = j_nf_evar evd j in
621624 check_evars env sigma evd (mkCast(j.uj_val,DEFAULTcast, j.uj_type));
622625 j
623626
624627 let understand_judgment_tcc evdref env c =
625 let j = pretype empty_tycon env evdref ([],[]) c in
628 let j = pretype true empty_tycon env evdref ([],[]) c in
626629 j_nf_evar !evdref j
627630
628631 (* Raw calls to the unsafe inference machine: boolean says if we must
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
390390 let _ = e_cumul pb.env pb.evdref indt typ in
391391 current
392392 else
393 (evd_comb2 (Coercion.inh_conv_coerce_to dummy_loc pb.env)
393 (evd_comb2 (Coercion.inh_conv_coerce_to true dummy_loc pb.env)
394394 pb.evdref (make_judge current typ) (mk_tycon_type indt)).uj_val in
395395 let sigma = !(pb.evdref) in
396396 (current,try_find_ind pb.env sigma indt names))
16791679 let inh_conv_coerce_to_tycon loc env evdref j tycon =
16801680 match tycon with
16811681 | Some p ->
1682 let (evd',j) = Coercion.inh_conv_coerce_to loc env !evdref j p in
1682 let (evd',j) = Coercion.inh_conv_coerce_to true loc env !evdref j p in
16831683 evdref := evd';
16841684 j
16851685 | None -> j
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
3535 inserts a coercion into [j], if needed, in such a way it gets as
3636 type a product; it returns [j] if no coercion is applicable *)
3737 val inh_app_fun :
38 env -> evar_map -> unsafe_judgment -> evar_map * unsafe_judgment
38 bool -> env -> evar_map -> unsafe_judgment -> evar_map * unsafe_judgment
3939
4040 (* [inh_coerce_to_sort env evd j] coerces [j] to a type; i.e. it
4141 inserts a coercion into [j], if needed, in such a way it gets as
5656 (* [inh_conv_coerce_to loc env evd j t] coerces [j] to an object of type
5757 [t]; i.e. it inserts a coercion into [j], if needed, in such a way [t] and
5858 [j.uj_type] are convertible; it fails if no coercion is applicable *)
59 val inh_conv_coerce_to : loc ->
59 val inh_conv_coerce_to : bool -> loc ->
6060 env -> evar_map -> unsafe_judgment -> type_constraint_type -> evar_map * unsafe_judgment
6161
62 val inh_conv_coerce_rigid_to : loc ->
62 val inh_conv_coerce_rigid_to : bool -> loc ->
6363 env -> evar_map -> unsafe_judgment -> type_constraint_type -> evar_map * unsafe_judgment
6464
6565 (* [inh_conv_coerces_to loc env evd t t'] checks if an object of type [t]
139139 lookup_path_to_fun_from env evd j.uj_type in
140140 (evd,apply_coercion env evd p j t)
141141
142 let inh_app_fun env evd j =
142 let inh_app_fun resolve_tc env evd j =
143143 try inh_app_fun env evd j
144 with Not_found ->
144 with
145 | Not_found when not resolve_tc -> (evd, j)
146 | Not_found ->
145147 try inh_app_fun env (saturate_evd env evd) j
146148 with Not_found -> (evd, j)
147149
220222 | _ -> raise NoCoercion
221223
222224 (* Look for cj' obtained from cj by inserting coercions, s.t. cj'.typ = t *)
223 let inh_conv_coerce_to_gen rigidonly loc env evd cj (n, t) =
225 let inh_conv_coerce_to_gen resolve_tc rigidonly loc env evd cj (n, t) =
224226 match n with
225227 None ->
226228 let (evd', val') =
227229 try
228230 inh_conv_coerce_to_fail loc env evd rigidonly (Some cj.uj_val) cj.uj_type t
229 with NoCoercion ->
231 with
232 | NoCoercion when not resolve_tc -> error_actual_type_loc loc env evd cj t
233 | NoCoercion ->
230234 let evd = saturate_evd env evd in
231235 try
232236 inh_conv_coerce_to_fail loc env evd rigidonly (Some cj.uj_val) cj.uj_type t
237241 (evd',{ uj_val = val'; uj_type = t })
238242 | Some (init, cur) -> (evd, cj)
239243
240 let inh_conv_coerce_to = inh_conv_coerce_to_gen false
241 let inh_conv_coerce_rigid_to = inh_conv_coerce_to_gen true
244 let inh_conv_coerce_to resolve_tc = inh_conv_coerce_to_gen resolve_tc false
245 let inh_conv_coerce_rigid_to resolve_tc = inh_conv_coerce_to_gen resolve_tc true
242246
243247
244248 let inh_conv_coerces_to loc env (evd : evar_map) t (abs, t') =
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
1717 module type S = sig
1818 (** {6 Coercions. } *)
1919
20 (** [inh_app_fun env isevars j] coerces [j] to a function; i.e. it
20 (** [inh_app_fun resolve_tc env isevars j] coerces [j] to a function; i.e. it
2121 inserts a coercion into [j], if needed, in such a way it gets as
22 type a product; it returns [j] if no coercion is applicable *)
22 type a product; it returns [j] if no coercion is applicable.
23 resolve_tc=false disables resolving type classes (as the last
24 resort before failing) *)
2325 val inh_app_fun :
24 env -> evar_map -> unsafe_judgment -> evar_map * unsafe_judgment
26 bool -> env -> evar_map -> unsafe_judgment -> evar_map * unsafe_judgment
2527
2628 (** [inh_coerce_to_sort env isevars j] coerces [j] to a type; i.e. it
2729 inserts a coercion into [j], if needed, in such a way it gets as
3941 val inh_coerce_to_prod : loc ->
4042 env -> evar_map -> type_constraint_type -> evar_map * type_constraint_type
4143
42 (** [inh_conv_coerce_to loc env isevars j t] coerces [j] to an object of type
44 (** [inh_conv_coerce_to resolve_tc loc env isevars j t] coerces [j] to an object of type
4345 [t]; i.e. it inserts a coercion into [j], if needed, in such a way [t] and
44 [j.uj_type] are convertible; it fails if no coercion is applicable *)
45 val inh_conv_coerce_to : loc ->
46 [j.uj_type] are convertible; it fails if no coercion is applicable.
47 resolve_tc=false disables resolving type classes (as the last
48 resort before failing) *)
49 val inh_conv_coerce_to : bool -> loc ->
4650 env -> evar_map -> unsafe_judgment -> type_constraint_type -> evar_map * unsafe_judgment
4751
48 val inh_conv_coerce_rigid_to : loc ->
52 val inh_conv_coerce_rigid_to : bool -> loc ->
4953 env -> evar_map -> unsafe_judgment -> type_constraint_type -> evar_map * unsafe_judgment
5054
5155 (** [inh_conv_coerces_to loc env isevars t t'] checks if an object of type [t]
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
1818 open Evarutil
1919 open Libnames
2020 open Evd
21
22 let debug_unification = ref (false)
23 let _ = Goptions.declare_bool_option {
24 Goptions.optsync = true; Goptions.optdepr = false;
25 Goptions.optname =
26 "Print states sended to Evarconv unification";
27 Goptions.optkey = ["Debug";"Unification"];
28 Goptions.optread = (fun () -> !debug_unification);
29 Goptions.optwrite = (fun a -> debug_unification:=a);
30 }
31
2132
2233 type flex_kind_of_term =
2334 | Rigid of constr
209220 in
210221
211222 (* Evar must be undefined since we have flushed evars *)
223 let () = if !debug_unification then
224 let open Pp in
225 let pr_state (tm,l) =
226 h 0 (Termops.print_constr tm ++ str "|" ++ cut ()
227 ++ prlist_with_sep pr_semicolon
228 (fun x -> hov 1 (Termops.print_constr x)) l) in
229 pp (v 0 (pr_state appr1 ++ cut () ++ pr_state appr2 ++ cut ()) ++ fnl ()) in
212230 match (flex_kind_of_term term1 l1, flex_kind_of_term term2 l2) with
213231 | Flexible (sp1,al1 as ev1), Flexible (sp2,al2 as ev2) ->
214232 let f1 i =
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
11791179 * such that "hyps' |- ?e : T"
11801180 *)
11811181
1182 let filter_candidates evd evk filter candidates =
1182 let filter_effective_candidates evi filter candidates =
1183 match filter with
1184 | None -> candidates
1185 | Some filter ->
1186 let ids = List.map pi1 (list_filter_with filter (evar_context evi)) in
1187 List.filter (fun a -> list_subset (Idset.elements (collect_vars a)) ids)
1188 candidates
1189
1190 let filter_candidates evd evk filter candidates_update =
11831191 let evi = Evd.find_undefined evd evk in
1184 let candidates = match candidates with
1192 let candidates = match candidates_update with
11851193 | None -> evi.evar_candidates
1186 | Some _ -> candidates in
1187 match candidates,filter with
1188 | None,_ | _, None -> candidates
1189 | Some l, Some filter ->
1190 let ids = List.map pi1 (list_filter_with filter (evar_context evi)) in
1191 Some (List.filter (fun a ->
1192 list_subset (Idset.elements (collect_vars a)) ids) l)
1194 | Some _ -> candidates_update
1195 in
1196 match candidates with
1197 | None -> None
1198 | Some l ->
1199 let l' = filter_effective_candidates evi filter l in
1200 if List.length l = List.length l' && candidates_update = None then
1201 None
1202 else
1203 Some l'
11931204
11941205 let closure_of_filter evd evk filter =
11951206 let evi = Evd.find_undefined evd evk in
15291540
15301541 exception NotInvertibleUsingOurAlgorithm of constr
15311542 exception NotEnoughInformationToProgress of (identifier * evar_projection) list
1543 exception NotEnoughInformationEvarEvar of constr
15321544 exception OccurCheckIn of evar_map * constr
15331545
15341546 let rec invert_definition conv_algo choose env evd (evk,argsv as ev) rhs =
16071619 with
16081620 | EvarSolvedOnTheFly (evd,t) -> evdref:=evd; imitate envk t
16091621 | CannotProject filter' ->
1610 assert !progress;
1622 if not !progress then
1623 raise (NotEnoughInformationEvarEvar t);
16111624 (* Make the virtual left evar real *)
16121625 let ty = get_type_of env' !evdref t in
16131626 let (evd,evar'',ev'') =
17131726 with
17141727 | NotEnoughInformationToProgress sols ->
17151728 postpone_non_unique_projection env evd ev sols rhs
1729 | NotEnoughInformationEvarEvar t ->
1730 add_conv_pb (Reduction.CONV,env,mkEvar ev,t) evd
17161731 | NotInvertibleUsingOurAlgorithm t ->
17171732 error_not_clean env evd evk t (evar_source evk evd)
17181733 | OccurCheckIn (evd,rhs) ->
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
358358 PatCstr (loc,cstr,args,na)
359359 | _ -> raise Not_found
360360
361 let rec cases_pattern_of_glob_constr na = function
362 | GVar (loc,id) when na<>Anonymous ->
363 (* Unable to manage the presence of both an alias and a variable *)
364 raise Not_found
365 | GVar (loc,id) -> PatVar (loc,Name id)
366 | GHole (loc,_) -> PatVar (loc,na)
367 | GRef (loc,ConstructRef cstr) ->
368 PatCstr (loc,cstr,[],na)
369 | GApp (loc,GRef (_,ConstructRef (ind,_ as cstr)),args) ->
370 let mib,_ = Global.lookup_inductive ind in
371 let nparams = mib.Declarations.mind_nparams in
372 if nparams > List.length args then
373 user_err_loc (loc,"",Pp.str "Invalid notation for pattern.");
374 let params,args = list_chop nparams args in
375 List.iter (function GHole _ -> ()
376 | _ -> user_err_loc (loc,"",Pp.str"Invalid notation for pattern."))
377 params;
378 let args = List.map (cases_pattern_of_glob_constr Anonymous) args in
379 PatCstr (loc,cstr,args,na)
380 | _ -> raise Not_found
381
382361 (* Turn a closed cases pattern into a glob_constr *)
383362 let rec glob_constr_of_closed_cases_pattern_aux = function
384363 | PatCstr (loc,cstr,[],Anonymous) ->
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
199199 * Unused outside, but useful for debugging
200200 *)
201201 val pretype :
202 type_constraint -> env -> evar_map ref ->
202 bool -> type_constraint -> env -> evar_map ref ->
203203 ltac_var_map -> glob_constr -> unsafe_judgment
204204
205205 val pretype_type :
206 val_constraint -> env -> evar_map ref ->
206 bool -> val_constraint -> env -> evar_map ref ->
207207 ltac_var_map -> glob_constr -> unsafe_type_judgment
208208
209209 val pretype_gen :
258258 done
259259
260260 (* coerce to tycon if any *)
261 let inh_conv_coerce_to_tycon loc env evdref j = function
261 let inh_conv_coerce_to_tycon resolve_tc loc env evdref j = function
262262 | None -> j
263 | Some t -> evd_comb2 (Coercion.inh_conv_coerce_to loc env) evdref j t
263 | Some t -> evd_comb2 (Coercion.inh_conv_coerce_to resolve_tc loc env) evdref j t
264264
265265 let push_rels vars env = List.fold_right push_rel vars env
266266
296296 let c = substl subst c in
297297 { uj_val = c; uj_type = protected_get_type_of env sigma c }
298298 with Not_found ->
299 (* Check if [id] is a section or goal variable *)
300 try
301 let (_,_,typ) = lookup_named id env in
302 { uj_val = mkVar id; uj_type = typ }
303 with Not_found ->
304 (* [id] not found, build nice error message if [id] yet known from ltac *)
299 (* if [id] an ltac variable not bound to a term *)
300 (* build a nice error message *)
305301 try
306302 match List.assoc id unbndltacvars with
307303 | None -> user_err_loc (loc,"",
308304 str "Variable " ++ pr_id id ++ str " should be bound to a term.")
309305 | Some id0 -> Pretype_errors.error_var_not_found_loc loc id0
306 with Not_found ->
307 (* Check if [id] is a section or goal variable *)
308 try
309 let (_,_,typ) = lookup_named id env in
310 { uj_val = mkVar id; uj_type = typ }
310311 with Not_found ->
311312 (* [id] not found, standard error message *)
312313 error_var_not_found_loc loc id
345346 (* [pretype tycon env evdref lvar lmeta cstr] attempts to type [cstr] *)
346347 (* in environment [env], with existential variables [evdref] and *)
347348 (* the type constraint tycon *)
348 let rec pretype (tycon : type_constraint) env evdref lvar = function
349 let rec pretype resolve_tc (tycon : type_constraint) env evdref lvar t =
350 let pretype = pretype resolve_tc in
351 let pretype_type = pretype_type resolve_tc in
352 let inh_conv_coerce_to_tycon = inh_conv_coerce_to_tycon resolve_tc in
353 match t with
349354 | GRef (loc,ref) ->
350355 inh_conv_coerce_to_tycon loc env evdref
351356 (pretype_ref loc evdref env ref)
458463 | [] -> resj
459464 | c::rest ->
460465 let argloc = loc_of_glob_constr c in
461 let resj = evd_comb1 (Coercion.inh_app_fun env) evdref resj in
466 let resj = evd_comb1 (Coercion.inh_app_fun resolve_tc env) evdref resj in
462467 let resty = whd_betadeltaiota env !evdref resj.uj_type in
463468 match kind_of_term resty with
464469 | Prod (na,c1,c2) ->
695700 in inh_conv_coerce_to_tycon loc env evdref cj tycon
696701
697702 (* [pretype_type valcon env evdref lvar c] coerces [c] into a type *)
698 and pretype_type valcon env evdref lvar = function
703 and pretype_type resolve_tc valcon env evdref lvar = function
699704 | GHole loc ->
700705 (match valcon with
701706 | Some v ->
715720 { utj_val = e_new_evar evdref env ~src:loc (mkSort s);
716721 utj_type = s})
717722 | c ->
718 let j = pretype empty_tycon env evdref lvar c in
723 let j = pretype resolve_tc empty_tycon env evdref lvar c in
719724 let loc = loc_of_glob_constr c in
720725 let tj = evd_comb1 (Coercion.inh_coerce_to_sort loc env) evdref j in
721726 match valcon with
730735 let c' = match kind with
731736 | OfType exptyp ->
732737 let tycon = match exptyp with None -> empty_tycon | Some t -> mk_tycon t in
733 (pretype tycon env evdref lvar c).uj_val
738 (pretype resolve_classes tycon env evdref lvar c).uj_val
734739 | IsType ->
735 (pretype_type empty_valcon env evdref lvar c).utj_val
740 (pretype_type resolve_classes empty_valcon env evdref lvar c).utj_val
736741 in
737742 resolve_evars env evdref fail_evar resolve_classes;
738743 let c = if expand_evar then nf_evar !evdref c' else c' in
746751
747752 let understand_judgment sigma env c =
748753 let evdref = ref sigma in
749 let j = pretype empty_tycon env evdref ([],[]) c in
754 let j = pretype true empty_tycon env evdref ([],[]) c in
750755 resolve_evars env evdref true true;
751756 let j = j_nf_evar !evdref j in
752757 check_evars env sigma !evdref (mkCast(j.uj_val,DEFAULTcast, j.uj_type));
753758 j
754759
755760 let understand_judgment_tcc evdref env c =
756 let j = pretype empty_tycon env evdref ([],[]) c in
761 let j = pretype true empty_tycon env evdref ([],[]) c in
757762 resolve_evars env evdref false true;
758763 j_nf_evar !evdref j
759764
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
9292 (**/**)
9393 (** Internal of Pretyping... *)
9494 val pretype :
95 type_constraint -> env -> evar_map ref ->
95 bool -> type_constraint -> env -> evar_map ref ->
9696 ltac_var_map -> glob_constr -> unsafe_judgment
9797
9898 val pretype_type :
99 val_constraint -> env -> evar_map ref ->
99 bool -> val_constraint -> env -> evar_map ref ->
100100 ltac_var_map -> glob_constr -> unsafe_type_judgment
101101
102102 val pretype_gen :
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
793793
794794 let try_to_coerce env evd c cty tycon =
795795 let j = make_judge c cty in
796 let (evd',j') = inh_conv_coerce_rigid_to dummy_loc env evd j tycon in
796 let (evd',j') = inh_conv_coerce_rigid_to true dummy_loc env evd j tycon in
797797 let evd' = Evarconv.consider_remaining_unif_problems env evd' in
798798 let evd' = Evd.map_metas_fvalue (nf_evar evd') evd' in
799799 (evd',j'.uj_val)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
154154 let j = Environ.make_judge t my_type in
155155 let tycon = Evarutil.mk_tycon_type typ in
156156 let (new_defs,j') =
157 Coercion.Default.inh_conv_coerce_to (Util.dummy_loc) env !rdefs j tycon
157 Coercion.Default.inh_conv_coerce_to true (Util.dummy_loc) env !rdefs j tycon
158158 in
159159 rdefs := new_defs;
160160 j'.Environ.uj_val
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
5656 let rec make_compilation_args = function
5757 | [] -> []
5858 | file :: fl ->
59 let dirname = Filename.dirname file in
60 let basename = Filename.basename file in
61 let modulename =
62 if Filename.check_suffix basename ".v" then
63 Filename.chop_suffix basename ".v"
59 let name_no_suffix =
60 if Filename.check_suffix file ".v" then
61 Filename.chop_suffix file ".v"
6462 else
65 basename
63 file
6664 in
65 let modulename = Filename.basename name_no_suffix in
6766 check_module_name modulename;
68 let file = Filename.concat dirname modulename in
6967 (if !verbose then "-compile-verbose" else "-compile")
70 :: file :: (make_compilation_args fl)
68 :: name_no_suffix :: (make_compilation_args fl)
7169
7270 (* compilation of files [files] with command [command] and args [args] *)
7371
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
893893 (str"apply " ++ pr_constr c ++ str" ; trivial")
894894 | Unfold_nth c -> (str"unfold " ++ pr_evaluable_reference c)
895895 | Extern tac ->
896 (str "(*external*) " ++ Pptactic.pr_glob_tactic (Global.env()) tac)
896 let env =
897 try
898 let (_, env) = Pfedit.get_current_goal_context () in
899 env
900 with e when Errors.noncritical e -> Global.env ()
901 in
902 (str "(*external*) " ++ Pptactic.pr_glob_tactic env tac)
897903
898904 let pr_hint (id, v) =
899905 (pr_autotactic v.code ++ str"(level " ++ int v.pri ++ str", id " ++ int id ++ str ")" ++ spc ())
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
2929 [ "admit" ] -> [ admit_as_an_axiom ]
3030 END
3131
32 let replace_in_clause_maybe_by (sigma1,c1) c2 in_hyp tac =
32
33
34 let classes_dirpath =
35 make_dirpath (List.map id_of_string ["Classes";"Coq"])
36
37 let init_setoid () =
38 if Libnames.is_dirpath_prefix_of classes_dirpath (Lib.cwd ()) then ()
39 else Coqlib.check_required_library ["Coq";"Setoids";"Setoid"]
40
41
42 let occurrences_of occs =
43 let loccs = match occs with
44 | n::_ as nl when n < 0 -> (false,List.map (fun n -> ArgArg (abs n)) nl)
45 | nl ->
46 if List.exists (fun n -> n < 0) nl then
47 error "Illegal negative occurrence number.";
48 (true, List.map (fun n -> (ArgArg n)) nl)
49 in
50 init_setoid ();
51 {onhyps = Some []; concl_occs =loccs}
52
53 let replace_in_clause_maybe_by (sigma1,c1) c2 cl tac =
3354 Refiner.tclWITHHOLES false
34 (replace_in_clause_maybe_by c1 c2 (glob_in_arg_hyp_to_clause in_hyp))
55 (replace_in_clause_maybe_by c1 c2 cl)
3556 sigma1
3657 (Option.map Tacinterp.eval_tactic tac)
3758
4364
4465 TACTIC EXTEND replace
4566 ["replace" open_constr(c1) "with" constr(c2) in_arg_hyp(in_hyp) by_arg_tac(tac) ]
46 -> [ replace_in_clause_maybe_by c1 c2 in_hyp tac ]
47 END
67 -> [ replace_in_clause_maybe_by c1 c2 (glob_in_arg_hyp_to_clause in_hyp) tac ]
68 END
69
70 TACTIC EXTEND replace_at
71 ["replace" open_constr(c1) "with" constr(c2) "at" occurrences(occs) by_arg_tac(tac) ]
72 -> [ replace_in_clause_maybe_by c1 c2 (occurrences_of occs) tac ]
73 END
74
4875
4976 TACTIC EXTEND replace_term_left
5077 [ "replace" "->" open_constr(c) in_arg_hyp(in_hyp) ]
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
276276
277277 open Tacinterp
278278 let decompose_applied_relation_expr env sigma flags (is, (c,l)) left2right =
279 let sigma, cbl = Tacinterp.interp_open_constr_with_bindings is env sigma (c,l) in
279 let sigma, cbl = Tacinterp.interp_open_constr_with_bindings false is env sigma (c,l) in
280280 decompose_applied_relation env sigma flags (Some (is, (c,l))) cbl left2right
281281
282282 let rewrite_db = "rewrite"
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
966966 (intern_quantified_hypothesis ist (out_gen rawwit_quant_hyp x))
967967 | RedExprArgType ->
968968 in_gen globwit_red_expr (intern_red_expr ist (out_gen rawwit_red_expr x))
969 | OpenConstrArgType b ->
970 in_gen (globwit_open_constr_gen b)
971 ((),intern_constr ist (snd (out_gen (rawwit_open_constr_gen b) x)))
969 | OpenConstrArgType (b1,b2) ->
970 in_gen (globwit_open_constr_gen (b1,b2))
971 ((),intern_constr ist (snd (out_gen (rawwit_open_constr_gen (b1,b2)) x)))
972972 | ConstrWithBindingsArgType ->
973973 in_gen globwit_constr_with_bindings
974974 (intern_constr_with_bindings ist (out_gen rawwit_constr_with_bindings x))
12891289 let interp_open_constr_gen kind ist =
12901290 interp_gen kind ist false true false false
12911291
1292 let interp_open_constr ccl ist =
1293 interp_gen (OfType ccl) ist false true false (ccl<>None)
1292 (* wTC is for retrocompatibility: TC resolution started only if needed *)
1293 let interp_open_constr ccl wTC ist e s t =
1294 try interp_gen (OfType ccl) ist false true false (ccl<>None) e s t
1295 with ex when Pretype_errors.precatchable_exception ex && ccl = None && wTC ->
1296 interp_gen (OfType ccl) ist false true false true e s t
12941297
12951298 let interp_pure_open_constr ist =
12961299 interp_gen (OfType None) ist false false false false
13321335
13331336 let interp_open_constr_list =
13341337 interp_constr_in_compound_list (fun x -> x) (fun x -> x)
1335 (interp_open_constr None)
1338 (interp_open_constr None false)
13361339
13371340 let interp_auto_lemmas ist env sigma lems =
13381341 let local_sigma, lems = interp_open_constr_list ist env sigma lems in
15251528 with Not_found -> NamedHyp id
15261529
15271530 let interp_binding ist env sigma (loc,b,c) =
1528 let sigma, c = interp_open_constr None ist env sigma c in
1531 let sigma, c = interp_open_constr None false ist env sigma c in
15291532 sigma, (loc,interp_binding_name ist b,c)
15301533
15311534 let interp_bindings ist env sigma = function
15401543
15411544 let interp_constr_with_bindings ist env sigma (c,bl) =
15421545 let sigma, bl = interp_bindings ist env sigma bl in
1543 let sigma, c = interp_open_constr None ist env sigma c in
1546 let sigma, c = interp_open_constr None false ist env sigma c in
15441547 sigma, (c,bl)
15451548
1546 let interp_open_constr_with_bindings ist env sigma (c,bl) =
1549 let interp_open_constr_with_bindings wTC ist env sigma (c,bl) =
15471550 let sigma, bl = interp_bindings ist env sigma bl in
1548 let sigma, c = interp_open_constr None ist env sigma c in
1551 let sigma, c = interp_open_constr None wTC ist env sigma c in
15491552 sigma, (c, bl)
15501553
15511554 let loc_of_bindings = function
15531556 | ImplicitBindings l -> loc_of_glob_constr (fst (list_last l))
15541557 | ExplicitBindings l -> pi1 (list_last l)
15551558
1556 let interp_open_constr_with_bindings_loc ist env sigma ((c,_),bl as cb) =
1559 let interp_open_constr_with_bindings_loc wTC ist env sigma ((c,_),bl as cb) =
15571560 let loc1 = loc_of_glob_constr c in
15581561 let loc2 = loc_of_bindings bl in
15591562 let loc = if loc2 = dummy_loc then loc1 else join_loc loc1 loc2 in
1560 let sigma, cb = interp_open_constr_with_bindings ist env sigma cb in
1563 let sigma, cb = interp_open_constr_with_bindings wTC ist env sigma cb in
15611564 sigma, (loc,cb)
15621565
15631566 let interp_induction_arg ist gl arg =
17531756 let mk_constr_value ist gl c =
17541757 let (sigma,c_interp) = pf_interp_constr ist gl c in
17551758 sigma,VConstr ([],c_interp)
1756 let mk_open_constr_value ist gl c =
1757 let (sigma,c_interp) = pf_apply (interp_open_constr None ist) gl c in
1759 let mk_open_constr_value wTC ist gl c =
1760 let (sigma,c_interp) = pf_apply (interp_open_constr None wTC ist) gl c in
17581761 sigma,VConstr ([],c_interp)
17591762 let mk_hyp_value ist gl c = VConstr ([],mkVar (interp_hyp ist gl c))
17601763 let mk_int_or_var_value ist c = VInteger (interp_int_or_var ist c)
21312134 let (sigma,r_interp) = pf_interp_red_expr ist gl (out_gen globwit_red_expr x) in
21322135 evdref := sigma;
21332136 in_gen wit_red_expr r_interp
2134 | OpenConstrArgType casted ->
2135 in_gen (wit_open_constr_gen casted)
2136 (interp_open_constr (if casted then Some (pf_concl gl) else None)
2137 | OpenConstrArgType (casted,wTC) ->
2138 in_gen (wit_open_constr_gen (casted,wTC))
2139 (interp_open_constr (if casted then Some (pf_concl gl) else None) wTC
21372140 ist (pf_env gl) (project gl)
2138 (snd (out_gen (globwit_open_constr_gen casted) x)))
2141 (snd (out_gen (globwit_open_constr_gen (casted,wTC)) x)))
21392142 | ConstrWithBindingsArgType ->
21402143 in_gen wit_constr_with_bindings
21412144 (pack_sigma (interp_constr_with_bindings ist (pf_env gl) (project gl)
23312334 (h_vm_cast_no_check c_interp)
23322335 | TacApply (a,ev,cb,cl) ->
23332336 let sigma, l =
2334 list_fold_map (interp_open_constr_with_bindings_loc ist env) sigma cb
2337 list_fold_map (interp_open_constr_with_bindings_loc true ist env) sigma cb
23352338 in
23362339 let tac = match cl with
23372340 | None -> h_apply a ev
25462549 (* Equality and inversion *)
25472550 | TacRewrite (ev,l,cl,by) ->
25482551 let l = List.map (fun (b,m,c) ->
2549 let f env sigma = interp_open_constr_with_bindings ist env sigma c in
2552 let f env sigma = interp_open_constr_with_bindings false ist env sigma c in
25502553 (b,m,f)) l in
25512554 let cl = interp_clause ist gl cl in
25522555 Equality.general_multi_multi_rewrite ev l cl
26092612 let (sigma,v) = mk_constr_value ist gl (out_gen globwit_constr x) in
26102613 evdref := sigma;
26112614 v
2612 | OpenConstrArgType false ->
2613 let (sigma,v) = mk_open_constr_value ist gl (snd (out_gen globwit_open_constr x)) in
2615 | OpenConstrArgType (false,true) ->
2616 let (sigma,v) = mk_open_constr_value true ist gl (snd (out_gen globwit_open_constr_wTC x)) in
2617 evdref := sigma;
2618 v
2619 | OpenConstrArgType (false,false) ->
2620 let (sigma,v) = mk_open_constr_value false ist gl (snd (out_gen globwit_open_constr x)) in
26142621 evdref := sigma;
26152622 v
26162623 | ConstrMayEvalArgType ->
30123019 (out_gen globwit_quant_hyp x))
30133020 | RedExprArgType ->
30143021 in_gen globwit_red_expr (subst_redexp subst (out_gen globwit_red_expr x))
3015 | OpenConstrArgType b ->
3016 in_gen (globwit_open_constr_gen b)
3017 ((),subst_glob_constr subst (snd (out_gen (globwit_open_constr_gen b) x)))
3022 | OpenConstrArgType (b1,b2) ->
3023 in_gen (globwit_open_constr_gen (b1,b2))
3024 ((),subst_glob_constr subst (snd (out_gen (globwit_open_constr_gen (b1,b2)) x)))
30183025 | ConstrWithBindingsArgType ->
30193026 in_gen globwit_constr_with_bindings
30203027 (subst_glob_with_bindings subst (out_gen globwit_constr_with_bindings x))
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
133133
134134 val interp_bindings : interp_sign -> Environ.env -> Evd.evar_map -> glob_constr_and_expr Glob_term.bindings -> Evd.evar_map * constr Glob_term.bindings
135135
136 val interp_open_constr_with_bindings : interp_sign -> Environ.env -> Evd.evar_map ->
136 (* first arguments mean wTC (with type classes resolution) *)
137 val interp_open_constr_with_bindings : bool -> interp_sign -> Environ.env -> Evd.evar_map ->
137138 glob_constr_and_expr Glob_term.with_bindings -> Evd.evar_map * constr Glob_term.with_bindings
138139
139140 (** Initial call for interpretation *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
1010 x1 : nat
1111 x4 : nat
1212 x0 : nat
13 H : forall x x3 : nat, x + x1 = x4 + x3
13 H : forall x5 x6 : nat, x5 + x1 = x4 + x6
1414 ============================
1515 x + x1 = x4 + x0
1616 1 subgoal
5050 x1 : nat
5151 x4 : nat
5252 x0 : nat
53 H : forall x x3 : nat,
54 x + x1 = x4 + x3 ->
55 forall x0 x4 x5 S0 : nat, x0 + S0 = x4 + x5 + (S x + x1)
53 H : forall x5 x6 : nat,
54 x5 + x1 = x4 + x6 ->
55 forall x7 x8 x9 S : nat, x7 + S = x8 + x9 + (Datatypes.S x5 + x1)
5656 H0 : x + x1 = x4 + x0
5757 ============================
5858 forall x5 x6 x7 S : nat, x5 + S = x6 + x7 + Datatypes.S x
6363 x1 : nat
6464 x4 : nat
6565 x0 : nat
66 H : forall x x3 : nat,
67 x + x1 = x4 + x3 ->
68 forall x0 x4 x5 S0 : nat, x0 + S0 = x4 + x5 + (S x + x1)
66 H : forall x5 x6 : nat,
67 x5 + x1 = x4 + x6 ->
68 forall x7 x8 x9 S : nat, x7 + S = x8 + x9 + (Datatypes.S x5 + x1)
6969 H0 : x + x1 = x4 + x0
7070 x5 : nat
7171 x6 : nat
7777
7878 x3 : nat
7979 a : nat
80 H : a = 0 -> forall a : nat, a = 0
80 H : a = 0 -> forall a0 : nat, a0 = 0
8181 ============================
8282 a = 0
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (* -*- coding: utf-8 -*- *)
11 (************************************************************************)
22 (* v * The Coq Proof Assistant / The Coq Development Team *)
3 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
3 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
44 (* \VV/ **************************************************************)
55 (* // * This file is distributed under the terms of the *)
66 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (* -*- coding: utf-8 -*- *)
11 (************************************************************************)
22 (* v * The Coq Proof Assistant / The Coq Development Team *)
3 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
3 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
44 (* \VV/ **************************************************************)
55 (* // * This file is distributed under the terms of the *)
66 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (* -*- coding: utf-8 -*- *)
11 (************************************************************************)
22 (* v * The Coq Proof Assistant / The Coq Development Team *)
3 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
3 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
44 (* \VV/ **************************************************************)
55 (* // * This file is distributed under the terms of the *)
66 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (* -*- coding: utf-8 -*- *)
11 (************************************************************************)
22 (* v * The Coq Proof Assistant / The Coq Development Team *)
3 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
3 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
44 (* \VV/ **************************************************************)
55 (* // * This file is distributed under the terms of the *)
66 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (* -*- coding: utf-8 -*- *)
11 (************************************************************************)
22 (* v * The Coq Proof Assistant / The Coq Development Team *)
3 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
3 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
44 (* \VV/ **************************************************************)
55 (* // * This file is distributed under the terms of the *)
66 (* * GNU Lesser General Public License Version 2.1 *)
00 (* -*- coding: utf-8 -*- *)
11 (************************************************************************)
22 (* v * The Coq Proof Assistant / The Coq Development Team *)
3 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
3 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
44 (* \VV/ **************************************************************)
55 (* // * This file is distributed under the terms of the *)
66 (* * GNU Lesser General Public License Version 2.1 *)
00 (* -*- coding: utf-8 -*- *)
11 (************************************************************************)
22 (* v * The Coq Proof Assistant / The Coq Development Team *)
3 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
3 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
44 (* \VV/ **************************************************************)
55 (* // * This file is distributed under the terms of the *)
66 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (* -*- coding: utf-8 -*- *)
11 (************************************************************************)
22 (* v * The Coq Proof Assistant / The Coq Development Team *)
3 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
3 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
44 (* \VV/ **************************************************************)
55 (* // * This file is distributed under the terms of the *)
66 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (* -*- coding: utf-8 -*- *)
11 (************************************************************************)
22 (* v * The Coq Proof Assistant / The Coq Development Team *)
3 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
3 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
44 (* \VV/ **************************************************************)
55 (* // * This file is distributed under the terms of the *)
66 (* * GNU Lesser General Public License Version 2.1 *)
00 (* -*- coding: utf-8 -*- *)
11 (************************************************************************)
22 (* v * The Coq Proof Assistant / The Coq Development Team *)
3 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
3 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
44 (* \VV/ **************************************************************)
55 (* // * This file is distributed under the terms of the *)
66 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (* -*- coding: utf-8 -*- *)
11 (************************************************************************)
22 (* v * The Coq Proof Assistant / The Coq Development Team *)
3 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
3 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
44 (* \VV/ **************************************************************)
55 (* // * This file is distributed under the terms of the *)
66 (* * GNU Lesser General Public License Version 2.1 *)
00 (* -*- coding: utf-8 -*- *)
11 (************************************************************************)
22 (* v * The Coq Proof Assistant / The Coq Development Team *)
3 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
3 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
44 (* \VV/ **************************************************************)
55 (* // * This file is distributed under the terms of the *)
66 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (* -*- coding: utf-8 -*- *)
11 (************************************************************************)
22 (* v * The Coq Proof Assistant / The Coq Development Team *)
3 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
3 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
44 (* \VV/ **************************************************************)
55 (* // * This file is distributed under the terms of the *)
66 (* * GNU Lesser General Public License Version 2.1 *)
00 (* -*- coding: utf-8 -*- *)
11 (************************************************************************)
22 (* v * The Coq Proof Assistant / The Coq Development Team *)
3 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
3 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
44 (* \VV/ **************************************************************)
55 (* // * This file is distributed under the terms of the *)
66 (* * GNU Lesser General Public License Version 2.1 *)
00 (* -*- coding: utf-8 -*- *)
11 (************************************************************************)
22 (* v * The Coq Proof Assistant / The Coq Development Team *)
3 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
3 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
44 (* \VV/ **************************************************************)
55 (* // * This file is distributed under the terms of the *)
66 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
0 (************************************************************************)
1 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
3 (* \VV/ **************************************************************)
4 (* // * This file is distributed under the terms of the *)
5 (* * GNU Lesser General Public License Version 2.1 *)
6 (************************************************************************)
7
08 Require Import Fourier.
19 Require Import Rbase.
210 Require Import Rtrigo1.
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (* -*- coding: utf-8 -*- *)
11 (************************************************************************)
22 (* v * The Coq Proof Assistant / The Coq Development Team *)
3 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
3 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
44 (* \VV/ **************************************************************)
55 (* // * This file is distributed under the terms of the *)
66 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
0 (************************************************************************)
1 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
3 (* \VV/ **************************************************************)
4 (* // * This file is distributed under the terms of the *)
5 (* * GNU Lesser General Public License Version 2.1 *)
6 (************************************************************************)
7
08 Require Import Rbase.
19 Require Import Ranalysis_reg.
210 Require Import Rfunctions.
13441352 field ; apply Rgt_not_eq ; apply Rabs_pos_lt ; assumption.
13451353 assumption.
13461354 field. assumption.
1347 Qed.
1355 Qed.
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
0 (************************************************************************)
1 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
3 (* \VV/ **************************************************************)
4 (* // * This file is distributed under the terms of the *)
5 (* * GNU Lesser General Public License Version 2.1 *)
6 (************************************************************************)
7
08 Require Import Fourier.
19 Require Import Rbase.
210 Require Import PSeries_reg.
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (* -*- coding: utf-8 -*- *)
11 (************************************************************************)
22 (* v * The Coq Proof Assistant / The Coq Development Team *)
3 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
3 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
44 (* \VV/ **************************************************************)
55 (* // * This file is distributed under the terms of the *)
66 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (* -*- coding: utf-8 -*- *)
11 (************************************************************************)
22 (* v * The Coq Proof Assistant / The Coq Development Team *)
3 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
3 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
44 (* \VV/ **************************************************************)
55 (* // * This file is distributed under the terms of the *)
66 (* * GNU Lesser General Public License Version 2.1 *)
00 (* -*- coding: utf-8 -*- *)
11 (************************************************************************)
22 (* v * The Coq Proof Assistant / The Coq Development Team *)
3 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
3 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
44 (* \VV/ **************************************************************)
55 (* // * This file is distributed under the terms of the *)
66 (* * GNU Lesser General Public License Version 2.1 *)
00 (* -*- coding:utf-8 -*- *)
11 (************************************************************************)
22 (* v * The Coq Proof Assistant / The Coq Development Team *)
3 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
3 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
44 (* \VV/ **************************************************************)
55 (* // * This file is distributed under the terms of the *)
66 (* * GNU Lesser General Public License Version 2.1 *)
00 (* -*- coding:utf-8 -*- *)
11 (************************************************************************)
22 (* v * The Coq Proof Assistant / The Coq Development Team *)
3 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
3 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
44 (* \VV/ **************************************************************)
55 (* // * This file is distributed under the terms of the *)
66 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (* -*- coding: utf-8 -*- *)
11 (************************************************************************)
22 (* v * The Coq Proof Assistant / The Coq Development Team *)
3 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
3 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
44 (* \VV/ **************************************************************)
55 (* // * This file is distributed under the terms of the *)
66 (* * GNU Lesser General Public License Version 2.1 *)
00 (* -*- coding: utf-8 -*- *)
11 (************************************************************************)
22 (* v * The Coq Proof Assistant / The Coq Development Team *)
3 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
3 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
44 (* \VV/ **************************************************************)
55 (* // * This file is distributed under the terms of the *)
66 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (* -*- coding: utf-8 -*- *)
11 (************************************************************************)
22 (* v * The Coq Proof Assistant / The Coq Development Team *)
3 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
3 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
44 (* \VV/ **************************************************************)
55 (* // * This file is distributed under the terms of the *)
66 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (* -*- coding: utf-8 -*- *)
11 (************************************************************************)
22 (* v * The Coq Proof Assistant / The Coq Development Team *)
3 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
3 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
44 (* \VV/ **************************************************************)
55 (* // * This file is distributed under the terms of the *)
66 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (* -*- coding: utf-8 -*- *)
11 (************************************************************************)
22 (* v * The Coq Proof Assistant / The Coq Development Team *)
3 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
3 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
44 (* \VV/ **************************************************************)
55 (* // * This file is distributed under the terms of the *)
66 (* * GNU Lesser General Public License Version 2.1 *)
00 (* -*- coding: utf-8 -*- *)
11 (************************************************************************)
22 (* v * The Coq Proof Assistant / The Coq Development Team *)
3 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
3 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
44 (* \VV/ **************************************************************)
55 (* // * This file is distributed under the terms of the *)
66 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (* -*- coding: utf-8 -*- *)
11 (************************************************************************)
22 (* v * The Coq Proof Assistant / The Coq Development Team *)
3 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
3 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
44 (* \VV/ **************************************************************)
55 (* // * This file is distributed under the terms of the *)
66 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (* -*- coding: utf-8 -*- *)
11 (************************************************************************)
22 (* v * The Coq Proof Assistant / The Coq Development Team *)
3 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
3 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
44 (* \VV/ **************************************************************)
55 (* // * This file is distributed under the terms of the *)
66 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (* -*- coding: utf-8 -*- *)
11 (************************************************************************)
22 (* v * The Coq Proof Assistant / The Coq Development Team *)
3 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
3 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
44 (* \VV/ **************************************************************)
55 (* // * This file is distributed under the terms of the *)
66 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
99 open System
1010 open Toplevel
1111
12 let (/) = Filename.concat
12 let (/) s1 s2 = s1 ^ "/" ^ s2
1313
1414 let set_debug () = Flags.debug := true
1515
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)
00 (************************************************************************)
11 (* v * The Coq Proof Assistant / The Coq Development Team *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2012 *)
2 (* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
33 (* \VV/ **************************************************************)
44 (* // * This file is distributed under the terms of the *)
55 (* * GNU Lesser General Public License Version 2.1 *)