aboutsummaryrefslogtreecommitdiff
path: root/kernel
diff options
context:
space:
mode:
authorPierre Courtieu2016-04-15 16:45:14 +0200
committerPierre Courtieu2016-04-15 16:45:14 +0200
commitcaa1f67de10614984fa7e1c68aa8adf0ff90196a (patch)
tree3c265ac5e16851c5dc1ca917ddb03725e09ea0ff /kernel
parentbe824224cc76f729872e9d803fc64831b95aee94 (diff)
parent3b3d98acd58e91c960a2e11cd47ac19b2b34f86b (diff)
Merge remote-tracking branch 'OFFICIAL/trunk' into morefresh
Diffstat (limited to 'kernel')
-rw-r--r--kernel/byterun/coq_interp.c79
-rw-r--r--kernel/byterun/coq_memory.c16
-rw-r--r--kernel/byterun/coq_values.c1
-rw-r--r--kernel/byterun/coq_values.h10
-rw-r--r--kernel/byterun/int64_emul.h270
-rw-r--r--kernel/byterun/int64_native.h48
-rw-r--r--kernel/cbytecodes.ml63
-rw-r--r--kernel/cbytecodes.mli44
-rw-r--r--kernel/cbytegen.ml246
-rw-r--r--kernel/cbytegen.mli10
-rw-r--r--kernel/cemitcodes.ml43
-rw-r--r--kernel/cemitcodes.mli8
-rw-r--r--kernel/closure.ml40
-rw-r--r--kernel/closure.mli4
-rw-r--r--kernel/constr.ml72
-rw-r--r--kernel/constr.mli48
-rw-r--r--kernel/context.ml520
-rw-r--r--kernel/context.mli308
-rw-r--r--kernel/conv_oracle.ml17
-rw-r--r--kernel/conv_oracle.mli2
-rw-r--r--kernel/cooking.ml24
-rw-r--r--kernel/cooking.mli2
-rw-r--r--kernel/csymtable.ml65
-rw-r--r--kernel/csymtable.mli2
-rw-r--r--kernel/declarations.mli38
-rw-r--r--kernel/declareops.ml119
-rw-r--r--kernel/declareops.mli15
-rw-r--r--kernel/entries.mli31
-rw-r--r--kernel/environ.ml135
-rw-r--r--kernel/environ.mli55
-rw-r--r--kernel/esubst.ml2
-rw-r--r--kernel/esubst.mli2
-rw-r--r--kernel/evar.ml2
-rw-r--r--kernel/evar.mli2
-rw-r--r--kernel/fast_typeops.ml28
-rw-r--r--kernel/fast_typeops.mli2
-rw-r--r--kernel/indtypes.ml313
-rw-r--r--kernel/indtypes.mli6
-rw-r--r--kernel/inductive.ml113
-rw-r--r--kernel/inductive.mli11
-rw-r--r--kernel/kernel.mllib1
-rw-r--r--kernel/mod_subst.ml4
-rw-r--r--kernel/mod_subst.mli7
-rw-r--r--kernel/mod_typing.ml229
-rw-r--r--kernel/mod_typing.mli33
-rw-r--r--kernel/modops.ml19
-rw-r--r--kernel/modops.mli8
-rw-r--r--kernel/names.ml107
-rw-r--r--kernel/names.mli95
-rw-r--r--kernel/nativecode.ml28
-rw-r--r--kernel/nativecode.mli2
-rw-r--r--kernel/nativeconv.ml123
-rw-r--r--kernel/nativeconv.mli8
-rw-r--r--kernel/nativeinstr.mli2
-rw-r--r--kernel/nativelambda.ml41
-rw-r--r--kernel/nativelambda.mli2
-rw-r--r--kernel/nativelib.ml14
-rw-r--r--kernel/nativelib.mli2
-rw-r--r--kernel/nativelibrary.ml2
-rw-r--r--kernel/nativelibrary.mli2
-rw-r--r--kernel/nativevalues.ml31
-rw-r--r--kernel/nativevalues.mli2
-rw-r--r--kernel/opaqueproof.ml9
-rw-r--r--kernel/opaqueproof.mli8
-rw-r--r--kernel/pre_env.ml36
-rw-r--r--kernel/pre_env.mli20
-rw-r--r--kernel/primitives.ml2
-rw-r--r--kernel/primitives.mli2
-rw-r--r--kernel/reduction.ml220
-rw-r--r--kernel/reduction.mli84
-rw-r--r--kernel/retroknowledge.ml2
-rw-r--r--kernel/retroknowledge.mli2
-rw-r--r--kernel/safe_typing.ml228
-rw-r--r--kernel/safe_typing.mli53
-rw-r--r--kernel/sorts.ml4
-rw-r--r--kernel/sorts.mli2
-rw-r--r--kernel/subtyping.ml17
-rw-r--r--kernel/subtyping.mli2
-rw-r--r--kernel/term.ml219
-rw-r--r--kernel/term.mli97
-rw-r--r--kernel/term_typing.ml328
-rw-r--r--kernel/term_typing.mli43
-rw-r--r--kernel/type_errors.ml2
-rw-r--r--kernel/type_errors.mli2
-rw-r--r--kernel/typeops.ml70
-rw-r--r--kernel/typeops.mli7
-rw-r--r--kernel/uGraph.ml906
-rw-r--r--kernel/uGraph.mli63
-rw-r--r--kernel/univ.ml913
-rw-r--r--kernel/univ.mli65
-rw-r--r--kernel/vars.ml52
-rw-r--r--kernel/vars.mli96
-rw-r--r--kernel/vconv.ml174
-rw-r--r--kernel/vconv.mli12
-rw-r--r--kernel/vm.ml305
-rw-r--r--kernel/vm.mli27
96 files changed, 4339 insertions, 3311 deletions
diff --git a/kernel/byterun/coq_interp.c b/kernel/byterun/coq_interp.c
index 399fa843f1..bf383a33ac 100644
--- a/kernel/byterun/coq_interp.c
+++ b/kernel/byterun/coq_interp.c
@@ -22,18 +22,10 @@
#include "coq_memory.h"
#include "coq_values.h"
-/*spiwack : imports support functions for 64-bit integers */
-#include <caml/config.h>
-#ifdef ARCH_INT64_TYPE
-#include "int64_native.h"
-#else
-#include "int64_emul.h"
-#endif
-
/* spiwack: I append here a few macros for value/number manipulation */
#define uint32_of_value(val) (((uint32_t)val >> 1))
#define value_of_uint32(i) ((value)(((uint32_t)(i) << 1) | 1))
-#define UI64_of_uint32(lo) ((uint64_t)(I64_literal(0,(uint32_t)(lo))))
+#define UI64_of_uint32(lo) ((uint64_t)(lo))
#define UI64_of_value(val) (UI64_of_uint32(uint32_of_value(val)))
/* /spiwack */
@@ -341,6 +333,7 @@ value coq_interprete
/* Fallthrough */
Instruct(ENVACC){
print_instr("ENVACC");
+ print_int(*pc);
accu = Field(coq_env, *pc++);
Next;
}
@@ -371,6 +364,10 @@ value coq_interprete
sp[1] = (value)pc;
sp[2] = coq_env;
sp[3] = Val_long(coq_extra_args);
+ print_instr("call stack=");
+ print_lint(sp[1]);
+ print_lint(sp[2]);
+ print_lint(sp[3]);
pc = Code_val(accu);
coq_env = accu;
coq_extra_args = 0;
@@ -458,6 +455,7 @@ value coq_interprete
sp[0] = arg1;
sp[1] = arg2;
pc = Code_val(accu);
+ print_lint(accu);
coq_env = accu;
coq_extra_args += 1;
goto check_stacks;
@@ -481,11 +479,18 @@ value coq_interprete
print_instr("RETURN");
print_int(*pc);
sp += *pc++;
+ print_instr("stack=");
+ print_lint(sp[0]);
+ print_lint(sp[1]);
+ print_lint(sp[2]);
if (coq_extra_args > 0) {
+ print_instr("extra args > 0");
+ print_lint(coq_extra_args);
coq_extra_args--;
pc = Code_val(accu);
coq_env = accu;
} else {
+ print_instr("extra args = 0");
pc = (code_t)(sp[0]);
coq_env = sp[1];
coq_extra_args = Long_val(sp[2]);
@@ -585,7 +590,10 @@ value coq_interprete
Alloc_small(accu, 1 + nvars, Closure_tag);
Code_val(accu) = pc + *pc;
pc++;
- for (i = 0; i < nvars; i++) Field(accu, i + 1) = sp[i];
+ for (i = 0; i < nvars; i++) {
+ print_lint(sp[i]);
+ Field(accu, i + 1) = sp[i];
+ }
sp += nvars;
Next;
}
@@ -720,6 +728,7 @@ value coq_interprete
/* Fallthrough */
Instruct(GETGLOBAL){
print_instr("GETGLOBAL");
+ print_int(*pc);
accu = Field(coq_global_data, *pc);
pc++;
Next;
@@ -732,7 +741,7 @@ value coq_interprete
tag_t tag = *pc++;
mlsize_t i;
value block;
- print_instr("MAKEBLOCK");
+ print_instr("MAKEBLOCK, tag=");
Alloc_small(block, wosize, tag);
Field(block, 0) = accu;
for (i = 1; i < wosize; i++) Field(block, i) = *sp++;
@@ -743,7 +752,8 @@ value coq_interprete
tag_t tag = *pc++;
value block;
- print_instr("MAKEBLOCK1");
+ print_instr("MAKEBLOCK1, tag=");
+ print_int(tag);
Alloc_small(block, 1, tag);
Field(block, 0) = accu;
accu = block;
@@ -753,7 +763,8 @@ value coq_interprete
tag_t tag = *pc++;
value block;
- print_instr("MAKEBLOCK2");
+ print_instr("MAKEBLOCK2, tag=");
+ print_int(tag);
Alloc_small(block, 2, tag);
Field(block, 0) = accu;
Field(block, 1) = sp[0];
@@ -764,7 +775,8 @@ value coq_interprete
Instruct(MAKEBLOCK3) {
tag_t tag = *pc++;
value block;
- print_instr("MAKEBLOCK3");
+ print_instr("MAKEBLOCK3, tag=");
+ print_int(tag);
Alloc_small(block, 3, tag);
Field(block, 0) = accu;
Field(block, 1) = sp[0];
@@ -776,7 +788,8 @@ value coq_interprete
Instruct(MAKEBLOCK4) {
tag_t tag = *pc++;
value block;
- print_instr("MAKEBLOCK4");
+ print_instr("MAKEBLOCK4, tag=");
+ print_int(tag);
Alloc_small(block, 4, tag);
Field(block, 0) = accu;
Field(block, 1) = sp[0];
@@ -844,7 +857,6 @@ value coq_interprete
}
Instruct(SETFIELD1){
- int i, j, size, size_aux;
print_instr("SETFIELD1");
caml_modify(&Field(accu, 1),*sp);
sp++;
@@ -884,19 +896,17 @@ value coq_interprete
Instruct(PROJ){
print_instr("PROJ");
if (Is_accu (accu)) {
+ value block;
/* Skip over the index of projected field */
pc++;
- /* Save the argument on the stack */
- *--sp = accu;
/* Create atom */
- Alloc_small(accu, 2, ATOM_PROJ_TAG);
- Field(accu, 0) = Field(coq_global_data, *pc);
- Field(accu, 1) = sp[0];
- sp[0] = accu;
+ Alloc_small(block, 2, ATOM_PROJ_TAG);
+ Field(block, 0) = Field(coq_global_data, *pc);
+ Field(block, 1) = accu;
/* Create accumulator */
Alloc_small(accu, 2, Accu_tag);
Code_val(accu) = accumulate;
- Field(accu,1) = *sp++;
+ Field(accu, 1) = block;
} else {
accu = Field(accu, *pc++);
}
@@ -943,6 +953,7 @@ value coq_interprete
/* Fallthrough */
Instruct(CONSTINT) {
print_instr("CONSTINT");
+ print_int(*pc);
accu = Val_int(*pc);
pc++;
Next;
@@ -1110,7 +1121,6 @@ value coq_interprete
/* returns the sum plus one with a carry */
uint32_t s;
s = (uint32_t)accu + (uint32_t)*sp++ + 1;
- value block;
if( (uint32_t)s <= (uint32_t)accu ) {
/* carry */
Alloc_small(accu, 1, 2); /* ( _ , arity, tag ) */
@@ -1183,8 +1193,8 @@ value coq_interprete
print_instr("MULCINT31");
uint64_t p;
/*accu = 2v+1, *sp=2w+1 ==> p = 2v*w */
- p = I64_mul (UI64_of_value (accu), UI64_of_uint32 ((*sp++)^1));
- if ( I64_is_zero(p) ) {
+ p = UI64_of_value (accu) * UI64_of_uint32 ((*sp++)^1);
+ if (p == 0) {
accu = (value)1;
}
else {
@@ -1193,8 +1203,8 @@ value coq_interprete
of the non-constant constructor is then 1 */
Alloc_small(accu, 2, 1); /* ( _ , arity, tag ) */
/*unsigned shift*/
- Field(accu, 0) = (value)(I64_lsr(p,31)|1) ; /*higher part*/
- Field(accu, 1) = (value)(I64_to_int32(p)|1); /*lower part*/
+ Field(accu, 0) = (value)((p >> 31)|1) ; /*higher part*/
+ Field(accu, 1) = (value)((int32_t)p|1); /*lower part*/
}
Next;
}
@@ -1206,19 +1216,20 @@ value coq_interprete
int62 by the int31 */
uint64_t bigint;
bigint = UI64_of_value(accu);
- bigint = I64_or(I64_lsl(bigint, 31),UI64_of_value(*sp++));
+ bigint = (bigint << 31) | UI64_of_value(*sp++);
uint64_t divisor;
divisor = UI64_of_value(*sp++);
Alloc_small(accu, 2, 1); /* ( _ , arity, tag ) */
- if (I64_is_zero (divisor)) {
+ if (divisor == 0) {
Field(accu, 0) = 1; /* 2*0+1 */
Field(accu, 1) = 1; /* 2*0+1 */
}
else {
uint64_t quo, mod;
- I64_udivmod(bigint, divisor, &quo, &mod);
- Field(accu, 0) = value_of_uint32(I64_to_int32(quo));
- Field(accu, 1) = value_of_uint32(I64_to_int32(mod));
+ quo = bigint / divisor;
+ mod = bigint % divisor;
+ Field(accu, 0) = value_of_uint32((uint32_t)(quo));
+ Field(accu, 1) = value_of_uint32((uint32_t)(mod));
}
Next;
}
@@ -1271,7 +1282,7 @@ value coq_interprete
Instruct (COMPAREINT31) {
/* returns Eq if equal, Lt if accu is less than *sp, Gt otherwise */
- /* assumes Inudctive _ : _ := Eq | Lt | Gt */
+ /* assumes Inductive _ : _ := Eq | Lt | Gt */
print_instr("COMPAREINT31");
if ((uint32_t)accu == (uint32_t)*sp) {
accu = 1; /* 2*0+1 */
diff --git a/kernel/byterun/coq_memory.c b/kernel/byterun/coq_memory.c
index 8d03829ab0..c9bcdc32ff 100644
--- a/kernel/byterun/coq_memory.c
+++ b/kernel/byterun/coq_memory.c
@@ -26,7 +26,6 @@ asize_t coq_max_stack_size = Coq_max_stack_size;
value coq_global_data;
-int coq_all_transp;
value coq_atom_tbl;
int drawinstr;
@@ -104,7 +103,6 @@ static int coq_vm_initialized = 0;
value init_coq_vm(value unit) /* ML */
{
- int i;
if (coq_vm_initialized == 1) {
fprintf(stderr,"already open \n");fflush(stderr);}
else {
@@ -117,7 +115,6 @@ value init_coq_vm(value unit) /* ML */
init_coq_global_data(Coq_global_data_Size);
init_coq_atom_tbl(40);
/* Initialing the interpreter */
- coq_all_transp = 0;
init_coq_interpreter();
/* Some predefined pointer code */
@@ -137,7 +134,6 @@ void realloc_coq_stack(asize_t required_space)
{
asize_t size;
value * new_low, * new_high, * new_sp;
- value * p;
size = coq_stack_high - coq_stack_low;
do {
size *= 2;
@@ -207,18 +203,6 @@ value realloc_coq_atom_tbl(value size) /* ML */
return Val_unit;
}
-
-value coq_set_transp_value(value transp)
-{
- coq_all_transp = (transp == Val_true);
- return Val_unit;
-}
-
-value get_coq_transp_value(value unit)
-{
- return Val_bool(coq_all_transp);
-}
-
value coq_set_drawinstr(value unit)
{
drawinstr = 1;
diff --git a/kernel/byterun/coq_values.c b/kernel/byterun/coq_values.c
index 007f61b27c..528babebfc 100644
--- a/kernel/byterun/coq_values.c
+++ b/kernel/byterun/coq_values.c
@@ -21,7 +21,6 @@
value coq_kind_of_closure(value v) {
opcode_t * c;
- int res;
int is_app = 0;
c = Code_val(v);
if (Is_instruction(c, GRAB)) return Val_int(0);
diff --git a/kernel/byterun/coq_values.h b/kernel/byterun/coq_values.h
index 1590a2141d..bb0f0eb5e4 100644
--- a/kernel/byterun/coq_values.h
+++ b/kernel/byterun/coq_values.h
@@ -17,23 +17,17 @@
#define Default_tag 0
#define Accu_tag 0
-
-
#define ATOM_ID_TAG 0
-#define ATOM_IDDEF_TAG 1
-#define ATOM_INDUCTIVE_TAG 2
+#define ATOM_INDUCTIVE_TAG 1
+#define ATOM_TYPE_TAG 2
#define ATOM_PROJ_TAG 3
#define ATOM_FIX_TAG 4
#define ATOM_SWITCH_TAG 5
#define ATOM_COFIX_TAG 6
#define ATOM_COFIXEVALUATED_TAG 7
-
-
/* Les blocs accumulate */
#define Is_accu(v) (Is_block(v) && (Tag_val(v) == Accu_tag))
#define IS_EVALUATED_COFIX(v) (Is_accu(v) && Is_block(Field(v,1)) && (Tag_val(Field(v,1)) == ATOM_COFIXEVALUATED_TAG))
#endif /* _COQ_VALUES_ */
-
-
diff --git a/kernel/byterun/int64_emul.h b/kernel/byterun/int64_emul.h
deleted file mode 100644
index 86bee72edb..0000000000
--- a/kernel/byterun/int64_emul.h
+++ /dev/null
@@ -1,270 +0,0 @@
-/***********************************************************************/
-/* */
-/* Objective Caml */
-/* */
-/* Xavier Leroy, projet Cristal, INRIA Rocquencourt */
-/* */
-/* Copyright 2002 Institut National de Recherche en Informatique et */
-/* en Automatique. All rights reserved. This file is distributed */
-/* under the terms of the GNU Library General Public License, with */
-/* the special exception on linking described in file ../LICENSE. */
-/* */
-/***********************************************************************/
-
-/* Software emulation of 64-bit integer arithmetic, for C compilers
- that do not support it. */
-
-#ifndef CAML_INT64_EMUL_H
-#define CAML_INT64_EMUL_H
-
-#include <math.h>
-
-#ifdef ARCH_BIG_ENDIAN
-#define I64_literal(hi,lo) { hi, lo }
-#else
-#define I64_literal(hi,lo) { lo, hi }
-#endif
-
-/* Unsigned comparison */
-static int I64_ucompare(uint64 x, uint64 y)
-{
- if (x.h > y.h) return 1;
- if (x.h < y.h) return -1;
- if (x.l > y.l) return 1;
- if (x.l < y.l) return -1;
- return 0;
-}
-
-#define I64_ult(x, y) (I64_ucompare(x, y) < 0)
-
-/* Signed comparison */
-static int I64_compare(int64 x, int64 y)
-{
- if ((int32)x.h > (int32)y.h) return 1;
- if ((int32)x.h < (int32)y.h) return -1;
- if (x.l > y.l) return 1;
- if (x.l < y.l) return -1;
- return 0;
-}
-
-/* Negation */
-static int64 I64_neg(int64 x)
-{
- int64 res;
- res.l = -x.l;
- res.h = ~x.h;
- if (res.l == 0) res.h++;
- return res;
-}
-
-/* Addition */
-static int64 I64_add(int64 x, int64 y)
-{
- int64 res;
- res.l = x.l + y.l;
- res.h = x.h + y.h;
- if (res.l < x.l) res.h++;
- return res;
-}
-
-/* Subtraction */
-static int64 I64_sub(int64 x, int64 y)
-{
- int64 res;
- res.l = x.l - y.l;
- res.h = x.h - y.h;
- if (x.l < y.l) res.h--;
- return res;
-}
-
-/* Multiplication */
-static int64 I64_mul(int64 x, int64 y)
-{
- int64 res;
- uint32 prod00 = (x.l & 0xFFFF) * (y.l & 0xFFFF);
- uint32 prod10 = (x.l >> 16) * (y.l & 0xFFFF);
- uint32 prod01 = (x.l & 0xFFFF) * (y.l >> 16);
- uint32 prod11 = (x.l >> 16) * (y.l >> 16);
- res.l = prod00;
- res.h = prod11 + (prod01 >> 16) + (prod10 >> 16);
- prod01 = prod01 << 16; res.l += prod01; if (res.l < prod01) res.h++;
- prod10 = prod10 << 16; res.l += prod10; if (res.l < prod10) res.h++;
- res.h += x.l * y.h + x.h * y.l;
- return res;
-}
-
-#define I64_is_zero(x) (((x).l | (x).h) == 0)
-
-#define I64_is_negative(x) ((int32) (x).h < 0)
-
-/* Bitwise operations */
-static int64 I64_and(int64 x, int64 y)
-{
- int64 res;
- res.l = x.l & y.l;
- res.h = x.h & y.h;
- return res;
-}
-
-static int64 I64_or(int64 x, int64 y)
-{
- int64 res;
- res.l = x.l | y.l;
- res.h = x.h | y.h;
- return res;
-}
-
-static int64 I64_xor(int64 x, int64 y)
-{
- int64 res;
- res.l = x.l ^ y.l;
- res.h = x.h ^ y.h;
- return res;
-}
-
-/* Shifts */
-static int64 I64_lsl(int64 x, int s)
-{
- int64 res;
- s = s & 63;
- if (s == 0) return x;
- if (s < 32) {
- res.l = x.l << s;
- res.h = (x.h << s) | (x.l >> (32 - s));
- } else {
- res.l = 0;
- res.h = x.l << (s - 32);
- }
- return res;
-}
-
-static int64 I64_lsr(int64 x, int s)
-{
- int64 res;
- s = s & 63;
- if (s == 0) return x;
- if (s < 32) {
- res.l = (x.l >> s) | (x.h << (32 - s));
- res.h = x.h >> s;
- } else {
- res.l = x.h >> (s - 32);
- res.h = 0;
- }
- return res;
-}
-
-static int64 I64_asr(int64 x, int s)
-{
- int64 res;
- s = s & 63;
- if (s == 0) return x;
- if (s < 32) {
- res.l = (x.l >> s) | (x.h << (32 - s));
- res.h = (int32) x.h >> s;
- } else {
- res.l = (int32) x.h >> (s - 32);
- res.h = (int32) x.h >> 31;
- }
- return res;
-}
-
-/* Division and modulus */
-
-#define I64_SHL1(x) x.h = (x.h << 1) | (x.l >> 31); x.l <<= 1
-#define I64_SHR1(x) x.l = (x.l >> 1) | (x.h << 31); x.h >>= 1
-
-static void I64_udivmod(uint64 modulus, uint64 divisor,
- uint64 * quo, uint64 * mod)
-{
- int64 quotient, mask;
- int cmp;
-
- quotient.h = 0; quotient.l = 0;
- mask.h = 0; mask.l = 1;
- while ((int32) divisor.h >= 0) {
- cmp = I64_ucompare(divisor, modulus);
- I64_SHL1(divisor);
- I64_SHL1(mask);
- if (cmp >= 0) break;
- }
- while (mask.l | mask.h) {
- if (I64_ucompare(modulus, divisor) >= 0) {
- quotient.h |= mask.h; quotient.l |= mask.l;
- modulus = I64_sub(modulus, divisor);
- }
- I64_SHR1(mask);
- I64_SHR1(divisor);
- }
- *quo = quotient;
- *mod = modulus;
-}
-
-static int64 I64_div(int64 x, int64 y)
-{
- int64 q, r;
- int32 sign;
-
- sign = x.h ^ y.h;
- if ((int32) x.h < 0) x = I64_neg(x);
- if ((int32) y.h < 0) y = I64_neg(y);
- I64_udivmod(x, y, &q, &r);
- if (sign < 0) q = I64_neg(q);
- return q;
-}
-
-static int64 I64_mod(int64 x, int64 y)
-{
- int64 q, r;
- int32 sign;
-
- sign = x.h;
- if ((int32) x.h < 0) x = I64_neg(x);
- if ((int32) y.h < 0) y = I64_neg(y);
- I64_udivmod(x, y, &q, &r);
- if (sign < 0) r = I64_neg(r);
- return r;
-}
-
-/* Coercions */
-
-static int64 I64_of_int32(int32 x)
-{
- int64 res;
- res.l = x;
- res.h = x >> 31;
- return res;
-}
-
-#define I64_to_int32(x) ((int32) (x).l)
-
-/* Note: we assume sizeof(intnat) = 4 here, which is true otherwise
- autoconfiguration would have selected native 64-bit integers */
-#define I64_of_intnat I64_of_int32
-#define I64_to_intnat I64_to_int32
-
-static double I64_to_double(int64 x)
-{
- double res;
- int32 sign = x.h;
- if (sign < 0) x = I64_neg(x);
- res = ldexp((double) x.h, 32) + x.l;
- if (sign < 0) res = -res;
- return res;
-}
-
-static int64 I64_of_double(double f)
-{
- int64 res;
- double frac, integ;
- int neg;
-
- neg = (f < 0);
- f = fabs(f);
- frac = modf(ldexp(f, -32), &integ);
- res.h = (uint32) integ;
- res.l = (uint32) ldexp(frac, 32);
- if (neg) res = I64_neg(res);
- return res;
-}
-
-#endif /* CAML_INT64_EMUL_H */
diff --git a/kernel/byterun/int64_native.h b/kernel/byterun/int64_native.h
deleted file mode 100644
index 657d0a07e0..0000000000
--- a/kernel/byterun/int64_native.h
+++ /dev/null
@@ -1,48 +0,0 @@
-/***********************************************************************/
-/* */
-/* Objective Caml */
-/* */
-/* Xavier Leroy, projet Cristal, INRIA Rocquencourt */
-/* */
-/* Copyright 2002 Institut National de Recherche en Informatique et */
-/* en Automatique. All rights reserved. This file is distributed */
-/* under the terms of the GNU Library General Public License, with */
-/* the special exception on linking described in file ../LICENSE. */
-/* */
-/***********************************************************************/
-
-/* Wrapper macros around native 64-bit integer arithmetic,
- so that it has the same interface as the software emulation
- provided in int64_emul.h */
-
-#ifndef CAML_INT64_NATIVE_H
-#define CAML_INT64_NATIVE_H
-
-#define I64_literal(hi,lo) ((int64_t)(hi) << 32 | (lo))
-#define I64_compare(x,y) (((x) > (y)) - ((x) < (y)))
-#define I64_ult(x,y) ((uint64_t)(x) < (uint64_t)(y))
-#define I64_neg(x) (-(x))
-#define I64_add(x,y) ((x) + (y))
-#define I64_sub(x,y) ((x) - (y))
-#define I64_mul(x,y) ((x) * (y))
-#define I64_is_zero(x) ((x) == 0)
-#define I64_is_negative(x) ((x) < 0)
-#define I64_div(x,y) ((x) / (y))
-#define I64_mod(x,y) ((x) % (y))
-#define I64_udivmod(x,y,quo,rem) \
- (*(rem) = (uint64_t)(x) % (uint64_t)(y), \
- *(quo) = (uint64_t)(x) / (uint64_t)(y))
-#define I64_and(x,y) ((x) & (y))
-#define I64_or(x,y) ((x) | (y))
-#define I64_xor(x,y) ((x) ^ (y))
-#define I64_lsl(x,y) ((x) << (y))
-#define I64_asr(x,y) ((x) >> (y))
-#define I64_lsr(x,y) ((uint64_t)(x) >> (y))
-#define I64_to_intnat(x) ((intnat) (x))
-#define I64_of_intnat(x) ((intnat) (x))
-#define I64_to_int32(x) ((int32_t) (x))
-#define I64_of_int32(x) ((int64_t) (x))
-#define I64_to_double(x) ((double)(x))
-#define I64_of_double(x) ((int64_t)(x))
-
-#endif /* CAML_INT64_NATIVE_H */
diff --git a/kernel/cbytecodes.ml b/kernel/cbytecodes.ml
index 891d95378b..f9cf2691ee 100644
--- a/kernel/cbytecodes.ml
+++ b/kernel/cbytecodes.ml
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -17,23 +17,28 @@ open Term
type tag = int
-let id_tag = 0
-let iddef_tag = 1
-let ind_tag = 2
-let fix_tag = 3
-let switch_tag = 4
-let cofix_tag = 5
-let cofix_evaluated_tag = 6
+let accu_tag = 0
+
+let type_atom_tag = 2
+let max_atom_tag = 2
+let proj_tag = 3
+let fix_app_tag = 4
+let switch_tag = 5
+let cofix_tag = 6
+let cofix_evaluated_tag = 7
+
(* It would be great if OCaml exported this value,
So fixme if this happens in a new version of OCaml *)
let last_variant_tag = 245
type structured_constant =
| Const_sorts of sorts
- | Const_ind of pinductive
+ | Const_ind of inductive
| Const_proj of Constant.t
| Const_b0 of tag
| Const_bn of tag * structured_constant array
+ | Const_univ_level of Univ.universe_level
+ | Const_type of Univ.universe
type reloc_table = (tag * int) array
@@ -68,7 +73,8 @@ type instruction =
| Kclosure of Label.t * int
| Kclosurerec of int * int * Label.t array * Label.t array
| Kclosurecofix of int * int * Label.t array * Label.t array
- | Kgetglobal of pconstant
+ (* nb fv, init, lbl types, lbl bodies *)
+ | Kgetglobal of constant
| Kconst of structured_constant
| Kmakeblock of int * tag
| Kmakeprod
@@ -124,7 +130,10 @@ type instruction =
and bytecodes = instruction list
-type fv_elem = FVnamed of Id.t | FVrel of int
+type fv_elem =
+ | FVnamed of Id.t
+ | FVrel of int
+ | FVuniv_var of int
type fv = fv_elem array
@@ -142,18 +151,17 @@ type vm_env = {
type comp_env = {
- nb_stack : int; (* nbre de variables sur la pile *)
- in_stack : int list; (* position dans la pile *)
- nb_rec : int; (* nbre de fonctions mutuellement *)
- (* recursives = nbr *)
+ nb_uni_stack : int ; (* number of universes on the stack, *)
+ (* universes are always at the bottom. *)
+ nb_stack : int; (* number of variables on the stack *)
+ in_stack : int list; (* position in the stack *)
+ nb_rec : int; (* number of mutually recursive functions *)
pos_rec : instruction list; (* instruction d'acces pour les variables *)
(* de point fix ou de cofix *)
offset : int;
- in_env : vm_env ref
+ in_env : vm_env ref (* The free variables of the expression *)
}
-
-
(* --- Pretty print *)
open Pp
open Util
@@ -166,14 +174,24 @@ let pp_sort s =
let rec pp_struct_const = function
| Const_sorts s -> pp_sort s
- | Const_ind ((mind, i), u) -> pr_mind mind ++ str"#" ++ int i
+ | Const_ind (mind, i) -> pr_mind mind ++ str"#" ++ int i
| Const_proj p -> Constant.print p
| Const_b0 i -> int i
| Const_bn (i,t) ->
int i ++ surround (prvect_with_sep pr_comma pp_struct_const t)
+ | Const_univ_level l -> Univ.Level.pr l
+ | Const_type u -> str "Type@{" ++ Univ.pr_uni u ++ str "}"
let pp_lbl lbl = str "L" ++ int lbl
+let pp_pcon (id,u) =
+ pr_con id ++ str "@{" ++ Univ.Instance.pr Univ.Level.pr u ++ str "}"
+
+let pp_fv_elem = function
+ | FVnamed id -> str "FVnamed(" ++ Id.print id ++ str ")"
+ | FVrel i -> str "Rel(" ++ int i ++ str ")"
+ | FVuniv_var v -> str "FVuniv(" ++ int v ++ str ")"
+
let rec pp_instr i =
match i with
| Klabel _ | Ksequence _ -> assert false
@@ -207,8 +225,7 @@ let rec pp_instr i =
prlist_with_sep spc pp_lbl (Array.to_list lblt) ++
str " bodies = " ++
prlist_with_sep spc pp_lbl (Array.to_list lblb))
- | Kgetglobal (id,u) ->
- str "getglobal " ++ pr_con id ++ str "@{" ++ Univ.Instance.pr Univ.Level.pr u ++ str "}"
+ | Kgetglobal idu -> str "getglobal " ++ pr_con idu
| Kconst sc ->
str "const " ++ pp_struct_const sc
| Kmakeblock(n, m) ->
@@ -266,10 +283,6 @@ and pp_bytecodes c =
| i :: c ->
tab () ++ pp_instr i ++ fnl () ++ pp_bytecodes c
-let dump_bytecode c =
- pperrnl (pp_bytecodes c);
- flush_all()
-
(*spiwack: moved this type in this file because I needed it for
retroknowledge which can't depend from cbytegen *)
type block =
diff --git a/kernel/cbytecodes.mli b/kernel/cbytecodes.mli
index 8f594a45b5..6fa0841af9 100644
--- a/kernel/cbytecodes.mli
+++ b/kernel/cbytecodes.mli
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -13,21 +13,28 @@ open Term
type tag = int
-val id_tag : tag
-val iddef_tag : tag
-val ind_tag : tag
-val fix_tag : tag
+val accu_tag : tag
+
+val type_atom_tag : tag
+val max_atom_tag : tag
+val proj_tag : tag
+val fix_app_tag : tag
val switch_tag : tag
val cofix_tag : tag
val cofix_evaluated_tag : tag
-val last_variant_tag : tag
+
+val last_variant_tag : tag
type structured_constant =
| Const_sorts of sorts
- | Const_ind of pinductive
+ | Const_ind of inductive
| Const_proj of Constant.t
| Const_b0 of tag
| Const_bn of tag * structured_constant array
+ | Const_univ_level of Univ.universe_level
+ | Const_type of Univ.universe
+
+val pp_struct_const : structured_constant -> Pp.std_ppcmds
type reloc_table = (tag * int) array
@@ -62,9 +69,11 @@ type instruction =
(** nb fv, init, lbl types, lbl bodies *)
| Kclosurecofix of int * int * Label.t array * Label.t array
(** nb fv, init, lbl types, lbl bodies *)
- | Kgetglobal of pconstant (** accu = coq_global_data[c] *)
+ | Kgetglobal of constant
| Kconst of structured_constant
- | Kmakeblock of (* size: *) int * tag (** allocate an ocaml block *)
+ | Kmakeblock of (* size: *) int * tag (** allocate an ocaml block. Index 0
+ ** is accu, all others are popped from
+ ** the top of the stack *)
| Kmakeprod
| Kmakeswitchblock of Label.t * Label.t * annot_switch * int
| Kswitch of Label.t array * Label.t array (** consts,blocks *)
@@ -118,7 +127,10 @@ type instruction =
and bytecodes = instruction list
-type fv_elem = FVnamed of Id.t | FVrel of int
+type fv_elem =
+ FVnamed of Id.t
+| FVrel of int
+| FVuniv_var of int
type fv = fv_elem array
@@ -127,26 +139,28 @@ type fv = fv_elem array
closed terms. *)
exception NotClosed
-(*spiwack: both type have been moved from Cbytegen because I needed then
+(*spiwack: both type have been moved from Cbytegen because I needed them
for the retroknowledge *)
type vm_env = {
- size : int; (** longueur de la liste [n] *)
+ size : int; (** length of the list [n] *)
fv_rev : fv_elem list (** [fvn; ... ;fv1] *)
}
type comp_env = {
+ nb_uni_stack : int ; (** number of universes on the stack *)
nb_stack : int; (** number of variables on the stack *)
in_stack : int list; (** position in the stack *)
nb_rec : int; (** number of mutually recursive functions *)
- (** recursives = nbr *)
+ (** (= nbr) *)
pos_rec : instruction list; (** instruction d'acces pour les variables *)
(** de point fix ou de cofix *)
offset : int;
- in_env : vm_env ref
+ in_env : vm_env ref (** the variables that are accessed *)
}
-val dump_bytecode : bytecodes -> unit
+val pp_bytecodes : bytecodes -> Pp.std_ppcmds
+val pp_fv_elem : fv_elem -> Pp.std_ppcmds
(*spiwack: moved this here because I needed it for retroknowledge *)
type block =
diff --git a/kernel/cbytegen.ml b/kernel/cbytegen.ml
index 3462694d61..77eac9ee93 100644
--- a/kernel/cbytegen.ml
+++ b/kernel/cbytegen.ml
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -91,18 +91,20 @@ open Pre_env
(* In Cfxe_t accumulators, we need to store [fcofixi] for testing *)
(* conversion of cofixpoints (which is intentional). *)
+type argument = ArgConstr of Constr.t | ArgUniv of Univ.Level.t
let empty_fv = { size= 0; fv_rev = [] }
let fv r = !(r.in_env)
-let empty_comp_env ()=
- { nb_stack = 0;
+let empty_comp_env ?(univs=0) ()=
+ { nb_uni_stack = univs;
+ nb_stack = 0;
in_stack = [];
nb_rec = 0;
pos_rec = [];
offset = 0;
- in_env = ref empty_fv;
+ in_env = ref empty_fv
}
(*i Creation functions for comp_env *)
@@ -110,8 +112,9 @@ let empty_comp_env ()=
let rec add_param n sz l =
if Int.equal n 0 then l else add_param (n - 1) sz (n+sz::l)
-let comp_env_fun arity =
- { nb_stack = arity;
+let comp_env_fun ?(univs=0) arity =
+ { nb_uni_stack = univs ;
+ nb_stack = arity;
in_stack = add_param arity 0 [];
nb_rec = 0;
pos_rec = [];
@@ -120,8 +123,9 @@ let comp_env_fun arity =
}
-let comp_env_fix_type rfv =
- { nb_stack = 0;
+let comp_env_fix_type rfv =
+ { nb_uni_stack = 0;
+ nb_stack = 0;
in_stack = [];
nb_rec = 0;
pos_rec = [];
@@ -134,7 +138,8 @@ let comp_env_fix ndef curr_pos arity rfv =
for i = ndef downto 1 do
prec := Koffsetclosure (2 * (ndef - curr_pos - i)) :: !prec
done;
- { nb_stack = arity;
+ { nb_uni_stack = 0;
+ nb_stack = arity;
in_stack = add_param arity 0 [];
nb_rec = ndef;
pos_rec = !prec;
@@ -143,7 +148,8 @@ let comp_env_fix ndef curr_pos arity rfv =
}
let comp_env_cofix_type ndef rfv =
- { nb_stack = 0;
+ { nb_uni_stack = 0;
+ nb_stack = 0;
in_stack = [];
nb_rec = 0;
pos_rec = [];
@@ -156,7 +162,8 @@ let comp_env_cofix ndef arity rfv =
for i = 1 to ndef do
prec := Kenvacc i :: !prec
done;
- { nb_stack = arity;
+ { nb_uni_stack = 0;
+ nb_stack = arity;
in_stack = add_param arity 0 [];
nb_rec = ndef;
pos_rec = !prec;
@@ -176,8 +183,6 @@ let push_local sz r =
nb_stack = r.nb_stack + 1;
in_stack = (sz + 1) :: r.in_stack }
-
-
(*i Compilation of variables *)
let find_at f l =
let rec aux n = function
@@ -214,6 +219,22 @@ let pos_rel i r sz =
r.in_env := { size = pos+1; fv_rev = db:: env.fv_rev};
Kenvacc(r.offset + pos)
+let pos_universe_var i r sz =
+ if i < r.nb_uni_stack then
+ Kacc (sz - r.nb_stack - (r.nb_uni_stack - i))
+ else
+ let env = !(r.in_env) in
+ let f = function
+ | FVuniv_var u -> Int.equal i u
+ | _ -> false
+ in
+ try Kenvacc (r.offset + env.size - (find_at f env.fv_rev))
+ with Not_found ->
+ let pos = env.size in
+ let db = FVuniv_var i in
+ r.in_env := { size = pos + 1; fv_rev = db::env.fv_rev } ;
+ Kenvacc(r.offset + pos)
+
(*i Examination of the continuation *)
(* Discard all instructions up to the next label. *)
@@ -459,8 +480,9 @@ let rec str_const c =
end
| _ -> Bconstr c
end
- | Ind ind -> Bstrconst (Const_ind ind)
- | Construct (((kn,j),i),u) ->
+ | Ind (ind,u) when Univ.Instance.is_empty u ->
+ Bstrconst (Const_ind ind)
+ | Construct (((kn,j),i),_) ->
begin
(* spiwack: tries first to apply the run-time compilation
behavior of the constructor, as in 2/ above *)
@@ -513,6 +535,7 @@ let compile_fv_elem reloc fv sz cont =
match fv with
| FVrel i -> pos_rel i reloc sz :: cont
| FVnamed id -> pos_named id reloc :: cont
+ | FVuniv_var i -> pos_universe_var i reloc sz :: cont
let rec compile_fv reloc l sz cont =
match l with
@@ -524,18 +547,17 @@ let rec compile_fv reloc l sz cont =
(* Compiling constants *)
-let rec get_alias env (kn,u as p) =
+let rec get_alias env kn =
let cb = lookup_constant kn env in
let tps = cb.const_body_code in
match tps with
- | None -> p
+ | None -> kn
| Some tps ->
(match Cemitcodes.force tps with
- | BCalias (kn',u') -> get_alias env (kn', Univ.subst_instance_instance u u')
- | _ -> p)
-
-(* Compiling expressions *)
+ | BCalias kn' -> get_alias env kn'
+ | _ -> kn)
+(* sz is the size of the local stack *)
let rec compile_constr reloc c sz cont =
match kind_of_term c with
| Meta _ -> invalid_arg "Cbytegen.compile_constr : Meta"
@@ -552,9 +574,44 @@ let rec compile_constr reloc c sz cont =
| Rel i -> pos_rel i reloc sz :: cont
| Var id -> pos_named id reloc :: cont
| Const (kn,u) -> compile_const reloc kn u [||] sz cont
- | Sort _ | Ind _ | Construct _ ->
+ | Ind (ind,u) ->
+ let bcst = Bstrconst (Const_ind ind) in
+ if Univ.Instance.is_empty u then
+ compile_str_cst reloc bcst sz cont
+ else
+ comp_app compile_str_cst compile_universe reloc
+ bcst
+ (Univ.Instance.to_array u)
+ sz
+ cont
+ | Sort (Prop _) | Construct _ ->
compile_str_cst reloc (str_const c) sz cont
-
+ | Sort (Type u) ->
+ (* We separate global and local universes in [u]. The former will be part
+ of the structured constant, while the later (if any) will be applied as
+ arguments. *)
+ let open Univ in begin
+ let levels = Universe.levels u in
+ let global_levels =
+ LSet.filter (fun x -> Level.var_index x = None) levels
+ in
+ let local_levels =
+ List.map_filter (fun x -> Level.var_index x)
+ (LSet.elements levels)
+ in
+ (* We assume that [Universe.type0m] is a neutral element for [Universe.sup] *)
+ let uglob =
+ LSet.fold (fun lvl u -> Universe.sup u (Universe.make lvl)) global_levels Universe.type0m
+ in
+ if local_levels = [] then
+ compile_str_cst reloc (Bstrconst (Const_sorts (Type uglob))) sz cont
+ else
+ let compile_get_univ reloc idx sz cont =
+ compile_fv_elem reloc (FVuniv_var idx) sz cont
+ in
+ comp_app compile_str_cst compile_get_univ reloc
+ (Bstrconst (Const_type u)) (Array.of_list local_levels) sz cont
+ end
| LetIn(_,xb,_,body) ->
compile_constr reloc xb sz
(Kpush ::
@@ -663,7 +720,9 @@ let rec compile_constr reloc c sz cont =
let lbl_sw = Label.create () in
let sz_b,branch,is_tailcall =
match branch1 with
- | Kreturn k -> assert (Int.equal k sz); sz, branch1, true
+ | Kreturn k ->
+ assert (Int.equal k sz) ;
+ sz, branch1, true
| _ -> sz+3, Kjump, false
in
let annot = {ci = ci; rtbl = tbl; tailcall = is_tailcall} in
@@ -745,8 +804,20 @@ and compile_str_cst reloc sc sz cont =
(* spiwack : compilation of constants with their arguments.
Makes a special treatment with 31-bit integer addition *)
-and compile_const =
- fun reloc-> fun kn u -> fun args -> fun sz -> fun cont ->
+and compile_get_global reloc (kn,u) sz cont =
+ let kn = get_alias !global_env kn in
+ if Univ.Instance.is_empty u then
+ Kgetglobal kn :: cont
+ else
+ comp_app (fun _ _ _ cont -> Kgetglobal kn :: cont)
+ compile_universe reloc () (Univ.Instance.to_array u) sz cont
+
+and compile_universe reloc uni sz cont =
+ match Univ.Level.var_index uni with
+ | None -> Kconst (Const_univ_level uni) :: cont
+ | Some idx -> pos_universe_var idx reloc sz :: cont
+
+and compile_const reloc kn u args sz cont =
let nargs = Array.length args in
(* spiwack: checks if there is a specific way to compile the constant
if there is not, Not_found is raised, and the function
@@ -756,31 +827,84 @@ and compile_const =
(mkConstU (kn,u)) reloc args sz cont
with Not_found ->
if Int.equal nargs 0 then
- Kgetglobal (get_alias !global_env (kn, u)) :: cont
+ compile_get_global reloc (kn,u) sz cont
else
- comp_app (fun _ _ _ cont ->
- Kgetglobal (get_alias !global_env (kn,u)) :: cont)
- compile_constr reloc () args sz cont
-
-let compile fail_on_error env c =
+ if Univ.Instance.is_empty u then
+ (* normal compilation *)
+ comp_app (fun _ _ sz cont ->
+ compile_get_global reloc (kn,u) sz cont)
+ compile_constr reloc () args sz cont
+ else
+ let compile_arg reloc constr_or_uni sz cont =
+ match constr_or_uni with
+ | ArgConstr cst -> compile_constr reloc cst sz cont
+ | ArgUniv uni -> compile_universe reloc uni sz cont
+ in
+ let u = Univ.Instance.to_array u in
+ let lu = Array.length u in
+ let all =
+ Array.init (lu + Array.length args)
+ (fun i -> if i < lu then ArgUniv u.(i) else ArgConstr args.(i-lu))
+ in
+ comp_app (fun _ _ _ cont -> Kgetglobal kn :: cont)
+ compile_arg reloc () all sz cont
+
+let is_univ_copy max u =
+ let u = Univ.Instance.to_array u in
+ if Array.length u = max then
+ Array.fold_left_i (fun i acc u ->
+ if acc then
+ match Univ.Level.var_index u with
+ | None -> false
+ | Some l -> l = i
+ else false) true u
+ else
+ false
+
+let dump_bytecodes init code fvs =
+ let open Pp in
+ (str "code =" ++ fnl () ++
+ pp_bytecodes init ++ fnl () ++
+ pp_bytecodes code ++ fnl () ++
+ str "fv = " ++
+ prlist_with_sep (fun () -> str "; ") pp_fv_elem fvs ++
+ fnl ())
+
+let compile fail_on_error ?universes:(universes=0) env c =
set_global_env env;
init_fun_code ();
Label.reset_label_counter ();
- let reloc = empty_comp_env () in
- try
- let init_code = compile_constr reloc c 0 [Kstop] in
- let fv = List.rev (!(reloc.in_env).fv_rev) in
- let pp_v v =
- match v with
- | FVnamed id -> Pp.str (Id.to_string id)
- | FVrel i -> Pp.str (string_of_int i)
+ let cont = [Kstop] in
+ try
+ let reloc, init_code =
+ if Int.equal universes 0 then
+ let reloc = empty_comp_env () in
+ reloc, compile_constr reloc c 0 cont
+ else
+ (* We are going to generate a lambda, but merge the universe closure
+ * with the function closure if it exists.
+ *)
+ let reloc = empty_comp_env () in
+ let arity , body =
+ match kind_of_term c with
+ | Lambda _ ->
+ let params, body = decompose_lam c in
+ List.length params , body
+ | _ -> 0 , c
+ in
+ let full_arity = arity + universes in
+ let r_fun = comp_env_fun ~univs:universes arity in
+ let lbl_fun = Label.create () in
+ let cont_fun =
+ compile_constr r_fun body full_arity [Kreturn full_arity]
+ in
+ fun_code := [Ksequence(add_grab full_arity lbl_fun cont_fun,!fun_code)];
+ let fv = fv r_fun in
+ reloc, compile_fv reloc fv.fv_rev 0 (Kclosure(lbl_fun,fv.size) :: cont)
in
- let open Pp in
- if !Flags.dump_bytecode then
- (dump_bytecode init_code;
- dump_bytecode !fun_code;
- Pp.msg_debug (Pp.str "fv = " ++
- Pp.prlist_with_sep (fun () -> Pp.str "; ") pp_v fv ++ Pp.fnl ()));
+ let fv = List.rev (!(reloc.in_env).fv_rev) in
+ (if !Flags.dump_bytecode then
+ Pp.msg_debug (dump_bytecodes init_code !fun_code fv)) ;
Some (init_code,!fun_code, Array.of_list fv)
with TooLargeInductive tname ->
let fn = if fail_on_error then Errors.errorlabstrm "compile" else Pp.msg_warning in
@@ -789,28 +913,33 @@ let compile fail_on_error env c =
Id.print tname ++ str str_max_constructors));
None)
-let compile_constant_body fail_on_error env = function
+let compile_constant_body fail_on_error env univs = function
| Undef _ | OpaqueDef _ -> Some BCconstant
| Def sb ->
let body = Mod_subst.force_constr sb in
+ let instance_size =
+ match univs with
+ | None -> 0
+ | Some univ -> Univ.UContext.size univ
+ in
match kind_of_term body with
- | Const (kn',u) ->
+ | Const (kn',u) when is_univ_copy instance_size u ->
(* we use the canonical name of the constant*)
let con= constant_of_kn (canonical_con kn') in
- Some (BCalias (get_alias env (con,u)))
+ Some (BCalias (get_alias env con))
| _ ->
- let res = compile fail_on_error env body in
+ let res = compile fail_on_error ~universes:instance_size env body in
Option.map (fun x -> BCdefined (to_memory x)) res
(* Shortcut of the previous function used during module strengthening *)
-let compile_alias (kn,u) = BCalias (constant_of_kn (canonical_con kn), u)
+let compile_alias kn = BCalias (constant_of_kn (canonical_con kn))
(* spiwack: additional function which allow different part of compilation of the
31-bit integers *)
let make_areconst n else_lbl cont =
- if n <=0 then
+ if n <= 0 then
cont
else
Kareconst (n, else_lbl)::cont
@@ -902,14 +1031,14 @@ let op2_compilation op =
3/ if at least one is not, branches to the normal behavior:
Kgetglobal (get_alias !global_env kn) *)
let op_compilation n op =
- let code_construct kn cont =
+ let code_construct reloc kn sz cont =
let f_cont =
let else_lbl = Label.create () in
Kareconst(n, else_lbl):: Kacc 0:: Kpop 1::
op:: Kreturn 0:: Klabel else_lbl::
(* works as comp_app with nargs = n and tailcall cont [Kreturn 0]*)
- Kgetglobal (get_alias !global_env kn)::
- Kappterm(n, n):: [] (* = discard_dead_code [Kreturn 0] *)
+ compile_get_global reloc kn sz (
+ Kappterm(n, n):: []) (* = discard_dead_code [Kreturn 0] *)
in
let lbl = Label.create () in
fun_code := [Ksequence (add_grab n lbl f_cont, !fun_code)];
@@ -926,12 +1055,11 @@ let op_compilation n op =
(*Kaddint31::escape::Klabel else_lbl::Kpush::*)
(op::escape::Klabel else_lbl::Kpush::
(* works as comp_app with nargs = n and non-tailcall cont*)
- Kgetglobal (get_alias !global_env kn)::
- Kapply n::labeled_cont)))
+ compile_get_global reloc kn sz (Kapply n::labeled_cont))))
else if Int.equal nargs 0 then
- code_construct kn cont
+ code_construct reloc kn sz cont
else
- comp_app (fun _ _ _ cont -> code_construct kn cont)
+ comp_app (fun reloc _ sz cont -> code_construct reloc kn sz cont)
compile_constr reloc () args sz cont
let int31_escape_before_match fc cont =
diff --git a/kernel/cbytegen.mli b/kernel/cbytegen.mli
index 1128f0d0b7..c0f48641ce 100644
--- a/kernel/cbytegen.mli
+++ b/kernel/cbytegen.mli
@@ -4,17 +4,17 @@ open Term
open Declarations
open Pre_env
-
+(** Should only be used for monomorphic terms *)
val compile : bool -> (* Fail on error with a nice user message, otherwise simply a warning *)
- env -> constr -> (bytecodes * bytecodes * fv) option
+ ?universes:int -> env -> constr -> (bytecodes * bytecodes * fv) option
(** init, fun, fv *)
-val compile_constant_body : bool ->
- env -> constant_def -> body_code option
+val compile_constant_body : bool ->
+ env -> constant_universes option -> constant_def -> body_code option
(** Shortcut of the previous function used during module strengthening *)
-val compile_alias : pconstant -> body_code
+val compile_alias : Names.constant -> body_code
(** spiwack: this function contains the information needed to perform
the static compilation of int31 (trying and obtaining
diff --git a/kernel/cemitcodes.ml b/kernel/cemitcodes.ml
index 9b275cb6c3..d779a81ff6 100644
--- a/kernel/cemitcodes.ml
+++ b/kernel/cemitcodes.ml
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -19,7 +19,7 @@ open Mod_subst
type reloc_info =
| Reloc_annot of annot_switch
| Reloc_const of structured_constant
- | Reloc_getglobal of pconstant
+ | Reloc_getglobal of Names.constant
type patch = reloc_info * int
@@ -29,11 +29,19 @@ let patch_char4 buff pos c1 c2 c3 c4 =
String.unsafe_set buff (pos + 2) c3;
String.unsafe_set buff (pos + 3) c4
-let patch_int buff pos n =
+let patch buff (pos, n) =
patch_char4 buff pos
(Char.unsafe_chr n) (Char.unsafe_chr (n asr 8)) (Char.unsafe_chr (n asr 16))
(Char.unsafe_chr (n asr 24))
+let patch_int buff patches =
+ (* copy code *before* patching because of nested evaluations:
+ the code we are patching might be called (and thus "concurrently" patched)
+ and results in wrong results. Side-effects... *)
+ let buff = String.copy buff in
+ let () = List.iter (fun p -> patch buff p) patches in
+ buff
+
(* Buffering of bytecode *)
let out_buffer = ref(String.create 1024)
@@ -127,11 +135,11 @@ let slot_for_const c =
enter (Reloc_const c);
out_int 0
-and slot_for_annot a =
+let slot_for_annot a =
enter (Reloc_annot a);
out_int 0
-and slot_for_getglobal p =
+let slot_for_getglobal p =
enter (Reloc_getglobal p);
out_int 0
@@ -190,7 +198,7 @@ let emit_instr = function
Array.iter (out_label_with_orig org) lbl_bodies
| Kgetglobal q ->
out opGETGLOBAL; slot_for_getglobal q
- | Kconst((Const_b0 i)) ->
+ | Kconst (Const_b0 i) ->
if i >= 0 && i <= 3
then out (opCONST0 + i)
else (out opCONSTINT; out_int i)
@@ -298,8 +306,6 @@ let init () =
type emitcodes = string
-let copy = String.copy
-
let length = String.length
type to_patch = emitcodes * (patch list) * fv
@@ -307,10 +313,10 @@ type to_patch = emitcodes * (patch list) * fv
(* Substitution *)
let rec subst_strcst s sc =
match sc with
- | Const_sorts _ | Const_b0 _ -> sc
+ | Const_sorts _ | Const_b0 _ | Const_univ_level _ | Const_type _ -> sc
| Const_proj p -> Const_proj (subst_constant s p)
| Const_bn(tag,args) -> Const_bn(tag,Array.map (subst_strcst s) args)
- | Const_ind(ind,u) -> let kn,i = ind in Const_ind((subst_mind s kn, i), u)
+ | Const_ind ind -> let kn,i = ind in Const_ind (subst_mind s kn, i)
let subst_patch s (ri,pos) =
match ri with
@@ -319,21 +325,19 @@ let subst_patch s (ri,pos) =
let ci = {a.ci with ci_ind = (subst_mind s kn,i)} in
(Reloc_annot {a with ci = ci},pos)
| Reloc_const sc -> (Reloc_const (subst_strcst s sc), pos)
- | Reloc_getglobal kn -> (Reloc_getglobal (subst_pcon s kn), pos)
+ | Reloc_getglobal kn -> (Reloc_getglobal (subst_constant s kn), pos)
let subst_to_patch s (code,pl,fv) =
code,List.rev_map (subst_patch s) pl,fv
-let subst_pconstant s (kn, u) = (fst (subst_con_kn s kn), u)
-
type body_code =
| BCdefined of to_patch
- | BCalias of pconstant
+ | BCalias of Names.constant
| BCconstant
type to_patch_substituted =
| PBCdefined of to_patch substituted
-| PBCalias of pconstant substituted
+| PBCalias of Names.constant substituted
| PBCconstant
let from_val = function
@@ -343,7 +347,7 @@ let from_val = function
let force = function
| PBCdefined tp -> BCdefined (force subst_to_patch tp)
-| PBCalias cu -> BCalias (force subst_pconstant cu)
+| PBCalias cu -> BCalias (force subst_constant cu)
| PBCconstant -> BCconstant
let subst_to_patch_subst s = function
@@ -366,6 +370,8 @@ let to_memory (init_code, fun_code, fv) =
emit fun_code;
let code = String.create !out_position in
String.unsafe_blit !out_buffer 0 code 0 !out_position;
+ (** Later uses of this string are all purely functional *)
+ let code = CString.hcons code in
let reloc = List.rev !reloc_info in
Array.iter (fun lbl ->
(match lbl with
@@ -373,8 +379,3 @@ let to_memory (init_code, fun_code, fv) =
| Label_undefined patchlist ->
assert (patchlist = []))) !label_table;
(code, reloc, fv)
-
-
-
-
-
diff --git a/kernel/cemitcodes.mli b/kernel/cemitcodes.mli
index 54b92b9121..c80edd5965 100644
--- a/kernel/cemitcodes.mli
+++ b/kernel/cemitcodes.mli
@@ -4,7 +4,7 @@ open Cbytecodes
type reloc_info =
| Reloc_annot of annot_switch
| Reloc_const of structured_constant
- | Reloc_getglobal of constant Univ.puniverses
+ | Reloc_getglobal of constant
type patch = reloc_info * int
@@ -13,11 +13,9 @@ val subst_patch : Mod_subst.substitution -> patch -> patch
type emitcodes
-val copy : emitcodes -> emitcodes
-
val length : emitcodes -> int
-val patch_int : emitcodes -> (*pos*)int -> int -> unit
+val patch_int : emitcodes -> ((*pos*)int * int) list -> emitcodes
type to_patch = emitcodes * (patch list) * fv
@@ -25,7 +23,7 @@ val subst_to_patch : Mod_subst.substitution -> to_patch -> to_patch
type body_code =
| BCdefined of to_patch
- | BCalias of constant Univ.puniverses
+ | BCalias of constant
| BCconstant
diff --git a/kernel/closure.ml b/kernel/closure.ml
index ea9b2755f2..bf3801e547 100644
--- a/kernel/closure.ml
+++ b/kernel/closure.ml
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -245,10 +245,12 @@ and 'a infos = {
let info_flags info = info.i_flags
let info_env info = info.i_cache.i_env
+open Context.Named.Declaration
+
let rec assoc_defined id = function
| [] -> raise Not_found
-| (_, None, _) :: ctxt -> assoc_defined id ctxt
-| (id', Some c, _) :: ctxt ->
+| LocalAssum _ :: ctxt -> assoc_defined id ctxt
+| LocalDef (id', c, _) :: ctxt ->
if Id.equal id id' then c else assoc_defined id ctxt
let ref_value_cache ({i_cache = cache} as infos) ref =
@@ -285,9 +287,10 @@ let defined_rels flags env =
let ctx = rel_context env in
let len = List.length ctx in
let ans = Array.make len None in
- let iter i (_, b, _) = match b with
- | None -> ()
- | Some _ -> Array.unsafe_set ans i b
+ let open Context.Rel.Declaration in
+ let iter i = function
+ | LocalAssum _ -> ()
+ | LocalDef (_,b,_) -> Array.unsafe_set ans i (Some b)
in
let () = List.iteri iter ctx in
ans
@@ -346,7 +349,6 @@ and fterm =
| FProj of projection * fconstr
| FFix of fixpoint * fconstr subs
| FCoFix of cofixpoint * fconstr subs
- | FCase of case_info * fconstr * fconstr * fconstr array
| FCaseT of case_info * constr * fconstr * constr array * fconstr subs (* predicate and branches are closures *)
| FLambda of int * (Name.t * constr) list * constr * fconstr subs
| FProd of Name.t * fconstr * fconstr
@@ -376,7 +378,6 @@ let update v1 no t =
type stack_member =
| Zapp of fconstr array
- | Zcase of case_info * fconstr * fconstr array
| ZcaseT of case_info * constr * constr array * fconstr subs
| Zproj of int * int * constant
| Zfix of fconstr * stack
@@ -569,10 +570,6 @@ let rec to_constr constr_fun lfts v =
| FFlex (ConstKey op) -> mkConstU op
| FInd op -> mkIndU op
| FConstruct op -> mkConstructU op
- | FCase (ci,p,c,ve) ->
- mkCase (ci, constr_fun lfts p,
- constr_fun lfts c,
- CArray.Fun1.map constr_fun lfts ve)
| FCaseT (ci,p,c,ve,env) ->
mkCase (ci, constr_fun lfts (mk_clos env p),
constr_fun lfts c,
@@ -646,9 +643,6 @@ let rec zip m stk =
match stk with
| [] -> m
| Zapp args :: s -> zip {norm=neutr m.norm; term=FApp(m, args)} s
- | Zcase(ci,p,br)::s ->
- let t = FCase(ci, p, m, br) in
- zip {norm=neutr m.norm; term=t} s
| ZcaseT(ci,p,br,e)::s ->
let t = FCaseT(ci, p, m, br, e) in
zip {norm=neutr m.norm; term=t} s
@@ -731,7 +725,7 @@ let rec get_args n tys f e stk =
(* Eta expansion: add a reference to implicit surrounding lambda at end of stack *)
let rec eta_expand_stack = function
- | (Zapp _ | Zfix _ | Zcase _ | ZcaseT _ | Zproj _
+ | (Zapp _ | Zfix _ | ZcaseT _ | Zproj _
| Zshift _ | Zupdate _ as e) :: s ->
e :: eta_expand_stack s
| [] ->
@@ -771,7 +765,7 @@ let drop_parameters depth n argstk =
(* we know that n < stack_args_size(argstk) (if well-typed term) *)
anomaly (Pp.str "ill-typed term: found a match on a partially applied constructor")
-(** [eta_expand_ind_stack env ind c s t] computes stacks correspoding
+(** [eta_expand_ind_stack env ind c s t] computes stacks corresponding
to the conversion of the eta expansion of t, considered as an inhabitant
of ind, and the Constructor c of this inductive type applied to arguments
s.
@@ -784,7 +778,7 @@ let eta_expand_ind_stack env ind m s (f, s') =
let mib = lookup_mind (fst ind) env in
match mib.Declarations.mind_record with
| Some (Some (_,projs,pbs)) when
- mib.Declarations.mind_finite <> Decl_kinds.CoFinite ->
+ mib.Declarations.mind_finite == Decl_kinds.BiFinite ->
(* (Construct, pars1 .. parsm :: arg1...argn :: []) ~= (f, s') ->
arg1..argn ~= (proj1 t...projn t) where t = zip (f,s') *)
let pars = mib.Declarations.mind_nparams in
@@ -842,7 +836,6 @@ let rec knh info m stk =
| FCLOS(t,e) -> knht info e t (zupdate m stk)
| FLOCKED -> assert false
| FApp(a,b) -> knh info a (append_stack b (zupdate m stk))
- | FCase(ci,p,t,br) -> knh info t (Zcase(ci,p,br)::zupdate m stk)
| FCaseT(ci,p,t,br,e) -> knh info t (ZcaseT(ci,p,br,e)::zupdate m stk)
| FFix(((ri,n),(_,_,_)),_) ->
(match get_nth_arg m ri.(n) stk with
@@ -904,10 +897,6 @@ let rec knr info m stk =
| None -> (set_norm m; (m,stk)))
| FConstruct((ind,c),u) when red_set info.i_flags fIOTA ->
(match strip_update_shift_app m stk with
- (depth, args, Zcase(ci,_,br)::s) ->
- assert (ci.ci_npar>=0);
- let rargs = drop_parameters depth ci.ci_npar args in
- kni info br.(c-1) (rargs@s)
| (depth, args, ZcaseT(ci,_,br,e)::s) ->
assert (ci.ci_npar>=0);
let rargs = drop_parameters depth ci.ci_npar args in
@@ -924,7 +913,7 @@ let rec knr info m stk =
| (_,args,s) -> (m,args@s))
| FCoFix _ when red_set info.i_flags fIOTA ->
(match strip_update_shift_app m stk with
- (_, args, (((Zcase _|ZcaseT _|Zproj _)::_) as stk')) ->
+ (_, args, (((ZcaseT _|Zproj _)::_) as stk')) ->
let (fxe,fxbd) = contract_fix_vect m.term in
knit info fxe fxbd (args@stk')
| (_,args,s) -> (m,args@s))
@@ -953,9 +942,6 @@ let rec zip_term zfun m stk =
| [] -> m
| Zapp args :: s ->
zip_term zfun (mkApp(m, Array.map zfun args)) s
- | Zcase(ci,p,br)::s ->
- let t = mkCase(ci, zfun p, m, Array.map zfun br) in
- zip_term zfun t s
| ZcaseT(ci,p,br,e)::s ->
let t = mkCase(ci, zfun (mk_clos e p), m,
Array.map (fun b -> zfun (mk_clos e b)) br) in
diff --git a/kernel/closure.mli b/kernel/closure.mli
index a3b0e0f301..07176cb7de 100644
--- a/kernel/closure.mli
+++ b/kernel/closure.mli
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -119,7 +119,6 @@ type fterm =
| FProj of projection * fconstr
| FFix of fixpoint * fconstr subs
| FCoFix of cofixpoint * fconstr subs
- | FCase of case_info * fconstr * fconstr * fconstr array
| FCaseT of case_info * constr * fconstr * constr array * fconstr subs (* predicate and branches are closures *)
| FLambda of int * (Name.t * constr) list * constr * fconstr subs
| FProd of Name.t * fconstr * fconstr
@@ -136,7 +135,6 @@ type fterm =
type stack_member =
| Zapp of fconstr array
- | Zcase of case_info * fconstr * fconstr array
| ZcaseT of case_info * constr * constr array * fconstr subs
| Zproj of int * int * constant
| Zfix of fconstr * stack
diff --git a/kernel/constr.ml b/kernel/constr.ml
index e2b1d3fd9c..ce20751abc 100644
--- a/kernel/constr.ml
+++ b/kernel/constr.ml
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -41,12 +41,24 @@ type case_printing =
{ ind_tags : bool list; (** tell whether letin or lambda in the arity of the inductive type *)
cstr_tags : bool list array; (* whether each pattern var of each constructor is a let-in (true) or not (false) *)
style : case_style }
+
+(* INVARIANT:
+ * - Array.length ci_cstr_ndecls = Array.length ci_cstr_nargs
+ * - forall (i : 0 .. pred (Array.length ci_cstr_ndecls)),
+ * ci_cstr_ndecls.(i) >= ci_cstr_nargs.(i)
+ *)
type case_info =
- { ci_ind : inductive;
- ci_npar : int;
- ci_cstr_ndecls : int array; (* number of pattern vars of each constructor (with let's)*)
- ci_cstr_nargs : int array; (* number of pattern vars of each constructor (w/o let's) *)
- ci_pp_info : case_printing (* not interpreted by the kernel *)
+ { ci_ind : inductive; (* inductive type to which belongs the value that is being matched *)
+ ci_npar : int; (* number of parameters of the above inductive type *)
+ ci_cstr_ndecls : int array; (* For each constructor, the corresponding integer determines
+ the number of values that can be bound in a match-construct.
+ NOTE: parameters of the inductive type are therefore excluded from the count *)
+ ci_cstr_nargs : int array; (* for each constructor, the corresponding integers determines
+ the number of values that can be applied to the constructor,
+ in addition to the parameters of the related inductive type
+ NOTE: "lets" are therefore excluded from the count
+ NOTE: parameters of the inductive type are also excluded from the count *)
+ ci_pp_info : case_printing (* not interpreted by the kernel *)
}
(********************************************************************)
@@ -545,8 +557,8 @@ let equal m n = eq_constr m n (* to avoid tracing a recursive fun *)
let eq_constr_univs univs m n =
if m == n then true
else
- let eq_universes _ = Univ.Instance.check_eq univs in
- let eq_sorts s1 s2 = s1 == s2 || Univ.check_eq univs (Sorts.univ_of_sort s1) (Sorts.univ_of_sort s2) in
+ let eq_universes _ = UGraph.check_eq_instances univs in
+ let eq_sorts s1 s2 = s1 == s2 || UGraph.check_eq univs (Sorts.univ_of_sort s1) (Sorts.univ_of_sort s2) in
let rec eq_constr' m n =
m == n || compare_head_gen eq_universes eq_sorts eq_constr' m n
in compare_head_gen eq_universes eq_sorts eq_constr' m n
@@ -554,11 +566,11 @@ let eq_constr_univs univs m n =
let leq_constr_univs univs m n =
if m == n then true
else
- let eq_universes _ = Univ.Instance.check_eq univs in
+ let eq_universes _ = UGraph.check_eq_instances univs in
let eq_sorts s1 s2 = s1 == s2 ||
- Univ.check_eq univs (Sorts.univ_of_sort s1) (Sorts.univ_of_sort s2) in
+ UGraph.check_eq univs (Sorts.univ_of_sort s1) (Sorts.univ_of_sort s2) in
let leq_sorts s1 s2 = s1 == s2 ||
- Univ.check_leq univs (Sorts.univ_of_sort s1) (Sorts.univ_of_sort s2) in
+ UGraph.check_leq univs (Sorts.univ_of_sort s1) (Sorts.univ_of_sort s2) in
let rec eq_constr' m n =
m == n || compare_head_gen eq_universes eq_sorts eq_constr' m n
in
@@ -571,12 +583,12 @@ let eq_constr_univs_infer univs m n =
if m == n then true, Constraint.empty
else
let cstrs = ref Constraint.empty in
- let eq_universes strict = Univ.Instance.check_eq univs in
+ let eq_universes strict = UGraph.check_eq_instances univs in
let eq_sorts s1 s2 =
if Sorts.equal s1 s2 then true
else
let u1 = Sorts.univ_of_sort s1 and u2 = Sorts.univ_of_sort s2 in
- if Univ.check_eq univs u1 u2 then true
+ if UGraph.check_eq univs u1 u2 then true
else
(cstrs := Univ.enforce_eq u1 u2 !cstrs;
true)
@@ -591,12 +603,12 @@ let leq_constr_univs_infer univs m n =
if m == n then true, Constraint.empty
else
let cstrs = ref Constraint.empty in
- let eq_universes strict l l' = Univ.Instance.check_eq univs l l' in
+ let eq_universes strict l l' = UGraph.check_eq_instances univs l l' in
let eq_sorts s1 s2 =
if Sorts.equal s1 s2 then true
else
let u1 = Sorts.univ_of_sort s1 and u2 = Sorts.univ_of_sort s2 in
- if Univ.check_eq univs u1 u2 then true
+ if UGraph.check_eq univs u1 u2 then true
else (cstrs := Univ.enforce_eq u1 u2 !cstrs;
true)
in
@@ -604,7 +616,7 @@ let leq_constr_univs_infer univs m n =
if Sorts.equal s1 s2 then true
else
let u1 = Sorts.univ_of_sort s1 and u2 = Sorts.univ_of_sort s2 in
- if Univ.check_leq univs u1 u2 then true
+ if UGraph.check_leq univs u1 u2 then true
else
(cstrs := Univ.enforce_leq u1 u2 !cstrs;
true)
@@ -732,12 +744,10 @@ let hasheq t1 t2 =
n1 == n2 && b1 == b2 && t1 == t2 && c1 == c2
| App (c1,l1), App (c2,l2) -> c1 == c2 && array_eqeq l1 l2
| Proj (p1,c1), Proj(p2,c2) -> p1 == p2 && c1 == c2
- | Evar (e1,l1), Evar (e2,l2) -> Evar.equal e1 e2 && array_eqeq l1 l2
+ | Evar (e1,l1), Evar (e2,l2) -> e1 == e2 && array_eqeq l1 l2
| Const (c1,u1), Const (c2,u2) -> c1 == c2 && u1 == u2
- | Ind ((sp1,i1),u1), Ind ((sp2,i2),u2) ->
- sp1 == sp2 && Int.equal i1 i2 && u1 == u2
- | Construct (((sp1,i1),j1),u1), Construct (((sp2,i2),j2),u2) ->
- sp1 == sp2 && Int.equal i1 i2 && Int.equal j1 j2 && u1 == u2
+ | Ind (ind1,u1), Ind (ind2,u2) -> ind1 == ind2 && u1 == u2
+ | Construct (cstr1,u1), Construct (cstr2,u2) -> cstr1 == cstr2 && u1 == u2
| Case (ci1,p1,c1,bl1), Case (ci2,p2,c2,bl2) ->
ci1 == ci2 && p1 == p2 && c1 == c2 && array_eqeq bl1 bl2
| Fix ((ln1, i1),(lna1,tl1,bl1)), Fix ((ln2, i2),(lna2,tl2,bl2)) ->
@@ -757,10 +767,10 @@ let hasheq t1 t2 =
once and for all the table we'll use for hash-consing all constr *)
module HashsetTerm =
- Hashset.Make(struct type t = constr let equal = hasheq end)
+ Hashset.Make(struct type t = constr let eq = hasheq end)
module HashsetTermArray =
- Hashset.Make(struct type t = constr array let equal = array_eqeq end)
+ Hashset.Make(struct type t = constr array let eq = array_eqeq end)
let term_table = HashsetTerm.create 19991
(* The associative table to hashcons terms. *)
@@ -815,19 +825,19 @@ let hashcons (sh_sort,sh_ci,sh_construct,sh_ind,sh_con,sh_na,sh_id) =
| Proj (p,c) ->
let c, hc = sh_rec c in
let p' = Projection.hcons p in
- (Proj (p', c), combinesmall 17 (combine (Projection.hash p') hc))
+ (Proj (p', c), combinesmall 17 (combine (Projection.SyntacticOrd.hash p') hc))
| Const (c,u) ->
let c' = sh_con c in
let u', hu = sh_instance u in
- (Const (c', u'), combinesmall 9 (combine (Constant.hash c) hu))
- | Ind ((kn,i) as ind,u) ->
+ (Const (c', u'), combinesmall 9 (combine (Constant.SyntacticOrd.hash c) hu))
+ | Ind (ind,u) ->
let u', hu = sh_instance u in
(Ind (sh_ind ind, u'),
- combinesmall 10 (combine (ind_hash ind) hu))
- | Construct ((((kn,i),j) as c,u))->
+ combinesmall 10 (combine (ind_syntactic_hash ind) hu))
+ | Construct (c,u) ->
let u', hu = sh_instance u in
(Construct (sh_construct c, u'),
- combinesmall 11 (combine (constructor_hash c) hu))
+ combinesmall 11 (combine (constructor_syntactic_hash c) hu))
| Case (ci,p,c,bl) ->
let p, hp = sh_rec p
and c, hc = sh_rec c in
@@ -930,7 +940,7 @@ struct
List.equal (==) info1.ind_tags info2.ind_tags &&
Array.equal (List.equal (==)) info1.cstr_tags info2.cstr_tags &&
info1.style == info2.style
- let equal ci ci' =
+ let eq ci ci' =
ci.ci_ind == ci'.ci_ind &&
Int.equal ci.ci_npar ci'.ci_npar &&
Array.equal Int.equal ci.ci_cstr_ndecls ci'.ci_cstr_ndecls && (* we use [Array.equal] on purpose *)
@@ -972,7 +982,7 @@ module Hsorts =
let hashcons huniv = function
Prop c -> Prop c
| Type u -> Type (huniv u)
- let equal s1 s2 =
+ let eq s1 s2 =
s1 == s2 ||
match (s1,s2) with
(Prop c1, Prop c2) -> c1 == c2
diff --git a/kernel/constr.mli b/kernel/constr.mli
index e6a3e71f89..f76b5ae4f0 100644
--- a/kernel/constr.mli
+++ b/kernel/constr.mli
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -30,13 +30,23 @@ type case_printing =
cstr_tags : bool list array; (** tell whether letin or lambda in the signature of each constructor *)
style : case_style }
-(** the integer is the number of real args, needed for reduction *)
+(* INVARIANT:
+ * - Array.length ci_cstr_ndecls = Array.length ci_cstr_nargs
+ * - forall (i : 0 .. pred (Array.length ci_cstr_ndecls)),
+ * ci_cstr_ndecls.(i) >= ci_cstr_nargs.(i)
+ *)
type case_info =
- { ci_ind : inductive;
- ci_npar : int;
- ci_cstr_ndecls : int array; (* number of pattern vars of each constructor (with let's)*)
- ci_cstr_nargs : int array; (* number of pattern vars of each constructor (w/o let's) *)
- ci_pp_info : case_printing (** not interpreted by the kernel *)
+ { ci_ind : inductive; (* inductive type to which belongs the value that is being matched *)
+ ci_npar : int; (* number of parameters of the above inductive type *)
+ ci_cstr_ndecls : int array; (* For each constructor, the corresponding integer determines
+ the number of values that can be bound in a match-construct.
+ NOTE: parameters of the inductive type are therefore excluded from the count *)
+ ci_cstr_nargs : int array; (* for each constructor, the corresponding integers determines
+ the number of values that can be applied to the constructor,
+ in addition to the parameters of the related inductive type
+ NOTE: "lets" are therefore excluded from the count
+ NOTE: parameters of the inductive type are also excluded from the count *)
+ ci_pp_info : case_printing (* not interpreted by the kernel *)
}
(** {6 The type of constructions } *)
@@ -93,8 +103,9 @@ val mkLambda : Name.t * types * constr -> constr
(** Constructs the product [let x = t1 : t2 in t3] *)
val mkLetIn : Name.t * constr * types * constr -> constr
-(** [mkApp (f,[| t_1; ...; t_n |]] constructs the application
- {% $(f~t_1~\dots~t_n)$ %}. *)
+(** [mkApp (f, [|t1; ...; tN|]] constructs the application
+ {%html:(f t<sub>1</sub> ... t<sub>n</sub>)%}
+ {%latex:$(f~t_1\dots f_n)$%}. *)
val mkApp : constr * constr array -> constr
val map_puniverses : ('a -> 'b) -> 'a puniverses -> 'b puniverses
@@ -181,10 +192,13 @@ type ('constr, 'types) kind_of_term =
| Evar of 'constr pexistential
| Sort of Sorts.t
| Cast of 'constr * cast_kind * 'types
- | Prod of Name.t * 'types * 'types
- | Lambda of Name.t * 'types * 'constr
- | LetIn of Name.t * 'constr * 'types * 'constr
- | App of 'constr * 'constr array
+ | Prod of Name.t * 'types * 'types (** Concrete syntax ["forall A:B,C"] is represented as [Prod (A,B,C)]. *)
+ | Lambda of Name.t * 'types * 'constr (** Concrete syntax ["fun A:B => C"] is represented as [Lambda (A,B,C)]. *)
+ | LetIn of Name.t * 'constr * 'types * 'constr (** Concrete syntax ["let A:B := C in D"] is represented as [LetIn (A,B,C,D)]. *)
+ | App of 'constr * 'constr array (** Concrete syntax ["(F P1 P2 ... Pn)"] is represented as [App (F, [|P1; P2; ...; Pn|])].
+ The {!mkApp} constructor also enforces the following invariant:
+ - [F] itself is not {!App}
+ - and [[|P1;..;Pn|]] is not empty. *)
| Const of constant puniverses
| Ind of inductive puniverses
| Construct of constructor puniverses
@@ -205,19 +219,19 @@ val equal : constr -> constr -> bool
(** [eq_constr_univs u a b] is [true] if [a] equals [b] modulo alpha, casts,
application grouping and the universe equalities in [u]. *)
-val eq_constr_univs : constr Univ.check_function
+val eq_constr_univs : constr UGraph.check_function
(** [leq_constr_univs u a b] is [true] if [a] is convertible to [b] modulo
alpha, casts, application grouping and the universe inequalities in [u]. *)
-val leq_constr_univs : constr Univ.check_function
+val leq_constr_univs : constr UGraph.check_function
(** [eq_constr_univs u a b] is [true] if [a] equals [b] modulo alpha, casts,
application grouping and the universe equalities in [u]. *)
-val eq_constr_univs_infer : Univ.universes -> constr -> constr -> bool Univ.constrained
+val eq_constr_univs_infer : UGraph.t -> constr -> constr -> bool Univ.constrained
(** [leq_constr_univs u a b] is [true] if [a] is convertible to [b] modulo
alpha, casts, application grouping and the universe inequalities in [u]. *)
-val leq_constr_univs_infer : Univ.universes -> constr -> constr -> bool Univ.constrained
+val leq_constr_univs_infer : UGraph.t -> constr -> constr -> bool Univ.constrained
(** [eq_constr_univs a b] [true, c] if [a] equals [b] modulo alpha, casts,
application grouping and ignoring universe instances. *)
diff --git a/kernel/context.ml b/kernel/context.ml
index 796f06d37e..4e53b73a28 100644
--- a/kernel/context.ml
+++ b/kernel/context.ml
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -15,123 +15,409 @@
(* This file defines types and combinators regarding indexes-based and
names-based contexts *)
-open Util
-open Names
-
-(***************************************************************************)
-(* Type of assumptions *)
-(***************************************************************************)
-
-type named_declaration = Id.t * Constr.t option * Constr.t
-type named_list_declaration = Id.t list * Constr.t option * Constr.t
-type rel_declaration = Name.t * Constr.t option * Constr.t
-
-let map_named_declaration_skel f (id, (v : Constr.t option), ty) =
- (id, Option.map f v, f ty)
-let map_named_list_declaration = map_named_declaration_skel
-let map_named_declaration = map_named_declaration_skel
-
-let map_rel_declaration = map_named_declaration
-
-let fold_named_declaration f (_, v, ty) a = f ty (Option.fold_right f v a)
-let fold_rel_declaration = fold_named_declaration
-
-let exists_named_declaration f (_, v, ty) = Option.cata f false v || f ty
-let exists_rel_declaration f (_, v, ty) = Option.cata f false v || f ty
-
-let for_all_named_declaration f (_, v, ty) = Option.cata f true v && f ty
-let for_all_rel_declaration f (_, v, ty) = Option.cata f true v && f ty
-
-let eq_named_declaration (i1, c1, t1) (i2, c2, t2) =
- Id.equal i1 i2 && Option.equal Constr.equal c1 c2 && Constr.equal t1 t2
-
-let eq_rel_declaration (n1, c1, t1) (n2, c2, t2) =
- Name.equal n1 n2 && Option.equal Constr.equal c1 c2 && Constr.equal t1 t2
-
-(***************************************************************************)
-(* Type of local contexts (telescopes) *)
-(***************************************************************************)
-
-(*s Signatures of ordered optionally named variables, intended to be
- accessed by de Bruijn indices (to represent bound variables) *)
-
-type rel_context = rel_declaration list
-
-let empty_rel_context = []
-
-let add_rel_decl d ctxt = d::ctxt
+(** The modules defined below represent a {e local context}
+ as defined by Chapter 4 in the Reference Manual:
-let rec lookup_rel n sign =
- match n, sign with
- | 1, decl :: _ -> decl
- | n, _ :: sign -> lookup_rel (n-1) sign
- | _, [] -> raise Not_found
+ A {e local context} is an ordered list of of {e local declarations}
+ of names that we call {e variables}.
-let rel_context_length = List.length
+ A {e local declaration} of some variable can be either:
+ - a {e local assumption}, or
+ - a {e local definition}.
+*)
-let rel_context_nhyps hyps =
- let rec nhyps acc = function
- | [] -> acc
- | (_,None,_)::hyps -> nhyps (1+acc) hyps
- | (_,Some _,_)::hyps -> nhyps acc hyps in
- nhyps 0 hyps
-
-let rel_context_tags ctx =
- let rec aux l = function
- | [] -> l
- | (_,Some _,_)::ctx -> aux (true::l) ctx
- | (_,None,_)::ctx -> aux (false::l) ctx
- in aux [] ctx
-
-(*s Signatures of named hypotheses. Used for section variables and
- goal assumptions. *)
-
-type named_context = named_declaration list
-type named_list_context = named_list_declaration list
-
-let empty_named_context = []
-
-let add_named_decl d sign = d::sign
-
-let rec lookup_named id = function
- | (id',_,_ as decl) :: _ when Id.equal id id' -> decl
- | _ :: sign -> lookup_named id sign
- | [] -> raise Not_found
-
-let named_context_length = List.length
-let named_context_equal = List.equal eq_named_declaration
-
-let vars_of_named_context ctx =
- List.fold_left (fun accu (id, _, _) -> Id.Set.add id accu) Id.Set.empty ctx
-
-let instance_from_named_context sign =
- let filter = function
- | (id, None, _) -> Some (Constr.mkVar id)
- | (_, Some _, _) -> None
- in
- List.map_filter filter sign
-
-let fold_named_context f l ~init = List.fold_right f l init
-let fold_named_list_context f l ~init = List.fold_right f l init
-let fold_named_context_reverse f ~init l = List.fold_left f init l
-
-(*s Signatures of ordered section variables *)
-type section_context = named_context
-
-let fold_rel_context f l ~init:x = List.fold_right f l x
-let fold_rel_context_reverse f ~init:x l = List.fold_left f x l
-
-let map_context f l =
- let map_decl (n, body_o, typ as decl) =
- let body_o' = Option.smartmap f body_o in
- let typ' = f typ in
- if body_o' == body_o && typ' == typ then decl else
- (n, body_o', typ')
- in
- List.smartmap map_decl l
-
-let map_rel_context = map_context
-let map_named_context = map_context
+open Util
+open Names
-let iter_rel_context f = List.iter (fun (_,b,t) -> f t; Option.iter f b)
-let iter_named_context f = List.iter (fun (_,b,t) -> f t; Option.iter f b)
+(** Representation of contexts that can capture anonymous as well as non-anonymous variables.
+ Individual declarations are then designated by de Bruijn indexes. *)
+module Rel =
+struct
+ (** Representation of {e local declarations}. *)
+ module Declaration =
+ struct
+ (* local declaration *)
+ type t =
+ | LocalAssum of Name.t * Constr.t (** name, type *)
+ | LocalDef of Name.t * Constr.t * Constr.t (** name, value, type *)
+
+ (** Return the name bound by a given declaration. *)
+ let get_name = function
+ | LocalAssum (na,_)
+ | LocalDef (na,_,_) -> na
+
+ (** Return [Some value] for local-declarations and [None] for local-assumptions. *)
+ let get_value = function
+ | LocalAssum _ -> None
+ | LocalDef (_,v,_) -> Some v
+
+ (** Return the type of the name bound by a given declaration. *)
+ let get_type = function
+ | LocalAssum (_,ty)
+ | LocalDef (_,_,ty) -> ty
+
+ (** Set the name that is bound by a given declaration. *)
+ let set_name na = function
+ | LocalAssum (_,ty) -> LocalAssum (na, ty)
+ | LocalDef (_,v,ty) -> LocalDef (na, v, ty)
+
+ (** Set the type of the bound variable in a given declaration. *)
+ let set_type ty = function
+ | LocalAssum (na,_) -> LocalAssum (na, ty)
+ | LocalDef (na,v,_) -> LocalDef (na, v, ty)
+
+ (** Return [true] iff a given declaration is a local assumption. *)
+ let is_local_assum = function
+ | LocalAssum _ -> true
+ | LocalDef _ -> false
+
+ (** Return [true] iff a given declaration is a local definition. *)
+ let is_local_def = function
+ | LocalAssum _ -> false
+ | LocalDef _ -> true
+
+ (** Check whether any term in a given declaration satisfies a given predicate. *)
+ let exists f = function
+ | LocalAssum (_, ty) -> f ty
+ | LocalDef (_, v, ty) -> f v || f ty
+
+ (** Check whether all terms in a given declaration satisfy a given predicate. *)
+ let for_all f = function
+ | LocalAssum (_, ty) -> f ty
+ | LocalDef (_, v, ty) -> f v && f ty
+
+ (** Check whether the two given declarations are equal. *)
+ let equal decl1 decl2 =
+ match decl1, decl2 with
+ | LocalAssum (n1,ty1), LocalAssum (n2, ty2) ->
+ Name.equal n1 n2 && Constr.equal ty1 ty2
+ | LocalDef (n1,v1,ty1), LocalDef (n2,v2,ty2) ->
+ Name.equal n1 n2 && Constr.equal v1 v2 && Constr.equal ty1 ty2
+ | _ ->
+ false
+
+ (** Map the name bound by a given declaration. *)
+ let map_name f = function
+ | LocalAssum (na, ty) as decl ->
+ let na' = f na in
+ if na == na' then decl else LocalAssum (na', ty)
+ | LocalDef (na, v, ty) as decl ->
+ let na' = f na in
+ if na == na' then decl else LocalDef (na', v, ty)
+
+ (** For local assumptions, this function returns the original local assumptions.
+ For local definitions, this function maps the value in the local definition. *)
+ let map_value f = function
+ | LocalAssum _ as decl -> decl
+ | LocalDef (na, v, t) as decl ->
+ let v' = f v in
+ if v == v' then decl else LocalDef (na, v', t)
+
+ (** Map the type of the name bound by a given declaration. *)
+ let map_type f = function
+ | LocalAssum (na, ty) as decl ->
+ let ty' = f ty in
+ if ty == ty' then decl else LocalAssum (na, ty')
+ | LocalDef (na, v, ty) as decl ->
+ let ty' = f ty in
+ if ty == ty' then decl else LocalDef (na, v, ty')
+
+ (** Map all terms in a given declaration. *)
+ let map_constr f = function
+ | LocalAssum (na, ty) as decl ->
+ let ty' = f ty in
+ if ty == ty' then decl else LocalAssum (na, ty')
+ | LocalDef (na, v, ty) as decl ->
+ let v' = f v in
+ let ty' = f ty in
+ if v == v' && ty == ty' then decl else LocalDef (na, v', ty')
+
+ (** Perform a given action on all terms in a given declaration. *)
+ let iter_constr f = function
+ | LocalAssum (_,ty) -> f ty
+ | LocalDef (_,v,ty) -> f v; f ty
+
+ (** Reduce all terms in a given declaration to a single value. *)
+ let fold f decl acc =
+ match decl with
+ | LocalAssum (n,ty) -> f ty acc
+ | LocalDef (n,v,ty) -> f ty (f v acc)
+
+ let to_tuple = function
+ | LocalAssum (na, ty) -> na, None, ty
+ | LocalDef (na, v, ty) -> na, Some v, ty
+
+ let of_tuple = function
+ | n, None, ty -> LocalAssum (n,ty)
+ | n, Some v, ty -> LocalDef (n,v,ty)
+ end
+
+ (** Rel-context is represented as a list of declarations.
+ Inner-most declarations are at the beginning of the list.
+ Outer-most declarations are at the end of the list. *)
+ type t = Declaration.t list
+
+ (** empty rel-context *)
+ let empty = []
+
+ (** Return a new rel-context enriched by with a given inner-most declaration. *)
+ let add d ctx = d :: ctx
+
+ (** Return the number of {e local declarations} in a given context. *)
+ let length = List.length
+
+ (** [extended_rel_list n Γ] builds an instance [args] such that [Γ,Δ ⊢ args:Γ]
+ with n = |Δ| and with the local definitions of [Γ] skipped in
+ [args]. Example: for [x:T,y:=c,z:U] and [n]=2, it gives [Rel 5, Rel 3]. *)
+ let nhyps =
+ let open Declaration in
+ let rec nhyps acc = function
+ | [] -> acc
+ | LocalAssum _ :: hyps -> nhyps (succ acc) hyps
+ | LocalDef _ :: hyps -> nhyps acc hyps
+ in
+ nhyps 0
+
+ (** Return a declaration designated by a given de Bruijn index.
+ @raise Not_found if the designated de Bruijn index is not present in the designated rel-context. *)
+ let rec lookup n ctx =
+ match n, ctx with
+ | 1, decl :: _ -> decl
+ | n, _ :: sign -> lookup (n-1) sign
+ | _, [] -> raise Not_found
+
+ (** Check whether given two rel-contexts are equal. *)
+ let equal = List.equal Declaration.equal
+
+ (** Map all terms in a given rel-context. *)
+ let map f = List.smartmap (Declaration.map_constr f)
+
+ (** Perform a given action on every declaration in a given rel-context. *)
+ let iter f = List.iter (Declaration.iter_constr f)
+
+ (** Reduce all terms in a given rel-context to a single value.
+ Innermost declarations are processed first. *)
+ let fold_inside f ~init = List.fold_left f init
+
+ (** Reduce all terms in a given rel-context to a single value.
+ Outermost declarations are processed first. *)
+ let fold_outside f l ~init = List.fold_right f l init
+
+ (** Map a given rel-context to a list where each {e local assumption} is mapped to [true]
+ and each {e local definition} is mapped to [false]. *)
+ let to_tags =
+ let rec aux l = function
+ | [] -> l
+ | Declaration.LocalDef _ :: ctx -> aux (true::l) ctx
+ | Declaration.LocalAssum _ :: ctx -> aux (false::l) ctx
+ in aux []
+
+ (** [extended_list n Γ] builds an instance [args] such that [Γ,Δ ⊢ args:Γ]
+ with n = |Δ| and with the {e local definitions} of [Γ] skipped in
+ [args]. Example: for [x:T, y:=c, z:U] and [n]=2, it gives [Rel 5, Rel 3]. *)
+ let to_extended_list n =
+ let rec reln l p = function
+ | Declaration.LocalAssum _ :: hyps -> reln (Constr.mkRel (n+p) :: l) (p+1) hyps
+ | Declaration.LocalDef _ :: hyps -> reln l (p+1) hyps
+ | [] -> l
+ in
+ reln [] 1
+
+ (** [extended_vect n Γ] does the same, returning instead an array. *)
+ let to_extended_vect n hyps = Array.of_list (to_extended_list n hyps)
+end
+
+(** This module represents contexts that can capture non-anonymous variables.
+ Individual declarations are then designated by the identifiers they bind. *)
+module Named =
+struct
+ (** Representation of {e local declarations}. *)
+ module Declaration =
+ struct
+ (** local declaration *)
+ type t =
+ | LocalAssum of Id.t * Constr.t (** identifier, type *)
+ | LocalDef of Id.t * Constr.t * Constr.t (** identifier, value, type *)
+
+ (** Return the identifier bound by a given declaration. *)
+ let get_id = function
+ | LocalAssum (id,_) -> id
+ | LocalDef (id,_,_) -> id
+
+ (** Return [Some value] for local-declarations and [None] for local-assumptions. *)
+ let get_value = function
+ | LocalAssum _ -> None
+ | LocalDef (_,v,_) -> Some v
+
+ (** Return the type of the name bound by a given declaration. *)
+ let get_type = function
+ | LocalAssum (_,ty)
+ | LocalDef (_,_,ty) -> ty
+
+ (** Set the identifier that is bound by a given declaration. *)
+ let set_id id = function
+ | LocalAssum (_,ty) -> LocalAssum (id, ty)
+ | LocalDef (_, v, ty) -> LocalDef (id, v, ty)
+
+ (** Set the type of the bound variable in a given declaration. *)
+ let set_type ty = function
+ | LocalAssum (id,_) -> LocalAssum (id, ty)
+ | LocalDef (id,v,_) -> LocalDef (id, v, ty)
+
+ (** Return [true] iff a given declaration is a local assumption. *)
+ let is_local_assum = function
+ | LocalAssum _ -> true
+ | LocalDef _ -> false
+
+ (** Return [true] iff a given declaration is a local definition. *)
+ let is_local_def = function
+ | LocalDef _ -> true
+ | LocalAssum _ -> false
+
+ (** Check whether any term in a given declaration satisfies a given predicate. *)
+ let exists f = function
+ | LocalAssum (_, ty) -> f ty
+ | LocalDef (_, v, ty) -> f v || f ty
+
+ (** Check whether all terms in a given declaration satisfy a given predicate. *)
+ let for_all f = function
+ | LocalAssum (_, ty) -> f ty
+ | LocalDef (_, v, ty) -> f v && f ty
+
+ (** Check whether the two given declarations are equal. *)
+ let equal decl1 decl2 =
+ match decl1, decl2 with
+ | LocalAssum (id1, ty1), LocalAssum (id2, ty2) ->
+ Id.equal id1 id2 && Constr.equal ty1 ty2
+ | LocalDef (id1, v1, ty1), LocalDef (id2, v2, ty2) ->
+ Id.equal id1 id2 && Constr.equal v1 v2 && Constr.equal ty1 ty2
+ | _ ->
+ false
+
+ (** Map the identifier bound by a given declaration. *)
+ let map_id f = function
+ | LocalAssum (id, ty) as decl ->
+ let id' = f id in
+ if id == id' then decl else LocalAssum (id', ty)
+ | LocalDef (id, v, ty) as decl ->
+ let id' = f id in
+ if id == id' then decl else LocalDef (id', v, ty)
+
+ (** For local assumptions, this function returns the original local assumptions.
+ For local definitions, this function maps the value in the local definition. *)
+ let map_value f = function
+ | LocalAssum _ as decl -> decl
+ | LocalDef (na, v, t) as decl ->
+ let v' = f v in
+ if v == v' then decl else LocalDef (na, v', t)
+
+ (** Map the type of the name bound by a given declaration. *)
+ let map_type f = function
+ | LocalAssum (id, ty) as decl ->
+ let ty' = f ty in
+ if ty == ty' then decl else LocalAssum (id, ty')
+ | LocalDef (id, v, ty) as decl ->
+ let ty' = f ty in
+ if ty == ty' then decl else LocalDef (id, v, ty')
+
+ (** Map all terms in a given declaration. *)
+ let map_constr f = function
+ | LocalAssum (id, ty) as decl ->
+ let ty' = f ty in
+ if ty == ty' then decl else LocalAssum (id, ty')
+ | LocalDef (id, v, ty) as decl ->
+ let v' = f v in
+ let ty' = f ty in
+ if v == v' && ty == ty' then decl else LocalDef (id, v', ty')
+
+ (** Perform a given action on all terms in a given declaration. *)
+ let iter_constr f = function
+ | LocalAssum (_, ty) -> f ty
+ | LocalDef (_, v, ty) -> f v; f ty
+
+ (** Reduce all terms in a given declaration to a single value. *)
+ let fold f decl a =
+ match decl with
+ | LocalAssum (_, ty) -> f ty a
+ | LocalDef (_, v, ty) -> a |> f v |> f ty
+
+ let to_tuple = function
+ | LocalAssum (id, ty) -> id, None, ty
+ | LocalDef (id, v, ty) -> id, Some v, ty
+
+ let of_tuple = function
+ | id, None, ty -> LocalAssum (id, ty)
+ | id, Some v, ty -> LocalDef (id, v, ty)
+ end
+
+ (** Named-context is represented as a list of declarations.
+ Inner-most declarations are at the beginning of the list.
+ Outer-most declarations are at the end of the list. *)
+ type t = Declaration.t list
+
+ (** empty named-context *)
+ let empty = []
+
+ (** empty named-context *)
+ let add d ctx = d :: ctx
+
+ (** Return the number of {e local declarations} in a given named-context. *)
+ let length = List.length
+
+(** Return a declaration designated by a given de Bruijn index.
+ @raise Not_found if the designated identifier is not present in the designated named-context. *) let rec lookup id = function
+ | decl :: _ when Id.equal id (Declaration.get_id decl) -> decl
+ | _ :: sign -> lookup id sign
+ | [] -> raise Not_found
+
+ (** Check whether given two named-contexts are equal. *)
+ let equal = List.equal Declaration.equal
+
+ (** Map all terms in a given named-context. *)
+ let map f = List.smartmap (Declaration.map_constr f)
+
+ (** Perform a given action on every declaration in a given named-context. *)
+ let iter f = List.iter (Declaration.iter_constr f)
+
+ (** Reduce all terms in a given named-context to a single value.
+ Innermost declarations are processed first. *)
+ let fold_inside f ~init = List.fold_left f init
+
+ (** Reduce all terms in a given named-context to a single value.
+ Outermost declarations are processed first. *)
+ let fold_outside f l ~init = List.fold_right f l init
+
+ (** Return the set of all identifiers bound in a given named-context. *)
+ let to_vars =
+ List.fold_left (fun accu decl -> Id.Set.add (Declaration.get_id decl) accu) Id.Set.empty
+
+ (** [instance_from_named_context Ω] builds an instance [args] such
+ that [Ω ⊢ args:Ω] where [Ω] is a named context and with the local
+ definitions of [Ω] skipped. Example: for [id1:T,id2:=c,id3:U], it
+ gives [Var id1, Var id3]. All [idj] are supposed distinct. *)
+ let to_instance =
+ let filter = function
+ | Declaration.LocalAssum (id, _) -> Some (Constr.mkVar id)
+ | _ -> None
+ in
+ List.map_filter filter
+ end
+
+module NamedList =
+ struct
+ module Declaration =
+ struct
+ type t = Id.t list * Constr.t option * Constr.t
+
+ let map_constr f (ids, copt, ty as decl) =
+ let copt' = Option.map f copt in
+ let ty' = f ty in
+ if copt == copt' && ty == ty' then decl else (ids, copt', ty')
+ end
+
+ type t = Declaration.t list
+
+ let fold f l ~init = List.fold_right f l init
+ end
+
+type section_context = Named.t
diff --git a/kernel/context.mli b/kernel/context.mli
index 5279aefb6b..b5f3904d22 100644
--- a/kernel/context.mli
+++ b/kernel/context.mli
@@ -1,122 +1,260 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
+(** The modules defined below represent a {e local context}
+ as defined by Chapter 4 in the Reference Manual:
+
+ A {e local context} is an ordered list of of {e local declarations}
+ of names that we call {e variables}.
+
+ A {e local declaration} of some variable can be either:
+ - a {e local assumption}, or
+ - a {e local definition}.
+
+ {e Local assumptions} are denoted in the Reference Manual as [(name : typ)] and
+ {e local definitions} are there denoted as [(name := value : typ)].
+*)
+
open Names
-(** TODO: cleanup *)
+(** Representation of contexts that can capture anonymous as well as non-anonymous variables.
+ Individual declarations are then designated by de Bruijn indexes. *)
+module Rel :
+sig
+ module Declaration :
+ sig
+ (* local declaration *)
+ type t = LocalAssum of Name.t * Constr.t (** name, type *)
+ | LocalDef of Name.t * Constr.t * Constr.t (** name, value, type *)
+
+ (** Return the name bound by a given declaration. *)
+ val get_name : t -> Name.t
+
+ (** Return [Some value] for local-declarations and [None] for local-assumptions. *)
+ val get_value : t -> Constr.t option
+
+ (** Return the type of the name bound by a given declaration. *)
+ val get_type : t -> Constr.t
+
+ (** Set the name that is bound by a given declaration. *)
+ val set_name : Name.t -> t -> t
+
+ (** Set the type of the bound variable in a given declaration. *)
+ val set_type : Constr.t -> t -> t
+
+ (** Return [true] iff a given declaration is a local assumption. *)
+ val is_local_assum : t -> bool
+
+ (** Return [true] iff a given declaration is a local definition. *)
+ val is_local_def : t -> bool
+
+ (** Check whether any term in a given declaration satisfies a given predicate. *)
+ val exists : (Constr.t -> bool) -> t -> bool
+
+ (** Check whether all terms in a given declaration satisfy a given predicate. *)
+ val for_all : (Constr.t -> bool) -> t -> bool
+
+ (** Check whether the two given declarations are equal. *)
+ val equal : t -> t -> bool
+
+ (** Map the name bound by a given declaration. *)
+ val map_name : (Name.t -> Name.t) -> t -> t
+
+ (** For local assumptions, this function returns the original local assumptions.
+ For local definitions, this function maps the value in the local definition. *)
+ val map_value : (Constr.t -> Constr.t) -> t -> t
+
+ (** Map the type of the name bound by a given declaration. *)
+ val map_type : (Constr.t -> Constr.t) -> t -> t
+
+ (** Map all terms in a given declaration. *)
+ val map_constr : (Constr.t -> Constr.t) -> t -> t
+
+ (** Perform a given action on all terms in a given declaration. *)
+ val iter_constr : (Constr.t -> unit) -> t -> unit
+
+ (** Reduce all terms in a given declaration to a single value. *)
+ val fold : (Constr.t -> 'a -> 'a) -> t -> 'a -> 'a
+
+ val to_tuple : t -> Name.t * Constr.t option * Constr.t
+ val of_tuple : Name.t * Constr.t option * Constr.t -> t
+ end
+
+ (** Rel-context is represented as a list of declarations.
+ Inner-most declarations are at the beginning of the list.
+ Outer-most declarations are at the end of the list. *)
+ type t = Declaration.t list
+
+ (** empty rel-context *)
+ val empty : t
+
+ (** Return a new rel-context enriched by with a given inner-most declaration. *)
+ val add : Declaration.t -> t -> t
+
+ (** Return the number of {e local declarations} in a given context. *)
+ val length : t -> int
+
+ (** Check whether given two rel-contexts are equal. *)
+ val equal : t -> t -> bool
+
+ (** Return the number of {e local assumptions} in a given rel-context. *)
+ val nhyps : t -> int
+
+ (** Return a declaration designated by a given de Bruijn index.
+ @raise Not_found if the designated de Bruijn index outside the range. *)
+ val lookup : int -> t -> Declaration.t
+
+ (** Map all terms in a given rel-context. *)
+ val map : (Constr.t -> Constr.t) -> t -> t
+
+ (** Perform a given action on every declaration in a given rel-context. *)
+ val iter : (Constr.t -> unit) -> t -> unit
+
+ (** Reduce all terms in a given rel-context to a single value.
+ Innermost declarations are processed first. *)
+ val fold_inside : ('a -> Declaration.t -> 'a) -> init:'a -> t -> 'a
+
+ (** Reduce all terms in a given rel-context to a single value.
+ Outermost declarations are processed first. *)
+ val fold_outside : (Declaration.t -> 'a -> 'a) -> t -> init:'a -> 'a
+
+ (** Map a given rel-context to a list where each {e local assumption} is mapped to [true]
+ and each {e local definition} is mapped to [false]. *)
+ val to_tags : t -> bool list
+
+ (** [extended_list n Γ] builds an instance [args] such that [Γ,Δ ⊢ args:Γ]
+ with n = |Δ| and with the {e local definitions} of [Γ] skipped in
+ [args]. Example: for [x:T, y:=c, z:U] and [n]=2, it gives [Rel 5, Rel 3]. *)
+ val to_extended_list : int -> t -> Constr.t list
+
+ (** [extended_vect n Γ] does the same, returning instead an array. *)
+ val to_extended_vect : int -> t -> Constr.t array
+end
+
+(** This module represents contexts that can capture non-anonymous variables.
+ Individual declarations are then designated by the identifiers they bind. *)
+module Named :
+sig
+ (** Representation of {e local declarations}. *)
+ module Declaration :
+ sig
+ type t = LocalAssum of Id.t * Constr.t (** identifier, type *)
+ | LocalDef of Id.t * Constr.t * Constr.t (** identifier, value, type *)
+
+ (** Return the identifier bound by a given declaration. *)
+ val get_id : t -> Id.t
+
+ (** Return [Some value] for local-declarations and [None] for local-assumptions. *)
+ val get_value : t -> Constr.t option
+
+ (** Return the type of the name bound by a given declaration. *)
+ val get_type : t -> Constr.t
+
+ (** Set the identifier that is bound by a given declaration. *)
+ val set_id : Id.t -> t -> t
+
+ (** Set the type of the bound variable in a given declaration. *)
+ val set_type : Constr.t -> t -> t
-(** {6 Declarations} *)
-(** A {e declaration} has the form [(name,body,type)]. It is either an
- {e assumption} if [body=None] or a {e definition} if
- [body=Some actualbody]. It is referred by {e name} if [na] is an
- identifier or by {e relative index} if [na] is not an identifier
- (in the latter case, [na] is of type [name] but just for printing
- purpose) *)
+ (** Return [true] iff a given declaration is a local assumption. *)
+ val is_local_assum : t -> bool
-type named_declaration = Id.t * Constr.t option * Constr.t
-type named_list_declaration = Id.t list * Constr.t option * Constr.t
-type rel_declaration = Name.t * Constr.t option * Constr.t
+ (** Return [true] iff a given declaration is a local definition. *)
+ val is_local_def : t -> bool
-val map_named_declaration :
- (Constr.t -> Constr.t) -> named_declaration -> named_declaration
-val map_named_list_declaration :
- (Constr.t -> Constr.t) -> named_list_declaration -> named_list_declaration
-val map_rel_declaration :
- (Constr.t -> Constr.t) -> rel_declaration -> rel_declaration
+ (** Check whether any term in a given declaration satisfies a given predicate. *)
+ val exists : (Constr.t -> bool) -> t -> bool
-val fold_named_declaration :
- (Constr.t -> 'a -> 'a) -> named_declaration -> 'a -> 'a
-val fold_rel_declaration :
- (Constr.t -> 'a -> 'a) -> rel_declaration -> 'a -> 'a
+ (** Check whether all terms in a given declaration satisfy a given predicate. *)
+ val for_all : (Constr.t -> bool) -> t -> bool
-val exists_named_declaration :
- (Constr.t -> bool) -> named_declaration -> bool
-val exists_rel_declaration :
- (Constr.t -> bool) -> rel_declaration -> bool
+ (** Check whether the two given declarations are equal. *)
+ val equal : t -> t -> bool
-val for_all_named_declaration :
- (Constr.t -> bool) -> named_declaration -> bool
-val for_all_rel_declaration :
- (Constr.t -> bool) -> rel_declaration -> bool
+ (** Map the identifier bound by a given declaration. *)
+ val map_id : (Id.t -> Id.t) -> t -> t
-val eq_named_declaration :
- named_declaration -> named_declaration -> bool
+ (** For local assumptions, this function returns the original local assumptions.
+ For local definitions, this function maps the value in the local definition. *)
+ val map_value : (Constr.t -> Constr.t) -> t -> t
-val eq_rel_declaration :
- rel_declaration -> rel_declaration -> bool
+ (** Map the type of the name bound by a given declaration. *)
+ val map_type : (Constr.t -> Constr.t) -> t -> t
-(** {6 Signatures of ordered named declarations } *)
+ (** Map all terms in a given declaration. *)
+ val map_constr : (Constr.t -> Constr.t) -> t -> t
-type named_context = named_declaration list
-type section_context = named_context
-type named_list_context = named_list_declaration list
-type rel_context = rel_declaration list
-(** In [rel_context], more recent declaration is on top *)
+ (** Perform a given action on all terms in a given declaration. *)
+ val iter_constr : (Constr.t -> unit) -> t -> unit
-val empty_named_context : named_context
-val add_named_decl : named_declaration -> named_context -> named_context
-val vars_of_named_context : named_context -> Id.Set.t
+ (** Reduce all terms in a given declaration to a single value. *)
+ val fold : (Constr.t -> 'a -> 'a) -> t -> 'a -> 'a
-val lookup_named : Id.t -> named_context -> named_declaration
+ val to_tuple : t -> Id.t * Constr.t option * Constr.t
+ val of_tuple : Id.t * Constr.t option * Constr.t -> t
+ end
-(** number of declarations *)
-val named_context_length : named_context -> int
+ (** Rel-context is represented as a list of declarations.
+ Inner-most declarations are at the beginning of the list.
+ Outer-most declarations are at the end of the list. *)
+ type t = Declaration.t list
-(** named context equality *)
-val named_context_equal : named_context -> named_context -> bool
+ (** empty named-context *)
+ val empty : t
-(** {6 Recurrence on [named_context]: older declarations processed first } *)
-val fold_named_context :
- (named_declaration -> 'a -> 'a) -> named_context -> init:'a -> 'a
+ (** Return a new rel-context enriched by with a given inner-most declaration. *)
+ val add : Declaration.t -> t -> t
-val fold_named_list_context :
- (named_list_declaration -> 'a -> 'a) -> named_list_context -> init:'a -> 'a
+ (** Return the number of {e local declarations} in a given named-context. *)
+ val length : t -> int
-(** newer declarations first *)
-val fold_named_context_reverse :
- ('a -> named_declaration -> 'a) -> init:'a -> named_context -> 'a
+ (** Return a declaration designated by an identifier of the variable bound in that declaration.
+ @raise Not_found if the designated identifier is not bound in a given named-context. *)
+ val lookup : Id.t -> t -> Declaration.t
-(** {6 Section-related auxiliary functions } *)
-val instance_from_named_context : named_context -> Constr.t list
+ (** Check whether given two rel-contexts are equal. *)
+ val equal : t -> t -> bool
-(** {6 ... } *)
-(** Signatures of ordered optionally named variables, intended to be
- accessed by de Bruijn indices *)
+ (** Map all terms in a given named-context. *)
+ val map : (Constr.t -> Constr.t) -> t -> t
-(** {6 Recurrence on [rel_context]: older declarations processed first } *)
-val fold_rel_context :
- (rel_declaration -> 'a -> 'a) -> rel_context -> init:'a -> 'a
+ (** Perform a given action on every declaration in a given named-context. *)
+ val iter : (Constr.t -> unit) -> t -> unit
-(** newer declarations first *)
-val fold_rel_context_reverse :
- ('a -> rel_declaration -> 'a) -> init:'a -> rel_context -> 'a
+ (** Reduce all terms in a given named-context to a single value.
+ Innermost declarations are processed first. *)
+ val fold_inside : ('a -> Declaration.t -> 'a) -> init:'a -> t -> 'a
-(** {6 Map function of [rel_context] } *)
-val map_rel_context : (Constr.t -> Constr.t) -> rel_context -> rel_context
+ (** Reduce all terms in a given named-context to a single value.
+ Outermost declarations are processed first. *)
+ val fold_outside : (Declaration.t -> 'a -> 'a) -> t -> init:'a -> 'a
-(** {6 Map function of [named_context] } *)
-val map_named_context : (Constr.t -> Constr.t) -> named_context -> named_context
+ (** Return the set of all identifiers bound in a given named-context. *)
+ val to_vars : t -> Id.Set.t
-(** {6 Map function of [rel_context] } *)
-val iter_rel_context : (Constr.t -> unit) -> rel_context -> unit
+ (** [instance_from_named_context Ω] builds an instance [args] such
+ that [Ω ⊢ args:Ω] where [Ω] is a named context and with the local
+ definitions of [Ω] skipped. Example: for [id1:T,id2:=c,id3:U], it
+ gives [Var id1, Var id3]. All [idj] are supposed distinct. *)
+ val to_instance : t -> Constr.t list
+end
-(** {6 Map function of [named_context] } *)
-val iter_named_context : (Constr.t -> unit) -> named_context -> unit
+module NamedList :
+sig
+ module Declaration :
+ sig
+ type t = Id.t list * Constr.t option * Constr.t
+ val map_constr : (Constr.t -> Constr.t) -> t -> t
+ end
-(** {6 Contexts of declarations referred to by de Bruijn indices } *)
+ type t = Declaration.t list
-val empty_rel_context : rel_context
-val add_rel_decl : rel_declaration -> rel_context -> rel_context
+ val fold : (Declaration.t -> 'a -> 'a) -> t -> init:'a -> 'a
+end
-val lookup_rel : int -> rel_context -> rel_declaration
-(** Size of the [rel_context] including LetIns *)
-val rel_context_length : rel_context -> int
-(** Size of the [rel_context] without LetIns *)
-val rel_context_nhyps : rel_context -> int
-(** Indicates whether a LetIn or a Lambda, starting from oldest declaration *)
-val rel_context_tags : rel_context -> bool list
+type section_context = Named.t
diff --git a/kernel/conv_oracle.ml b/kernel/conv_oracle.ml
index 3b01538b92..462413bd38 100644
--- a/kernel/conv_oracle.ml
+++ b/kernel/conv_oracle.ml
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -82,12 +82,17 @@ let fold_strategy f { var_opacity; cst_opacity; } accu =
let get_transp_state { var_trstate; cst_trstate } = (var_trstate, cst_trstate)
(* Unfold the first constant only if it is "more transparent" than the
- second one. In case of tie, expand the second one. *)
+ second one. In case of tie, use the recommended default. *)
let oracle_order f o l2r k1 k2 =
match get_strategy o f k1, get_strategy o f k2 with
- | Expand, _ -> true
- | Level n1, Opaque -> true
- | Level n1, Level n2 -> n1 < n2
- | _ -> l2r (* use recommended default *)
+ | Expand, Expand -> l2r
+ | Expand, (Opaque | Level _) -> true
+ | (Opaque | Level _), Expand -> false
+ | Opaque, Opaque -> l2r
+ | Level _, Opaque -> true
+ | Opaque, Level _ -> false
+ | Level n1, Level n2 ->
+ if Int.equal n1 n2 then l2r
+ else n1 < n2
let get_strategy o = get_strategy o (fun x -> x)
diff --git a/kernel/conv_oracle.mli b/kernel/conv_oracle.mli
index 629912220f..70f02b54d7 100644
--- a/kernel/conv_oracle.mli
+++ b/kernel/conv_oracle.mli
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
diff --git a/kernel/cooking.ml b/kernel/cooking.ml
index be71bd7b41..6dc2a617dd 100644
--- a/kernel/cooking.ml
+++ b/kernel/cooking.ml
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -44,15 +44,15 @@ module RefHash =
struct
type t = my_global_reference
let equal gr1 gr2 = match gr1, gr2 with
- | ConstRef c1, ConstRef c2 -> Constant.CanOrd.equal c1 c2
- | IndRef i1, IndRef i2 -> eq_ind i1 i2
- | ConstructRef c1, ConstructRef c2 -> eq_constructor c1 c2
+ | ConstRef c1, ConstRef c2 -> Constant.SyntacticOrd.equal c1 c2
+ | IndRef i1, IndRef i2 -> eq_syntactic_ind i1 i2
+ | ConstructRef c1, ConstructRef c2 -> eq_syntactic_constructor c1 c2
| _ -> false
open Hashset.Combine
let hash = function
- | ConstRef c -> combinesmall 1 (Constant.hash c)
- | IndRef i -> combinesmall 2 (ind_hash i)
- | ConstructRef c -> combinesmall 3 (constructor_hash c)
+ | ConstRef c -> combinesmall 1 (Constant.SyntacticOrd.hash c)
+ | IndRef i -> combinesmall 2 (ind_syntactic_hash i)
+ | ConstructRef c -> combinesmall 3 (constructor_syntactic_hash c)
end
module RefTable = Hashtbl.Make(RefHash)
@@ -173,7 +173,7 @@ let expmod_constr_subst cache modlist subst c =
let cook_constr { Opaqueproof.modlist ; abstract } c =
let cache = RefTable.create 13 in
let expmod = expmod_constr_subst cache modlist (pi2 abstract) in
- let hyps = Context.map_named_context expmod (pi1 abstract) in
+ let hyps = Context.Named.map expmod (pi1 abstract) in
abstract_constant_body (expmod c) hyps
let lift_univs cb subst =
@@ -195,14 +195,16 @@ let cook_constant env { from = cb; info } =
let abstract, usubst, abs_ctx = abstract in
let usubst, univs = lift_univs cb usubst in
let expmod = expmod_constr_subst cache modlist usubst in
- let hyps = Context.map_named_context expmod abstract in
+ let hyps = Context.Named.map expmod abstract in
let body = on_body modlist (hyps, usubst, abs_ctx)
(fun c -> abstract_constant_body (expmod c) hyps)
cb.const_body
in
let const_hyps =
- Context.fold_named_context (fun (h,_,_) hyps ->
- List.filter (fun (id,_,_) -> not (Id.equal id h)) hyps)
+ Context.Named.fold_outside (fun decl hyps ->
+ let open Context.Named.Declaration in
+ List.filter (fun decl' -> not (Id.equal (get_id decl) (get_id decl')))
+ hyps)
hyps ~init:cb.const_hyps in
let typ = match cb.const_type with
| RegularArity t ->
diff --git a/kernel/cooking.mli b/kernel/cooking.mli
index 441c9dd2d5..327e697d23 100644
--- a/kernel/cooking.mli
+++ b/kernel/cooking.mli
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
diff --git a/kernel/csymtable.ml b/kernel/csymtable.ml
index b3f0ba5b58..047da682ad 100644
--- a/kernel/csymtable.ml
+++ b/kernel/csymtable.ml
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -15,7 +15,6 @@
open Util
open Names
open Term
-open Context
open Vm
open Cemitcodes
open Cbytecodes
@@ -58,7 +57,7 @@ let set_global v =
let rec eq_structured_constant c1 c2 = match c1, c2 with
| Const_sorts s1, Const_sorts s2 -> Sorts.equal s1 s2
| Const_sorts _, _ -> false
-| Const_ind i1, Const_ind i2 -> Univ.eq_puniverses eq_ind i1 i2
+| Const_ind i1, Const_ind i2 -> eq_ind i1 i2
| Const_ind _, _ -> false
| Const_proj p1, Const_proj p2 -> Constant.equal p1 p2
| Const_proj _, _ -> false
@@ -67,18 +66,24 @@ let rec eq_structured_constant c1 c2 = match c1, c2 with
| Const_bn (t1, a1), Const_bn (t2, a2) ->
Int.equal t1 t2 && Array.equal eq_structured_constant a1 a2
| Const_bn _, _ -> false
+| Const_univ_level l1 , Const_univ_level l2 -> Univ.eq_levels l1 l2
+| Const_univ_level _ , _ -> false
+| Const_type u1 , Const_type u2 -> Univ.Universe.equal u1 u2
+| Const_type _ , _ -> false
let rec hash_structured_constant c =
let open Hashset.Combine in
match c with
| Const_sorts s -> combinesmall 1 (Sorts.hash s)
- | Const_ind (i,u) -> combinesmall 2 (combine (ind_hash i) (Univ.Instance.hash u))
+ | Const_ind i -> combinesmall 2 (ind_hash i)
| Const_proj p -> combinesmall 3 (Constant.hash p)
| Const_b0 t -> combinesmall 4 (Int.hash t)
| Const_bn (t, a) ->
let fold h c = combine h (hash_structured_constant c) in
let h = Array.fold_left fold 0 a in
combinesmall 5 (combine (Int.hash t) h)
+ | Const_univ_level l -> combinesmall 6 (Univ.Level.hash l)
+ | Const_type u -> combinesmall 7 (Univ.Universe.hash u)
module SConstTable = Hashtbl.Make (struct
type t = structured_constant
@@ -124,9 +129,9 @@ exception NotEvaluated
let key rk =
match !rk with
| None -> raise NotEvaluated
- | Some k -> (*Pp.msgnl (str"found at: "++int k);*)
- try Ephemeron.get k
- with Ephemeron.InvalidKey -> raise NotEvaluated
+ | Some k ->
+ try CEphemeron.get k
+ with CEphemeron.InvalidKey -> raise NotEvaluated
(************************)
(* traduction des patch *)
@@ -148,25 +153,24 @@ let slot_for_annot key =
AnnotTable.add annot_tbl key n;
n
-let rec slot_for_getglobal env (kn,u) =
+let rec slot_for_getglobal env kn =
let (cb,(_,rk)) = lookup_constant_key kn env in
try key rk
with NotEvaluated ->
(* Pp.msgnl(str"not yet evaluated");*)
let pos =
match cb.const_body_code with
- | None -> set_global (val_of_constant (kn,u))
+ | None -> set_global (val_of_constant kn)
| Some code ->
match Cemitcodes.force code with
| BCdefined(code,pl,fv) ->
- if Univ.Instance.is_empty u then
- let v = eval_to_patch env (code,pl,fv) in
- set_global v
- else set_global (val_of_constant (kn,u))
+ let v = eval_to_patch env (code,pl,fv) in
+ set_global v
| BCalias kn' -> slot_for_getglobal env kn'
- | BCconstant -> set_global (val_of_constant (kn,u)) in
+ | BCconstant -> set_global (val_of_constant kn)
+ in
(*Pp.msgnl(str"value stored at: "++int pos);*)
- rk := Some (Ephemeron.create pos);
+ rk := Some (CEphemeron.create pos);
pos
and slot_for_fv env fv =
@@ -185,36 +189,33 @@ and slot_for_fv env fv =
let nv = Pre_env.lookup_named_val id env in
begin match force_lazy_val nv with
| None ->
- let _, b, _ = Context.lookup_named id env.env_named_context in
- fill_fv_cache nv id val_of_named idfun b
+ let open Context.Named in
+ let open Declaration in
+ env.env_named_context |> lookup id |> get_value |> fill_fv_cache nv id val_of_named idfun
| Some (v, _) -> v
end
| FVrel i ->
let rv = Pre_env.lookup_rel_val i env in
begin match force_lazy_val rv with
| None ->
- let _, b, _ = lookup_rel i env.env_rel_context in
- fill_fv_cache rv i val_of_rel env_of_rel b
+ let open Context.Rel in
+ let open Declaration in
+ env.env_rel_context |> lookup i |> get_value |> fill_fv_cache rv i val_of_rel env_of_rel
| Some (v, _) -> v
end
+ | FVuniv_var idu ->
+ assert false
and eval_to_patch env (buff,pl,fv) =
- (* copy code *before* patching because of nested evaluations:
- the code we are patching might be called (and thus "concurrently" patched)
- and results in wrong results. Side-effects... *)
- let buff = Cemitcodes.copy buff in
let patch = function
- | Reloc_annot a, pos -> patch_int buff pos (slot_for_annot a)
- | Reloc_const sc, pos -> patch_int buff pos (slot_for_str_cst sc)
- | Reloc_getglobal kn, pos ->
-(* Pp.msgnl (str"patching global: "++str(debug_string_of_con kn));*)
- patch_int buff pos (slot_for_getglobal env kn);
-(* Pp.msgnl (str"patch done: "++str(debug_string_of_con kn))*)
+ | Reloc_annot a, pos -> (pos, slot_for_annot a)
+ | Reloc_const sc, pos -> (pos, slot_for_str_cst sc)
+ | Reloc_getglobal kn, pos -> (pos, slot_for_getglobal env kn)
in
- List.iter patch pl;
+ let patches = List.map_left patch pl in
+ let buff = patch_int buff patches in
let vm_env = Array.map (slot_for_fv env) fv in
let tc = tcode_of_code buff (length buff) in
-(*Pp.msgnl (str"execute code");*)
eval_tcode tc vm_env
and val_of_constr env c =
@@ -232,5 +233,3 @@ and val_of_constr env c =
let set_transparent_const kn = () (* !?! *)
let set_opaque_const kn = () (* !?! *)
-
-
diff --git a/kernel/csymtable.mli b/kernel/csymtable.mli
index ca5f8ac297..cd561148bf 100644
--- a/kernel/csymtable.mli
+++ b/kernel/csymtable.mli
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
diff --git a/kernel/declarations.mli b/kernel/declarations.mli
index 561c66b422..1b77d5b7cb 100644
--- a/kernel/declarations.mli
+++ b/kernel/declarations.mli
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -8,7 +8,6 @@
open Names
open Term
-open Context
(** This module defines the internal representation of global
declarations. This includes global constants/axioms, mutual
@@ -38,7 +37,7 @@ type ('a, 'b) declaration_arity =
| RegularArity of 'a
| TemplateArity of 'b
-type constant_type = (types, rel_context * template_arity) declaration_arity
+type constant_type = (types, Context.Rel.t * template_arity) declaration_arity
(** Inlining level of parameters at functor applications.
None means no inlining *)
@@ -79,12 +78,6 @@ type constant_body = {
const_proj : projection_body option;
const_inline_code : bool }
-type seff_env = [ `Nothing | `Opaque of Constr.t * Univ.universe_context_set ]
-
-type side_effect =
- | SEsubproof of constant * constant_body * seff_env
- | SEscheme of (inductive * constant * constant_body * seff_env) list * string
-
(** {6 Representation of mutual inductive types in the kernel } *)
type recarg =
@@ -123,7 +116,7 @@ type one_inductive_body = {
mind_typename : Id.t; (** Name of the type: [Ii] *)
- mind_arity_ctxt : rel_context; (** Arity context of [Ii] with parameters: [forall params, Ui] *)
+ mind_arity_ctxt : Context.Rel.t; (** Arity context of [Ii] with parameters: [forall params, Ui] *)
mind_arity : inductive_arity; (** Arity sort and original user arity *)
@@ -177,7 +170,7 @@ type mutual_inductive_body = {
mind_nparams_rec : int; (** Number of recursively uniform (i.e. ordinary) parameters *)
- mind_params_ctxt : rel_context; (** The context of parameters (includes let-in declaration) *)
+ mind_params_ctxt : Context.Rel.t; (** The context of parameters (includes let-in declaration) *)
mind_polymorphic : bool; (** Is it polymorphic or not *)
@@ -244,17 +237,26 @@ and module_body =
{ mod_mp : module_path; (** absolute path of the module *)
mod_expr : module_implementation; (** implementation *)
mod_type : module_signature; (** expanded type *)
- (** algebraic type, kept if it's relevant for extraction *)
- mod_type_alg : module_expression option;
- (** set of all constraints in the module *)
- mod_constraints : Univ.constraints;
- (** quotiented set of equivalent constants and inductive names *)
- mod_delta : Mod_subst.delta_resolver;
+ mod_type_alg : module_expression option; (** algebraic type *)
+ mod_constraints : Univ.ContextSet.t; (**
+ set of all universes constraints in the module *)
+ mod_delta : Mod_subst.delta_resolver; (**
+ quotiented set of equivalent constants and inductive names *)
mod_retroknowledge : Retroknowledge.action list }
+(** For a module, there are five possible situations:
+ - [Declare Module M : T] then [mod_expr = Abstract; mod_type_alg = Some T]
+ - [Module M := E] then [mod_expr = Algebraic E; mod_type_alg = None]
+ - [Module M : T := E] then [mod_expr = Algebraic E; mod_type_alg = Some T]
+ - [Module M. ... End M] then [mod_expr = FullStruct; mod_type_alg = None]
+ - [Module M : T. ... End M] then [mod_expr = Struct; mod_type_alg = Some T]
+ And of course, all these situations may be functors or not. *)
+
(** A [module_type_body] is just a [module_body] with no
implementation ([mod_expr] always [Abstract]) and also
- an empty [mod_retroknowledge] *)
+ an empty [mod_retroknowledge]. Its [mod_type_alg] contains
+ the algebraic definition of this module type, or [None]
+ if it has been built interactively. *)
and module_type_body = module_body
diff --git a/kernel/declareops.ml b/kernel/declareops.ml
index a7051d5c13..a09a8b7862 100644
--- a/kernel/declareops.ml
+++ b/kernel/declareops.ml
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -9,6 +9,7 @@
open Declarations
open Mod_subst
open Util
+open Context.Rel.Declaration
(** Operations concernings types in [Declarations] :
[constant_body], [mutual_inductive_body], [module_body] ... *)
@@ -87,10 +88,8 @@ let is_opaque cb = match cb.const_body with
(** {7 Constant substitutions } *)
-let subst_rel_declaration sub (id,copt,t as x) =
- let copt' = Option.smartmap (subst_mps sub) copt in
- let t' = subst_mps sub t in
- if copt == copt' && t == t' then x else (id,copt',t')
+let subst_rel_declaration sub =
+ map_constr (subst_mps sub)
let subst_rel_context sub = List.smartmap (subst_rel_declaration sub)
@@ -140,11 +139,8 @@ let subst_const_body sub cb =
share internal fields (e.g. constr), and not the records
themselves. But would it really bring substantial gains ? *)
-let hcons_rel_decl ((n,oc,t) as d) =
- let n' = Names.Name.hcons n
- and oc' = Option.smartmap Term.hcons_constr oc
- and t' = Term.hcons_types t
- in if n' == n && oc' == oc && t' == t then d else (n',oc',t')
+let hcons_rel_decl =
+ map_type Term.hcons_types % map_value Term.hcons_constr % map_name Names.Name.hcons
let hcons_rel_context l = List.smartmap hcons_rel_decl l
@@ -254,7 +250,7 @@ let subst_mind_body sub mib =
mind_nparams = mib.mind_nparams;
mind_nparams_rec = mib.mind_nparams_rec;
mind_params_ctxt =
- Context.map_rel_context (subst_mps sub) mib.mind_params_ctxt;
+ Context.Rel.map (subst_mps sub) mib.mind_params_ctxt;
mind_packets = Array.smartmap (subst_mind_packet sub) mib.mind_packets ;
mind_polymorphic = mib.mind_polymorphic;
mind_universes = mib.mind_universes;
@@ -304,17 +300,90 @@ let hcons_mind mib =
(** {6 Stm machinery } *)
-let string_of_side_effect = function
- | SEsubproof (c,_,_) -> Names.string_of_con c
- | SEscheme (cl,_) ->
- String.concat ", " (List.map (fun (_,c,_,_) -> Names.string_of_con c) cl)
-type side_effects = side_effect list
-let no_seff = ([] : side_effects)
-let iter_side_effects f l = List.iter f (List.rev l)
-let fold_side_effects f a l = List.fold_left f a l
-let uniquize_side_effects l = List.rev (CList.uniquize (List.rev l))
-let union_side_effects l1 l2 = l1 @ l2
-let flatten_side_effects l = List.flatten l
-let side_effects_of_list l = l
-let cons_side_effects x l = x :: l
-let side_effects_is_empty = List.is_empty
+let string_of_side_effect { Entries.eff } = match eff with
+ | Entries.SEsubproof (c,_,_) -> "P(" ^ Names.string_of_con c ^ ")"
+ | Entries.SEscheme (cl,_) ->
+ "S(" ^ String.concat ", " (List.map (fun (_,c,_,_) -> Names.string_of_con c) cl) ^ ")"
+
+(** Hashconsing of modules *)
+
+let hcons_functorize hty he hself f = match f with
+| NoFunctor e ->
+ let e' = he e in
+ if e == e' then f else NoFunctor e'
+| MoreFunctor (mid, ty, nf) ->
+ (** FIXME *)
+ let mid' = mid in
+ let ty' = hty ty in
+ let nf' = hself nf in
+ if mid == mid' && ty == ty' && nf == nf' then f
+ else MoreFunctor (mid, ty', nf')
+
+let hcons_module_alg_expr me = me
+
+let rec hcons_structure_field_body sb = match sb with
+| SFBconst cb ->
+ let cb' = hcons_const_body cb in
+ if cb == cb' then sb else SFBconst cb'
+| SFBmind mib ->
+ let mib' = hcons_mind mib in
+ if mib == mib' then sb else SFBmind mib'
+| SFBmodule mb ->
+ let mb' = hcons_module_body mb in
+ if mb == mb' then sb else SFBmodule mb'
+| SFBmodtype mb ->
+ let mb' = hcons_module_body mb in
+ if mb == mb' then sb else SFBmodtype mb'
+
+and hcons_structure_body sb =
+ (** FIXME *)
+ let map (l, sfb as fb) =
+ let l' = Names.Label.hcons l in
+ let sfb' = hcons_structure_field_body sfb in
+ if l == l' && sfb == sfb' then fb else (l', sfb')
+ in
+ List.smartmap map sb
+
+and hcons_module_signature ms =
+ hcons_functorize hcons_module_body hcons_structure_body hcons_module_signature ms
+
+and hcons_module_expression me =
+ hcons_functorize hcons_module_body hcons_module_alg_expr hcons_module_expression me
+
+and hcons_module_implementation mip = match mip with
+| Abstract -> Abstract
+| Algebraic me ->
+ let me' = hcons_module_expression me in
+ if me == me' then mip else Algebraic me'
+| Struct ms ->
+ let ms' = hcons_module_signature ms in
+ if ms == ms' then mip else Struct ms
+| FullStruct -> FullStruct
+
+and hcons_module_body mb =
+ let mp' = mb.mod_mp in
+ let expr' = hcons_module_implementation mb.mod_expr in
+ let type' = hcons_module_signature mb.mod_type in
+ let type_alg' = mb.mod_type_alg in
+ let constraints' = Univ.hcons_universe_context_set mb.mod_constraints in
+ let delta' = mb.mod_delta in
+ let retroknowledge' = mb.mod_retroknowledge in
+
+ if
+ mb.mod_mp == mp' &&
+ mb.mod_expr == expr' &&
+ mb.mod_type == type' &&
+ mb.mod_type_alg == type_alg' &&
+ mb.mod_constraints == constraints' &&
+ mb.mod_delta == delta' &&
+ mb.mod_retroknowledge == retroknowledge'
+ then mb
+ else {
+ mod_mp = mp';
+ mod_expr = expr';
+ mod_type = type';
+ mod_type_alg = type_alg';
+ mod_constraints = constraints';
+ mod_delta = delta';
+ mod_retroknowledge = retroknowledge';
+ }
diff --git a/kernel/declareops.mli b/kernel/declareops.mli
index ce65af975e..ad2b5d0a6c 100644
--- a/kernel/declareops.mli
+++ b/kernel/declareops.mli
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -9,6 +9,7 @@
open Declarations
open Mod_subst
open Univ
+open Entries
(** Operations concerning types in [Declarations] :
[constant_body], [mutual_inductive_body], [module_body] ... *)
@@ -49,17 +50,6 @@ val is_opaque : constant_body -> bool
val string_of_side_effect : side_effect -> string
-type side_effects
-val no_seff : side_effects
-val iter_side_effects : (side_effect -> unit) -> side_effects -> unit
-val fold_side_effects : ('a -> side_effect -> 'a) -> 'a -> side_effects -> 'a
-val uniquize_side_effects : side_effects -> side_effects
-val union_side_effects : side_effects -> side_effects -> side_effects
-val flatten_side_effects : side_effects list -> side_effects
-val side_effects_of_list : side_effect list -> side_effects
-val cons_side_effects : side_effect -> side_effects -> side_effects
-val side_effects_is_empty : side_effects -> bool
-
(** {6 Inductive types} *)
val eq_recarg : recarg -> recarg -> bool
@@ -87,3 +77,4 @@ val inductive_context : mutual_inductive_body -> universe_context
val hcons_const_body : constant_body -> constant_body
val hcons_mind : mutual_inductive_body -> mutual_inductive_body
+val hcons_module_body : module_body -> module_body
diff --git a/kernel/entries.mli b/kernel/entries.mli
index 303d27d355..d07ca2103f 100644
--- a/kernel/entries.mli
+++ b/kernel/entries.mli
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -18,8 +18,8 @@ open Term
(** {6 Local entries } *)
type local_entry =
- | LocalDef of constr
- | LocalAssum of constr
+ | LocalDefEntry of constr
+ | LocalAssumEntry of constr
(** {6 Declaration of inductive types. } *)
@@ -54,11 +54,11 @@ type mutual_inductive_entry = {
mind_entry_private : bool option }
(** {6 Constants (Definition/Axiom) } *)
-type proof_output = constr Univ.in_universe_context_set * Declareops.side_effects
-type const_entry_body = proof_output Future.computation
+type 'a proof_output = constr Univ.in_universe_context_set * 'a
+type 'a const_entry_body = 'a proof_output Future.computation
-type definition_entry = {
- const_entry_body : const_entry_body;
+type 'a definition_entry = {
+ const_entry_body : 'a const_entry_body;
(* List of section variables *)
const_entry_secctx : Context.section_context option;
(* State id on which the completion of type checking is reported *)
@@ -78,8 +78,8 @@ type projection_entry = {
proj_entry_ind : mutual_inductive;
proj_entry_arg : int }
-type constant_entry =
- | DefinitionEntry of definition_entry
+type 'a constant_entry =
+ | DefinitionEntry of 'a definition_entry
| ParameterEntry of parameter_entry
| ProjectionEntry of projection_entry
@@ -96,3 +96,16 @@ type module_entry =
| MType of module_params_entry * module_struct_entry
| MExpr of
module_params_entry * module_struct_entry * module_struct_entry option
+
+type seff_env = [ `Nothing | `Opaque of Constr.t * Univ.universe_context_set ]
+
+type side_eff =
+ | SEsubproof of constant * Declarations.constant_body * seff_env
+ | SEscheme of (inductive * constant * Declarations.constant_body * seff_env) list * string
+
+type side_effect = {
+ from_env : Declarations.structure_body CEphemeron.key;
+ eff : side_eff;
+}
+
+type side_effects = side_effect list
diff --git a/kernel/environ.ml b/kernel/environ.ml
index 109e3830c2..d8493d9baf 100644
--- a/kernel/environ.ml
+++ b/kernel/environ.ml
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -24,10 +24,10 @@ open Errors
open Util
open Names
open Term
-open Context
open Vars
open Declarations
open Pre_env
+open Context.Rel.Declaration
(* The type of environments. *)
@@ -70,21 +70,19 @@ let empty_context env =
(* Rel context *)
let lookup_rel n env =
- lookup_rel n env.env_rel_context
+ Context.Rel.lookup n env.env_rel_context
let evaluable_rel n env =
- match lookup_rel n env with
- | (_,Some _,_) -> true
- | _ -> false
+ is_local_def (lookup_rel n env)
let nb_rel env = env.env_nb_rel
let push_rel = push_rel
-let push_rel_context ctxt x = Context.fold_rel_context push_rel ctxt ~init:x
+let push_rel_context ctxt x = Context.Rel.fold_outside push_rel ctxt ~init:x
let push_rec_types (lna,typarray,_) env =
- let ctxt = Array.map2_i (fun i na t -> (na, None, lift i t)) lna typarray in
+ let ctxt = Array.map2_i (fun i na t -> LocalAssum (na, lift i t)) lna typarray in
Array.fold_left (fun e assum -> push_rel assum e) env ctxt
let fold_rel_context f env ~init =
@@ -108,19 +106,10 @@ let named_vals_of_val = snd
(* [map_named_val f ctxt] apply [f] to the body and the type of
each declarations.
*** /!\ *** [f t] should be convertible with t *)
-let map_named_val f (ctxt,ctxtv) =
- let rec map ctx = match ctx with
- | [] -> []
- | (id, body, typ) :: rem ->
- let body' = Option.smartmap f body in
- let typ' = f typ in
- let rem' = map rem in
- if body' == body && typ' == typ && rem' == rem then ctx
- else (id, body', typ') :: rem'
- in
- (map ctxt, ctxtv)
+let map_named_val f =
+ on_fst (Context.Named.map f)
-let empty_named_context = empty_named_context
+let empty_named_context = Context.Named.empty
let push_named = push_named
let push_named_context = List.fold_right push_named
@@ -130,19 +119,21 @@ let val_of_named_context ctxt =
List.fold_right push_named_context_val ctxt empty_named_context_val
-let lookup_named id env = Context.lookup_named id env.env_named_context
-let lookup_named_val id (ctxt,_) = Context.lookup_named id ctxt
+let lookup_named id env = Context.Named.lookup id env.env_named_context
+let lookup_named_val id (ctxt,_) = Context.Named.lookup id ctxt
let eq_named_context_val c1 c2 =
- c1 == c2 || named_context_equal (named_context_of_val c1) (named_context_of_val c2)
+ c1 == c2 || Context.Named.equal (named_context_of_val c1) (named_context_of_val c2)
(* A local const is evaluable if it is defined *)
+open Context.Named.Declaration
+
let named_type id env =
- let (_,_,t) = lookup_named id env in t
+ get_type (lookup_named id env)
let named_body id env =
- let (_,b,_) = lookup_named id env in b
+ get_value (lookup_named id env)
let evaluable_named id env =
match named_body id env with
@@ -153,7 +144,7 @@ let reset_with_named_context (ctxt,ctxtv) env =
{ env with
env_named_context = ctxt;
env_named_vals = ctxtv;
- env_rel_context = empty_rel_context;
+ env_rel_context = Context.Rel.empty;
env_rel_val = [];
env_nb_rel = 0 }
@@ -176,31 +167,49 @@ let fold_named_context f env ~init =
in fold_right env
let fold_named_context_reverse f ~init env =
- Context.fold_named_context_reverse f ~init:init (named_context env)
+ Context.Named.fold_inside f ~init:init (named_context env)
(* Universe constraints *)
-let add_constraints c env =
- if Univ.Constraint.is_empty c then
- env
- else
- let s = env.env_stratification in
+let map_universes f env =
+ let s = env.env_stratification in
{ env with env_stratification =
- { s with env_universes = Univ.merge_constraints c s.env_universes } }
+ { s with env_universes = f s.env_universes } }
+
+let add_constraints c env =
+ if Univ.Constraint.is_empty c then env
+ else map_universes (UGraph.merge_constraints c) env
let check_constraints c env =
- Univ.check_constraints c env.env_stratification.env_universes
-
-let set_engagement c env = (* Unsafe *)
- { env with env_stratification =
- { env.env_stratification with env_engagement = c } }
+ UGraph.check_constraints c env.env_stratification.env_universes
let push_constraints_to_env (_,univs) env =
add_constraints univs env
-let push_context ctx env = add_constraints (Univ.UContext.constraints ctx) env
-let push_context_set ctx env = add_constraints (Univ.ContextSet.constraints ctx) env
+let add_universes strict ctx g =
+ let g = Array.fold_left
+ (* Be lenient, module typing reintroduces universes and constraints due to includes *)
+ (fun g v -> try UGraph.add_universe v strict g with UGraph.AlreadyDeclared -> g)
+ g (Univ.Instance.to_array (Univ.UContext.instance ctx))
+ in
+ UGraph.merge_constraints (Univ.UContext.constraints ctx) g
+
+let push_context ?(strict=false) ctx env =
+ map_universes (add_universes strict ctx) env
+
+let add_universes_set strict ctx g =
+ let g = Univ.LSet.fold
+ (fun v g -> try UGraph.add_universe v strict g with UGraph.AlreadyDeclared -> g)
+ (Univ.ContextSet.levels ctx) g
+ in UGraph.merge_constraints (Univ.ContextSet.constraints ctx) g
+
+let push_context_set ?(strict=false) ctx env =
+ map_universes (add_universes_set strict ctx) env
+
+let set_engagement c env = (* Unsafe *)
+ { env with env_stratification =
+ { env.env_stratification with env_engagement = c } }
(* Global constants *)
@@ -371,11 +380,11 @@ let add_mind kn mib env =
let lookup_constant_variables c env =
let cmap = lookup_constant c env in
- Context.vars_of_named_context cmap.const_hyps
+ Context.Named.to_vars cmap.const_hyps
let lookup_inductive_variables (kn,i) env =
let mis = lookup_mind kn env in
- Context.vars_of_named_context mis.mind_hyps
+ Context.Named.to_vars mis.mind_hyps
let lookup_constructor_variables (ind,_) env =
lookup_inductive_variables ind env
@@ -409,15 +418,15 @@ let global_vars_set env constr =
contained in the types of the needed variables. *)
let really_needed env needed =
- Context.fold_named_context_reverse
- (fun need (id,copt,t) ->
- if Id.Set.mem id need then
+ Context.Named.fold_inside
+ (fun need decl ->
+ if Id.Set.mem (get_id decl) need then
let globc =
- match copt with
- | None -> Id.Set.empty
- | Some c -> global_vars_set env c in
+ match decl with
+ | LocalAssum _ -> Id.Set.empty
+ | LocalDef (_,c,_) -> global_vars_set env c in
Id.Set.union
- (global_vars_set env t)
+ (global_vars_set env (get_type decl))
(Id.Set.union globc need)
else need)
~init:needed
@@ -425,9 +434,9 @@ let really_needed env needed =
let keep_hyps env needed =
let really_needed = really_needed env needed in
- Context.fold_named_context
- (fun (id,_,_ as d) nsign ->
- if Id.Set.mem id really_needed then add_named_decl d nsign
+ Context.Named.fold_outside
+ (fun d nsign ->
+ if Id.Set.mem (get_id d) really_needed then Context.Named.add d nsign
else nsign)
(named_context env)
~init:empty_named_context
@@ -479,9 +488,9 @@ exception Hyp_not_found
let apply_to_hyp (ctxt,vals) id f =
let rec aux rtail ctxt vals =
match ctxt, vals with
- | (idc,c,ct as d)::ctxt, v::vals ->
- if Id.equal idc id then
- (f ctxt d rtail)::ctxt, v::vals
+ | d::ctxt, v::vals ->
+ if Id.equal (get_id d) id then
+ (f ctxt d rtail)::ctxt, v::vals
else
let ctxt',vals' = aux (d::rtail) ctxt vals in
d::ctxt', v::vals'
@@ -492,8 +501,8 @@ let apply_to_hyp (ctxt,vals) id f =
let apply_to_hyp_and_dependent_on (ctxt,vals) id f g =
let rec aux ctxt vals =
match ctxt,vals with
- | (idc,c,ct as d)::ctxt, v::vals ->
- if Id.equal idc id then
+ | d::ctxt, v::vals ->
+ if Id.equal (get_id d) id then
let sign = ctxt,vals in
push_named_context_val (f d sign) sign
else
@@ -506,8 +515,8 @@ let apply_to_hyp_and_dependent_on (ctxt,vals) id f g =
let insert_after_hyp (ctxt,vals) id d check =
let rec aux ctxt vals =
match ctxt, vals with
- | (idc,c,ct)::ctxt', v::vals' ->
- if Id.equal idc id then begin
+ | decl::ctxt', v::vals' ->
+ if Id.equal (get_id decl) id then begin
check ctxt;
push_named_context_val d (ctxt,vals)
end else
@@ -523,9 +532,8 @@ let remove_hyps ids check_context check_value (ctxt, vals) =
let rec remove_hyps ctxt vals = match ctxt, vals with
| [], [] -> [], []
| d :: rctxt, (nid, v) :: rvals ->
- let (id, _, _) = d in
let ans = remove_hyps rctxt rvals in
- if Id.Set.mem id ids then ans
+ if Id.Set.mem (get_id d) ids then ans
else
let (rctxt', rvals') = ans in
let d' = check_context d in
@@ -584,7 +592,10 @@ let dispatch =
Array.init 31 (fun n -> mkConstruct
(digit_ind, nth_digit_plus_one i (30-n)))
in
- mkApp(mkConstruct(ind, 1), array_of_int tag)
+ (* We check that no bit above 31 is set to one. This assertion used to
+ fail in the VM, and led to conversion tests failing at Qed. *)
+ assert (Int.equal (tag lsr 31) 0);
+ mkApp(mkConstruct(ind, 1), array_of_int tag)
in
(* subfunction which dispatches the compiling information of an
diff --git a/kernel/environ.mli b/kernel/environ.mli
index 4ad0327fc7..5203899546 100644
--- a/kernel/environ.mli
+++ b/kernel/environ.mli
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -8,7 +8,6 @@
open Names
open Term
-open Context
open Declarations
open Univ
@@ -41,9 +40,9 @@ val eq_named_context_val : named_context_val -> named_context_val -> bool
val empty_env : env
-val universes : env -> Univ.universes
-val rel_context : env -> rel_context
-val named_context : env -> named_context
+val universes : env -> UGraph.t
+val rel_context : env -> Context.Rel.t
+val named_context : env -> Context.Named.t
val named_context_val : env -> named_context_val
val opaque_tables : env -> Opaqueproof.opaquetab
@@ -60,25 +59,25 @@ val empty_context : env -> bool
(** {5 Context of de Bruijn variables ([rel_context]) } *)
val nb_rel : env -> int
-val push_rel : rel_declaration -> env -> env
-val push_rel_context : rel_context -> env -> env
+val push_rel : Context.Rel.Declaration.t -> env -> env
+val push_rel_context : Context.Rel.t -> env -> env
val push_rec_types : rec_declaration -> env -> env
(** Looks up in the context of local vars referred by indice ([rel_context])
raises [Not_found] if the index points out of the context *)
-val lookup_rel : int -> env -> rel_declaration
+val lookup_rel : int -> env -> Context.Rel.Declaration.t
val evaluable_rel : int -> env -> bool
(** {6 Recurrence on [rel_context] } *)
val fold_rel_context :
- (env -> rel_declaration -> 'a -> 'a) -> env -> init:'a -> 'a
+ (env -> Context.Rel.Declaration.t -> 'a -> 'a) -> env -> init:'a -> 'a
(** {5 Context of variables (section variables and goal assumptions) } *)
-val named_context_of_val : named_context_val -> named_context
+val named_context_of_val : named_context_val -> Context.Named.t
val named_vals_of_val : named_context_val -> Pre_env.named_vals
-val val_of_named_context : named_context -> named_context_val
+val val_of_named_context : Context.Named.t -> named_context_val
val empty_named_context_val : named_context_val
@@ -88,18 +87,18 @@ val empty_named_context_val : named_context_val
val map_named_val :
(constr -> constr) -> named_context_val -> named_context_val
-val push_named : named_declaration -> env -> env
-val push_named_context : named_context -> env -> env
+val push_named : Context.Named.Declaration.t -> env -> env
+val push_named_context : Context.Named.t -> env -> env
val push_named_context_val :
- named_declaration -> named_context_val -> named_context_val
+ Context.Named.Declaration.t -> named_context_val -> named_context_val
(** Looks up in the context of local vars referred by names ([named_context])
raises [Not_found] if the Id.t is not found *)
-val lookup_named : variable -> env -> named_declaration
-val lookup_named_val : variable -> named_context_val -> named_declaration
+val lookup_named : variable -> env -> Context.Named.Declaration.t
+val lookup_named_val : variable -> named_context_val -> Context.Named.Declaration.t
val evaluable_named : variable -> env -> bool
val named_type : variable -> env -> types
val named_body : variable -> env -> constr option
@@ -107,11 +106,11 @@ val named_body : variable -> env -> constr option
(** {6 Recurrence on [named_context]: older declarations processed first } *)
val fold_named_context :
- (env -> named_declaration -> 'a -> 'a) -> env -> init:'a -> 'a
+ (env -> Context.Named.Declaration.t -> 'a -> 'a) -> env -> init:'a -> 'a
(** Recurrence on [named_context] starting from younger decl *)
val fold_named_context_reverse :
- ('a -> named_declaration -> 'a) -> init:'a -> env -> 'a
+ ('a -> Context.Named.Declaration.t -> 'a) -> init:'a -> env -> 'a
(** This forgets named and rel contexts *)
val reset_context : env -> env
@@ -208,8 +207,8 @@ val add_constraints : Univ.constraints -> env -> env
(** Check constraints are satifiable in the environment. *)
val check_constraints : Univ.constraints -> env -> bool
-val push_context : Univ.universe_context -> env -> env
-val push_context_set : Univ.universe_context_set -> env -> env
+val push_context : ?strict:bool -> Univ.universe_context -> env -> env
+val push_context_set : ?strict:bool -> Univ.universe_context_set -> env -> env
val push_constraints_to_env : 'a Univ.constrained -> env -> env
val set_engagement : engagement -> env -> env
@@ -228,7 +227,7 @@ val vars_of_global : env -> constr -> Id.Set.t
val really_needed : env -> Id.Set.t -> Id.Set.t
(** like [really_needed] but computes a well ordered named context *)
-val keep_hyps : env -> Id.Set.t -> section_context
+val keep_hyps : env -> Id.Set.t -> Context.section_context
(** {5 Unsafe judgments. }
We introduce here the pre-type of judgments, which is
@@ -250,7 +249,7 @@ type unsafe_type_judgment = {
(** {6 Compilation of global declaration } *)
-val compile_constant_body : env -> constant_def -> Cemitcodes.body_code option
+val compile_constant_body : env -> constant_universes option -> constant_def -> Cemitcodes.body_code option
exception Hyp_not_found
@@ -258,22 +257,22 @@ exception Hyp_not_found
return [tail::(f head (id,_,_) (rev tail))::head].
the value associated to id should not change *)
val apply_to_hyp : named_context_val -> variable ->
- (named_context -> named_declaration -> named_context -> named_declaration) ->
+ (Context.Named.t -> Context.Named.Declaration.t -> Context.Named.t -> Context.Named.Declaration.t) ->
named_context_val
(** [apply_to_hyp_and_dependent_on sign id f g] split [sign] into
[tail::(id,_,_)::head] and
return [(g tail)::(f (id,_,_))::head]. *)
val apply_to_hyp_and_dependent_on : named_context_val -> variable ->
- (named_declaration -> named_context_val -> named_declaration) ->
- (named_declaration -> named_context_val -> named_declaration) ->
+ (Context.Named.Declaration.t -> named_context_val -> Context.Named.Declaration.t) ->
+ (Context.Named.Declaration.t -> named_context_val -> Context.Named.Declaration.t) ->
named_context_val
val insert_after_hyp : named_context_val -> variable ->
- named_declaration ->
- (named_context -> unit) -> named_context_val
+ Context.Named.Declaration.t ->
+ (Context.Named.t -> unit) -> named_context_val
-val remove_hyps : Id.Set.t -> (named_declaration -> named_declaration) -> (Pre_env.lazy_val -> Pre_env.lazy_val) -> named_context_val -> named_context_val
+val remove_hyps : Id.Set.t -> (Context.Named.Declaration.t -> Context.Named.Declaration.t) -> (Pre_env.lazy_val -> Pre_env.lazy_val) -> named_context_val -> named_context_val
diff --git a/kernel/esubst.ml b/kernel/esubst.ml
index 42ca48ef3b..1dc389c649 100644
--- a/kernel/esubst.ml
+++ b/kernel/esubst.ml
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
diff --git a/kernel/esubst.mli b/kernel/esubst.mli
index 2b34da4da0..533d1c68c3 100644
--- a/kernel/esubst.mli
+++ b/kernel/esubst.mli
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
diff --git a/kernel/evar.ml b/kernel/evar.ml
index 54f15df4da..b972fc1147 100644
--- a/kernel/evar.ml
+++ b/kernel/evar.ml
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
diff --git a/kernel/evar.mli b/kernel/evar.mli
index 2c94db3f05..f28a13640f 100644
--- a/kernel/evar.mli
+++ b/kernel/evar.mli
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
diff --git a/kernel/fast_typeops.ml b/kernel/fast_typeops.ml
index d22abff10c..7f4ba8ecbe 100644
--- a/kernel/fast_typeops.ml
+++ b/kernel/fast_typeops.ml
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -33,7 +33,7 @@ let check_constraints cst env =
if Environ.check_constraints cst env then ()
else error_unsatisfied_constraints env cst
-(* This should be a type (a priori without intension to be an assumption) *)
+(* This should be a type (a priori without intention to be an assumption) *)
let type_judgment env c t =
match kind_of_term(whd_betadeltaiota env t) with
| Sort s -> {utj_val = c; utj_type = s }
@@ -52,8 +52,8 @@ let assumption_of_judgment env t ty =
error_assumption env (make_judge t ty)
(************************************************)
-(* Incremental typing rules: builds a typing judgement given the *)
-(* judgements for the subterms. *)
+(* Incremental typing rules: builds a typing judgment given the *)
+(* judgments for the subterms. *)
(*s Type of sorts *)
@@ -73,8 +73,8 @@ let judge_of_type u =
let judge_of_relative env n =
try
- let (_,_,typ) = lookup_rel n env in
- lift n typ
+ let open Context.Rel.Declaration in
+ env |> lookup_rel n |> get_type |> lift n
with Not_found ->
error_unbound_rel env n
@@ -90,8 +90,11 @@ let judge_of_variable env id =
variables of the current env *)
(* TODO: check order? *)
let check_hyps_inclusion env f c sign =
- Context.fold_named_context
- (fun (id,_,ty1) () ->
+ Context.Named.fold_outside
+ (fun decl () ->
+ let open Context.Named.Declaration in
+ let id = get_id decl in
+ let ty1 = get_type decl in
try
let ty2 = named_type id env in
if not (eq_constr ty2 ty1) then raise Exit
@@ -227,7 +230,7 @@ let judge_of_cast env c ct k expected_type =
default_conv ~l2r:true CUMUL env ct expected_type
| NATIVEcast ->
let sigma = Nativelambda.empty_evars in
- native_conv CUMUL sigma env ct expected_type
+ Nativeconv.native_conv CUMUL sigma env ct expected_type
with NotConvertible ->
error_actual_type env (make_judge c ct) expected_type
@@ -325,6 +328,7 @@ let type_fixpoint env lna lar vdef vdeft =
Ind et Constructsi un jour cela devient des constructions
arbitraires et non plus des variables *)
let rec execute env cstr =
+ let open Context.Rel.Declaration in
match kind_of_term cstr with
(* Atomic terms *)
| Sort (Prop c) ->
@@ -368,13 +372,13 @@ let rec execute env cstr =
| Lambda (name,c1,c2) ->
let _ = execute_is_type env c1 in
- let env1 = push_rel (name,None,c1) env in
+ let env1 = push_rel (LocalAssum (name,c1)) env in
let c2t = execute env1 c2 in
judge_of_abstraction env name c1 c2t
| Prod (name,c1,c2) ->
let vars = execute_is_type env c1 in
- let env1 = push_rel (name,None,c1) env in
+ let env1 = push_rel (LocalAssum (name,c1)) env in
let vars' = execute_is_type env1 c2 in
judge_of_product env name vars vars'
@@ -382,7 +386,7 @@ let rec execute env cstr =
let c1t = execute env c1 in
let _c2s = execute_is_type env c2 in
let _ = judge_of_cast env c1 c1t DEFAULTcast c2 in
- let env1 = push_rel (name,Some c1,c2) env in
+ let env1 = push_rel (LocalDef (name,c1,c2)) env in
let c3t = execute env1 c3 in
subst1 c1 c3t
diff --git a/kernel/fast_typeops.mli b/kernel/fast_typeops.mli
index 90d9c55f16..05d52b2d3c 100644
--- a/kernel/fast_typeops.mli
+++ b/kernel/fast_typeops.mli
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
diff --git a/kernel/indtypes.ml b/kernel/indtypes.ml
index e3457006d0..edb758f078 100644
--- a/kernel/indtypes.ml
+++ b/kernel/indtypes.ml
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -12,7 +12,6 @@ open Names
open Univ
open Term
open Vars
-open Context
open Declarations
open Declareops
open Inductive
@@ -21,6 +20,17 @@ open Reduction
open Typeops
open Entries
open Pp
+open Context.Rel.Declaration
+
+(* Terminology:
+paramdecls (ou paramsctxt?)
+args = params + realargs (called vargs when an array, largs when a list)
+params = recparams + nonrecparams
+nonrecargs = nonrecparams + realargs
+env_ar = initial env + declaration of inductive types
+env_ar_par = env_ar + declaration of parameters
+nmr = ongoing computation of recursive parameters
+*)
(* Tell if indices (aka real arguments) contribute to size of inductive type *)
(* If yes, this is compatible with the univalent model *)
@@ -117,18 +127,18 @@ let is_unit constrsinfos =
| [] -> (* type without constructors *) true
| _ -> false
-let infos_and_sort env ctx t =
- let rec aux env ctx t max =
+let infos_and_sort env t =
+ let rec aux env t max =
let t = whd_betadeltaiota env t in
match kind_of_term t with
| Prod (name,c1,c2) ->
let varj = infer_type env c1 in
- let env1 = Environ.push_rel (name,None,varj.utj_val) env in
+ let env1 = Environ.push_rel (LocalAssum (name,varj.utj_val)) env in
let max = Universe.sup max (univ_of_sort varj.utj_type) in
- aux env1 ctx c2 max
+ aux env1 c2 max
| _ when is_constructor_head t -> max
| _ -> (* don't fail if not positive, it is tested later *) max
- in aux env ctx t Universe.type0m
+ in aux env t Universe.type0m
(* Computing the levels of polymorphic inductive types
@@ -153,14 +163,14 @@ let infos_and_sort env ctx t =
(* This (re)computes informations relevant to extraction and the sort of an
arity or type constructor; we do not to recompute universes constraints *)
-let infer_constructor_packet env_ar_par ctx params lc =
+let infer_constructor_packet env_ar_par params lc =
(* type-check the constructors *)
let jlc = List.map (infer_type env_ar_par) lc in
let jlc = Array.of_list jlc in
(* generalize the constructor over the parameters *)
let lc'' = Array.map (fun j -> it_mkProd_or_LetIn j.utj_val params) jlc in
(* compute the max of the sorts of the products of the constructors types *)
- let levels = List.map (infos_and_sort env_ar_par ctx) lc in
+ let levels = List.map (infos_and_sort env_ar_par) lc in
let isunit = is_unit levels in
let min = if Array.length jlc > 1 then Universe.type0 else Universe.type0m in
let level = List.fold_left (fun max l -> Universe.sup max l) min levels in
@@ -169,12 +179,14 @@ let infer_constructor_packet env_ar_par ctx params lc =
(* If indices matter *)
let cumulate_arity_large_levels env sign =
fst (List.fold_right
- (fun (_,b,t as d) (lev,env) ->
- if Option.is_empty b then
+ (fun d (lev,env) ->
+ match d with
+ | LocalAssum (_,t) ->
let tj = infer_type env t in
let u = univ_of_sort tj.utj_type in
(Universe.sup u lev, push_rel d env)
- else lev, push_rel d env)
+ | LocalDef _ ->
+ lev, push_rel d env)
sign (Universe.type0m,env))
let is_impredicative env u =
@@ -184,15 +196,16 @@ let is_impredicative env u =
polymorphism. The elements x_k is None if the k-th parameter (starting
from the most recent and ignoring let-definitions) is not contributing
or is Some u_k if its level is u_k and is contributing. *)
-let param_ccls params =
- let fold acc = function (_, None, p) ->
+let param_ccls paramsctxt =
+ let fold acc = function
+ | (LocalAssum (_, p)) ->
(let c = strip_prod_assum p in
match kind_of_term c with
| Sort (Type u) -> Univ.Universe.level u
| _ -> None) :: acc
- | _ -> acc
+ | LocalDef _ -> acc
in
- List.fold_left fold [] params
+ List.fold_left fold [] paramsctxt
(* Type-check an inductive definition. Does not check positivity
conditions. *)
@@ -208,7 +221,7 @@ let typecheck_inductive env mie =
mind_check_names mie;
(* Params are typed-checked here *)
let env' = push_context mie.mind_entry_universes env in
- let (env_params, params) = infer_local_decls env' mie.mind_entry_params in
+ let (env_params,paramsctxt) = infer_local_decls env' mie.mind_entry_params in
(* We first type arity of each inductive definition *)
(* This allows building the environment of arities and to share *)
(* the set of constraints *)
@@ -247,27 +260,26 @@ let typecheck_inductive env mie =
later, after the validation of the inductive definition,
full_arity is used as argument or subject to cast, an
upper universe will be generated *)
- let full_arity = it_mkProd_or_LetIn arity params in
+ let full_arity = it_mkProd_or_LetIn arity paramsctxt in
let id = ind.mind_entry_typename in
let env_ar' =
- push_rel (Name id, None, full_arity) env_ar in
+ push_rel (LocalAssum (Name id, full_arity)) env_ar in
(* (add_constraints cst2 env_ar) in *)
- (env_ar', (id,full_arity,sign @ params,expltype,deflev,inflev)::l))
+ (env_ar', (id,full_arity,sign @ paramsctxt,expltype,deflev,inflev)::l))
(env',[])
mie.mind_entry_inds in
let arity_list = List.rev rev_arity_list in
(* builds the typing context "Gamma, I1:A1, ... In:An, params" *)
- let env_ar_par = push_rel_context params env_arities in
+ let env_ar_par = push_rel_context paramsctxt env_arities in
(* Now, we type the constructors (without params) *)
let inds =
List.fold_right2
(fun ind arity_data inds ->
let (lc',cstrs_univ) =
- infer_constructor_packet env_ar_par ContextSet.empty
- params ind.mind_entry_lc in
+ infer_constructor_packet env_ar_par paramsctxt ind.mind_entry_lc in
let consnames = ind.mind_entry_consnames in
let ind' = (arity_data,consnames,lc',cstrs_univ) in
ind'::inds)
@@ -290,8 +302,7 @@ let typecheck_inductive env mie =
let full_polymorphic () =
let defu = Term.univ_of_sort def_level in
let is_natural =
- type_in_type env || (check_leq (universes env') infu defu &&
- not (is_type0m_univ defu && not is_unit))
+ type_in_type env || (UGraph.check_leq (universes env') infu defu)
in
let _ =
(** Impredicative sort, always allow *)
@@ -317,14 +328,14 @@ let typecheck_inductive env mie =
(* conclusions of the parameters *)
(* We enforce [u >= lev] in case [lev] has a strict upper *)
(* constraints over [u] *)
- let b = type_in_type env || check_leq (universes env') infu u in
+ let b = type_in_type env || UGraph.check_leq (universes env') infu u in
if not b then
anomaly ~label:"check_inductive"
(Pp.str"Incorrect universe " ++
Universe.pr u ++ Pp.str " declared for inductive type, inferred level is "
++ Universe.pr clev)
else
- TemplateArity (param_ccls params, infu)
+ TemplateArity (param_ccls paramsctxt, infu)
| _ (* Not an explicit occurrence of Type *) ->
full_polymorphic ()
in
@@ -334,7 +345,7 @@ let typecheck_inductive env mie =
in
(id,cn,lc,(sign,arity)))
inds
- in (env_arities, env_ar_par, params, inds)
+ in (env_arities, env_ar_par, paramsctxt, inds)
(************************************************************************)
(************************************************************************)
@@ -343,7 +354,7 @@ let typecheck_inductive env mie =
type ill_formed_ind =
| LocalNonPos of int
| LocalNotEnoughArgs of int
- | LocalNotConstructor
+ | LocalNotConstructor of Context.Rel.t * int
| LocalNonPar of int * int * int
exception IllFormedInd of ill_formed_ind
@@ -354,20 +365,22 @@ exception IllFormedInd of ill_formed_ind
let mind_extract_params = decompose_prod_n_assum
-let explain_ind_err id ntyp env nbpar c nargs err =
- let (lpar,c') = mind_extract_params nbpar c in
+let explain_ind_err id ntyp env nparamsctxt c err =
+ let (lparams,c') = mind_extract_params nparamsctxt c in
match err with
| LocalNonPos kt ->
- raise (InductiveError (NonPos (env,c',mkRel (kt+nbpar))))
+ raise (InductiveError (NonPos (env,c',mkRel (kt+nparamsctxt))))
| LocalNotEnoughArgs kt ->
raise (InductiveError
- (NotEnoughArgs (env,c',mkRel (kt+nbpar))))
- | LocalNotConstructor ->
+ (NotEnoughArgs (env,c',mkRel (kt+nparamsctxt))))
+ | LocalNotConstructor (paramsctxt,nargs)->
+ let nparams = Context.Rel.nhyps paramsctxt in
raise (InductiveError
- (NotConstructor (env,id,c',mkRel (ntyp+nbpar),nbpar,nargs)))
+ (NotConstructor (env,id,c',mkRel (ntyp+nparamsctxt),
+ nparams,nargs)))
| LocalNonPar (n,i,l) ->
raise (InductiveError
- (NonPar (env,c',n,mkRel i, mkRel (l+nbpar))))
+ (NonPar (env,c',n,mkRel i,mkRel (l+nparamsctxt))))
let failwith_non_pos n ntypes c =
for k = n to n + ntypes - 1 do
@@ -383,43 +396,50 @@ let failwith_non_pos_list n ntypes l =
anomaly ~label:"failwith_non_pos_list" (Pp.str "some k in [n;n+ntypes-1] should occur")
(* Check the inductive type is called with the expected parameters *)
-let check_correct_par (env,n,ntypes,_) hyps l largs =
- let nparams = rel_context_nhyps hyps in
- let largs = Array.of_list largs in
- if Array.length largs < nparams then
- raise (IllFormedInd (LocalNotEnoughArgs l));
- let (lpar,largs') = Array.chop nparams largs in
- let nhyps = List.length hyps in
- let rec check k index = function
+(* [n] is the index of the last inductive type in [env] *)
+let check_correct_par (env,n,ntypes,_) paramdecls ind_index args =
+ let nparams = Context.Rel.nhyps paramdecls in
+ let args = Array.of_list args in
+ if Array.length args < nparams then
+ raise (IllFormedInd (LocalNotEnoughArgs ind_index));
+ let (params,realargs) = Array.chop nparams args in
+ let nparamdecls = List.length paramdecls in
+ let rec check param_index paramdecl_index = function
| [] -> ()
- | (_,Some _,_)::hyps -> check k (index+1) hyps
- | _::hyps ->
- match kind_of_term (whd_betadeltaiota env lpar.(k)) with
- | Rel w when Int.equal w index -> check (k-1) (index+1) hyps
- | _ -> raise (IllFormedInd (LocalNonPar (k+1, index, l)))
- in check (nparams-1) (n-nhyps) hyps;
- if not (Array.for_all (noccur_between n ntypes) largs') then
- failwith_non_pos_vect n ntypes largs'
-
-(* Computes the maximum number of recursive parameters :
- the first parameters which are constant in recursive arguments
- n is the current depth, nmr is the maximum number of possible
- recursive parameters *)
-
-let compute_rec_par (env,n,_,_) hyps nmr largs =
+ | LocalDef _ :: paramdecls ->
+ check param_index (paramdecl_index+1) paramdecls
+ | _::paramdecls ->
+ match kind_of_term (whd_betadeltaiota env params.(param_index)) with
+ | Rel w when Int.equal w paramdecl_index ->
+ check (param_index-1) (paramdecl_index+1) paramdecls
+ | _ ->
+ let paramdecl_index_in_env = paramdecl_index-n+nparamdecls+1 in
+ let err =
+ LocalNonPar (param_index+1, paramdecl_index_in_env, ind_index) in
+ raise (IllFormedInd err)
+ in check (nparams-1) (n-nparamdecls) paramdecls;
+ if not (Array.for_all (noccur_between n ntypes) realargs) then
+ failwith_non_pos_vect n ntypes realargs
+
+(* Computes the maximum number of recursive parameters:
+ the first parameters which are constant in recursive arguments
+ [n] is the current depth, [nmr] is the maximum number of possible
+ recursive parameters *)
+
+let compute_rec_par (env,n,_,_) paramsctxt nmr largs =
if Int.equal nmr 0 then 0 else
-(* start from 0, hyps will be in reverse order *)
+(* start from 0, params will be in reverse order *)
let (lpar,_) = List.chop nmr largs in
let rec find k index =
function
([],_) -> nmr
- | (_,[]) -> assert false (* |hyps|>=nmr *)
- | (lp,(_,Some _,_)::hyps) -> find k (index-1) (lp,hyps)
- | (p::lp,_::hyps) ->
+ | (_,[]) -> assert false (* |paramsctxt|>=nmr *)
+ | (lp, LocalDef _ :: paramsctxt) -> find k (index-1) (lp,paramsctxt)
+ | (p::lp,_::paramsctxt) ->
( match kind_of_term (whd_betadeltaiota env p) with
- | Rel w when Int.equal w index -> find (k+1) (index-1) (lp,hyps)
+ | Rel w when Int.equal w index -> find (k+1) (index-1) (lp,paramsctxt)
| _ -> k)
- in find 0 (n-1) (lpar,List.rev hyps)
+ in find 0 (n-1) (lpar,List.rev paramsctxt)
(* [env] is the typing environment
[n] is the dB of the last inductive type
@@ -428,15 +448,15 @@ if Int.equal nmr 0 then 0 else
[lra] is the list of recursive tree of each variable
*)
let ienv_push_var (env, n, ntypes, lra) (x,a,ra) =
- (push_rel (x,None,a) env, n+1, ntypes, (Norec,ra)::lra)
+ (push_rel (LocalAssum (x,a)) env, n+1, ntypes, (Norec,ra)::lra)
-let ienv_push_inductive (env, n, ntypes, ra_env) ((mi,u),lpar) =
+let ienv_push_inductive (env, n, ntypes, ra_env) ((mi,u),lrecparams) =
let auxntyp = 1 in
let specif = (lookup_mind_specif env mi, u) in
let ty = type_of_inductive env specif in
let env' =
- push_rel (Anonymous,None,
- hnf_prod_applist env ty lpar) env in
+ let decl = LocalAssum (Anonymous, hnf_prod_applist env ty lrecparams) in
+ push_rel decl env in
let ra_env' =
(Imbr mi,(Rtree.mk_rec_calls 1).(0)) ::
List.map (fun (r,t) -> (r,Rtree.lift 1 t)) ra_env in
@@ -456,7 +476,7 @@ let rec ienv_decompose_prod (env,_,_,_ as ienv) n c =
let array_min nmr a = if Int.equal nmr 0 then 0 else
Array.fold_left (fun k (nmri,_) -> min k nmri) nmr a
-(** [check_positivity_one ienv hyps (mind,i) nargs lcnames indlc]
+(** [check_positivity_one ienv paramsctxt (mind,i) nnonrecargs lcnames indlc]
checks the positivity of the [i]-th member of the mutually
inductive definition [mind]. It returns an [Rtree.t] which
represents the position of the recursive calls of inductive in [i]
@@ -464,9 +484,9 @@ let array_min nmr a = if Int.equal nmr 0 then 0 else
considered sub-terms) as well as the number of of non-uniform
arguments (used to generate induction schemes, so a priori less
relevant to the kernel). *)
-let check_positivity_one (env,_,ntypes,_ as ienv) hyps (_,i as ind) nargs lcnames indlc =
- let lparams = rel_context_length hyps in
- let nmr = rel_context_nhyps hyps in
+let check_positivity_one recursive (env,_,ntypes,_ as ienv) paramsctxt (_,i as ind) nnonrecargs lcnames indlc =
+ let nparamsctxt = Context.Rel.length paramsctxt in
+ let nmr = Context.Rel.nhyps paramsctxt in
(** Positivity of one argument [c] of a constructor (i.e. the
constructor [cn] has a type of the shape [… -> c … -> P], where,
more generally, the arrows may be dependent). *)
@@ -489,7 +509,7 @@ let check_positivity_one (env,_,ntypes,_ as ienv) hyps (_,i as ind) nargs lcname
let largs = List.map (whd_betadeltaiota env) largs in
let nmr1 =
(match ra with
- Mrec _ -> compute_rec_par ienv hyps nmr largs
+ Mrec _ -> compute_rec_par ienv paramsctxt nmr largs
| _ -> nmr)
in
(** The case where one of the inductives of the mutually
@@ -524,27 +544,27 @@ let check_positivity_one (env,_,ntypes,_ as ienv) hyps (_,i as ind) nargs lcname
(* accesses to the environment are not factorised, but is it worth? *)
and check_positive_nested (env,n,ntypes,ra_env as ienv) nmr ((mi,u), largs) =
let (mib,mip) = lookup_mind_specif env mi in
- let auxnpar = mib.mind_nparams_rec in
- let nonrecpar = mib.mind_nparams - auxnpar in
- let (lpar,auxlargs) =
- try List.chop auxnpar largs
+ let auxnrecpar = mib.mind_nparams_rec in
+ let auxnnonrecpar = mib.mind_nparams - auxnrecpar in
+ let (auxrecparams,auxnonrecargs) =
+ try List.chop auxnrecpar largs
with Failure _ -> raise (IllFormedInd (LocalNonPos n)) in
(** Inductives of the inductive block being defined are only
allowed to appear nested in the parameters of another inductive
type. Not in the proper indices. *)
- if not (List.for_all (noccur_between n ntypes) auxlargs) then
- failwith_non_pos_list n ntypes auxlargs;
+ if not (List.for_all (noccur_between n ntypes) auxnonrecargs) then
+ failwith_non_pos_list n ntypes auxnonrecargs;
(* Nested mutual inductive types are not supported *)
let auxntyp = mib.mind_ntypes in
if not (Int.equal auxntyp 1) then raise (IllFormedInd (LocalNonPos n));
(* The nested inductive type with parameters removed *)
- let auxlcvect = abstract_mind_lc auxntyp auxnpar mip.mind_nf_lc in
+ let auxlcvect = abstract_mind_lc auxntyp auxnrecpar mip.mind_nf_lc in
(* Extends the environment with a variable corresponding to
the inductive def *)
- let (env',_,_,_ as ienv') = ienv_push_inductive ienv ((mi,u),lpar) in
+ let (env',_,_,_ as ienv') = ienv_push_inductive ienv ((mi,u),auxrecparams) in
(* Parameters expressed in env' *)
- let lpar' = List.map (lift auxntyp) lpar in
+ let auxrecparams' = List.map (lift auxntyp) auxrecparams in
let irecargs_nmr =
(** Checks that the "nesting" inductive type is covariant in
the relevant parameters. In other words, that the
@@ -553,9 +573,9 @@ let check_positivity_one (env,_,ntypes,_ as ienv) hyps (_,i as ind) nargs lcname
positively in the types of the nested constructors. *)
Array.map
(function c ->
- let c' = hnf_prod_applist env' c lpar' in
+ let c' = hnf_prod_applist env' c auxrecparams' in
(* skip non-recursive parameters *)
- let (ienv',c') = ienv_decompose_prod ienv' nonrecpar c' in
+ let (ienv',c') = ienv_decompose_prod ienv' auxnnonrecpar c' in
check_constructors ienv' false nmr c')
auxlcvect
in
@@ -579,6 +599,8 @@ let check_positivity_one (env,_,ntypes,_ as ienv) hyps (_,i as ind) nargs lcname
| Prod (na,b,d) ->
let () = assert (List.is_empty largs) in
+ if not recursive && not (noccur_between n ntypes b) then
+ raise (InductiveError BadEntry);
let nmr',recarg = check_pos ienv nmr b in
let ienv' = ienv_push_var ienv (na,b,mk_norec) in
check_constr_rec ienv' nmr' (recarg::lrec) d
@@ -587,8 +609,8 @@ let check_positivity_one (env,_,ntypes,_ as ienv) hyps (_,i as ind) nargs lcname
if check_head then
begin match hd with
| Rel j when Int.equal j (n + ntypes - i - 1) ->
- check_correct_par ienv hyps (ntypes - i) largs
- | _ -> raise (IllFormedInd LocalNotConstructor)
+ check_correct_par ienv paramsctxt (ntypes - i) largs
+ | _ -> raise (IllFormedInd (LocalNotConstructor(paramsctxt,nnonrecargs)))
end
else
if not (List.for_all (noccur_between n ntypes) largs)
@@ -600,31 +622,32 @@ let check_positivity_one (env,_,ntypes,_ as ienv) hyps (_,i as ind) nargs lcname
let irecargs_nmr =
Array.map2
(fun id c ->
- let _,rawc = mind_extract_params lparams c in
+ let _,rawc = mind_extract_params nparamsctxt c in
try
check_constructors ienv true nmr rawc
with IllFormedInd err ->
- explain_ind_err id (ntypes-i) env lparams c nargs err)
+ explain_ind_err id (ntypes-i) env nparamsctxt c err)
(Array.of_list lcnames) indlc
in
let irecargs = Array.map snd irecargs_nmr
and nmr' = array_min nmr irecargs_nmr
in (nmr', mk_paths (Mrec ind) irecargs)
-(** [check_positivity kn env_ar params] checks that the mutually
+(** [check_positivity kn env_ar paramsctxt inds] checks that the mutually
inductive block [inds] is strictly positive. *)
-let check_positivity kn env_ar params inds =
+let check_positivity kn env_ar_par paramsctxt finite inds =
let ntypes = Array.length inds in
+ let recursive = finite != Decl_kinds.BiFinite in
let rc = Array.mapi (fun j t -> (Mrec (kn,j),t)) (Rtree.mk_rec_calls ntypes) in
- let lra_ind = Array.rev_to_list rc in
- let lparams = rel_context_length params in
- let nmr = rel_context_nhyps params in
+ let ra_env_ar = Array.rev_to_list rc in
+ let nparamsctxt = Context.Rel.length paramsctxt in
+ let nmr = Context.Rel.nhyps paramsctxt in
let check_one i (_,lcnames,lc,(sign,_)) =
- let ra_env =
- List.init lparams (fun _ -> (Norec,mk_norec)) @ lra_ind in
- let ienv = (env_ar, 1+lparams, ntypes, ra_env) in
- let nargs = rel_context_nhyps sign - nmr in
- check_positivity_one ienv params (kn,i) nargs lcnames lc
+ let ra_env_ar_par =
+ List.init nparamsctxt (fun _ -> (Norec,mk_norec)) @ ra_env_ar in
+ let ienv = (env_ar_par, 1+nparamsctxt, ntypes, ra_env_ar_par) in
+ let nnonrecargs = Context.Rel.nhyps sign - nmr in
+ check_positivity_one recursive ienv paramsctxt (kn,i) nnonrecargs lcnames lc
in
let irecargs_nmr = Array.mapi check_one inds in
let irecargs = Array.map snd irecargs_nmr
@@ -681,6 +704,7 @@ let used_section_variables env inds =
keep_hyps env ids
let rel_vect n m = Array.init m (fun i -> mkRel(n+m-i))
+let rel_list n m = Array.to_list (rel_vect n m)
exception UndefinableExpansion
@@ -695,24 +719,21 @@ let compute_projections ((kn, _ as ind), u as indu) n x nparamargs params
that typechecking projections requires just a substitution and not
matching with a parameter context. *)
let indty, paramsletsubst =
- let subst, inst =
- List.fold_right
- (fun (na, b, t) (subst, inst) ->
- match b with
- | None -> (mkRel 1 :: List.map (lift 1) subst,
- mkRel 1 :: List.map (lift 1) inst)
- | Some b -> (substl subst b) :: subst, List.map (lift 1) inst)
- paramslet ([], [])
- in
+ (* [ty] = [Ind inst] is typed in context [params] *)
+ let inst = Context.Rel.to_extended_vect 0 paramslet in
+ let ty = mkApp (mkIndU indu, inst) in
+ (* [Ind inst] is typed in context [params-wo-let] *)
+ let inst' = rel_list 0 nparamargs in
+ (* {params-wo-let |- subst:params] *)
+ let subst = subst_of_rel_context_instance paramslet inst' in
+ (* {params-wo-let, x:Ind inst' |- subst':(params,x:Ind inst)] *)
let subst = (* For the record parameter: *)
- mkRel 1 :: List.map (lift 1) subst
- in
- let ty = mkApp (mkIndU indu, CArray.rev_of_list inst) in
+ mkRel 1 :: List.map (lift 1) subst in
ty, subst
in
let ci =
let print_info =
- { ind_tags = []; cstr_tags = [|rel_context_tags ctx|]; style = LetStyle } in
+ { ind_tags = []; cstr_tags = [|Context.Rel.to_tags ctx|]; style = LetStyle } in
{ ci_ind = ind;
ci_npar = nparamargs;
ci_cstr_ndecls = mind_consnrealdecls;
@@ -730,16 +751,39 @@ let compute_projections ((kn, _ as ind), u as indu) n x nparamargs params
let body = mkCase (ci, p, mkRel 1, [|lift 1 branch|]) in
it_mkLambda_or_LetIn (mkLambda (x,indty,body)) params
in
- let projections (na, b, t) (i, j, kns, pbs, subst, letsubst) =
- match b with
- | Some c -> (i, j+1, kns, pbs, substl subst c :: subst,
- substl letsubst c :: subst)
- | None ->
+ let projections decl (i, j, kns, pbs, subst, letsubst) =
+ match decl with
+ | LocalDef (na,c,t) ->
+ (* From [params, field1,..,fieldj |- c(params,field1,..,fieldj)]
+ to [params, x:I, field1,..,fieldj |- c(params,field1,..,fieldj)] *)
+ let c = liftn 1 j c in
+ (* From [params, x:I, field1,..,fieldj |- c(params,field1,..,fieldj)]
+ to [params, x:I |- c(params,proj1 x,..,projj x)] *)
+ let c1 = substl subst c in
+ (* From [params, x:I |- subst:field1,..,fieldj]
+ to [params, x:I |- subst:field1,..,fieldj+1] where [subst]
+ is represented with instance of field1 last *)
+ let subst = c1 :: subst in
+ (* From [params, x:I, field1,..,fieldj |- c(params,field1,..,fieldj)]
+ to [params-wo-let, x:I |- c(params,proj1 x,..,projj x)] *)
+ let c2 = substl letsubst c in
+ (* From [params-wo-let, x:I |- subst:(params, x:I, field1,..,fieldj)]
+ to [params-wo-let, x:I |- subst:(params, x:I, field1,..,fieldj+1)] *)
+ let letsubst = c2 :: letsubst in
+ (i, j+1, kns, pbs, subst, letsubst)
+ | LocalAssum (na,t) ->
match na with
| Name id ->
let kn = Constant.make1 (KerName.make mp dp (Label.of_id id)) in
- let projty = substl letsubst (liftn 1 j t) in
- let ty = substl subst (liftn 1 j t) in
+ (* from [params, field1,..,fieldj |- t(params,field1,..,fieldj)]
+ to [params, x:I, field1,..,fieldj |- t(params,field1,..,fieldj] *)
+ let t = liftn 1 j t in
+ (* from [params, x:I, field1,..,fieldj |- t(params,field1,..,fieldj)]
+ to [params-wo-let, x:I |- t(params,proj1 x,..,projj x)] *)
+ let projty = substl letsubst t in
+ (* from [params, x:I, field1,..,fieldj |- t(field1,..,fieldj)]
+ to [params, x:I |- t(proj1 x,..,projj x)] *)
+ let ty = substl subst t in
let term = mkProj (Projection.make kn true, mkRel 1) in
let fterm = mkProj (Projection.make kn false, mkRel 1) in
let compat = compat_body ty (j - 1) in
@@ -758,14 +802,14 @@ let compute_projections ((kn, _ as ind), u as indu) n x nparamargs params
Array.of_list (List.rev kns),
Array.of_list (List.rev pbs)
-let build_inductive env p prv ctx env_ar params kn isrecord isfinite inds nmr recargs =
+let build_inductive env p prv ctx env_ar paramsctxt kn isrecord isfinite inds nmr recargs =
let ntypes = Array.length inds in
(* Compute the set of used section variables *)
let hyps = used_section_variables env inds in
- let nparamargs = rel_context_nhyps params in
- let nparamdecls = rel_context_length params in
+ let nparamargs = Context.Rel.nhyps paramsctxt in
+ let nparamsctxt = Context.Rel.length paramsctxt in
let subst, ctx = Univ.abstract_universes p ctx in
- let params = Vars.subst_univs_level_context subst params in
+ let paramsctxt = Vars.subst_univs_level_context subst paramsctxt in
let env_ar =
let ctx = Environ.rel_context env_ar in
let ctx' = Vars.subst_univs_level_context subst ctx in
@@ -778,10 +822,10 @@ let build_inductive env p prv ctx env_ar params kn isrecord isfinite inds nmr re
let splayed_lc = Array.map (dest_prod_assum env_ar) lc in
let nf_lc = Array.map (fun (d,b) -> it_mkProd_or_LetIn b d) splayed_lc in
let consnrealdecls =
- Array.map (fun (d,_) -> rel_context_length d - rel_context_length params)
+ Array.map (fun (d,_) -> Context.Rel.length d - nparamsctxt)
splayed_lc in
let consnrealargs =
- Array.map (fun (d,_) -> rel_context_nhyps d - rel_context_nhyps params)
+ Array.map (fun (d,_) -> Context.Rel.nhyps d - nparamargs)
splayed_lc in
(* Elimination sorts *)
let arkind,kelim =
@@ -814,8 +858,8 @@ let build_inductive env p prv ctx env_ar params kn isrecord isfinite inds nmr re
{ mind_typename = id;
mind_arity = arkind;
mind_arity_ctxt = Vars.subst_univs_level_context subst ar_sign;
- mind_nrealargs = rel_context_nhyps ar_sign - nparamargs;
- mind_nrealdecls = rel_context_length ar_sign - nparamdecls;
+ mind_nrealargs = Context.Rel.nhyps ar_sign - nparamargs;
+ mind_nrealdecls = Context.Rel.length ar_sign - nparamsctxt;
mind_kelim = kelim;
mind_consnames = Array.of_list cnames;
mind_consnrealdecls = consnrealdecls;
@@ -828,10 +872,11 @@ let build_inductive env p prv ctx env_ar params kn isrecord isfinite inds nmr re
mind_reloc_tbl = rtbl;
} in
let packets = Array.map2 build_one_packet inds recargs in
- let pkt = packets.(0) in
+ let pkt = packets.(0) in
let isrecord =
match isrecord with
- | Some (Some rid) when pkt.mind_kelim == all_sorts && Array.length pkt.mind_consnames == 1
+ | Some (Some rid) when pkt.mind_kelim == all_sorts
+ && Array.length pkt.mind_consnames == 1
&& pkt.mind_consnrealargs.(0) > 0 ->
(** The elimination criterion ensures that all projections can be defined. *)
let u =
@@ -844,7 +889,7 @@ let build_inductive env p prv ctx env_ar params kn isrecord isfinite inds nmr re
(try
let fields, paramslet = List.chop pkt.mind_consnrealdecls.(0) rctx in
let kns, projs =
- compute_projections indsp pkt.mind_typename rid nparamargs params
+ compute_projections indsp pkt.mind_typename rid nparamargs paramsctxt
pkt.mind_consnrealdecls pkt.mind_consnrealargs paramslet fields
in Some (Some (rid, kns, projs))
with UndefinableExpansion -> Some None)
@@ -858,7 +903,7 @@ let build_inductive env p prv ctx env_ar params kn isrecord isfinite inds nmr re
mind_hyps = hyps;
mind_nparams = nparamargs;
mind_nparams_rec = nmr;
- mind_params_ctxt = params;
+ mind_params_ctxt = paramsctxt;
mind_packets = packets;
mind_polymorphic = p;
mind_universes = ctx;
@@ -870,11 +915,11 @@ let build_inductive env p prv ctx env_ar params kn isrecord isfinite inds nmr re
let check_inductive env kn mie =
(* First type-check the inductive definition *)
- let (env_ar, env_ar_par, params, inds) = typecheck_inductive env mie in
+ let (env_ar, env_ar_par, paramsctxt, inds) = typecheck_inductive env mie in
(* Then check positivity conditions *)
- let (nmr,recargs) = check_positivity kn env_ar_par params inds in
+ let (nmr,recargs) = check_positivity kn env_ar_par paramsctxt mie.mind_entry_finite inds in
(* Build the inductive packets *)
build_inductive env mie.mind_entry_polymorphic mie.mind_entry_private
mie.mind_entry_universes
- env_ar params kn mie.mind_entry_record mie.mind_entry_finite
+ env_ar paramsctxt kn mie.mind_entry_record mie.mind_entry_finite
inds nmr recargs
diff --git a/kernel/indtypes.mli b/kernel/indtypes.mli
index 01acdce5c8..5b4615399d 100644
--- a/kernel/indtypes.mli
+++ b/kernel/indtypes.mli
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -42,6 +42,6 @@ val enforce_indices_matter : unit -> unit
val is_indices_matter : unit -> bool
val compute_projections : pinductive -> Id.t -> Id.t ->
- int -> Context.rel_context -> int array -> int array ->
- Context.rel_context -> Context.rel_context ->
+ int -> Context.Rel.t -> int array -> int array ->
+ Context.Rel.t -> Context.Rel.t ->
(constant array * projection_body array)
diff --git a/kernel/inductive.ml b/kernel/inductive.ml
index 87c139f48d..499cbf0dfd 100644
--- a/kernel/inductive.ml
+++ b/kernel/inductive.ml
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -12,12 +12,12 @@ open Names
open Univ
open Term
open Vars
-open Context
open Declarations
open Declareops
open Environ
open Reduction
open Type_errors
+open Context.Rel.Declaration
type mind_specif = mutual_inductive_body * one_inductive_body
@@ -77,11 +77,11 @@ let instantiate_params full t u args sign =
let fail () =
anomaly ~label:"instantiate_params" (Pp.str "type, ctxt and args mismatch") in
let (rem_args, subs, ty) =
- Context.fold_rel_context
- (fun (_,copt,_) (largs,subs,ty) ->
- match (copt, largs, kind_of_term ty) with
- | (None, a::args, Prod(_,_,t)) -> (args, a::subs, t)
- | (Some b,_,LetIn(_,_,_,t)) ->
+ Context.Rel.fold_outside
+ (fun decl (largs,subs,ty) ->
+ match (decl, largs, kind_of_term ty) with
+ | (LocalAssum _, a::args, Prod(_,_,t)) -> (args, a::subs, t)
+ | (LocalDef (_,b,_), _, LetIn(_,_,_,t)) ->
(largs, (substl subs (subst_instance_constr u b))::subs, t)
| (_,[],_) -> if full then fail() else ([], subs, ty)
| _ -> fail ())
@@ -151,9 +151,9 @@ let remember_subst u subst =
(* Bind expected levels of parameters to actual levels *)
(* Propagate the new levels in the signature *)
-let rec make_subst env =
+let make_subst env =
let rec make subst = function
- | (_,Some _,_)::sign, exp, args ->
+ | LocalDef _ :: sign, exp, args ->
make subst (sign, exp, args)
| d::sign, None::exp, args ->
let args = match args with _::args -> args | [] -> [] in
@@ -165,8 +165,8 @@ let rec make_subst env =
(* to be greater than the level of the argument; this is probably *)
(* a useless extra constraint *)
let s = sort_as_univ (snd (dest_arity env (Lazy.force a))) in
- make (cons_subst u s subst) (sign, exp, args)
- | (na,None,t)::sign, Some u::exp, [] ->
+ make (cons_subst u s subst) (sign, exp, args)
+ | LocalAssum (na,t) :: sign, Some u::exp, [] ->
(* No more argument here: we add the remaining universes to the *)
(* substitution (when [u] is distinct from all other universes in the *)
(* template, it is identity substitution otherwise (ie. when u is *)
@@ -270,18 +270,6 @@ let type_of_constructors (ind,u) (mib,mip) =
(* Type of case predicates *)
-let local_rels ctxt =
- let (rels,_) =
- Context.fold_rel_context_reverse
- (fun (rels,n) (_,copt,_) ->
- match copt with
- None -> (mkRel n :: rels, n+1)
- | Some _ -> (rels, n+1))
- ~init:([],1)
- ctxt
- in
- rels
-
(* Get type of inductive, with parameters instantiated *)
let inductive_sort_family mip =
@@ -304,20 +292,12 @@ let is_primitive_record (mib,_) =
| Some (Some _) -> true
| _ -> false
-let extended_rel_list n hyps =
- let rec reln l p = function
- | (_,None,_) :: hyps -> reln (mkRel (n+p) :: l) (p+1) hyps
- | (_,Some _,_) :: hyps -> reln l (p+1) hyps
- | [] -> l
- in
- reln [] 1 hyps
-
let build_dependent_inductive ind (_,mip) params =
let realargs,_ = List.chop mip.mind_nrealdecls mip.mind_arity_ctxt in
applist
(mkIndU ind,
List.map (lift mip.mind_nrealdecls) params
- @ extended_rel_list 0 realargs)
+ @ Context.Rel.to_extended_list 0 realargs)
(* This exception is local *)
exception LocalArity of (sorts_family * sorts_family * arity_error) option
@@ -335,14 +315,14 @@ let is_correct_arity env c pj ind specif params =
let rec srec env pt ar =
let pt' = whd_betadeltaiota env pt in
match kind_of_term pt', ar with
- | Prod (na1,a1,t), (_,None,a1')::ar' ->
+ | Prod (na1,a1,t), (LocalAssum (_,a1'))::ar' ->
let () =
try conv env a1 a1'
with NotConvertible -> raise (LocalArity None) in
- srec (push_rel (na1,None,a1) env) t ar'
+ srec (push_rel (LocalAssum (na1,a1)) env) t ar'
(* The last Prod domain is the type of the scrutinee *)
| Prod (na1,a1,a2), [] -> (* whnf of t was not needed here! *)
- let env' = push_rel (na1,None,a1) env in
+ let env' = push_rel (LocalAssum (na1,a1)) env in
let ksort = match kind_of_term (whd_betadeltaiota env' a2) with
| Sort s -> family_of_sort s
| _ -> raise (LocalArity None) in
@@ -351,7 +331,7 @@ let is_correct_arity env c pj ind specif params =
try conv env a1 dep_ind
with NotConvertible -> raise (LocalArity None) in
check_allowed_sort ksort specif
- | _, (_,Some _,_ as d)::ar' ->
+ | _, (LocalDef _ as d)::ar' ->
srec (push_rel d env) (lift 1 pt') ar'
| _ ->
raise (LocalArity None)
@@ -369,22 +349,22 @@ let is_correct_arity env c pj ind specif params =
let build_branches_type (ind,u) (_,mip as specif) params p =
let build_one_branch i cty =
let typi = full_constructor_instantiate (ind,u,specif,params) cty in
- let (args,ccl) = decompose_prod_assum typi in
- let nargs = rel_context_length args in
+ let (cstrsign,ccl) = decompose_prod_assum typi in
+ let nargs = Context.Rel.length cstrsign in
let (_,allargs) = decompose_app ccl in
let (lparams,vargs) = List.chop (inductive_params specif) allargs in
let cargs =
let cstr = ith_constructor_of_inductive ind (i+1) in
- let dep_cstr = applist (mkConstructU (cstr,u),lparams@(local_rels args)) in
+ let dep_cstr = applist (mkConstructU (cstr,u),lparams@(Context.Rel.to_extended_list 0 cstrsign)) in
vargs @ [dep_cstr] in
- let base = betazeta_appvect mip.mind_nrealdecls (lift nargs p) (Array.of_list cargs) in
- it_mkProd_or_LetIn base args in
+ let base = lambda_appvect_assum (mip.mind_nrealdecls+1) (lift nargs p) (Array.of_list cargs) in
+ it_mkProd_or_LetIn base cstrsign in
Array.mapi build_one_branch mip.mind_nf_lc
(* [p] is the predicate, [c] is the match object, [realargs] is the
list of real args of the inductive type *)
let build_case_type env n p c realargs =
- whd_betaiota env (betazeta_appvect (n+1) p (Array.of_list (realargs@[c])))
+ whd_betaiota env (lambda_appvect_assum (n+1) p (Array.of_list (realargs@[c])))
let type_case_branches env (pind,largs) pj c =
let specif = lookup_mind_specif env (fst pind) in
@@ -500,10 +480,10 @@ type guard_env =
let make_renv env recarg tree =
{ env = env;
rel_min = recarg+2; (* recarg = 0 ==> Rel 1 -> recarg; Rel 2 -> fix *)
- genv = [Lazy.lazy_from_val(Subterm(Large,tree))] }
+ genv = [Lazy.from_val(Subterm(Large,tree))] }
let push_var renv (x,ty,spec) =
- { env = push_rel (x,None,ty) renv.env;
+ { env = push_rel (LocalAssum (x,ty)) renv.env;
rel_min = renv.rel_min+1;
genv = spec:: renv.genv }
@@ -519,7 +499,7 @@ let subterm_var p renv =
with Failure _ | Invalid_argument _ -> Not_subterm
let push_ctxt_renv renv ctxt =
- let n = rel_context_length ctxt in
+ let n = Context.Rel.length ctxt in
{ env = push_rel_context ctxt renv.env;
rel_min = renv.rel_min+n;
genv = iterate (fun ge -> lazy Not_subterm::ge) n renv.genv }
@@ -589,14 +569,14 @@ let check_inductive_codomain env p =
(* The following functions are almost duplicated from indtypes.ml, except
that they carry here a poorer environment (containing less information). *)
let ienv_push_var (env, lra) (x,a,ra) =
- (push_rel (x,None,a) env, (Norec,ra)::lra)
+ (push_rel (LocalAssum (x,a)) env, (Norec,ra)::lra)
let ienv_push_inductive (env, ra_env) ((mind,u),lpar) =
let mib = Environ.lookup_mind mind env in
let ntypes = mib.mind_ntypes in
let push_ind specif env =
- push_rel (Anonymous,None,
- hnf_prod_applist env (type_of_inductive env ((mib,specif),u)) lpar) env
+ let decl = LocalAssum (Anonymous, hnf_prod_applist env (type_of_inductive env ((mib,specif),u)) lpar) in
+ push_rel decl env
in
let env = Array.fold_right push_ind mib.mind_packets env in
let rc = Array.mapi (fun j t -> (Imbr (mind,j),t)) (Rtree.mk_rec_calls ntypes) in
@@ -721,7 +701,7 @@ let restrict_spec env spec p =
else let absctx, ar = dest_lam_assum env p in
(* Optimization: if the predicate is not dependent, no restriction is needed
and we avoid building the recargs tree. *)
- if noccur_with_meta 1 (rel_context_length absctx) ar then spec
+ if noccur_with_meta 1 (Context.Rel.length absctx) ar then spec
else
let env = push_rel_context absctx env in
let arctx, s = dest_prod_assum env ar in
@@ -814,7 +794,15 @@ let rec subterm_specif renv stack t =
| Proj (p, c) ->
let subt = subterm_specif renv stack c in
(match subt with
- | Subterm (s, wf) -> Subterm (Strict, wf)
+ | Subterm (s, wf) ->
+ (* We take the subterm specs of the constructor of the record *)
+ let wf_args = (dest_subterms wf).(0) in
+ (* We extract the tree of the projected argument *)
+ let kn = Projection.constant p in
+ let cb = lookup_constant kn renv.env in
+ let pb = Option.get cb.const_proj in
+ let n = pb.proj_arg in
+ Subterm (Strict, List.nth wf_args n)
| Dead_code -> Dead_code
| Not_subterm -> Not_subterm)
@@ -829,7 +817,7 @@ and stack_element_specif = function
|SArg x -> x
and extract_stack renv a = function
- | [] -> Lazy.lazy_from_val Not_subterm , []
+ | [] -> Lazy.from_val Not_subterm , []
| h::t -> stack_element_specif h, t
(* Check term c can be applied to one of the mutual fixpoints. *)
@@ -863,13 +851,13 @@ let filter_stack_domain env ci p stack =
let absctx, ar = dest_lam_assum env p in
(* Optimization: if the predicate is not dependent, no restriction is needed
and we avoid building the recargs tree. *)
- if noccur_with_meta 1 (rel_context_length absctx) ar then stack
+ if noccur_with_meta 1 (Context.Rel.length absctx) ar then stack
else let env = push_rel_context absctx env in
let rec filter_stack env ar stack =
let t = whd_betadeltaiota env ar in
match stack, kind_of_term t with
| elt :: stack', Prod (n,a,c0) ->
- let d = (n,None,a) in
+ let d = LocalAssum (n,a) in
let ty, args = decompose_app (whd_betadeltaiota env a) in
let elt = match kind_of_term ty with
| Ind ind ->
@@ -926,10 +914,10 @@ let check_one_fix renv recpos trees def =
end
else
begin
- match pi2 (lookup_rel p renv.env) with
- | None ->
+ match lookup_rel p renv.env with
+ | LocalAssum _ ->
List.iter (check_rec_call renv []) l
- | Some c ->
+ | LocalDef (_,c,_) ->
try List.iter (check_rec_call renv []) l
with FixGuardError _ ->
check_rec_call renv stack (applist(lift p c,l))
@@ -1004,10 +992,11 @@ let check_one_fix renv recpos trees def =
| Var id ->
begin
- match pi2 (lookup_named id renv.env) with
- | None ->
+ let open Context.Named.Declaration in
+ match lookup_named id renv.env with
+ | LocalAssum _ ->
List.iter (check_rec_call renv []) l
- | Some c ->
+ | LocalDef (_,c,_) ->
try List.iter (check_rec_call renv []) l
with (FixGuardError _) ->
check_rec_call renv stack (applist(c,l))
@@ -1061,7 +1050,7 @@ let inductive_of_mutfix env ((nvect,bodynum),(names,types,bodies as recdef)) =
match kind_of_term (whd_betadeltaiota env def) with
| Lambda (x,a,b) ->
if noccur_with_meta n nbfix a then
- let env' = push_rel (x, None, a) env in
+ let env' = push_rel (LocalAssum (x,a)) env in
if Int.equal n (k + 1) then
(* get the inductive type of the fixpoint *)
let (mind, _) =
@@ -1111,7 +1100,7 @@ let rec codomain_is_coind env c =
let b = whd_betadeltaiota env c in
match kind_of_term b with
| Prod (x,a,b) ->
- codomain_is_coind (push_rel (x, None, a) env) b
+ codomain_is_coind (push_rel (LocalAssum (x,a)) env) b
| _ ->
(try find_coinductive env b
with Not_found ->
@@ -1152,7 +1141,7 @@ let check_one_cofix env nbfix def deftype =
| Lambda (x,a,b) ->
let () = assert (List.is_empty args) in
if noccur_with_meta n nbfix a then
- let env' = push_rel (x, None, a) env in
+ let env' = push_rel (LocalAssum (x,a)) env in
check_rec_call env' alreadygrd (n+1) tree vlra b
else
raise (CoFixGuardError (env,RecCallInTypeOfAbstraction a))
diff --git a/kernel/inductive.mli b/kernel/inductive.mli
index 5847d25f6f..c0d18bc6eb 100644
--- a/kernel/inductive.mli
+++ b/kernel/inductive.mli
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -8,7 +8,6 @@
open Names
open Term
-open Context
open Univ
open Declarations
open Environ
@@ -35,7 +34,7 @@ val lookup_mind_specif : env -> inductive -> mind_specif
(** {6 Functions to build standard types related to inductive } *)
val ind_subst : mutual_inductive -> mutual_inductive_body -> universe_instance -> constr list
-val inductive_paramdecls : mutual_inductive_body puniverses -> rel_context
+val inductive_paramdecls : mutual_inductive_body puniverses -> Context.Rel.t
val instantiate_inductive_constraints :
mutual_inductive_body -> universe_instance -> constraints
@@ -86,7 +85,7 @@ val build_branches_type :
constr list -> constr -> types array
(** Return the arity of an inductive type *)
-val mind_arity : one_inductive_body -> rel_context * sorts_family
+val mind_arity : one_inductive_body -> Context.Rel.t * sorts_family
val inductive_sort_family : one_inductive_body -> sorts_family
@@ -111,8 +110,8 @@ exception SingletonInductiveBecomesProp of Id.t
val max_inductive_sort : sorts array -> universe
-val instantiate_universes : env -> rel_context ->
- template_arity -> constr Lazy.t array -> rel_context * sorts
+val instantiate_universes : env -> Context.Rel.t ->
+ template_arity -> constr Lazy.t array -> Context.Rel.t * sorts
(** {6 Debug} *)
diff --git a/kernel/kernel.mllib b/kernel/kernel.mllib
index 29fe887d75..f7220c94a1 100644
--- a/kernel/kernel.mllib
+++ b/kernel/kernel.mllib
@@ -1,6 +1,7 @@
Names
Uint31
Univ
+UGraph
Esubst
Sorts
Evar
diff --git a/kernel/mod_subst.ml b/kernel/mod_subst.ml
index f7ae30e7af..95990bea6a 100644
--- a/kernel/mod_subst.ml
+++ b/kernel/mod_subst.ml
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -122,7 +122,7 @@ let add_kn_delta_resolver kn kn' =
let add_mp_delta_resolver mp1 mp2 = Deltamap.add_mp mp1 mp2
-(** Extending a [substitution] *)
+(** Extending a [substitution] without sequential composition *)
let add_mbid mbid mp resolve s = Umap.add_mbi mbid (mp,resolve) s
let add_mp mp1 mp2 resolve s = Umap.add_mp mp1 (mp2,resolve) s
diff --git a/kernel/mod_subst.mli b/kernel/mod_subst.mli
index fc2b0441ca..6d86b94167 100644
--- a/kernel/mod_subst.mli
+++ b/kernel/mod_subst.mli
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -68,8 +68,9 @@ val empty_subst : substitution
val is_empty_subst : substitution -> bool
-(** add_* add [arg2/arg1]\{arg3\} to the substitution with no
- sequential composition *)
+(** add_* add [arg2/arg1]\{arg3\} to the substitution with no sequential
+ composition. Most often this is not what you want. For sequential
+ composition, try [join (map_mbid mp delta) subs] **)
val add_mbid :
MBId.t -> module_path -> delta_resolver -> substitution -> substitution
val add_mp :
diff --git a/kernel/mod_typing.ml b/kernel/mod_typing.ml
index 4f20e5f62a..ff44f0f540 100644
--- a/kernel/mod_typing.ml
+++ b/kernel/mod_typing.ml
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -21,7 +21,7 @@ open Modops
open Mod_subst
type 'alg translation =
- module_signature * 'alg option * delta_resolver * Univ.constraints
+ module_signature * 'alg * delta_resolver * Univ.ContextSet.t
let rec mp_from_mexpr = function
| MEident mp -> mp
@@ -52,7 +52,7 @@ let rec rebuild_mp mp l =
| []-> mp
| i::r -> rebuild_mp (MPdot(mp,Label.of_id i)) r
-let (+++) = Univ.Constraint.union
+let (+++) = Univ.ContextSet.union
let rec check_with_def env struc (idl,(c,ctx)) mp equiv =
let lab,idl = match idl with
@@ -72,33 +72,71 @@ let rec check_with_def env struc (idl,(c,ctx)) mp equiv =
(* In the spirit of subtyping.check_constant, we accept
any implementations of parameters and opaques terms,
as long as they have the right type *)
- let ccst = Declareops.constraints_of_constant (opaque_tables env) cb in
- let env' = Environ.add_constraints ccst env' in
- let newus, cst = Univ.UContext.dest ctx in
- let env' = Environ.add_constraints cst env' in
- let c',cst = match cb.const_body with
- | Undef _ | OpaqueDef _ ->
- let j = Typeops.infer env' c in
- let typ = Typeops.type_of_constant_type env' cb.const_type in
- let cst' = Reduction.infer_conv_leq env' (Environ.universes env')
- j.uj_type typ in
- j.uj_val,cst' +++ cst
- | Def cs ->
- let cst' = Reduction.infer_conv env' (Environ.universes env') c
- (Mod_subst.force_constr cs) in
- let cst = (*FIXME MS: what to check here? subtyping of polymorphic constants... *)
- if cb.const_polymorphic then cst' +++ cst
- else cst' +++ cst
+ let uctx = Declareops.universes_of_constant (opaque_tables env) cb in
+ let uctx = (* Context of the spec *)
+ if cb.const_polymorphic then
+ Univ.instantiate_univ_context uctx
+ else uctx
+ in
+ let c', univs, ctx' =
+ if not cb.const_polymorphic then
+ let env' = Environ.push_context ~strict:true uctx env' in
+ let env' = Environ.push_context ~strict:true ctx env' in
+ let c',cst = match cb.const_body with
+ | Undef _ | OpaqueDef _ ->
+ let j = Typeops.infer env' c in
+ let typ = Typeops.type_of_constant_type env' cb.const_type in
+ let cst' = Reduction.infer_conv_leq env' (Environ.universes env')
+ j.uj_type typ in
+ j.uj_val, cst'
+ | Def cs ->
+ let c' = Mod_subst.force_constr cs in
+ c, Reduction.infer_conv env' (Environ.universes env') c c'
+ in c', ctx, Univ.ContextSet.add_constraints cst (Univ.ContextSet.of_context ctx)
+ else
+ let cus, ccst = Univ.UContext.dest uctx in
+ let newus, cst = Univ.UContext.dest ctx in
+ let () =
+ if not (Univ.Instance.length cus == Univ.Instance.length newus) then
+ error_incorrect_with_constraint lab
+ in
+ let inst = Univ.Instance.append cus newus in
+ let csti = Univ.enforce_eq_instances cus newus cst in
+ let csta = Univ.Constraint.union csti ccst in
+ let env' = Environ.push_context ~strict:false (Univ.UContext.make (inst, csta)) env in
+ let () = if not (UGraph.check_constraints cst (Environ.universes env')) then
+ error_incorrect_with_constraint lab
in
- c, cst
+ let cst = match cb.const_body with
+ | Undef _ | OpaqueDef _ ->
+ let j = Typeops.infer env' c in
+ let typ = Typeops.type_of_constant_type env' cb.const_type in
+ let typ = Vars.subst_instance_constr cus typ in
+ let cst' = Reduction.infer_conv_leq env' (Environ.universes env')
+ j.uj_type typ in
+ cst'
+ | Def cs ->
+ let c' = Vars.subst_instance_constr cus (Mod_subst.force_constr cs) in
+ let cst' = Reduction.infer_conv env' (Environ.universes env') c c' in
+ cst'
+ in
+ if not (Univ.Constraint.is_empty cst) then
+ error_incorrect_with_constraint lab;
+ let subst, ctx = Univ.abstract_universes true ctx in
+ Vars.subst_univs_level_constr subst c, ctx, Univ.ContextSet.empty
in
let def = Def (Mod_subst.from_val c') in
- let ctx' = Univ.UContext.make (newus, cst) in
+(* let ctx' = Univ.UContext.make (newus, cst) in *)
+ let univs =
+ if cb.const_polymorphic then Some cb.const_universes
+ else None
+ in
let cb' =
{ cb with
const_body = def;
- const_body_code = Option.map Cemitcodes.from_val (compile_constant_body env' def);
- const_universes = ctx' }
+ const_universes = ctx ;
+ const_body_code = Option.map Cemitcodes.from_val
+ (compile_constant_body env' univs def) }
in
before@(lab,SFBconst(cb'))::after, c', ctx'
else
@@ -145,9 +183,11 @@ let rec check_with_mod env struc (idl,mp1) mp equiv =
begin
try
let mtb_old = module_type_of_module old in
- Subtyping.check_subtypes env' mtb_mp1 mtb_old
- +++ old.mod_constraints
- with Failure _ -> error_incorrect_with_constraint lab
+ let chk_cst = Subtyping.check_subtypes env' mtb_mp1 mtb_old in
+ Univ.ContextSet.add_constraints chk_cst old.mod_constraints
+ with Failure _ ->
+ (* TODO: where can a Failure come from ??? *)
+ error_incorrect_with_constraint lab
end
| Algebraic (NoFunctor (MEident(mp'))) ->
check_modpath_equiv env' mp1 mp';
@@ -194,138 +234,157 @@ let rec check_with_mod env struc (idl,mp1) mp equiv =
| Algebraic (NoFunctor (MEident mp0)) ->
let mpnew = rebuild_mp mp0 idl in
check_modpath_equiv env' mpnew mp;
- before@(lab,spec)::after, equiv, Univ.Constraint.empty
+ before@(lab,spec)::after, equiv, Univ.ContextSet.empty
| _ -> error_generative_module_expected lab
end
with
| Not_found -> error_no_such_label lab
| Reduction.NotConvertible -> error_incorrect_with_constraint lab
-let mk_alg_with alg wd = Option.map (fun a -> MEwith (a,wd)) alg
-
let check_with env mp (sign,alg,reso,cst) = function
|WithDef(idl,c) ->
let struc = destr_nofunctor sign in
let struc',c',cst' = check_with_def env struc (idl,c) mp reso in
- let alg' = mk_alg_with alg (WithDef (idl,(c',cst'))) in
- (NoFunctor struc'),alg',reso, cst+++(Univ.UContext.constraints cst')
+ let wd' = WithDef (idl,(c',Univ.ContextSet.to_context cst')) in
+ NoFunctor struc', MEwith (alg,wd'), reso, cst+++cst'
|WithMod(idl,mp1) as wd ->
let struc = destr_nofunctor sign in
let struc',reso',cst' = check_with_mod env struc (idl,mp1) mp reso in
- let alg' = mk_alg_with alg wd in
- (NoFunctor struc'),alg',reso', cst+++cst'
+ NoFunctor struc', MEwith (alg,wd), reso', cst+++cst'
-let mk_alg_app mpo alg arg = match mpo, alg with
- | Some _, Some alg -> Some (MEapply (alg,arg))
- | _ -> None
+let translate_apply env inl (sign,alg,reso,cst1) mp1 mkalg =
+ let farg_id, farg_b, fbody_b = destr_functor sign in
+ let mtb = module_type_of_module (lookup_module mp1 env) in
+ let cst2 = Subtyping.check_subtypes env mtb farg_b in
+ let mp_delta = discr_resolver mtb in
+ let mp_delta = inline_delta_resolver env inl mp1 farg_id farg_b mp_delta in
+ let subst = map_mbid farg_id mp1 mp_delta in
+ let body = subst_signature subst fbody_b in
+ let alg' = mkalg alg mp1 in
+ let reso' = subst_codom_delta_resolver subst reso in
+ body,alg',reso', Univ.ContextSet.add_constraints cst2 cst1
(** Translation of a module struct entry :
- We translate to a module when a [module_path] is given,
otherwise to a module type.
- The first output is the expanded signature
- The second output is the algebraic expression, kept for the extraction.
- It is never None when translating to a module, but for module type
- it could not be contain [SEBapply] or [SEBfunctor].
*)
+let mk_alg_app alg arg = MEapply (alg,arg)
+
let rec translate_mse env mpo inl = function
- |MEident mp1 ->
- let sign,reso = match mpo with
- |Some mp ->
- let mb = strengthen_and_subst_mb (lookup_module mp1 env) mp false in
- mb.mod_type, mb.mod_delta
- |None ->
- let mtb = lookup_modtype mp1 env in
- mtb.mod_type, mtb.mod_delta
+ |MEident mp1 as me ->
+ let mb = match mpo with
+ |Some mp -> strengthen_and_subst_mb (lookup_module mp1 env) mp false
+ |None -> lookup_modtype mp1 env
in
- sign,Some (MEident mp1),reso,Univ.Constraint.empty
+ mb.mod_type, me, mb.mod_delta, Univ.ContextSet.empty
|MEapply (fe,mp1) ->
- translate_apply env inl (translate_mse env mpo inl fe) mp1 (mk_alg_app mpo)
+ translate_apply env inl (translate_mse env mpo inl fe) mp1 mk_alg_app
|MEwith(me, with_decl) ->
assert (mpo == None); (* No 'with' syntax for modules *)
let mp = mp_from_mexpr me in
check_with env mp (translate_mse env None inl me) with_decl
-and translate_apply env inl (sign,alg,reso,cst1) mp1 mkalg =
- let farg_id, farg_b, fbody_b = destr_functor sign in
- let mtb = module_type_of_module (lookup_module mp1 env) in
- let cst2 = Subtyping.check_subtypes env mtb farg_b in
- let mp_delta = discr_resolver mtb in
- let mp_delta = inline_delta_resolver env inl mp1 farg_id farg_b mp_delta in
- let subst = map_mbid farg_id mp1 mp_delta in
- let body = subst_signature subst fbody_b in
- let alg' = mkalg alg mp1 in
- let reso' = subst_codom_delta_resolver subst reso in
- body,alg',reso', cst1 +++ cst2
-
-let mk_alg_funct mpo mbid mtb alg = match mpo, alg with
- | Some _, Some alg -> Some (MoreFunctor (mbid,mtb,alg))
- | _ -> None
-
-let mk_mod mp e ty ty' cst reso =
+let mk_mod mp e ty cst reso =
{ mod_mp = mp;
mod_expr = e;
mod_type = ty;
- mod_type_alg = ty';
+ mod_type_alg = None;
mod_constraints = cst;
mod_delta = reso;
mod_retroknowledge = [] }
-let mk_modtype mp ty cst reso = mk_mod mp Abstract ty None cst reso
+let mk_modtype mp ty cst reso = mk_mod mp Abstract ty cst reso
let rec translate_mse_funct env mpo inl mse = function
|[] ->
let sign,alg,reso,cst = translate_mse env mpo inl mse in
- sign, Option.map (fun a -> NoFunctor a) alg, reso, cst
+ sign, NoFunctor alg, reso, cst
|(mbid, ty) :: params ->
let mp_id = MPbound mbid in
let mtb = translate_modtype env mp_id inl ([],ty) in
let env' = add_module_type mp_id mtb env in
let sign,alg,reso,cst = translate_mse_funct env' mpo inl mse params in
- let alg' = mk_alg_funct mpo mbid mtb alg in
+ let alg' = MoreFunctor (mbid,mtb,alg) in
MoreFunctor (mbid, mtb, sign), alg',reso, cst +++ mtb.mod_constraints
and translate_modtype env mp inl (params,mte) =
let sign,alg,reso,cst = translate_mse_funct env None inl mte params in
let mtb = mk_modtype (mp_from_mexpr mte) sign cst reso in
let mtb' = subst_modtype_and_resolver mtb mp in
- { mtb' with mod_type_alg = alg }
+ { mtb' with mod_type_alg = Some alg }
(** [finalize_module] :
- from an already-translated (or interactive) implementation
- and a signature entry, produce a final [module_expr] *)
+ from an already-translated (or interactive) implementation and
+ an (optional) signature entry, produces a final [module_body] *)
let finalize_module env mp (sign,alg,reso,cst) restype = match restype with
|None ->
let impl = match alg with Some e -> Algebraic e | None -> FullStruct in
- mk_mod mp impl sign None cst reso
+ mk_mod mp impl sign cst reso
|Some (params_mte,inl) ->
let res_mtb = translate_modtype env mp inl params_mte in
- let auto_mtb = mk_modtype mp sign Univ.Constraint.empty reso in
+ let auto_mtb = mk_modtype mp sign Univ.ContextSet.empty reso in
let cst' = Subtyping.check_subtypes env auto_mtb res_mtb in
let impl = match alg with Some e -> Algebraic e | None -> Struct sign in
{ res_mtb with
mod_mp = mp;
mod_expr = impl;
- (** cst from module body typing, cst' from subtyping,
- and constraints from module type. *)
- mod_constraints = cst +++ cst' +++ res_mtb.mod_constraints }
+ (** cst from module body typing,
+ cst' from subtyping,
+ constraints from module type. *)
+ mod_constraints =
+ Univ.ContextSet.add_constraints cst' (cst +++ res_mtb.mod_constraints) }
let translate_module env mp inl = function
|MType (params,ty) ->
let mtb = translate_modtype env mp inl (params,ty) in
module_body_of_type mp mtb
|MExpr (params,mse,oty) ->
- let t = translate_mse_funct env (Some mp) inl mse params in
+ let (sg,alg,reso,cst) = translate_mse_funct env (Some mp) inl mse params in
let restype = Option.map (fun ty -> ((params,ty),inl)) oty in
- finalize_module env mp t restype
+ finalize_module env mp (sg,Some alg,reso,cst) restype
+
+(** We now forbid any Include of functors with restricted signatures.
+ Otherwise, we could end with the creation of undesired axioms
+ (see #3746). Note that restricted non-functorized modules are ok,
+ thanks to strengthening. *)
-let rec translate_mse_incl env mp inl = function
+let rec unfunct = function
+ |NoFunctor me -> me
+ |MoreFunctor(_,_,me) -> unfunct me
+
+let rec forbid_incl_signed_functor env = function
+ |MEapply(fe,_) -> forbid_incl_signed_functor env fe
+ |MEwith _ -> assert false (* No 'with' syntax for modules *)
+ |MEident mp1 ->
+ let mb = lookup_module mp1 env in
+ match mb.mod_type, mb.mod_type_alg, mb.mod_expr with
+ |MoreFunctor _, Some _, _ ->
+ (* functor + restricted signature = error *)
+ error_include_restricted_functor mp1
+ |MoreFunctor _, None, Algebraic me ->
+ (* functor, no signature yet, a definition which may be restricted *)
+ forbid_incl_signed_functor env (unfunct me)
+ |_ -> ()
+
+let rec translate_mse_inclmod env mp inl = function
|MEident mp1 ->
let mb = strengthen_and_subst_mb (lookup_module mp1 env) mp true in
let sign = clean_bounded_mod_expr mb.mod_type in
- sign,None,mb.mod_delta,Univ.Constraint.empty
+ sign,(),mb.mod_delta,Univ.ContextSet.empty
|MEapply (fe,arg) ->
- let ftrans = translate_mse_incl env mp inl fe in
- translate_apply env inl ftrans arg (fun _ _ -> None)
- |_ -> Modops.error_higher_order_include ()
+ let ftrans = translate_mse_inclmod env mp inl fe in
+ translate_apply env inl ftrans arg (fun _ _ -> ())
+ |MEwith _ -> assert false (* No 'with' syntax for modules *)
+
+let translate_mse_incl is_mod env mp inl me =
+ if is_mod then
+ let () = forbid_incl_signed_functor env me in
+ translate_mse_inclmod env mp inl me
+ else
+ let mtb = translate_modtype env mp inl ([],me) in
+ let sign = clean_bounded_mod_expr mtb.mod_type in
+ sign,(),mtb.mod_delta,mtb.mod_constraints
diff --git a/kernel/mod_typing.mli b/kernel/mod_typing.mli
index b39e821254..5949dad08c 100644
--- a/kernel/mod_typing.mli
+++ b/kernel/mod_typing.mli
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -14,9 +14,18 @@ open Names
(** Main functions for translating module entries *)
+(** [translate_module] produces a [module_body] out of a [module_entry].
+ In the output fields:
+ - [mod_expr] is [Abstract] for a [MType] entry, or [Algebraic] for [MExpr].
+ - [mod_type_alg] is [None] only for a [MExpr] without explicit signature.
+*)
+
val translate_module :
env -> module_path -> inline -> module_entry -> module_body
+(** [translate_modtype] produces a [module_type_body] whose [mod_type_alg]
+ cannot be [None] (and of course [mod_expr] is [Abstract]). *)
+
val translate_modtype :
env -> module_path -> inline -> module_type_entry -> module_type_body
@@ -24,23 +33,27 @@ val translate_modtype :
- We translate to a module when a [module_path] is given,
otherwise to a module type.
- The first output is the expanded signature
- - The second output is the algebraic expression, kept for the extraction.
- It is never None when translating to a module, but for module type
- it could not be contain applications or functors.
-*)
+ - The second output is the algebraic expression, kept mostly for
+ the extraction. *)
type 'alg translation =
- module_signature * 'alg option * delta_resolver * Univ.constraints
+ module_signature * 'alg * delta_resolver * Univ.ContextSet.t
val translate_mse :
env -> module_path option -> inline -> module_struct_entry ->
module_alg_expr translation
-val translate_mse_incl :
- env -> module_path -> inline -> module_struct_entry ->
- module_alg_expr translation
+(** From an already-translated (or interactive) implementation and
+ an (optional) signature entry, produces a final [module_body] *)
val finalize_module :
- env -> module_path -> module_expression translation ->
+ env -> module_path -> (module_expression option) translation ->
(module_type_entry * inline) option ->
module_body
+
+(** [translate_mse_incl] translate the mse of a module or
+ module type given to an Include *)
+
+val translate_mse_incl :
+ bool -> env -> module_path -> inline -> module_struct_entry ->
+ unit translation
diff --git a/kernel/modops.ml b/kernel/modops.ml
index d52fe611c0..6fe7e382c6 100644
--- a/kernel/modops.ml
+++ b/kernel/modops.ml
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -67,16 +67,13 @@ type module_typing_error =
| IncorrectWithConstraint of Label.t
| GenerativeModuleExpected of Label.t
| LabelMissing of Label.t * string
- | HigherOrderInclude
+ | IncludeRestrictedFunctor of module_path
exception ModuleTypingError of module_typing_error
let error_existing_label l =
raise (ModuleTypingError (LabelAlreadyDeclared l))
-let error_application_to_not_path mexpr =
- raise (ModuleTypingError (ApplicationToNotPath mexpr))
-
let error_not_a_functor () =
raise (ModuleTypingError NotAFunctor)
@@ -113,8 +110,8 @@ let error_generative_module_expected l =
let error_no_such_label_sub l l1 =
raise (ModuleTypingError (LabelMissing (l,l1)))
-let error_higher_order_include () =
- raise (ModuleTypingError HigherOrderInclude)
+let error_include_restricted_functor mp =
+ raise (ModuleTypingError (IncludeRestrictedFunctor mp))
(** {6 Operations on functors } *)
@@ -331,13 +328,15 @@ let strengthen_const mp_from l cb resolver =
let kn = KerName.make2 mp_from l in
let con = constant_of_delta_kn resolver kn in
let u =
- if cb.const_polymorphic then
- Univ.UContext.instance cb.const_universes
+ if cb.const_polymorphic then
+ let u = Univ.UContext.instance cb.const_universes in
+ let s = Univ.make_instance_subst u in
+ Univ.subst_univs_level_instance s u
else Univ.Instance.empty
in
{ cb with
const_body = Def (Mod_subst.from_val (mkConstU (con,u)));
- const_body_code = Some (Cemitcodes.from_val (Cbytegen.compile_alias (con,u))) }
+ const_body_code = Some (Cemitcodes.from_val (Cbytegen.compile_alias con)) }
let rec strengthen_mod mp_from mp_to mb =
if mp_in_delta mb.mod_mp mb.mod_delta then mb
diff --git a/kernel/modops.mli b/kernel/modops.mli
index 6fbcd81d03..e9f3db6e91 100644
--- a/kernel/modops.mli
+++ b/kernel/modops.mli
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -126,14 +126,12 @@ type module_typing_error =
| IncorrectWithConstraint of Label.t
| GenerativeModuleExpected of Label.t
| LabelMissing of Label.t * string
- | HigherOrderInclude
+ | IncludeRestrictedFunctor of module_path
exception ModuleTypingError of module_typing_error
val error_existing_label : Label.t -> 'a
-val error_application_to_not_path : module_struct_entry -> 'a
-
val error_incompatible_modtypes :
module_type_body -> module_type_body -> 'a
@@ -154,4 +152,4 @@ val error_generative_module_expected : Label.t -> 'a
val error_no_such_label_sub : Label.t->string->'a
-val error_higher_order_include : unit -> 'a
+val error_include_restricted_functor : module_path -> 'a
diff --git a/kernel/names.ml b/kernel/names.ml
index ae2b3b6389..8e0237863f 100644
--- a/kernel/names.ml
+++ b/kernel/names.ml
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -23,6 +23,7 @@ open Util
(** {6 Identifiers } *)
+(** Representation and operations on identifiers. *)
module Id =
struct
type t = string
@@ -74,10 +75,18 @@ struct
end
-
+(** Representation and operations on identifiers that are allowed to be anonymous
+ (i.e. "_" in concrete syntax). *)
module Name =
struct
- type t = Name of Id.t | Anonymous
+ type t = Anonymous (** anonymous identifier *)
+ | Name of Id.t (** non-anonymous identifier *)
+
+ let is_anonymous = function
+ | Anonymous -> true
+ | Name _ -> false
+
+ let is_name = not % is_anonymous
let compare n1 n2 = match n1, n2 with
| Anonymous, Anonymous -> 0
@@ -102,7 +111,7 @@ struct
let hashcons hident = function
| Name id -> Name (hident id)
| n -> n
- let equal n1 n2 =
+ let eq n1 n2 =
n1 == n2 ||
match (n1,n2) with
| (Name id1, Name id2) -> id1 == id2
@@ -117,8 +126,8 @@ struct
end
-type name = Name.t = Name of Id.t | Anonymous
(** Alias, to import constructors. *)
+type name = Name.t = Anonymous | Name of Id.t
(** {6 Various types based on identifiers } *)
@@ -204,7 +213,7 @@ struct
DirPath.to_string p ^ "." ^ s
let debug_to_string (i, s, p) =
- "<"(*^string_of_dirpath p ^"#"^*) ^ s ^"#"^ string_of_int i^">"
+ "<"^DirPath.to_string p ^"#" ^ s ^"#"^ string_of_int i^">"
let compare (x : t) (y : t) =
if x == y then 0
@@ -236,7 +245,7 @@ struct
type t = _t
type u = (Id.t -> Id.t) * (DirPath.t -> DirPath.t)
let hashcons (hid,hdir) (n,s,dir) = (n,hid s,hdir dir)
- let equal ((n1,s1,dir1) as x) ((n2,s2,dir2) as y) =
+ let eq ((n1,s1,dir1) as x) ((n2,s2,dir2) as y) =
(x == y) ||
(Int.equal n1 n2 && s1 == s2 && dir1 == dir2)
let hash = hash
@@ -282,6 +291,11 @@ module ModPath = struct
| MPbound uid -> MBId.to_string uid
| MPdot (mp,l) -> to_string mp ^ "." ^ Label.to_string l
+ let rec debug_to_string = function
+ | MPfile sl -> DirPath.to_string sl
+ | MPbound uid -> MBId.debug_to_string uid
+ | MPdot (mp,l) -> debug_to_string mp ^ "." ^ Label.to_string l
+
(** we compare labels first if both are MPdots *)
let rec compare mp1 mp2 =
if mp1 == mp2 then 0
@@ -327,7 +341,7 @@ module ModPath = struct
| MPfile dir -> MPfile (hdir dir)
| MPbound m -> MPbound (huniqid m)
| MPdot (md,l) -> MPdot (hashcons hfuns md, hstr l)
- let rec equal d1 d2 =
+ let rec eq d1 d2 =
d1 == d2 ||
match d1,d2 with
| MPfile dir1, MPfile dir2 -> dir1 == dir2
@@ -375,12 +389,16 @@ module KerName = struct
let modpath kn = kn.modpath
let label kn = kn.knlabel
- let to_string kn =
+ let to_string_gen mp_to_string kn =
let dp =
if DirPath.is_empty kn.dirpath then "."
else "#" ^ DirPath.to_string kn.dirpath ^ "#"
in
- ModPath.to_string kn.modpath ^ dp ^ Label.to_string kn.knlabel
+ mp_to_string kn.modpath ^ dp ^ Label.to_string kn.knlabel
+
+ let to_string kn = to_string_gen ModPath.to_string kn
+
+ let debug_to_string kn = to_string_gen ModPath.debug_to_string kn
let print kn = str (to_string kn)
@@ -423,7 +441,7 @@ module KerName = struct
let hashcons (hmod,hdir,hstr) kn =
let { modpath = mp; dirpath = dp; knlabel = l; refhash; } = kn in
{ modpath = hmod mp; dirpath = hdir dp; knlabel = hstr l; refhash; canary; }
- let equal kn1 kn2 =
+ let eq kn1 kn2 =
kn1.modpath == kn2.modpath && kn1.dirpath == kn2.dirpath &&
kn1.knlabel == kn2.knlabel
let hash = hash
@@ -477,7 +495,7 @@ module KerPair = struct
| Dual (kn,_) -> kn
let same kn = Same kn
- let make knu knc = if knu == knc then Same knc else Dual (knu,knc)
+ let make knu knc = if KerName.equal knu knc then Same knc else Dual (knu,knc)
let make1 = same
let make2 mp l = same (KerName.make2 mp l)
@@ -500,9 +518,9 @@ module KerPair = struct
let print kp = str (to_string kp)
let debug_to_string = function
- | Same kn -> "(" ^ KerName.to_string kn ^ ")"
+ | Same kn -> "(" ^ KerName.debug_to_string kn ^ ")"
| Dual (knu,knc) ->
- "(" ^ KerName.to_string knu ^ "," ^ KerName.to_string knc ^ ")"
+ "(" ^ KerName.debug_to_string knu ^ "," ^ KerName.debug_to_string knc ^ ")"
let debug_print kp = str (debug_to_string kp)
@@ -524,6 +542,23 @@ module KerPair = struct
let hash x = KerName.hash (canonical x)
end
+ module SyntacticOrd = struct
+ type t = kernel_pair
+ let compare x y = match x, y with
+ | Same knx, Same kny -> KerName.compare knx kny
+ | Dual (knux,kncx), Dual (knuy,kncy) ->
+ let c = KerName.compare knux knuy in
+ if not (Int.equal c 0) then c
+ else KerName.compare kncx kncy
+ | Same _, _ -> -1
+ | Dual _, _ -> 1
+ let equal x y = x == y || compare x y = 0
+ let hash = function
+ | Same kn -> KerName.hash kn
+ | Dual (knu, knc) ->
+ Hashset.Combine.combine (KerName.hash knu) (KerName.hash knc)
+ end
+
(** Default (logical) comparison and hash is on the canonical part *)
let equal = CanOrd.equal
let hash = CanOrd.hash
@@ -535,7 +570,7 @@ module KerPair = struct
let hashcons hkn = function
| Same kn -> Same (hkn kn)
| Dual (knu,knc) -> make (hkn knu) (hkn knc)
- let equal x y = (* physical comparison on subterms *)
+ let eq x y = (* physical comparison on subterms *)
x == y ||
match x,y with
| Same x, Same y -> x == y
@@ -573,11 +608,16 @@ module Mindmap = HMap.Make(MutInd.CanOrd)
module Mindset = Mindmap.Set
module Mindmap_env = HMap.Make(MutInd.UserOrd)
-(** Beware: first inductive has index 0 *)
-(** Beware: first constructor has index 1 *)
+(** Designation of a (particular) inductive type. *)
+type inductive = MutInd.t (* the name of the inductive type *)
+ * int (* the position of this inductive type
+ within the block of mutually-recursive inductive types.
+ BEWARE: indexing starts from 0. *)
-type inductive = MutInd.t * int
-type constructor = inductive * int
+(** Designation of a (particular) constructor of a (particular) inductive type. *)
+type constructor = inductive (* designates the inductive type *)
+ * int (* the index of the constructor
+ BEWARE: indexing starts from 1. *)
let ind_modpath (mind,_) = MutInd.modpath mind
let constr_modpath (ind,_) = ind_modpath ind
@@ -590,6 +630,8 @@ let index_of_constructor (ind, i) = i
let eq_ind (m1, i1) (m2, i2) = Int.equal i1 i2 && MutInd.equal m1 m2
let eq_user_ind (m1, i1) (m2, i2) =
Int.equal i1 i2 && MutInd.UserOrd.equal m1 m2
+let eq_syntactic_ind (m1, i1) (m2, i2) =
+ Int.equal i1 i2 && MutInd.SyntacticOrd.equal m1 m2
let ind_ord (m1, i1) (m2, i2) =
let c = Int.compare i1 i2 in
@@ -597,15 +639,22 @@ let ind_ord (m1, i1) (m2, i2) =
let ind_user_ord (m1, i1) (m2, i2) =
let c = Int.compare i1 i2 in
if Int.equal c 0 then MutInd.UserOrd.compare m1 m2 else c
+let ind_syntactic_ord (m1, i1) (m2, i2) =
+ let c = Int.compare i1 i2 in
+ if Int.equal c 0 then MutInd.SyntacticOrd.compare m1 m2 else c
let ind_hash (m, i) =
Hashset.Combine.combine (MutInd.hash m) (Int.hash i)
let ind_user_hash (m, i) =
Hashset.Combine.combine (MutInd.UserOrd.hash m) (Int.hash i)
+let ind_syntactic_hash (m, i) =
+ Hashset.Combine.combine (MutInd.SyntacticOrd.hash m) (Int.hash i)
let eq_constructor (ind1, j1) (ind2, j2) = Int.equal j1 j2 && eq_ind ind1 ind2
let eq_user_constructor (ind1, j1) (ind2, j2) =
Int.equal j1 j2 && eq_user_ind ind1 ind2
+let eq_syntactic_constructor (ind1, j1) (ind2, j2) =
+ Int.equal j1 j2 && eq_syntactic_ind ind1 ind2
let constructor_ord (ind1, j1) (ind2, j2) =
let c = Int.compare j1 j2 in
@@ -613,11 +662,16 @@ let constructor_ord (ind1, j1) (ind2, j2) =
let constructor_user_ord (ind1, j1) (ind2, j2) =
let c = Int.compare j1 j2 in
if Int.equal c 0 then ind_user_ord ind1 ind2 else c
+let constructor_syntactic_ord (ind1, j1) (ind2, j2) =
+ let c = Int.compare j1 j2 in
+ if Int.equal c 0 then ind_syntactic_ord ind1 ind2 else c
let constructor_hash (ind, i) =
Hashset.Combine.combine (ind_hash ind) (Int.hash i)
let constructor_user_hash (ind, i) =
Hashset.Combine.combine (ind_user_hash ind) (Int.hash i)
+let constructor_syntactic_hash (ind, i) =
+ Hashset.Combine.combine (ind_syntactic_hash ind) (Int.hash i)
module InductiveOrdered = struct
type t = inductive
@@ -662,7 +716,7 @@ module Hind = Hashcons.Make(
type t = inductive
type u = MutInd.t -> MutInd.t
let hashcons hmind (mind, i) = (hmind mind, i)
- let equal (mind1,i1) (mind2,i2) = mind1 == mind2 && Int.equal i1 i2
+ let eq (mind1,i1) (mind2,i2) = mind1 == mind2 && Int.equal i1 i2
let hash = ind_hash
end)
@@ -671,7 +725,7 @@ module Hconstruct = Hashcons.Make(
type t = constructor
type u = inductive -> inductive
let hashcons hind (ind, j) = (hind ind, j)
- let equal (ind1, j1) (ind2, j2) = ind1 == ind2 && Int.equal j1 j2
+ let eq (ind1, j1) (ind2, j2) = ind1 == ind2 && Int.equal j1 j2
let hash = constructor_hash
end)
@@ -805,13 +859,22 @@ struct
let hash (c, b) = (if b then 0 else 1) + Constant.hash c
+ module SyntacticOrd = struct
+ type t = constant * bool
+ let compare (c, b) (c', b') =
+ if b = b' then Constant.SyntacticOrd.compare c c' else -1
+ let equal (c, b as x) (c', b' as x') =
+ x == x' || b = b' && Constant.SyntacticOrd.equal c c'
+ let hash (c, b) = (if b then 0 else 1) + Constant.SyntacticOrd.hash c
+ end
+
module Self_Hashcons =
struct
type _t = t
type t = _t
type u = Constant.t -> Constant.t
let hashcons hc (c,b) = (hc c,b)
- let equal ((c,b) as x) ((c',b') as y) =
+ let eq ((c,b) as x) ((c',b') as y) =
x == y || (c == c' && b == b')
let hash = hash
end
diff --git a/kernel/names.mli b/kernel/names.mli
index 7cc4443752..1dfdd82903 100644
--- a/kernel/names.mli
+++ b/kernel/names.mli
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -10,30 +10,33 @@ open Util
(** {6 Identifiers } *)
+(** Representation and operations on identifiers. *)
module Id :
sig
type t
- (** Type of identifiers *)
+ (** Values of this type represent (Coq) identifiers. *)
val equal : t -> t -> bool
- (** Equality over identifiers *)
+ (** Equality over identifiers. *)
val compare : t -> t -> int
- (** Comparison over identifiers *)
+ (** Comparison over identifiers. *)
val hash : t -> int
- (** Hash over identifiers *)
+ (** Hash over identifiers. *)
val is_valid : string -> bool
- (** Check that a string may be converted to an identifier. *)
+ (** Check that a string may be converted to an identifier.
+ @raise Unicode.Unsupported if the provided string contains unsupported UTF-8 characters. *)
val of_string : string -> t
- (** Converts a string into an identifier. May raise [UserError _] if the
- string is not valid, or echo a warning if it contains invalid identifier
- characters. *)
+ (** Converts a string into an identifier.
+ @raise UserError if the string is not valid, or echo a warning if it contains invalid identifier characters.
+ @raise Unicode.Unsupported if the provided string contains unsupported UTF-8 characters. *)
val of_string_soft : string -> t
- (** Same as {!of_string} except that no warning is ever issued. *)
+ (** Same as {!of_string} except that no warning is ever issued.
+ @raise Unicode.Unsupported if the provided string contains unsupported UTF-8 characters. *)
val to_string : t -> string
(** Converts a identifier into an string. *)
@@ -58,10 +61,18 @@ sig
end
+(** Representation and operations on identifiers that are allowed to be anonymous
+ (i.e. "_" in concrete syntax). *)
module Name :
sig
- type t = Name of Id.t | Anonymous
- (** A name is either undefined, either an identifier. *)
+ type t = Anonymous (** anonymous identifier *)
+ | Name of Id.t (** non-anonymous identifier *)
+
+ val is_anonymous : t -> bool
+ (** Return [true] iff a given name is [Anonymous]. *)
+
+ val is_name : t -> bool
+ (** Return [true] iff a given name is [Name _]. *)
val compare : t -> t -> int
(** Comparison over names. *)
@@ -79,7 +90,7 @@ end
(** {6 Type aliases} *)
-type name = Name.t = Name of Id.t | Anonymous
+type name = Name.t = Anonymous | Name of Id.t
type variable = Id.t
type module_ident = Id.t
@@ -160,6 +171,8 @@ sig
module Set : Set.S with type elt = t
module Map : Map.ExtS with type key = t and module Set := Set
+ val hcons : t -> t
+
end
(** {6 Unique names for bound modules} *)
@@ -217,6 +230,9 @@ sig
val to_string : t -> string
+ val debug_to_string : t -> string
+ (** Same as [to_string], but outputs information related to debug. *)
+
val initial : t
(** Name of the toplevel structure ([= MPfile initial_dir]) *)
@@ -244,6 +260,10 @@ sig
(** Display *)
val to_string : t -> string
+
+ val debug_to_string : t -> string
+ (** Same as [to_string], but outputs information related to debug. *)
+
val print : t -> Pp.std_ppcmds
(** Comparisons *)
@@ -305,6 +325,12 @@ sig
val hash : t -> int
end
+ module SyntacticOrd : sig
+ val compare : t -> t -> int
+ val equal : t -> t -> bool
+ val hash : t -> int
+ end
+
val equal : t -> t -> bool
(** Default comparison, alias for [CanOrd.equal] *)
@@ -379,6 +405,12 @@ sig
val hash : t -> int
end
+ module SyntacticOrd : sig
+ val compare : t -> t -> int
+ val equal : t -> t -> bool
+ val hash : t -> int
+ end
+
val equal : t -> t -> bool
(** Default comparison, alias for [CanOrd.equal] *)
@@ -395,18 +427,23 @@ end
module Mindset : CSig.SetS with type elt = MutInd.t
module Mindmap : Map.ExtS with type key = MutInd.t and module Set := Mindset
-module Mindmap_env : Map.S with type key = MutInd.t
+module Mindmap_env : CSig.MapS with type key = MutInd.t
-(** Beware: first inductive has index 0 *)
-type inductive = MutInd.t * int
+(** Designation of a (particular) inductive type. *)
+type inductive = MutInd.t (* the name of the inductive type *)
+ * int (* the position of this inductive type
+ within the block of mutually-recursive inductive types.
+ BEWARE: indexing starts from 0. *)
-(** Beware: first constructor has index 1 *)
-type constructor = inductive * int
+(** Designation of a (particular) constructor of a (particular) inductive type. *)
+type constructor = inductive (* designates the inductive type *)
+ * int (* the index of the constructor
+ BEWARE: indexing starts from 1. *)
-module Indmap : Map.S with type key = inductive
-module Constrmap : Map.S with type key = constructor
-module Indmap_env : Map.S with type key = inductive
-module Constrmap_env : Map.S with type key = constructor
+module Indmap : CSig.MapS with type key = inductive
+module Constrmap : CSig.MapS with type key = constructor
+module Indmap_env : CSig.MapS with type key = inductive
+module Constrmap_env : CSig.MapS with type key = constructor
val ind_modpath : inductive -> ModPath.t
val constr_modpath : constructor -> ModPath.t
@@ -417,16 +454,22 @@ val inductive_of_constructor : constructor -> inductive
val index_of_constructor : constructor -> int
val eq_ind : inductive -> inductive -> bool
val eq_user_ind : inductive -> inductive -> bool
+val eq_syntactic_ind : inductive -> inductive -> bool
val ind_ord : inductive -> inductive -> int
val ind_hash : inductive -> int
val ind_user_ord : inductive -> inductive -> int
val ind_user_hash : inductive -> int
+val ind_syntactic_ord : inductive -> inductive -> int
+val ind_syntactic_hash : inductive -> int
val eq_constructor : constructor -> constructor -> bool
val eq_user_constructor : constructor -> constructor -> bool
+val eq_syntactic_constructor : constructor -> constructor -> bool
val constructor_ord : constructor -> constructor -> int
-val constructor_user_ord : constructor -> constructor -> int
val constructor_hash : constructor -> int
+val constructor_user_ord : constructor -> constructor -> int
val constructor_user_hash : constructor -> int
+val constructor_syntactic_ord : constructor -> constructor -> int
+val constructor_syntactic_hash : constructor -> int
(** Better to have it here that in Closure, since required in grammar.cma *)
type evaluable_global_reference =
@@ -640,6 +683,12 @@ module Projection : sig
val make : constant -> bool -> t
+ module SyntacticOrd : sig
+ val compare : t -> t -> int
+ val equal : t -> t -> bool
+ val hash : t -> int
+ end
+
val constant : t -> constant
val unfolded : t -> bool
val unfold : t -> t
diff --git a/kernel/nativecode.ml b/kernel/nativecode.ml
index 687b740f67..dabe905dee 100644
--- a/kernel/nativecode.ml
+++ b/kernel/nativecode.ml
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -8,7 +8,6 @@
open Errors
open Names
open Term
-open Context
open Declarations
open Util
open Nativevalues
@@ -481,7 +480,7 @@ and eq_mllam_branches gn1 gn2 n env1 env2 br1 br2 =
in
Array.equal eq_branch br1 br2
-(* hash_mllambda gn n env t computes the hash for t ignoring occurences of gn *)
+(* hash_mllambda gn n env t computes the hash for t ignoring occurrences of gn *)
let rec hash_mllambda gn n env t =
match t with
| MLlocal ln -> combinesmall 1 (LNmap.find ln env)
@@ -979,7 +978,7 @@ let compile_prim decl cond paux =
let args = Array.map opt_prim_aux args in
app_prim (Coq_primitive(op,None)) args
(*
- TODO: check if this inling was useful
+ TODO: check if this inlining was useful
begin match op with
| Int31lt ->
if Sys.word_size = 64 then
@@ -1826,31 +1825,32 @@ and apply_fv env sigma univ (fv_named,fv_rel) auxdefs ml =
in
let auxdefs = List.fold_right get_rel_val fv_rel auxdefs in
let auxdefs = List.fold_right get_named_val fv_named auxdefs in
- let lvl = rel_context_length env.env_rel_context in
+ let lvl = Context.Rel.length env.env_rel_context in
let fv_rel = List.map (fun (n,_) -> MLglobal (Grel (lvl-n))) fv_rel in
let fv_named = List.map (fun (id,_) -> MLglobal (Gnamed id)) fv_named in
let aux_name = fresh_lname Anonymous in
auxdefs, MLlet(aux_name, ml, mkMLapp (MLlocal aux_name) (Array.of_list (fv_rel@fv_named)))
and compile_rel env sigma univ auxdefs n =
- let (_,body,_) = lookup_rel n env.env_rel_context in
- let n = rel_context_length env.env_rel_context - n in
- match body with
- | Some t ->
+ let open Context.Rel in
+ let n = length env.env_rel_context - n in
+ let open Declaration in
+ match lookup n env.env_rel_context with
+ | LocalDef (_,t,_) ->
let code = lambda_of_constr env sigma t in
let auxdefs,code = compile_with_fv env sigma univ auxdefs None code in
Glet(Grel n, code)::auxdefs
- | None ->
+ | LocalAssum _ ->
Glet(Grel n, MLprimitive (Mk_rel n))::auxdefs
and compile_named env sigma univ auxdefs id =
- let (_,body,_) = lookup_named id env.env_named_context in
- match body with
- | Some t ->
+ let open Context.Named.Declaration in
+ match Context.Named.lookup id env.env_named_context with
+ | LocalDef (_,t,_) ->
let code = lambda_of_constr env sigma t in
let auxdefs,code = compile_with_fv env sigma univ auxdefs None code in
Glet(Gnamed id, code)::auxdefs
- | None ->
+ | LocalAssum _ ->
Glet(Gnamed id, MLprimitive (Mk_var id))::auxdefs
let compile_constant env sigma prefix ~interactive con cb =
diff --git a/kernel/nativecode.mli b/kernel/nativecode.mli
index 5d4c9e1e2d..77d9c33f8d 100644
--- a/kernel/nativecode.mli
+++ b/kernel/nativecode.mli
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
diff --git a/kernel/nativeconv.ml b/kernel/nativeconv.ml
index d0aa96fd15..7ac5b8d7b2 100644
--- a/kernel/nativeconv.ml
+++ b/kernel/nativeconv.ml
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -16,21 +16,21 @@ open Nativecode
(** This module implements the conversion test by compiling to OCaml code *)
-let rec conv_val env pb lvl cu v1 v2 =
- if v1 == v2 then ()
+let rec conv_val env pb lvl v1 v2 cu =
+ if v1 == v2 then cu
else
match kind_of_value v1, kind_of_value v2 with
| Vfun f1, Vfun f2 ->
let v = mk_rel_accu lvl in
- conv_val env CONV (lvl+1) cu (f1 v) (f2 v)
+ conv_val env CONV (lvl+1) (f1 v) (f2 v) cu
| Vfun f1, _ ->
- conv_val env CONV lvl cu v1 (fun x -> v2 x)
+ conv_val env CONV lvl v1 (fun x -> v2 x) cu
| _, Vfun f2 ->
- conv_val env CONV lvl cu (fun x -> v1 x) v2
+ conv_val env CONV lvl (fun x -> v1 x) v2 cu
| Vaccu k1, Vaccu k2 ->
- conv_accu env pb lvl cu k1 k2
+ conv_accu env pb lvl k1 k2 cu
| Vconst i1, Vconst i2 ->
- if not (Int.equal i1 i2) then raise NotConvertible
+ if Int.equal i1 i2 then cu else raise NotConvertible
| Vblock b1, Vblock b2 ->
let n1 = block_size b1 in
let n2 = block_size b2 in
@@ -38,76 +38,78 @@ let rec conv_val env pb lvl cu v1 v2 =
raise NotConvertible;
let rec aux lvl max b1 b2 i cu =
if Int.equal i max then
- conv_val env CONV lvl cu (block_field b1 i) (block_field b2 i)
+ conv_val env CONV lvl (block_field b1 i) (block_field b2 i) cu
else
- (conv_val env CONV lvl cu (block_field b1 i) (block_field b2 i);
- aux lvl max b1 b2 (i+1) cu)
+ let cu = conv_val env CONV lvl (block_field b1 i) (block_field b2 i) cu in
+ aux lvl max b1 b2 (i+1) cu
in
aux lvl (n1-1) b1 b2 0 cu
| Vaccu _, _ | Vconst _, _ | Vblock _, _ -> raise NotConvertible
-and conv_accu env pb lvl cu k1 k2 =
+and conv_accu env pb lvl k1 k2 cu =
let n1 = accu_nargs k1 in
let n2 = accu_nargs k2 in
if not (Int.equal n1 n2) then raise NotConvertible;
if Int.equal n1 0 then
conv_atom env pb lvl (atom_of_accu k1) (atom_of_accu k2) cu
else
- (conv_atom env pb lvl (atom_of_accu k1) (atom_of_accu k2) cu;
- List.iter2 (conv_val env CONV lvl cu) (args_of_accu k1) (args_of_accu k2))
+ let cu = conv_atom env pb lvl (atom_of_accu k1) (atom_of_accu k2) cu in
+ List.fold_right2 (conv_val env CONV lvl) (args_of_accu k1) (args_of_accu k2) cu
and conv_atom env pb lvl a1 a2 cu =
- if a1 == a2 then ()
+ if a1 == a2 then cu
else
match a1, a2 with
| Ameta _, _ | _, Ameta _ | Aevar _, _ | _, Aevar _ -> assert false
| Arel i1, Arel i2 ->
- if not (Int.equal i1 i2) then raise NotConvertible
- | Aind ind1, Aind ind2 ->
- if not (eq_puniverses eq_ind ind1 ind2) then raise NotConvertible
- | Aconstant c1, Aconstant c2 ->
- if not (eq_puniverses eq_constant c1 c2) then raise NotConvertible
+ if Int.equal i1 i2 then cu else raise NotConvertible
+ | Aind (ind1,u1), Aind (ind2,u2) ->
+ if eq_ind ind1 ind2 then convert_instances ~flex:false u1 u2 cu
+ else raise NotConvertible
+ | Aconstant (c1,u1), Aconstant (c2,u2) ->
+ if Constant.equal c1 c2 then convert_instances ~flex:true u1 u2 cu
+ else raise NotConvertible
| Asort s1, Asort s2 ->
- check_sort_cmp_universes env pb s1 s2 cu
+ sort_cmp_universes env pb s1 s2 cu
| Avar id1, Avar id2 ->
- if not (Id.equal id1 id2) then raise NotConvertible
+ if Id.equal id1 id2 then cu else raise NotConvertible
| Acase(a1,ac1,p1,bs1), Acase(a2,ac2,p2,bs2) ->
if not (eq_ind a1.asw_ind a2.asw_ind) then raise NotConvertible;
- conv_accu env CONV lvl cu ac1 ac2;
+ let cu = conv_accu env CONV lvl ac1 ac2 cu in
let tbl = a1.asw_reloc in
let len = Array.length tbl in
- if Int.equal len 0 then conv_val env CONV lvl cu p1 p2
+ if Int.equal len 0 then conv_val env CONV lvl p1 p2 cu
else begin
- conv_val env CONV lvl cu p1 p2;
- let max = len - 1 in
- let rec aux i =
- let tag,arity = tbl.(i) in
- let ci =
- if Int.equal arity 0 then mk_const tag
- else mk_block tag (mk_rels_accu lvl arity) in
- let bi1 = bs1 ci and bi2 = bs2 ci in
- if Int.equal i max then conv_val env CONV (lvl + arity) cu bi1 bi2
- else (conv_val env CONV (lvl + arity) cu bi1 bi2; aux (i+1)) in
- aux 0
+ let cu = conv_val env CONV lvl p1 p2 cu in
+ let max = len - 1 in
+ let rec aux i cu =
+ let tag,arity = tbl.(i) in
+ let ci =
+ if Int.equal arity 0 then mk_const tag
+ else mk_block tag (mk_rels_accu lvl arity) in
+ let bi1 = bs1 ci and bi2 = bs2 ci in
+ if Int.equal i max then conv_val env CONV (lvl + arity) bi1 bi2 cu
+ else aux (i+1) (conv_val env CONV (lvl + arity) bi1 bi2 cu) in
+ aux 0 cu
end
| Afix(t1,f1,rp1,s1), Afix(t2,f2,rp2,s2) ->
if not (Int.equal s1 s2) || not (Array.equal Int.equal rp1 rp2) then raise NotConvertible;
- if f1 == f2 then ()
+ if f1 == f2 then cu
else conv_fix env lvl t1 f1 t2 f2 cu
| (Acofix(t1,f1,s1,_) | Acofixe(t1,f1,s1,_)),
(Acofix(t2,f2,s2,_) | Acofixe(t2,f2,s2,_)) ->
if not (Int.equal s1 s2) then raise NotConvertible;
- if f1 == f2 then ()
+ if f1 == f2 then cu
else
if not (Int.equal (Array.length f1) (Array.length f2)) then raise NotConvertible
else conv_fix env lvl t1 f1 t2 f2 cu
| Aprod(_,d1,c1), Aprod(_,d2,c2) ->
- conv_val env CONV lvl cu d1 d2;
- let v = mk_rel_accu lvl in
- conv_val env pb (lvl + 1) cu (d1 v) (d2 v)
+ let cu = conv_val env CONV lvl d1 d2 cu in
+ let v = mk_rel_accu lvl in
+ conv_val env pb (lvl + 1) (d1 v) (d2 v) cu
| Aproj(p1,ac1), Aproj(p2,ac2) ->
if not (Constant.equal p1 p2) then raise NotConvertible
- else conv_accu env CONV lvl cu ac1 ac2
+ else conv_accu env CONV lvl ac1 ac2 cu
| Arel _, _ | Aind _, _ | Aconstant _, _ | Asort _, _ | Avar _, _
| Acase _, _ | Afix _, _ | Acofix _, _ | Acofixe _, _ | Aprod _, _
| Aproj _, _ -> raise NotConvertible
@@ -118,21 +120,15 @@ and conv_fix env lvl t1 f1 t2 f2 cu =
let max = len - 1 in
let fargs = mk_rels_accu lvl len in
let flvl = lvl + len in
- let rec aux i =
- conv_val env CONV lvl cu t1.(i) t2.(i);
+ let rec aux i cu =
+ let cu = conv_val env CONV lvl t1.(i) t2.(i) cu in
let fi1 = napply f1.(i) fargs in
let fi2 = napply f2.(i) fargs in
- if Int.equal i max then conv_val env CONV flvl cu fi1 fi2
- else (conv_val env CONV flvl cu fi1 fi2; aux (i+1)) in
- aux 0
+ if Int.equal i max then conv_val env CONV flvl fi1 fi2 cu
+ else aux (i+1) (conv_val env CONV flvl fi1 fi2 cu) in
+ aux 0 cu
-let native_conv pb sigma env t1 t2 =
- if Coq_config.no_native_compiler then begin
- let msg = "Native compiler is disabled, falling back to VM conversion test." in
- Pp.msg_warning (Pp.str msg);
- vm_conv pb env t1 t2
- end
- else
+let native_conv_gen pb sigma env univs t1 t2 =
let penv = Environ.pre_env env in
let ml_filename, prefix = get_ml_filename () in
let code, upds = mk_conv_code penv sigma prefix t1 t2 in
@@ -146,8 +142,25 @@ let native_conv pb sigma env t1 t2 =
let time_info = Format.sprintf "Evaluation done in %.5f@." (t1 -. t0) in
if !Flags.debug then Pp.msg_debug (Pp.str time_info);
(* TODO change 0 when we can have deBruijn *)
- conv_val env pb 0 (Environ.universes env) !rt1 !rt2
+ fst (conv_val env pb 0 !rt1 !rt2 univs)
end
| _ -> anomaly (Pp.str "Compilation failure")
-let _ = set_nat_conv native_conv
+(* Wrapper for [native_conv] above *)
+let native_conv cv_pb sigma env t1 t2 =
+ if Coq_config.no_native_compiler then begin
+ let msg = "Native compiler is disabled, falling back to VM conversion test." in
+ Pp.msg_warning (Pp.str msg);
+ vm_conv cv_pb env t1 t2
+ end
+ else
+ let univs = Environ.universes env in
+ let b =
+ if cv_pb = CUMUL then Constr.leq_constr_univs univs t1 t2
+ else Constr.eq_constr_univs univs t1 t2
+ in
+ if not b then
+ let univs = (univs, checked_universes) in
+ let t1 = Term.it_mkLambda_or_LetIn t1 (Environ.rel_context env) in
+ let t2 = Term.it_mkLambda_or_LetIn t2 (Environ.rel_context env) in
+ let _ = native_conv_gen cv_pb sigma env univs t1 t2 in ()
diff --git a/kernel/nativeconv.mli b/kernel/nativeconv.mli
index 318a7d830b..63b1eb0586 100644
--- a/kernel/nativeconv.mli
+++ b/kernel/nativeconv.mli
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -11,4 +11,8 @@ open Nativelambda
(** This module implements the conversion test by compiling to OCaml code *)
-val native_conv : conv_pb -> evars -> types conversion_function
+val native_conv : conv_pb -> evars -> types kernel_conversion_function
+
+(** A conversion function parametrized by a universe comparator. Used outside of
+ the kernel. *)
+val native_conv_gen : conv_pb -> evars -> (types, 'a) generic_conversion_function
diff --git a/kernel/nativeinstr.mli b/kernel/nativeinstr.mli
index b7d3dadcd8..41e79a5355 100644
--- a/kernel/nativeinstr.mli
+++ b/kernel/nativeinstr.mli
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
diff --git a/kernel/nativelambda.ml b/kernel/nativelambda.ml
index cb08b5058f..91b40be7e9 100644
--- a/kernel/nativelambda.ml
+++ b/kernel/nativelambda.ml
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -277,7 +277,7 @@ and reduce_lapp substf lids body substa largs =
| [], _::_ -> simplify_app substf body substa (Array.of_list largs)
-(* [occurence kind k lam]:
+(* [occurrence kind k lam]:
If [kind] is [true] return [true] if the variable [k] does not appear in
[lam], return [false] if the variable appear one time and not
under a lambda, a fixpoint, a cofixpoint; else raise Not_found.
@@ -285,7 +285,7 @@ and reduce_lapp substf lids body substa largs =
else raise [Not_found]
*)
-let rec occurence k kind lam =
+let rec occurrence k kind lam =
match lam with
| Lrel (_,n) ->
if Int.equal n k then
@@ -294,35 +294,35 @@ let rec occurence k kind lam =
| Lvar _ | Lconst _ | Lproj _ | Luint _ | Lval _ | Lsort _ | Lind _
| Lconstruct _ | Llazy | Lforce | Lmeta _ | Levar _ -> kind
| Lprod(dom, codom) ->
- occurence k (occurence k kind dom) codom
+ occurrence k (occurrence k kind dom) codom
| Llam(ids,body) ->
- let _ = occurence (k+Array.length ids) false body in kind
+ let _ = occurrence (k+Array.length ids) false body in kind
| Llet(_,def,body) ->
- occurence (k+1) (occurence k kind def) body
+ occurrence (k+1) (occurrence k kind def) body
| Lapp(f, args) ->
- occurence_args k (occurence k kind f) args
+ occurrence_args k (occurrence k kind f) args
| Lprim(_,_,_,args) | Lmakeblock(_,_,_,args) ->
- occurence_args k kind args
+ occurrence_args k kind args
| Lcase(_,t,a,br) ->
- let kind = occurence k (occurence k kind t) a in
+ let kind = occurrence k (occurrence k kind t) a in
let r = ref kind in
Array.iter (fun (_,ids,c) ->
- r := occurence (k+Array.length ids) kind c && !r) br;
+ r := occurrence (k+Array.length ids) kind c && !r) br;
!r
| Lif (t, bt, bf) ->
- let kind = occurence k kind t in
- kind && occurence k kind bt && occurence k kind bf
+ let kind = occurrence k kind t in
+ kind && occurrence k kind bt && occurrence k kind bf
| Lfix(_,(ids,ltypes,lbodies))
| Lcofix(_,(ids,ltypes,lbodies)) ->
- let kind = occurence_args k kind ltypes in
- let _ = occurence_args (k+Array.length ids) false lbodies in
+ let kind = occurrence_args k kind ltypes in
+ let _ = occurrence_args (k+Array.length ids) false lbodies in
kind
-and occurence_args k kind args =
- Array.fold_left (occurence k) kind args
+and occurrence_args k kind args =
+ Array.fold_left (occurrence k) kind args
let occur_once lam =
- try let _ = occurence 1 true lam in true
+ try let _ = occurrence 1 true lam in true
with Not_found -> false
(* [remove_let lam] remove let expression in [lam] if the variable is *)
@@ -379,7 +379,7 @@ let rec get_alias env (kn, u as p) =
| None -> p
| Some tps ->
match Cemitcodes.force tps with
- | Cemitcodes.BCalias kn' -> get_alias env kn'
+ | Cemitcodes.BCalias kn' -> get_alias env (kn', u)
| _ -> p
(*i Global environment *)
@@ -485,7 +485,7 @@ module Renv =
let pop env = Vect.pop env.name_rel
let popn env n =
- for i = 1 to n do pop env done
+ for _i = 1 to n do pop env done
let get env n =
Lrel (Vect.get_last env.name_rel (n-1), n)
@@ -727,7 +727,8 @@ let optimize lam =
let lambda_of_constr env sigma c =
set_global_env env;
let env = Renv.make () in
- let ids = List.rev_map (fun (id, _, _) -> id) !global_env.env_rel_context in
+ let open Context.Rel.Declaration in
+ let ids = List.rev_map get_name !global_env.env_rel_context in
Renv.push_rels env (Array.of_list ids);
let lam = lambda_of_constr env sigma c in
(* if Flags.vm_draw_opt () then begin
diff --git a/kernel/nativelambda.mli b/kernel/nativelambda.mli
index 3b6fafbbc9..c33574408b 100644
--- a/kernel/nativelambda.mli
+++ b/kernel/nativelambda.mli
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
diff --git a/kernel/nativelib.ml b/kernel/nativelib.ml
index 331598d85a..4296b73abe 100644
--- a/kernel/nativelib.ml
+++ b/kernel/nativelib.ml
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -74,7 +74,17 @@ let call_compiler ml_filename =
::include_dirs
@ ["-impl"; ml_filename] in
if !Flags.debug then Pp.msg_debug (Pp.str (ocamlfind () ^ " " ^ (String.concat " " args)));
- try CUnix.sys_command (ocamlfind ()) args = Unix.WEXITED 0, link_filename
+ try
+ let res = CUnix.sys_command (ocamlfind ()) args in
+ let res = match res with
+ | Unix.WEXITED 0 -> true
+ | Unix.WEXITED n ->
+ Pp.(msg_warning (str "command exited with status " ++ int n)); false
+ | Unix.WSIGNALED n ->
+ Pp.(msg_warning (str "command killed by signal " ++ int n)); false
+ | Unix.WSTOPPED n ->
+ Pp.(msg_warning (str "command stopped by signal " ++ int n)); false in
+ res, link_filename
with Unix.Unix_error (e,_,_) ->
Pp.(msg_warning (str (Unix.error_message e)));
false, link_filename
diff --git a/kernel/nativelib.mli b/kernel/nativelib.mli
index 0941dc56ce..12ad3cf2f7 100644
--- a/kernel/nativelib.mli
+++ b/kernel/nativelib.mli
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
diff --git a/kernel/nativelibrary.ml b/kernel/nativelibrary.ml
index 443cd8c2a0..9d159be64c 100644
--- a/kernel/nativelibrary.ml
+++ b/kernel/nativelibrary.ml
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
diff --git a/kernel/nativelibrary.mli b/kernel/nativelibrary.mli
index 29368d1408..7d01640b29 100644
--- a/kernel/nativelibrary.mli
+++ b/kernel/nativelibrary.mli
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
diff --git a/kernel/nativevalues.ml b/kernel/nativevalues.ml
index e4a7799933..d6fdfefa02 100644
--- a/kernel/nativevalues.ml
+++ b/kernel/nativevalues.ml
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -10,7 +10,7 @@ open Names
open Errors
open Util
-(** This modules defines the representation of values internally used by
+(** This module defines the representation of values internally used by
the native compiler *)
type t = t -> t
@@ -78,8 +78,6 @@ let accumulate_code (k:accumulator) (x:t) =
let rec accumulate (x:t) =
accumulate_code (Obj.magic accumulate) x
-let raccumulate = ref accumulate
-
let mk_accu_gen rcode (a:atom) =
(* Format.eprintf "size rcode =%i\n" (Obj.size (Obj.magic rcode)); *)
let r = Obj.new_block 0 3 in
@@ -160,31 +158,6 @@ let is_accu x =
let o = Obj.repr x in
Obj.is_block o && Int.equal (Obj.tag o) accumulate_tag
-(*let accumulate_fix_code (k:accumulator) (a:t) =
- match atom_of_accu k with
- | Afix(frec,_,rec_pos,_,_) ->
- let nargs = accu_nargs k in
- if nargs <> rec_pos || is_accu a then
- accumulate_code k a
- else
- let r = ref frec in
- for i = 0 to nargs - 1 do
- r := !r (arg_of_accu k i)
- done;
- !r a
- | _ -> assert false
-
-
-let rec accumulate_fix (x:t) =
- accumulate_fix_code (Obj.magic accumulate_fix) x
-
-let raccumulate_fix = ref accumulate_fix *)
-
-let is_atom_fix (a:atom) =
- match a with
- | Afix _ -> true
- | _ -> false
-
let mk_fix_accu rec_pos pos types bodies =
mk_accu_gen accumulate (Afix(types,bodies,rec_pos, pos))
diff --git a/kernel/nativevalues.mli b/kernel/nativevalues.mli
index 79e35d4a04..f4396659ec 100644
--- a/kernel/nativevalues.mli
+++ b/kernel/nativevalues.mli
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
diff --git a/kernel/opaqueproof.ml b/kernel/opaqueproof.ml
index 9f4361f401..0c8772d8d2 100644
--- a/kernel/opaqueproof.ml
+++ b/kernel/opaqueproof.ml
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -16,7 +16,7 @@ type work_list = (Instance.t * Id.t array) Cmap.t *
type cooking_info = {
modlist : work_list;
- abstract : Context.named_context * Univ.universe_level_subst * Univ.UContext.t }
+ abstract : Context.Named.t * Univ.universe_level_subst * Univ.UContext.t }
type proofterm = (constr * Univ.universe_context_set) Future.computation
type opaque =
| Indirect of substitution list * DirPath.t * int (* subst, lib, index *)
@@ -43,7 +43,10 @@ let set_indirect_univ_accessor f = (get_univ := f)
let create cu = Direct ([],cu)
let turn_indirect dp o (prfs,odp) = match o with
- | Indirect _ -> Errors.anomaly (Pp.str "Already an indirect opaque")
+ | Indirect (_,_,i) ->
+ if not (Int.Map.mem i prfs)
+ then Errors.anomaly (Pp.str "Indirect in a different table")
+ else Errors.anomaly (Pp.str "Already an indirect opaque")
| Direct (d,cu) ->
let cu = Future.chain ~pure:true cu (fun (c, u) -> hcons_constr c, u) in
let id = Int.Map.cardinal prfs in
diff --git a/kernel/opaqueproof.mli b/kernel/opaqueproof.mli
index 0609c8517e..5139cf0512 100644
--- a/kernel/opaqueproof.mli
+++ b/kernel/opaqueproof.mli
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -11,9 +11,9 @@ open Term
open Mod_subst
(** This module implements the handling of opaque proof terms.
- Opauqe proof terms are special since:
+ Opaque proof terms are special since:
- they can be lazily computed and substituted
- - they are stoked in an optionally loaded segment of .vo files
+ - they are stored in an optionally loaded segment of .vo files
An [opaque] proof terms holds the real data until fully discharged.
In this case it is called [direct].
When it is [turn_indirect] the data is relocated to an opaque table
@@ -48,7 +48,7 @@ type work_list = (Univ.Instance.t * Id.t array) Cmap.t *
type cooking_info = {
modlist : work_list;
- abstract : Context.named_context * Univ.universe_level_subst * Univ.UContext.t }
+ abstract : Context.Named.t * Univ.universe_level_subst * Univ.UContext.t }
(* The type has two caveats:
1) cook_constr is defined after
diff --git a/kernel/pre_env.ml b/kernel/pre_env.ml
index 5f3f559a2c..0e56e76aa5 100644
--- a/kernel/pre_env.ml
+++ b/kernel/pre_env.ml
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -15,17 +15,17 @@
open Util
open Names
-open Context
open Univ
open Term
open Declarations
+open Context.Named.Declaration
(* The type of environments. *)
(* The key attached to each constant is used by the VM to retrieve previous *)
(* evaluations of the constant. It is essentially an index in the symbols table *)
(* used by the VM. *)
-type key = int Ephemeron.key option ref
+type key = int CEphemeron.key option ref
(** Linking information for the native compiler. *)
@@ -45,30 +45,30 @@ type globals = {
env_modtypes : module_type_body MPmap.t}
type stratification = {
- env_universes : universes;
+ env_universes : UGraph.t;
env_engagement : engagement
}
type val_kind =
- | VKvalue of (values * Id.Set.t) Ephemeron.key
+ | VKvalue of (values * Id.Set.t) CEphemeron.key
| VKnone
type lazy_val = val_kind ref
let force_lazy_val vk = match !vk with
| VKnone -> None
-| VKvalue v -> try Some (Ephemeron.get v) with Ephemeron.InvalidKey -> None
+| VKvalue v -> try Some (CEphemeron.get v) with CEphemeron.InvalidKey -> None
let dummy_lazy_val () = ref VKnone
-let build_lazy_val vk key = vk := VKvalue (Ephemeron.create key)
+let build_lazy_val vk key = vk := VKvalue (CEphemeron.create key)
type named_vals = (Id.t * lazy_val) list
type env = {
env_globals : globals;
- env_named_context : named_context;
+ env_named_context : Context.Named.t;
env_named_vals : named_vals;
- env_rel_context : rel_context;
+ env_rel_context : Context.Rel.t;
env_rel_val : lazy_val list;
env_nb_rel : int;
env_stratification : stratification;
@@ -77,7 +77,7 @@ type env = {
indirect_pterms : Opaqueproof.opaquetab;
}
-type named_context_val = named_context * named_vals
+type named_context_val = Context.Named.t * named_vals
let empty_named_context_val = [],[]
@@ -87,13 +87,13 @@ let empty_env = {
env_inductives = Mindmap_env.empty;
env_modules = MPmap.empty;
env_modtypes = MPmap.empty};
- env_named_context = empty_named_context;
+ env_named_context = Context.Named.empty;
env_named_vals = [];
- env_rel_context = empty_rel_context;
+ env_rel_context = Context.Rel.empty;
env_rel_val = [];
env_nb_rel = 0;
env_stratification = {
- env_universes = initial_universes;
+ env_universes = UGraph.initial_universes;
env_engagement = (PredicativeSet,StratifiedType) };
env_conv_oracle = Conv_oracle.empty;
retroknowledge = Retroknowledge.initial_retroknowledge;
@@ -107,7 +107,7 @@ let nb_rel env = env.env_nb_rel
let push_rel d env =
let rval = ref VKnone in
{ env with
- env_rel_context = add_rel_decl d env.env_rel_context;
+ env_rel_context = Context.Rel.add d env.env_rel_context;
env_rel_val = rval :: env.env_rel_val;
env_nb_rel = env.env_nb_rel + 1 }
@@ -125,18 +125,16 @@ let env_of_rel n env =
(* Named context *)
let push_named_context_val d (ctxt,vals) =
- let id,_,_ = d in
let rval = ref VKnone in
- add_named_decl d ctxt, (id,rval)::vals
+ Context.Named.add d ctxt, (get_id d,rval)::vals
let push_named d env =
(* if not (env.env_rel_context = []) then raise (ASSERT env.env_rel_context);
assert (env.env_rel_context = []); *)
- let id,body,_ = d in
let rval = ref VKnone in
{ env_globals = env.env_globals;
- env_named_context = Context.add_named_decl d env.env_named_context;
- env_named_vals = (id, rval) :: env.env_named_vals;
+ env_named_context = Context.Named.add d env.env_named_context;
+ env_named_vals = (get_id d, rval) :: env.env_named_vals;
env_rel_context = env.env_rel_context;
env_rel_val = env.env_rel_val;
env_nb_rel = env.env_nb_rel;
diff --git a/kernel/pre_env.mli b/kernel/pre_env.mli
index 0ce0bed235..353c461120 100644
--- a/kernel/pre_env.mli
+++ b/kernel/pre_env.mli
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -8,9 +8,7 @@
open Names
open Term
-open Context
open Declarations
-open Univ
(** The type of environments. *)
@@ -19,7 +17,7 @@ type link_info =
| LinkedInteractive of string
| NotLinked
-type key = int Ephemeron.key option ref
+type key = int CEphemeron.key option ref
type constant_key = constant_body * (link_info ref * key)
@@ -32,7 +30,7 @@ type globals = {
env_modtypes : module_type_body MPmap.t}
type stratification = {
- env_universes : universes;
+ env_universes : UGraph.t;
env_engagement : engagement
}
@@ -46,9 +44,9 @@ type named_vals = (Id.t * lazy_val) list
type env = {
env_globals : globals;
- env_named_context : named_context;
+ env_named_context : Context.Named.t;
env_named_vals : named_vals;
- env_rel_context : rel_context;
+ env_rel_context : Context.Rel.t;
env_rel_val : lazy_val list;
env_nb_rel : int;
env_stratification : stratification;
@@ -57,7 +55,7 @@ type env = {
indirect_pterms : Opaqueproof.opaquetab;
}
-type named_context_val = named_context * named_vals
+type named_context_val = Context.Named.t * named_vals
val empty_named_context_val : named_context_val
@@ -66,15 +64,15 @@ val empty_env : env
(** Rel context *)
val nb_rel : env -> int
-val push_rel : rel_declaration -> env -> env
+val push_rel : Context.Rel.Declaration.t -> env -> env
val lookup_rel_val : int -> env -> lazy_val
val env_of_rel : int -> env -> env
(** Named context *)
val push_named_context_val :
- named_declaration -> named_context_val -> named_context_val
-val push_named : named_declaration -> env -> env
+ Context.Named.Declaration.t -> named_context_val -> named_context_val
+val push_named : Context.Named.Declaration.t -> env -> env
val lookup_named_val : Id.t -> env -> lazy_val
val env_of_named : Id.t -> env -> env
diff --git a/kernel/primitives.ml b/kernel/primitives.ml
index 649eb125e4..27732c00cb 100644
--- a/kernel/primitives.ml
+++ b/kernel/primitives.ml
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
diff --git a/kernel/primitives.mli b/kernel/primitives.mli
index 9f99264afa..86e86a5e5a 100644
--- a/kernel/primitives.mli
+++ b/kernel/primitives.mli
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
diff --git a/kernel/reduction.ml b/kernel/reduction.ml
index 3253cddf7c..cfc286135d 100644
--- a/kernel/reduction.ml
+++ b/kernel/reduction.ml
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -20,13 +20,11 @@ open Util
open Names
open Term
open Vars
-open Context
open Univ
open Environ
open Closure
open Esubst
-
-let left2right = ref false
+open Context.Rel.Declaration
let rec is_empty_stack = function
[] -> true
@@ -56,8 +54,7 @@ let compare_stack_shape stk1 stk2 =
| (_, Zapp l2::s2) -> compare_rec (bal-Array.length l2) stk1 s2
| (Zproj (n1,m1,p1)::s1, Zproj (n2,m2,p2)::s2) ->
Int.equal bal 0 && compare_rec 0 s1 s2
- | ((Zcase(c1,_,_)|ZcaseT(c1,_,_,_))::s1,
- (Zcase(c2,_,_)|ZcaseT(c2,_,_,_))::s2) ->
+ | (ZcaseT(c1,_,_,_)::s1, ZcaseT(c2,_,_,_)::s2) ->
Int.equal bal 0 (* && c1.ci_ind = c2.ci_ind *) && compare_rec 0 s1 s2
| (Zfix(_,a1)::s1, Zfix(_,a2)::s2) ->
Int.equal bal 0 && compare_rec 0 a1 a2 && compare_rec 0 s1 s2
@@ -91,9 +88,8 @@ let pure_stack lfts stk =
let (lfx,pa) = pure_rec l a in
(l, Zlfix((lfx,fx),pa)::pstk)
| (ZcaseT(ci,p,br,e),(l,pstk)) ->
- (l,Zlcase(ci,l,mk_clos e p,Array.map (mk_clos e) br)::pstk)
- | (Zcase(ci,p,br),(l,pstk)) ->
- (l,Zlcase(ci,l,p,br)::pstk)) in
+ (l,Zlcase(ci,l,mk_clos e p,Array.map (mk_clos e) br)::pstk))
+ in
snd (pure_rec lfts stk)
(****************************************************************************)
@@ -124,34 +120,20 @@ let whd_betadeltaiota_nolet env t =
Prod _|Lambda _|Fix _|CoFix _|LetIn _) -> t
| _ -> whd_val (create_clos_infos betadeltaiotanolet env) (inject t)
-(* Beta *)
-
-let beta_appvect c v =
- let rec stacklam env t stack =
- match kind_of_term t, stack with
- Lambda(_,_,c), arg::stacktl -> stacklam (arg::env) c stacktl
- | _ -> applist (substl env t, stack) in
- stacklam [] c (Array.to_list v)
-
-let betazeta_appvect n c v =
- let rec stacklam n env t stack =
- if Int.equal n 0 then applist (substl env t, stack) else
- match kind_of_term t, stack with
- Lambda(_,_,c), arg::stacktl -> stacklam (n-1) (arg::env) c stacktl
- | LetIn(_,b,_,c), _ -> stacklam (n-1) (b::env) c stack
- | _ -> anomaly (Pp.str "Not enough lambda/let's") in
- stacklam n [] c (Array.to_list v)
-
(********************************************************************)
(* Conversion *)
(********************************************************************)
(* Conversion utility functions *)
-type 'a conversion_function = env -> 'a -> 'a -> unit
-type 'a trans_conversion_function = Names.transparent_state -> 'a conversion_function
-type 'a universe_conversion_function = env -> Univ.universes -> 'a -> 'a -> unit
-type 'a trans_universe_conversion_function =
- Names.transparent_state -> 'a universe_conversion_function
+
+(* functions of this type are called from the kernel *)
+type 'a kernel_conversion_function = env -> 'a -> 'a -> unit
+
+(* functions of this type can be called from outside the kernel *)
+type 'a extended_conversion_function =
+ ?l2r:bool -> ?reds:Names.transparent_state -> env ->
+ ?evars:((existential->constr option) * UGraph.t) ->
+ 'a -> 'a -> unit
exception NotConvertible
exception NotConvertibleVect of int
@@ -175,20 +157,22 @@ let is_cumul = function CUMUL -> true | CONV -> false
type 'a universe_compare =
{ (* Might raise NotConvertible *)
compare : env -> conv_pb -> sorts -> sorts -> 'a -> 'a;
- compare_instances: bool -> Univ.Instance.t -> Univ.Instance.t -> 'a -> 'a;
+ compare_instances: flex:bool -> Univ.Instance.t -> Univ.Instance.t -> 'a -> 'a;
}
type 'a universe_state = 'a * 'a universe_compare
type ('a,'b) generic_conversion_function = env -> 'b universe_state -> 'a -> 'a -> 'b
-type 'a infer_conversion_function = env -> Univ.universes -> 'a -> 'a -> Univ.constraints
+type 'a infer_conversion_function = env -> UGraph.t -> 'a -> 'a -> Univ.constraints
let sort_cmp_universes env pb s0 s1 (u, check) =
(check.compare env pb s0 s1 u, check)
-let convert_instances flex u u' (s, check) =
- (check.compare_instances flex u u' s, check)
+(* [flex] should be true for constants, false for inductive types and
+ constructors. *)
+let convert_instances ~flex u u' (s, check) =
+ (check.compare_instances ~flex u u' s, check)
let conv_table_key infos k1 k2 cuniv =
if k1 == k2 then cuniv else
@@ -198,7 +182,7 @@ let conv_table_key infos k1 k2 cuniv =
else
let flex = evaluable_constant cst (info_env infos)
&& RedFlags.red_set (info_flags infos) (RedFlags.fCONST cst)
- in convert_instances flex u u' cuniv
+ in convert_instances ~flex u u' cuniv
| VarKey id, VarKey id' when Id.equal id id' -> cuniv
| RelKey n, RelKey n' when Int.equal n n' -> cuniv
| _ -> raise NotConvertible
@@ -210,9 +194,7 @@ let compare_stacks f fmind lft1 stk1 lft2 stk2 cuniv =
let cu1 = cmp_rec s1 s2 cuniv in
(match (z1,z2) with
| (Zlapp a1,Zlapp a2) ->
- if !left2right then
- Array.fold_left2 (fun cu x y -> f x y cu) cu1 a1 a2
- else Array.fold_right2 f a1 a2 cu1
+ Array.fold_right2 f a1 a2 cu1
| (Zlproj (c1,l1),Zlproj (c2,l2)) ->
if not (eq_constant c1 c2) then
raise NotConvertible
@@ -237,7 +219,6 @@ let rec no_arg_available = function
| Zshift _ :: stk -> no_arg_available stk
| Zapp v :: stk -> Int.equal (Array.length v) 0 && no_arg_available stk
| Zproj _ :: _ -> true
- | Zcase _ :: _ -> true
| ZcaseT _ :: _ -> true
| Zfix _ :: _ -> true
@@ -250,7 +231,6 @@ let rec no_nth_arg_available n = function
if n >= k then no_nth_arg_available (n-k) stk
else false
| Zproj _ :: _ -> true
- | Zcase _ :: _ -> true
| ZcaseT _ :: _ -> true
| Zfix _ :: _ -> true
@@ -260,13 +240,12 @@ let rec no_case_available = function
| Zshift _ :: stk -> no_case_available stk
| Zapp _ :: stk -> no_case_available stk
| Zproj (_,_,p) :: _ -> false
- | Zcase _ :: _ -> false
| ZcaseT _ :: _ -> false
| Zfix _ :: _ -> true
let in_whnf (t,stk) =
match fterm_of t with
- | (FLetIn _ | FCase _ | FCaseT _ | FApp _
+ | (FLetIn _ | FCaseT _ | FApp _
| FCLOS _ | FLIFT _ | FCast _) -> false
| FLambda _ -> no_arg_available stk
| FConstruct _ -> no_case_available stk
@@ -532,8 +511,8 @@ and eqappr cv_pb l2r infos (lft1,st1) (lft2,st2) cuniv =
else raise NotConvertible
(* Should not happen because both (hd1,v1) and (hd2,v2) are in whnf *)
- | ( (FLetIn _, _) | (FCase _,_) | (FCaseT _,_) | (FApp _,_) | (FCLOS _,_) | (FLIFT _,_)
- | (_, FLetIn _) | (_,FCase _) | (_,FCaseT _) | (_,FApp _) | (_,FCLOS _) | (_,FLIFT _)
+ | ( (FLetIn _, _) | (FCaseT _,_) | (FApp _,_) | (FCLOS _,_) | (FLIFT _,_)
+ | (_, FLetIn _) | (_,FCaseT _) | (_,FApp _) | (_,FCLOS _) | (_,FLIFT _)
| (FLOCKED,_) | (_,FLOCKED) ) -> assert false
(* In all other cases, terms are not convertible *)
@@ -558,17 +537,17 @@ and convert_vect l2r infos lft1 lft2 v1 v2 cuniv =
fold 0 cuniv
else raise NotConvertible
-let clos_fconv trans cv_pb l2r evars env univs t1 t2 =
+let clos_gen_conv trans cv_pb l2r evars env univs t1 t2 =
let reds = Closure.RedFlags.red_add_transparent betaiotazeta trans in
let infos = create_clos_infos ~evars reds env in
ccnv cv_pb l2r infos el_id el_id (inject t1) (inject t2) univs
let check_eq univs u u' =
- if not (check_eq univs u u') then raise NotConvertible
+ if not (UGraph.check_eq univs u u') then raise NotConvertible
let check_leq univs u u' =
- if not (check_leq univs u u') then raise NotConvertible
+ if not (UGraph.check_leq univs u u') then raise NotConvertible
let check_sort_cmp_universes env pb s0 s1 univs =
match (s0,s1) with
@@ -594,8 +573,8 @@ let check_sort_cmp_universes env pb s0 s1 univs =
let checked_sort_cmp_universes env pb s0 s1 univs =
check_sort_cmp_universes env pb s0 s1 univs; univs
-let check_convert_instances _flex u u' univs =
- if Univ.Instance.check_eq univs u u' then univs
+let check_convert_instances ~flex u u' univs =
+ if UGraph.check_eq_instances univs u u' then univs
else raise NotConvertible
let checked_universes =
@@ -603,12 +582,12 @@ let checked_universes =
compare_instances = check_convert_instances }
let infer_eq (univs, cstrs as cuniv) u u' =
- if Univ.check_eq univs u u' then cuniv
+ if UGraph.check_eq univs u u' then cuniv
else
univs, (Univ.enforce_eq u u' cstrs)
let infer_leq (univs, cstrs as cuniv) u u' =
- if Univ.check_leq univs u u' then cuniv
+ if UGraph.check_leq univs u u' then cuniv
else
let cstrs' = Univ.enforce_leq u u' cstrs in
univs, cstrs'
@@ -634,60 +613,38 @@ let infer_cmp_universes env pb s0 s1 univs =
| CONV -> infer_eq univs u1 u2)
else univs
-let infer_convert_instances flex u u' (univs,cstrs) =
+let infer_convert_instances ~flex u u' (univs,cstrs) =
(univs, Univ.enforce_eq_instances u u' cstrs)
-let infered_universes : (Univ.universes * Univ.Constraint.t) universe_compare =
+let inferred_universes : (UGraph.t * Univ.Constraint.t) universe_compare =
{ compare = infer_cmp_universes;
compare_instances = infer_convert_instances }
-let trans_fconv_universes reds cv_pb l2r evars env univs t1 t2 =
+let gen_conv cv_pb l2r reds env evars univs t1 t2 =
let b =
if cv_pb = CUMUL then leq_constr_univs univs t1 t2
else eq_constr_univs univs t1 t2
in
if b then ()
else
- let _ = clos_fconv reds cv_pb l2r evars env (univs, checked_universes) t1 t2 in
+ let _ = clos_gen_conv reds cv_pb l2r evars env (univs, checked_universes) t1 t2 in
()
(* Profiling *)
-let trans_fconv_universes =
+let gen_conv cv_pb ?(l2r=false) ?(reds=full_transparent_state) env ?(evars=(fun _->None), universes env) =
+ let evars, univs = evars in
if Flags.profile then
- let trans_fconv_universes_key = Profile.declare_profile "trans_fconv_universes" in
- Profile.profile8 trans_fconv_universes_key trans_fconv_universes
- else trans_fconv_universes
-
-let trans_fconv reds cv_pb l2r evars env =
- trans_fconv_universes reds cv_pb l2r evars env (universes env)
-
-let trans_conv_cmp ?(l2r=false) conv reds = trans_fconv reds conv l2r (fun _->None)
-let trans_conv ?(l2r=false) ?(evars=fun _->None) reds = trans_fconv reds CONV l2r evars
-let trans_conv_leq ?(l2r=false) ?(evars=fun _->None) reds = trans_fconv reds CUMUL l2r evars
-
-let trans_conv_universes ?(l2r=false) ?(evars=fun _->None) reds =
- trans_fconv_universes reds CONV l2r evars
-let trans_conv_leq_universes ?(l2r=false) ?(evars=fun _->None) reds =
- trans_fconv_universes reds CUMUL l2r evars
-
-let fconv = trans_fconv (Id.Pred.full, Cpred.full)
-
-let conv_cmp ?(l2r=false) cv_pb = fconv cv_pb l2r (fun _->None)
-let conv ?(l2r=false) ?(evars=fun _->None) = fconv CONV l2r evars
-let conv_leq ?(l2r=false) ?(evars=fun _->None) = fconv CUMUL l2r evars
-
-let conv_leq_vecti ?(l2r=false) ?(evars=fun _->None) env v1 v2 =
- Array.fold_left2_i
- (fun i _ t1 t2 ->
- try conv_leq ~l2r ~evars env t1 t2
- with NotConvertible -> raise (NotConvertibleVect i))
- ()
- v1
- v2
-
-let generic_conv cv_pb l2r evars reds env univs t1 t2 =
+ let fconv_universes_key = Profile.declare_profile "trans_fconv_universes" in
+ Profile.profile8 fconv_universes_key gen_conv cv_pb l2r reds env evars univs
+ else gen_conv cv_pb l2r reds env evars univs
+
+let conv = gen_conv CONV
+
+let conv_leq = gen_conv CUMUL
+
+let generic_conv cv_pb ~l2r evars reds env univs t1 t2 =
let (s, _) =
- clos_fconv reds cv_pb l2r evars env univs t1 t2
+ clos_gen_conv reds cv_pb l2r evars env univs t1 t2
in s
let infer_conv_universes cv_pb l2r evars reds env univs t1 t2 =
@@ -697,8 +654,8 @@ let infer_conv_universes cv_pb l2r evars reds env univs t1 t2 =
in
if b then cstrs
else
- let univs = ((univs, Univ.Constraint.empty), infered_universes) in
- let ((_,cstrs), _) = clos_fconv reds cv_pb l2r evars env univs t1 t2 in
+ let univs = ((univs, Univ.Constraint.empty), inferred_universes) in
+ let ((_,cstrs), _) = clos_gen_conv reds cv_pb l2r evars env univs t1 t2 in
cstrs
(* Profiling *)
@@ -716,39 +673,20 @@ let infer_conv_leq ?(l2r=false) ?(evars=fun _ -> None) ?(ts=full_transparent_sta
env univs t1 t2 =
infer_conv_universes CUMUL l2r evars ts env univs t1 t2
-(* option for conversion *)
-let nat_conv = ref (fun cv_pb sigma ->
- fconv cv_pb false (sigma.Nativelambda.evars_val))
-let set_nat_conv f = nat_conv := f
-
-let native_conv cv_pb sigma env t1 t2 =
- if eq_constr t1 t2 then ()
- else begin
- let t1 = (it_mkLambda_or_LetIn t1 (rel_context env)) in
- let t2 = (it_mkLambda_or_LetIn t2 (rel_context env)) in
- !nat_conv cv_pb sigma env t1 t2
- end
-
-let vm_conv = ref (fun cv_pb -> fconv cv_pb false (fun _->None))
-let set_vm_conv f = vm_conv := f
+(* This reference avoids always having to link C code with the kernel *)
+let vm_conv = ref (fun cv_pb env ->
+ gen_conv cv_pb env ~evars:((fun _->None), universes env))
+
+let set_vm_conv (f:conv_pb -> Term.types kernel_conversion_function) = vm_conv := f
let vm_conv cv_pb env t1 t2 =
try
!vm_conv cv_pb env t1 t2
with Not_found | Invalid_argument _ ->
Pp.msg_warning (Pp.str "Bytecode compilation failed, falling back to standard conversion");
- fconv cv_pb false (fun _->None) env t1 t2
-
-
-let default_conv = ref (fun cv_pb ?(l2r=false) -> fconv cv_pb l2r (fun _->None))
-
-let set_default_conv f = default_conv := f
+ gen_conv cv_pb env t1 t2
let default_conv cv_pb ?(l2r=false) env t1 t2 =
- try
- !default_conv ~l2r cv_pb env t1 t2
- with Not_found | Invalid_argument _ ->
- Pp.msg_warning (Pp.str "Compilation failed, falling back to standard conversion");
- fconv cv_pb false (fun _->None) env t1 t2
+ gen_conv cv_pb env t1 t2
let default_conv_leq = default_conv CUMUL
(*
@@ -761,12 +699,28 @@ let conv env t1 t2 =
Profile.profile4 convleqkey conv env t1 t2;;
*)
+(* Application with on-the-fly reduction *)
+
+let beta_applist c l =
+ let rec app subst c l =
+ match kind_of_term c, l with
+ | Lambda(_,_,c), arg::l -> app (arg::subst) c l
+ | _ -> applist (substl subst c, l) in
+ app [] c l
+
+let beta_appvect c v = beta_applist c (Array.to_list v)
+
+let beta_app c a = beta_applist c [a]
+
+(* Compatibility *)
+let betazeta_appvect = lambda_appvect_assum
+
(********************************************************************)
(* Special-Purpose Reduction *)
(********************************************************************)
(* pseudo-reduction rule:
- * [hnf_prod_app env s (Prod(_,B)) N --> B[N]
+ * [hnf_prod_app env (Prod(_,B)) N --> B[N]
* with an HNF on the first argument to produce a product.
* if this does not work, then we use the string S as part of our
* error message. *)
@@ -786,11 +740,11 @@ let dest_prod env =
let t = whd_betadeltaiota env c in
match kind_of_term t with
| Prod (n,a,c0) ->
- let d = (n,None,a) in
- decrec (push_rel d env) (add_rel_decl d m) c0
+ let d = LocalAssum (n,a) in
+ decrec (push_rel d env) (Context.Rel.add d m) c0
| _ -> m,t
in
- decrec env empty_rel_context
+ decrec env Context.Rel.empty
(* The same but preserving lets in the context, not internal ones. *)
let dest_prod_assum env =
@@ -798,33 +752,33 @@ let dest_prod_assum env =
let rty = whd_betadeltaiota_nolet env ty in
match kind_of_term rty with
| Prod (x,t,c) ->
- let d = (x,None,t) in
- prodec_rec (push_rel d env) (add_rel_decl d l) c
+ let d = LocalAssum (x,t) in
+ prodec_rec (push_rel d env) (Context.Rel.add d l) c
| LetIn (x,b,t,c) ->
- let d = (x,Some b,t) in
- prodec_rec (push_rel d env) (add_rel_decl d l) c
+ let d = LocalDef (x,b,t) in
+ prodec_rec (push_rel d env) (Context.Rel.add d l) c
| Cast (c,_,_) -> prodec_rec env l c
| _ ->
let rty' = whd_betadeltaiota env rty in
if Term.eq_constr rty' rty then l, rty
else prodec_rec env l rty'
in
- prodec_rec env empty_rel_context
+ prodec_rec env Context.Rel.empty
let dest_lam_assum env =
let rec lamec_rec env l ty =
let rty = whd_betadeltaiota_nolet env ty in
match kind_of_term rty with
| Lambda (x,t,c) ->
- let d = (x,None,t) in
- lamec_rec (push_rel d env) (add_rel_decl d l) c
+ let d = LocalAssum (x,t) in
+ lamec_rec (push_rel d env) (Context.Rel.add d l) c
| LetIn (x,b,t,c) ->
- let d = (x,Some b,t) in
- lamec_rec (push_rel d env) (add_rel_decl d l) c
+ let d = LocalDef (x,b,t) in
+ lamec_rec (push_rel d env) (Context.Rel.add d l) c
| Cast (c,_,_) -> lamec_rec env l c
| _ -> l,rty
in
- lamec_rec env empty_rel_context
+ lamec_rec env Context.Rel.empty
exception NotArity
diff --git a/kernel/reduction.mli b/kernel/reduction.mli
index 6ced5c4985..1b5e5e32a3 100644
--- a/kernel/reduction.mli
+++ b/kernel/reduction.mli
@@ -1,17 +1,14 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
open Term
-open Context
open Environ
-val left2right : bool ref
-
(***********************************************************************
s Reduction functions *)
@@ -28,18 +25,18 @@ val nf_betaiota : env -> constr -> constr
exception NotConvertible
exception NotConvertibleVect of int
-type 'a conversion_function = env -> 'a -> 'a -> unit
-type 'a trans_conversion_function = Names.transparent_state -> 'a conversion_function
-type 'a universe_conversion_function = env -> Univ.universes -> 'a -> 'a -> unit
-type 'a trans_universe_conversion_function =
- Names.transparent_state -> 'a universe_conversion_function
+type 'a kernel_conversion_function = env -> 'a -> 'a -> unit
+type 'a extended_conversion_function =
+ ?l2r:bool -> ?reds:Names.transparent_state -> env ->
+ ?evars:((existential->constr option) * UGraph.t) ->
+ 'a -> 'a -> unit
type conv_pb = CONV | CUMUL
type 'a universe_compare =
{ (* Might raise NotConvertible *)
compare : env -> conv_pb -> sorts -> sorts -> 'a -> 'a;
- compare_instances: bool (* Instance of a flexible constant? *) ->
+ compare_instances: flex:bool ->
Univ.Instance.t -> Univ.Instance.t -> 'a -> 'a;
}
@@ -47,74 +44,63 @@ type 'a universe_state = 'a * 'a universe_compare
type ('a,'b) generic_conversion_function = env -> 'b universe_state -> 'a -> 'a -> 'b
-type 'a infer_conversion_function = env -> Univ.universes -> 'a -> 'a -> Univ.constraints
-
-val check_sort_cmp_universes :
- env -> conv_pb -> sorts -> sorts -> Univ.universes -> unit
+type 'a infer_conversion_function = env -> UGraph.t -> 'a -> 'a -> Univ.constraints
-(* val sort_cmp : *)
-(* conv_pb -> sorts -> sorts -> Univ.constraints -> Univ.constraints *)
+val sort_cmp_universes : env -> conv_pb -> sorts -> sorts ->
+ 'a * 'a universe_compare -> 'a * 'a universe_compare
-(* val conv_sort : sorts conversion_function *)
-(* val conv_sort_leq : sorts conversion_function *)
+(* [flex] should be true for constants, false for inductive types and
+constructors. *)
+val convert_instances : flex:bool -> Univ.Instance.t -> Univ.Instance.t ->
+ 'a * 'a universe_compare -> 'a * 'a universe_compare
-val trans_conv_cmp : ?l2r:bool -> conv_pb -> constr trans_conversion_function
-val trans_conv :
- ?l2r:bool -> ?evars:(existential->constr option) -> constr trans_conversion_function
-val trans_conv_leq :
- ?l2r:bool -> ?evars:(existential->constr option) -> types trans_conversion_function
+val checked_universes : UGraph.t universe_compare
+val inferred_universes : (UGraph.t * Univ.Constraint.t) universe_compare
-val trans_conv_universes :
- ?l2r:bool -> ?evars:(existential->constr option) -> constr trans_universe_conversion_function
-val trans_conv_leq_universes :
- ?l2r:bool -> ?evars:(existential->constr option) -> types trans_universe_conversion_function
+val conv : constr extended_conversion_function
-val conv_cmp : ?l2r:bool -> conv_pb -> constr conversion_function
-val conv :
- ?l2r:bool -> ?evars:(existential->constr option) -> constr conversion_function
-val conv_leq :
- ?l2r:bool -> ?evars:(existential->constr option) -> types conversion_function
-val conv_leq_vecti :
- ?l2r:bool -> ?evars:(existential->constr option) -> types array conversion_function
+val conv_leq : types extended_conversion_function
+(** These conversion functions are used by module subtyping, which needs to infer
+ universe constraints inside the kernel *)
val infer_conv : ?l2r:bool -> ?evars:(existential->constr option) ->
?ts:Names.transparent_state -> constr infer_conversion_function
val infer_conv_leq : ?l2r:bool -> ?evars:(existential->constr option) ->
?ts:Names.transparent_state -> types infer_conversion_function
-val generic_conv : conv_pb -> bool -> (existential->constr option) ->
+val generic_conv : conv_pb -> l2r:bool -> (existential->constr option) ->
Names.transparent_state -> (constr,'a) generic_conversion_function
(** option for conversion *)
-val set_vm_conv : (conv_pb -> types conversion_function) -> unit
-val vm_conv : conv_pb -> types conversion_function
+val set_vm_conv : (conv_pb -> types kernel_conversion_function) -> unit
+val vm_conv : conv_pb -> types kernel_conversion_function
-val set_nat_conv :
- (conv_pb -> Nativelambda.evars -> types conversion_function) -> unit
-val native_conv : conv_pb -> Nativelambda.evars -> types conversion_function
-
-val set_default_conv : (conv_pb -> ?l2r:bool -> types conversion_function) -> unit
-val default_conv : conv_pb -> ?l2r:bool -> types conversion_function
-val default_conv_leq : ?l2r:bool -> types conversion_function
+val default_conv : conv_pb -> ?l2r:bool -> types kernel_conversion_function
+val default_conv_leq : ?l2r:bool -> types kernel_conversion_function
(************************************************************************)
(** Builds an application node, reducing beta redexes it may produce. *)
+val beta_applist : constr -> constr list -> constr
+
+(** Builds an application node, reducing beta redexes it may produce. *)
val beta_appvect : constr -> constr array -> constr
-(** Builds an application node, reducing the [n] first beta-zeta redexes. *)
-val betazeta_appvect : int -> constr -> constr array -> constr
+(** Builds an application node, reducing beta redexe it may produce. *)
+val beta_app : constr -> constr -> constr
(** Pseudo-reduction rule Prod(x,A,B) a --> B[x\a] *)
val hnf_prod_applist : env -> types -> constr list -> types
+(** Compatibility alias for Term.lambda_appvect_assum *)
+val betazeta_appvect : int -> constr -> constr array -> constr
(***********************************************************************
s Recognizing products and arities modulo reduction *)
-val dest_prod : env -> types -> rel_context * types
-val dest_prod_assum : env -> types -> rel_context * types
-val dest_lam_assum : env -> types -> rel_context * types
+val dest_prod : env -> types -> Context.Rel.t * types
+val dest_prod_assum : env -> types -> Context.Rel.t * types
+val dest_lam_assum : env -> types -> Context.Rel.t * types
exception NotArity
diff --git a/kernel/retroknowledge.ml b/kernel/retroknowledge.ml
index cc307f1456..970bc0fcc5 100644
--- a/kernel/retroknowledge.ml
+++ b/kernel/retroknowledge.ml
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
diff --git a/kernel/retroknowledge.mli b/kernel/retroknowledge.mli
index 9a63deb7e3..905a05fe53 100644
--- a/kernel/retroknowledge.mli
+++ b/kernel/retroknowledge.mli
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
diff --git a/kernel/safe_typing.ml b/kernel/safe_typing.ml
index 907ad2a1d4..ce05190b6f 100644
--- a/kernel/safe_typing.ml
+++ b/kernel/safe_typing.ml
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -60,6 +60,7 @@
open Util
open Names
open Declarations
+open Context.Named.Declaration
(** {6 Safe environments }
@@ -81,8 +82,7 @@ open Declarations
These fields could be deduced from [revstruct], but they allow faster
name freshness checks.
- [univ] and [future_cst] : current and future universe constraints
- - [engagement] : are we Set-impredicative?
- - [type_in_type] : does the universe hierarchy collapse?
+ - [engagement] : are we Set-impredicative? does the universe hierarchy collapse?
- [required] : names and digests of Require'd libraries since big-bang.
This field will only grow
- [loads] : list of libraries Require'd inside the current module.
@@ -119,10 +119,9 @@ type safe_environment =
revstruct : structure_body;
modlabels : Label.Set.t;
objlabels : Label.Set.t;
- univ : Univ.constraints;
- future_cst : Univ.constraints Future.computation list;
+ univ : Univ.ContextSet.t;
+ future_cst : Univ.ContextSet.t Future.computation list;
engagement : engagement option;
- type_in_type : bool;
required : vodigest DPMap.t;
loads : (module_path * module_body) list;
local_retroknowledge : Retroknowledge.action list;
@@ -150,9 +149,8 @@ let empty_environment =
modlabels = Label.Set.empty;
objlabels = Label.Set.empty;
future_cst = [];
- univ = Univ.Constraint.empty;
+ univ = Univ.ContextSet.empty;
engagement = None;
- type_in_type = false;
required = DPMap.empty;
loads = [];
local_retroknowledge = [];
@@ -210,36 +208,72 @@ let get_opaque_body env cbo =
(Opaqueproof.force_proof (Environ.opaque_tables env) opaque,
Opaqueproof.force_constraints (Environ.opaque_tables env) opaque)
-let sideff_of_con env c =
+type private_constant = Entries.side_effect
+type private_constants = private_constant list
+
+type private_constant_role = Term_typing.side_effect_role =
+ | Subproof
+ | Schema of inductive * string
+
+let empty_private_constants = []
+let add_private x xs = x :: xs
+let concat_private xs ys = xs @ ys
+let mk_pure_proof = Term_typing.mk_pure_proof
+let inline_private_constants_in_constr = Term_typing.inline_side_effects
+let inline_private_constants_in_definition_entry = Term_typing.inline_entry_side_effects
+let side_effects_of_private_constants x = Term_typing.uniq_seff (List.rev x)
+
+let private_con_of_con env c =
let cbo = Environ.lookup_constant c env.env in
- SEsubproof (c, cbo, get_opaque_body env.env cbo)
-let sideff_of_scheme kind env cl =
- SEscheme(
- List.map (fun (i,c) ->
- let cbo = Environ.lookup_constant c env.env in
- i, c, cbo, get_opaque_body env.env cbo) cl,
- kind)
+ { Entries.from_env = CEphemeron.create env.revstruct;
+ Entries.eff = Entries.SEsubproof (c,cbo,get_opaque_body env.env cbo) }
+
+let private_con_of_scheme ~kind env cl =
+ { Entries.from_env = CEphemeron.create env.revstruct;
+ Entries.eff = Entries.SEscheme(
+ List.map (fun (i,c) ->
+ let cbo = Environ.lookup_constant c env.env in
+ i, c, cbo, get_opaque_body env.env cbo) cl,
+ kind) }
+
+let universes_of_private eff =
+ let open Declarations in
+ List.fold_left (fun acc { Entries.eff } ->
+ match eff with
+ | Entries.SEscheme (l,s) ->
+ List.fold_left (fun acc (_,_,cb,c) ->
+ let acc = match c with
+ | `Nothing -> acc
+ | `Opaque (_, ctx) -> ctx :: acc in
+ if cb.const_polymorphic then acc
+ else (Univ.ContextSet.of_context cb.const_universes) :: acc)
+ acc l
+ | Entries.SEsubproof (c, cb, e) ->
+ if cb.const_polymorphic then acc
+ else Univ.ContextSet.of_context cb.const_universes :: acc)
+ [] eff
let env_of_safe_env senv = senv.env
let env_of_senv = env_of_safe_env
type constraints_addition =
- Now of Univ.constraints | Later of Univ.constraints Future.computation
+ | Now of bool * Univ.ContextSet.t
+ | Later of Univ.ContextSet.t Future.computation
let add_constraints cst senv =
match cst with
| Later fc ->
{senv with future_cst = fc :: senv.future_cst}
- | Now cst ->
+ | Now (poly,cst) ->
{ senv with
- env = Environ.add_constraints cst senv.env;
- univ = Univ.Constraint.union cst senv.univ }
+ env = Environ.push_context_set ~strict:(not poly) cst senv.env;
+ univ = Univ.ContextSet.union cst senv.univ }
let add_constraints_list cst senv =
- List.fold_right add_constraints cst senv
+ List.fold_left (fun acc c -> add_constraints c acc) senv cst
-let push_context_set ctx = add_constraints (Now (Univ.ContextSet.constraints ctx))
-let push_context ctx = add_constraints (Now (Univ.UContext.constraints ctx))
+let push_context_set poly ctx = add_constraints (Now (poly,ctx))
+let push_context poly ctx = add_constraints (Now (poly,Univ.ContextSet.of_context ctx))
let is_curmod_library senv =
match senv.modvariant with LIBRARY -> true | _ -> false
@@ -249,7 +283,7 @@ let join_safe_environment ?(except=Future.UUIDSet.empty) e =
List.fold_left
(fun e fc ->
if Future.UUIDSet.mem (Future.uuid fc) except then e
- else add_constraints (Now (Future.join fc)) e)
+ else add_constraints (Now (false, Future.join fc)) e)
{e with future_cst = []} e.future_cst
let is_joined_environment e = List.is_empty e.future_cst
@@ -329,7 +363,8 @@ let check_required current_libs needed =
hypothesis many many times, and the check performed here would
cost too much. *)
-let safe_push_named (id,_,_ as d) env =
+let safe_push_named d env =
+ let id = get_id d in
let _ =
try
let _ = Environ.lookup_named id env in
@@ -339,23 +374,24 @@ let safe_push_named (id,_,_ as d) env =
let push_named_def (id,de) senv =
- let c,typ,univs = Term_typing.translate_local_def senv.env id de in
- let senv' = push_context univs senv in
- let c, senv' = match c with
- | Def c -> Mod_subst.force_constr c, senv'
+ let c,typ,univs = Term_typing.translate_local_def senv.revstruct senv.env id de in
+ let poly = de.Entries.const_entry_polymorphic in
+ let univs = Univ.ContextSet.of_context univs in
+ let c, univs = match c with
+ | Def c -> Mod_subst.force_constr c, univs
| OpaqueDef o ->
- Opaqueproof.force_proof (Environ.opaque_tables senv'.env) o,
- push_context_set
- (Opaqueproof.force_constraints (Environ.opaque_tables senv'.env) o)
- senv'
+ Opaqueproof.force_proof (Environ.opaque_tables senv.env) o,
+ Univ.ContextSet.union univs
+ (Opaqueproof.force_constraints (Environ.opaque_tables senv.env) o)
| _ -> assert false in
- let env'' = safe_push_named (id,Some c,typ) senv'.env in
- {senv' with env=env''}
+ let senv' = push_context_set poly univs senv in
+ let env'' = safe_push_named (LocalDef (id,c,typ)) senv'.env in
+ univs, {senv' with env=env''}
-let push_named_assum ((id,t),ctx) senv =
- let senv' = push_context_set ctx senv in
+let push_named_assum ((id,t,poly),ctx) senv =
+ let senv' = push_context_set poly ctx senv in
let t = Term_typing.translate_local_assum senv'.env t in
- let env'' = safe_push_named (id,None,t) senv'.env in
+ let env'' = safe_push_named (LocalAssum (id,t)) senv'.env in
{senv' with env=env''}
@@ -376,10 +412,10 @@ let labels_of_mib mib =
let globalize_constant_universes env cb =
if cb.const_polymorphic then
- [Now Univ.Constraint.empty]
+ [Now (true, Univ.ContextSet.empty)]
else
- let cstrs = Univ.UContext.constraints cb.const_universes in
- Now cstrs ::
+ let cstrs = Univ.ContextSet.of_context cb.const_universes in
+ Now (false, cstrs) ::
(match cb.const_body with
| (Undef _ | Def _) -> []
| OpaqueDef lc ->
@@ -387,23 +423,21 @@ let globalize_constant_universes env cb =
| None -> []
| Some fc ->
match Future.peek_val fc with
- | None -> [Later (Future.chain
- ~greedy:(not (Future.is_exn fc))
- ~pure:true fc Univ.ContextSet.constraints)]
- | Some c -> [Now (Univ.ContextSet.constraints c)])
+ | None -> [Later fc]
+ | Some c -> [Now (false, c)])
let globalize_mind_universes mb =
if mb.mind_polymorphic then
- [Now Univ.Constraint.empty]
+ [Now (true, Univ.ContextSet.empty)]
else
- [Now (Univ.UContext.constraints mb.mind_universes)]
+ [Now (false, Univ.ContextSet.of_context mb.mind_universes)]
let constraints_of_sfb env sfb =
match sfb with
| SFBconst cb -> globalize_constant_universes env cb
| SFBmind mib -> globalize_mind_universes mib
- | SFBmodtype mtb -> [Now mtb.mod_constraints]
- | SFBmodule mb -> [Now mb.mod_constraints]
+ | SFBmodtype mtb -> [Now (false, mtb.mod_constraints)]
+ | SFBmodule mb -> [Now (false, mb.mod_constraints)]
(** A generic function for adding a new field in a same environment.
It also performs the corresponding [add_constraints]. *)
@@ -445,19 +479,16 @@ let update_resolver f senv = { senv with modresolver = f senv.modresolver }
(** Insertion of constants and parameters in environment *)
type global_declaration =
- | ConstantEntry of Entries.constant_entry
+ | ConstantEntry of bool * private_constants Entries.constant_entry
| GlobalRecipe of Cooking.recipe
-let add_constant dir l decl senv =
- let kn = make_con senv.modpath dir l in
- let cb = match decl with
- | ConstantEntry ce -> Term_typing.translate_constant senv.env kn ce
- | GlobalRecipe r ->
- let cb = Term_typing.translate_recipe senv.env kn r in
- if DirPath.is_empty dir then Declareops.hcons_const_body cb else cb
- in
+type exported_private_constant =
+ constant * private_constants Entries.constant_entry * private_constant_role
+
+let add_constant_aux no_section senv (kn, cb) =
+ let l = pi3 (Constant.repr3 kn) in
let cb, otab = match cb.const_body with
- | OpaqueDef lc when DirPath.is_empty dir ->
+ | OpaqueDef lc when no_section ->
(* In coqc, opaque constants outside sections will be stored
indirectly in a specific table *)
let od, otab =
@@ -474,7 +505,32 @@ let add_constant dir l decl senv =
(Mod_subst.add_inline_delta_resolver (user_con kn) (lev,None)) senv'
| _ -> senv'
in
- kn, senv''
+ senv''
+
+let add_constant dir l decl senv =
+ let kn = make_con senv.modpath dir l in
+ let no_section = DirPath.is_empty dir in
+ let seff_to_export, decl =
+ match decl with
+ | ConstantEntry (true, ce) ->
+ let exports, ce =
+ Term_typing.export_side_effects senv.revstruct senv.env ce in
+ exports, ConstantEntry (false, ce)
+ | _ -> [], decl
+ in
+ let senv =
+ List.fold_left (add_constant_aux no_section) senv
+ (List.map (fun (kn,cb,_,_) -> kn, cb) seff_to_export) in
+ let senv =
+ let cb =
+ match decl with
+ | ConstantEntry (export_seff,ce) ->
+ Term_typing.translate_constant senv.revstruct senv.env kn ce
+ | GlobalRecipe r ->
+ let cb = Term_typing.translate_recipe senv.env kn r in
+ if no_section then Declareops.hcons_const_body cb else cb in
+ add_constant_aux no_section senv (kn, cb) in
+ (kn, List.map (fun (kn,_,ce,r) -> kn, ce, r) seff_to_export), senv
(** Insertion of inductive types *)
@@ -500,19 +556,20 @@ let add_mind dir l mie senv =
let add_modtype l params_mte inl senv =
let mp = MPdot(senv.modpath, l) in
let mtb = Mod_typing.translate_modtype senv.env mp inl params_mte in
+ let mtb = Declareops.hcons_module_body mtb in
let senv' = add_field (l,SFBmodtype mtb) MT senv in
mp, senv'
(** full_add_module adds module with universes and constraints *)
let full_add_module mb senv =
- let senv = add_constraints (Now mb.mod_constraints) senv in
+ let senv = add_constraints (Now (false, mb.mod_constraints)) senv in
let dp = ModPath.dp mb.mod_mp in
let linkinfo = Nativecode.link_info_of_dirpath dp in
{ senv with env = Modops.add_linked_module mb linkinfo senv.env }
let full_add_module_type mp mt senv =
- let senv = add_constraints (Now mt.mod_constraints) senv in
+ let senv = add_constraints (Now (false, mt.mod_constraints)) senv in
{ senv with env = Modops.add_module_type mp mt senv.env }
(** Insertion of modules *)
@@ -520,6 +577,7 @@ let full_add_module_type mp mt senv =
let add_module l me inl senv =
let mp = MPdot(senv.modpath, l) in
let mb = Mod_typing.translate_module senv.env mp inl me in
+ let mb = Declareops.hcons_module_body mb in
let senv' = add_field (l,SFBmodule mb) M senv in
let senv'' =
if Modops.is_functor mb.mod_type then senv'
@@ -620,8 +678,8 @@ let propagate_senv newdef newenv newresolver senv oldsenv =
modlabels = Label.Set.add (fst newdef) oldsenv.modlabels;
univ =
List.fold_left (fun acc cst ->
- Univ.Constraint.union acc (Future.force cst))
- (Univ.Constraint.union senv.univ oldsenv.univ)
+ Univ.ContextSet.union acc (Future.force cst))
+ (Univ.ContextSet.union senv.univ oldsenv.univ)
now_cst;
future_cst = later_cst @ oldsenv.future_cst;
(* engagement is propagated to the upper level *)
@@ -644,8 +702,8 @@ let end_module l restype senv =
let senv'=
propagate_loads { senv with
env = newenv;
- univ = Univ.Constraint.union senv.univ mb.mod_constraints} in
- let newenv = Environ.add_constraints mb.mod_constraints senv'.env in
+ univ = Univ.ContextSet.union senv.univ mb.mod_constraints} in
+ let newenv = Environ.push_context_set ~strict:true mb.mod_constraints senv'.env in
let newenv = Modops.add_module mb newenv in
let newresolver =
if Modops.is_functor mb.mod_type then oldsenv.modresolver
@@ -670,7 +728,7 @@ let end_modtype l senv =
let () = check_empty_context senv in
let mbids = List.rev_map fst params in
let newenv = Environ.set_opaque_tables oldsenv.env (Environ.opaque_tables senv.env) in
- let newenv = Environ.add_constraints senv.univ newenv in
+ let newenv = Environ.push_context_set ~strict:true senv.univ newenv in
let newenv = set_engagement_opt newenv senv.engagement in
let senv' = propagate_loads {senv with env=newenv} in
let auto_tb = functorize params (NoFunctor (List.rev senv.revstruct)) in
@@ -685,38 +743,32 @@ let end_modtype l senv =
let add_include me is_module inl senv =
let open Mod_typing in
let mp_sup = senv.modpath in
- let sign,cst,resolver =
- if is_module then
- let sign,_,reso,cst = translate_mse_incl senv.env mp_sup inl me in
- sign,cst,reso
- else
- let mtb = translate_modtype senv.env mp_sup inl ([],me) in
- mtb.mod_type,mtb.mod_constraints,mtb.mod_delta
+ let sign,(),resolver,cst =
+ translate_mse_incl is_module senv.env mp_sup inl me
in
- let senv = add_constraints (Now cst) senv in
+ let senv = add_constraints (Now (false, cst)) senv in
(* Include Self support *)
let rec compute_sign sign mb resolver senv =
match sign with
| MoreFunctor(mbid,mtb,str) ->
let cst_sub = Subtyping.check_subtypes senv.env mb mtb in
- let senv = add_constraints (Now cst_sub) senv in
+ let senv =
+ add_constraints
+ (Now (false, Univ.ContextSet.add_constraints cst_sub Univ.ContextSet.empty))
+ senv in
let mpsup_delta =
Modops.inline_delta_resolver senv.env inl mp_sup mbid mtb mb.mod_delta
in
let subst = Mod_subst.map_mbid mbid mp_sup mpsup_delta in
let resolver = Mod_subst.subst_codom_delta_resolver subst resolver in
compute_sign (Modops.subst_signature subst str) mb resolver senv
- | str -> resolver,str,senv
+ | NoFunctor str -> resolver,str,senv
in
- let resolver,sign,senv =
+ let resolver,str,senv =
let struc = NoFunctor (List.rev senv.revstruct) in
- let mtb = build_mtb mp_sup struc Univ.Constraint.empty senv.modresolver in
+ let mtb = build_mtb mp_sup struc Univ.ContextSet.empty senv.modresolver in
compute_sign sign mtb resolver senv
in
- let str = match sign with
- | NoFunctor struc -> struc
- | MoreFunctor _ -> Modops.error_higher_order_include ()
- in
let senv = update_resolver (Mod_subst.add_delta_resolver resolver) senv
in
let add senv ((l,elem) as field) =
@@ -804,8 +856,10 @@ let import lib cst vodigest senv =
check_engagement senv.env lib.comp_enga;
let mp = MPfile lib.comp_name in
let mb = lib.comp_mod in
- let env = Environ.add_constraints mb.mod_constraints senv.env in
- let env = Environ.push_context_set cst env in
+ let env = Environ.push_context_set ~strict:true
+ (Univ.ContextSet.union mb.mod_constraints cst)
+ senv.env
+ in
mp,
{ senv with
env =
@@ -858,7 +912,9 @@ let register_inline kn senv =
let env = { env with env_globals = new_globals } in
{ senv with env = env_of_pre_env env }
-let add_constraints c = add_constraints (Now c)
+let add_constraints c =
+ add_constraints
+ (Now (false, Univ.ContextSet.add_constraints c Univ.ContextSet.empty))
(* NB: The next old comment probably refers to [propagate_loads] above.
diff --git a/kernel/safe_typing.mli b/kernel/safe_typing.mli
index 2b4324b96f..71dac321f6 100644
--- a/kernel/safe_typing.mli
+++ b/kernel/safe_typing.mli
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -39,10 +39,30 @@ type 'a safe_transformer = safe_environment -> 'a * safe_environment
(** {6 Stm machinery } *)
-val sideff_of_con : safe_environment -> constant -> Declarations.side_effect
-val sideff_of_scheme :
- string -> safe_environment -> (inductive * constant) list ->
- Declarations.side_effect
+type private_constant
+type private_constants
+
+type private_constant_role =
+ | Subproof
+ | Schema of inductive * string
+
+val side_effects_of_private_constants :
+ private_constants -> Entries.side_effects
+
+val empty_private_constants : private_constants
+val add_private : private_constant -> private_constants -> private_constants
+val concat_private : private_constants -> private_constants -> private_constants
+
+val private_con_of_con : safe_environment -> constant -> private_constant
+val private_con_of_scheme : kind:string -> safe_environment -> (inductive * constant) list -> private_constant
+
+val mk_pure_proof : Constr.constr -> private_constants Entries.proof_output
+val inline_private_constants_in_constr :
+ Environ.env -> Constr.constr -> private_constants -> Constr.constr
+val inline_private_constants_in_definition_entry :
+ Environ.env -> private_constants Entries.definition_entry -> private_constants Entries.definition_entry
+
+val universes_of_private : private_constants -> Univ.universe_context_set list
val is_curmod_library : safe_environment -> bool
@@ -57,18 +77,29 @@ val is_joined_environment : safe_environment -> bool
(** Insertion of local declarations (Local or Variables) *)
val push_named_assum :
- (Id.t * Term.types) Univ.in_universe_context_set -> safe_transformer0
+ (Id.t * Term.types * bool (* polymorphic *))
+ Univ.in_universe_context_set -> safe_transformer0
+
+(** Returns the full universe context necessary to typecheck the definition
+ (futures are forced) *)
val push_named_def :
- Id.t * Entries.definition_entry -> safe_transformer0
+ Id.t * private_constants Entries.definition_entry -> Univ.universe_context_set safe_transformer
(** Insertion of global axioms or definitions *)
type global_declaration =
- | ConstantEntry of Entries.constant_entry
+ (* bool: export private constants *)
+ | ConstantEntry of bool * private_constants Entries.constant_entry
| GlobalRecipe of Cooking.recipe
+type exported_private_constant =
+ constant * private_constants Entries.constant_entry * private_constant_role
+
+(** returns the main constant plus a list of auxiliary constants (empty
+ unless one requires the side effects to be exported) *)
val add_constant :
- DirPath.t -> Label.t -> global_declaration -> constant safe_transformer
+ DirPath.t -> Label.t -> global_declaration ->
+ (constant * exported_private_constant list) safe_transformer
(** Adding an inductive type *)
@@ -88,10 +119,10 @@ val add_modtype :
(** Adding universe constraints *)
val push_context_set :
- Univ.universe_context_set -> safe_transformer0
+ bool -> Univ.universe_context_set -> safe_transformer0
val push_context :
- Univ.universe_context -> safe_transformer0
+ bool -> Univ.universe_context -> safe_transformer0
val add_constraints :
Univ.constraints -> safe_transformer0
diff --git a/kernel/sorts.ml b/kernel/sorts.ml
index e2854abfd3..62013b38f1 100644
--- a/kernel/sorts.ml
+++ b/kernel/sorts.ml
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -98,7 +98,7 @@ module Hsorts =
let u' = huniv u in
if u' == u then c else Type u'
| s -> s
- let equal s1 s2 = match (s1,s2) with
+ let eq s1 s2 = match (s1,s2) with
| (Prop c1, Prop c2) -> c1 == c2
| (Type u1, Type u2) -> u1 == u2
|_ -> false
diff --git a/kernel/sorts.mli b/kernel/sorts.mli
index cd65b23152..eb4697ad6d 100644
--- a/kernel/sorts.mli
+++ b/kernel/sorts.mli
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
diff --git a/kernel/subtyping.ml b/kernel/subtyping.ml
index db155e6c86..5efc1078ee 100644
--- a/kernel/subtyping.ml
+++ b/kernel/subtyping.ml
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -311,15 +311,19 @@ let check_constant cst env mp1 l info1 cb2 spec2 subst1 subst2 =
try
(* The environment with the expected universes plus equality
of the body instances with the expected instance *)
- let env = Environ.add_constraints cstrs env in
- (* Check that the given definition does not add any constraint over
- the expected ones, so that it can be used in place of the original. *)
- if Univ.check_constraints ctx1 (Environ.universes env) then
+ let ctxi = Univ.Instance.append inst1 inst2 in
+ let ctx = Univ.UContext.make (ctxi, cstrs) in
+ let env = Environ.push_context ctx env in
+ (* Check that the given definition does not add any constraint over
+ the expected ones, so that it can be used in place of
+ the original. *)
+ if UGraph.check_constraints ctx1 (Environ.universes env) then
cstrs, env, inst2
else error (IncompatibleConstraints ctx1)
with Univ.UniverseInconsistency incon ->
error (IncompatibleUniverses incon)
- else cst, env, Univ.Instance.empty
+ else
+ cst, env, Univ.Instance.empty
in
(* Now check types *)
let typ1 = Typeops.type_of_constant_type env' cb1.const_type in
@@ -456,6 +460,7 @@ and check_modtypes cst env mtb1 mtb2 subst1 subst2 equiv =
let check_subtypes env sup super =
let env = add_module_type sup.mod_mp sup env in
+ let env = Environ.push_context_set ~strict:true super.mod_constraints env in
check_modtypes Univ.Constraint.empty env
(strengthen sup sup.mod_mp) super empty_subst
(map_mp super.mod_mp sup.mod_mp sup.mod_delta) false
diff --git a/kernel/subtyping.mli b/kernel/subtyping.mli
index 443f5037f5..a00eb87329 100644
--- a/kernel/subtyping.mli
+++ b/kernel/subtyping.mli
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
diff --git a/kernel/term.ml b/kernel/term.ml
index 7bf4c8182d..4416770fe4 100644
--- a/kernel/term.ml
+++ b/kernel/term.ml
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -10,7 +10,6 @@ open Util
open Pp
open Errors
open Names
-open Context
open Vars
(**********************************************************************)
@@ -384,40 +383,46 @@ let mkNamedLambda id typ c = mkLambda (Name id, typ, subst_var id c)
let mkNamedLetIn id c1 t c2 = mkLetIn (Name id, c1, t, subst_var id c2)
(* Constructs either [(x:t)c] or [[x=b:t]c] *)
-let mkProd_or_LetIn (na,body,t) c =
- match body with
- | None -> mkProd (na, t, c)
- | Some b -> mkLetIn (na, b, t, c)
-
-let mkNamedProd_or_LetIn (id,body,t) c =
- match body with
- | None -> mkNamedProd id t c
- | Some b -> mkNamedLetIn id b t c
+let mkProd_or_LetIn decl c =
+ let open Context.Rel.Declaration in
+ match decl with
+ | LocalAssum (na,t) -> mkProd (na, t, c)
+ | LocalDef (na,b,t) -> mkLetIn (na, b, t, c)
+
+let mkNamedProd_or_LetIn decl c =
+ let open Context.Named.Declaration in
+ match decl with
+ | LocalAssum (id,t) -> mkNamedProd id t c
+ | LocalDef (id,b,t) -> mkNamedLetIn id b t c
(* Constructs either [(x:t)c] or [c] where [x] is replaced by [b] *)
-let mkProd_wo_LetIn (na,body,t) c =
- match body with
- | None -> mkProd (na, t, c)
- | Some b -> subst1 b c
-
-let mkNamedProd_wo_LetIn (id,body,t) c =
- match body with
- | None -> mkNamedProd id t c
- | Some b -> subst1 b (subst_var id c)
+let mkProd_wo_LetIn decl c =
+ let open Context.Rel.Declaration in
+ match decl with
+ | LocalAssum (na,t) -> mkProd (na, t, c)
+ | LocalDef (na,b,t) -> subst1 b c
+
+let mkNamedProd_wo_LetIn decl c =
+ let open Context.Named.Declaration in
+ match decl with
+ | LocalAssum (id,t) -> mkNamedProd id t c
+ | LocalDef (id,b,t) -> subst1 b (subst_var id c)
(* non-dependent product t1 -> t2 *)
let mkArrow t1 t2 = mkProd (Anonymous, t1, t2)
(* Constructs either [[x:t]c] or [[x=b:t]c] *)
-let mkLambda_or_LetIn (na,body,t) c =
- match body with
- | None -> mkLambda (na, t, c)
- | Some b -> mkLetIn (na, b, t, c)
-
-let mkNamedLambda_or_LetIn (id,body,t) c =
- match body with
- | None -> mkNamedLambda id t c
- | Some b -> mkNamedLetIn id b t c
+let mkLambda_or_LetIn decl c =
+ let open Context.Rel.Declaration in
+ match decl with
+ | LocalAssum (na,t) -> mkLambda (na, t, c)
+ | LocalDef (na,b,t) -> mkLetIn (na, b, t, c)
+
+let mkNamedLambda_or_LetIn decl c =
+ let open Context.Named.Declaration in
+ match decl with
+ | LocalAssum (id,t) -> mkNamedLambda id t c
+ | LocalDef (id,b,t) -> mkNamedLetIn id b t c
(* prodn n [xn:Tn;..;x1:T1;Gamma] b = (x1:T1)..(xn:Tn)b *)
let prodn n env b =
@@ -471,26 +476,58 @@ let rec to_prod n lam =
| Cast (c,_,_) -> to_prod n c
| _ -> errorlabstrm "to_prod" (mt ())
-(* pseudo-reduction rule:
- * [prod_app s (Prod(_,B)) N --> B[N]
- * with an strip_outer_cast on the first argument to produce a product *)
+let it_mkProd_or_LetIn = List.fold_left (fun c d -> mkProd_or_LetIn d c)
+let it_mkLambda_or_LetIn = List.fold_left (fun c d -> mkLambda_or_LetIn d c)
-let prod_app t n =
- match kind_of_term (strip_outer_cast t) with
- | Prod (_,_,b) -> subst1 n b
- | _ ->
- errorlabstrm "prod_app"
- (str"Needed a product, but didn't find one" ++ fnl ())
+(* Application with expected on-the-fly reduction *)
+let lambda_applist c l =
+ let rec app subst c l =
+ match kind_of_term c, l with
+ | Lambda(_,_,c), arg::l -> app (arg::subst) c l
+ | _, [] -> substl subst c
+ | _ -> anomaly (Pp.str "Not enough lambda's") in
+ app [] c l
-(* prod_appvect T [| a1 ; ... ; an |] -> (T a1 ... an) *)
-let prod_appvect t nL = Array.fold_left prod_app t nL
+let lambda_appvect c v = lambda_applist c (Array.to_list v)
+
+let lambda_applist_assum n c l =
+ let rec app n subst t l =
+ if Int.equal n 0 then
+ if l == [] then substl subst t
+ else anomaly (Pp.str "Not enough arguments")
+ else match kind_of_term t, l with
+ | Lambda(_,_,c), arg::l -> app (n-1) (arg::subst) c l
+ | LetIn(_,b,_,c), _ -> app (n-1) (substl subst b::subst) c l
+ | _ -> anomaly (Pp.str "Not enough lambda/let's") in
+ app n [] c l
+
+let lambda_appvect_assum n c v = lambda_applist_assum n c (Array.to_list v)
(* prod_applist T [ a1 ; ... ; an ] -> (T a1 ... an) *)
-let prod_applist t nL = List.fold_left prod_app t nL
+let prod_applist c l =
+ let rec app subst c l =
+ match kind_of_term c, l with
+ | Prod(_,_,c), arg::l -> app (arg::subst) c l
+ | _, [] -> substl subst c
+ | _ -> anomaly (Pp.str "Not enough prod's") in
+ app [] c l
-let it_mkProd_or_LetIn = List.fold_left (fun c d -> mkProd_or_LetIn d c)
-let it_mkLambda_or_LetIn = List.fold_left (fun c d -> mkLambda_or_LetIn d c)
+(* prod_appvect T [| a1 ; ... ; an |] -> (T a1 ... an) *)
+let prod_appvect c v = prod_applist c (Array.to_list v)
+
+let prod_applist_assum n c l =
+ let rec app n subst t l =
+ if Int.equal n 0 then
+ if l == [] then substl subst t
+ else anomaly (Pp.str "Not enough arguments")
+ else match kind_of_term t, l with
+ | Prod(_,_,c), arg::l -> app (n-1) (arg::subst) c l
+ | LetIn(_,b,_,c), _ -> app (n-1) (substl subst b::subst) c l
+ | _ -> anomaly (Pp.str "Not enough prod/let's") in
+ app n [] c l
+
+let prod_appvect_assum n c v = prod_applist_assum n c (Array.to_list v)
(*********************************)
(* Other term destructors *)
@@ -545,77 +582,84 @@ let decompose_lam_n n =
(* Transforms a product term (x1:T1)..(xn:Tn)T into the pair
([(xn,Tn);...;(x1,T1)],T), where T is not a product *)
let decompose_prod_assum =
+ let open Context.Rel.Declaration in
let rec prodec_rec l c =
match kind_of_term c with
- | Prod (x,t,c) -> prodec_rec (add_rel_decl (x,None,t) l) c
- | LetIn (x,b,t,c) -> prodec_rec (add_rel_decl (x,Some b,t) l) c
+ | Prod (x,t,c) -> prodec_rec (Context.Rel.add (LocalAssum (x,t)) l) c
+ | LetIn (x,b,t,c) -> prodec_rec (Context.Rel.add (LocalDef (x,b,t)) l) c
| Cast (c,_,_) -> prodec_rec l c
| _ -> l,c
in
- prodec_rec empty_rel_context
+ prodec_rec Context.Rel.empty
(* Transforms a lambda term [x1:T1]..[xn:Tn]T into the pair
([(xn,Tn);...;(x1,T1)],T), where T is not a lambda *)
let decompose_lam_assum =
let rec lamdec_rec l c =
+ let open Context.Rel.Declaration in
match kind_of_term c with
- | Lambda (x,t,c) -> lamdec_rec (add_rel_decl (x,None,t) l) c
- | LetIn (x,b,t,c) -> lamdec_rec (add_rel_decl (x,Some b,t) l) c
+ | Lambda (x,t,c) -> lamdec_rec (Context.Rel.add (LocalAssum (x,t)) l) c
+ | LetIn (x,b,t,c) -> lamdec_rec (Context.Rel.add (LocalDef (x,b,t)) l) c
| Cast (c,_,_) -> lamdec_rec l c
| _ -> l,c
in
- lamdec_rec empty_rel_context
+ lamdec_rec Context.Rel.empty
-(* Given a positive integer n, transforms a product term (x1:T1)..(xn:Tn)T
- into the pair ([(xn,Tn);...;(x1,T1)],T) *)
+(* Given a positive integer n, decompose a product or let-in term
+ of the form [forall (x1:T1)..(xi:=ci:Ti)..(xn:Tn), T] into the pair
+ of the quantifying context [(xn,None,Tn);..;(xi,Some
+ ci,Ti);..;(x1,None,T1)] and of the inner type [T]) *)
let decompose_prod_n_assum n =
if n < 0 then
error "decompose_prod_n_assum: integer parameter must be positive";
let rec prodec_rec l n c =
if Int.equal n 0 then l,c
- else match kind_of_term c with
- | Prod (x,t,c) -> prodec_rec (add_rel_decl (x,None,t) l) (n-1) c
- | LetIn (x,b,t,c) -> prodec_rec (add_rel_decl (x,Some b,t) l) (n-1) c
- | Cast (c,_,_) -> prodec_rec l n c
- | c -> error "decompose_prod_n_assum: not enough assumptions"
+ else
+ let open Context.Rel.Declaration in
+ match kind_of_term c with
+ | Prod (x,t,c) -> prodec_rec (Context.Rel.add (LocalAssum (x,t)) l) (n-1) c
+ | LetIn (x,b,t,c) -> prodec_rec (Context.Rel.add (LocalDef (x,b,t)) l) (n-1) c
+ | Cast (c,_,_) -> prodec_rec l n c
+ | c -> error "decompose_prod_n_assum: not enough assumptions"
in
- prodec_rec empty_rel_context n
+ prodec_rec Context.Rel.empty n
-(* Given a positive integer n, transforms a lambda term [x1:T1]..[xn:Tn]T
- into the pair ([(xn,Tn);...;(x1,T1)],T)
+(* Given a positive integer n, decompose a lambda or let-in term [fun
+ (x1:T1)..(xi:=ci:Ti)..(xn:Tn) => T] into the pair of the abstracted
+ context [(xn,None,Tn);...;(xi,Some ci,Ti);...;(x1,None,T1)] and of
+ the inner body [T].
Lets in between are not expanded but turn into local definitions,
- but n is the actual number of destructurated lambdas. *)
+ but n is the actual number of destructurated lambdas. *)
let decompose_lam_n_assum n =
if n < 0 then
error "decompose_lam_n_assum: integer parameter must be positive";
let rec lamdec_rec l n c =
if Int.equal n 0 then l,c
- else match kind_of_term c with
- | Lambda (x,t,c) -> lamdec_rec (add_rel_decl (x,None,t) l) (n-1) c
- | LetIn (x,b,t,c) -> lamdec_rec (add_rel_decl (x,Some b,t) l) n c
- | Cast (c,_,_) -> lamdec_rec l n c
- | c -> error "decompose_lam_n_assum: not enough abstractions"
+ else
+ let open Context.Rel.Declaration in
+ match kind_of_term c with
+ | Lambda (x,t,c) -> lamdec_rec (Context.Rel.add (LocalAssum (x,t)) l) (n-1) c
+ | LetIn (x,b,t,c) -> lamdec_rec (Context.Rel.add (LocalDef (x,b,t)) l) n c
+ | Cast (c,_,_) -> lamdec_rec l n c
+ | c -> error "decompose_lam_n_assum: not enough abstractions"
in
- lamdec_rec empty_rel_context n
-
-(* (nb_lam [na1:T1]...[nan:Tan]c) where c is not an abstraction
- * gives n (casts are ignored) *)
-let nb_lam =
- let rec nbrec n c = match kind_of_term c with
- | Lambda (_,_,c) -> nbrec (n+1) c
- | Cast (c,_,_) -> nbrec n c
- | _ -> n
- in
- nbrec 0
-
-(* similar to nb_lam, but gives the number of products instead *)
-let nb_prod =
- let rec nbrec n c = match kind_of_term c with
- | Prod (_,_,c) -> nbrec (n+1) c
- | Cast (c,_,_) -> nbrec n c
- | _ -> n
+ lamdec_rec Context.Rel.empty n
+
+(* Same, counting let-in *)
+let decompose_lam_n_decls n =
+ if n < 0 then
+ error "decompose_lam_n_decls: integer parameter must be positive";
+ let rec lamdec_rec l n c =
+ if Int.equal n 0 then l,c
+ else
+ let open Context.Rel.Declaration in
+ match kind_of_term c with
+ | Lambda (x,t,c) -> lamdec_rec (Context.Rel.add (LocalAssum (x,t)) l) (n-1) c
+ | LetIn (x,b,t,c) -> lamdec_rec (Context.Rel.add (LocalDef (x,b,t)) l) (n-1) c
+ | Cast (c,_,_) -> lamdec_rec l n c
+ | c -> error "decompose_lam_n_decls: not enough abstractions"
in
- nbrec 0
+ lamdec_rec Context.Rel.empty n
let prod_assum t = fst (decompose_prod_assum t)
let prod_n_assum n t = fst (decompose_prod_n_assum n t)
@@ -636,13 +680,14 @@ let strip_lam_n n t = snd (decompose_lam_n n t)
Such a term can canonically be seen as the pair of a context of types
and of a sort *)
-type arity = rel_context * sorts
+type arity = Context.Rel.t * sorts
let destArity =
+ let open Context.Rel.Declaration in
let rec prodec_rec l c =
match kind_of_term c with
- | Prod (x,t,c) -> prodec_rec ((x,None,t)::l) c
- | LetIn (x,b,t,c) -> prodec_rec ((x,Some b,t)::l) c
+ | Prod (x,t,c) -> prodec_rec (LocalAssum (x,t) :: l) c
+ | LetIn (x,b,t,c) -> prodec_rec (LocalDef (x,b,t) :: l) c
| Cast (c,_,_) -> prodec_rec l c
| Sort s -> l,s
| _ -> anomaly ~label:"destArity" (Pp.str "not an arity")
diff --git a/kernel/term.mli b/kernel/term.mli
index 501aaf741e..32267f6c4c 100644
--- a/kernel/term.mli
+++ b/kernel/term.mli
@@ -1,13 +1,12 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
open Names
-open Context
(** {5 Redeclaration of types from module Constr and Sorts}
@@ -203,7 +202,7 @@ val destCoFix : constr -> cofixpoint
(** non-dependent product [t1 -> t2], an alias for
[forall (_:t1), t2]. Beware [t_2] is NOT lifted.
- Eg: in context [A:Prop], [A->A] is built by [(mkArrow (mkRel 0) (mkRel 1))]
+ Eg: in context [A:Prop], [A->A] is built by [(mkArrow (mkRel 1) (mkRel 2))]
*)
val mkArrow : types -> types -> constr
@@ -213,14 +212,14 @@ val mkNamedLetIn : Id.t -> constr -> types -> constr -> constr
val mkNamedProd : Id.t -> types -> types -> types
(** Constructs either [(x:t)c] or [[x=b:t]c] *)
-val mkProd_or_LetIn : rel_declaration -> types -> types
-val mkProd_wo_LetIn : rel_declaration -> types -> types
-val mkNamedProd_or_LetIn : named_declaration -> types -> types
-val mkNamedProd_wo_LetIn : named_declaration -> types -> types
+val mkProd_or_LetIn : Context.Rel.Declaration.t -> types -> types
+val mkProd_wo_LetIn : Context.Rel.Declaration.t -> types -> types
+val mkNamedProd_or_LetIn : Context.Named.Declaration.t -> types -> types
+val mkNamedProd_wo_LetIn : Context.Named.Declaration.t -> types -> types
(** Constructs either [[x:t]c] or [[x=b:t]c] *)
-val mkLambda_or_LetIn : rel_declaration -> constr -> constr
-val mkNamedLambda_or_LetIn : named_declaration -> constr -> constr
+val mkLambda_or_LetIn : Context.Rel.Declaration.t -> constr -> constr
+val mkNamedLambda_or_LetIn : Context.Named.Declaration.t -> constr -> constr
(** {5 Other term constructors. } *)
@@ -262,14 +261,34 @@ val to_lambda : int -> constr -> constr
where [l] is [fun (x_1:T_1)...(x_n:T_n) => T] *)
val to_prod : int -> constr -> constr
+val it_mkLambda_or_LetIn : constr -> Context.Rel.t -> constr
+val it_mkProd_or_LetIn : types -> Context.Rel.t -> types
+
+(** In [lambda_applist c args], [c] is supposed to have the form
+ [λΓ.c] with [Γ] without let-in; it returns [c] with the variables
+ of [Γ] instantiated by [args]. *)
+val lambda_applist : constr -> constr list -> constr
+val lambda_appvect : constr -> constr array -> constr
+
+(** In [lambda_applist_assum n c args], [c] is supposed to have the
+ form [λΓ.c] with [Γ] of length [m] and possibly with let-ins; it
+ returns [c] with the assumptions of [Γ] instantiated by [args] and
+ the local definitions of [Γ] expanded. *)
+val lambda_applist_assum : int -> constr -> constr list -> constr
+val lambda_appvect_assum : int -> constr -> constr array -> constr
+
(** pseudo-reduction rule *)
(** [prod_appvect] [forall (x1:B1;...;xn:Bn), B] [a1...an] @return [B[a1...an]] *)
val prod_appvect : constr -> constr array -> constr
val prod_applist : constr -> constr list -> constr
-val it_mkLambda_or_LetIn : constr -> rel_context -> constr
-val it_mkProd_or_LetIn : types -> rel_context -> types
+(** In [prod_appvect_assum n c args], [c] is supposed to have the
+ form [∀Γ.c] with [Γ] of length [m] and possibly with let-ins; it
+ returns [c] with the assumptions of [Γ] instantiated by [args] and
+ the local definitions of [Γ] expanded. *)
+val prod_appvect_assum : int -> constr -> constr array -> constr
+val prod_applist_assum : int -> constr -> constr list -> constr
(** {5 Other term destructors. } *)
@@ -281,40 +300,42 @@ val decompose_prod : constr -> (Name.t*constr) list * constr
{% $ %}([(x_n,T_n);...;(x_1,T_1)],T){% $ %}, where {% $ %}T{% $ %} is not a lambda. *)
val decompose_lam : constr -> (Name.t*constr) list * constr
-(** Given a positive integer n, transforms a product term
+(** Given a positive integer n, decompose a product term
{% $ %}(x_1:T_1)..(x_n:T_n)T{% $ %}
- into the pair {% $ %}([(xn,Tn);...;(x1,T1)],T){% $ %}. *)
+ into the pair {% $ %}([(xn,Tn);...;(x1,T1)],T){% $ %}.
+ Raise a user error if not enough products. *)
val decompose_prod_n : int -> constr -> (Name.t * constr) list * constr
-(** Given a positive integer {% $ %}n{% $ %}, transforms a lambda term
- {% $ %}[x_1:T_1]..[x_n:T_n]T{% $ %} into the pair {% $ %}([(x_n,T_n);...;(x_1,T_1)],T){% $ %} *)
+(** Given a positive integer {% $ %}n{% $ %}, decompose a lambda term
+ {% $ %}[x_1:T_1]..[x_n:T_n]T{% $ %} into the pair {% $ %}([(x_n,T_n);...;(x_1,T_1)],T){% $ %}.
+ Raise a user error if not enough lambdas. *)
val decompose_lam_n : int -> constr -> (Name.t * constr) list * constr
(** Extract the premisses and the conclusion of a term of the form
"(xi:Ti) ... (xj:=cj:Tj) ..., T" where T is not a product nor a let *)
-val decompose_prod_assum : types -> rel_context * types
+val decompose_prod_assum : types -> Context.Rel.t * types
(** Idem with lambda's *)
-val decompose_lam_assum : constr -> rel_context * constr
+val decompose_lam_assum : constr -> Context.Rel.t * constr
+
+(** Idem but extract the first [n] premisses, counting let-ins. *)
+val decompose_prod_n_assum : int -> types -> Context.Rel.t * types
-(** Idem but extract the first [n] premisses *)
-val decompose_prod_n_assum : int -> types -> rel_context * types
-val decompose_lam_n_assum : int -> constr -> rel_context * constr
+(** Idem for lambdas, _not_ counting let-ins *)
+val decompose_lam_n_assum : int -> constr -> Context.Rel.t * constr
-(** [nb_lam] {% $ %}[x_1:T_1]...[x_n:T_n]c{% $ %} where {% $ %}c{% $ %} is not an abstraction
- gives {% $ %}n{% $ %} (casts are ignored) *)
-val nb_lam : constr -> int
+(** Idem, counting let-ins *)
+val decompose_lam_n_decls : int -> constr -> Context.Rel.t * constr
-(** Similar to [nb_lam], but gives the number of products instead *)
-val nb_prod : constr -> int
+(** Return the premisses/parameters of a type/term (let-in included) *)
+val prod_assum : types -> Context.Rel.t
+val lam_assum : constr -> Context.Rel.t
-(** Returns the premisses/parameters of a type/term (let-in included) *)
-val prod_assum : types -> rel_context
-val lam_assum : constr -> rel_context
+(** Return the first n-th premisses/parameters of a type (let included and counted) *)
+val prod_n_assum : int -> types -> Context.Rel.t
-(** Returns the first n-th premisses/parameters of a type/term (let included)*)
-val prod_n_assum : int -> types -> rel_context
-val lam_n_assum : int -> constr -> rel_context
+(** Return the first n-th premisses/parameters of a term (let included but not counted) *)
+val lam_n_assum : int -> constr -> Context.Rel.t
(** Remove the premisses/parameters of a type/term *)
val strip_prod : types -> types
@@ -328,11 +349,11 @@ val strip_lam_n : int -> constr -> constr
val strip_prod_assum : types -> types
val strip_lam_assum : constr -> constr
-(** flattens application lists *)
+(** Flattens application lists *)
val collapse_appl : constr -> constr
-(** Removes recursively the casts around a term i.e.
+(** Remove recursively the casts around a term i.e.
[strip_outer_cast (Cast (Cast ... (Cast c, t) ... ))] is [c]. *)
val strip_outer_cast : constr -> constr
@@ -347,15 +368,15 @@ val under_outer_cast : (constr -> constr) -> constr -> constr
Such a term can canonically be seen as the pair of a context of types
and of a sort *)
-type arity = rel_context * sorts
+type arity = Context.Rel.t * sorts
(** Build an "arity" from its canonical form *)
val mkArity : arity -> types
-(** Destructs an "arity" into its canonical form *)
+(** Destruct an "arity" into its canonical form *)
val destArity : types -> arity
-(** Tells if a term has the form of an arity *)
+(** Tell if a term has the form of an arity *)
val isArity : types -> bool
(** {5 Kind of type} *)
@@ -427,11 +448,11 @@ val eq_constr : constr -> constr -> bool
(** [eq_constr_univs u a b] is [true] if [a] equals [b] modulo alpha, casts,
application grouping and the universe constraints in [u]. *)
-val eq_constr_univs : constr Univ.check_function
+val eq_constr_univs : constr UGraph.check_function
(** [leq_constr_univs u a b] is [true] if [a] is convertible to [b] modulo
alpha, casts, application grouping and the universe constraints in [u]. *)
-val leq_constr_univs : constr Univ.check_function
+val leq_constr_univs : constr UGraph.check_function
(** [eq_constr_univs a b] [true, c] if [a] equals [b] modulo alpha, casts,
application grouping and ignoring universe instances. *)
diff --git a/kernel/term_typing.ml b/kernel/term_typing.ml
index 83e566041f..3839135a86 100644
--- a/kernel/term_typing.ml
+++ b/kernel/term_typing.ml
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -16,7 +16,6 @@ open Errors
open Util
open Names
open Term
-open Context
open Declarations
open Environ
open Entries
@@ -43,10 +42,29 @@ let map_option_typ = function None -> `None | Some x -> `Some x
(* Insertion of constants and parameters in environment. *)
-let mk_pure_proof c = (c, Univ.ContextSet.empty), Declareops.no_seff
+let mk_pure_proof c = (c, Univ.ContextSet.empty), []
-let handle_side_effects env body side_eff =
- let handle_sideff t se =
+let equal_eff e1 e2 =
+ let open Entries in
+ match e1, e2 with
+ | { eff = SEsubproof (c1,_,_) }, { eff = SEsubproof (c2,_,_) } ->
+ Names.Constant.equal c1 c2
+ | { eff = SEscheme (cl1,_) }, { eff = SEscheme (cl2,_) } ->
+ CList.for_all2eq
+ (fun (_,c1,_,_) (_,c2,_,_) -> Names.Constant.equal c1 c2)
+ cl1 cl2
+ | _ -> false
+
+let rec uniq_seff = function
+ | [] -> []
+ | x :: xs -> x :: uniq_seff (List.filter (fun y -> not (equal_eff x y)) xs)
+(* The list of side effects is in reverse order (most recent first).
+ * To keep the "topological" order between effects we have to uniq-ize from
+ * the tail *)
+let uniq_seff l = List.rev (uniq_seff (List.rev l))
+
+let inline_side_effects env body ctx side_eff =
+ let handle_sideff (t,ctx,sl) { eff = se; from_env = mb } =
let cbl = match se with
| SEsubproof (c,cb,b) -> [c,cb,b]
| SEscheme (cl,_) -> List.map (fun (_,c,cb,b) -> c,cb,b) cl in
@@ -65,8 +83,8 @@ let handle_side_effects env body side_eff =
let rec sub_body c u b i x = match kind_of_term x with
| Const (c',u') when eq_constant c c' ->
Vars.subst_instance_constr u' b
- | _ -> map_constr_with_binders ((+) 1) (fun i x -> sub_body c u b i x) i x in
- let fix_body (c,cb,b) t =
+ | _ -> map_constr_with_binders ((+) 1) (sub_body c u b) i x in
+ let fix_body (c,cb,b) (t,ctx) =
match cb.const_body, b with
| Def b, _ ->
let b = Mod_subst.force_constr b in
@@ -74,37 +92,87 @@ let handle_side_effects env body side_eff =
if not poly then
let b_ty = Typeops.type_of_constant_type env cb.const_type in
let t = sub c 1 (Vars.lift 1 t) in
- mkLetIn (cname c, b, b_ty, t)
+ mkLetIn (cname c, b, b_ty, t),
+ Univ.ContextSet.union ctx
+ (Univ.ContextSet.of_context cb.const_universes)
else
let univs = cb.const_universes in
- sub_body c (Univ.UContext.instance univs) b 1 (Vars.lift 1 t)
+ sub_body c (Univ.UContext.instance univs) b 1 (Vars.lift 1 t), ctx
| OpaqueDef _, `Opaque (b,_) ->
let poly = cb.const_polymorphic in
if not poly then
let b_ty = Typeops.type_of_constant_type env cb.const_type in
let t = sub c 1 (Vars.lift 1 t) in
- mkApp (mkLambda (cname c, b_ty, t), [|b|])
+ mkApp (mkLambda (cname c, b_ty, t), [|b|]),
+ Univ.ContextSet.union ctx
+ (Univ.ContextSet.of_context cb.const_universes)
else
let univs = cb.const_universes in
- sub_body c (Univ.UContext.instance univs) b 1 (Vars.lift 1 t)
+ sub_body c (Univ.UContext.instance univs) b 1 (Vars.lift 1 t), ctx
| _ -> assert false
in
- List.fold_right fix_body cbl t
+ let t, ctx = List.fold_right fix_body cbl (t,ctx) in
+ t, ctx, (mb,List.length cbl) :: sl
in
(* CAVEAT: we assure a proper order *)
- Declareops.fold_side_effects handle_sideff body
- (Declareops.uniquize_side_effects side_eff)
+ List.fold_left handle_sideff (body,ctx,[]) (uniq_seff side_eff)
+
+(* Given the list of signatures of side effects, checks if they match.
+ * I.e. if they are ordered descendants of the current revstruct *)
+let check_signatures curmb sl =
+ let is_direct_ancestor (sl, curmb) (mb, how_many) =
+ match curmb with
+ | None -> None, None
+ | Some curmb ->
+ try
+ let mb = CEphemeron.get mb in
+ match sl with
+ | None -> sl, None
+ | Some n ->
+ if List.length mb >= how_many && CList.skipn how_many mb == curmb
+ then Some (n + how_many), Some mb
+ else None, None
+ with CEphemeron.InvalidKey -> None, None in
+ let sl, _ = List.fold_left is_direct_ancestor (Some 0,Some curmb) sl in
+ sl
+
+let skip_trusted_seff sl b e =
+ let rec aux sl b e acc =
+ let open Context.Rel.Declaration in
+ match sl, kind_of_term b with
+ | (None|Some 0), _ -> b, e, acc
+ | Some sl, LetIn (n,c,ty,bo) ->
+ aux (Some (sl-1)) bo
+ (Environ.push_rel (LocalDef (n,c,ty)) e) (`Let(n,c,ty)::acc)
+ | Some sl, App(hd,arg) ->
+ begin match kind_of_term hd with
+ | Lambda (n,ty,bo) ->
+ aux (Some (sl-1)) bo
+ (Environ.push_rel (LocalAssum (n,ty)) e) (`Cut(n,ty,arg)::acc)
+ | _ -> assert false
+ end
+ | _ -> assert false
+ in
+ aux sl b e []
+
+let rec unzip ctx j =
+ match ctx with
+ | [] -> j
+ | `Let (n,c,ty) :: ctx ->
+ unzip ctx { j with uj_val = mkLetIn (n,c,ty,j.uj_val) }
+ | `Cut (n,ty,arg) :: ctx ->
+ unzip ctx { j with uj_val = mkApp (mkLambda (n,ty,j.uj_val),arg) }
let hcons_j j =
{ uj_val = hcons_constr j.uj_val; uj_type = hcons_constr j.uj_type}
let feedback_completion_typecheck =
Option.iter (fun state_id -> Pp.feedback ~state_id Feedback.Complete)
-
-let infer_declaration env kn dcl =
+
+let infer_declaration ~trust env kn dcl =
match dcl with
| ParameterEntry (ctx,poly,(t,uctx),nl) ->
- let env = push_context uctx env in
+ let env = push_context ~strict:(not poly) uctx env in
let j = infer env t in
let abstract = poly && not (Option.is_empty kn) in
let usubst, univs = Univ.abstract_universes abstract uctx in
@@ -115,34 +183,41 @@ let infer_declaration env kn dcl =
| DefinitionEntry ({ const_entry_type = Some typ;
const_entry_opaque = true;
const_entry_polymorphic = false} as c) ->
- let env = push_context c.const_entry_universes env in
+ let env = push_context ~strict:true c.const_entry_universes env in
let { const_entry_body = body; const_entry_feedback = feedback_id } = c in
let tyj = infer_type env typ in
let proofterm =
- Future.chain ~greedy:true ~pure:true body (fun ((body, ctx),side_eff) ->
- let body = handle_side_effects env body side_eff in
- let env' = push_context_set ctx env in
- let j = infer env' body in
+ Future.chain ~greedy:true ~pure:true body (fun ((body,uctx),side_eff) ->
+ let body, uctx, signatures =
+ inline_side_effects env body uctx side_eff in
+ let valid_signatures = check_signatures trust signatures in
+ let env' = push_context_set uctx env in
+ let j =
+ let body,env',ectx = skip_trusted_seff valid_signatures body env' in
+ let j = infer env' body in
+ unzip ectx j in
let j = hcons_j j in
let subst = Univ.LMap.empty in
let _typ = constrain_type env' j c.const_entry_polymorphic subst
(`SomeWJ (typ,tyj)) in
feedback_completion_typecheck feedback_id;
- j.uj_val, ctx) in
+ j.uj_val, uctx) in
let def = OpaqueDef (Opaqueproof.create proofterm) in
def, RegularArity typ, None, c.const_entry_polymorphic,
c.const_entry_universes,
c.const_entry_inline_code, c.const_entry_secctx
| DefinitionEntry c ->
- let env = push_context c.const_entry_universes env in
let { const_entry_type = typ; const_entry_opaque = opaque } = c in
let { const_entry_body = body; const_entry_feedback = feedback_id } = c in
let (body, ctx), side_eff = Future.join body in
- assert(Univ.ContextSet.is_empty ctx);
- let body = handle_side_effects env body side_eff in
+ let univsctx = Univ.ContextSet.of_context c.const_entry_universes in
+ let body, ctx, _ = inline_side_effects env body
+ (Univ.ContextSet.union univsctx ctx) side_eff in
+ let env = push_context_set ~strict:(not c.const_entry_polymorphic) ctx env in
let abstract = c.const_entry_polymorphic && not (Option.is_empty kn) in
- let usubst, univs = Univ.abstract_universes abstract c.const_entry_universes in
+ let usubst, univs =
+ Univ.abstract_universes abstract (Univ.ContextSet.to_context ctx) in
let j = infer env body in
let typ = constrain_type env j c.const_entry_polymorphic usubst (map_option_typ typ) in
let def = hcons_constr (Vars.subst_univs_level_constr usubst j.uj_val) in
@@ -171,24 +246,30 @@ let infer_declaration env kn dcl =
let global_vars_set_constant_type env = function
| RegularArity t -> global_vars_set env t
| TemplateArity (ctx,_) ->
- Context.fold_rel_context
- (fold_rel_declaration
+ Context.Rel.fold_outside
+ (Context.Rel.Declaration.fold
(fun t c -> Id.Set.union (global_vars_set env t) c))
ctx ~init:Id.Set.empty
-let record_aux env s1 s2 =
+let record_aux env s_ty s_bo suggested_expr =
+ let open Context.Named.Declaration in
+ let in_ty = keep_hyps env s_ty in
let v =
String.concat " "
- (List.map (fun (id, _,_) -> Id.to_string id)
- (keep_hyps env (Id.Set.union s1 s2))) in
- Aux_file.record_in_aux "context_used" v
+ (CList.map_filter (fun decl ->
+ let id = get_id decl in
+ if List.exists (Id.equal id % get_id) in_ty then None
+ else Some (Id.to_string id))
+ (keep_hyps env s_bo)) in
+ Aux_file.record_in_aux "context_used" (v ^ ";" ^ suggested_expr)
-let suggest_proof_using = ref (fun _ _ _ _ _ -> ())
+let suggest_proof_using = ref (fun _ _ _ _ _ -> "")
let set_suggest_proof_using f = suggest_proof_using := f
let build_constant_declaration kn env (def,typ,proj,poly,univs,inline_code,ctx) =
+ let open Context.Named.Declaration in
let check declared inferred =
- let mk_set l = List.fold_right Id.Set.add (List.map pi1 l) Id.Set.empty in
+ let mk_set l = List.fold_right Id.Set.add (List.map get_id l) Id.Set.empty in
let inferred_set, declared_set = mk_set inferred, mk_set declared in
if not (Id.Set.subset inferred_set declared_set) then
let l = Id.Set.elements (Idset.diff inferred_set declared_set) in
@@ -198,9 +279,14 @@ let build_constant_declaration kn env (def,typ,proj,poly,univs,inline_code,ctx)
str " " ++ str (String.conjugate_verb_to_be n) ++
str " used but not declared:" ++
fnl () ++ pr_sequence Id.print (List.rev l) ++ str ".")) in
+ let sort evn l =
+ List.filter (fun decl ->
+ let id = get_id decl in
+ List.exists (Names.Id.equal id % get_id) l)
+ (named_context env) in
(* We try to postpone the computation of used section variables *)
let hyps, def =
- let context_ids = List.map pi1 (named_context env) in
+ let context_ids = List.map get_id (named_context env) in
match ctx with
| None when not (List.is_empty context_ids) ->
(* No declared section vars, and non-empty section context:
@@ -215,19 +301,21 @@ let build_constant_declaration kn env (def,typ,proj,poly,univs,inline_code,ctx)
(Opaqueproof.force_proof (opaque_tables env) lc) in
(* we force so that cst are added to the env immediately after *)
ignore(Opaqueproof.force_constraints (opaque_tables env) lc);
- !suggest_proof_using kn env vars ids_typ context_ids;
+ let expr =
+ !suggest_proof_using (Constant.to_string kn)
+ env vars ids_typ context_ids in
if !Flags.compilation_mode = Flags.BuildVo then
- record_aux env ids_typ vars;
+ record_aux env ids_typ vars expr;
vars
in
keep_hyps env (Idset.union ids_typ ids_def), def
| None ->
if !Flags.compilation_mode = Flags.BuildVo then
- record_aux env Id.Set.empty Id.Set.empty;
+ record_aux env Id.Set.empty Id.Set.empty "";
[], def (* Empty section context: no need to check *)
| Some declared ->
(* We use the declared set and chain a check of correctness *)
- declared,
+ sort env declared,
match def with
| Undef _ as x -> x (* nothing to check *)
| Def cs as x ->
@@ -243,33 +331,30 @@ let build_constant_declaration kn env (def,typ,proj,poly,univs,inline_code,ctx)
let inferred = keep_hyps env (Idset.union ids_typ ids_def) in
check declared inferred) lc) in
let tps =
- (* FIXME: incompleteness of the bytecode vm: we compile polymorphic
- constants like opaque definitions. *)
- if poly then Some (Cemitcodes.from_val Cemitcodes.BCconstant)
- else
- let res =
- match proj with
- | None -> compile_constant_body env def
- | Some pb ->
+ let res =
+ let comp_univs = if poly then Some univs else None in
+ match proj with
+ | None -> compile_constant_body env comp_univs def
+ | Some pb ->
(* The compilation of primitive projections is a bit tricky, because
they refer to themselves (the body of p looks like fun c =>
Proj(p,c)). We break the cycle by building an ad-hoc compilation
environment. A cleaner solution would be that kernel projections are
simply Proj(i,c) with i an int and c a constr, but we would have to
get rid of the compatibility layer. *)
- let cb =
- { const_hyps = hyps;
- const_body = def;
- const_type = typ;
- const_proj = proj;
- const_body_code = None;
- const_polymorphic = poly;
- const_universes = univs;
- const_inline_code = inline_code }
- in
- let env = add_constant kn cb env in
- compile_constant_body env def
- in Option.map Cemitcodes.from_val res
+ let cb =
+ { const_hyps = hyps;
+ const_body = def;
+ const_type = typ;
+ const_proj = proj;
+ const_body_code = None;
+ const_polymorphic = poly;
+ const_universes = univs;
+ const_inline_code = inline_code }
+ in
+ let env = add_constant kn cb env in
+ compile_constant_body env comp_univs def
+ in Option.map Cemitcodes.from_val res
in
{ const_hyps = hyps;
const_body = def;
@@ -282,8 +367,98 @@ let build_constant_declaration kn env (def,typ,proj,poly,univs,inline_code,ctx)
(*s Global and local constant declaration. *)
-let translate_constant env kn ce =
- build_constant_declaration kn env (infer_declaration env (Some kn) ce)
+let translate_constant mb env kn ce =
+ build_constant_declaration kn env
+ (infer_declaration ~trust:mb env (Some kn) ce)
+
+let constant_entry_of_side_effect cb u =
+ let pt =
+ match cb.const_body, u with
+ | OpaqueDef _, `Opaque (b, c) -> b, c
+ | Def b, `Nothing -> Mod_subst.force_constr b, Univ.ContextSet.empty
+ | _ -> assert false in
+ DefinitionEntry {
+ const_entry_body = Future.from_val (pt, []);
+ const_entry_secctx = None;
+ const_entry_feedback = None;
+ const_entry_type =
+ (match cb.const_type with RegularArity t -> Some t | _ -> None);
+ const_entry_polymorphic = cb.const_polymorphic;
+ const_entry_universes = cb.const_universes;
+ const_entry_opaque = Declareops.is_opaque cb;
+ const_entry_inline_code = cb.const_inline_code }
+;;
+
+let turn_direct (kn,cb,u,r as orig) =
+ match cb.const_body, u with
+ | OpaqueDef _, `Opaque (b,c) ->
+ let pt = Future.from_val (b,c) in
+ kn, { cb with const_body = OpaqueDef (Opaqueproof.create pt) }, u, r
+ | _ -> orig
+;;
+
+type side_effect_role =
+ | Subproof
+ | Schema of inductive * string
+
+type exported_side_effect =
+ constant * constant_body * side_effects constant_entry * side_effect_role
+
+let export_side_effects mb env ce =
+ match ce with
+ | ParameterEntry _ | ProjectionEntry _ -> [], ce
+ | DefinitionEntry c ->
+ let { const_entry_body = body } = c in
+ let _, eff = Future.force body in
+ let ce = DefinitionEntry { c with
+ const_entry_body = Future.chain ~greedy:true ~pure:true body
+ (fun (b_ctx, _) -> b_ctx, []) } in
+ let not_exists (c,_,_,_) =
+ try ignore(Environ.lookup_constant c env); false
+ with Not_found -> true in
+ let aux (acc,sl) { eff = se; from_env = mb } =
+ let cbl = match se with
+ | SEsubproof (c,cb,b) -> [c,cb,b,Subproof]
+ | SEscheme (cl,k) ->
+ List.map (fun (i,c,cb,b) -> c,cb,b,Schema(i,k)) cl in
+ let cbl = List.filter not_exists cbl in
+ if cbl = [] then acc, sl
+ else cbl :: acc, (mb,List.length cbl) :: sl in
+ let seff, signatures = List.fold_left aux ([],[]) (uniq_seff eff) in
+ let trusted = check_signatures mb signatures in
+ let push_seff env = function
+ | kn, cb, `Nothing, _ ->
+ let env = Environ.add_constant kn cb env in
+ if not cb.const_polymorphic then
+ Environ.push_context ~strict:true cb.const_universes env
+ else env
+ | kn, cb, `Opaque(_, ctx), _ ->
+ let env = Environ.add_constant kn cb env in
+ if not cb.const_polymorphic then
+ let env = Environ.push_context ~strict:true cb.const_universes env in
+ Environ.push_context_set ~strict:true ctx env
+ else env in
+ let rec translate_seff sl seff acc env =
+ match sl, seff with
+ | _, [] -> List.rev acc, ce
+ | (None | Some 0), cbs :: rest ->
+ let env, cbs =
+ List.fold_left (fun (env,cbs) (kn, ocb, u, r) ->
+ let ce = constant_entry_of_side_effect ocb u in
+ let cb = translate_constant mb env kn ce in
+ (push_seff env (kn, cb,`Nothing, Subproof),(kn,cb,ce,r) :: cbs))
+ (env,[]) cbs in
+ translate_seff sl rest (cbs @ acc) env
+ | Some sl, cbs :: rest ->
+ let cbs_len = List.length cbs in
+ let cbs = List.map turn_direct cbs in
+ let env = List.fold_left push_seff env cbs in
+ let ecbs = List.map (fun (kn,cb,u,r) ->
+ kn, cb, constant_entry_of_side_effect cb u, r) cbs in
+ translate_seff (Some (sl-cbs_len)) rest (ecbs @ acc) env
+ in
+ translate_seff trusted seff [] env
+;;
let translate_local_assum env t =
let j = infer env t in
@@ -293,18 +468,37 @@ let translate_local_assum env t =
let translate_recipe env kn r =
build_constant_declaration kn env (Cooking.cook_constant env r)
-let translate_local_def env id centry =
+let translate_local_def mb env id centry =
let def,typ,proj,poly,univs,inline_code,ctx =
- infer_declaration env None (DefinitionEntry centry) in
+ infer_declaration ~trust:mb env None (DefinitionEntry centry) in
let typ = type_of_constant_type env typ in
+ if ctx = None && !Flags.compilation_mode = Flags.BuildVo then begin
+ match def with
+ | Undef _ -> ()
+ | Def _ -> ()
+ | OpaqueDef lc ->
+ let open Context.Named.Declaration in
+ let context_ids = List.map get_id (named_context env) in
+ let ids_typ = global_vars_set env typ in
+ let ids_def = global_vars_set env
+ (Opaqueproof.force_proof (opaque_tables env) lc) in
+ let expr =
+ !suggest_proof_using (Id.to_string id)
+ env ids_def ids_typ context_ids in
+ record_aux env ids_typ ids_def expr
+ end;
def, typ, univs
(* Insertion of inductive types. *)
let translate_mind env kn mie = Indtypes.check_inductive env kn mie
-let handle_entry_side_effects env ce = { ce with
+let inline_entry_side_effects env ce = { ce with
const_entry_body = Future.chain ~greedy:true ~pure:true
ce.const_entry_body (fun ((body, ctx), side_eff) ->
- (handle_side_effects env body side_eff, ctx), Declareops.no_seff);
+ let body, ctx',_ = inline_side_effects env body ctx side_eff in
+ (body, ctx'), []);
}
+
+let inline_side_effects env body side_eff =
+ pi1 (inline_side_effects env body Univ.ContextSet.empty side_eff)
diff --git a/kernel/term_typing.mli b/kernel/term_typing.mli
index 1b54b1ea1e..fcd95576c0 100644
--- a/kernel/term_typing.mli
+++ b/kernel/term_typing.mli
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -12,23 +12,46 @@ open Environ
open Declarations
open Entries
-val translate_local_def : env -> Id.t -> definition_entry ->
+val translate_local_def : structure_body -> env -> Id.t -> side_effects definition_entry ->
constant_def * types * constant_universes
val translate_local_assum : env -> types -> types
-val mk_pure_proof : constr -> proof_output
+val mk_pure_proof : constr -> side_effects proof_output
-val handle_side_effects : env -> constr -> Declareops.side_effects -> constr
+val inline_side_effects : env -> constr -> side_effects -> constr
(** Returns the term where side effects have been turned into let-ins or beta
redexes. *)
-val handle_entry_side_effects : env -> definition_entry -> definition_entry
-(** Same as {!handle_side_effects} but applied to entries. Only modifies the
+val inline_entry_side_effects :
+ env -> side_effects definition_entry -> side_effects definition_entry
+(** Same as {!inline_side_effects} but applied to entries. Only modifies the
{!Entries.const_entry_body} field. It is meant to get a term out of a not
yet type checked proof. *)
-val translate_constant : env -> constant -> constant_entry -> constant_body
+val uniq_seff : side_effects -> side_effects
+
+val translate_constant :
+ structure_body -> env -> constant -> side_effects constant_entry ->
+ constant_body
+
+type side_effect_role =
+ | Subproof
+ | Schema of inductive * string
+
+type exported_side_effect =
+ constant * constant_body * side_effects constant_entry * side_effect_role
+
+(* Given a constant entry containing side effects it exports them (either
+ * by re-checking them or trusting them). Returns the constant bodies to
+ * be pushed in the safe_env by safe typing. The main constant entry
+ * needs to be translated as usual after this step. *)
+val export_side_effects :
+ structure_body -> env -> side_effects constant_entry ->
+ exported_side_effect list * side_effects constant_entry
+
+val constant_entry_of_side_effect :
+ constant_body -> seff_env -> side_effects constant_entry
val translate_mind :
env -> mutual_inductive -> mutual_inductive_entry -> mutual_inductive_body
@@ -37,11 +60,11 @@ val translate_recipe : env -> constant -> Cooking.recipe -> constant_body
(** Internal functions, mentioned here for debug purpose only *)
-val infer_declaration : env -> constant option ->
- constant_entry -> Cooking.result
+val infer_declaration : trust:structure_body -> env -> constant option ->
+ side_effects constant_entry -> Cooking.result
val build_constant_declaration :
constant -> env -> Cooking.result -> constant_body
val set_suggest_proof_using :
- (constant -> env -> Id.Set.t -> Id.Set.t -> Id.t list -> unit) -> unit
+ (string -> env -> Id.Set.t -> Id.Set.t -> Id.t list -> string) -> unit
diff --git a/kernel/type_errors.ml b/kernel/type_errors.ml
index 33c4172e5e..5071f0ad5d 100644
--- a/kernel/type_errors.ml
+++ b/kernel/type_errors.ml
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
diff --git a/kernel/type_errors.mli b/kernel/type_errors.mli
index 7b3d2f1c6d..0c3a952b8d 100644
--- a/kernel/type_errors.mli
+++ b/kernel/type_errors.mli
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
diff --git a/kernel/typeops.ml b/kernel/typeops.ml
index fe82d85d5d..0ea68e2bcc 100644
--- a/kernel/typeops.ml
+++ b/kernel/typeops.ml
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -12,13 +12,13 @@ open Names
open Univ
open Term
open Vars
-open Context
open Declarations
open Environ
open Entries
open Reduction
open Inductive
open Type_errors
+open Context.Rel.Declaration
let conv_leq l2r env x y = default_conv CUMUL ~l2r env x y
@@ -79,7 +79,7 @@ let judge_of_type u =
let judge_of_relative env n =
try
- let (_,_,typ) = lookup_rel n env in
+ let typ = get_type (lookup_rel n env) in
{ uj_val = mkRel n;
uj_type = lift n typ }
with Not_found ->
@@ -99,18 +99,20 @@ let judge_of_variable env id =
variables of the current env.
Order does not have to be checked assuming that all names are distinct *)
let check_hyps_inclusion env c sign =
- Context.fold_named_context
- (fun (id,b1,ty1) () ->
+ Context.Named.fold_outside
+ (fun d1 () ->
+ let open Context.Named.Declaration in
+ let id = get_id d1 in
try
- let (_,b2,ty2) = lookup_named id env in
- conv env ty2 ty1;
- (match b2,b1 with
- | None, None -> ()
- | None, Some _ ->
+ let d2 = lookup_named id env in
+ conv env (get_type d2) (get_type d1);
+ (match d2,d1 with
+ | LocalAssum _, LocalAssum _ -> ()
+ | LocalAssum _, LocalDef _ ->
(* This is wrong, because we don't know if the body is
needed or not for typechecking: *) ()
- | Some _, None -> raise NotConvertible
- | Some b2, Some b1 -> conv env b2 b1);
+ | LocalDef _, LocalAssum _ -> raise NotConvertible
+ | LocalDef (_,b2,_), LocalDef (_,b1,_) -> conv env b2 b1);
with Not_found | NotConvertible | Option.Heterogeneous ->
error_reference_variables env id c)
sign
@@ -125,19 +127,25 @@ let extract_level env p =
match kind_of_term c with Sort (Type u) -> Univ.Universe.level u | _ -> None
let extract_context_levels env l =
- let fold l (_, b, p) = match b with
- | None -> extract_level env p :: l
- | _ -> l
+ let fold l = function
+ | LocalAssum (_,p) -> extract_level env p :: l
+ | LocalDef _ -> l
in
List.fold_left fold [] l
let make_polymorphic_if_constant_for_ind env {uj_val = c; uj_type = t} =
let params, ccl = dest_prod_assum env t in
match kind_of_term ccl with
- | Sort (Type u) when isInd (fst (decompose_app (whd_betadeltaiota env c))) ->
- let param_ccls = extract_context_levels env params in
- let s = { template_param_levels = param_ccls; template_level = u} in
- TemplateArity (params,s)
+ | Sort (Type u) ->
+ let ind, l = decompose_app (whd_betadeltaiota env c) in
+ if isInd ind && List.is_empty l then
+ let mis = lookup_mind_specif env (fst (destInd ind)) in
+ let nparams = Inductive.inductive_params mis in
+ let paramsl = CList.lastn nparams params in
+ let param_ccls = extract_context_levels env paramsl in
+ let s = { template_param_levels = param_ccls; template_level = u} in
+ TemplateArity (params,s)
+ else RegularArity t
| _ ->
RegularArity t
@@ -294,7 +302,7 @@ let judge_of_cast env cj k tj =
match k with
| VMcast ->
mkCast (cj.uj_val, k, expected_type),
- vm_conv CUMUL env cj.uj_type expected_type
+ Reduction.vm_conv CUMUL env cj.uj_type expected_type
| DEFAULTcast ->
mkCast (cj.uj_val, k, expected_type),
default_conv ~l2r:false CUMUL env cj.uj_type expected_type
@@ -304,7 +312,7 @@ let judge_of_cast env cj k tj =
| NATIVEcast ->
let sigma = Nativelambda.empty_evars in
mkCast (cj.uj_val, k, expected_type),
- native_conv CUMUL sigma env cj.uj_type expected_type
+ Nativeconv.native_conv CUMUL sigma env cj.uj_type expected_type
in
{ uj_val = c;
uj_type = expected_type }
@@ -453,13 +461,13 @@ let rec execute env cstr =
| Lambda (name,c1,c2) ->
let varj = execute_type env c1 in
- let env1 = push_rel (name,None,varj.utj_val) env in
+ let env1 = push_rel (LocalAssum (name,varj.utj_val)) env in
let j' = execute env1 c2 in
judge_of_abstraction env name varj j'
| Prod (name,c1,c2) ->
let varj = execute_type env c1 in
- let env1 = push_rel (name,None,varj.utj_val) env in
+ let env1 = push_rel (LocalAssum (name,varj.utj_val)) env in
let varj' = execute_type env1 c2 in
judge_of_product env name varj varj'
@@ -467,11 +475,11 @@ let rec execute env cstr =
let j1 = execute env c1 in
let j2 = execute_type env c2 in
let _ = judge_of_cast env j1 DEFAULTcast j2 in
- let env1 = push_rel (name,Some j1.uj_val,j2.utj_val) env in
+ let env1 = push_rel (LocalDef (name,j1.uj_val,j2.utj_val)) env in
let j' = execute env1 c3 in
judge_of_letin env name j1 j2 j'
- | Cast (c,k, t) ->
+ | Cast (c,k,t) ->
let cj = execute env c in
let tj = execute_type env t in
judge_of_cast env cj k tj
@@ -543,18 +551,18 @@ let infer_v env cv =
(* Typing of several terms. *)
let infer_local_decl env id = function
- | LocalDef c ->
+ | LocalDefEntry c ->
let j = infer env c in
- (Name id, Some j.uj_val, j.uj_type)
- | LocalAssum c ->
+ LocalDef (Name id, j.uj_val, j.uj_type)
+ | LocalAssumEntry c ->
let j = infer env c in
- (Name id, None, assumption_of_judgment env j)
+ LocalAssum (Name id, assumption_of_judgment env j)
let infer_local_decls env decls =
let rec inferec env = function
| (id, d) :: l ->
let (env, l) = inferec env l in
let d = infer_local_decl env id d in
- (push_rel d env, add_rel_decl d l)
- | [] -> (env, empty_rel_context) in
+ (push_rel d env, Context.Rel.add d l)
+ | [] -> (env, Context.Rel.empty) in
inferec env decls
diff --git a/kernel/typeops.mli b/kernel/typeops.mli
index 010b2b6f03..2112284ea6 100644
--- a/kernel/typeops.mli
+++ b/kernel/typeops.mli
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -9,7 +9,6 @@
open Names
open Univ
open Term
-open Context
open Environ
open Entries
open Declarations
@@ -28,7 +27,7 @@ val infer_v : env -> constr array -> unsafe_judgment array
val infer_type : env -> types -> unsafe_type_judgment
val infer_local_decls :
- env -> (Id.t * local_entry) list -> (env * rel_context)
+ env -> (Id.t * local_entry) list -> (env * Context.Rel.t)
(** {6 Basic operations of the typing machine. } *)
@@ -128,4 +127,4 @@ val make_polymorphic_if_constant_for_ind : env -> unsafe_judgment ->
constant_type
(** Check that hyps are included in env and fails with error otherwise *)
-val check_hyps_inclusion : env -> constr -> section_context -> unit
+val check_hyps_inclusion : env -> constr -> Context.section_context -> unit
diff --git a/kernel/uGraph.ml b/kernel/uGraph.ml
new file mode 100644
index 0000000000..00883ddd84
--- /dev/null
+++ b/kernel/uGraph.ml
@@ -0,0 +1,906 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Pp
+open Errors
+open Util
+open Univ
+
+(* Created in Caml by Gérard Huet for CoC 4.8 [Dec 1988] *)
+(* Functional code by Jean-Christophe Filliâtre for Coq V7.0 [1999] *)
+(* Extension with algebraic universes by HH for Coq V7.0 [Sep 2001] *)
+(* Additional support for sort-polymorphic inductive types by HH [Mar 2006] *)
+(* Support for universe polymorphism by MS [2014] *)
+
+(* Revisions by Bruno Barras, Hugo Herbelin, Pierre Letouzey, Matthieu
+ Sozeau, Pierre-Marie Pédrot, Jacques-Henri Jourdan *)
+
+let error_inconsistency o u v (p:explanation option) =
+ raise (UniverseInconsistency (o,Universe.make u,Universe.make v,p))
+
+(* Universes are stratified by a partial ordering $\le$.
+ Let $\~{}$ be the associated equivalence. We also have a strict ordering
+ $<$ between equivalence classes, and we maintain that $<$ is acyclic,
+ and contained in $\le$ in the sense that $[U]<[V]$ implies $U\le V$.
+
+ At every moment, we have a finite number of universes, and we
+ maintain the ordering in the presence of assertions $U<V$ and $U\le V$.
+
+ The equivalence $\~{}$ is represented by a tree structure, as in the
+ union-find algorithm. The assertions $<$ and $\le$ are represented by
+ adjacency lists.
+
+ We use the algorithm described in the paper:
+
+ Bender, M. A., Fineman, J. T., Gilbert, S., & Tarjan, R. E. (2011). A
+ new approach to incremental cycle detection and related
+ problems. arXiv preprint arXiv:1112.0784.
+
+ *)
+
+open Universe
+
+module UMap = LMap
+
+type status = NoMark | Visited | WeakVisited | ToMerge
+
+(* Comparison on this type is pointer equality *)
+type canonical_node =
+ { univ: Level.t;
+ ltle: bool UMap.t; (* true: strict (lt) constraint.
+ false: weak (le) constraint. *)
+ gtge: LSet.t;
+ rank : int;
+ klvl: int;
+ ilvl: int;
+ mutable status: status
+ }
+
+let big_rank = 1000000
+
+(* A Level.t is either an alias for another one, or a canonical one,
+ for which we know the universes that are above *)
+
+type univ_entry =
+ Canonical of canonical_node
+ | Equiv of Level.t
+
+type universes =
+ { entries : univ_entry UMap.t;
+ index : int;
+ n_nodes : int; n_edges : int }
+
+type t = universes
+
+(** Used to cleanup universes if a traversal function is interrupted before it
+ has the opportunity to do it itself. *)
+let unsafe_cleanup_universes g =
+ let iter _ n = match n with
+ | Equiv _ -> ()
+ | Canonical n -> n.status <- NoMark
+ in
+ UMap.iter iter g.entries
+
+let rec cleanup_universes g =
+ try unsafe_cleanup_universes g
+ with e ->
+ (** The only way unsafe_cleanup_universes may raise an exception is when
+ a serious error (stack overflow, out of memory) occurs, or a signal is
+ sent. In this unlikely event, we relaunch the cleanup until we finally
+ succeed. *)
+ cleanup_universes g; raise e
+
+(* Every Level.t has a unique canonical arc representative *)
+
+(* Low-level function : makes u an alias for v.
+ Does not removes edges from n_edges, but decrements n_nodes.
+ u should be entered as canonical before. *)
+let enter_equiv g u v =
+ { entries =
+ UMap.modify u (fun _ a ->
+ match a with
+ | Canonical n ->
+ n.status <- NoMark;
+ Equiv v
+ | _ -> assert false) g.entries;
+ index = g.index;
+ n_nodes = g.n_nodes - 1;
+ n_edges = g.n_edges }
+
+(* Low-level function : changes data associated with a canonical node.
+ Resets the mutable fields in the old record, in order to avoid breaking
+ invariants for other users of this record.
+ n.univ should already been inserted as a canonical node. *)
+let change_node g n =
+ { g with entries =
+ UMap.modify n.univ
+ (fun _ a ->
+ match a with
+ | Canonical n' ->
+ n'.status <- NoMark;
+ Canonical n
+ | _ -> assert false)
+ g.entries }
+
+(* repr : universes -> Level.t -> canonical_node *)
+(* canonical representative : we follow the Equiv links *)
+let rec repr g u =
+ let a =
+ try UMap.find u g.entries
+ with Not_found -> anomaly ~label:"Univ.repr"
+ (str"Universe " ++ Level.pr u ++ str" undefined")
+ in
+ match a with
+ | Equiv v -> repr g v
+ | Canonical arc -> arc
+
+let get_set_arc g = repr g Level.set
+let is_set_arc u = Level.is_set u.univ
+let is_prop_arc u = Level.is_prop u.univ
+
+exception AlreadyDeclared
+
+(* Reindexes the given universe, using the next available index. *)
+let use_index g u =
+ let u = repr g u in
+ let g = change_node g { u with ilvl = g.index } in
+ assert (g.index > min_int);
+ { g with index = g.index - 1 }
+
+(* [safe_repr] is like [repr] but if the graph doesn't contain the
+ searched universe, we add it. *)
+let safe_repr g u =
+ let rec safe_repr_rec entries u =
+ match UMap.find u entries with
+ | Equiv v -> safe_repr_rec entries v
+ | Canonical arc -> arc
+ in
+ try g, safe_repr_rec g.entries u
+ with Not_found ->
+ let can =
+ { univ = u;
+ ltle = UMap.empty; gtge = LSet.empty;
+ rank = if Level.is_small u then big_rank else 0;
+ klvl = 0; ilvl = 0;
+ status = NoMark }
+ in
+ let g = { g with
+ entries = UMap.add u (Canonical can) g.entries;
+ n_nodes = g.n_nodes + 1 }
+ in
+ let g = use_index g u in
+ g, repr g u
+
+(* Returns 1 if u is higher than v in topological order.
+ -1 lower
+ 0 if u = v *)
+let topo_compare u v =
+ if u.klvl > v.klvl then 1
+ else if u.klvl < v.klvl then -1
+ else if u.ilvl > v.ilvl then 1
+ else if u.ilvl < v.ilvl then -1
+ else (assert (u==v); 0)
+
+(* Checks most of the invariants of the graph. For debugging purposes. *)
+let check_universes_invariants g =
+ let n_edges = ref 0 in
+ let n_nodes = ref 0 in
+ UMap.iter (fun l u ->
+ match u with
+ | Canonical u ->
+ UMap.iter (fun v strict ->
+ incr n_edges;
+ let v = repr g v in
+ assert (topo_compare u v = -1);
+ if u.klvl = v.klvl then
+ assert (LSet.mem u.univ v.gtge ||
+ LSet.exists (fun l -> u == repr g l) v.gtge))
+ u.ltle;
+ LSet.iter (fun v ->
+ let v = repr g v in
+ assert (v.klvl = u.klvl &&
+ (UMap.mem u.univ v.ltle ||
+ UMap.exists (fun l _ -> u == repr g l) v.ltle))
+ ) u.gtge;
+ assert (u.status = NoMark);
+ assert (Level.equal l u.univ);
+ assert (u.ilvl > g.index);
+ assert (not (UMap.mem u.univ u.ltle));
+ incr n_nodes
+ | Equiv _ -> assert (not (Level.is_small l)))
+ g.entries;
+ assert (!n_edges = g.n_edges);
+ assert (!n_nodes = g.n_nodes)
+
+let clean_ltle g ltle =
+ UMap.fold (fun u strict acc ->
+ let uu = (repr g u).univ in
+ if Level.equal uu u then acc
+ else (
+ let acc = UMap.remove u (fst acc) in
+ if not strict && UMap.mem uu acc then (acc, true)
+ else (UMap.add uu strict acc, true)))
+ ltle (ltle, false)
+
+let clean_gtge g gtge =
+ LSet.fold (fun u acc ->
+ let uu = (repr g u).univ in
+ if Level.equal uu u then acc
+ else LSet.add uu (LSet.remove u (fst acc)), true)
+ gtge (gtge, false)
+
+(* [get_ltle] and [get_gtge] return ltle and gtge arcs.
+ Moreover, if one of these lists is dirty (e.g. points to a
+ non-canonical node), these functions clean this node in the
+ graph by removing some duplicate edges *)
+let get_ltle g u =
+ let ltle, chgt_ltle = clean_ltle g u.ltle in
+ if not chgt_ltle then u.ltle, u, g
+ else
+ let sz = UMap.cardinal u.ltle in
+ let sz2 = UMap.cardinal ltle in
+ let u = { u with ltle } in
+ let g = change_node g u in
+ let g = { g with n_edges = g.n_edges + sz2 - sz } in
+ u.ltle, u, g
+
+let get_gtge g u =
+ let gtge, chgt_gtge = clean_gtge g u.gtge in
+ if not chgt_gtge then u.gtge, u, g
+ else
+ let u = { u with gtge } in
+ let g = change_node g u in
+ u.gtge, u, g
+
+(* [revert_graph] rollbacks the changes made to mutable fields in
+ nodes in the graph.
+ [to_revert] contains the touched nodes. *)
+let revert_graph to_revert g =
+ List.iter (fun t ->
+ match UMap.find t g.entries with
+ | Equiv _ -> ()
+ | Canonical t ->
+ t.status <- NoMark) to_revert
+
+exception AbortBackward of universes
+exception CycleDetected
+
+(* Implementation of the algorithm described in § 5.1 of the following paper:
+
+ Bender, M. A., Fineman, J. T., Gilbert, S., & Tarjan, R. E. (2011). A
+ new approach to incremental cycle detection and related
+ problems. arXiv preprint arXiv:1112.0784.
+
+ The "STEP X" comments contained in this file refers to the
+ corresponding step numbers of the algorithm described in Section
+ 5.1 of this paper. *)
+
+(* [delta] is the timeout for backward search. It might be
+ useful to tune a multiplicative constant. *)
+let get_delta g =
+ int_of_float
+ (min (float_of_int g.n_edges ** 0.5)
+ (float_of_int g.n_nodes ** (2./.3.)))
+
+let rec backward_traverse to_revert b_traversed count g x =
+ let x = repr g x in
+ let count = count - 1 in
+ if count < 0 then begin
+ revert_graph to_revert g;
+ raise (AbortBackward g)
+ end;
+ if x.status = NoMark then begin
+ x.status <- Visited;
+ let to_revert = x.univ::to_revert in
+ let gtge, x, g = get_gtge g x in
+ let to_revert, b_traversed, count, g =
+ LSet.fold (fun y (to_revert, b_traversed, count, g) ->
+ backward_traverse to_revert b_traversed count g y)
+ gtge (to_revert, b_traversed, count, g)
+ in
+ to_revert, x.univ::b_traversed, count, g
+ end
+ else to_revert, b_traversed, count, g
+
+let rec forward_traverse f_traversed g v_klvl x y =
+ let y = repr g y in
+ if y.klvl < v_klvl then begin
+ let y = { y with klvl = v_klvl;
+ gtge = if x == y then LSet.empty
+ else LSet.singleton x.univ }
+ in
+ let g = change_node g y in
+ let ltle, y, g = get_ltle g y in
+ let f_traversed, g =
+ UMap.fold (fun z _ (f_traversed, g) ->
+ forward_traverse f_traversed g v_klvl y z)
+ ltle (f_traversed, g)
+ in
+ y.univ::f_traversed, g
+ end else if y.klvl = v_klvl && x != y then
+ let g = change_node g
+ { y with gtge = LSet.add x.univ y.gtge } in
+ f_traversed, g
+ else f_traversed, g
+
+let rec find_to_merge to_revert g x v =
+ let x = repr g x in
+ match x.status with
+ | Visited -> false, to_revert | ToMerge -> true, to_revert
+ | NoMark ->
+ let to_revert = x::to_revert in
+ if Level.equal x.univ v then
+ begin x.status <- ToMerge; true, to_revert end
+ else
+ begin
+ let merge, to_revert = LSet.fold
+ (fun y (merge, to_revert) ->
+ let merge', to_revert = find_to_merge to_revert g y v in
+ merge' || merge, to_revert) x.gtge (false, to_revert)
+ in
+ x.status <- if merge then ToMerge else Visited;
+ merge, to_revert
+ end
+ | _ -> assert false
+
+let get_new_edges g to_merge =
+ (* Computing edge sets. *)
+ let to_merge_lvl =
+ List.fold_left (fun acc u -> UMap.add u.univ u acc)
+ UMap.empty to_merge
+ in
+ let ltle =
+ UMap.fold (fun _ n acc ->
+ UMap.merge (fun _ strict1 strict2 ->
+ match strict1, strict2 with
+ | Some true, _ | _, Some true -> Some true
+ | _, _ -> Some false)
+ acc n.ltle)
+ to_merge_lvl UMap.empty
+ in
+ let ltle, _ = clean_ltle g ltle in
+ let ltle =
+ UMap.merge (fun _ a strict ->
+ match a, strict with
+ | Some _, Some true ->
+ (* There is a lt edge inside the new component. This is a
+ "bad cycle". *)
+ raise CycleDetected
+ | Some _, Some false -> None
+ | _, _ -> strict
+ ) to_merge_lvl ltle
+ in
+ let gtge =
+ UMap.fold (fun _ n acc -> LSet.union acc n.gtge)
+ to_merge_lvl LSet.empty
+ in
+ let gtge, _ = clean_gtge g gtge in
+ let gtge = LSet.diff gtge (UMap.domain to_merge_lvl) in
+ (ltle, gtge)
+
+
+let reorder g u v =
+ (* STEP 2: backward search in the k-level of u. *)
+ let delta = get_delta g in
+
+ (* [v_klvl] is the chosen future level for u, v and all
+ traversed nodes. *)
+ let b_traversed, v_klvl, g =
+ try
+ let to_revert, b_traversed, _, g = backward_traverse [] [] delta g u in
+ revert_graph to_revert g;
+ let v_klvl = (repr g u).klvl in
+ b_traversed, v_klvl, g
+ with AbortBackward g ->
+ (* Backward search was too long, use the next k-level. *)
+ let v_klvl = (repr g u).klvl + 1 in
+ [], v_klvl, g
+ in
+ let f_traversed, g =
+ (* STEP 3: forward search. Contrary to what is described in
+ the paper, we do not test whether v_klvl = u.klvl nor we assign
+ v_klvl to v.klvl. Indeed, the first call to forward_traverse
+ will do all that. *)
+ forward_traverse [] g v_klvl (repr g v) v
+ in
+
+ (* STEP 4: merge nodes if needed. *)
+ let to_merge, b_reindex, f_reindex =
+ if (repr g u).klvl = v_klvl then
+ begin
+ let merge, to_revert = find_to_merge [] g u v in
+ let r =
+ if merge then
+ List.filter (fun u -> u.status = ToMerge) to_revert,
+ List.filter (fun u -> (repr g u).status <> ToMerge) b_traversed,
+ List.filter (fun u -> (repr g u).status <> ToMerge) f_traversed
+ else [], b_traversed, f_traversed
+ in
+ List.iter (fun u -> u.status <- NoMark) to_revert;
+ r
+ end
+ else [], b_traversed, f_traversed
+ in
+ let to_reindex, g =
+ match to_merge with
+ | [] -> List.rev_append f_reindex b_reindex, g
+ | n0::q0 ->
+ (* Computing new root. *)
+ let root, rank_rest =
+ List.fold_left (fun ((best, rank_rest) as acc) n ->
+ if n.rank >= best.rank then n, best.rank else acc)
+ (n0, min_int) q0
+ in
+ let ltle, gtge = get_new_edges g to_merge in
+ (* Inserting the new root. *)
+ let g = change_node g
+ { root with ltle; gtge;
+ rank = max root.rank (rank_rest + 1); }
+ in
+
+ (* Inserting shortcuts for old nodes. *)
+ let g = List.fold_left (fun g n ->
+ if Level.equal n.univ root.univ then g else enter_equiv g n.univ root.univ)
+ g to_merge
+ in
+
+ (* Updating g.n_edges *)
+ let oldsz =
+ List.fold_left (fun sz u -> sz+UMap.cardinal u.ltle)
+ 0 to_merge
+ in
+ let sz = UMap.cardinal ltle in
+ let g = { g with n_edges = g.n_edges + sz - oldsz } in
+
+ (* Not clear in the paper: we have to put the newly
+ created component just between B and F. *)
+ List.rev_append f_reindex (root.univ::b_reindex), g
+
+ in
+
+ (* STEP 5: reindex traversed nodes. *)
+ List.fold_left use_index g to_reindex
+
+(* Assumes [u] and [v] are already in the graph. *)
+(* Does NOT assume that ucan != vcan. *)
+let insert_edge strict ucan vcan g =
+ try
+ let u = ucan.univ and v = vcan.univ in
+ (* STEP 1: do we need to reorder nodes ? *)
+ let g = if topo_compare ucan vcan <= 0 then g else reorder g u v in
+
+ (* STEP 6: insert the new edge in the graph. *)
+ let u = repr g u in
+ let v = repr g v in
+ if u == v then
+ if strict then raise CycleDetected else g
+ else
+ let g =
+ try let oldstrict = UMap.find v.univ u.ltle in
+ if strict && not oldstrict then
+ change_node g { u with ltle = UMap.add v.univ true u.ltle }
+ else g
+ with Not_found ->
+ { (change_node g { u with ltle = UMap.add v.univ strict u.ltle })
+ with n_edges = g.n_edges + 1 }
+ in
+ if u.klvl <> v.klvl || LSet.mem u.univ v.gtge then g
+ else
+ let v = { v with gtge = LSet.add u.univ v.gtge } in
+ change_node g v
+ with
+ | CycleDetected as e -> raise e
+ | e ->
+ (** Unlikely event: fatal error or signal *)
+ let () = cleanup_universes g in
+ raise e
+
+let add_universe vlev strict g =
+ try
+ let _arcv = UMap.find vlev g.entries in
+ raise AlreadyDeclared
+ with Not_found ->
+ assert (g.index > min_int);
+ let v = {
+ univ = vlev;
+ ltle = LMap.empty;
+ gtge = LSet.empty;
+ rank = 0;
+ klvl = 0;
+ ilvl = g.index;
+ status = NoMark;
+ }
+ in
+ let entries = UMap.add vlev (Canonical v) g.entries in
+ let g = { entries; index = g.index - 1; n_nodes = g.n_nodes + 1; n_edges = g.n_edges } in
+ insert_edge strict (get_set_arc g) v g
+
+exception Found_explanation of explanation
+
+let get_explanation strict u v g =
+ let v = repr g v in
+ let visited_strict = ref UMap.empty in
+ let rec traverse strict u =
+ if u == v then
+ if strict then None else Some []
+ else if topo_compare u v = 1 then None
+ else
+ let visited =
+ try not (UMap.find u.univ !visited_strict) || strict
+ with Not_found -> false
+ in
+ if visited then None
+ else begin
+ visited_strict := UMap.add u.univ strict !visited_strict;
+ try
+ UMap.iter (fun u' strictu' ->
+ match traverse (strict && not strictu') (repr g u') with
+ | None -> ()
+ | Some exp ->
+ let typ = if strictu' then Lt else Le in
+ raise (Found_explanation ((typ, make u') :: exp)))
+ u.ltle;
+ None
+ with Found_explanation exp -> Some exp
+ end
+ in
+ let u = repr g u in
+ if u == v then [(Eq, make v.univ)]
+ else match traverse strict u with Some exp -> exp | None -> assert false
+
+let get_explanation strict u v g =
+ if !Flags.univ_print then Some (get_explanation strict u v g)
+ else None
+
+(* To compare two nodes, we simply do a forward search.
+ We implement two improvements:
+ - we ignore nodes that are higher than the destination;
+ - we do a BFS rather than a DFS because we expect to have a short
+ path (typically, the shortest path has length 1)
+*)
+exception Found of canonical_node list
+let search_path strict u v g =
+ let rec loop to_revert todo next_todo =
+ match todo, next_todo with
+ | [], [] -> to_revert (* No path found *)
+ | [], _ -> loop to_revert next_todo []
+ | (u, strict)::todo, _ ->
+ if u.status = Visited || (u.status = WeakVisited && strict)
+ then loop to_revert todo next_todo
+ else
+ let to_revert =
+ if u.status = NoMark then u::to_revert else to_revert
+ in
+ u.status <- if strict then WeakVisited else Visited;
+ if try UMap.find v.univ u.ltle || not strict
+ with Not_found -> false
+ then raise (Found to_revert)
+ else
+ begin
+ let next_todo =
+ UMap.fold (fun u strictu next_todo ->
+ let strict = not strictu && strict in
+ let u = repr g u in
+ if u == v && not strict then raise (Found to_revert)
+ else if topo_compare u v = 1 then next_todo
+ else (u, strict)::next_todo)
+ u.ltle next_todo
+ in
+ loop to_revert todo next_todo
+ end
+ in
+ if u == v then not strict
+ else
+ try
+ let res, to_revert =
+ try false, loop [] [u, strict] []
+ with Found to_revert -> true, to_revert
+ in
+ List.iter (fun u -> u.status <- NoMark) to_revert;
+ res
+ with e ->
+ (** Unlikely event: fatal error or signal *)
+ let () = cleanup_universes g in
+ raise e
+
+(** Uncomment to debug the cycle detection algorithm. *)
+(*let insert_edge strict ucan vcan g =
+ check_universes_invariants g;
+ let g = insert_edge strict ucan vcan g in
+ check_universes_invariants g;
+ let ucan = repr g ucan.univ in
+ let vcan = repr g vcan.univ in
+ assert (search_path strict ucan vcan g);
+ g*)
+
+(** First, checks on universe levels *)
+
+let check_equal g u v =
+ let arcu = repr g u and arcv = repr g v in
+ arcu == arcv
+
+let check_eq_level g u v = u == v || check_equal g u v
+
+let check_smaller g strict u v =
+ let arcu = repr g u and arcv = repr g v in
+ if strict then
+ search_path true arcu arcv g
+ else
+ is_prop_arc arcu
+ || (is_set_arc arcu && not (is_prop_arc arcv))
+ || search_path false arcu arcv g
+
+(** Then, checks on universes *)
+
+type 'a check_function = universes -> 'a -> 'a -> bool
+
+let check_equal_expr g x y =
+ x == y || (let (u, n) = x and (v, m) = y in
+ Int.equal n m && check_equal g u v)
+
+let check_eq_univs g l1 l2 =
+ let f x1 x2 = check_equal_expr g x1 x2 in
+ let exists x1 l = Universe.exists (fun x2 -> f x1 x2) l in
+ Universe.for_all (fun x1 -> exists x1 l2) l1
+ && Universe.for_all (fun x2 -> exists x2 l1) l2
+
+let check_eq g u v =
+ Universe.equal u v || check_eq_univs g u v
+
+let check_smaller_expr g (u,n) (v,m) =
+ let diff = n - m in
+ match diff with
+ | 0 -> check_smaller g false u v
+ | 1 -> check_smaller g true u v
+ | x when x < 0 -> check_smaller g false u v
+ | _ -> false
+
+let exists_bigger g ul l =
+ Universe.exists (fun ul' ->
+ check_smaller_expr g ul ul') l
+
+let real_check_leq g u v =
+ Universe.for_all (fun ul -> exists_bigger g ul v) u
+
+let check_leq g u v =
+ Universe.equal u v ||
+ is_type0m_univ u ||
+ check_eq_univs g u v || real_check_leq g u v
+
+(* enforce_univ_eq g u v will force u=v if possible, will fail otherwise *)
+
+let rec enforce_univ_eq u v g =
+ let ucan = repr g u in
+ let vcan = repr g v in
+ if topo_compare ucan vcan = 1 then enforce_univ_eq v u g
+ else
+ let g = insert_edge false ucan vcan g in (* Cannot fail *)
+ try insert_edge false vcan ucan g
+ with CycleDetected ->
+ error_inconsistency Eq v u (get_explanation true u v g)
+
+(* enforce_univ_leq g u v will force u<=v if possible, will fail otherwise *)
+let enforce_univ_leq u v g =
+ let ucan = repr g u in
+ let vcan = repr g v in
+ try insert_edge false ucan vcan g
+ with CycleDetected ->
+ error_inconsistency Le u v (get_explanation true v u g)
+
+(* enforce_univ_lt u v will force u<v if possible, will fail otherwise *)
+let enforce_univ_lt u v g =
+ let ucan = repr g u in
+ let vcan = repr g v in
+ try insert_edge true ucan vcan g
+ with CycleDetected ->
+ error_inconsistency Lt u v (get_explanation false v u g)
+
+let empty_universes =
+ let set_arc = Canonical {
+ univ = Level.set;
+ ltle = LMap.empty;
+ gtge = LSet.empty;
+ rank = big_rank;
+ klvl = 0;
+ ilvl = (-1);
+ status = NoMark;
+ } in
+ let prop_arc = Canonical {
+ univ = Level.prop;
+ ltle = LMap.empty;
+ gtge = LSet.empty;
+ rank = big_rank;
+ klvl = 0;
+ ilvl = 0;
+ status = NoMark;
+ } in
+ let entries = UMap.add Level.set set_arc (UMap.singleton Level.prop prop_arc) in
+ let empty = { entries; index = (-2); n_nodes = 2; n_edges = 0 } in
+ enforce_univ_lt Level.prop Level.set empty
+
+(* Prop = Set is forbidden here. *)
+let initial_universes = empty_universes
+
+let is_initial_universes g = UMap.equal (==) g.entries initial_universes.entries
+
+let enforce_constraint cst g =
+ match cst with
+ | (u,Lt,v) -> enforce_univ_lt u v g
+ | (u,Le,v) -> enforce_univ_leq u v g
+ | (u,Eq,v) -> enforce_univ_eq u v g
+
+let merge_constraints c g =
+ Constraint.fold enforce_constraint c g
+
+let check_constraint g (l,d,r) =
+ match d with
+ | Eq -> check_equal g l r
+ | Le -> check_smaller g false l r
+ | Lt -> check_smaller g true l r
+
+let check_constraints c g =
+ Constraint.for_all (check_constraint g) c
+
+(* Normalization *)
+
+(** [normalize_universes g] returns a graph where all edges point
+ directly to the canonical representent of their target. The output
+ graph should be equivalent to the input graph from a logical point
+ of view, but optimized. We maintain the invariant that the key of
+ a [Canonical] element is its own name, by keeping [Equiv] edges. *)
+let normalize_universes g =
+ let g =
+ { g with
+ entries = UMap.map (fun entry ->
+ match entry with
+ | Equiv u -> Equiv ((repr g u).univ)
+ | Canonical ucan -> Canonical { ucan with rank = 1 })
+ g.entries }
+ in
+ UMap.fold (fun _ u g ->
+ match u with
+ | Equiv u -> g
+ | Canonical u ->
+ let _, u, g = get_ltle g u in
+ let _, _, g = get_gtge g u in
+ g)
+ g.entries g
+
+let constraints_of_universes g =
+ let constraints_of u v acc =
+ match v with
+ | Canonical {univ=u; ltle} ->
+ UMap.fold (fun v strict acc->
+ let typ = if strict then Lt else Le in
+ Constraint.add (u,typ,v) acc) ltle acc
+ | Equiv v -> Constraint.add (u,Eq,v) acc
+ in
+ UMap.fold constraints_of g.entries Constraint.empty
+
+let constraints_of_universes g =
+ constraints_of_universes (normalize_universes g)
+
+(** [sort_universes g] builds a totally ordered universe graph. The
+ output graph should imply the input graph (and the implication
+ will be strict most of the time), but is not necessarily minimal.
+ Moreover, it adds levels [Type.n] to identify universes at level
+ n. An artificial constraint Set < Type.2 is added to ensure that
+ Type.n and small universes are not merged. Note: the result is
+ unspecified if the input graph already contains [Type.n] nodes
+ (calling a module Type is probably a bad idea anyway). *)
+let sort_universes g =
+ let cans =
+ UMap.fold (fun _ u l ->
+ match u with
+ | Equiv _ -> l
+ | Canonical can -> can :: l
+ ) g.entries []
+ in
+ let cans = List.sort topo_compare cans in
+ let lowest_levels =
+ UMap.mapi (fun u _ -> if Level.is_small u then 0 else 2)
+ (UMap.filter
+ (fun _ u -> match u with Equiv _ -> false | Canonical _ -> true)
+ g.entries)
+ in
+ let lowest_levels =
+ List.fold_left (fun lowest_levels can ->
+ let lvl = UMap.find can.univ lowest_levels in
+ UMap.fold (fun u' strict lowest_levels ->
+ let cost = if strict then 1 else 0 in
+ let u' = (repr g u').univ in
+ UMap.modify u' (fun _ lvl0 -> max lvl0 (lvl+cost)) lowest_levels)
+ can.ltle lowest_levels)
+ lowest_levels cans
+ in
+ let max_lvl = UMap.fold (fun _ a b -> max a b) lowest_levels 0 in
+ let mp = Names.DirPath.make [Names.Id.of_string "Type"] in
+ let types = Array.init (max_lvl + 1) (function
+ | 0 -> Level.prop
+ | 1 -> Level.set
+ | n -> Level.make mp (n-2))
+ in
+ let g = Array.fold_left (fun g u ->
+ let g, u = safe_repr g u in
+ change_node g { u with rank = big_rank }) g types
+ in
+ let g = if max_lvl >= 2 then enforce_univ_lt Level.set types.(2) g else g in
+ let g =
+ UMap.fold (fun u lvl g -> enforce_univ_eq u (types.(lvl)) g)
+ lowest_levels g
+ in
+ normalize_universes g
+
+(** Instances *)
+
+let check_eq_instances g t1 t2 =
+ let t1 = Instance.to_array t1 in
+ let t2 = Instance.to_array t2 in
+ t1 == t2 ||
+ (Int.equal (Array.length t1) (Array.length t2) &&
+ let rec aux i =
+ (Int.equal i (Array.length t1)) || (check_eq_level g t1.(i) t2.(i) && aux (i + 1))
+ in aux 0)
+
+(** Pretty-printing *)
+
+let pr_arc prl = function
+ | _, Canonical {univ=u; ltle} ->
+ if UMap.is_empty ltle then mt ()
+ else
+ prl u ++ str " " ++
+ v 0
+ (pr_sequence (fun (v, strict) ->
+ (if strict then str "< " else str "<= ") ++ prl v)
+ (UMap.bindings ltle)) ++
+ fnl ()
+ | u, Equiv v ->
+ prl u ++ str " = " ++ prl v ++ fnl ()
+
+let pr_universes prl g =
+ let graph = UMap.fold (fun u a l -> (u,a)::l) g.entries [] in
+ prlist (pr_arc prl) graph
+
+(* Dumping constraints to a file *)
+
+let dump_universes output g =
+ let dump_arc u = function
+ | Canonical {univ=u; ltle} ->
+ let u_str = Level.to_string u in
+ UMap.iter (fun v strict ->
+ let typ = if strict then Lt else Le in
+ output typ u_str (Level.to_string v)) ltle;
+ | Equiv v ->
+ output Eq (Level.to_string u) (Level.to_string v)
+ in
+ UMap.iter dump_arc g.entries
+
+(** Profiling *)
+
+let merge_constraints =
+ if Flags.profile then
+ let key = Profile.declare_profile "merge_constraints" in
+ Profile.profile2 key merge_constraints
+ else merge_constraints
+let check_constraints =
+ if Flags.profile then
+ let key = Profile.declare_profile "check_constraints" in
+ Profile.profile2 key check_constraints
+ else check_constraints
+
+let check_eq =
+ if Flags.profile then
+ let check_eq_key = Profile.declare_profile "check_eq" in
+ Profile.profile3 check_eq_key check_eq
+ else check_eq
+
+let check_leq =
+ if Flags.profile then
+ let check_leq_key = Profile.declare_profile "check_leq" in
+ Profile.profile3 check_leq_key check_leq
+ else check_leq
diff --git a/kernel/uGraph.mli b/kernel/uGraph.mli
new file mode 100644
index 0000000000..e95cf4d1cb
--- /dev/null
+++ b/kernel/uGraph.mli
@@ -0,0 +1,63 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Univ
+
+(** {6 Graphs of universes. } *)
+
+type t
+
+type universes = t
+
+type 'a check_function = universes -> 'a -> 'a -> bool
+val check_leq : universe check_function
+val check_eq : universe check_function
+
+(** The empty graph of universes *)
+val empty_universes : universes
+
+(** The initial graph of universes: Prop < Set *)
+val initial_universes : universes
+
+val is_initial_universes : universes -> bool
+
+val sort_universes : universes -> universes
+
+(** Adds a universe to the graph, ensuring it is >= or > Set.
+ @raises AlreadyDeclared if the level is already declared in the graph. *)
+
+exception AlreadyDeclared
+
+val add_universe : universe_level -> bool -> universes -> universes
+
+(** {6 ... } *)
+(** Merge of constraints in a universes graph.
+ The function [merge_constraints] merges a set of constraints in a given
+ universes graph. It raises the exception [UniverseInconsistency] if the
+ constraints are not satisfiable. *)
+
+val enforce_constraint : univ_constraint -> universes -> universes
+val merge_constraints : constraints -> universes -> universes
+
+val constraints_of_universes : universes -> constraints
+
+val check_constraint : universes -> univ_constraint -> bool
+val check_constraints : constraints -> universes -> bool
+
+val check_eq_instances : Instance.t check_function
+(** Check equality of instances w.r.t. a universe graph *)
+
+(** {6 Pretty-printing of universes. } *)
+
+val pr_universes : (Level.t -> Pp.std_ppcmds) -> universes -> Pp.std_ppcmds
+
+(** {6 Dumping to a file } *)
+
+val dump_universes :
+ (constraint_type -> string -> string -> unit) ->
+ universes -> unit
diff --git a/kernel/univ.ml b/kernel/univ.ml
index 336cdb653e..126f95f1f1 100644
--- a/kernel/univ.ml
+++ b/kernel/univ.ml
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -12,8 +12,8 @@
(* Additional support for sort-polymorphic inductive types by HH [Mar 2006] *)
(* Support for universe polymorphism by MS [2014] *)
-(* Revisions by Bruno Barras, Hugo Herbelin, Pierre Letouzey, Matthieu Sozeau,
- Pierre-Marie Pédrot *)
+(* Revisions by Bruno Barras, Hugo Herbelin, Pierre Letouzey, Matthieu
+ Sozeau, Pierre-Marie Pédrot *)
open Pp
open Errors
@@ -35,7 +35,7 @@ module type Hashconsed =
sig
type t
val hash : t -> int
- val equal : t -> t -> bool
+ val eq : t -> t -> bool
val hcons : t -> t
end
@@ -53,7 +53,7 @@ struct
type t = _t
type u = (M.t -> M.t)
let hash = function Nil -> 0 | Cons (_, h, _) -> h
- let equal l1 l2 = match l1, l2 with
+ let eq l1 l2 = match l1, l2 with
| Nil, Nil -> true
| Cons (x1, _, l1), Cons (x2, _, l2) -> x1 == x2 && l1 == l2
| _ -> false
@@ -135,12 +135,12 @@ module HList = struct
let rec remove x = function
| Nil -> nil
| Cons (y, _, l) ->
- if H.equal x y then l
+ if H.eq x y then l
else cons y (remove x l)
let rec mem x = function
| Nil -> false
- | Cons (y, _, l) -> H.equal x y || mem x l
+ | Cons (y, _, l) -> H.eq x y || mem x l
let rec compare cmp l1 l2 = match l1, l2 with
| Nil, Nil -> 0
@@ -251,7 +251,7 @@ module Level = struct
type _t = t
type t = _t
type u = unit
- let equal x y = x.hash == y.hash && RawLevel.hequal x.data y.data
+ let eq x y = x.hash == y.hash && RawLevel.hequal x.data y.data
let hash x = x.hash
let hashcons () x =
let data' = RawLevel.hcons x.data in
@@ -270,8 +270,10 @@ module Level = struct
let is_small x =
match data x with
| Level _ -> false
- | _ -> true
-
+ | Var _ -> false
+ | Prop -> true
+ | Set -> true
+
let is_prop x =
match data x with
| Prop -> true
@@ -398,7 +400,7 @@ struct
let hashcons hdir (b,n as x) =
let b' = hdir b in
if b' == b then x else (b',n)
- let equal l1 l2 =
+ let eq l1 l2 =
l1 == l2 ||
match l1,l2 with
| (b,n), (b',n') -> b == b' && n == n'
@@ -417,7 +419,7 @@ struct
let hcons =
Hashcons.simple_hcons H.generate H.hcons Level.hcons
let hash = ExprHash.hash
- let equal x y = x == y ||
+ let eq x y = x == y ||
(let (u,n) = x and (v,n') = y in
Int.equal n n' && Level.equal u v)
@@ -551,6 +553,10 @@ struct
| Cons (l, _, Nil) -> Expr.is_level l
| _ -> false
+ let rec is_levels l = match l with
+ | Cons (l, _, r) -> Expr.is_level l && is_levels r
+ | Nil -> true
+
let level l = match l with
| Cons (l, _, Nil) -> Expr.level l
| _ -> None
@@ -647,153 +653,6 @@ open Universe
let universe_level = Universe.level
-type status = Unset | SetLe | SetLt
-
-(* Comparison on this type is pointer equality *)
-type canonical_arc =
- { univ: Level.t;
- lt: Level.t list;
- le: Level.t list;
- rank : int;
- predicative : bool;
- mutable status : status;
- (** Guaranteed to be unset out of the [compare_neq] functions. It is used
- to do an imperative traversal of the graph, ensuring a O(1) check that
- a node has already been visited. Quite performance critical indeed. *)
- }
-
-let arc_is_le arc = match arc.status with
-| Unset -> false
-| SetLe | SetLt -> true
-
-let arc_is_lt arc = match arc.status with
-| Unset | SetLe -> false
-| SetLt -> true
-
-let terminal u = {univ=u; lt=[]; le=[]; rank=0; predicative=false; status = Unset}
-
-module UMap :
-sig
- type key = Level.t
- type +'a t
- val empty : 'a t
- val add : key -> 'a -> 'a t -> 'a t
- val find : key -> 'a t -> 'a
- val equal : ('a -> 'a -> bool) -> 'a t -> 'a t -> bool
- val fold : (key -> 'a -> 'b -> 'b) -> 'a t -> 'b -> 'b
- val iter : (key -> 'a -> unit) -> 'a t -> unit
- val mapi : (key -> 'a -> 'b) -> 'a t -> 'b t
-end = HMap.Make(Level)
-
-(* A Level.t is either an alias for another one, or a canonical one,
- for which we know the universes that are above *)
-
-type univ_entry =
- Canonical of canonical_arc
- | Equiv of Level.t
-
-type universes = univ_entry UMap.t
-
-(** Used to cleanup universes if a traversal function is interrupted before it
- has the opportunity to do it itself. *)
-let unsafe_cleanup_universes g =
- let iter _ arc = match arc with
- | Equiv _ -> ()
- | Canonical arc -> arc.status <- Unset
- in
- UMap.iter iter g
-
-let rec cleanup_universes g =
- try unsafe_cleanup_universes g
- with e ->
- (** The only way unsafe_cleanup_universes may raise an exception is when
- a serious error (stack overflow, out of memory) occurs, or a signal is
- sent. In this unlikely event, we relaunch the cleanup until we finally
- succeed. *)
- cleanup_universes g; raise e
-
-let enter_equiv_arc u v g =
- UMap.add u (Equiv v) g
-
-let enter_arc ca g =
- UMap.add ca.univ (Canonical ca) g
-
-(* Every Level.t has a unique canonical arc representative *)
-
-(* repr : universes -> Level.t -> canonical_arc *)
-(* canonical representative : we follow the Equiv links *)
-
-let rec repr g u =
- let a =
- try UMap.find u g
- with Not_found -> anomaly ~label:"Univ.repr"
- (str"Universe " ++ Level.pr u ++ str" undefined")
- in
- match a with
- | Equiv v -> repr g v
- | Canonical arc -> arc
-
-(* [safe_repr] also search for the canonical representative, but
- if the graph doesn't contain the searched universe, we add it. *)
-
-let safe_repr g u =
- let rec safe_repr_rec g u =
- match UMap.find u g with
- | Equiv v -> safe_repr_rec g v
- | Canonical arc -> arc
- in
- try g, safe_repr_rec g u
- with Not_found ->
- let can = terminal u in
- enter_arc can g, can
-
-(* reprleq : canonical_arc -> canonical_arc list *)
-(* All canonical arcv such that arcu<=arcv with arcv#arcu *)
-let reprleq g arcu =
- let rec searchrec w = function
- | [] -> w
- | v :: vl ->
- let arcv = repr g v in
- if List.memq arcv w || arcu==arcv then
- searchrec w vl
- else
- searchrec (arcv :: w) vl
- in
- searchrec [] arcu.le
-
-
-(* between : Level.t -> canonical_arc -> canonical_arc list *)
-(* between u v = { w | u<=w<=v, w canonical } *)
-(* between is the most costly operation *)
-
-let between g arcu arcv =
- (* good are all w | u <= w <= v *)
- (* bad are all w | u <= w ~<= v *)
- (* find good and bad nodes in {w | u <= w} *)
- (* explore b u = (b or "u is good") *)
- let rec explore ((good, bad, b) as input) arcu =
- if List.memq arcu good then
- (good, bad, true) (* b or true *)
- else if List.memq arcu bad then
- input (* (good, bad, b or false) *)
- else
- let leq = reprleq g arcu in
- (* is some universe >= u good ? *)
- let good, bad, b_leq =
- List.fold_left explore (good, bad, false) leq
- in
- if b_leq then
- arcu::good, bad, true (* b or true *)
- else
- good, arcu::bad, b (* b or false *)
- in
- let good,_,_ = explore ([arcv],[],false) arcu in
- good
-
-(* We assume compare(u,v) = LE with v canonical (see compare below).
- In this case List.hd(between g u v) = repr u
- Otherwise, between g u v = []
- *)
type constraint_type = Lt | Le | Eq
@@ -808,361 +667,6 @@ let constraint_type_ord c1 c2 = match c1, c2 with
| Eq, Eq -> 0
| Eq, _ -> 1
-(** [fast_compare_neq] : is [arcv] in the transitive upward closure of [arcu] ?
-
- In [strict] mode, we fully distinguish between LE and LT, while in
- non-strict mode, we simply answer LE for both situations.
-
- If [arcv] is encountered in a LT part, we could directly answer
- without visiting unneeded parts of this transitive closure.
- In [strict] mode, if [arcv] is encountered in a LE part, we could only
- change the default answer (1st arg [c]) from NLE to LE, since a strict
- constraint may appear later. During the recursive traversal,
- [lt_done] and [le_done] are universes we have already visited,
- they do not contain [arcv]. The 4rd arg is [(lt_todo,le_todo)],
- two lists of universes not yet considered, known to be above [arcu],
- strictly or not.
-
- We use depth-first search, but the presence of [arcv] in [new_lt]
- is checked as soon as possible : this seems to be slightly faster
- on a test.
-
- We do the traversal imperatively, setting the [status] flag on visited nodes.
- This ensures O(1) check, but it also requires unsetting the flag when leaving
- the function. Some special care has to be taken in order to ensure we do not
- recover a messed up graph at the end. This occurs in particular when the
- traversal raises an exception. Even though the code below is exception-free,
- OCaml may still raise random exceptions, essentially fatal exceptions or
- signal handlers. Therefore we ensure the cleanup by a catch-all clause. Note
- also that the use of an imperative solution does make this function
- thread-unsafe. For now we do not check universes in different threads, but if
- ever this is to be done, we would need some lock somewhere.
-
-*)
-
-let get_explanation strict g arcu arcv =
- (* [c] characterizes whether (and how) arcv has already been related
- to arcu among the lt_done,le_done universe *)
- let rec cmp c to_revert lt_todo le_todo = match lt_todo, le_todo with
- | [],[] -> (to_revert, c)
- | (arc,p)::lt_todo, le_todo ->
- if arc_is_lt arc then
- cmp c to_revert lt_todo le_todo
- else
- let rec find lt_todo lt le = match le with
- | [] ->
- begin match lt with
- | [] ->
- let () = arc.status <- SetLt in
- cmp c (arc :: to_revert) lt_todo le_todo
- | u :: lt ->
- let arc = repr g u in
- let p = (Lt, make u) :: p in
- if arc == arcv then
- if strict then (to_revert, p) else (to_revert, p)
- else find ((arc, p) :: lt_todo) lt le
- end
- | u :: le ->
- let arc = repr g u in
- let p = (Le, make u) :: p in
- if arc == arcv then
- if strict then (to_revert, p) else (to_revert, p)
- else find ((arc, p) :: lt_todo) lt le
- in
- find lt_todo arc.lt arc.le
- | [], (arc,p)::le_todo ->
- if arc == arcv then
- (* No need to continue inspecting universes above arc:
- if arcv is strictly above arc, then we would have a cycle.
- But we cannot answer LE yet, a stronger constraint may
- come later from [le_todo]. *)
- if strict then cmp p to_revert [] le_todo else (to_revert, p)
- else
- if arc_is_le arc then
- cmp c to_revert [] le_todo
- else
- let rec find lt_todo lt = match lt with
- | [] ->
- let fold accu u =
- let p = (Le, make u) :: p in
- let node = (repr g u, p) in
- node :: accu
- in
- let le_new = List.fold_left fold le_todo arc.le in
- let () = arc.status <- SetLe in
- cmp c (arc :: to_revert) lt_todo le_new
- | u :: lt ->
- let arc = repr g u in
- let p = (Lt, make u) :: p in
- if arc == arcv then
- if strict then (to_revert, p) else (to_revert, p)
- else find ((arc, p) :: lt_todo) lt
- in
- find [] arc.lt
- in
- try
- let (to_revert, c) = cmp [] [] [] [(arcu, [])] in
- (** Reset all the touched arcs. *)
- let () = List.iter (fun arc -> arc.status <- Unset) to_revert in
- List.rev c
- with e ->
- (** Unlikely event: fatal error or signal *)
- let () = cleanup_universes g in
- raise e
-
-let get_explanation strict g arcu arcv =
- if !Flags.univ_print then Some (get_explanation strict g arcu arcv)
- else None
-
-type fast_order = FastEQ | FastLT | FastLE | FastNLE
-
-let fast_compare_neq strict g arcu arcv =
- (* [c] characterizes whether arcv has already been related
- to arcu among the lt_done,le_done universe *)
- let rec cmp c to_revert lt_todo le_todo = match lt_todo, le_todo with
- | [],[] -> (to_revert, c)
- | arc::lt_todo, le_todo ->
- if arc_is_lt arc then
- cmp c to_revert lt_todo le_todo
- else
- let () = arc.status <- SetLt in
- process_lt c (arc :: to_revert) lt_todo le_todo arc.lt arc.le
- | [], arc::le_todo ->
- if arc == arcv then
- (* No need to continue inspecting universes above arc:
- if arcv is strictly above arc, then we would have a cycle.
- But we cannot answer LE yet, a stronger constraint may
- come later from [le_todo]. *)
- if strict then cmp FastLE to_revert [] le_todo else (to_revert, FastLE)
- else
- if arc_is_le arc then
- cmp c to_revert [] le_todo
- else
- let () = arc.status <- SetLe in
- process_le c (arc :: to_revert) [] le_todo arc.lt arc.le
-
- and process_lt c to_revert lt_todo le_todo lt le = match le with
- | [] ->
- begin match lt with
- | [] -> cmp c to_revert lt_todo le_todo
- | u :: lt ->
- let arc = repr g u in
- if arc == arcv then
- if strict then (to_revert, FastLT) else (to_revert, FastLE)
- else process_lt c to_revert (arc :: lt_todo) le_todo lt le
- end
- | u :: le ->
- let arc = repr g u in
- if arc == arcv then
- if strict then (to_revert, FastLT) else (to_revert, FastLE)
- else process_lt c to_revert (arc :: lt_todo) le_todo lt le
-
- and process_le c to_revert lt_todo le_todo lt le = match lt with
- | [] ->
- let fold accu u =
- let node = repr g u in
- node :: accu
- in
- let le_new = List.fold_left fold le_todo le in
- cmp c to_revert lt_todo le_new
- | u :: lt ->
- let arc = repr g u in
- if arc == arcv then
- if strict then (to_revert, FastLT) else (to_revert, FastLE)
- else process_le c to_revert (arc :: lt_todo) le_todo lt le
-
- in
-
- try
- let (to_revert, c) = cmp FastNLE [] [] [arcu] in
- (** Reset all the touched arcs. *)
- let () = List.iter (fun arc -> arc.status <- Unset) to_revert in
- c
- with e ->
- (** Unlikely event: fatal error or signal *)
- let () = cleanup_universes g in
- raise e
-
-let get_explanation_strict g arcu arcv = get_explanation true g arcu arcv
-
-let fast_compare g arcu arcv =
- if arcu == arcv then FastEQ else fast_compare_neq true g arcu arcv
-
-let is_leq g arcu arcv =
- arcu == arcv ||
- (match fast_compare_neq false g arcu arcv with
- | FastNLE -> false
- | (FastEQ|FastLE|FastLT) -> true)
-
-let is_lt g arcu arcv =
- if arcu == arcv then false
- else
- match fast_compare_neq true g arcu arcv with
- | FastLT -> true
- | (FastEQ|FastLE|FastNLE) -> false
-
-(* Invariants : compare(u,v) = EQ <=> compare(v,u) = EQ
- compare(u,v) = LT or LE => compare(v,u) = NLE
- compare(u,v) = NLE => compare(v,u) = NLE or LE or LT
-
- Adding u>=v is consistent iff compare(v,u) # LT
- and then it is redundant iff compare(u,v) # NLE
- Adding u>v is consistent iff compare(v,u) = NLE
- and then it is redundant iff compare(u,v) = LT *)
-
-(** * Universe checks [check_eq] and [check_leq], used in coqchk *)
-
-(** First, checks on universe levels *)
-
-let check_equal g u v =
- let g, arcu = safe_repr g u in
- let _, arcv = safe_repr g v in
- arcu == arcv
-
-let check_eq_level g u v = u == v || check_equal g u v
-
-let is_set_arc u = Level.is_set u.univ
-let is_prop_arc u = Level.is_prop u.univ
-let get_prop_arc g = snd (safe_repr g Level.prop)
-
-let check_smaller g strict u v =
- let g, arcu = safe_repr g u in
- let g, arcv = safe_repr g v in
- if strict then
- is_lt g arcu arcv
- else
- is_prop_arc arcu
- || (is_set_arc arcu && arcv.predicative)
- || is_leq g arcu arcv
-
-(** Then, checks on universes *)
-
-type 'a check_function = universes -> 'a -> 'a -> bool
-
-let check_equal_expr g x y =
- x == y || (let (u, n) = x and (v, m) = y in
- Int.equal n m && check_equal g u v)
-
-let check_eq_univs g l1 l2 =
- let f x1 x2 = check_equal_expr g x1 x2 in
- let exists x1 l = Huniv.exists (fun x2 -> f x1 x2) l in
- Huniv.for_all (fun x1 -> exists x1 l2) l1
- && Huniv.for_all (fun x2 -> exists x2 l1) l2
-
-let check_eq g u v =
- Universe.equal u v || check_eq_univs g u v
-
-let check_smaller_expr g (u,n) (v,m) =
- let diff = n - m in
- match diff with
- | 0 -> check_smaller g false u v
- | 1 -> check_smaller g true u v
- | x when x < 0 -> check_smaller g false u v
- | _ -> false
-
-let exists_bigger g ul l =
- Huniv.exists (fun ul' ->
- check_smaller_expr g ul ul') l
-
-let real_check_leq g u v =
- Huniv.for_all (fun ul -> exists_bigger g ul v) u
-
-let check_leq g u v =
- Universe.equal u v ||
- Universe.is_type0m u ||
- check_eq_univs g u v || real_check_leq g u v
-
-(** Enforcing new constraints : [setlt], [setleq], [merge], [merge_disc] *)
-
-(** To speed up tests of Set </<= i *)
-let set_predicative g arcv =
- enter_arc {arcv with predicative = true} g
-
-(* setlt : Level.t -> Level.t -> reason -> unit *)
-(* forces u > v *)
-(* this is normally an update of u in g rather than a creation. *)
-let setlt g arcu arcv =
- let arcu' = {arcu with lt=arcv.univ::arcu.lt} in
- let g =
- if is_set_arc arcu then set_predicative g arcv
- else g
- in
- enter_arc arcu' g, arcu'
-
-(* checks that non-redundant *)
-let setlt_if (g,arcu) v =
- let arcv = repr g v in
- if is_lt g arcu arcv then g, arcu
- else setlt g arcu arcv
-
-(* setleq : Level.t -> Level.t -> unit *)
-(* forces u >= v *)
-(* this is normally an update of u in g rather than a creation. *)
-let setleq g arcu arcv =
- let arcu' = {arcu with le=arcv.univ::arcu.le} in
- let g =
- if is_set_arc arcu' then
- set_predicative g arcv
- else g
- in
- enter_arc arcu' g, arcu'
-
-(* checks that non-redundant *)
-let setleq_if (g,arcu) v =
- let arcv = repr g v in
- if is_leq g arcu arcv then g, arcu
- else setleq g arcu arcv
-
-(* merge : Level.t -> Level.t -> unit *)
-(* we assume compare(u,v) = LE *)
-(* merge u v forces u ~ v with repr u as canonical repr *)
-let merge g arcu arcv =
- (* we find the arc with the biggest rank, and we redirect all others to it *)
- let arcu, g, v =
- let best_ranked (max_rank, old_max_rank, best_arc, rest) arc =
- if Level.is_small arc.univ || arc.rank >= max_rank
- then (arc.rank, max_rank, arc, best_arc::rest)
- else (max_rank, old_max_rank, best_arc, arc::rest)
- in
- match between g arcu arcv with
- | [] -> anomaly (str "Univ.between")
- | arc::rest ->
- let (max_rank, old_max_rank, best_arc, rest) =
- List.fold_left best_ranked (arc.rank, min_int, arc, []) rest in
- if max_rank > old_max_rank then best_arc, g, rest
- else begin
- (* one redirected node also has max_rank *)
- let arcu = {best_arc with rank = max_rank + 1} in
- arcu, enter_arc arcu g, rest
- end
- in
- let redirect (g,w,w') arcv =
- let g' = enter_equiv_arc arcv.univ arcu.univ g in
- (g',List.unionq arcv.lt w,arcv.le@w')
- in
- let (g',w,w') = List.fold_left redirect (g,[],[]) v in
- let g_arcu = (g',arcu) in
- let g_arcu = List.fold_left setlt_if g_arcu w in
- let g_arcu = List.fold_left setleq_if g_arcu w' in
- fst g_arcu
-
-(* merge_disc : Level.t -> Level.t -> unit *)
-(* we assume compare(u,v) = compare(v,u) = NLE *)
-(* merge_disc u v forces u ~ v with repr u as canonical repr *)
-let merge_disc g arc1 arc2 =
- let arcu, arcv = if arc1.rank < arc2.rank then arc2, arc1 else arc1, arc2 in
- let arcu, g =
- if not (Int.equal arc1.rank arc2.rank) then arcu, g
- else
- let arcu = {arcu with rank = succ arcu.rank} in
- arcu, enter_arc arcu g
- in
- let g' = enter_equiv_arc arcv.univ arcu.univ g in
- let g_arcu = (g',arcu) in
- let g_arcu = List.fold_left setlt_if g_arcu arcv.lt in
- let g_arcu = List.fold_left setleq_if g_arcu arcv.le in
- fst g_arcu
-
(* Universe inconsistency: error raised when trying to enforce a relation
that would create a cycle in the graph of universes. *)
@@ -1173,81 +677,10 @@ exception UniverseInconsistency of univ_inconsistency
let error_inconsistency o u v (p:explanation option) =
raise (UniverseInconsistency (o,make u,make v,p))
-(* enforc_univ_eq : Level.t -> Level.t -> unit *)
-(* enforc_univ_eq u v will force u=v if possible, will fail otherwise *)
-
-let enforce_univ_eq u v g =
- let g,arcu = safe_repr g u in
- let g,arcv = safe_repr g v in
- match fast_compare g arcu arcv with
- | FastEQ -> g
- | FastLT ->
- let p = get_explanation_strict g arcu arcv in
- error_inconsistency Eq v u p
- | FastLE -> merge g arcu arcv
- | FastNLE ->
- (match fast_compare g arcv arcu with
- | FastLT ->
- let p = get_explanation_strict g arcv arcu in
- error_inconsistency Eq u v p
- | FastLE -> merge g arcv arcu
- | FastNLE -> merge_disc g arcu arcv
- | FastEQ -> anomaly (Pp.str "Univ.compare"))
-
-(* enforce_univ_leq : Level.t -> Level.t -> unit *)
-(* enforce_univ_leq u v will force u<=v if possible, will fail otherwise *)
-let enforce_univ_leq u v g =
- let g,arcu = safe_repr g u in
- let g,arcv = safe_repr g v in
- if is_leq g arcu arcv then g
- else
- match fast_compare g arcv arcu with
- | FastLT ->
- let p = get_explanation_strict g arcv arcu in
- error_inconsistency Le u v p
- | FastLE -> merge g arcv arcu
- | FastNLE -> fst (setleq g arcu arcv)
- | FastEQ -> anomaly (Pp.str "Univ.compare")
-
-(* enforce_univ_lt u v will force u<v if possible, will fail otherwise *)
-let enforce_univ_lt u v g =
- let g,arcu = safe_repr g u in
- let g,arcv = safe_repr g v in
- match fast_compare g arcu arcv with
- | FastLT -> g
- | FastLE -> fst (setlt g arcu arcv)
- | FastEQ -> error_inconsistency Lt u v (Some [(Eq,make v)])
- | FastNLE ->
- match fast_compare_neq false g arcv arcu with
- FastNLE -> fst (setlt g arcu arcv)
- | FastEQ -> anomaly (Pp.str "Univ.compare")
- | (FastLE|FastLT) ->
- let p = get_explanation false g arcv arcu in
- error_inconsistency Lt u v p
-
-let empty_universes = UMap.empty
-
-(* Prop = Set is forbidden here. *)
-let initial_universes = enforce_univ_lt Level.prop Level.set UMap.empty
-
-let is_initial_universes g = UMap.equal (==) g initial_universes
-
-let add_universe vlev g =
- let v = terminal vlev in
- let proparc = get_prop_arc g in
- enter_arc {proparc with le=vlev::proparc.le}
- (enter_arc v g)
-
(* Constraints and sets of constraints. *)
type univ_constraint = Level.t * constraint_type * Level.t
-let enforce_constraint cst g =
- match cst with
- | (u,Lt,v) -> enforce_univ_lt u v g
- | (u,Le,v) -> enforce_univ_leq u v g
- | (u,Eq,v) -> enforce_univ_eq u v g
-
let pr_constraint_type op =
let op_str = match op with
| Lt -> " < "
@@ -1282,8 +715,6 @@ end
let empty_constraint = Constraint.empty
let union_constraint = Constraint.union
let eq_constraint = Constraint.equal
-let merge_constraints c g =
- Constraint.fold enforce_constraint c g
type constraints = Constraint.t
@@ -1293,7 +724,7 @@ module Hconstraint =
type t = univ_constraint
type u = universe_level -> universe_level
let hashcons hul (l1,k,l2) = (hul l1, k, hul l2)
- let equal (l1,k,l2) (l1',k',l2') =
+ let eq (l1,k,l2) (l1',k',l2') =
l1 == l1' && k == k' && l2 == l2'
let hash = Hashtbl.hash
end)
@@ -1305,7 +736,7 @@ module Hconstraints =
type u = univ_constraint -> univ_constraint
let hashcons huc s =
Constraint.fold (fun x -> Constraint.add (huc x)) s Constraint.empty
- let equal s s' =
+ let eq s s' =
List.for_all2eq (==)
(Constraint.elements s)
(Constraint.elements s')
@@ -1370,10 +801,12 @@ let check_univ_leq u v =
let enforce_leq u v c =
let open Universe.Huniv in
+ let rec aux acc v =
match v with
- | Cons (v, _, Nil) ->
- fold (fun u -> constraint_add_leq u v) u c
- | _ -> anomaly (Pp.str"A universe bound can only be a variable")
+ | Cons (v, _, l) ->
+ aux (fold (fun u -> constraint_add_leq u v) u c) l
+ | Nil -> acc
+ in aux c v
let enforce_leq u v c =
if check_univ_leq u v then c
@@ -1382,219 +815,12 @@ let enforce_leq u v c =
let enforce_leq_level u v c =
if Level.equal u v then c else Constraint.add (u,Le,v) c
-let check_constraint g (l,d,r) =
- match d with
- | Eq -> check_equal g l r
- | Le -> check_smaller g false l r
- | Lt -> check_smaller g true l r
-
-let check_constraints c g =
- Constraint.for_all (check_constraint g) c
-
let enforce_univ_constraint (u,d,v) =
match d with
| Eq -> enforce_eq u v
| Le -> enforce_leq u v
| Lt -> enforce_leq (super u) v
-(* Normalization *)
-
-let lookup_level u g =
- try Some (UMap.find u g) with Not_found -> None
-
-(** [normalize_universes g] returns a graph where all edges point
- directly to the canonical representent of their target. The output
- graph should be equivalent to the input graph from a logical point
- of view, but optimized. We maintain the invariant that the key of
- a [Canonical] element is its own name, by keeping [Equiv] edges
- (see the assertion)... I (Stéphane Glondu) am not sure if this
- plays a role in the rest of the module. *)
-let normalize_universes g =
- let rec visit u arc cache = match lookup_level u cache with
- | Some x -> x, cache
- | None -> match Lazy.force arc with
- | None ->
- u, UMap.add u u cache
- | Some (Canonical {univ=v; lt=_; le=_}) ->
- v, UMap.add u v cache
- | Some (Equiv v) ->
- let v, cache = visit v (lazy (lookup_level v g)) cache in
- v, UMap.add u v cache
- in
- let cache = UMap.fold
- (fun u arc cache -> snd (visit u (Lazy.lazy_from_val (Some arc)) cache))
- g UMap.empty
- in
- let repr x = UMap.find x cache in
- let lrepr us = List.fold_left
- (fun e x -> LSet.add (repr x) e) LSet.empty us
- in
- let canonicalize u = function
- | Equiv _ -> Equiv (repr u)
- | Canonical {univ=v; lt=lt; le=le; rank=rank} ->
- assert (u == v);
- (* avoid duplicates and self-loops *)
- let lt = lrepr lt and le = lrepr le in
- let le = LSet.filter
- (fun x -> x != u && not (LSet.mem x lt)) le
- in
- LSet.iter (fun x -> assert (x != u)) lt;
- Canonical {
- univ = v;
- lt = LSet.elements lt;
- le = LSet.elements le;
- rank = rank;
- predicative = false;
- status = Unset;
- }
- in
- UMap.mapi canonicalize g
-
-let constraints_of_universes g =
- let constraints_of u v acc =
- match v with
- | Canonical {univ=u; lt=lt; le=le} ->
- let acc = List.fold_left (fun acc v -> Constraint.add (u,Lt,v) acc) acc lt in
- let acc = List.fold_left (fun acc v -> Constraint.add (u,Le,v) acc) acc le in
- acc
- | Equiv v -> Constraint.add (u,Eq,v) acc
- in
- UMap.fold constraints_of g Constraint.empty
-
-let constraints_of_universes g =
- constraints_of_universes (normalize_universes g)
-
-(** Longest path algorithm. This is used to compute the minimal number of
- universes required if the only strict edge would be the Lt one. This
- algorithm assumes that the given universes constraints are a almost DAG, in
- the sense that there may be {Eq, Le}-cycles. This is OK for consistent
- universes, which is the only case where we use this algorithm. *)
-
-(** Adjacency graph *)
-type graph = constraint_type LMap.t LMap.t
-
-exception Connected
-
-(** Check connectedness *)
-let connected x y (g : graph) =
- let rec connected x target seen g =
- if Level.equal x target then raise Connected
- else if not (LSet.mem x seen) then
- let seen = LSet.add x seen in
- let fold z _ seen = connected z target seen g in
- let neighbours = try LMap.find x g with Not_found -> LMap.empty in
- LMap.fold fold neighbours seen
- else seen
- in
- try ignore(connected x y LSet.empty g); false with Connected -> true
-
-let add_edge x y v (g : graph) =
- try
- let neighbours = LMap.find x g in
- let neighbours = LMap.add y v neighbours in
- LMap.add x neighbours g
- with Not_found ->
- LMap.add x (LMap.singleton y v) g
-
-(** We want to keep the graph DAG. If adding an edge would cause a cycle, that
- would necessarily be an {Eq, Le}-cycle, otherwise there would have been a
- universe inconsistency. Therefore we may omit adding such a cycling edge
- without changing the compacted graph. *)
-let add_eq_edge x y v g = if connected y x g then g else add_edge x y v g
-
-(** Construct the DAG and its inverse at the same time. *)
-let make_graph g : (graph * graph) =
- let fold u arc accu = match arc with
- | Equiv v ->
- let (dir, rev) = accu in
- (add_eq_edge u v Eq dir, add_eq_edge v u Eq rev)
- | Canonical { univ; lt; le; } ->
- let () = assert (u == univ) in
- let fold_lt (dir, rev) v = (add_edge u v Lt dir, add_edge v u Lt rev) in
- let fold_le (dir, rev) v = (add_eq_edge u v Le dir, add_eq_edge v u Le rev) in
- (** Order is important : lt after le, because of the possible redundancy
- between [le] and [lt] in a canonical arc. This way, the [lt] constraint
- is the last one set, which is correct because it implies [le]. *)
- let accu = List.fold_left fold_le accu le in
- let accu = List.fold_left fold_lt accu lt in
- accu
- in
- UMap.fold fold g (LMap.empty, LMap.empty)
-
-(** Construct a topological order out of a DAG. *)
-let rec topological_fold u g rem seen accu =
- let is_seen =
- try
- let status = LMap.find u seen in
- assert status; (** If false, not a DAG! *)
- true
- with Not_found -> false
- in
- if not is_seen then
- let rem = LMap.remove u rem in
- let seen = LMap.add u false seen in
- let neighbours = try LMap.find u g with Not_found -> LMap.empty in
- let fold v _ (rem, seen, accu) = topological_fold v g rem seen accu in
- let (rem, seen, accu) = LMap.fold fold neighbours (rem, seen, accu) in
- (rem, LMap.add u true seen, u :: accu)
- else (rem, seen, accu)
-
-let rec topological g rem seen accu =
- let node = try Some (LMap.choose rem) with Not_found -> None in
- match node with
- | None -> accu
- | Some (u, _) ->
- let rem, seen, accu = topological_fold u g rem seen accu in
- topological g rem seen accu
-
-(** Compute the longest path from any vertex. *)
-let constraint_cost = function
-| Eq | Le -> 0
-| Lt -> 1
-
-(** This algorithm browses the graph in topological order, computing for each
- encountered node the length of the longest path leading to it. Should be
- O(|V|) or so (modulo map representation). *)
-let rec flatten_graph rem (rev : graph) map mx = match rem with
-| [] -> map, mx
-| u :: rem ->
- let prev = try LMap.find u rev with Not_found -> LMap.empty in
- let fold v cstr accu =
- let v_cost = LMap.find v map in
- max (v_cost + constraint_cost cstr) accu
- in
- let u_cost = LMap.fold fold prev 0 in
- let map = LMap.add u u_cost map in
- flatten_graph rem rev map (max mx u_cost)
-
-(** [sort_universes g] builds a map from universes in [g] to natural
- numbers. It outputs a graph containing equivalence edges from each
- level appearing in [g] to [Type.n], and [lt] edges between the
- [Type.n]s. The output graph should imply the input graph (and the
- [Type.n]s. The output graph should imply the input graph (and the
- implication will be strict most of the time), but is not
- necessarily minimal. Note: the result is unspecified if the input
- graph already contains [Type.n] nodes (calling a module Type is
- probably a bad idea anyway). *)
-let sort_universes orig =
- let (dir, rev) = make_graph orig in
- let order = topological dir dir LMap.empty [] in
- let compact, max = flatten_graph order rev LMap.empty 0 in
- let mp = Names.DirPath.make [Names.Id.of_string "Type"] in
- let types = Array.init (max + 1) (fun n -> Level.make mp n) in
- (** Old universes are made equal to [Type.n] *)
- let fold u level accu = UMap.add u (Equiv types.(level)) accu in
- let sorted = LMap.fold fold compact UMap.empty in
- (** Add all [Type.n] nodes *)
- let fold i accu u =
- if 0 < i then
- let pred = types.(i - 1) in
- let arc = {univ = u; lt = [pred]; le = []; rank = 0; predicative = false; status = Unset; } in
- UMap.add u (Canonical arc) accu
- else accu
- in
- Array.fold_left_i fold sorted types
-
(* Miscellaneous functions to remove or test local univ assumed to
occur in a universe *)
@@ -1650,7 +876,6 @@ module Instance : sig
val pr : (Level.t -> Pp.std_ppcmds) -> t -> Pp.std_ppcmds
val levels : t -> LSet.t
- val check_eq : t check_function
end =
struct
type t = Level.t array
@@ -1676,7 +901,7 @@ struct
a
end
- let equal t1 t2 =
+ let eq t1 t2 =
t1 == t2 ||
(Int.equal (Array.length t1) (Array.length t2) &&
let rec aux i =
@@ -1712,7 +937,9 @@ struct
else if Array.length y = 0 then x
else Array.append x y
- let of_array a = a
+ let of_array a =
+ assert(Array.for_all (fun x -> not (Level.is_prop x)) a);
+ a
let to_array a = a
@@ -1720,7 +947,7 @@ struct
let subst_fn fn t =
let t' = CArray.smartmap fn t in
- if t' == t then t else t'
+ if t' == t then t else of_array t'
let levels x = LSet.of_array x
@@ -1734,13 +961,6 @@ struct
(* Necessary as universe instances might come from different modules and
unmarshalling doesn't preserve sharing *))
- let check_eq g t1 t2 =
- t1 == t2 ||
- (Int.equal (Array.length t1) (Array.length t2) &&
- let rec aux i =
- (Int.equal i (Array.length t1)) || (check_eq_level g t1.(i) t2.(i) && aux (i + 1))
- in aux 0)
-
end
let enforce_eq_instances x y =
@@ -1759,7 +979,7 @@ let eq_puniverses f (x, u) (y, u') =
f x y && Instance.equal u u'
(** A context of universe levels with universe constraints,
- representiong local universe variables and constraints *)
+ representing local universe variables and constraints *)
module UContext =
struct
@@ -1773,7 +993,7 @@ struct
let pr prl (univs, cst as ctx) =
if is_empty ctx then mt() else
- Instance.pr prl univs ++ str " |= " ++ v 0 (Constraint.pr prl cst)
+ h 0 (Instance.pr prl univs ++ str " |= ") ++ h 0 (v 0 (Constraint.pr prl cst))
let hcons (univs, cst) =
(Instance.hcons univs, hcons_constraints cst)
@@ -1783,8 +1003,11 @@ struct
let union (univs, cst) (univs', cst') =
Instance.append univs univs', Constraint.union cst cst'
-
+
let dest x = x
+
+ let size (x,_) = Instance.length x
+
end
type universe_context = UContext.t
@@ -1802,6 +1025,9 @@ struct
let empty = (LSet.empty, Constraint.empty)
let is_empty (univs, cst) = LSet.is_empty univs && Constraint.is_empty cst
+ let equal (univs, cst as x) (univs', cst' as y) =
+ x == y || (LSet.equal univs univs' && Constraint.equal cst cst')
+
let of_set s = (s, Constraint.empty)
let singleton l = of_set (LSet.singleton l)
let of_instance i = of_set (Instance.levels i)
@@ -1841,7 +1067,7 @@ struct
let pr prl (univs, cst as ctx) =
if is_empty ctx then mt() else
- LSet.pr prl univs ++ str " |= " ++ v 0 (Constraint.pr prl cst)
+ h 0 (LSet.pr prl univs ++ str " |= ") ++ h 0 (v 0 (Constraint.pr prl cst))
let constraints (univs, cst) = cst
let levels (univs, cst) = univs
@@ -1990,27 +1216,6 @@ let abstract_universes poly ctx =
(** Pretty-printing *)
-let pr_arc prl = function
- | _, Canonical {univ=u; lt=[]; le=[]} ->
- mt ()
- | _, Canonical {univ=u; lt=lt; le=le} ->
- let opt_sep = match lt, le with
- | [], _ | _, [] -> mt ()
- | _ -> spc ()
- in
- prl u ++ str " " ++
- v 0
- (pr_sequence (fun v -> str "< " ++ prl v) lt ++
- opt_sep ++
- pr_sequence (fun v -> str "<= " ++ prl v) le) ++
- fnl ()
- | u, Equiv v ->
- prl u ++ str " = " ++ prl v ++ fnl ()
-
-let pr_universes prl g =
- let graph = UMap.fold (fun u a l -> (u,a)::l) g [] in
- prlist (pr_arc prl) graph
-
let pr_constraints prl = Constraint.pr prl
let pr_universe_context = UContext.pr
@@ -2023,19 +1228,6 @@ let pr_universe_subst =
let pr_universe_level_subst =
LMap.pr (fun u -> str" := " ++ Level.pr u ++ spc ())
-(* Dumping constraints to a file *)
-
-let dump_universes output g =
- let dump_arc u = function
- | Canonical {univ=u; lt=lt; le=le} ->
- let u_str = Level.to_string u in
- List.iter (fun v -> output Lt (Level.to_string v) u_str) lt;
- List.iter (fun v -> output Le (Level.to_string v) u_str) le
- | Equiv v ->
- output Eq (Level.to_string u) (Level.to_string v)
- in
- UMap.iter dump_arc g
-
module Huniverse_set =
Hashcons.Make(
struct
@@ -2043,7 +1235,7 @@ module Huniverse_set =
type u = universe_level -> universe_level
let hashcons huc s =
LSet.fold (fun x -> LSet.add (huc x)) s LSet.empty
- let equal s s' =
+ let eq s s' =
LSet.equal s s'
let hash = Hashtbl.hash
end)
@@ -2083,26 +1275,3 @@ let subst_instance_constraints =
let key = Profile.declare_profile "subst_instance_constraints" in
Profile.profile2 key subst_instance_constraints
else subst_instance_constraints
-
-let merge_constraints =
- if Flags.profile then
- let key = Profile.declare_profile "merge_constraints" in
- Profile.profile2 key merge_constraints
- else merge_constraints
-let check_constraints =
- if Flags.profile then
- let key = Profile.declare_profile "check_constraints" in
- Profile.profile2 key check_constraints
- else check_constraints
-
-let check_eq =
- if Flags.profile then
- let check_eq_key = Profile.declare_profile "check_eq" in
- Profile.profile3 check_eq_key check_eq
- else check_eq
-
-let check_leq =
- if Flags.profile then
- let check_leq_key = Profile.declare_profile "check_leq" in
- Profile.profile3 check_leq_key check_leq
- else check_leq
diff --git a/kernel/univ.mli b/kernel/univ.mli
index 7aaf2ffe61..1ccdebd501 100644
--- a/kernel/univ.mli
+++ b/kernel/univ.mli
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -20,7 +20,11 @@ sig
val is_small : t -> bool
(** Is the universe set or prop? *)
-
+
+ val is_prop : t -> bool
+ val is_set : t -> bool
+ (** Is it specifically Prop or Set *)
+
val compare : t -> t -> int
(** Comparison function *)
@@ -36,6 +40,9 @@ sig
val pr : t -> Pp.std_ppcmds
(** Pretty-printing *)
+ val to_string : t -> string
+ (** Debug printing *)
+
val var : int -> t
val var_index : t -> int option
@@ -87,6 +94,9 @@ sig
val is_level : t -> bool
(** Test if the universe is a level or an algebraic universe. *)
+ val is_levels : t -> bool
+ (** Test if the universe is a lub of levels or contains +n's. *)
+
val level : t -> Level.t option
(** Try to get a level out of a universe, returns [None] if it
is an algebraic universe. *)
@@ -108,6 +118,9 @@ sig
val type1 : t
(** the universe of the type of Prop/Set *)
+
+ val exists : (Level.t * int -> bool) -> t -> bool
+ val for_all : (Level.t * int -> bool) -> t -> bool
end
type universe = Universe.t
@@ -141,27 +154,6 @@ val univ_level_mem : universe_level -> universe -> bool
val univ_level_rem : universe_level -> universe -> universe -> universe
-(** {6 Graphs of universes. } *)
-
-type universes
-
-type 'a check_function = universes -> 'a -> 'a -> bool
-val check_leq : universe check_function
-val check_eq : universe check_function
-
-(** The empty graph of universes *)
-val empty_universes : universes
-
-(** The initial graph of universes: Prop < Set *)
-val initial_universes : universes
-
-val is_initial_universes : universes -> bool
-
-val sort_universes : universes -> universes
-
-(** Adds a universe to the graph, ensuring it is >= Prop. *)
-val add_universe : universe_level -> universes -> universes
-
(** {6 Constraints. } *)
type constraint_type = Lt | Le | Eq
@@ -192,12 +184,6 @@ val enforce_leq : universe constraint_function
val enforce_eq_level : universe_level constraint_function
val enforce_leq_level : universe_level constraint_function
-(** {6 ... } *)
-(** Merge of constraints in a universes graph.
- The function [merge_constraints] merges a set of constraints in a given
- universes graph. It raises the exception [UniverseInconsistency] if the
- constraints are not satisfiable. *)
-
(** Type explanation is used to decorate error messages to provide
useful explanation why a given constraint is rejected. It is composed
of a path of universes and relation kinds [(r1,u1);..;(rn,un)] means
@@ -215,14 +201,6 @@ type univ_inconsistency = constraint_type * universe * universe * explanation op
exception UniverseInconsistency of univ_inconsistency
-val enforce_constraint : univ_constraint -> universes -> universes
-val merge_constraints : constraints -> universes -> universes
-
-val constraints_of_universes : universes -> constraints
-
-val check_constraint : universes -> univ_constraint -> bool
-val check_constraints : constraints -> universes -> bool
-
(** {6 Support for universe polymorphism } *)
(** Polymorphic maps from universe levels to 'a *)
@@ -298,8 +276,6 @@ sig
val levels : t -> LSet.t
(** The set of levels in the instance *)
- val check_eq : t check_function
- (** Check equality of instances w.r.t. a universe graph *)
end
type universe_instance = Instance.t
@@ -332,6 +308,9 @@ sig
(** Keeps the order of the instances *)
val union : t -> t -> t
+ (* the number of universes in the context *)
+ val size : t -> int
+
end
type universe_context = UContext.t
@@ -349,6 +328,7 @@ sig
val of_instance : Instance.t -> t
val of_set : universe_set -> t
+ val equal : t -> t -> bool
val union : t -> t -> t
val append : t -> t -> t
@@ -413,7 +393,6 @@ val instantiate_univ_constraints : universe_instance -> universe_context -> cons
(** {6 Pretty-printing of universes. } *)
-val pr_universes : (Level.t -> Pp.std_ppcmds) -> universes -> Pp.std_ppcmds
val pr_constraint_type : constraint_type -> Pp.std_ppcmds
val pr_constraints : (Level.t -> Pp.std_ppcmds) -> constraints -> Pp.std_ppcmds
val pr_universe_context : (Level.t -> Pp.std_ppcmds) -> universe_context -> Pp.std_ppcmds
@@ -424,12 +403,6 @@ val explain_universe_inconsistency : (Level.t -> Pp.std_ppcmds) ->
val pr_universe_level_subst : universe_level_subst -> Pp.std_ppcmds
val pr_universe_subst : universe_subst -> Pp.std_ppcmds
-(** {6 Dumping to a file } *)
-
-val dump_universes :
- (constraint_type -> string -> string -> unit) ->
- universes -> unit
-
(** {6 Hash-consing } *)
val hcons_univ : universe -> universe
diff --git a/kernel/vars.ml b/kernel/vars.ml
index 88c1e1038c..b935ab6b91 100644
--- a/kernel/vars.ml
+++ b/kernel/vars.ml
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -8,7 +8,7 @@
open Names
open Esubst
-open Context
+open Context.Rel.Declaration
(*********************)
(* Occurring *)
@@ -151,20 +151,33 @@ let make_subst = function
done;
subst
+(* The type of substitutions, with term substituting most recent
+ binder at the head *)
+
+type substl = Constr.t list
+
let substnl laml n c = substn_many (make_subst laml) n c
let substl laml c = substn_many (make_subst laml) 0 c
let subst1 lam c = substn_many [|make_substituend lam|] 0 c
-let substnl_decl laml k r = map_rel_declaration (fun c -> substnl laml k c) r
-let substl_decl laml r = map_rel_declaration (fun c -> substnl laml 0 c) r
-let subst1_decl lam r = map_rel_declaration (fun c -> subst1 lam c) r
+let substnl_decl laml k r = map_constr (fun c -> substnl laml k c) r
+let substl_decl laml r = map_constr (fun c -> substnl laml 0 c) r
+let subst1_decl lam r = map_constr (fun c -> subst1 lam c) r
+
+(* Build a substitution from an instance, inserting missing let-ins *)
+
+let subst_of_rel_context_instance sign l =
+ let rec aux subst sign l =
+ match sign, l with
+ | LocalAssum _ :: sign', a::args' -> aux (a::subst) sign' args'
+ | LocalDef (_,c,_)::sign', args' ->
+ aux (substl subst c :: subst) sign' args'
+ | [], [] -> subst
+ | _ -> Errors.anomaly (Pp.str "Instance and signature do not match")
+ in aux [] (List.rev sign) l
-let substnl_named_decl laml k d =
- map_named_declaration (fun c -> substnl laml k c) d
-let substl_named_decl laml d =
- map_named_declaration (fun c -> substnl laml 0 c) d
-let subst1_named_decl lam d =
- map_named_declaration (fun c -> subst1 lam c) d
+let adjust_subst_to_rel_context sign l =
+ List.rev (subst_of_rel_context_instance sign l)
(* (thin_val sigma) removes identity substitutions from sigma *)
@@ -197,15 +210,10 @@ let replace_vars var_alist x =
in
substrec 0 x
-(*
-let repvarkey = Profile.declare_profile "replace_vars";;
-let replace_vars vl c = Profile.profile2 repvarkey replace_vars vl c ;;
-*)
-
-(* (subst_var str t) substitute (VAR str) by (Rel 1) in t *)
+(* (subst_var str t) substitute (Var str) by (Rel 1) in t *)
let subst_var str t = replace_vars [(str, Constr.mkRel 1)] t
-(* (subst_vars [id1;...;idn] t) substitute (VAR idj) by (Rel j) in t *)
+(* (subst_vars [id1;...;idn] t) substitute (Var idj) by (Rel j) in t *)
let substn_vars p vars c =
let _,subst =
List.fold_left (fun (n,l) var -> ((n+1),(var,Constr.mkRel n)::l)) (p,[]) vars
@@ -294,7 +302,7 @@ let subst_univs_level_constr subst c =
if !changed then c' else c
let subst_univs_level_context s =
- map_rel_context (subst_univs_level_constr s)
+ Context.Rel.map (subst_univs_level_constr s)
let subst_instance_constr subst c =
if Univ.Instance.is_empty subst then c
@@ -335,7 +343,7 @@ let subst_instance_constr subst c =
let subst_instance_context s ctx =
if Univ.Instance.is_empty s then ctx
- else map_rel_context (fun x -> subst_instance_constr s x) ctx
+ else Context.Rel.map (fun x -> subst_instance_constr s x) ctx
-type id_key = pconstant tableKey
-let eq_id_key x y = Names.eq_table_key (Univ.eq_puniverses Constant.equal) x y
+type id_key = constant tableKey
+let eq_id_key x y = Names.eq_table_key Constant.equal x y
diff --git a/kernel/vars.mli b/kernel/vars.mli
index fdd4603b5b..574d50eccb 100644
--- a/kernel/vars.mli
+++ b/kernel/vars.mli
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -8,7 +8,6 @@
open Names
open Constr
-open Context
(** {6 Occur checks } *)
@@ -42,32 +41,85 @@ val liftn : int -> int -> constr -> constr
(** [lift n c] lifts by [n] the positive indexes in [c] *)
val lift : int -> constr -> constr
-(** [substnl [a1;...;an] k c] substitutes in parallel [a1],...,[an]
+(** The type [substl] is the type of substitutions [u₁..un] of type
+ some context Δ and defined in some environment Γ. Typing of
+ substitutions is defined by:
+ - Γ ⊢ ∅ : ∅,
+ - Γ ⊢ u₁..u{_n-1} : Δ and Γ ⊢ u{_n} : An\[u₁..u{_n-1}\] implies
+ Γ ⊢ u₁..u{_n} : Δ,x{_n}:A{_n}
+ - Γ ⊢ u₁..u{_n-1} : Δ and Γ ⊢ un : A{_n}\[u₁..u{_n-1}\] implies
+ Γ ⊢ u₁..u{_n} : Δ,x{_n}:=c{_n}:A{_n} when Γ ⊢ u{_n} ≡ c{_n}\[u₁..u{_n-1}\]
+
+ Note that [u₁..un] is represented as a list with [un] at the head of
+ the list, i.e. as [[un;...;u₁]]. *)
+
+type substl = constr list
+
+(** Let [Γ] be a context interleaving declarations [x₁:T₁..xn:Tn]
+ and definitions [y₁:=c₁..yp:=cp] in some context [Γ₀]. Let
+ [u₁..un] be an {e instance} of [Γ], i.e. an instance in [Γ₀]
+ of the [xi]. Then, [subst_of_rel_context_instance Γ u₁..un]
+ returns the corresponding {e substitution} of [Γ], i.e. the
+ appropriate interleaving [σ] of the [u₁..un] with the [c₁..cp],
+ all of them in [Γ₀], so that a derivation [Γ₀, Γ, Γ₁|- t:T]
+ can be instantiated into a derivation [Γ₀, Γ₁ |- t[σ]:T[σ]] using
+ [substnl σ |Γ₁| t].
+ Note that the instance [u₁..un] is represented starting with [u₁],
+ as if usable in [applist] while the substitution is
+ represented the other way round, i.e. ending with either [u₁] or
+ [c₁], as if usable for [substl]. *)
+val subst_of_rel_context_instance : Context.Rel.t -> constr list -> substl
+
+(** For compatibility: returns the substitution reversed *)
+val adjust_subst_to_rel_context : Context.Rel.t -> constr list -> constr list
+
+(** [substnl [a₁;...;an] k c] substitutes in parallel [a₁],...,[an]
for respectively [Rel(k+1)],...,[Rel(k+n)] in [c]; it relocates
- accordingly indexes in [a1],...,[an] and [c] *)
-val substnl : constr list -> int -> constr -> constr
-val substl : constr list -> constr -> constr
+ accordingly indexes in [an],...,[a1] and [c]. In terms of typing, if
+ Γ ⊢ a{_n}..a₁ : Δ and Γ, Δ, Γ' ⊢ c : T with |Γ'|=k, then
+ Γ, Γ' ⊢ [substnl [a₁;...;an] k c] : [substnl [a₁;...;an] k T]. *)
+val substnl : substl -> int -> constr -> constr
+
+(** [substl σ c] is a short-hand for [substnl σ 0 c] *)
+val substl : substl -> constr -> constr
+
+(** [substl a c] is a short-hand for [substnl [a] 0 c] *)
val subst1 : constr -> constr -> constr
-val substnl_decl : constr list -> int -> rel_declaration -> rel_declaration
-val substl_decl : constr list -> rel_declaration -> rel_declaration
-val subst1_decl : constr -> rel_declaration -> rel_declaration
+(** [substnl_decl [a₁;...;an] k Ω] substitutes in parallel [a₁], ..., [an]
+ for respectively [Rel(k+1)], ..., [Rel(k+n)] in [Ω]; it relocates
+ accordingly indexes in [a₁],...,[an] and [c]. In terms of typing, if
+ Γ ⊢ a{_n}..a₁ : Δ and Γ, Δ, Γ', Ω ⊢ with |Γ'|=[k], then
+ Γ, Γ', [substnl_decl [a₁;...;an]] k Ω ⊢. *)
+val substnl_decl : substl -> int -> Context.Rel.Declaration.t -> Context.Rel.Declaration.t
-val substnl_named_decl : constr list -> int -> named_declaration -> named_declaration
-val subst1_named_decl : constr -> named_declaration -> named_declaration
-val substl_named_decl : constr list -> named_declaration -> named_declaration
+(** [substl_decl σ Ω] is a short-hand for [substnl_decl σ 0 Ω] *)
+val substl_decl : substl -> Context.Rel.Declaration.t -> Context.Rel.Declaration.t
+(** [subst1_decl a Ω] is a short-hand for [substnl_decl [a] 0 Ω] *)
+val subst1_decl : constr -> Context.Rel.Declaration.t -> Context.Rel.Declaration.t
+
+(** [replace_vars k [(id₁,c₁);...;(idn,cn)] t] substitutes [Var idj] by
+ [cj] in [t]. *)
val replace_vars : (Id.t * constr) list -> constr -> constr
-(** (subst_var str t) substitute (VAR str) by (Rel 1) in t *)
-val subst_var : Id.t -> constr -> constr
-(** [subst_vars [id1;...;idn] t] substitute [VAR idj] by [Rel j] in [t]
- if two names are identical, the one of least indice is kept *)
+(** [substn_vars k [id₁;...;idn] t] substitutes [Var idj] by [Rel j+k-1] in [t].
+ If two names are identical, the one of least index is kept. In terms of
+ typing, if Γ,x{_n}:U{_n},...,x₁:U₁,Γ' ⊢ t:T, together with id{_j}:T{_j} and
+ Γ,x{_n}:U{_n},...,x₁:U₁,Γ' ⊢ T{_j}\[id{_j+1}..id{_n}:=x{_j+1}..x{_n}\] ≡ Uj,
+ then Γ\\{id₁,...,id{_n}\},x{_n}:U{_n},...,x₁:U₁,Γ' ⊢ [substn_vars
+ (|Γ'|+1) [id₁;...;idn] t] : [substn_vars (|Γ'|+1) [id₁;...;idn]
+ T]. *)
+val substn_vars : int -> Id.t list -> constr -> constr
+
+(** [subst_vars [id1;...;idn] t] is a short-hand for [substn_vars
+ [id1;...;idn] 1 t]: it substitutes [Var idj] by [Rel j] in [t]. If
+ two names are identical, the one of least index is kept. *)
val subst_vars : Id.t list -> constr -> constr
-(** [substn_vars n [id1;...;idn] t] substitute [VAR idj] by [Rel j+n-1] in [t]
- if two names are identical, the one of least indice is kept *)
-val substn_vars : int -> Id.t list -> constr -> constr
+(** [subst_var id t] is a short-hand for [substn_vars [id] 1 t]: it
+ substitutes [Var id] by [Rel 1] in [t]. *)
+val subst_var : Id.t -> constr -> constr
(** {3 Substitution of universes} *)
@@ -82,11 +134,11 @@ val subst_univs_constr : universe_subst -> constr -> constr
(** Level substitutions for polymorphism. *)
val subst_univs_level_constr : universe_level_subst -> constr -> constr
-val subst_univs_level_context : Univ.universe_level_subst -> rel_context -> rel_context
+val subst_univs_level_context : Univ.universe_level_subst -> Context.Rel.t -> Context.Rel.t
(** Instance substitution for polymorphism. *)
val subst_instance_constr : universe_instance -> constr -> constr
-val subst_instance_context : universe_instance -> rel_context -> rel_context
+val subst_instance_context : universe_instance -> Context.Rel.t -> Context.Rel.t
-type id_key = pconstant tableKey
+type id_key = constant tableKey
val eq_id_key : id_key -> id_key -> bool
diff --git a/kernel/vconv.ml b/kernel/vconv.ml
index a03a67db8b..4610dbcb07 100644
--- a/kernel/vconv.ml
+++ b/kernel/vconv.ml
@@ -40,15 +40,20 @@ let conv_vect fconv vect1 vect2 cu =
!rcu
else raise NotConvertible
-let infos = ref (create_clos_infos betaiotazeta Environ.empty_env)
-
let rec conv_val env pb k v1 v2 cu =
if v1 == v2 then cu
else conv_whd env pb k (whd_val v1) (whd_val v2) cu
and conv_whd env pb k whd1 whd2 cu =
+(* Pp.(msg_debug (str "conv_whd(" ++ pr_whd whd1 ++ str ", " ++ pr_whd whd2 ++ str ")")) ; *)
match whd1, whd2 with
- | Vsort s1, Vsort s2 -> check_sort_cmp_universes env pb s1 s2 cu; cu
+ | Vsort s1, Vsort s2 -> sort_cmp_universes env pb s1 s2 cu
+ | Vuniv_level _ , _
+ | _ , Vuniv_level _ ->
+ (** Both of these are invalid since universes are handled via
+ ** special cases in the code.
+ **)
+ assert false
| Vprod p1, Vprod p2 ->
let cu = conv_val env CONV k (dom p1) (dom p2) cu in
conv_fun env pb k (codom p1) (codom p2) cu
@@ -76,53 +81,53 @@ and conv_whd env pb k whd1 whd2 cu =
| Vatom_stk(a1,stk1), Vatom_stk(a2,stk2) ->
conv_atom env pb k a1 stk1 a2 stk2 cu
| Vfun _, _ | _, Vfun _ ->
- conv_val env CONV (k+1) (eta_whd k whd1) (eta_whd k whd2) cu
- | _, Vatom_stk(Aiddef(_,v),stk) ->
- conv_whd env pb k whd1 (force_whd v stk) cu
- | Vatom_stk(Aiddef(_,v),stk), _ ->
- conv_whd env pb k (force_whd v stk) whd2 cu
+ conv_val env CONV (k+1) (apply_whd k whd1) (apply_whd k whd2) cu
| Vsort _, _ | Vprod _, _ | Vfix _, _ | Vcofix _, _ | Vconstr_const _, _
| Vconstr_block _, _ | Vatom_stk _, _ -> raise NotConvertible
and conv_atom env pb k a1 stk1 a2 stk2 cu =
+(* Pp.(msg_debug (str "conv_atom(" ++ pr_atom a1 ++ str ", " ++ pr_atom a2 ++ str ")")) ; *)
match a1, a2 with
- | Aind ind1, Aind ind2 ->
- if eq_puniverses eq_ind ind1 ind2 && compare_stack stk1 stk2
+ | Aind ((mi,i) as ind1) , Aind ind2 ->
+ if eq_ind ind1 ind2 && compare_stack stk1 stk2
then
- conv_stack env k stk1 stk2 cu
+ if Environ.polymorphic_ind ind1 env
+ then
+ let mib = Environ.lookup_mind mi env in
+ let ulen = Univ.UContext.size mib.Declarations.mind_universes in
+ match stk1 , stk2 with
+ | [], [] -> assert (Int.equal ulen 0); cu
+ | Zapp args1 :: stk1' , Zapp args2 :: stk2' ->
+ assert (ulen <= nargs args1);
+ assert (ulen <= nargs args2);
+ let u1 = Array.init ulen (fun i -> uni_lvl_val (arg args1 i)) in
+ let u2 = Array.init ulen (fun i -> uni_lvl_val (arg args2 i)) in
+ let u1 = Univ.Instance.of_array u1 in
+ let u2 = Univ.Instance.of_array u2 in
+ let cu = convert_instances ~flex:false u1 u2 cu in
+ conv_arguments env ~from:ulen k args1 args2
+ (conv_stack env k stk1' stk2' cu)
+ | _, _ -> assert false (* Should not happen if problem is well typed *)
+ else
+ conv_stack env k stk1 stk2 cu
else raise NotConvertible
| Aid ik1, Aid ik2 ->
if Vars.eq_id_key ik1 ik2 && compare_stack stk1 stk2 then
conv_stack env k stk1 stk2 cu
else raise NotConvertible
- | Aiddef(ik1,v1), Aiddef(ik2,v2) ->
- begin
- try
- if Vars.eq_id_key ik1 ik2 && compare_stack stk1 stk2 then
- conv_stack env k stk1 stk2 cu
- else raise NotConvertible
- with NotConvertible ->
- if oracle_order Univ.out_punivs (oracle_of_infos !infos)
- false ik1 ik2 then
- conv_whd env pb k (whd_stack v1 stk1) (Vatom_stk(a2,stk2)) cu
- else conv_whd env pb k (Vatom_stk(a1,stk1)) (whd_stack v2 stk2) cu
- end
- | Aiddef(ik1,v1), _ ->
- conv_whd env pb k (force_whd v1 stk1) (Vatom_stk(a2,stk2)) cu
- | _, Aiddef(ik2,v2) ->
- conv_whd env pb k (Vatom_stk(a1,stk1)) (force_whd v2 stk2) cu
+ | Atype _ , _ | _, Atype _ -> assert false
| Aind _, _ | Aid _, _ -> raise NotConvertible
-and conv_stack env k stk1 stk2 cu =
+and conv_stack env ?from:(from=0) k stk1 stk2 cu =
match stk1, stk2 with
| [], [] -> cu
| Zapp args1 :: stk1, Zapp args2 :: stk2 ->
- conv_stack env k stk1 stk2 (conv_arguments env k args1 args2 cu)
+ conv_stack env k stk1 stk2 (conv_arguments env ~from:from k args1 args2 cu)
| Zfix(f1,args1) :: stk1, Zfix(f2,args2) :: stk2 ->
conv_stack env k stk1 stk2
- (conv_arguments env k args1 args2 (conv_fix env k f1 f2 cu))
+ (conv_arguments env ~from:from k args1 args2 (conv_fix env k f1 f2 cu))
| Zswitch sw1 :: stk1, Zswitch sw2 :: stk2 ->
if check_switch sw1 sw2 then
let vt1,vt2 = type_of_switch sw1, type_of_switch sw2 in
@@ -166,98 +171,37 @@ and conv_cofix env k cf1 cf2 cu =
conv_vect (conv_val env CONV (k + Array.length tcf1)) bcf1 bcf2 cu
else raise NotConvertible
-and conv_arguments env k args1 args2 cu =
+and conv_arguments env ?from:(from=0) k args1 args2 cu =
if args1 == args2 then cu
else
let n = nargs args1 in
if Int.equal n (nargs args2) then
let rcu = ref cu in
- for i = 0 to n - 1 do
+ for i = from to n - 1 do
rcu := conv_val env CONV k (arg args1 i) (arg args2 i) !rcu
done;
!rcu
else raise NotConvertible
-let rec eq_puniverses f (x,l1) (y,l2) cu =
- if f x y then conv_universes l1 l2 cu
- else raise NotConvertible
-
-and conv_universes l1 l2 cu =
- if Univ.Instance.equal l1 l2 then cu else raise NotConvertible
-
-let rec conv_eq env pb t1 t2 cu =
- if t1 == t2 then cu
- else
- match kind_of_term t1, kind_of_term t2 with
- | Rel n1, Rel n2 ->
- if Int.equal n1 n2 then cu else raise NotConvertible
- | Meta m1, Meta m2 ->
- if Int.equal m1 m2 then cu else raise NotConvertible
- | Var id1, Var id2 ->
- if Id.equal id1 id2 then cu else raise NotConvertible
- | Sort s1, Sort s2 -> check_sort_cmp_universes env pb s1 s2 cu; cu
- | Cast (c1,_,_), _ -> conv_eq env pb c1 t2 cu
- | _, Cast (c2,_,_) -> conv_eq env pb t1 c2 cu
- | Prod (_,t1,c1), Prod (_,t2,c2) ->
- conv_eq env pb c1 c2 (conv_eq env CONV t1 t2 cu)
- | Lambda (_,t1,c1), Lambda (_,t2,c2) -> conv_eq env CONV c1 c2 cu
- | LetIn (_,b1,t1,c1), LetIn (_,b2,t2,c2) ->
- conv_eq env pb c1 c2 (conv_eq env CONV b1 b2 cu)
- | App (c1,l1), App (c2,l2) ->
- conv_eq_vect env l1 l2 (conv_eq env CONV c1 c2 cu)
- | Evar (e1,l1), Evar (e2,l2) ->
- if Evar.equal e1 e2 then conv_eq_vect env l1 l2 cu
- else raise NotConvertible
- | Const c1, Const c2 -> eq_puniverses eq_constant c1 c2 cu
- | Proj (p1,c1), Proj (p2,c2) ->
- if eq_constant (Projection.constant p1) (Projection.constant p2) then
- conv_eq env pb c1 c2 cu
- else raise NotConvertible
- | Ind c1, Ind c2 ->
- eq_puniverses eq_ind c1 c2 cu
- | Construct c1, Construct c2 ->
- eq_puniverses eq_constructor c1 c2 cu
- | Case (_,p1,c1,bl1), Case (_,p2,c2,bl2) ->
- let pcu = conv_eq env CONV p1 p2 cu in
- let ccu = conv_eq env CONV c1 c2 pcu in
- conv_eq_vect env bl1 bl2 ccu
- | Fix ((ln1, i1),(_,tl1,bl1)), Fix ((ln2, i2),(_,tl2,bl2)) ->
- if Int.equal i1 i2 && Array.equal Int.equal ln1 ln2 then conv_eq_vect env tl1 tl2 (conv_eq_vect env bl1 bl2 cu)
- else raise NotConvertible
- | CoFix(ln1,(_,tl1,bl1)), CoFix(ln2,(_,tl2,bl2)) ->
- if Int.equal ln1 ln2 then conv_eq_vect env tl1 tl2 (conv_eq_vect env bl1 bl2 cu)
- else raise NotConvertible
- | _ -> raise NotConvertible
-
-and conv_eq_vect env vt1 vt2 cu =
- let len = Array.length vt1 in
- if Int.equal len (Array.length vt2) then
- let rcu = ref cu in
- for i = 0 to len-1 do
- rcu := conv_eq env CONV vt1.(i) vt2.(i) !rcu
- done; !rcu
- else raise NotConvertible
-
-let vconv pb env t1 t2 =
- infos := create_clos_infos betaiotazeta env;
- let _cu =
- try conv_eq env pb t1 t2 (universes env)
- with NotConvertible ->
- let v1 = val_of_constr env t1 in
- let v2 = val_of_constr env t2 in
- let cu = conv_val env pb (nb_rel env) v1 v2 (universes env) in
- cu
- in ()
-
-let _ = Reduction.set_vm_conv vconv
-
-let use_vm = ref false
-
-let set_use_vm b =
- use_vm := b;
- if b then Reduction.set_default_conv (fun cv_pb ?(l2r=false) -> vconv cv_pb)
- else Reduction.set_default_conv (fun cv_pb ?(l2r=false) -> Reduction.conv_cmp cv_pb)
-
-let use_vm _ = !use_vm
-
-
+let vm_conv_gen cv_pb env univs t1 t2 =
+ try
+ let v1 = val_of_constr env t1 in
+ let v2 = val_of_constr env t2 in
+ fst (conv_val env cv_pb (nb_rel env) v1 v2 univs)
+ with Not_found | Invalid_argument _ ->
+ (Pp.msg_warning
+ (Pp.str "Bytecode compilation failed, falling back to default conversion");
+ Reduction.generic_conv cv_pb ~l2r:false (fun _ -> None)
+ full_transparent_state env univs t1 t2)
+
+let vm_conv cv_pb env t1 t2 =
+ let univs = Environ.universes env in
+ let b =
+ if cv_pb = CUMUL then Constr.leq_constr_univs univs t1 t2
+ else Constr.eq_constr_univs univs t1 t2
+ in
+ if not b then
+ let univs = (univs, checked_universes) in
+ let _ = vm_conv_gen cv_pb env univs t1 t2 in ()
+
+let _ = Reduction.set_vm_conv vm_conv
diff --git a/kernel/vconv.mli b/kernel/vconv.mli
index 096d31ac81..ff01735c0b 100644
--- a/kernel/vconv.mli
+++ b/kernel/vconv.mli
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -12,9 +12,11 @@ open Reduction
(**********************************************************************
s conversion functions *)
-val use_vm : unit -> bool
-val set_use_vm : bool -> unit
-val vconv : conv_pb -> types conversion_function
+val vm_conv : conv_pb -> types kernel_conversion_function
-val val_of_constr : env -> constr -> values
+(** A conversion function parametrized by a universe comparator. Used outside of
+ the kernel. *)
+val vm_conv_gen : conv_pb -> (types, 'a) generic_conversion_function
+(** Precompute a VM value from a constr *)
+val val_of_constr : env -> constr -> values
diff --git a/kernel/vm.ml b/kernel/vm.ml
index a822f92eb3..7029876438 100644
--- a/kernel/vm.ml
+++ b/kernel/vm.ml
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -19,8 +19,6 @@ external set_drawinstr : unit -> unit = "coq_set_drawinstr"
external offset_closure : Obj.t -> int -> Obj.t = "coq_offset_closure"
external offset : Obj.t -> int = "coq_offset"
-let accu_tag = 0
-
(*******************************************)
(* Initalization of the abstract machine ***)
(*******************************************)
@@ -29,9 +27,6 @@ external init_vm : unit -> unit = "init_coq_vm"
let _ = init_vm ()
-external transp_values : unit -> bool = "get_coq_transp_value"
-external set_transp_values : bool -> unit = "coq_set_transp_value"
-
(*******************************************)
(* Machine code *** ************************)
(*******************************************)
@@ -126,12 +121,12 @@ type vswitch = {
(* *)
(* + Accumulators : At_[accumulate| accu | arg1 | ... | argn ] *)
(* - representation of [accu] : tag_[....] *)
-(* -- tag <= 2 : encoding atom type (sorts, free vars, etc.) *)
-(* -- 3_[accu|proj name] : a projection blocked by an accu *)
-(* -- 4_[accu|fix_app] : a fixpoint blocked by an accu *)
-(* -- 5_[accu|vswitch] : a match blocked by an accu *)
-(* -- 6_[fcofix] : a cofix function *)
-(* -- 7_[fcofix|val] : a cofix function, val represent the value *)
+(* -- tag <= 3 : encoding atom type (sorts, free vars, etc.) *)
+(* -- 10_[accu|proj name] : a projection blocked by an accu *)
+(* -- 11_[accu|fix_app] : a fixpoint blocked by an accu *)
+(* -- 12_[accu|vswitch] : a match blocked by an accu *)
+(* -- 13_[fcofix] : a cofix function *)
+(* -- 14_[fcofix|val] : a cofix function, val represent the value *)
(* of the function applied to arg1 ... argn *)
(* The [arguments] type, which is abstracted as an array, represents : *)
(* tag[ _ | _ |v1|... | vn] *)
@@ -141,8 +136,8 @@ type vswitch = {
type atom =
| Aid of Vars.id_key
- | Aiddef of Vars.id_key * values
- | Aind of pinductive
+ | Aind of inductive
+ | Atype of Univ.universe
(* Zippers *)
@@ -165,31 +160,112 @@ type whd =
| Vconstr_const of int
| Vconstr_block of vblock
| Vatom_stk of atom * stack
+ | Vuniv_level of Univ.universe_level
+
+(************************************************)
+(* Abstract machine *****************************)
+(************************************************)
+
+(* gestion de la pile *)
+external push_ra : tcode -> unit = "coq_push_ra"
+external push_val : values -> unit = "coq_push_val"
+external push_arguments : arguments -> unit = "coq_push_arguments"
+external push_vstack : vstack -> unit = "coq_push_vstack"
+
+
+(* interpreteur *)
+external interprete : tcode -> values -> vm_env -> int -> values =
+ "coq_interprete_ml"
+
+
+
+(* Functions over arguments *)
+let nargs : arguments -> int = fun args -> (Obj.size (Obj.repr args)) - 2
+let arg args i =
+ if 0 <= i && i < (nargs args) then
+ val_of_obj (Obj.field (Obj.repr args) (i+2))
+ else invalid_arg
+ ("Vm.arg size = "^(string_of_int (nargs args))^
+ " acces "^(string_of_int i))
+
+(* Apply a value to arguments contained in [vargs] *)
+let apply_arguments vf vargs =
+ let n = nargs vargs in
+ if Int.equal n 0 then vf
+ else
+ begin
+ push_ra stop;
+ push_arguments vargs;
+ interprete (fun_code vf) vf (Obj.magic vf) (n - 1)
+ end
+
+(* Apply value [vf] to an array of argument values [varray] *)
+let apply_varray vf varray =
+ let n = Array.length varray in
+ if Int.equal n 0 then vf
+ else
+ begin
+ push_ra stop;
+ push_vstack varray;
+ interprete (fun_code vf) vf (Obj.magic vf) (n - 1)
+ end
(*************************************************)
(* Destructors ***********************************)
(*************************************************)
+let uni_lvl_val (v : values) : Univ.universe_level =
+ let whd = Obj.magic v in
+ match whd with
+ | Vuniv_level lvl -> lvl
+ | _ ->
+ let pr =
+ let open Pp in
+ match whd with
+ | Vsort _ -> str "Vsort"
+ | Vprod _ -> str "Vprod"
+ | Vfun _ -> str "Vfun"
+ | Vfix _ -> str "Vfix"
+ | Vcofix _ -> str "Vcofix"
+ | Vconstr_const i -> str "Vconstr_const"
+ | Vconstr_block b -> str "Vconstr_block"
+ | Vatom_stk (a,stk) -> str "Vatom_stk"
+ | _ -> assert false
+ in
+ Errors.anomaly
+ Pp.( strbrk "Parsing virtual machine value expected universe level, got "
+ ++ pr)
+
let rec whd_accu a stk =
let stk =
if Int.equal (Obj.size a) 2 then stk
else Zapp (Obj.obj a) :: stk in
let at = Obj.field a 1 in
match Obj.tag at with
- | i when i <= 2 ->
+ | i when Int.equal i type_atom_tag ->
+ begin match stk with
+ | [Zapp args] ->
+ let u = ref (Obj.obj (Obj.field at 0)) in
+ for i = 0 to nargs args - 1 do
+ u := Univ.Universe.sup !u (Univ.Universe.make (uni_lvl_val (arg args i)))
+ done;
+ Vsort (Type !u)
+ | _ -> assert false
+ end
+ | i when i <= max_atom_tag ->
Vatom_stk(Obj.magic at, stk)
- | 3 (* proj tag *) ->
+ | i when Int.equal i proj_tag ->
let zproj = Zproj (Obj.obj (Obj.field at 0)) in
whd_accu (Obj.field at 1) (zproj :: stk)
- | 4 (* fix_app tag *) ->
+ | i when Int.equal i fix_app_tag ->
let fa = Obj.field at 1 in
let zfix =
Zfix (Obj.obj (Obj.field fa 1), Obj.obj fa) in
whd_accu (Obj.field at 0) (zfix :: stk)
- | 5 (* switch tag *) ->
+ | i when Int.equal i switch_tag ->
let zswitch = Zswitch (Obj.obj (Obj.field at 1)) in
whd_accu (Obj.field at 0) (zswitch :: stk)
- | 6 (* cofix_tag *) ->
+ | i when Int.equal i cofix_tag ->
let vcfx = Obj.obj (Obj.field at 0) in
let to_up = Obj.obj a in
begin match stk with
@@ -197,7 +273,7 @@ let rec whd_accu a stk =
| [Zapp args] -> Vcofix(vcfx, to_up, Some args)
| _ -> assert false
end
- | 7 (* cofix_evaluated_tag *) ->
+ | i when Int.equal i cofix_evaluated_tag ->
let vcofix = Obj.obj (Obj.field at 0) in
let res = Obj.obj a in
begin match stk with
@@ -205,7 +281,9 @@ let rec whd_accu a stk =
| [Zapp args] -> Vcofix(vcofix, res, Some args)
| _ -> assert false
end
- | _ -> assert false
+ | tg ->
+ Errors.anomaly
+ Pp.(strbrk "Failed to parse VM value. Tag = " ++ int tg)
external kind_of_closure : Obj.t -> int = "coq_kind_of_closure"
@@ -218,65 +296,19 @@ let whd_val : values -> whd =
if tag = accu_tag then
(
if Int.equal (Obj.size o) 1 then Obj.obj o (* sort *)
- else
+ else
if is_accumulate (fun_code o) then whd_accu o []
- else (Vprod(Obj.obj o)))
+ else Vprod(Obj.obj o))
else
if tag = Obj.closure_tag || tag = Obj.infix_tag then
- ( match kind_of_closure o with
+ (match kind_of_closure o with
| 0 -> Vfun(Obj.obj o)
| 1 -> Vfix(Obj.obj o, None)
| 2 -> Vfix(Obj.obj (Obj.field o 1), Some (Obj.obj o))
| 3 -> Vatom_stk(Aid(RelKey(int_tcode (fun_code o) 1)), [])
| _ -> Errors.anomaly ~label:"Vm.whd " (Pp.str "kind_of_closure does not work"))
- else
- Vconstr_block(Obj.obj o)
-
-(************************************************)
-(* Abstrct machine ******************************)
-(************************************************)
-
-(* gestion de la pile *)
-external push_ra : tcode -> unit = "coq_push_ra"
-external push_val : values -> unit = "coq_push_val"
-external push_arguments : arguments -> unit = "coq_push_arguments"
-external push_vstack : vstack -> unit = "coq_push_vstack"
-
-
-(* interpreteur *)
-external interprete : tcode -> values -> vm_env -> int -> values =
- "coq_interprete_ml"
-
-
-
-(* Functions over arguments *)
-let nargs : arguments -> int = fun args -> (Obj.size (Obj.repr args)) - 2
-let arg args i =
- if 0 <= i && i < (nargs args) then
- val_of_obj (Obj.field (Obj.repr args) (i+2))
- else invalid_arg
- ("Vm.arg size = "^(string_of_int (nargs args))^
- " acces "^(string_of_int i))
-
-let apply_arguments vf vargs =
- let n = nargs vargs in
- if Int.equal n 0 then vf
- else
- begin
- push_ra stop;
- push_arguments vargs;
- interprete (fun_code vf) vf (Obj.magic vf) (n - 1)
- end
-
-let apply_vstack vf vstk =
- let n = Array.length vstk in
- if Int.equal n 0 then vf
- else
- begin
- push_ra stop;
- push_vstack vstk;
- interprete (fun_code vf) vf (Obj.magic vf) (n - 1)
- end
+ else
+ Vconstr_block(Obj.obj o)
(**********************************************)
(* Constructors *******************************)
@@ -303,6 +335,8 @@ let rec obj_of_str_const str =
Obj.set_field res i (obj_of_str_const args.(i))
done;
res
+ | Const_univ_level l -> Obj.repr (Vuniv_level l)
+ | Const_type u -> obj_of_atom (Atype u)
let val_of_obj o = ((Obj.obj o) : values)
@@ -310,13 +344,22 @@ let val_of_str_const str = val_of_obj (obj_of_str_const str)
let val_of_atom a = val_of_obj (obj_of_atom a)
+let atom_of_proj kn v =
+ let r = Obj.new_block proj_tag 2 in
+ Obj.set_field r 0 (Obj.repr kn);
+ Obj.set_field r 1 (Obj.repr v);
+ ((Obj.obj r) : atom)
+
+let val_of_proj kn v =
+ val_of_atom (atom_of_proj kn v)
+
module IdKeyHash =
struct
- type t = pconstant tableKey
- let equal = Names.eq_table_key (Univ.eq_puniverses Constant.equal)
+ type t = constant tableKey
+ let equal = Names.eq_table_key Constant.equal
open Hashset.Combine
let hash = function
- | ConstKey (c,u) -> combinesmall 1 (Constant.hash c)
+ | ConstKey c -> combinesmall 1 (Constant.hash c)
| VarKey id -> combinesmall 2 (Id.hash id)
| RelKey i -> combinesmall 3 (Int.hash i)
end
@@ -360,14 +403,14 @@ external closure_arity : vfun -> int = "coq_closure_arity"
let body_of_vfun k vf =
let vargs = mkrel_vstack k 1 in
- apply_vstack (Obj.magic vf) vargs
+ apply_varray (Obj.magic vf) vargs
let decompose_vfun2 k vf1 vf2 =
let arity = min (closure_arity vf1) (closure_arity vf2) in
assert (0 < arity && arity < Sys.max_array_length);
let vargs = mkrel_vstack k arity in
- let v1 = apply_vstack (Obj.magic vf1) vargs in
- let v2 = apply_vstack (Obj.magic vf2) vargs in
+ let v1 = apply_varray (Obj.magic vf1) vargs in
+ let v2 = apply_varray (Obj.magic vf2) vargs in
arity, v1, v2
(* Functions over fixpoint *)
@@ -497,7 +540,7 @@ let reduce_cofix k vcf =
let self = Obj.new_block accu_tag 2 in
Obj.set_field self 0 (Obj.repr accumulate);
Obj.set_field self 1 (Obj.repr atom);
- apply_vstack (Obj.obj e) [|Obj.obj self|] in
+ apply_varray (Obj.obj e) [|Obj.obj self|] in
(Array.init ndef cofix_body, ftyp)
@@ -550,62 +593,13 @@ let branch_of_switch k sw =
in
Array.map eval_branch sw.sw_annot.rtbl
-
-
-(* Evaluation *)
-
-let rec whd_stack v stk =
- match stk with
- | [] -> whd_val v
- | Zapp args :: stkt -> whd_stack (apply_arguments v args) stkt
- | Zfix (f,args) :: stkt ->
- let o = Obj.repr v in
- if Obj.is_block o && Obj.tag o = accu_tag then
- whd_accu (Obj.repr v) stk
- else
- let v', stkt =
- match stkt with
- | Zapp args' :: stkt ->
- push_ra stop;
- push_arguments args';
- push_val v;
- push_arguments args;
- let v' =
- interprete (fun_code f) (Obj.magic f) (Obj.magic f)
- (nargs args+ nargs args') in
- v', stkt
- | _ ->
- push_ra stop;
- push_val v;
- push_arguments args;
- let v' =
- interprete (fun_code f) (Obj.magic f) (Obj.magic f)
- (nargs args) in
- v', stkt
- in
- whd_stack v' stkt
- | Zswitch sw :: stkt ->
- let o = Obj.repr v in
- if Obj.is_block o && Obj.tag o = accu_tag then
- if Obj.tag (Obj.field o 1) < cofix_tag then whd_accu (Obj.repr v) stk
- else
- let to_up =
- match whd_accu (Obj.repr v) [] with
- | Vcofix (_, to_up, _) -> to_up
- | _ -> assert false in
- whd_stack (apply_switch sw to_up) stkt
- else whd_stack (apply_switch sw v) stkt
-
-let rec force_whd v stk =
- match whd_stack v stk with
- | Vatom_stk(Aiddef(_,v),stk) -> force_whd v stk
- | res -> res
-
-
-let rec eta_stack a stk v =
+(* Apply the term represented by a under stack stk to argument v *)
+(* t = a stk --> t v *)
+let rec apply_stack a stk v =
match stk with
- | [] -> apply_vstack a [|v|]
- | Zapp args :: stk -> eta_stack (apply_arguments a args) stk v
+ | [] -> apply_varray a [|v|]
+ | Zapp args :: stk -> apply_stack (apply_arguments a args) stk v
+ | Zproj kn :: stk -> apply_stack (val_of_proj kn a) stk v
| Zfix(f,args) :: stk ->
let a,stk =
match stk with
@@ -626,11 +620,11 @@ let rec eta_stack a stk v =
interprete (fun_code f) (Obj.magic f) (Obj.magic f)
(nargs args) in
a, stk in
- eta_stack a stk v
+ apply_stack a stk v
| Zswitch sw :: stk ->
- eta_stack (apply_switch sw a) stk v
+ apply_stack (apply_switch sw a) stk v
-let eta_whd k whd =
+let apply_whd k whd =
let v = val_of_rel k in
match whd with
| Vsort _ | Vprod _ | Vconstr_const _ | Vconstr_block _ -> assert false
@@ -649,8 +643,35 @@ let eta_whd k whd =
push_val v;
interprete (fun_code to_up) (Obj.magic to_up) (Obj.magic to_up) 0
| Vatom_stk(a,stk) ->
- eta_stack (val_of_atom a) stk v
-
-
-
-
+ apply_stack (val_of_atom a) stk v
+ | Vuniv_level lvl -> assert false
+
+let rec pr_atom a =
+ Pp.(match a with
+ | Aid c -> str "Aid(" ++ (match c with
+ | ConstKey c -> Names.pr_con c
+ | RelKey i -> str "#" ++ int i
+ | _ -> str "...") ++ str ")"
+ | Aind (mi,i) -> str "Aind(" ++ Names.pr_mind mi ++ str "#" ++ int i ++ str ")"
+ | Atype _ -> str "Atype(")
+and pr_whd w =
+ Pp.(match w with
+ | Vsort _ -> str "Vsort"
+ | Vprod _ -> str "Vprod"
+ | Vfun _ -> str "Vfun"
+ | Vfix _ -> str "Vfix"
+ | Vcofix _ -> str "Vcofix"
+ | Vconstr_const i -> str "Vconstr_const(" ++ int i ++ str ")"
+ | Vconstr_block b -> str "Vconstr_block"
+ | Vatom_stk (a,stk) -> str "Vatom_stk(" ++ pr_atom a ++ str ", " ++ pr_stack stk ++ str ")"
+ | Vuniv_level _ -> assert false)
+and pr_stack stk =
+ Pp.(match stk with
+ | [] -> str "[]"
+ | s :: stk -> pr_zipper s ++ str " :: " ++ pr_stack stk)
+and pr_zipper z =
+ Pp.(match z with
+ | Zapp args -> str "Zapp(len = " ++ int (nargs args) ++ str ")"
+ | Zfix (f,args) -> str "Zfix(..., len=" ++ int (nargs args) ++ str ")"
+ | Zswitch s -> str "Zswitch(...)"
+ | Zproj c -> str "Zproj(" ++ Names.pr_con c ++ str ")")
diff --git a/kernel/vm.mli b/kernel/vm.mli
index d31448ee13..6e9579aa46 100644
--- a/kernel/vm.mli
+++ b/kernel/vm.mli
@@ -2,13 +2,10 @@ open Names
open Term
open Cbytecodes
-(** Efficient Virtual Machine *)
+(** Debug printing *)
val set_drawinstr : unit -> unit
-val transp_values : unit -> bool
-val set_transp_values : bool -> unit
-
(** Machine code *)
type tcode
@@ -25,8 +22,8 @@ type arguments
type atom =
| Aid of Vars.id_key
- | Aiddef of Vars.id_key * values
- | Aind of pinductive
+ | Aind of inductive
+ | Atype of Univ.universe
(** Zippers *)
@@ -49,19 +46,27 @@ type whd =
| Vconstr_const of int
| Vconstr_block of vblock
| Vatom_stk of atom * stack
+ | Vuniv_level of Univ.universe_level
+
+(** For debugging purposes only *)
+
+val pr_atom : atom -> Pp.std_ppcmds
+val pr_whd : whd -> Pp.std_ppcmds
+val pr_stack : stack -> Pp.std_ppcmds
(** Constructors *)
val val_of_str_const : structured_constant -> values
val val_of_rel : int -> values
val val_of_named : Id.t -> values
-val val_of_constant : pconstant -> values
+val val_of_constant : constant -> values
external val_of_annot_switch : annot_switch -> values = "%identity"
(** Destructors *)
val whd_val : values -> whd
+val uni_lvl_val : values -> Univ.universe_level
(** Arguments *)
@@ -106,10 +111,6 @@ val case_info : vswitch -> case_info
val type_of_switch : vswitch -> values
val branch_of_switch : int -> vswitch -> (int * values) array
-(** Evaluation *)
-
-val whd_stack : values -> stack -> whd
-val force_whd : values -> stack -> whd
-
-val eta_whd : int -> whd -> values
+(** Apply a value *)
+val apply_whd : int -> whd -> values