aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
-rw-r--r--.gitattributes79
-rw-r--r--CHANGES3
-rw-r--r--dev/ci/README.md5
-rwxr-xr-xdev/ci/ci-compcert.sh4
-rw-r--r--engine/proofview.ml2
-rw-r--r--engine/termops.ml11
-rw-r--r--engine/termops.mli5
-rw-r--r--kernel/constr.ml28
-rw-r--r--kernel/constr.mli9
-rw-r--r--kernel/declarations.ml2
-rw-r--r--kernel/nativecode.ml152
-rw-r--r--kernel/nativecode.mli2
-rw-r--r--kernel/nativeconv.ml4
-rw-r--r--kernel/nativeinstr.mli2
-rw-r--r--kernel/nativelambda.ml7
-rw-r--r--kernel/nativevalues.ml2
-rw-r--r--kernel/nativevalues.mli4
-rw-r--r--plugins/ssrmatching/ssrmatching.ml48
-rw-r--r--pretyping/nativenorm.ml44
-rw-r--r--pretyping/unification.ml9
-rw-r--r--pretyping/unification.mli3
-rw-r--r--pretyping/vnorm.ml20
-rw-r--r--proofs/pfedit.ml3
-rw-r--r--proofs/proof_global.ml4
-rw-r--r--tactics/autorewrite.ml22
-rw-r--r--tactics/eqschemes.ml11
-rw-r--r--tactics/inv.ml2
-rw-r--r--tactics/leminv.ml5
-rw-r--r--tactics/tactics.ml45
-rw-r--r--test-suite/Makefile15
-rw-r--r--test-suite/README.md14
-rw-r--r--test-suite/bugs/closed/7068.v6
-rw-r--r--test-suite/bugs/closed/7076.v4
-rw-r--r--test-suite/bugs/closed/7631.v21
-rw-r--r--theories/Logic/Berardi.v7
-rw-r--r--vernac/auto_ind_decl.ml17
-rw-r--r--vernac/comInductive.ml3
37 files changed, 380 insertions, 204 deletions
diff --git a/.gitattributes b/.gitattributes
index e087e1737..a5edcdb5b 100644
--- a/.gitattributes
+++ b/.gitattributes
@@ -5,48 +5,51 @@
# Because our commit hook automatically does [apply whitespace=fix] we
# disable whitespace checking for all files except those where we want
# it. Otherwise rogue global configuration and forgotten local
-# configuration can break commits.
+# configuration can break commits. Note that git cannot fix but can
+# detect a blank-at-eof when it comes from removing a chunk of text at
+# the end of the file, leaving an extra newline from before that
+# chunk, so we disable blank-at-eof.
* -whitespace
# tabs are allowed in Makefiles.
-Makefile* whitespace=trailing-space
-tools/CoqMakefile.in whitespace=trailing-space
+Makefile* whitespace=blank-at-eol
+tools/CoqMakefile.in whitespace=blank-at-eol
# in general we don't want tabs.
-*.asciidoc whitespace=trailing-space,tab-in-indent
-*.bib whitespace=trailing-space,tab-in-indent
-*.c whitespace=trailing-space,tab-in-indent
-*.css whitespace=trailing-space,tab-in-indent
-*.dtd whitespace=trailing-space,tab-in-indent
-*.el whitespace=trailing-space,tab-in-indent
-*.g whitespace=trailing-space,tab-in-indent
-*.h whitespace=trailing-space,tab-in-indent
-*.html whitespace=trailing-space,tab-in-indent
-*.hva whitespace=trailing-space,tab-in-indent
-*.js whitespace=trailing-space,tab-in-indent
-*.json whitespace=trailing-space,tab-in-indent
-*.lang whitespace=trailing-space,tab-in-indent
-*.md whitespace=trailing-space,tab-in-indent
-*.merlin whitespace=trailing-space,tab-in-indent
-*.ml whitespace=trailing-space,tab-in-indent
-*.ml4 whitespace=trailing-space,tab-in-indent
-*.mli whitespace=trailing-space,tab-in-indent
-*.mll whitespace=trailing-space,tab-in-indent
-*.mllib whitespace=trailing-space,tab-in-indent
-*.mlp whitespace=trailing-space,tab-in-indent
-*.mlpack whitespace=trailing-space,tab-in-indent
-*.nsh whitespace=trailing-space,tab-in-indent
-*.nsi whitespace=trailing-space,tab-in-indent
-*.py whitespace=trailing-space,tab-in-indent
-*.rst whitespace=trailing-space,tab-in-indent
-*.sh whitespace=trailing-space,tab-in-indent
-*.sty whitespace=trailing-space,tab-in-indent
-*.tex whitespace=trailing-space,tab-in-indent
-*.tokens whitespace=trailing-space,tab-in-indent
-*.txt whitespace=trailing-space,tab-in-indent
-*.v whitespace=trailing-space,tab-in-indent
-*.xml whitespace=trailing-space,tab-in-indent
-*.yml whitespace=trailing-space,tab-in-indent
+*.asciidoc whitespace=blank-at-eol,tab-in-indent
+*.bib whitespace=blank-at-eol,tab-in-indent
+*.c whitespace=blank-at-eol,tab-in-indent
+*.css whitespace=blank-at-eol,tab-in-indent
+*.dtd whitespace=blank-at-eol,tab-in-indent
+*.el whitespace=blank-at-eol,tab-in-indent
+*.g whitespace=blank-at-eol,tab-in-indent
+*.h whitespace=blank-at-eol,tab-in-indent
+*.html whitespace=blank-at-eol,tab-in-indent
+*.hva whitespace=blank-at-eol,tab-in-indent
+*.js whitespace=blank-at-eol,tab-in-indent
+*.json whitespace=blank-at-eol,tab-in-indent
+*.lang whitespace=blank-at-eol,tab-in-indent
+*.md whitespace=blank-at-eol,tab-in-indent
+*.merlin whitespace=blank-at-eol,tab-in-indent
+*.ml whitespace=blank-at-eol,tab-in-indent
+*.ml4 whitespace=blank-at-eol,tab-in-indent
+*.mli whitespace=blank-at-eol,tab-in-indent
+*.mll whitespace=blank-at-eol,tab-in-indent
+*.mllib whitespace=blank-at-eol,tab-in-indent
+*.mlp whitespace=blank-at-eol,tab-in-indent
+*.mlpack whitespace=blank-at-eol,tab-in-indent
+*.nsh whitespace=blank-at-eol,tab-in-indent
+*.nsi whitespace=blank-at-eol,tab-in-indent
+*.py whitespace=blank-at-eol,tab-in-indent
+*.rst whitespace=blank-at-eol,tab-in-indent
+*.sh whitespace=blank-at-eol,tab-in-indent
+*.sty whitespace=blank-at-eol,tab-in-indent
+*.tex whitespace=blank-at-eol,tab-in-indent
+*.tokens whitespace=blank-at-eol,tab-in-indent
+*.txt whitespace=blank-at-eol,tab-in-indent
+*.v whitespace=blank-at-eol,tab-in-indent
+*.xml whitespace=blank-at-eol,tab-in-indent
+*.yml whitespace=blank-at-eol,tab-in-indent
# CR is desired for these Windows files.
-*.bat whitespace=cr-at-eol,trailing-space,tab-in-indent
+*.bat whitespace=cr-at-eol,blank-at-eol,tab-in-indent
diff --git a/CHANGES b/CHANGES
index e8718bbab..787c9ba12 100644
--- a/CHANGES
+++ b/CHANGES
@@ -25,6 +25,9 @@ Tactics
- Deprecated the Implicit Tactic family of commands.
+- The `simple apply` tactic now respects the `Opaque` flag when called from
+ Ltac (`auto` still does not respect it).
+
Tools
- Coq_makefile lets one override or extend the following variables from
diff --git a/dev/ci/README.md b/dev/ci/README.md
index 697a160ca..dc586c61f 100644
--- a/dev/ci/README.md
+++ b/dev/ci/README.md
@@ -89,6 +89,11 @@ We are currently running tests on the following platforms:
- AppVeyor is used to test the compilation of Coq and run the test-suite on
Windows.
+GitLab CI and Travis CI and AppVeyor support putting `[ci skip]` in a commit
+message to bypass CI. Do not use this unless your commit only changes files
+that are not compiled (e.g. Markdown files like this one, or files under
+[`.github/`](/.github/)).
+
You can anticipate the results of most of these tests prior to submitting your
PR by running GitLab CI on your private branches. To do so follow these steps:
diff --git a/dev/ci/ci-compcert.sh b/dev/ci/ci-compcert.sh
index 8d490591b..0b0e06e29 100755
--- a/dev/ci/ci-compcert.sh
+++ b/dev/ci/ci-compcert.sh
@@ -5,6 +5,10 @@ ci_dir="$(dirname "$0")"
CompCert_CI_DIR="${CI_BUILD_DIR}/CompCert"
+# Temporal workaround, to be removed when upstream decides what to do
+# with their problem.
+opam install -y menhir.20180528
+
git_checkout "${CompCert_CI_BRANCH}" "${CompCert_CI_GITURL}" "${CompCert_CI_DIR}"
( cd "${CompCert_CI_DIR}" && ./configure -ignore-coq-version x86_32-linux && make && make check-proof )
diff --git a/engine/proofview.ml b/engine/proofview.ml
index fdb0a215d..b4afb6415 100644
--- a/engine/proofview.ml
+++ b/engine/proofview.ml
@@ -39,7 +39,7 @@ let proofview p =
let compact el ({ solution } as pv) =
let nf c = Evarutil.nf_evar solution c in
- let nf0 c = EConstr.Unsafe.to_constr (Evarutil.nf_evar solution (EConstr.of_constr c)) in
+ let nf0 c = EConstr.(to_constr solution (of_constr c)) in
let size = Evd.fold (fun _ _ i -> i+1) solution 0 in
let new_el = List.map (fun (t,ty) -> nf t, nf ty) el in
let pruned_solution = Evd.drop_all_defined solution in
diff --git a/engine/termops.ml b/engine/termops.ml
index 0c567754a..eacc36107 100644
--- a/engine/termops.ml
+++ b/engine/termops.ml
@@ -781,24 +781,23 @@ let map_constr_with_full_binders sigma g f l cstr =
let fold_constr_with_full_binders sigma g f n acc c =
let open RelDecl in
- let inj c = EConstr.Unsafe.to_constr c in
match EConstr.kind sigma c with
| (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _
| Construct _) -> acc
| Cast (c,_, t) -> f n (f n acc c) t
- | Prod (na,t,c) -> f (g (LocalAssum (na, inj t)) n) (f n acc t) c
- | Lambda (na,t,c) -> f (g (LocalAssum (na, inj t)) n) (f n acc t) c
- | LetIn (na,b,t,c) -> f (g (LocalDef (na, inj b, inj t)) n) (f n (f n acc b) t) c
+ | Prod (na,t,c) -> f (g (LocalAssum (na, t)) n) (f n acc t) c
+ | Lambda (na,t,c) -> f (g (LocalAssum (na, t)) n) (f n acc t) c
+ | LetIn (na,b,t,c) -> f (g (LocalDef (na, b, t)) n) (f n (f n acc b) t) c
| App (c,l) -> Array.fold_left (f n) (f n acc c) l
| Proj (p,c) -> f n acc c
| Evar (_,l) -> Array.fold_left (f n) acc l
| Case (_,p,c,bl) -> Array.fold_left (f n) (f n (f n acc p) c) bl
| Fix (_,(lna,tl,bl)) ->
- let n' = CArray.fold_left2 (fun c n t -> g (LocalAssum (n, inj t)) c) n lna tl in
+ let n' = CArray.fold_left2 (fun c n t -> g (LocalAssum (n, t)) c) n lna tl in
let fd = Array.map2 (fun t b -> (t,b)) tl bl in
Array.fold_left (fun acc (t,b) -> f n' (f n acc t) b) acc fd
| CoFix (_,(lna,tl,bl)) ->
- let n' = CArray.fold_left2 (fun c n t -> g (LocalAssum (n, inj t)) c) n lna tl in
+ let n' = CArray.fold_left2 (fun c n t -> g (LocalAssum (n, t)) c) n lna tl in
let fd = Array.map2 (fun t b -> (t,b)) tl bl in
Array.fold_left (fun acc (t,b) -> f n' (f n acc t) b) acc fd
diff --git a/engine/termops.mli b/engine/termops.mli
index 6e63539ca..255494031 100644
--- a/engine/termops.mli
+++ b/engine/termops.mli
@@ -75,8 +75,9 @@ val fold_constr_with_binders : Evd.evar_map ->
('a -> 'a) -> ('a -> 'b -> constr -> 'b) -> 'a -> 'b -> constr -> 'b
val fold_constr_with_full_binders : Evd.evar_map ->
- (Context.Rel.Declaration.t -> 'a -> 'a) -> ('a -> 'b -> constr -> 'b) ->
- 'a -> 'b -> constr -> 'b
+ (rel_declaration -> 'a -> 'a) ->
+ ('a -> 'b -> constr -> 'b) ->
+ 'a -> 'b -> constr -> 'b
val iter_constr_with_full_binders : Evd.evar_map ->
(rel_declaration -> 'a -> 'a) ->
diff --git a/kernel/constr.ml b/kernel/constr.ml
index 8f83d6baa..c11b9ebf4 100644
--- a/kernel/constr.ml
+++ b/kernel/constr.ml
@@ -479,6 +479,34 @@ let iter_with_binders g f n c = match kind c with
Array.Fun1.iter f n tl;
Array.Fun1.iter f (iterate g (Array.length tl) n) bl
+(* [fold_constr_with_binders g f n acc c] folds [f n] on the immediate
+ subterms of [c] starting from [acc] and proceeding from left to
+ right according to the usual representation of the constructions as
+ [fold_constr] but it carries an extra data [n] (typically a lift
+ index) which is processed by [g] (which typically add 1 to [n]) at
+ each binder traversal; it is not recursive *)
+
+let fold_constr_with_binders g f n acc c =
+ match kind c with
+ | (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _
+ | Construct _) -> acc
+ | Cast (c,_, t) -> f n (f n acc c) t
+ | Prod (na,t,c) -> f (g n) (f n acc t) c
+ | Lambda (na,t,c) -> f (g n) (f n acc t) c
+ | LetIn (na,b,t,c) -> f (g n) (f n (f n acc b) t) c
+ | App (c,l) -> Array.fold_left (f n) (f n acc c) l
+ | Proj (p,c) -> f n acc c
+ | Evar (_,l) -> Array.fold_left (f n) acc l
+ | Case (_,p,c,bl) -> Array.fold_left (f n) (f n (f n acc p) c) bl
+ | Fix (_,(lna,tl,bl)) ->
+ let n' = CArray.fold_left2 (fun c n t -> g c) n lna tl in
+ let fd = Array.map2 (fun t b -> (t,b)) tl bl in
+ Array.fold_left (fun acc (t,b) -> f n' (f n acc t) b) acc fd
+ | CoFix (_,(lna,tl,bl)) ->
+ let n' = CArray.fold_left2 (fun c n t -> g c) n lna tl in
+ let fd = Array.map2 (fun t b -> (t,b)) tl bl in
+ Array.fold_left (fun acc (t,b) -> f n' (f n acc t) b) acc fd
+
(* [map f c] maps [f] on the immediate subterms of [c]; it is
not recursive and the order with which subterms are processed is
not specified *)
diff --git a/kernel/constr.mli b/kernel/constr.mli
index b35ea6653..742a13919 100644
--- a/kernel/constr.mli
+++ b/kernel/constr.mli
@@ -402,6 +402,15 @@ val iter : (constr -> unit) -> constr -> unit
val iter_with_binders :
('a -> 'a) -> ('a -> constr -> unit) -> 'a -> constr -> unit
+(** [iter_with_binders g f n c] iters [f n] on the immediate
+ subterms of [c]; it carries an extra data [n] (typically a lift
+ index) which is processed by [g] (which typically add 1 to [n]) at
+ each binder traversal; it is not recursive and the order with which
+ subterms are processed is not specified *)
+
+val fold_constr_with_binders :
+ ('a -> 'a) -> ('a -> 'b -> constr -> 'b) -> 'a -> 'b -> constr -> 'b
+
type constr_compare_fn = int -> constr -> constr -> bool
(** [compare_head f c1 c2] compare [c1] and [c2] using [f] to compare
diff --git a/kernel/declarations.ml b/kernel/declarations.ml
index 913c13173..7bd70c050 100644
--- a/kernel/declarations.ml
+++ b/kernel/declarations.ml
@@ -52,7 +52,7 @@ type inline = int option
type projection_body = {
proj_ind : MutInd.t;
proj_npars : int;
- proj_arg : int;
+ proj_arg : int; (** Projection index, starting from 0 *)
proj_type : types; (* Type under params *)
proj_eta : constr * types; (* Eta-expanded term and type *)
proj_body : constr; (* For compatibility with VMs only, the match version *)
diff --git a/kernel/nativecode.ml b/kernel/nativecode.ml
index 036cd4847..8257dc8b8 100644
--- a/kernel/nativecode.ml
+++ b/kernel/nativecode.ml
@@ -53,7 +53,7 @@ type gname =
| Gind of string * inductive (* prefix, inductive name *)
| Gconstruct of string * constructor (* prefix, constructor name *)
| Gconstant of string * Constant.t (* prefix, constant name *)
- | Gproj of string * Constant.t (* prefix, constant name *)
+ | Gproj of string * inductive * int (* prefix, inductive, index (start from 0) *)
| Gcase of Label.t option * int
| Gpred of Label.t option * int
| Gfixtype of Label.t option * int
@@ -108,7 +108,7 @@ let gname_hash gn = match gn with
| Ginternal s -> combinesmall 9 (String.hash s)
| Grel i -> combinesmall 10 (Int.hash i)
| Gnamed id -> combinesmall 11 (Id.hash id)
-| Gproj (s, p) -> combinesmall 12 (combine (String.hash s) (Constant.hash p))
+| Gproj (s, p, i) -> combinesmall 12 (combine (String.hash s) (combine (ind_hash p) i))
let case_ctr = ref (-1)
@@ -152,6 +152,7 @@ type symbol =
| SymbMeta of metavariable
| SymbEvar of Evar.t
| SymbLevel of Univ.Level.t
+ | SymbProj of (inductive * int)
let dummy_symb = SymbValue (dummy_value ())
@@ -166,6 +167,7 @@ let eq_symbol sy1 sy2 =
| SymbMeta m1, SymbMeta m2 -> Int.equal m1 m2
| SymbEvar evk1, SymbEvar evk2 -> Evar.equal evk1 evk2
| SymbLevel l1, SymbLevel l2 -> Univ.Level.equal l1 l2
+ | SymbProj (i1, k1), SymbProj (i2, k2) -> eq_ind i1 i2 && Int.equal k1 k2
| _, _ -> false
let hash_symbol symb =
@@ -179,6 +181,7 @@ let hash_symbol symb =
| SymbMeta m -> combinesmall 7 m
| SymbEvar evk -> combinesmall 8 (Evar.hash evk)
| SymbLevel l -> combinesmall 9 (Univ.Level.hash l)
+ | SymbProj (i, k) -> combinesmall 10 (combine (ind_hash i) k)
module HashedTypeSymbol = struct
type t = symbol
@@ -241,6 +244,11 @@ let get_level tbl i =
| SymbLevel u -> u
| _ -> anomaly (Pp.str "get_level failed.")
+let get_proj tbl i =
+ match tbl.(i) with
+ | SymbProj p -> p
+ | _ -> anomaly (Pp.str "get_proj failed.")
+
let push_symbol x =
try HashtblSymbol.find symb_tbl x
with Not_found ->
@@ -885,6 +893,10 @@ let get_level_code i =
MLapp (MLglobal (Ginternal "get_level"),
[|MLglobal symbols_tbl_name; MLint i|])
+let get_proj_code i =
+ MLapp (MLglobal (Ginternal "get_proj"),
+ [|MLglobal symbols_tbl_name; MLint i|])
+
type rlist =
| Rnil
| Rcons of (constructor * lname option array) list ref * LNset.t * mllambda * rlist'
@@ -1070,7 +1082,7 @@ let ml_of_instance instance u =
| Lconst (prefix, (c, u)) ->
let args = ml_of_instance env.env_univ u in
mkMLapp (MLglobal(Gconstant (prefix, c))) args
- | Lproj (prefix,c) -> MLglobal(Gproj (prefix,c))
+ | Lproj (prefix, ind, i) -> MLglobal(Gproj (prefix, ind, i))
| Lprim _ ->
let decl,cond,paux = extract_prim (ml_of_lam env l) t in
compile_prim decl cond paux
@@ -1544,8 +1556,8 @@ let string_of_gname g =
Format.sprintf "%sconstruct_%s_%i_%i" prefix (string_of_mind mind) i (j-1)
| Gconstant (prefix, c) ->
Format.sprintf "%sconst_%s" prefix (string_of_con c)
- | Gproj (prefix, c) ->
- Format.sprintf "%sproj_%s" prefix (string_of_con c)
+ | Gproj (prefix, (mind, n), i) ->
+ Format.sprintf "%sproj_%s_%i_%i" prefix (string_of_mind mind) n i
| Gcase (l,i) ->
Format.sprintf "case_%s_%i" (string_of_label_def l) i
| Gpred (l,i) ->
@@ -1858,8 +1870,6 @@ and compile_named env sigma univ auxdefs id =
Glet(Gnamed id, MLprimitive (Mk_var id))::auxdefs
let compile_constant env sigma prefix ~interactive con cb =
- match cb.const_proj with
- | false ->
let no_univs =
match cb.const_universes with
| Monomorphic_const _ -> true
@@ -1903,39 +1913,6 @@ let compile_constant env sigma prefix ~interactive con cb =
if interactive then LinkedInteractive prefix
else Linked prefix
end
- | true ->
- let pb = lookup_projection (Projection.make con false) env in
- let mind = pb.proj_ind in
- let ind = (mind,0) in
- let mib = lookup_mind mind env in
- let oib = mib.mind_packets.(0) in
- let tbl = oib.mind_reloc_tbl in
- (* Building info *)
- let prefix = get_mind_prefix env mind in
- let ci = { ci_ind = ind; ci_npar = mib.mind_nparams;
- ci_cstr_nargs = [|0|];
- ci_cstr_ndecls = [||] (*FIXME*);
- ci_pp_info = { ind_tags = []; cstr_tags = [||] (*FIXME*); style = RegularStyle } } in
- let asw = { asw_ind = ind; asw_prefix = prefix; asw_ci = ci;
- asw_reloc = tbl; asw_finite = true } in
- let c_uid = fresh_lname Anonymous in
- let cf_uid = fresh_lname Anonymous in
- let _, arity = tbl.(0) in
- let ci_uid = fresh_lname Anonymous in
- let cargs = Array.init arity
- (fun i -> if Int.equal i pb.proj_arg then Some ci_uid else None)
- in
- let i = push_symbol (SymbConst con) in
- let accu = MLapp (MLprimitive Cast_accu, [|MLlocal cf_uid|]) in
- let accu_br = MLapp (MLprimitive Mk_proj, [|get_const_code i;accu|]) in
- let code = MLmatch(asw,MLlocal cf_uid,accu_br,[|[((ind,1),cargs)],MLlocal ci_uid|]) in
- let code = MLlet(cf_uid, MLapp (MLprimitive Force_cofix, [|MLlocal c_uid|]), code) in
- let gn = Gproj ("",con) in
- let fargs = Array.init (pb.proj_npars + 1) (fun _ -> fresh_lname Anonymous) in
- let arg = fargs.(pb.proj_npars) in
- Glet(Gconstant ("", con), mkMLlam fargs (MLapp (MLglobal gn, [|MLlocal
- arg|])))::
- [Glet(gn, mkMLlam [|c_uid|] code)], Linked prefix
module StringOrd = struct type t = string let compare = String.compare end
module StringSet = Set.Make(StringOrd)
@@ -1962,10 +1939,12 @@ let arg_name = Name (Id.of_string "arg")
let compile_mind prefix ~interactive mb mind stack =
let u = Declareops.inductive_polymorphic_context mb in
+ (** Generate data for every block *)
let f i stack ob =
- let gtype = Gtype((mind, i), Array.map snd ob.mind_reloc_tbl) in
- let j = push_symbol (SymbInd (mind,i)) in
- let name = Gind ("", (mind, i)) in
+ let ind = (mind, i) in
+ let gtype = Gtype(ind, Array.map snd ob.mind_reloc_tbl) in
+ let j = push_symbol (SymbInd ind) in
+ let name = Gind ("", ind) in
let accu =
let args =
if Int.equal (Univ.AUContext.size u) 0 then
@@ -1979,12 +1958,41 @@ let compile_mind prefix ~interactive mb mind stack =
Array.init nparams (fun i -> {lname = param_name; luid = i}) in
let add_construct j acc (_,arity) =
let args = Array.init arity (fun k -> {lname = arg_name; luid = k}) in
- let c = (mind,i), (j+1) in
+ let c = ind, (j+1) in
Glet(Gconstruct ("", c),
mkMLlam (Array.append params args)
(MLconstruct("", c, Array.map (fun id -> MLlocal id) args)))::acc
in
- Array.fold_left_i add_construct (gtype::accu::stack) ob.mind_reloc_tbl
+ let constructors = Array.fold_left_i add_construct [] ob.mind_reloc_tbl in
+ let add_proj j acc pb =
+ let tbl = ob.mind_reloc_tbl in
+ (* Building info *)
+ let ci = { ci_ind = ind; ci_npar = nparams;
+ ci_cstr_nargs = [|0|];
+ ci_cstr_ndecls = [||] (*FIXME*);
+ ci_pp_info = { ind_tags = []; cstr_tags = [||] (*FIXME*); style = RegularStyle } } in
+ let asw = { asw_ind = ind; asw_prefix = ""; asw_ci = ci;
+ asw_reloc = tbl; asw_finite = true } in
+ let c_uid = fresh_lname Anonymous in
+ let cf_uid = fresh_lname Anonymous in
+ let _, arity = tbl.(0) in
+ let ci_uid = fresh_lname Anonymous in
+ let cargs = Array.init arity
+ (fun i -> if Int.equal i pb.proj_arg then Some ci_uid else None)
+ in
+ let i = push_symbol (SymbProj (ind, j)) in
+ let accu = MLapp (MLprimitive Cast_accu, [|MLlocal cf_uid|]) in
+ let accu_br = MLapp (MLprimitive Mk_proj, [|get_proj_code i;accu|]) in
+ let code = MLmatch(asw,MLlocal cf_uid,accu_br,[|[((ind,1),cargs)],MLlocal ci_uid|]) in
+ let code = MLlet(cf_uid, MLapp (MLprimitive Force_cofix, [|MLlocal c_uid|]), code) in
+ let gn = Gproj ("", (pb.proj_ind, j), pb.proj_arg) in
+ Glet (gn, mkMLlam [|c_uid|] code) :: acc
+ in
+ let projs = match mb.mind_record with
+ | None | Some None -> []
+ | Some (Some (id, kns, pbs)) -> Array.fold_left_i add_proj [] pbs
+ in
+ projs @ constructors @ gtype :: accu :: stack
in
Array.fold_left_i f stack mb.mind_packets
@@ -2017,25 +2025,22 @@ let compile_mind_deps env prefix ~interactive
(* This function compiles all necessary dependencies of t, and generates code in
reverse order, as well as linking information updates *)
-let rec compile_deps env sigma prefix ~interactive init t =
+let compile_deps env sigma prefix ~interactive init t =
+ let rec aux env lvl init t =
match kind t with
| Ind ((mind,_),u) -> compile_mind_deps env prefix ~interactive init mind
| Const c ->
- let c,u = get_alias env c in
- let cb,(nameref,_) = lookup_constant_key c env in
- let (_, (_, const_updates)) = init in
- if is_code_loaded ~interactive nameref
- || (Cmap_env.mem c const_updates)
- then init
- else
+ let c,u = get_alias env c in
+ let cb,(nameref,_) = lookup_constant_key c env in
+ let (_, (_, const_updates)) = init in
+ if is_code_loaded ~interactive nameref
+ || (Cmap_env.mem c const_updates)
+ then init
+ else
let comp_stack, (mind_updates, const_updates) =
- match cb.const_proj, cb.const_body with
- | false, Def t ->
- compile_deps env sigma prefix ~interactive init (Mod_subst.force_constr t)
- | true, _ ->
- let pb = lookup_projection (Projection.make c false) env in
- let mind = pb.proj_ind in
- compile_mind_deps env prefix ~interactive init mind
+ match cb.const_body with
+ | Def t ->
+ aux env lvl init (Mod_subst.force_constr t)
| _ -> init
in
let code, name =
@@ -2046,13 +2051,32 @@ let rec compile_deps env sigma prefix ~interactive init t =
comp_stack, (mind_updates, const_updates)
| Construct (((mind,_),_),u) -> compile_mind_deps env prefix ~interactive init mind
| Proj (p,c) ->
- let term = mkApp (mkConst (Projection.constant p), [|c|]) in
- compile_deps env sigma prefix ~interactive init term
+ let pb = lookup_projection p env in
+ let init = compile_mind_deps env prefix ~interactive init pb.proj_ind in
+ aux env lvl init c
| Case (ci, p, c, ac) ->
let mind = fst ci.ci_ind in
let init = compile_mind_deps env prefix ~interactive init mind in
- Constr.fold (compile_deps env sigma prefix ~interactive) init t
- | _ -> Constr.fold (compile_deps env sigma prefix ~interactive) init t
+ fold_constr_with_binders succ (aux env) lvl init t
+ | Var id ->
+ let open Context.Named.Declaration in
+ begin match lookup_named id env with
+ | LocalDef (_,t,_) ->
+ aux env lvl init t
+ | _ -> init
+ end
+ | Rel n when n > lvl ->
+ let open Context.Rel.Declaration in
+ let decl = lookup_rel n env in
+ let env = env_of_rel n env in
+ begin match decl with
+ | LocalDef (_,t,_) ->
+ aux env lvl init t
+ | LocalAssum _ -> init
+ end
+ | _ -> fold_constr_with_binders succ (aux env) lvl init t
+ in
+ aux env 0 init t
let compile_constant_field env prefix con acc cb =
let (gl, _) =
diff --git a/kernel/nativecode.mli b/kernel/nativecode.mli
index 42f2cbc2e..684983a87 100644
--- a/kernel/nativecode.mli
+++ b/kernel/nativecode.mli
@@ -50,6 +50,8 @@ val get_evar : symbols -> int -> Evar.t
val get_level : symbols -> int -> Univ.Level.t
+val get_proj : symbols -> int -> inductive * int
+
val get_symbols : unit -> symbols
type code_location_update
diff --git a/kernel/nativeconv.ml b/kernel/nativeconv.ml
index c07025660..e97dbd0d6 100644
--- a/kernel/nativeconv.ml
+++ b/kernel/nativeconv.ml
@@ -114,8 +114,8 @@ and conv_atom env pb lvl a1 a2 cu =
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
+ | Aproj((ind1, i1), ac1), Aproj((ind2, i2), ac2) ->
+ if not (eq_ind ind1 ind2 && Int.equal i1 i2) then raise NotConvertible
else conv_accu env CONV lvl ac1 ac2 cu
| Arel _, _ | Aind _, _ | Aconstant _, _ | Asort _, _ | Avar _, _
| Acase _, _ | Afix _, _ | Acofix _, _ | Acofixe _, _ | Aprod _, _
diff --git a/kernel/nativeinstr.mli b/kernel/nativeinstr.mli
index c319be32d..eaad8ee0c 100644
--- a/kernel/nativeinstr.mli
+++ b/kernel/nativeinstr.mli
@@ -31,7 +31,7 @@ and lambda =
| Llet of Name.t * lambda * lambda
| Lapp of lambda * lambda array
| Lconst of prefix * pconstant
- | Lproj of prefix * Constant.t (* prefix, projection name *)
+ | Lproj of prefix * inductive * int (* prefix, inductive, index starting from 0 *)
| Lprim of prefix * Constant.t * CPrimitives.t * lambda array
| Lcase of annot_sw * lambda * lambda * lam_branches
(* annotations, term being matched, accu, branches *)
diff --git a/kernel/nativelambda.ml b/kernel/nativelambda.ml
index 6db75c89a..0325a00b4 100644
--- a/kernel/nativelambda.ml
+++ b/kernel/nativelambda.ml
@@ -519,8 +519,11 @@ let rec lambda_of_constr env sigma c =
| Construct _ -> lambda_of_app env sigma c empty_args
| Proj (p, c) ->
- let kn = Projection.constant p in
- mkLapp (Lproj (get_const_prefix !global_env kn, kn)) [|lambda_of_constr env sigma c|]
+ let pb = lookup_projection p !global_env in
+ (** FIXME: handle mutual records *)
+ let ind = (pb.proj_ind, 0) in
+ let prefix = get_mind_prefix !global_env (fst ind) in
+ mkLapp (Lproj (prefix, ind, pb.proj_arg)) [|lambda_of_constr env sigma c|]
| Case(ci,t,a,branches) ->
let (mind,i as ind) = ci.ci_ind in
diff --git a/kernel/nativevalues.ml b/kernel/nativevalues.ml
index cfcb0a485..da4413a0a 100644
--- a/kernel/nativevalues.ml
+++ b/kernel/nativevalues.ml
@@ -64,7 +64,7 @@ type atom =
| Aprod of Name.t * t * (t -> t)
| Ameta of metavariable * t
| Aevar of Evar.t * t * t array
- | Aproj of Constant.t * accumulator
+ | Aproj of (inductive * int) * accumulator
let accumulate_tag = 0
diff --git a/kernel/nativevalues.mli b/kernel/nativevalues.mli
index 4a58a3c7d..649853f06 100644
--- a/kernel/nativevalues.mli
+++ b/kernel/nativevalues.mli
@@ -54,7 +54,7 @@ type atom =
| Aprod of Name.t * t * (t -> t)
| Ameta of metavariable * t
| Aevar of Evar.t * t (* type *) * t array (* arguments *)
- | Aproj of Constant.t * accumulator
+ | Aproj of (inductive * int) * accumulator
(* Constructors *)
@@ -71,7 +71,7 @@ val mk_fix_accu : rec_pos -> int -> t array -> t array -> t
val mk_cofix_accu : int -> t array -> t array -> t
val mk_meta_accu : metavariable -> t
val mk_evar_accu : Evar.t -> t -> t array -> t
-val mk_proj_accu : Constant.t -> accumulator -> t
+val mk_proj_accu : (inductive * int) -> accumulator -> t
val upd_cofix : t -> t -> unit
val force_cofix : t -> t
val mk_const : tag -> t
diff --git a/plugins/ssrmatching/ssrmatching.ml4 b/plugins/ssrmatching/ssrmatching.ml4
index 93c63d522..b1c5e131f 100644
--- a/plugins/ssrmatching/ssrmatching.ml4
+++ b/plugins/ssrmatching/ssrmatching.ml4
@@ -312,20 +312,22 @@ let unif_HO_args env ise0 pa i ca =
(* for HO evars, though hopefully Miller patterns can pick up some of *)
(* those cases, and HO matching will mop up the rest. *)
let flags_FO env =
+ let oracle = Environ.oracle env in
+ let ts = Conv_oracle.get_transp_state oracle in
let flags =
- { (Unification.default_no_delta_unify_flags ()).Unification.core_unify_flags
+ { (Unification.default_no_delta_unify_flags ts).Unification.core_unify_flags
with
Unification.modulo_conv_on_closed_terms = None;
Unification.modulo_eta = true;
Unification.modulo_betaiota = true;
- Unification.modulo_delta_types = Conv_oracle.get_transp_state (Environ.oracle env)}
+ Unification.modulo_delta_types = ts }
in
{ Unification.core_unify_flags = flags;
Unification.merge_unify_flags = flags;
Unification.subterm_unify_flags = flags;
Unification.allow_K_in_toplevel_higher_order_unification = false;
Unification.resolve_evars =
- (Unification.default_no_delta_unify_flags ()).Unification.resolve_evars
+ (Unification.default_no_delta_unify_flags ts).Unification.resolve_evars
}
let unif_FO env ise p c =
Unification.w_unify env ise Reduction.CONV ~flags:(flags_FO env)
diff --git a/pretyping/nativenorm.ml b/pretyping/nativenorm.ml
index 978ceed1e..4b8e0e096 100644
--- a/pretyping/nativenorm.ml
+++ b/pretyping/nativenorm.ml
@@ -188,6 +188,14 @@ let branch_of_switch lvl ans bs =
bs ci in
Array.init (Array.length tbl) branch
+let get_proj env ((mind, _n), i) =
+ let mib = Environ.lookup_mind mind env in
+ match mib.mind_record with
+ | None | Some None ->
+ CErrors.anomaly (Pp.strbrk "Return type is not a primitive record")
+ | Some (Some (_, projs, _)) ->
+ Projection.make projs.(i) true
+
let rec nf_val env sigma v typ =
match kind_of_value v with
| Vaccu accu -> nf_accu env sigma accu
@@ -279,9 +287,10 @@ and nf_atom env sigma atom =
let codom = nf_type env sigma (codom vn) in
mkProd(n,dom,codom)
| Ameta (mv,_) -> mkMeta mv
- | Aproj(p,c) ->
+ | Aproj (p, c) ->
let c = nf_accu env sigma c in
- mkProj(Projection.make p true,c)
+ let p = get_proj env p in
+ mkProj(p, c)
| _ -> fst (nf_atom_type env sigma atom)
and nf_atom_type env sigma atom =
@@ -303,10 +312,10 @@ and nf_atom_type env sigma atom =
let (mib,mip) = Inductive.lookup_mind_specif env (fst ind) in
let nparams = mib.mind_nparams in
let params,realargs = Array.chop nparams allargs in
+ let nparamdecls = Context.Rel.length (Inductive.inductive_paramdecls (mib,u)) in
let pT =
- hnf_prod_applist env
+ hnf_prod_applist_assum env nparamdecls
(Inductiveops.type_of_inductive env ind) (Array.to_list params) in
- let pT = whd_all env pT in
let dep, p = nf_predicate env sigma ind mip params p pT in
(* Calcul du type des branches *)
let btypes = build_branches_type env sigma (fst ind) mib mip u params dep p in
@@ -357,25 +366,30 @@ and nf_atom_type env sigma atom =
| Aproj(p,c) ->
let c,tc = nf_accu_type env sigma c in
let cj = make_judge c tc in
- let uj = Typeops.judge_of_projection env (Projection.make p true) cj in
+ let p = get_proj env p in
+ let uj = Typeops.judge_of_projection env p cj in
uj.uj_val, uj.uj_type
and nf_predicate env sigma ind mip params v pT =
- match kind_of_value v, kind pT with
- | Vfun f, Prod _ ->
+ match kind (whd_allnolet env pT) with
+ | LetIn (name,b,t,pT) ->
+ let dep,body =
+ nf_predicate (push_rel (LocalDef (name,b,t)) env) sigma ind mip params v pT in
+ dep, mkLetIn (name,b,t,body)
+ | Prod (name,dom,codom) -> begin
+ match kind_of_value v with
+ | Vfun f ->
let k = nb_rel env in
let vb = f (mk_rel_accu k) in
- let name,dom,codom =
- try decompose_prod env pT with
- DestKO ->
- CErrors.anomaly
- (Pp.strbrk "Returned a functional value in a type not recognized as a product type.")
- in
let dep,body =
nf_predicate (push_rel (LocalAssum (name,dom)) env) sigma ind mip params vb codom in
dep, mkLambda(name,dom,body)
- | Vfun f, _ ->
+ | _ -> false, nf_type env sigma v
+ end
+ | _ ->
+ match kind_of_value v with
+ | Vfun f ->
let k = nb_rel env in
let vb = f (mk_rel_accu k) in
let name = Name (Id.of_string "c") in
@@ -385,7 +399,7 @@ and nf_predicate env sigma ind mip params v pT =
let dom = mkApp(mkIndU ind,Array.append params rargs) in
let body = nf_type (push_rel (LocalAssum (name,dom)) env) sigma vb in
true, mkLambda(name,dom,body)
- | _, _ -> false, nf_type env sigma v
+ | _ -> false, nf_type env sigma v
and nf_evar env sigma evk ty args =
let evi = try Evd.find sigma evk with Not_found -> assert false in
diff --git a/pretyping/unification.ml b/pretyping/unification.ml
index a8a4003dc..5cf6e4b26 100644
--- a/pretyping/unification.ml
+++ b/pretyping/unification.ml
@@ -398,8 +398,13 @@ let default_no_delta_core_unify_flags () = { (default_core_unify_flags ()) with
modulo_betaiota = false;
}
-let default_no_delta_unify_flags () =
- let flags = default_no_delta_core_unify_flags () in {
+let default_no_delta_unify_flags ts =
+ let flags = default_no_delta_core_unify_flags () in
+ let flags = { flags with
+ modulo_conv_on_closed_terms = Some ts;
+ modulo_delta_types = ts
+ } in
+ {
core_unify_flags = flags;
merge_unify_flags = flags;
subterm_unify_flags = flags;
diff --git a/pretyping/unification.mli b/pretyping/unification.mli
index 16ce5c93d..e2e261ae7 100644
--- a/pretyping/unification.mli
+++ b/pretyping/unification.mli
@@ -8,6 +8,7 @@
(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
+open Names
open Constr
open EConstr
open Environ
@@ -40,7 +41,7 @@ val default_core_unify_flags : unit -> core_unify_flags
val default_no_delta_core_unify_flags : unit -> core_unify_flags
val default_unify_flags : unit -> unify_flags
-val default_no_delta_unify_flags : unit -> unify_flags
+val default_no_delta_unify_flags : transparent_state -> unify_flags
val elim_flags : unit -> unify_flags
val elim_no_delta_flags : unit -> unify_flags
diff --git a/pretyping/vnorm.ml b/pretyping/vnorm.ml
index a1ba4a6a9..14c9f49b1 100644
--- a/pretyping/vnorm.ml
+++ b/pretyping/vnorm.ml
@@ -266,7 +266,6 @@ and nf_stk ?from:(from=0) env sigma c t stk =
let nparamdecls = Context.Rel.length (Inductive.inductive_paramdecls (mib,u)) in
let pT =
hnf_prod_applist_assum env nparamdecls (type_of_ind env (ind,u)) (Array.to_list params) in
- let pT = whd_all env pT in
let dep, p = nf_predicate env sigma (ind,u) mip params (type_of_switch sw) pT in
(* Calcul du type des branches *)
let btypes = build_branches_type env sigma ind mib mip u params dep p in
@@ -288,15 +287,24 @@ and nf_stk ?from:(from=0) env sigma c t stk =
nf_stk env sigma (mkProj(p',c)) ty stk
and nf_predicate env sigma ind mip params v pT =
- match whd_val v, kind pT with
- | Vfun f, Prod _ ->
+ match kind (whd_allnolet env pT) with
+ | LetIn (name,b,t,pT) ->
+ let dep,body =
+ nf_predicate (push_rel (LocalDef (name,b,t)) env) sigma ind mip params v pT in
+ dep, mkLetIn (name,b,t,body)
+ | Prod (name,dom,codom) -> begin
+ match whd_val v with
+ | Vfun f ->
let k = nb_rel env in
let vb = reduce_fun k f in
- let name,dom,codom = decompose_prod env pT in
let dep,body =
nf_predicate (push_rel (LocalAssum (name,dom)) env) sigma ind mip params vb codom in
dep, mkLambda(name,dom,body)
- | Vfun f, _ ->
+ | _ -> assert false
+ end
+ | _ ->
+ match whd_val v with
+ | Vfun f ->
let k = nb_rel env in
let vb = reduce_fun k f in
let name = Name (Id.of_string "c") in
@@ -306,7 +314,7 @@ and nf_predicate env sigma ind mip params v pT =
let dom = mkApp(mkIndU ind,Array.append params rargs) in
let body = nf_vtype (push_rel (LocalAssum (name,dom)) env) sigma vb in
true, mkLambda(name,dom,body)
- | _, _ -> false, nf_val env sigma v crazy_type
+ | _ -> false, nf_val env sigma v crazy_type
and nf_args env sigma vargs ?from:(f=0) t =
let t = ref t in
diff --git a/proofs/pfedit.ml b/proofs/pfedit.ml
index 03c0969fa..678c3ea3f 100644
--- a/proofs/pfedit.ml
+++ b/proofs/pfedit.ml
@@ -200,8 +200,7 @@ let refine_by_tactic env sigma ty tac =
| [c, _] -> c
| _ -> assert false
in
- let ans = Reductionops.nf_evar sigma ans in
- let ans = EConstr.Unsafe.to_constr ans in
+ let ans = EConstr.to_constr ~abort_on_undefined_evars:false sigma ans in
(** [neff] contains the freshly generated side-effects *)
let neff = Evd.eval_side_effects sigma in
(** Reset the old side-effects *)
diff --git a/proofs/proof_global.ml b/proofs/proof_global.ml
index 3abdd129e..946379356 100644
--- a/proofs/proof_global.ml
+++ b/proofs/proof_global.ml
@@ -437,8 +437,8 @@ let return_proof ?(allow_partial=false) () =
(** ppedrot: FIXME, this is surely wrong. There is no reason to duplicate
side-effects... This may explain why one need to uniquize side-effects
thereafter... *)
- let proofs =
- List.map (fun (c, _) -> (Evarutil.nf_evars_universes evd (EConstr.Unsafe.to_constr c), eff)) initial_goals in
+ let proofs =
+ List.map (fun (c, _) -> (EConstr.to_constr evd c, eff)) initial_goals in
proofs, Evd.evar_universe_context evd
let close_future_proof ~feedback_id proof =
diff --git a/tactics/autorewrite.ml b/tactics/autorewrite.ml
index 0b0e629ab..c8fd0b7a7 100644
--- a/tactics/autorewrite.ml
+++ b/tactics/autorewrite.ml
@@ -228,7 +228,7 @@ let decompose_applied_relation metas env sigma c ctype left2right =
if metas then eqclause
else clenv_pose_metas_as_evars eqclause (Evd.undefined_metas eqclause.evd)
in
- let (equiv, args) = decompose_app (EConstr.Unsafe.to_constr (Clenv.clenv_type eqclause)) in
+ let (equiv, args) = EConstr.decompose_app sigma (Clenv.clenv_type eqclause) in
let rec split_last_two = function
| [c1;c2] -> [],(c1, c2)
| x::y::z ->
@@ -236,17 +236,19 @@ let decompose_applied_relation metas env sigma c ctype left2right =
| _ -> raise Not_found
in
try
- let others,(c1,c2) = split_last_two args in
- let ty1, ty2 =
- Typing.unsafe_type_of env eqclause.evd (EConstr.of_constr c1), Typing.unsafe_type_of env eqclause.evd (EConstr.of_constr c2)
- in
- let ty = EConstr.Unsafe.to_constr ty in
- let ty1 = EConstr.Unsafe.to_constr ty1 in
+ let others,(c1,c2) = split_last_two args in
+ let ty1, ty2 = Typing.unsafe_type_of env eqclause.evd c1, Typing.unsafe_type_of env eqclause.evd c2 in
+ (* XXX: It looks like mk_clenv_from_env should be fixed instead? *)
+ let open EConstr in
+ let hyp_ty = Unsafe.to_constr ty in
+ let hyp_car = Unsafe.to_constr ty1 in
+ let hyp_prf = Unsafe.to_constr @@ Clenv.clenv_value eqclause in
+ let hyp_rel = Unsafe.to_constr @@ mkApp (equiv, Array.of_list others) in
+ let hyp_left = Unsafe.to_constr @@ c1 in
+ let hyp_right = Unsafe.to_constr @@ c2 in
(* if not (evd_convertible env eqclause.evd ty1 ty2) then None *)
(* else *)
- Some { hyp_cl=eqclause; hyp_prf=EConstr.Unsafe.to_constr (Clenv.clenv_value eqclause); hyp_ty = ty;
- hyp_car=ty1; hyp_rel=mkApp (equiv, Array.of_list others);
- hyp_l2r=left2right; hyp_left=c1; hyp_right=c2; }
+ Some { hyp_cl=eqclause; hyp_prf; hyp_ty; hyp_car; hyp_rel; hyp_l2r=left2right; hyp_left; hyp_right; }
with Not_found -> None
in
match find_rel ctype with
diff --git a/tactics/eqschemes.ml b/tactics/eqschemes.ml
index eede13329..ad5239116 100644
--- a/tactics/eqschemes.ml
+++ b/tactics/eqschemes.ml
@@ -108,9 +108,14 @@ let get_coq_eq ctx =
user_err Pp.(str "eq not found.")
let univ_of_eq env eq =
- let eq = EConstr.of_constr eq in
- match Constr.kind (EConstr.Unsafe.to_constr (Retyping.get_type_of env (Evd.from_env env) eq)) with
- | Prod (_,t,_) -> (match Constr.kind t with Sort (Type u) -> u | _ -> assert false)
+ let open EConstr in
+ let eq = of_constr eq in
+ let sigma = Evd.from_env env in
+ match kind sigma (Retyping.get_type_of env sigma eq) with
+ | Prod (_,t,_) -> (match kind sigma t with
+ Sort k ->
+ (match ESorts.kind sigma k with Type u -> u | _ -> assert false)
+ | _ -> assert false)
| _ -> assert false
(**********************************************************************)
diff --git a/tactics/inv.ml b/tactics/inv.ml
index 339abbc2e..102b8e54d 100644
--- a/tactics/inv.ml
+++ b/tactics/inv.ml
@@ -293,7 +293,7 @@ let error_too_many_names pats =
str "Unexpected " ++
str (String.plural (List.length pats) "introduction pattern") ++
str ": " ++ pr_enum (Miscprint.pr_intro_pattern
- (fun c -> Printer.pr_constr_env env sigma (EConstr.Unsafe.to_constr (snd (c env (Evd.from_env env)))))) pats ++
+ (fun c -> Printer.pr_econstr_env env sigma (snd (c env (Evd.from_env env))))) pats ++
str ".")
let get_names (allow_conj,issimple) ({CAst.loc;v=pat} as x) = match pat with
diff --git a/tactics/leminv.ml b/tactics/leminv.ml
index f47e6b2cd..10937322e 100644
--- a/tactics/leminv.ml
+++ b/tactics/leminv.ml
@@ -232,9 +232,8 @@ let inversion_scheme env sigma t sort dep_option inv_op =
let c = fill_holes pfterm in
(* warning: side-effect on ownSign *)
let invProof = it_mkNamedLambda_or_LetIn c !ownSign in
- let invProof = EConstr.Unsafe.to_constr invProof in
- let p = Evarutil.nf_evars_universes sigma invProof in
- p, sigma
+ let p = EConstr.to_constr sigma invProof in
+ p, sigma
let add_inversion_lemma ~poly name env sigma t sort dep inv_op =
let invProof, sigma = inversion_scheme env sigma t sort dep inv_op in
diff --git a/tactics/tactics.ml b/tactics/tactics.ml
index d7888f6ca..b571b347d 100644
--- a/tactics/tactics.ml
+++ b/tactics/tactics.ml
@@ -1637,13 +1637,11 @@ let tclORELSEOPT t k =
Proofview.tclZERO ~info e
| Some tac -> tac)
-let general_apply with_delta with_destruct with_evars clear_flag
- {CAst.loc;v=(c,lbind : EConstr.constr with_bindings)} =
+let general_apply ?(respect_opaque=false) with_delta with_destruct with_evars
+ clear_flag {CAst.loc;v=(c,lbind : EConstr.constr with_bindings)} =
Proofview.Goal.enter begin fun gl ->
let concl = Proofview.Goal.concl gl in
let sigma = Tacmach.New.project gl in
- let flags =
- if with_delta then default_unify_flags () else default_no_delta_unify_flags () in
(* The actual type of the theorem. It will be matched against the
goal. If this fails, then the head constant will be unfolded step by
step. *)
@@ -1652,7 +1650,12 @@ let general_apply with_delta with_destruct with_evars clear_flag
Proofview.Goal.enter begin fun gl ->
let env = Proofview.Goal.env gl in
let sigma = Tacmach.New.project gl in
-
+ let ts =
+ if respect_opaque then Conv_oracle.get_transp_state (oracle env)
+ else full_transparent_state
+ in
+ let flags =
+ if with_delta then default_unify_flags () else default_no_delta_unify_flags ts in
let thm_ty0 = nf_betaiota env sigma (Retyping.get_type_of env sigma c) in
let try_apply thm_ty nprod =
try
@@ -1718,14 +1721,14 @@ let rec apply_with_bindings_gen b e = function
(general_apply b b e k cb)
(apply_with_bindings_gen b e cbl)
-let apply_with_delayed_bindings_gen b e l =
+let apply_with_delayed_bindings_gen b e l =
let one k {CAst.loc;v=f} =
Proofview.Goal.enter begin fun gl ->
let sigma = Tacmach.New.project gl in
let env = Proofview.Goal.env gl in
let (sigma, cb) = f env sigma in
Tacticals.New.tclWITHHOLES e
- (general_apply b b e k CAst.(make ?loc cb)) sigma
+ (general_apply ~respect_opaque:(not b) b b e k CAst.(make ?loc cb)) sigma
end
in
let rec aux = function
@@ -1800,14 +1803,12 @@ let apply_in_once_main flags innerclause env sigma (loc,d,lbind) =
in
aux (make_clenv_binding env sigma (d,thm) lbind)
-let apply_in_once sidecond_first with_delta with_destruct with_evars naming
- id (clear_flag,{ CAst.loc; v= d,lbind}) tac =
+let apply_in_once ?(respect_opaque = false) sidecond_first with_delta
+ with_destruct with_evars naming id (clear_flag,{ CAst.loc; v= d,lbind}) tac =
let open Context.Rel.Declaration in
Proofview.Goal.enter begin fun gl ->
let env = Proofview.Goal.env gl in
let sigma = Tacmach.New.project gl in
- let flags =
- if with_delta then default_unify_flags () else default_no_delta_unify_flags () in
let t' = Tacmach.New.pf_get_hyp_typ id gl in
let innerclause = mk_clenv_from_env env sigma (Some 0) (mkVar id,t') in
let targetid = find_name true (LocalAssum (Anonymous,t')) naming gl in
@@ -1815,6 +1816,12 @@ let apply_in_once sidecond_first with_delta with_destruct with_evars naming
Proofview.Goal.enter begin fun gl ->
let env = Proofview.Goal.env gl in
let sigma = Tacmach.New.project gl in
+ let ts =
+ if respect_opaque then Conv_oracle.get_transp_state (oracle env)
+ else full_transparent_state
+ in
+ let flags =
+ if with_delta then default_unify_flags () else default_no_delta_unify_flags ts in
try
let clause = apply_in_once_main flags innerclause env sigma (loc,c,lbind) in
clenv_refine_in ~sidecond_first with_evars targetid id sigma clause
@@ -1834,14 +1841,14 @@ let apply_in_once sidecond_first with_delta with_destruct with_evars naming
aux [] with_destruct d
end
-let apply_in_delayed_once sidecond_first with_delta with_destruct with_evars naming
- id (clear_flag,{CAst.loc;v=f}) tac =
+let apply_in_delayed_once ?(respect_opaque = false) sidecond_first with_delta
+ with_destruct with_evars naming id (clear_flag,{CAst.loc;v=f}) tac =
Proofview.Goal.enter begin fun gl ->
let env = Proofview.Goal.env gl in
let sigma = Tacmach.New.project gl in
let (sigma, c) = f env sigma in
Tacticals.New.tclWITHHOLES with_evars
- (apply_in_once sidecond_first with_delta with_destruct with_evars
+ (apply_in_once ~respect_opaque sidecond_first with_delta with_destruct with_evars
naming id (clear_flag,CAst.(make ?loc c)) tac)
sigma
end
@@ -2531,11 +2538,11 @@ let assert_as first hd ipat t =
(* apply in as *)
-let general_apply_in sidecond_first with_delta with_destruct with_evars
- id lemmas ipat =
+let general_apply_in ?(respect_opaque=false) sidecond_first with_delta
+ with_destruct with_evars id lemmas ipat =
let tac (naming,lemma) tac id =
- apply_in_delayed_once sidecond_first with_delta with_destruct with_evars
- naming id lemma tac in
+ apply_in_delayed_once ~respect_opaque sidecond_first with_delta
+ with_destruct with_evars naming id lemma tac in
Proofview.Goal.enter begin fun gl ->
let destopt =
if with_evars then MoveLast (* evars would depend on the whole context *)
@@ -2566,7 +2573,7 @@ let apply_in simple with_evars id lemmas ipat =
general_apply_in false simple simple with_evars id lemmas ipat
let apply_delayed_in simple with_evars id lemmas ipat =
- general_apply_in false simple simple with_evars id lemmas ipat
+ general_apply_in ~respect_opaque:true false simple simple with_evars id lemmas ipat
(*****************************)
(* Tactics abstracting terms *)
diff --git a/test-suite/Makefile b/test-suite/Makefile
index f41fb5b1e..32e245e36 100644
--- a/test-suite/Makefile
+++ b/test-suite/Makefile
@@ -362,26 +362,33 @@ $(addsuffix .log,$(wildcard output/*.v)): %.v.log: %.v %.out $(PREREQUISITELOG)
@echo "TEST $< $(call get_coq_prog_args_in_parens,"$<")"
$(HIDE){ \
echo $(call log_intro,$<); \
- tmpoutput=`mktemp /tmp/coqcheck.XXXXXX`; \
+ output=$*.out.real; \
$(call get_command_based_on_flags,"$<") "$<" $(call get_coq_prog_args,"$<") 2>&1 \
| grep -v "Welcome to Coq" \
| grep -v "\[Loading ML file" \
| grep -v "Skipping rcfile loading" \
| grep -v "^<W>" \
| sed 's/File "[^"]*"/File "stdin"/' \
- > $$tmpoutput; \
- diff -u --strip-trailing-cr $*.out $$tmpoutput 2>&1; R=$$?; times; \
+ > $$output; \
+ diff -u --strip-trailing-cr $*.out $$output 2>&1; R=$$?; times; \
if [ $$R = 0 ]; then \
echo $(log_success); \
echo " $<...Ok"; \
+ rm $$output; \
else \
echo $(log_failure); \
echo " $<...Error! (unexpected output)"; \
$(FAIL); \
fi; \
- rm $$tmpoutput; \
} > "$@"
+.PHONY: approve-output
+approve-output: output
+ $(HIDE)for f in output/*.out.real; do \
+ mv "$$f" "$${f%.real}"; \
+ echo "Updated $${f%.real}!"; \
+ done
+
# the expected output for the MExtraction test is
# /plugins/micromega/micromega.ml except with additional newline
output/MExtraction.out: ../plugins/micromega/micromega.ml
diff --git a/test-suite/README.md b/test-suite/README.md
index 4572c98cf..ef2e574ec 100644
--- a/test-suite/README.md
+++ b/test-suite/README.md
@@ -76,3 +76,17 @@ There are also output tests in `test-suite/output` which consist of a `.v` file
There are unit tests of OCaml code in `test-suite/unit-tests`. These tests are contained in `.ml` files, and rely on the `OUnit`
unit-test framework, as described at http://ounit.forge.ocamlcore.org/. Use `make unit-tests' in the unit-tests directory to run them.
+
+## Fixing output tests
+
+When an output test `output/foo.v` fails, the output is stored in
+`output/foo.out.real`. Move that file to the reference file
+`output/foo.out` to update the test, approving the new output. Target
+`approve-output` will do this for all failing output tests
+automatically.
+
+Don't forget to check the updated `.out` files into git!
+
+Note that `output/MExtraction.out` is special: it is copied from
+`micromega/micromega.ml` in the plugin source directory. Automatic
+approval will incorrectly update the copy.
diff --git a/test-suite/bugs/closed/7068.v b/test-suite/bugs/closed/7068.v
new file mode 100644
index 000000000..9fadb195b
--- /dev/null
+++ b/test-suite/bugs/closed/7068.v
@@ -0,0 +1,6 @@
+(* These tests are only about a subset of #7068 *)
+(* The original issue is still open *)
+
+Inductive foo : let T := Type in T := .
+Definition bob1 := Eval vm_compute in foo_rect.
+Definition bob2 := Eval native_compute in foo_rect.
diff --git a/test-suite/bugs/closed/7076.v b/test-suite/bugs/closed/7076.v
new file mode 100644
index 000000000..0abc88c28
--- /dev/null
+++ b/test-suite/bugs/closed/7076.v
@@ -0,0 +1,4 @@
+(* These calls were raising an anomaly at some time *)
+Inductive A : nat -> id (nat->Type) := .
+Eval vm_compute in fun x => match x in A y z return y = z with end.
+Eval native_compute in fun x => match x in A y z return y = z with end.
diff --git a/test-suite/bugs/closed/7631.v b/test-suite/bugs/closed/7631.v
new file mode 100644
index 000000000..34eb8b867
--- /dev/null
+++ b/test-suite/bugs/closed/7631.v
@@ -0,0 +1,21 @@
+Module NamedContext.
+
+Definition foo := true.
+
+Section Foo.
+
+Let bar := foo.
+
+Eval native_compute in bar.
+
+End Foo.
+
+End NamedContext.
+
+Module RelContext.
+
+Definition foo := true.
+
+Definition bar (x := foo) := Eval native_compute in x.
+
+End RelContext.
diff --git a/theories/Logic/Berardi.v b/theories/Logic/Berardi.v
index c6836a1c7..ed4d69ab0 100644
--- a/theories/Logic/Berardi.v
+++ b/theories/Logic/Berardi.v
@@ -82,7 +82,7 @@ End Retracts.
(** This lemma is basically a commutation of implication and existential
quantification: (EX x | A -> P(x)) <=> (A -> EX x | P(x))
which is provable in classical logic ( => is already provable in
- intuitionnistic logic). *)
+ intuitionistic logic). *)
Lemma L1 : forall A B:Prop, retract_cond (pow A) (pow B).
Proof.
@@ -136,7 +136,7 @@ trivial.
Qed.
-Theorem classical_proof_irrelevence : T = F.
+Theorem classical_proof_irrelevance : T = F.
Proof.
generalize not_has_fixpoint.
unfold Not_b.
@@ -148,4 +148,7 @@ intros not_true is_true.
elim not_true; trivial.
Qed.
+
+Notation classical_proof_irrelevence := classical_proof_irrelevance (compat "8.8").
+
End Berardis_paradox.
diff --git a/vernac/auto_ind_decl.ml b/vernac/auto_ind_decl.ml
index 30a268a11..8b56275c7 100644
--- a/vernac/auto_ind_decl.ml
+++ b/vernac/auto_ind_decl.ml
@@ -186,10 +186,10 @@ let build_beq_scheme mode kn =
*)
let compute_A_equality rel_list nlist eqA ndx t =
let lifti = ndx in
- let sigma = Evd.empty (** FIXME *) in
let rec aux c =
- let (c,a) = Reductionops.whd_betaiota_stack Evd.empty c in
- match EConstr.kind sigma c with
+ let (c,a) = Reductionops.whd_betaiota_stack Evd.empty EConstr.(of_constr c) in
+ let (c,a) = EConstr.Unsafe.(to_constr c, List.map to_constr a) in
+ match Constr.kind c with
| Rel x -> mkRel (x-nlist+ndx), Safe_typing.empty_private_constants
| Var x ->
let eid = Id.of_string ("eq_"^(Id.to_string x)) in
@@ -198,7 +198,7 @@ let build_beq_scheme mode kn =
with Not_found -> raise (ParameterWithoutEquality (VarRef x))
in
mkVar eid, Safe_typing.empty_private_constants
- | Cast (x,_,_) -> aux (EConstr.applist (x,a))
+ | Cast (x,_,_) -> aux (Term.applist (x,a))
| App _ -> assert false
| Ind ((kn',i as ind'),u) (*FIXME: universes *) ->
if MutInd.equal kn kn' then mkRel(eqA-nlist-i+nb_ind-1), Safe_typing.empty_private_constants
@@ -213,8 +213,8 @@ let build_beq_scheme mode kn =
List.fold_left Safe_typing.concat_private eff (List.rev effs)
in
let args =
- Array.append
- (Array.of_list (List.map (fun x -> lift lifti (EConstr.Unsafe.to_constr x)) a)) eqa in
+ Array.append
+ (Array.of_list (List.map (fun x -> lift lifti x) a)) eqa in
if Int.equal (Array.length args) 0 then eq, eff
else mkApp (eq, args), eff
with Not_found -> raise(EqNotFound (ind', fst ind))
@@ -224,10 +224,9 @@ let build_beq_scheme mode kn =
| Lambda _-> raise (EqUnknown "abstraction")
| LetIn _ -> raise (EqUnknown "let-in")
| Const (kn, u) ->
- let u = EConstr.EInstance.kind sigma u in
(match Environ.constant_opt_value_in env (kn, u) with
| None -> raise (ParameterWithoutEquality (ConstRef kn))
- | Some c -> aux (EConstr.applist (EConstr.of_constr c,a)))
+ | Some c -> aux (Term.applist (c,a)))
| Proj _ -> raise (EqUnknown "projection")
| Construct _ -> raise (EqUnknown "constructor")
| Case _ -> raise (EqUnknown "match")
@@ -271,7 +270,7 @@ let build_beq_scheme mode kn =
nparrec
(nparrec+3+2*nb_cstr_args)
(nb_cstr_args+ndx+1)
- (EConstr.of_constr cc)
+ cc
in
eff := Safe_typing.concat_private eff' !eff;
Array.set eqs ndx
diff --git a/vernac/comInductive.ml b/vernac/comInductive.ml
index 101c14266..b93e8d9ac 100644
--- a/vernac/comInductive.ml
+++ b/vernac/comInductive.ml
@@ -27,7 +27,6 @@ open Impargs
open Reductionops
open Indtypes
open Pretyping
-open Evarutil
open Indschemes
open Context.Rel.Declaration
open Entries
@@ -158,7 +157,7 @@ let sign_level env evd sign =
| LocalDef _ -> lev, push_rel d env
| LocalAssum _ ->
let s = destSort (Reduction.whd_all env
- (EConstr.Unsafe.to_constr (nf_evar evd (Retyping.get_type_of env evd (EConstr.of_constr (RelDecl.get_type d))))))
+ (EConstr.to_constr evd (Retyping.get_type_of env evd (EConstr.of_constr (RelDecl.get_type d)))))
in
let u = univ_of_sort s in
(Univ.sup u lev, push_rel d env))