aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
-rw-r--r--.github/CODEOWNERS4
-rw-r--r--.gitlab-ci.yml3
-rw-r--r--CHANGES6
-rw-r--r--Makefile.ci3
-rw-r--r--README.md8
-rw-r--r--checker/cic.mli7
-rw-r--r--checker/indtypes.ml10
-rw-r--r--checker/inductive.ml14
-rw-r--r--checker/print.ml4
-rw-r--r--checker/reduction.ml41
-rw-r--r--checker/term.ml20
-rw-r--r--checker/typeops.ml14
-rw-r--r--checker/values.ml4
-rwxr-xr-xdev/ci/ci-basic-overlay.sh6
-rwxr-xr-xdev/ci/ci-bedrock2.sh11
-rw-r--r--dev/ci/user-overlays/07863-rm-sorts-contents.sh6
-rw-r--r--dev/doc/critical-bugs26
-rw-r--r--dev/top_printers.ml8
-rw-r--r--dev/vm_printers.ml4
-rw-r--r--doc/sphinx/language/gallina-extensions.rst8
-rw-r--r--engine/evarutil.ml30
-rw-r--r--engine/evarutil.mli12
-rw-r--r--engine/evd.ml5
-rw-r--r--engine/evd.mli4
-rw-r--r--engine/namegen.ml5
-rw-r--r--engine/termops.ml15
-rw-r--r--ide/wg_ProofView.ml4
-rw-r--r--interp/constrintern.ml10
-rw-r--r--interp/declare.ml21
-rw-r--r--kernel/cbytecodes.ml5
-rw-r--r--kernel/cbytecodes.mli1
-rw-r--r--kernel/cbytegen.ml33
-rw-r--r--kernel/constr.ml10
-rw-r--r--kernel/inductive.ml11
-rw-r--r--kernel/nativevalues.ml2
-rw-r--r--kernel/reduction.ml61
-rw-r--r--kernel/sorts.ml57
-rw-r--r--kernel/sorts.mli7
-rw-r--r--kernel/term.ml5
-rw-r--r--kernel/term.mli5
-rw-r--r--kernel/typeops.ml16
-rw-r--r--kernel/typeops.mli1
-rw-r--r--lib/coqProject_file.ml (renamed from lib/coqProject_file.ml4)64
-rw-r--r--plugins/cc/ccalgo.ml4
-rw-r--r--plugins/ltac/taccoerce.ml9
-rw-r--r--pretyping/coercion.ml4
-rw-r--r--pretyping/constr_matching.ml4
-rw-r--r--pretyping/detyping.ml4
-rw-r--r--pretyping/evarsolve.ml2
-rw-r--r--pretyping/inductiveops.ml22
-rw-r--r--pretyping/patternops.ml4
-rw-r--r--pretyping/retyping.ml13
-rw-r--r--pretyping/typing.ml11
-rw-r--r--tactics/class_tactics.ml2
-rw-r--r--tactics/equality.ml2
-rw-r--r--tactics/hints.ml74
-rw-r--r--tactics/inv.ml2
-rw-r--r--tactics/tactics.ml6
-rw-r--r--test-suite/bugs/closed/5696.v5
-rw-r--r--test-suite/bugs/closed/7723.v58
-rw-r--r--test-suite/bugs/closed/7903.v4
-rw-r--r--test-suite/success/mutual_record.v57
-rw-r--r--test-suite/success/vm_records.v40
-rw-r--r--vernac/comInductive.ml8
-rw-r--r--vernac/record.ml252
-rw-r--r--vernac/record.mli11
-rw-r--r--vernac/vernacentries.ml131
67 files changed, 863 insertions, 457 deletions
diff --git a/.github/CODEOWNERS b/.github/CODEOWNERS
index 3a762b42a..eb2797101 100644
--- a/.github/CODEOWNERS
+++ b/.github/CODEOWNERS
@@ -186,6 +186,10 @@
/pretyping/ @mattam82
# Secondary maintainer @gares
+/pretyping/vnorm.* @maximedenes
+/pretyping/nativenorm.* @maximedenes
+# Secondary maintainer @ppedrot
+
########## Pretty printer ##########
/printing/ @herbelin
diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml
index 971a281ea..4e159ebdc 100644
--- a/.gitlab-ci.yml
+++ b/.gitlab-ci.yml
@@ -306,6 +306,9 @@ validate:edge+flambda:
OPAM_SWITCH: edge
OPAM_VARIANT: "+flambda"
+ci-bedrock2:
+ <<: *ci-template
+
ci-bignums:
<<: *ci-template
diff --git a/CHANGES b/CHANGES
index 6ad2cc548..44186d30c 100644
--- a/CHANGES
+++ b/CHANGES
@@ -1,6 +1,10 @@
Changes from 8.8.2 to 8.9+beta1
===============================
+Kernel
+
+- Mutually defined records are now supported.
+
Tactics
- Added toplevel goal selector ! which expects a single focused goal.
@@ -88,6 +92,8 @@ Kernel
- Fix a critical bug with cofixpoints and vm_compute/native_compute (#7333).
- Fix a critical bug with inlining of polymorphic constants (#7615).
+- Fix a critical bug with universe polymorphism and vm_compute (#7723). Was
+ present since 8.5.
Notations
diff --git a/Makefile.ci b/Makefile.ci
index 7f63157fa..fce16906c 100644
--- a/Makefile.ci
+++ b/Makefile.ci
@@ -8,7 +8,8 @@
## # (see LICENSE file for the text of the license) ##
##########################################################################
-CI_TARGETS=ci-bignums \
+CI_TARGETS=ci-bedrock2 \
+ ci-bignums \
ci-color \
ci-compcert \
ci-coq-dpdgraph \
diff --git a/README.md b/README.md
index df4ca8e40..67f4f6fea 100644
--- a/README.md
+++ b/README.md
@@ -56,3 +56,11 @@ used, and include a complete source example leading to the bug.
## Contributing
Guidelines for contributing to Coq in various ways are listed in the [contributor's guide](CONTRIBUTING.md).
+
+## Supporting Coq
+
+Help the Coq community grow and prosper by becoming a sponsor! The [Coq
+Consortium](https://coq.inria.fr/consortium) can establish sponsorship contracts
+or receive donations. If you want to take an active role in shaping Coq's
+future, you can also become a Consortium member. If you are interested, please
+get in touch!
diff --git a/checker/cic.mli b/checker/cic.mli
index a890f2cef..4846a9af8 100644
--- a/checker/cic.mli
+++ b/checker/cic.mli
@@ -33,11 +33,10 @@ open Names
(** {6 The sorts of CCI. } *)
-type contents = Pos | Null
-
type sorts =
- | Prop of contents (** Prop and Set *)
- | Type of Univ.universe (** Type *)
+ | Prop
+ | Set
+ | Type of Univ.universe
(** {6 The sorts family of CCI. } *)
diff --git a/checker/indtypes.ml b/checker/indtypes.ml
index 916934a81..bcb71fe55 100644
--- a/checker/indtypes.ml
+++ b/checker/indtypes.ml
@@ -107,11 +107,11 @@ let rec sorts_of_constr_args env t =
(* Prop and Set are small *)
let is_small_sort = function
- | Prop _ -> true
+ | Prop | Set -> true
| _ -> false
let is_logic_sort = function
-| Prop Null -> true
+| Prop -> true
| _ -> false
(* [infos] is a sequence of pair [islogic,issmall] for each type in
@@ -186,10 +186,10 @@ let check_predicativity env s small level =
(* (universes env) in *)
if not (Univ.check_leq (universes env) level u) then
failwith "impredicative Type inductive type"
- | Prop Pos, ImpredicativeSet -> ()
- | Prop Pos, _ ->
+ | Set, ImpredicativeSet -> ()
+ | Set, _ ->
if not small then failwith "impredicative Set inductive type"
- | Prop Null,_ -> ()
+ | Prop,_ -> ()
let sort_of_ind = function
diff --git a/checker/inductive.ml b/checker/inductive.ml
index e1c6b135d..b1edec556 100644
--- a/checker/inductive.ml
+++ b/checker/inductive.ml
@@ -101,7 +101,7 @@ let instantiate_params full t u args sign =
substl subs ty
let full_inductive_instantiate mib u params sign =
- let dummy = Prop Null in
+ let dummy = Prop in
let t = mkArity (Term.subst_instance_context u sign,dummy) in
fst (destArity (instantiate_params true t u params mib.mind_params_ctxt))
@@ -137,8 +137,8 @@ Remark: Set (predicative) is encoded as Type(0)
let sort_as_univ = function
| Type u -> u
-| Prop Null -> Univ.type0m_univ
-| Prop Pos -> Univ.type0_univ
+| Prop -> Univ.type0m_univ
+| Set -> Univ.type0_univ
(* cons_subst add the mapping [u |-> su] in subst if [u] is not *)
(* in the domain or add [u |-> sup x su] if [u] is already mapped *)
@@ -195,9 +195,9 @@ let instantiate_universes env ctx ar argsorts =
let level = Univ.subst_univs_universe (Univ.make_subst subst) ar.template_level in
let ty =
(* Singleton type not containing types are interpretable in Prop *)
- if Univ.is_type0m_univ level then Prop Null
+ if Univ.is_type0m_univ level then Prop
(* Non singleton type not containing types are interpretable in Set *)
- else if Univ.is_type0_univ level then Prop Pos
+ else if Univ.is_type0_univ level then Set
(* This is a Type with constraints *)
else Type level
in
@@ -226,8 +226,8 @@ let type_of_inductive env mip =
(* The max of an array of universes *)
let cumulate_constructor_univ u = function
- | Prop Null -> u
- | Prop Pos -> Univ.sup Univ.type0_univ u
+ | Prop -> u
+ | Set -> Univ.sup Univ.type0_univ u
| Type u' -> Univ.sup u u'
let max_inductive_sort =
diff --git a/checker/print.ml b/checker/print.ml
index fc9cd687e..247c811f8 100644
--- a/checker/print.ml
+++ b/checker/print.ml
@@ -57,8 +57,8 @@ let print_pure_constr fmt csr =
fprintf fmt "Proj(%a,@,@[%a@])" sp_con_display (Projection.constant p) pp_term c
and pp_sort fmt = function
- | Prop(Pos) -> pp_print_string fmt "Set"
- | Prop(Null) -> pp_print_string fmt "Prop"
+ | Set -> pp_print_string fmt "Set"
+ | Prop -> pp_print_string fmt "Prop"
| Type u -> fprintf fmt "Type(%a)" chk_pp (Univ.pr_uni u)
and pp_name fmt = function
diff --git a/checker/reduction.ml b/checker/reduction.ml
index 4e508dc77..16c701213 100644
--- a/checker/reduction.ml
+++ b/checker/reduction.ml
@@ -210,29 +210,26 @@ let convert_constructors
let sort_cmp env univ pb s0 s1 =
match (s0,s1) with
- | (Prop c1, Prop c2) when pb = CUMUL -> if c1 = Pos && c2 = Null then raise NotConvertible
- | (Prop c1, Prop c2) -> if c1 <> c2 then raise NotConvertible
- | (Prop c1, Type u) ->
+ | Prop, Prop | Set, Set -> ()
+ | Prop, (Set | Type _) | Set, Type _ ->
+ if not (pb = CUMUL) then raise NotConvertible
+ | Type u1, Type u2 ->
+ (** FIXME: handle type-in-type option here *)
+ if (* snd (engagement env) == StratifiedType && *)
+ not
(match pb with
- CUMUL -> ()
- | _ -> raise NotConvertible)
- | (Type u1, Type u2) ->
- (** FIXME: handle type-in-type option here *)
- if (* snd (engagement env) == StratifiedType && *)
- not
- (match pb with
- | CONV -> Univ.check_eq univ u1 u2
- | CUMUL -> Univ.check_leq univ u1 u2)
- then begin
- if !Flags.debug then begin
- let op = match pb with CONV -> "=" | CUMUL -> "<=" in
- Format.eprintf "sort_cmp: @[%a@]\n%!" Pp.pp_with Pp.(
- str"Error: " ++ Univ.pr_uni u1 ++ str op ++ Univ.pr_uni u2 ++ str ":" ++ cut()
- ++ Univ.pr_universes univ)
- end;
- raise NotConvertible
- end
- | (_, _) -> raise NotConvertible
+ | CONV -> Univ.check_eq univ u1 u2
+ | CUMUL -> Univ.check_leq univ u1 u2)
+ then begin
+ if !Flags.debug then begin
+ let op = match pb with CONV -> "=" | CUMUL -> "<=" in
+ Format.eprintf "sort_cmp: @[%a@]\n%!" Pp.pp_with Pp.(
+ str"Error: " ++ Univ.pr_uni u1 ++ str op ++ Univ.pr_uni u2 ++ str ":" ++ cut()
+ ++ Univ.pr_universes univ)
+ end;
+ raise NotConvertible
+ end
+ | Set, Prop | Type _, (Prop | Set) -> raise NotConvertible
let rec no_arg_available = function
| [] -> true
diff --git a/checker/term.ml b/checker/term.ml
index 509634bdb..d84491b38 100644
--- a/checker/term.ml
+++ b/checker/term.ml
@@ -20,15 +20,15 @@ open Cic
(* Sorts. *)
let family_of_sort = function
- | Prop Null -> InProp
- | Prop Pos -> InSet
+ | Prop -> InProp
+ | Set -> InSet
| Type _ -> InType
let family_equal = (==)
let sort_of_univ u =
- if Univ.is_type0m_univ u then Prop Null
- else if Univ.is_type0_univ u then Prop Pos
+ if Univ.is_type0m_univ u then Prop
+ else if Univ.is_type0_univ u then Set
else Type u
(********************************************************************)
@@ -356,15 +356,11 @@ let rec isArity c =
(* alpha conversion : ignore print names and casts *)
let compare_sorts s1 s2 = match s1, s2 with
-| Prop c1, Prop c2 ->
- begin match c1, c2 with
- | Pos, Pos | Null, Null -> true
- | Pos, Null -> false
- | Null, Pos -> false
- end
+| Prop, Prop | Set, Set -> true
| Type u1, Type u2 -> Univ.Universe.equal u1 u2
-| Prop _, Type _ -> false
-| Type _, Prop _ -> false
+| Prop, Set | Set, Prop -> false
+| (Prop | Set), Type _ -> false
+| Type _, (Prop | Set) -> false
let eq_puniverses f (c1,u1) (c2,u2) =
Univ.Instance.equal u1 u2 && f c1 c2
diff --git a/checker/typeops.ml b/checker/typeops.ml
index 345ee5b8f..19ede4b9a 100644
--- a/checker/typeops.ml
+++ b/checker/typeops.ml
@@ -103,11 +103,11 @@ let judge_of_apply env (f,funj) argjv =
let sort_of_product env domsort rangsort =
match (domsort, rangsort) with
(* Product rule (s,Prop,Prop) *)
- | (_, Prop Null) -> rangsort
+ | _, Prop -> rangsort
(* Product rule (Prop/Set,Set,Set) *)
- | (Prop _, Prop Pos) -> rangsort
+ | (Prop | Set), Set -> rangsort
(* Product rule (Type,Set,?) *)
- | (Type u1, Prop Pos) ->
+ | Type u1, Set ->
if engagement env = ImpredicativeSet then
(* Rule is (Type,Set,Set) in the Set-impredicative calculus *)
rangsort
@@ -115,11 +115,11 @@ let sort_of_product env domsort rangsort =
(* Rule is (Type_i,Set,Type_i) in the Set-predicative calculus *)
Type (Univ.sup u1 Univ.type0_univ)
(* Product rule (Prop,Type_i,Type_i) *)
- | (Prop Pos, Type u2) -> Type (Univ.sup Univ.type0_univ u2)
+ | Set, Type u2 -> Type (Univ.sup Univ.type0_univ u2)
(* Product rule (Prop,Type_i,Type_i) *)
- | (Prop Null, Type _) -> rangsort
+ | Prop, Type _ -> rangsort
(* Product rule (Type_i,Type_i,Type_i) *)
- | (Type u1, Type u2) -> Type (Univ.sup u1 u2)
+ | Type u1, Type u2 -> Type (Univ.sup u1 u2)
(* Type of a type cast *)
@@ -239,7 +239,7 @@ let type_fixpoint env lna lar lbody vdefj =
let rec execute env cstr =
match cstr with
(* Atomic terms *)
- | Sort (Prop _) -> judge_of_prop
+ | Sort (Prop | Set) -> judge_of_prop
| Sort (Type u) -> judge_of_type u
diff --git a/checker/values.ml b/checker/values.ml
index 4f28d6e44..88cdb644d 100644
--- a/checker/values.ml
+++ b/checker/values.ml
@@ -15,7 +15,7 @@
To ensure this file is up-to-date, 'make' now compares the md5 of cic.mli
with a copy we maintain here:
-MD5 42fb0781dc5f7f2cbe3ca127f8249264 checker/cic.mli
+MD5 c395aa2dbfc18794b3b7192f3dc5b2e5 checker/cic.mli
*)
@@ -122,7 +122,7 @@ let v_context_set = v_tuple "universe_context_set" [|v_hset v_level;v_cstrs|]
(** kernel/term *)
-let v_sort = v_sum "sort" 0 [|[|v_enum "cnt" 2|];[|v_univ|]|]
+let v_sort = v_sum "sort" 2 (*Prop, Set*) [|[|v_univ(*Type*)|]|]
let v_sortfam = v_enum "sorts_family" 3
let v_puniverses v = v_tuple "punivs" [|v;v_instance|]
diff --git a/dev/ci/ci-basic-overlay.sh b/dev/ci/ci-basic-overlay.sh
index 2ebbf2cc4..28321d13e 100755
--- a/dev/ci/ci-basic-overlay.sh
+++ b/dev/ci/ci-basic-overlay.sh
@@ -148,6 +148,12 @@
: "${bignums_CI_GITURL:=https://github.com/coq/bignums}"
########################################################################
+# bedrock2
+########################################################################
+: "${bedrock2_CI_BRANCH:=master}"
+: "${bedrock2_CI_GITURL:=https://github.com/mit-plv/bedrock2}"
+
+########################################################################
# Equations
########################################################################
: "${Equations_CI_BRANCH:=master}"
diff --git a/dev/ci/ci-bedrock2.sh b/dev/ci/ci-bedrock2.sh
new file mode 100755
index 000000000..447076e09
--- /dev/null
+++ b/dev/ci/ci-bedrock2.sh
@@ -0,0 +1,11 @@
+#!/usr/bin/env bash
+
+ci_dir="$(dirname "$0")"
+. "${ci_dir}/ci-common.sh"
+
+bedrock2_CI_DIR="${CI_BUILD_DIR}/bedrock2"
+
+git_checkout "${bedrock2_CI_BRANCH}" "${bedrock2_CI_GITURL}" "${bedrock2_CI_DIR}"
+( cd "${bedrock2_CI_DIR}" && git submodule update --init --recursive )
+
+( cd "${bedrock2_CI_DIR}" && make )
diff --git a/dev/ci/user-overlays/07863-rm-sorts-contents.sh b/dev/ci/user-overlays/07863-rm-sorts-contents.sh
new file mode 100644
index 000000000..374a61484
--- /dev/null
+++ b/dev/ci/user-overlays/07863-rm-sorts-contents.sh
@@ -0,0 +1,6 @@
+#!/bin/sh
+
+if [ "$CI_PULL_REQUEST" = "7863" ] || [ "$CI_BRANCH" = "rm-sorts-contents" ]; then
+ Elpi_CI_BRANCH=fix-sorts-contents
+ Elpi_CI_GITURL=https://github.com/SkySkimmer/coq-elpi.git
+fi
diff --git a/dev/doc/critical-bugs b/dev/doc/critical-bugs
index 293b01f63..6166d24b7 100644
--- a/dev/doc/critical-bugs
+++ b/dev/doc/critical-bugs
@@ -81,6 +81,20 @@ Module system
GH issue number: #4294
risk: ?
+Module system
+
+ component: modules, universes
+ summary: universe constraints for module subtyping not stored in vo files
+ introduced: presumably 8.2 (b3d3b56)
+ impacted released versions: 8.2, 8.3, 8.4
+ impacted development branches: v8.5
+ impacted coqchk versions: none
+ fixed in: v8.2 (c1d9889), v8.3 (8056d02), v8.4 (a07deb4), trunk (0cd0a3e) Mar 5, 2014, Tassi
+ found by: Tassi by running coqchk on the mathematical components library
+ exploit: requires multiple files, no test provided
+ GH issue number: #3243
+ risk: could be exploited by mistake
+
Universes
component: template polymorphism
@@ -123,6 +137,18 @@ Primitive projections
Conversion machines
+ component: "lazy machine" (lazy krivine abstract machine)
+ summary: the invariant justifying some optimization was wrong for some combination of sharing side effects
+ introduced: prior to V7.0
+ impacted released versions: V8.0-V8.0pl4, V8.1-V8.1pl3
+ impacted development branches: none
+ impacted coqchk versions: (eefe63d52, Barras, 20 May 2008), was in beta-development for 8.2 at this time
+ fixed in: master/trunk/8.2 (f13aaec57/a8b034513, 15 May 2008, Barras), v8.1 (e7611477a, 15 May 2008, Barras), v8.0 (6ed40a8bc, 29 Nov 2016, Herbelin, backport)
+ found by: Gonthier
+ exploit: by Gonthier
+ GH issue number: none
+ risk: unrealistic to be exploited by chance
+
component: "virtual machine" (compilation to bytecode ran by a C-interpreter)
summary: collision between constructors when more than 256 constructors in a type
introduced: V8.1
diff --git a/dev/top_printers.ml b/dev/top_printers.ml
index 844ad9188..811abd67f 100644
--- a/dev/top_printers.ml
+++ b/dev/top_printers.ml
@@ -301,8 +301,8 @@ let constr_display csr =
incr cnt; pp (str "with " ++ int !cnt ++ str" " ++ Level.pr u ++ fnl ())
and sort_display = function
- | Prop(Pos) -> "Prop(Pos)"
- | Prop(Null) -> "Prop(Null)"
+ | Set -> "Set"
+ | Prop -> "Prop"
| Type u -> univ_display u;
"Type("^(string_of_int !cnt)^")"
@@ -423,8 +423,8 @@ let print_pure_constr csr =
Array.iter (fun u -> print_space (); pp (Level.pr u)) (Instance.to_array u)
and sort_display = function
- | Prop(Pos) -> print_string "Set"
- | Prop(Null) -> print_string "Prop"
+ | Set -> print_string "Set"
+ | Prop -> print_string "Prop"
| Type u -> open_hbox();
print_string "Type("; pp (pr_uni u); print_string ")"; close_box()
diff --git a/dev/vm_printers.ml b/dev/vm_printers.ml
index 7589e5348..c8385da61 100644
--- a/dev/vm_printers.ml
+++ b/dev/vm_printers.ml
@@ -26,8 +26,8 @@ let print_vfix_app () = print_string "vfix_app"
let print_vswith () = print_string "switch"
let ppsort = function
- | Prop(Pos) -> print_string "Set"
- | Prop(Null) -> print_string "Prop"
+ | Set -> print_string "Set"
+ | Prop -> print_string "Prop"
| Type u -> print_string "Type"
diff --git a/doc/sphinx/language/gallina-extensions.rst b/doc/sphinx/language/gallina-extensions.rst
index c21d8d4ec..294ac5799 100644
--- a/doc/sphinx/language/gallina-extensions.rst
+++ b/doc/sphinx/language/gallina-extensions.rst
@@ -23,8 +23,9 @@ expressions. In this sense, the :cmd:`Record` construction allows defining
.. _record_grammar:
.. productionlist:: `sentence`
- record : `record_keyword` `ident` [ `binders` ] [: `sort` ] := [ `ident` ] { [ `field` ; … ; `field` ] }.
+ record : `record_keyword` `record_body` with … with `record_body`
record_keyword : Record | Inductive | CoInductive
+ record_body : `ident` [ `binders` ] [: `sort` ] := [ `ident` ] { [ `field` ; … ; `field` ] }.
field : `ident` [ `binders` ] : `type` [ where `notation` ]
: | `ident` [ `binders` ] [: `type` ] := `term`
@@ -167,12 +168,13 @@ and the syntax `term.(@qualid` |term_1| |term_n| `)` to `@qualid` |term_1| `…`
In each case, `term` is the object projected and the
other arguments are the parameters of the inductive type.
+
.. note:: Records defined with the ``Record`` keyword are not allowed to be
recursive (references to the record's name in the type of its field
raises an error). To define recursive records, one can use the ``Inductive``
and ``CoInductive`` keywords, resulting in an inductive or co-inductive record.
- A *caveat*, however, is that records cannot appear in mutually inductive
- (or co-inductive) definitions.
+ Definition of mutal inductive or co-inductive records are also allowed, as long
+ as all of the types in the block are records.
.. note:: Induction schemes are automatically generated for inductive records.
Automatic generation of induction schemes for non-recursive records
diff --git a/engine/evarutil.ml b/engine/evarutil.ml
index 1625f6fc8..0c044f20d 100644
--- a/engine/evarutil.ml
+++ b/engine/evarutil.ml
@@ -426,10 +426,6 @@ let push_rel_context_to_named_context ?hypnaming env sigma typ =
let default_source = Loc.tag @@ Evar_kinds.InternalHole
-let restrict_evar evd evk filter ?src candidates =
- let evd, evk' = Evd.restrict evk filter ?candidates ?src evd in
- Evd.declare_future_goal evk' evd, evk'
-
let new_pure_evar_full evd evi =
let (evd, evk) = Evd.new_evar evd evi in
let evd = Evd.declare_future_goal evk evd in
@@ -547,11 +543,33 @@ let generalize_evar_over_rels sigma (ev,args) =
type clear_dependency_error =
| OccurHypInSimpleClause of Id.t option
| EvarTypingBreak of existential
+| NoCandidatesLeft of Evar.t
exception ClearDependencyError of Id.t * clear_dependency_error * GlobRef.t option
exception Depends of Id.t
+let set_of_evctx l =
+ List.fold_left (fun s decl -> Id.Set.add (NamedDecl.get_id decl) s) Id.Set.empty l
+
+let filter_effective_candidates evd evi filter candidates =
+ let ids = set_of_evctx (Filter.filter_list filter (evar_context evi)) in
+ List.filter (fun a -> Id.Set.subset (collect_vars evd a) ids) candidates
+
+let restrict_evar evd evk filter ?src candidates =
+ let evar_info = Evd.find_undefined evd evk in
+ let candidates = Option.map (filter_effective_candidates evd evar_info filter) candidates in
+ match candidates with
+ | Some [] -> raise (ClearDependencyError (*FIXME*)(Id.of_string "blah", (NoCandidatesLeft evk), None))
+ | _ ->
+ let evd, evk' = Evd.restrict evk filter ?candidates ?src evd in
+ (** Mark new evar as future goal, removing previous one,
+ circumventing Proofview.advance but making Proof.run_tactic catch these. *)
+ let future_goals = Evd.save_future_goals evd in
+ let future_goals = Evd.filter_future_goals (fun evk' -> not (Evar.equal evk evk')) future_goals in
+ let evd = Evd.restore_future_goals evd future_goals in
+ (Evd.declare_future_goal evk' evd, evk')
+
let rec check_and_clear_in_constr env evdref err ids global c =
(* returns a new constr where all the evars have been 'cleaned'
(ie the hypotheses ids have been removed from the contexts of
@@ -621,7 +639,9 @@ let rec check_and_clear_in_constr env evdref err ids global c =
let origfilter = Evd.evar_filter evi in
let filter = Evd.Filter.apply_subfilter origfilter filter in
let evd = !evdref in
- let (evd,_) = restrict_evar evd evk filter None in
+ let candidates = Evd.evar_candidates evi in
+ let candidates = Option.map (List.map EConstr.of_constr) candidates in
+ let (evd,_) = restrict_evar evd evk filter candidates in
evdref := evd;
Evd.existential_value0 !evdref ev
diff --git a/engine/evarutil.mli b/engine/evarutil.mli
index db638be9e..8ce1b625f 100644
--- a/engine/evarutil.mli
+++ b/engine/evarutil.mli
@@ -65,9 +65,6 @@ val new_type_evar :
val new_Type : ?rigid:rigid -> env -> evar_map -> evar_map * constr
-val restrict_evar : evar_map -> Evar.t -> Filter.t ->
- ?src:Evar_kinds.t Loc.located -> constr list option -> evar_map * Evar.t
-
(** Polymorphic constants *)
val new_global : evar_map -> GlobRef.t -> evar_map * constr
@@ -231,9 +228,18 @@ raise OccurHypInSimpleClause if the removal breaks dependencies *)
type clear_dependency_error =
| OccurHypInSimpleClause of Id.t option
| EvarTypingBreak of Constr.existential
+| NoCandidatesLeft of Evar.t
exception ClearDependencyError of Id.t * clear_dependency_error * GlobRef.t option
+(** Restrict an undefined evar according to a (sub)filter and candidates.
+ The evar will be defined if there is only one candidate left,
+@raise ClearDependencyError NoCandidatesLeft if the filter turns the candidates
+ into an empty list. *)
+
+val restrict_evar : evar_map -> Evar.t -> Filter.t ->
+ ?src:Evar_kinds.t Loc.located -> constr list option -> evar_map * Evar.t
+
val clear_hyps_in_evi : env -> evar_map -> named_context_val -> types ->
Id.Set.t -> evar_map * named_context_val * types
diff --git a/engine/evd.ml b/engine/evd.ml
index 714a0b645..761ae7983 100644
--- a/engine/evd.ml
+++ b/engine/evd.ml
@@ -171,6 +171,8 @@ let evar_context evi = named_context_of_val evi.evar_hyps
let evar_filtered_context evi =
Filter.filter_list (evar_filter evi) (evar_context evi)
+let evar_candidates evi = evi.evar_candidates
+
let evar_hyps evi = evi.evar_hyps
let evar_filtered_hyps evi = match Filter.repr (evar_filter evi) with
@@ -620,6 +622,7 @@ let merge_universe_context evd uctx' =
let set_universe_context evd uctx' =
{ evd with universes = uctx' }
+(* TODO: make unique *)
let add_conv_pb ?(tail=false) pb d =
if tail then {d with conv_pbs = d.conv_pbs @ [pb]}
else {d with conv_pbs = pb::d.conv_pbs}
@@ -852,7 +855,7 @@ let normalize_universe_instance evd l =
let normalize_sort evars s =
match s with
- | Prop _ -> s
+ | Prop | Set -> s
| Type u ->
let u' = normalize_universe evars u in
if u' == u then s else Type u'
diff --git a/engine/evd.mli b/engine/evd.mli
index d166fd804..64db70451 100644
--- a/engine/evd.mli
+++ b/engine/evd.mli
@@ -113,6 +113,7 @@ val evar_filtered_context : evar_info -> (econstr, etypes) Context.Named.pt
val evar_hyps : evar_info -> named_context_val
val evar_filtered_hyps : evar_info -> named_context_val
val evar_body : evar_info -> evar_body
+val evar_candidates : evar_info -> constr list option
val evar_filter : evar_info -> Filter.t
val evar_env : evar_info -> env
val evar_filtered_env : evar_info -> env
@@ -243,7 +244,8 @@ val evars_reset_evd : ?with_conv_pbs:bool -> ?with_univs:bool ->
val restrict : Evar.t-> Filter.t -> ?candidates:econstr list ->
?src:Evar_kinds.t located -> evar_map -> evar_map * Evar.t
(** Restrict an undefined evar into a new evar by filtering context and
- possibly limiting the instances to a set of candidates *)
+ possibly limiting the instances to a set of candidates (candidates
+ are filtered according to the filter) *)
val is_restricted_evar : evar_info -> Evar.t option
(** Tell if an evar comes from restriction of another evar, and if yes, which *)
diff --git a/engine/namegen.ml b/engine/namegen.ml
index 23c691139..978f33b68 100644
--- a/engine/namegen.ml
+++ b/engine/namegen.ml
@@ -137,8 +137,9 @@ let lowercase_first_char id = (* First character of a constr *)
s ^ Unicode.lowercase_first_char s'
let sort_hdchar = function
- | Prop(_) -> "P"
- | Type(_) -> "T"
+ | Prop -> "P"
+ | Set -> "S"
+ | Type _ -> "T"
let hdchar env sigma c =
let rec hdrec k c =
diff --git a/engine/termops.ml b/engine/termops.ml
index 2db2e07bf..2b179c43b 100644
--- a/engine/termops.ml
+++ b/engine/termops.ml
@@ -25,8 +25,8 @@ module CompactedDecl = Context.Compacted.Declaration
(* Sorts and sort family *)
let print_sort = function
- | Prop Pos -> (str "Set")
- | Prop Null -> (str "Prop")
+ | Set -> (str "Set")
+ | Prop -> (str "Prop")
| Type u -> (str "Type(" ++ Univ.Universe.pr u ++ str ")")
let pr_sort_family = function
@@ -1162,15 +1162,14 @@ let is_template_polymorphic env sigma f =
let base_sort_cmp pb s0 s1 =
match (s0,s1) with
- | (Prop c1, Prop c2) -> c1 == Null || c2 == Pos (* Prop <= Set *)
- | (Prop c1, Type u) -> pb == Reduction.CUMUL
- | (Type u1, Type u2) -> true
- | _ -> false
+ | Prop, Prop | Set, Set | Type _, Type _ -> true
+ | Prop, Set | Prop, Type _ | Set, Type _ -> pb == Reduction.CUMUL
+ | Set, Prop | Type _, Prop | Type _, Set -> false
let rec is_Prop sigma c = match EConstr.kind sigma c with
| Sort u ->
begin match EConstr.ESorts.kind sigma u with
- | Prop Null -> true
+ | Prop -> true
| _ -> false
end
| Cast (c,_,_) -> is_Prop sigma c
@@ -1179,7 +1178,7 @@ let rec is_Prop sigma c = match EConstr.kind sigma c with
let rec is_Set sigma c = match EConstr.kind sigma c with
| Sort u ->
begin match EConstr.ESorts.kind sigma u with
- | Prop Pos -> true
+ | Set -> true
| _ -> false
end
| Cast (c,_,_) -> is_Set sigma c
diff --git a/ide/wg_ProofView.ml b/ide/wg_ProofView.ml
index b3088ee28..9be562d3e 100644
--- a/ide/wg_ProofView.ml
+++ b/ide/wg_ProofView.ml
@@ -103,7 +103,7 @@ let mode_tactic sel_cb (proof : #GText.view_skel) goals ~unfoc_goals hints = mat
else []
in
proof#buffer#insert (goal_str ~shownum:true 1 goals_cnt);
- insert_xml proof#buffer (Richpp.richpp_of_pp width cur_goal);
+ insert_xml ~tags:[Tags.Proof.goal] proof#buffer (Richpp.richpp_of_pp width cur_goal);
proof#buffer#insert "\n"
in
(* Insert remaining goals (no hypotheses) *)
@@ -128,7 +128,7 @@ let mode_tactic sel_cb (proof : #GText.view_skel) goals ~unfoc_goals hints = mat
ignore(proof#buffer#place_cursor
~where:(proof#buffer#end_iter#backward_to_tag_toggle
(Some Tags.Proof.goal)));
- ignore(proof#scroll_to_mark ~use_align:true ~yalign:0.95 `INSERT)
+ ignore(proof#scroll_to_mark `INSERT)
let rec flatten = function
| [] -> []
diff --git a/interp/constrintern.ml b/interp/constrintern.ml
index 4e217b2cd..18d6c1a5b 100644
--- a/interp/constrintern.ml
+++ b/interp/constrintern.ml
@@ -710,10 +710,12 @@ let instantiate_notation_constr loc intern intern_pat ntnvars subst infos c =
let arg = match arg with
| None -> None
| Some arg ->
- let mk_env (c, (tmp_scope, subscopes)) =
+ let mk_env id (c, (tmp_scope, subscopes)) map =
let nenv = {env with tmp_scope; scopes = subscopes @ env.scopes} in
- let gc = intern nenv c in
- (gc, Some c)
+ try
+ let gc = intern nenv c in
+ Id.Map.add id (gc, Some c) map
+ with GlobalizationError _ -> map
in
let mk_env' (c, (onlyident,(tmp_scope,subscopes))) =
let nenv = {env with tmp_scope; scopes = subscopes @ env.scopes} in
@@ -725,7 +727,7 @@ let instantiate_notation_constr loc intern intern_pat ntnvars subst infos c =
| [pat] -> (glob_constr_of_cases_pattern pat, None)
| _ -> error_cannot_coerce_disjunctive_pattern_term ?loc:c.loc ()
in
- let terms = Id.Map.map mk_env terms in
+ let terms = Id.Map.fold mk_env terms Id.Map.empty in
let binders = Id.Map.map mk_env' binders in
let bindings = Id.Map.fold Id.Map.add terms binders in
Some (Genintern.generic_substitute_notation bindings arg)
diff --git a/interp/declare.ml b/interp/declare.ml
index e79cc6079..fcb62ac8c 100644
--- a/interp/declare.ml
+++ b/interp/declare.ml
@@ -383,13 +383,12 @@ let inInductive : inductive_obj -> obj =
rebuild_function = infer_inductive_subtyping }
let declare_projections univs mind =
- (** FIXME: handle mutual records *)
- let mind = (mind, 0) in
let env = Global.env () in
- let spec,_ = Inductive.lookup_mind_specif env mind in
- match spec.mind_record with
- | PrimRecord info ->
- let _, kns, _ = info.(0) in
+ let mib = Environ.lookup_mind mind env in
+ match mib.mind_record with
+ | PrimRecord info ->
+ let iter i (_, kns, _) =
+ let mind = (mind, i) in
let projs = Inductiveops.compute_projections env mind in
Array.iter2 (fun kn (term, types) ->
let id = Label.to_id (Constant.label kn) in
@@ -411,10 +410,12 @@ let declare_projections univs mind =
let entry = definition_entry ~types ~univs term in
let kn' = declare_constant id (DefinitionEntry entry, IsDefinition StructureComponent) in
assert (Constant.equal kn kn')
- ) kns projs;
- true, true
- | FakeRecord -> true,false
- | NotRecord -> false,false
+ ) kns projs
+ in
+ let () = Array.iteri iter info in
+ true, true
+ | FakeRecord -> true, false
+ | NotRecord -> false, false
(* for initial declaration *)
let declare_mind mie =
diff --git a/kernel/cbytecodes.ml b/kernel/cbytecodes.ml
index 521f540d2..3095ce148 100644
--- a/kernel/cbytecodes.ml
+++ b/kernel/cbytecodes.ml
@@ -217,6 +217,7 @@ type vm_env = {
type comp_env = {
+ arity : int; (* arity of the current function, 0 if none *)
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 *)
@@ -235,8 +236,8 @@ open Util
let pp_sort s =
let open Sorts in
match s with
- | Prop Null -> str "Prop"
- | Prop Pos -> str "Set"
+ | Prop -> str "Prop"
+ | Set -> str "Set"
| Type u -> str "Type@{" ++ Univ.pr_uni u ++ str "}"
let rec pp_struct_const = function
diff --git a/kernel/cbytecodes.mli b/kernel/cbytecodes.mli
index 238edc0af..de21401b3 100644
--- a/kernel/cbytecodes.mli
+++ b/kernel/cbytecodes.mli
@@ -159,6 +159,7 @@ type vm_env = {
type comp_env = {
+ arity : int; (* arity of the current function, 0 if none *)
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 *)
diff --git a/kernel/cbytegen.ml b/kernel/cbytegen.ml
index 7a27a3d20..6677db2fd 100644
--- a/kernel/cbytegen.ml
+++ b/kernel/cbytegen.ml
@@ -112,8 +112,9 @@ let push_fv d e = {
let fv r = !(r.in_env)
-let empty_comp_env ?(univs=0) ()=
- { nb_uni_stack = univs;
+let empty_comp_env ()=
+ { arity = 0;
+ nb_uni_stack = 0;
nb_stack = 0;
in_stack = [];
nb_rec = 0;
@@ -148,7 +149,8 @@ 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 ?(univs=0) arity =
- { nb_uni_stack = univs ;
+ { arity;
+ nb_uni_stack = univs ;
nb_stack = arity;
in_stack = add_param arity 0 [];
nb_rec = 0;
@@ -159,7 +161,8 @@ let comp_env_fun ?(univs=0) arity =
let comp_env_fix_type rfv =
- { nb_uni_stack = 0;
+ { arity = 0;
+ nb_uni_stack = 0;
nb_stack = 0;
in_stack = [];
nb_rec = 0;
@@ -173,7 +176,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_uni_stack = 0;
+ { arity;
+ nb_uni_stack = 0;
nb_stack = arity;
in_stack = add_param arity 0 [];
nb_rec = ndef;
@@ -183,7 +187,8 @@ let comp_env_fix ndef curr_pos arity rfv =
}
let comp_env_cofix_type ndef rfv =
- { nb_uni_stack = 0;
+ { arity = 0;
+ nb_uni_stack = 0;
nb_stack = 0;
in_stack = [];
nb_rec = 0;
@@ -197,7 +202,8 @@ let comp_env_cofix ndef arity rfv =
for i = 1 to ndef do
prec := Kenvacc i :: !prec
done;
- { nb_uni_stack = 0;
+ { arity;
+ nb_uni_stack = 0;
nb_stack = arity;
in_stack = add_param arity 0 [];
nb_rec = ndef;
@@ -249,8 +255,15 @@ let pos_rel i r sz =
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))
+ (* Compilation of a universe variable can happen either at toplevel (the
+ current closure correspond to a constant and has local universes) or in a
+ local closure (which has no local universes). *)
+ if r.nb_uni_stack != 0 then
+ (* Universe variables are represented by De Bruijn levels (not indices),
+ starting at 0. The shape of the stack will be [v1|..|vn|u1..up|arg1..argq]
+ with size = n + p + q, and q = r.arity. So Kacc (sz - r.arity - 1) will access
+ the last universe. *)
+ Kacc (sz - r.arity - (r.nb_uni_stack - i))
else
let env = !(r.in_env) in
let db = FVuniv_var i in
@@ -498,7 +511,7 @@ let rec compile_lam env cenv lam sz cont =
else comp_app compile_structured_constant compile_universe cenv
(Const_ind ind) (Univ.Instance.to_array u) sz cont
- | Lsort (Sorts.Prop _ as s) ->
+ | Lsort (Sorts.Prop | Sorts.Set as s) ->
compile_structured_constant cenv (Const_sort s) sz cont
| Lsort (Sorts.Type u) ->
(* We represent universes as a global constant with local universes
diff --git a/kernel/constr.ml b/kernel/constr.ml
index e68f906ec..45812b5a1 100644
--- a/kernel/constr.ml
+++ b/kernel/constr.ml
@@ -130,8 +130,8 @@ let mkProp = Sort Sorts.prop
let mkSet = Sort Sorts.set
let mkType u = Sort (Sorts.Type u)
let mkSort = function
- | Sorts.Prop Sorts.Null -> mkProp (* Easy sharing *)
- | Sorts.Prop Sorts.Pos -> mkSet
+ | Sorts.Prop -> mkProp (* Easy sharing *)
+ | Sorts.Set -> mkSet
| s -> Sort s
(* Constructs the term t1::t2, i.e. the term t1 casted with the type t2 *)
@@ -260,17 +260,17 @@ let isSort c = match kind c with
| _ -> false
let rec isprop c = match kind c with
- | Sort (Sorts.Prop _) -> true
+ | Sort (Sorts.Prop | Sorts.Set) -> true
| Cast (c,_,_) -> isprop c
| _ -> false
let rec is_Prop c = match kind c with
- | Sort (Sorts.Prop Sorts.Null) -> true
+ | Sort Sorts.Prop -> true
| Cast (c,_,_) -> is_Prop c
| _ -> false
let rec is_Set c = match kind c with
- | Sort (Sorts.Prop Sorts.Pos) -> true
+ | Sort Sorts.Set -> true
| Cast (c,_,_) -> is_Set c
| _ -> false
diff --git a/kernel/inductive.ml b/kernel/inductive.ml
index 9130b8778..584c1af03 100644
--- a/kernel/inductive.ml
+++ b/kernel/inductive.ml
@@ -130,11 +130,6 @@ where
Remark: Set (predicative) is encoded as Type(0)
*)
-let sort_as_univ = let open Sorts in function
-| Type u -> u
-| Prop Null -> Universe.type0m
-| Prop Pos -> Universe.type0
-
(* Template polymorphism *)
(* cons_subst add the mapping [u |-> su] in subst if [u] is not *)
@@ -168,7 +163,7 @@ let make_subst env =
(* arity is a global level which, at typing time, will be enforce *)
(* 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
+ let s = Sorts.univ_of_sort (snd (dest_arity env (Lazy.force a))) in
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 *)
@@ -236,8 +231,8 @@ let type_of_inductive_knowing_parameters env ?(polyprop=true) mip args =
(* The max of an array of universes *)
let cumulate_constructor_univ u = let open Sorts in function
- | Prop Null -> u
- | Prop Pos -> Universe.sup Universe.type0 u
+ | Prop -> u
+ | Set -> Universe.sup Universe.type0 u
| Type u' -> Universe.sup u u'
let max_inductive_sort =
diff --git a/kernel/nativevalues.ml b/kernel/nativevalues.ml
index da4413a0a..3901cb9ce 100644
--- a/kernel/nativevalues.ml
+++ b/kernel/nativevalues.ml
@@ -116,7 +116,7 @@ let mk_ind_accu ind u =
let mk_sort_accu s u =
let open Sorts in
match s with
- | Prop _ -> mk_accu (Asort s)
+ | Prop | Set -> mk_accu (Asort s)
| Type s ->
let u = Univ.Instance.of_array u in
let s = Univ.subst_instance_universe u s in
diff --git a/kernel/reduction.ml b/kernel/reduction.ml
index 2c61b7a01..3228a155f 100644
--- a/kernel/reduction.ml
+++ b/kernel/reduction.ml
@@ -649,23 +649,19 @@ let check_leq univs u u' =
let check_sort_cmp_universes env pb s0 s1 univs =
let open Sorts in
if not (type_in_type env) then
+ let check_pb u0 u1 =
+ match pb with
+ | CUMUL -> check_leq univs u0 u1
+ | CONV -> check_eq univs u0 u1
+ in
match (s0,s1) with
- | (Prop c1, Prop c2) when is_cumul pb ->
- begin match c1, c2 with
- | Null, _ | _, Pos -> () (* Prop <= Set *)
- | _ -> raise NotConvertible
- end
- | (Prop c1, Prop c2) -> if c1 != c2 then raise NotConvertible
- | (Prop c1, Type u) ->
- let u0 = univ_of_sort s0 in
- (match pb with
- | CUMUL -> check_leq univs u0 u
- | CONV -> check_eq univs u0 u)
- | (Type u, Prop c) -> raise NotConvertible
- | (Type u1, Type u2) ->
- (match pb with
- | CUMUL -> check_leq univs u1 u2
- | CONV -> check_eq univs u1 u2)
+ | Prop, Prop | Set, Set -> ()
+ | Prop, (Set | Type _) -> if not (is_cumul pb) then raise NotConvertible
+ | Set, Prop -> raise NotConvertible
+ | Set, Type u -> check_pb Univ.type0_univ u
+ | Type u, Prop -> raise NotConvertible
+ | Type u, Set -> check_pb u Univ.type0_univ
+ | Type u0, Type u1 -> check_pb u0 u1
let checked_sort_cmp_universes env pb s0 s1 univs =
check_sort_cmp_universes env pb s0 s1 univs; univs
@@ -697,26 +693,23 @@ let infer_leq (univs, cstrs as cuniv) u u' =
univs, Univ.Constraint.union cstrs cstrs'
let infer_cmp_universes env pb s0 s1 univs =
- let open Sorts in
- if type_in_type env then univs
+ if type_in_type env
+ then univs
else
+ let open Sorts in
+ let infer_pb u0 u1 =
+ match pb with
+ | CUMUL -> infer_leq univs u0 u1
+ | CONV -> infer_eq univs u0 u1
+ in
match (s0,s1) with
- | (Prop c1, Prop c2) when is_cumul pb ->
- begin match c1, c2 with
- | Null, _ | _, Pos -> univs (* Prop <= Set *)
- | _ -> raise NotConvertible
- end
- | (Prop c1, Prop c2) -> if c1 == c2 then univs else raise NotConvertible
- | (Prop c1, Type u) ->
- let u0 = univ_of_sort s0 in
- (match pb with
- | CUMUL -> infer_leq univs u0 u
- | CONV -> infer_eq univs u0 u)
- | (Type u, Prop c) -> raise NotConvertible
- | (Type u1, Type u2) ->
- (match pb with
- | CUMUL -> infer_leq univs u1 u2
- | CONV -> infer_eq univs u1 u2)
+ | Prop, Prop | Set, Set -> univs
+ | Prop, (Set | Type _) -> if not (is_cumul pb) then raise NotConvertible else univs
+ | Set, Prop -> raise NotConvertible
+ | Set, Type u -> infer_pb Univ.type0_univ u
+ | Type u, Prop -> raise NotConvertible
+ | Type u, Set -> infer_pb u Univ.type0_univ
+ | Type u0, Type u1 -> infer_pb u0 u1
let infer_convert_instances ~flex u u' (univs,cstrs) =
let cstrs' =
diff --git a/kernel/sorts.ml b/kernel/sorts.ml
index daeb90be7..a7bb08f5b 100644
--- a/kernel/sorts.ml
+++ b/kernel/sorts.ml
@@ -10,22 +10,21 @@
open Univ
-type contents = Pos | Null
-
type family = InProp | InSet | InType
type t =
- | Prop of contents (* proposition types *)
+ | Prop
+ | Set
| Type of Universe.t
-let prop = Prop Null
-let set = Prop Pos
+let prop = Prop
+let set = Set
let type1 = Type type1_univ
let univ_of_sort = function
| Type u -> u
- | Prop Pos -> Universe.type0
- | Prop Null -> Universe.type0m
+ | Set -> Universe.type0
+ | Prop -> Universe.type0m
let sort_of_univ u =
if is_type0m_univ u then prop
@@ -34,36 +33,34 @@ let sort_of_univ u =
let compare s1 s2 =
if s1 == s2 then 0 else
- match s1, s2 with
- | Prop c1, Prop c2 ->
- begin match c1, c2 with
- | Pos, Pos | Null, Null -> 0
- | Pos, Null -> -1
- | Null, Pos -> 1
- end
- | Type u1, Type u2 -> Universe.compare u1 u2
- | Prop _, Type _ -> -1
- | Type _, Prop _ -> 1
+ match s1, s2 with
+ | Prop, Prop -> 0
+ | Prop, _ -> -1
+ | Set, Prop -> 1
+ | Set, Set -> 0
+ | Set, _ -> -1
+ | Type u1, Type u2 -> Universe.compare u1 u2
+ | Type _, _ -> -1
let equal s1 s2 = Int.equal (compare s1 s2) 0
let is_prop = function
- | Prop Null -> true
+ | Prop -> true
| Type u when Universe.equal Universe.type0m u -> true
| _ -> false
let is_set = function
- | Prop Pos -> true
+ | Set -> true
| Type u when Universe.equal Universe.type0 u -> true
| _ -> false
let is_small = function
- | Prop _ -> true
+ | Prop | Set -> true
| Type u -> is_small_univ u
let family = function
- | Prop Null -> InProp
- | Prop Pos -> InSet
+ | Prop -> InProp
+ | Set -> InSet
| Type u when is_type0m_univ u -> InProp
| Type u when is_type0_univ u -> InSet
| Type _ -> InType
@@ -73,15 +70,11 @@ let family_equal = (==)
open Hashset.Combine
let hash = function
-| Prop p ->
- let h = match p with
- | Pos -> 0
- | Null -> 1
- in
- combinesmall 1 h
-| Type u ->
- let h = Univ.Universe.hash u in
- combinesmall 2 h
+ | Prop -> combinesmall 1 0
+ | Set -> combinesmall 1 1
+ | Type u ->
+ let h = Univ.Universe.hash u in
+ combinesmall 2 h
module List = struct
let mem = List.memq
@@ -101,7 +94,7 @@ module Hsorts =
if u' == u then c else Type u'
| s -> s
let eq s1 s2 = match (s1,s2) with
- | (Prop c1, Prop c2) -> c1 == c2
+ | Prop, Prop | Set, Set -> true
| (Type u1, Type u2) -> u1 == u2
|_ -> false
diff --git a/kernel/sorts.mli b/kernel/sorts.mli
index 1bbde2608..cac6229b9 100644
--- a/kernel/sorts.mli
+++ b/kernel/sorts.mli
@@ -10,13 +10,12 @@
(** {6 The sorts of CCI. } *)
-type contents = Pos | Null
-
type family = InProp | InSet | InType
type t =
-| Prop of contents (** Prop and Set *)
-| Type of Univ.Universe.t (** Type *)
+ | Prop
+ | Set
+ | Type of Univ.Universe.t
val set : t
val prop : t
diff --git a/kernel/term.ml b/kernel/term.ml
index b44e038e9..81e344e73 100644
--- a/kernel/term.ml
+++ b/kernel/term.ml
@@ -16,14 +16,11 @@ open Vars
open Constr
(* Deprecated *)
-type contents = Sorts.contents = Pos | Null
-[@@ocaml.deprecated "Alias for Sorts.contents"]
-
type sorts_family = Sorts.family = InProp | InSet | InType
[@@ocaml.deprecated "Alias for Sorts.family"]
type sorts = Sorts.t =
- | Prop of Sorts.contents (** Prop and Set *)
+ | Prop | Set
| Type of Univ.Universe.t (** Type *)
[@@ocaml.deprecated "Alias for Sorts.t"]
diff --git a/kernel/term.mli b/kernel/term.mli
index f651d1a58..4d340399d 100644
--- a/kernel/term.mli
+++ b/kernel/term.mli
@@ -190,13 +190,10 @@ type ('constr, 'types) kind_of_type =
val kind_of_type : types -> (constr, types) kind_of_type
(* Deprecated *)
-type contents = Sorts.contents = Pos | Null
-[@@ocaml.deprecated "Alias for Sorts.contents"]
-
type sorts_family = Sorts.family = InProp | InSet | InType
[@@ocaml.deprecated "Alias for Sorts.family"]
type sorts = Sorts.t =
- | Prop of Sorts.contents (** Prop and Set *)
+ | Prop | Set
| Type of Univ.Universe.t (** Type *)
[@@ocaml.deprecated "Alias for Sorts.t"]
diff --git a/kernel/typeops.ml b/kernel/typeops.ml
index 34ed2afb2..7c0057696 100644
--- a/kernel/typeops.ml
+++ b/kernel/typeops.ml
@@ -69,7 +69,7 @@ let type_of_type u =
mkType uu
let type_of_sort = function
- | Prop c -> type1
+ | Prop | Set -> type1
| Type u -> type_of_type u
(*s Type of a de Bruijn index. *)
@@ -178,11 +178,11 @@ let type_of_apply env func funt argsv argstv =
let sort_of_product env domsort rangsort =
match (domsort, rangsort) with
(* Product rule (s,Prop,Prop) *)
- | (_, Prop Null) -> rangsort
+ | (_, Prop) -> rangsort
(* Product rule (Prop/Set,Set,Set) *)
- | (Prop _, Prop Pos) -> rangsort
+ | ((Prop | Set), Set) -> rangsort
(* Product rule (Type,Set,?) *)
- | (Type u1, Prop Pos) ->
+ | (Type u1, Set) ->
if is_impredicative_set env then
(* Rule is (Type,Set,Set) in the Set-impredicative calculus *)
rangsort
@@ -190,9 +190,9 @@ let sort_of_product env domsort rangsort =
(* Rule is (Type_i,Set,Type_i) in the Set-predicative calculus *)
Type (Universe.sup Universe.type0 u1)
(* Product rule (Prop,Type_i,Type_i) *)
- | (Prop Pos, Type u2) -> Type (Universe.sup Universe.type0 u2)
+ | (Set, Type u2) -> Type (Universe.sup Universe.type0 u2)
(* Product rule (Prop,Type_i,Type_i) *)
- | (Prop Null, Type _) -> rangsort
+ | (Prop, Type _) -> rangsort
(* Product rule (Type_i,Type_i,Type_i) *)
| (Type u1, Type u2) -> Type (Universe.sup u1 u2)
@@ -481,10 +481,6 @@ let judge_of_prop = make_judge mkProp type1
let judge_of_set = make_judge mkSet type1
let judge_of_type u = make_judge (mkType u) (type_of_type u)
-let judge_of_prop_contents = function
- | Null -> judge_of_prop
- | Pos -> judge_of_set
-
let judge_of_relative env k = make_judge (mkRel k) (type_of_relative env k)
let judge_of_variable env x = make_judge (mkVar x) (type_of_variable env x)
diff --git a/kernel/typeops.mli b/kernel/typeops.mli
index 546f2d2b4..3b2abc777 100644
--- a/kernel/typeops.mli
+++ b/kernel/typeops.mli
@@ -43,7 +43,6 @@ val type1 : types
val type_of_sort : Sorts.t -> types
val judge_of_prop : unsafe_judgment
val judge_of_set : unsafe_judgment
-val judge_of_prop_contents : Sorts.contents -> unsafe_judgment
val judge_of_type : Universe.t -> unsafe_judgment
(** {6 Type of a bound variable. } *)
diff --git a/lib/coqProject_file.ml4 b/lib/coqProject_file.ml
index 61eb1dafd..c2bcd73ff 100644
--- a/lib/coqProject_file.ml4
+++ b/lib/coqProject_file.ml
@@ -90,30 +90,50 @@ let rec post_canonize f =
exception Parsing_error of string
-let rec parse_string = parser
- | [< '' ' | '\n' | '\t' >] -> ""
- | [< 'c; s >] -> (String.make 1 c)^(parse_string s)
- | [< >] -> ""
-and parse_string2 = parser
- | [< ''"' >] -> ""
- | [< 'c; s >] -> (String.make 1 c)^(parse_string2 s)
- | [< >] -> raise (Parsing_error "unterminated string")
-and parse_skip_comment = parser
- | [< ''\n'; s >] -> s
- | [< 'c; s >] -> parse_skip_comment s
- | [< >] -> [< >]
-and parse_args = parser
- | [< '' ' | '\n' | '\t'; s >] -> parse_args s
- | [< ''#'; s >] -> parse_args (parse_skip_comment s)
- | [< ''"'; str = parse_string2; s >] -> ("" ^ str) :: parse_args s
- | [< 'c; str = parse_string; s >] -> ((String.make 1 c) ^ str) :: (parse_args s)
- | [< >] -> []
+let buffer buf =
+ let ans = Buffer.contents buf in
+ let () = Buffer.clear buf in
+ ans
+
+let rec parse_string buf s = match Stream.next s with
+| ' ' | '\n' | '\t' -> buffer buf
+| c ->
+ let () = Buffer.add_char buf c in
+ parse_string buf s
+| exception Stream.Failure -> buffer buf
+
+and parse_string2 buf s = match Stream.next s with
+| '"' -> buffer buf
+| c ->
+ let () = Buffer.add_char buf c in
+ parse_string2 buf s
+| exception Stream.Failure -> raise (Parsing_error "unterminated string")
+
+and parse_skip_comment s = match Stream.next s with
+| '\n' -> ()
+| _ -> parse_skip_comment s
+| exception Stream.Failure -> ()
+
+and parse_args buf accu s = match Stream.next s with
+| ' ' | '\n' | '\t' -> parse_args buf accu s
+| '#' ->
+ let () = parse_skip_comment s in
+ parse_args buf accu s
+| '"' ->
+ let str = parse_string2 buf s in
+ parse_args buf (str :: accu) s
+| c ->
+ let () = Buffer.add_char buf c in
+ let str = parse_string buf s in
+ parse_args buf (str :: accu) s
+| exception Stream.Failure -> accu
let parse f =
let c = open_in f in
- let res = parse_args (Stream.of_channel c) in
+ let buf = Buffer.create 64 in
+ let res = parse_args buf [] (Stream.of_channel c) in
close_in c;
- res
+ List.rev res
;;
(* Copy from minisys.ml, since we don't see that file here *)
@@ -143,7 +163,7 @@ let process_cmd_line orig_dir proj args =
error "Use \"-install none\" instead of \"-no-install\""
| "-custom" :: _ ->
error "Use \"-extra[-phony] target deps command\" instead of \"-custom command deps target\""
-
+
| ("-no-opt"|"-byte") :: r -> aux { proj with use_ocamlopt = false } r
| ("-full"|"-opt") :: r -> aux { proj with use_ocamlopt = true } r
| "-install" :: d :: r ->
@@ -189,7 +209,7 @@ let process_cmd_line orig_dir proj args =
error "Output file must be in the current directory";
if proj.makefile <> None then
error "Option -o given more than once";
- aux { proj with makefile = Some file } r
+ aux { proj with makefile = Some file } r
| v :: "=" :: def :: r ->
aux { proj with defs = proj.defs @ [sourced (v,def)] } r
| "-arg" :: a :: r ->
diff --git a/plugins/cc/ccalgo.ml b/plugins/cc/ccalgo.ml
index 4c6156a38..4a691e442 100644
--- a/plugins/cc/ccalgo.ml
+++ b/plugins/cc/ccalgo.ml
@@ -130,8 +130,8 @@ type cinfo=
ci_nhyps: int} (* # projectable args *)
let family_eq f1 f2 = match f1, f2 with
- | Prop Pos, Prop Pos
- | Prop Null, Prop Null
+ | Set, Set
+ | Prop, Prop
| Type _, Type _ -> true
| _ -> false
diff --git a/plugins/ltac/taccoerce.ml b/plugins/ltac/taccoerce.ml
index cc9c2046d..84baea964 100644
--- a/plugins/ltac/taccoerce.ml
+++ b/plugins/ltac/taccoerce.ml
@@ -199,11 +199,12 @@ let id_of_name = function
let basename = Nametab.basename_of_global ref in
basename
| Sort s ->
- begin
+ begin
match ESorts.kind sigma s with
- | Sorts.Prop _ -> Label.to_id (Label.make "Prop")
- | Sorts.Type _ -> Label.to_id (Label.make "Type")
- end
+ | Sorts.Prop -> Label.to_id (Label.make "Prop")
+ | Sorts.Set -> Label.to_id (Label.make "Set")
+ | Sorts.Type _ -> Label.to_id (Label.make "Type")
+ end
| _ -> fail()
diff --git a/pretyping/coercion.ml b/pretyping/coercion.ml
index bf9e37aa7..5c4cbefad 100644
--- a/pretyping/coercion.ml
+++ b/pretyping/coercion.ml
@@ -209,8 +209,8 @@ and coerce ?loc env evdref (x : EConstr.constr) (y : EConstr.constr)
match (EConstr.kind !evdref x, EConstr.kind !evdref y) with
| Sort s, Sort s' ->
(match ESorts.kind !evdref s, ESorts.kind !evdref s' with
- | Prop x, Prop y when x == y -> None
- | Prop _, Type _ -> None
+ | Prop, Prop | Set, Set -> None
+ | (Prop | Set), Type _ -> None
| Type x, Type y when Univ.Universe.equal x y -> None (* false *)
| _ -> subco ())
| Prod (name, a, b), Prod (name', a', b') ->
diff --git a/pretyping/constr_matching.ml b/pretyping/constr_matching.ml
index 2bc603a90..d7118efd7 100644
--- a/pretyping/constr_matching.ml
+++ b/pretyping/constr_matching.ml
@@ -281,8 +281,8 @@ let matches_core env sigma allow_bound_rels
let open Glob_term in
begin match ps, ESorts.kind sigma s with
- | GProp, Prop Null -> subst
- | GSet, Prop Pos -> subst
+ | GProp, Prop -> subst
+ | GSet, Set -> subst
| GType _, Type _ -> subst
| _ -> raise PatternMatchingFailure
end
diff --git a/pretyping/detyping.ml b/pretyping/detyping.ml
index 23a985dc3..d0de2f8c0 100644
--- a/pretyping/detyping.ml
+++ b/pretyping/detyping.ml
@@ -597,8 +597,8 @@ let detype_universe sigma u =
Univ.Universe.map fn u
let detype_sort sigma = function
- | Prop Null -> GProp
- | Prop Pos -> GSet
+ | Prop -> GProp
+ | Set -> GSet
| Type u ->
GType
(if !print_universes
diff --git a/pretyping/evarsolve.ml b/pretyping/evarsolve.ml
index 8afb9b942..3f5d186d4 100644
--- a/pretyping/evarsolve.ml
+++ b/pretyping/evarsolve.ml
@@ -69,7 +69,7 @@ let refresh_universes ?(status=univ_rigid) ?(onlyalg=false) ?(refreshset=false)
if onlyalg && alg then
(evdref := Evd.make_flexible_variable !evdref ~algebraic:false l; t)
else t))
- | Prop Pos when refreshset && not direction ->
+ | Set when refreshset && not direction ->
(* Cannot make a universe "lower" than "Set",
only refreshing when we want higher universes. *)
refresh_sort status ~direction s
diff --git a/pretyping/inductiveops.ml b/pretyping/inductiveops.ml
index d599afe69..efb205182 100644
--- a/pretyping/inductiveops.ml
+++ b/pretyping/inductiveops.ml
@@ -465,22 +465,21 @@ let build_branch_type env sigma dep p cs =
let compute_projections env (kn, i as ind) =
let open Term in
let mib = Environ.lookup_mind kn env in
- let indu = match mib.mind_universes with
- | Monomorphic_ind _ -> mkInd ind
- | Polymorphic_ind ctx -> mkIndU (ind, make_abstract_instance ctx)
- | Cumulative_ind ctx ->
- mkIndU (ind, make_abstract_instance (ACumulativityInfo.univ_context ctx))
+ let u = match mib.mind_universes with
+ | Monomorphic_ind _ -> Instance.empty
+ | Polymorphic_ind auctx -> make_abstract_instance auctx
+ | Cumulative_ind acumi ->
+ make_abstract_instance (ACumulativityInfo.univ_context acumi)
in
let x = match mib.mind_record with
| NotRecord | FakeRecord ->
anomaly Pp.(str "Trying to build primitive projections for a non-primitive record")
| PrimRecord info-> Name (pi1 (info.(i)))
in
- (** FIXME: handle mutual records *)
- let pkt = mib.mind_packets.(0) in
- let { mind_consnrealargs; mind_consnrealdecls } = pkt in
+ let pkt = mib.mind_packets.(i) in
let { mind_nparams = nparamargs; mind_params_ctxt = params } = mib in
- let rctx, _ = decompose_prod_assum (subst1 indu pkt.mind_nf_lc.(0)) in
+ let subst = List.init mib.mind_ntypes (fun i -> mkIndU ((kn, mib.mind_ntypes - i - 1), u)) in
+ let rctx, _ = decompose_prod_assum (substl subst pkt.mind_nf_lc.(0)) in
let ctx, paramslet = List.chop pkt.mind_consnrealdecls.(0) rctx in
let mp, dp, l = MutInd.repr3 kn in
(** We build a substitution smashing the lets in the record parameters so
@@ -489,6 +488,7 @@ let compute_projections env (kn, i as ind) =
let indty =
(* [ty] = [Ind inst] is typed in context [params] *)
let inst = Context.Rel.to_extended_vect mkRel 0 paramslet in
+ let indu = mkIndU (ind, u) in
let ty = mkApp (indu, inst) in
(* [Ind inst] is typed in context [params-wo-let] *)
ty
@@ -498,8 +498,8 @@ let compute_projections env (kn, i as ind) =
{ ind_tags = []; cstr_tags = [|Context.Rel.to_tags ctx|]; style = LetStyle } in
{ ci_ind = ind;
ci_npar = nparamargs;
- ci_cstr_ndecls = mind_consnrealdecls;
- ci_cstr_nargs = mind_consnrealargs;
+ ci_cstr_ndecls = pkt.mind_consnrealdecls;
+ ci_cstr_nargs = pkt.mind_consnrealargs;
ci_pp_info = print_info }
in
let len = List.length ctx in
diff --git a/pretyping/patternops.ml b/pretyping/patternops.ml
index 622a8e982..685aa400b 100644
--- a/pretyping/patternops.ml
+++ b/pretyping/patternops.ml
@@ -150,8 +150,8 @@ let pattern_of_constr env sigma t =
| Rel n -> PRel n
| Meta n -> PMeta (Some (Id.of_string ("META" ^ string_of_int n)))
| Var id -> PVar id
- | Sort (Prop Null) -> PSort GProp
- | Sort (Prop Pos) -> PSort GSet
+ | Sort Prop -> PSort GProp
+ | Sort Set -> PSort GSet
| Sort (Type _) -> PSort (GType [])
| Cast (c,_,_) -> pattern_of_constr env c
| LetIn (na,c,t,b) -> PLetIn (na,pattern_of_constr env c,Some (pattern_of_constr env t),
diff --git a/pretyping/retyping.ml b/pretyping/retyping.ml
index 746a68b21..7e43c5e41 100644
--- a/pretyping/retyping.ml
+++ b/pretyping/retyping.ml
@@ -149,18 +149,13 @@ let retype ?(polyprop=true) sigma =
| Cast (c,_, s) when isSort sigma s -> destSort sigma s
| Sort s ->
begin match ESorts.kind sigma s with
- | Prop _ -> Sorts.type1
+ | Prop | Set -> Sorts.type1
| Type u -> Type (Univ.super u)
end
| Prod (name,t,c2) ->
- (match (sort_of env t, sort_of (push_rel (LocalAssum (name,t)) env) c2) with
- | _, (Prop Null as s) -> s
- | Prop _, (Prop Pos as s) -> s
- | Type _, (Prop Pos as s) when is_impredicative_set env -> s
- | Type u1, Prop Pos -> Type (Univ.sup u1 Univ.type0_univ)
- | Prop Pos, (Type u2) -> Type (Univ.sup Univ.type0_univ u2)
- | Prop Null, (Type _ as s) -> s
- | Type u1, Type u2 -> Type (Univ.sup u1 u2))
+ let dom = sort_of env t in
+ let rang = sort_of (push_rel (LocalAssum (name,t)) env) c2 in
+ Typeops.sort_of_product env dom rang
| App(f,args) when is_template_polymorphic env sigma f ->
let t = type_of_global_reference_knowing_parameters env f args in
sort_of_atomic_type env sigma t args
diff --git a/pretyping/typing.ml b/pretyping/typing.ml
index cf34ac016..a66ecaaac 100644
--- a/pretyping/typing.ml
+++ b/pretyping/typing.ml
@@ -241,10 +241,6 @@ let judge_of_set =
{ uj_val = EConstr.mkSet;
uj_type = EConstr.mkSort Sorts.type1 }
-let judge_of_prop_contents = function
- | Null -> judge_of_prop
- | Pos -> judge_of_set
-
let judge_of_type u =
let uu = Univ.Universe.super u in
{ uj_val = EConstr.mkType u;
@@ -333,10 +329,9 @@ let rec execute env sigma cstr =
| Sort s ->
begin match ESorts.kind sigma s with
- | Prop c ->
- sigma, judge_of_prop_contents c
- | Type u ->
- sigma, judge_of_type u
+ | Prop -> sigma, judge_of_prop
+ | Set -> sigma, judge_of_set
+ | Type u -> sigma, judge_of_type u
end
| Proj (p, c) ->
diff --git a/tactics/class_tactics.ml b/tactics/class_tactics.ml
index 773fc1520..9c5fdcd1c 100644
--- a/tactics/class_tactics.ml
+++ b/tactics/class_tactics.ml
@@ -477,7 +477,7 @@ let is_Prop env sigma concl =
match EConstr.kind sigma ty with
| Sort s ->
begin match ESorts.kind sigma s with
- | Prop Null -> true
+ | Prop -> true
| _ -> false
end
| _ -> false
diff --git a/tactics/equality.ml b/tactics/equality.ml
index 91c577405..0e3921570 100644
--- a/tactics/equality.ml
+++ b/tactics/equality.ml
@@ -942,7 +942,7 @@ let rec build_discriminator env sigma true_0 false_0 dirn c = function
let (cnum_nlams,cnum_env,kont) = descend_then env sigma c cnum in
let newc = mkRel(cnum_nlams-argnum) in
let subval = build_discriminator cnum_env sigma true_0 false_0 dirn newc l in
- kont sigma subval (false_0,mkSort (Prop Null))
+ kont sigma subval (false_0,mkProp)
(* Note: discrimination could be more clever: if some elimination is
not allowed because of a large impredicative constructor in the
diff --git a/tactics/hints.ml b/tactics/hints.ml
index 85ff02824..d28d4848c 100644
--- a/tactics/hints.ml
+++ b/tactics/hints.ml
@@ -828,38 +828,48 @@ let make_exact_entry env sigma info poly ?(name=PathAny) (c, cty, ctx) =
let make_apply_entry env sigma (eapply,hnf,verbose) info poly ?(name=PathAny) (c, cty, ctx) =
let cty = if hnf then hnf_constr env sigma cty else cty in
- match EConstr.kind sigma cty with
- | Prod _ ->
- let sigma' = Evd.merge_context_set univ_flexible sigma ctx in
- let ce = mk_clenv_from_env env sigma' None (c,cty) in
- let c' = clenv_type (* ~reduce:false *) ce in
- let pat = Patternops.pattern_of_constr env ce.evd (EConstr.to_constr ~abort_on_undefined_evars:false sigma c') in
- let hd =
- try head_pattern_bound pat
- with BoundPattern -> failwith "make_apply_entry" in
- let nmiss = List.length (clenv_missing ce) in
- let secvars = secvars_of_constr env sigma c in
- let pri = match info.hint_priority with None -> nb_hyp sigma' cty + nmiss | Some p -> p in
- let pat = match info.hint_pattern with
- | Some p -> snd p | None -> pat
- in
- if Int.equal nmiss 0 then
- (Some hd,
- { pri; poly; pat = Some pat; name;
- db = None;
- secvars;
- code = with_uid (Res_pf(c,cty,ctx)); })
- else begin
- if not eapply then failwith "make_apply_entry";
- if verbose then
- Feedback.msg_info (str "the hint: eapply " ++ pr_leconstr_env env sigma' c ++
- str " will only be used by eauto");
- (Some hd,
- { pri; poly; pat = Some pat; name;
- db = None; secvars;
- code = with_uid (ERes_pf(c,cty,ctx)); })
- end
- | _ -> failwith "make_apply_entry"
+ match EConstr.kind sigma cty with
+ | Prod _ ->
+ let sigma' = Evd.merge_context_set univ_flexible sigma ctx in
+ let ce = mk_clenv_from_env env sigma' None (c,cty) in
+ let c' = clenv_type (* ~reduce:false *) ce in
+ let pat = Patternops.pattern_of_constr env ce.evd (EConstr.to_constr ~abort_on_undefined_evars:false sigma c') in
+ let hd =
+ try head_pattern_bound pat
+ with BoundPattern -> failwith "make_apply_entry" in
+ let miss = clenv_missing ce in
+ let nmiss = List.length miss in
+ let secvars = secvars_of_constr env sigma c in
+ let pri = match info.hint_priority with None -> nb_hyp sigma' cty + nmiss | Some p -> p in
+ let pat = match info.hint_pattern with
+ | Some p -> snd p | None -> pat
+ in
+ if Int.equal nmiss 0 then
+ (Some hd,
+ { pri; poly; pat = Some pat; name;
+ db = None;
+ secvars;
+ code = with_uid (Res_pf(c,cty,ctx)); })
+ else begin
+ if not eapply then failwith "make_apply_entry";
+ if verbose then begin
+ let variables = str (CString.plural nmiss "variable") in
+ Feedback.msg_info (
+ strbrk "The hint " ++
+ pr_leconstr_env env sigma' c ++
+ strbrk " will only be used by eauto, because applying " ++
+ pr_leconstr_env env sigma' c ++
+ strbrk " would leave " ++ variables ++ Pp.spc () ++
+ Pp.prlist_with_sep Pp.pr_comma Name.print (List.map (Evd.meta_name ce.evd) miss) ++
+ strbrk " as unresolved existential " ++ variables ++ str "."
+ )
+ end;
+ (Some hd,
+ { pri; poly; pat = Some pat; name;
+ db = None; secvars;
+ code = with_uid (ERes_pf(c,cty,ctx)); })
+ end
+ | _ -> failwith "make_apply_entry"
(* flags is (e,h,v) with e=true if eapply and h=true if hnf and v=true if verbose
c is a constr
diff --git a/tactics/inv.ml b/tactics/inv.ml
index 755494c2d..e467eacd9 100644
--- a/tactics/inv.ml
+++ b/tactics/inv.ml
@@ -495,7 +495,7 @@ let raw_inversion inv_kind id status names =
(* Error messages of the inversion tactics *)
let wrap_inv_error id = function (e, info) -> match e with
| Indrec.RecursionSchemeError
- (Indrec.NotAllowedCaseAnalysis (_,(Type _ | Prop Pos as k),i)) ->
+ (Indrec.NotAllowedCaseAnalysis (_,(Type _ | Set as k),i)) ->
Proofview.tclENV >>= fun env ->
Proofview.tclEVARMAP >>= fun sigma ->
tclZEROMSG (
diff --git a/tactics/tactics.ml b/tactics/tactics.ml
index c430edf2e..928530744 100644
--- a/tactics/tactics.ml
+++ b/tactics/tactics.ml
@@ -212,6 +212,9 @@ let clear_dependency_msg env sigma id err inglobal =
str "Cannot remove " ++ Id.print id ++
strbrk " without breaking the typing of " ++
Printer.pr_existential env sigma ev ++ str"."
+ | Evarutil.NoCandidatesLeft ev ->
+ str "Cannot remove " ++ Id.print id ++ str " as it would leave the existential " ++
+ Printer.pr_existential_key sigma ev ++ str" without candidates."
let error_clear_dependency env sigma id err inglobal =
user_err (clear_dependency_msg env sigma id err inglobal)
@@ -228,6 +231,9 @@ let replacing_dependency_msg env sigma id err inglobal =
str "Cannot change " ++ Id.print id ++
strbrk " without breaking the typing of " ++
Printer.pr_existential env sigma ev ++ str"."
+ | Evarutil.NoCandidatesLeft ev ->
+ str "Cannot change " ++ Id.print id ++ str " as it would leave the existential " ++
+ Printer.pr_existential_key sigma ev ++ str" without candidates."
let error_replacing_dependency env sigma id err inglobal =
user_err (replacing_dependency_msg env sigma id err inglobal)
diff --git a/test-suite/bugs/closed/5696.v b/test-suite/bugs/closed/5696.v
new file mode 100644
index 000000000..a20ad1b4d
--- /dev/null
+++ b/test-suite/bugs/closed/5696.v
@@ -0,0 +1,5 @@
+(* Slightly improving interpretation of Ltac subterms in notations *)
+
+Notation "'var2' x .. y = z ; e" := (ltac:(exact z), (fun x => .. (fun y => e)
+..)) (at level 200, x binder, y binder, e at level 220).
+Check (var2 a = 1; a).
diff --git a/test-suite/bugs/closed/7723.v b/test-suite/bugs/closed/7723.v
new file mode 100644
index 000000000..216290123
--- /dev/null
+++ b/test-suite/bugs/closed/7723.v
@@ -0,0 +1,58 @@
+Set Universe Polymorphism.
+
+Module Segfault.
+
+Inductive decision_tree : Type := .
+
+Fixpoint first_satisfying_helper {A B} (f : A -> option B) (ls : list A) : option B
+ := match ls with
+ | nil => None
+ | cons x xs
+ => match f x with
+ | Some v => Some v
+ | None => first_satisfying_helper f xs
+ end
+ end.
+
+Axiom admit : forall {T}, T.
+Definition dtree4 : option decision_tree :=
+ match first_satisfying_helper (fun pat : nat => Some pat) (cons 0 nil)
+ with
+ | Some _ => admit
+ | None => admit
+ end
+.
+Definition dtree'' := Eval vm_compute in dtree4. (* segfault *)
+
+End Segfault.
+
+Module OtherExample.
+
+Definition bar@{i} := Type@{i}.
+Definition foo@{i j} (x y z : nat) :=
+ @id Type@{j} bar@{i}.
+Eval vm_compute in foo.
+
+End OtherExample.
+
+Module LocalClosure.
+
+Definition bar@{i} := Type@{i}.
+
+Definition foo@{i j} (x y z : nat) :=
+ @id (nat -> Type@{j}) (fun _ => Type@{i}).
+Eval vm_compute in foo.
+
+End LocalClosure.
+
+Require Import Hurkens.
+Polymorphic Inductive unit := tt.
+
+Polymorphic Definition foo :=
+ let x := id tt in (x, x, Type).
+
+Lemma bad : False.
+ refine (TypeNeqSmallType.paradox (snd foo) _).
+ vm_compute.
+ Fail reflexivity.
+Abort.
diff --git a/test-suite/bugs/closed/7903.v b/test-suite/bugs/closed/7903.v
new file mode 100644
index 000000000..55c7ee99a
--- /dev/null
+++ b/test-suite/bugs/closed/7903.v
@@ -0,0 +1,4 @@
+(* Slightly improving interpretation of Ltac subterms in notations *)
+
+Notation bar x f := (let z := ltac:(exact 1) in (fun x : nat => f)).
+Check bar x (x + x).
diff --git a/test-suite/success/mutual_record.v b/test-suite/success/mutual_record.v
new file mode 100644
index 000000000..77529733b
--- /dev/null
+++ b/test-suite/success/mutual_record.v
@@ -0,0 +1,57 @@
+Module M0.
+
+Inductive foo (A : Type) := Foo {
+ foo0 : option (bar A);
+ foo1 : nat;
+ foo2 := foo1 = 0;
+ foo3 : foo2;
+}
+
+with bar (A : Type) := Bar {
+ bar0 : A;
+ bar1 := 0;
+ bar2 : bar1 = 0;
+ bar3 : nat -> foo A;
+}.
+
+End M0.
+
+Module M1.
+
+Set Primitive Projections.
+
+Inductive foo (A : Type) := Foo {
+ foo0 : option (bar A);
+ foo1 : nat;
+ foo2 := foo1 = 0;
+ foo3 : foo2;
+}
+
+with bar (A : Type) := Bar {
+ bar0 : A;
+ bar1 := 0;
+ bar2 : bar1 = 0;
+ bar3 : nat -> foo A;
+}.
+
+End M1.
+
+Module M2.
+
+Set Primitive Projections.
+
+CoInductive foo (A : Type) := Foo {
+ foo0 : option (bar A);
+ foo1 : nat;
+ foo2 := foo1 = 0;
+ foo3 : foo2;
+}
+
+with bar (A : Type) := Bar {
+ bar0 : A;
+ bar1 := 0;
+ bar2 : bar1 = 0;
+ bar3 : nat -> foo A;
+}.
+
+End M2.
diff --git a/test-suite/success/vm_records.v b/test-suite/success/vm_records.v
new file mode 100644
index 000000000..8a1544c8e
--- /dev/null
+++ b/test-suite/success/vm_records.v
@@ -0,0 +1,40 @@
+Set Primitive Projections.
+
+Module M.
+
+CoInductive foo := Foo {
+ foo0 : foo;
+ foo1 : bar;
+}
+with bar := Bar {
+ bar0 : foo;
+ bar1 : bar;
+}.
+
+CoFixpoint f : foo := Foo f g
+with g : bar := Bar f g.
+
+Check (@eq_refl _ g.(bar0) <: f.(foo0).(foo0) = g.(bar0)).
+Check (@eq_refl _ g <: g.(bar1).(bar0).(foo1) = g).
+
+End M.
+
+Module N.
+
+Inductive foo := Foo {
+ foo0 : option foo;
+ foo1 : list bar;
+}
+with bar := Bar {
+ bar0 : option bar;
+ bar1 : list foo;
+}.
+
+Definition f_0 := Foo None nil.
+Definition g_0 := Bar None nil.
+
+Definition f := Foo (Some f_0) (cons g_0 nil).
+
+Check (@eq_refl _ f.(foo1) <: f.(foo1) = cons g_0 nil).
+
+End N.
diff --git a/vernac/comInductive.ml b/vernac/comInductive.ml
index 6057c05f5..0387b32ba 100644
--- a/vernac/comInductive.ml
+++ b/vernac/comInductive.ml
@@ -126,7 +126,7 @@ let make_conclusion_flexible sigma ty poly =
else sigma
let is_impredicative env u =
- u = Prop Null || (is_impredicative_set env && u = Prop Pos)
+ u = Prop || (is_impredicative_set env && u = Set)
let interp_ind_arity env sigma ind =
let c = intern_gen IsType env sigma ind.ind_arity in
@@ -245,7 +245,7 @@ let solve_constraints_system levels level_bounds =
let inductive_levels env evd poly arities inds =
let destarities = List.map (fun x -> x, Reduction.dest_arity env x) arities in
let levels = List.map (fun (x,(ctx,a)) ->
- if a = Prop Null then None
+ if a = Prop then None
else Some (univ_of_sort a)) destarities
in
let cstrs_levels, min_levels, sizes =
@@ -298,14 +298,14 @@ let inductive_levels env evd poly arities inds =
(** "Polymorphic" type constraint and more than one constructor,
should not land in Prop. Add constraint only if it would
land in Prop directly (no informative arguments as well). *)
- Evd.set_leq_sort env evd (Prop Pos) du
+ Evd.set_leq_sort env evd Set du
else evd
in
let duu = Sorts.univ_of_sort du in
let evd =
if not (Univ.is_small_univ duu) && Univ.Universe.equal cu duu then
if is_flexible_sort evd duu && not (Evd.check_leq evd Univ.type0_univ duu) then
- Evd.set_eq_sort env evd (Prop Null) du
+ Evd.set_eq_sort env evd Prop du
else evd
else Evd.set_eq_sort env evd (Type cu) du
in
diff --git a/vernac/record.ml b/vernac/record.ml
index 202c9b92f..7a8ce7d25 100644
--- a/vernac/record.ml
+++ b/vernac/record.ml
@@ -100,7 +100,7 @@ let binder_of_decl = function
let binders_of_decls = List.map binder_of_decl
-let typecheck_params_and_fields finite def id poly pl t ps nots fs =
+let typecheck_params_and_fields finite def poly pl ps records =
let env0 = Global.env () in
let sigma, decl = Constrexpr_ops.interp_univ_decl_opt env0 pl in
let _ =
@@ -117,7 +117,7 @@ let typecheck_params_and_fields finite def id poly pl t ps nots fs =
Loc.raise ?loc (Stream.Error "pattern with quote not allowed in record parameters")) ps
in
let sigma, (impls_env, ((env1,newps), imps)) = interp_context_evars env0 sigma ps in
- let sigma, typ, sort, template = match t with
+ let fold (sigma, template) (_, t, _, _) = match t with
| Some t ->
let env = EConstr.push_rel_context newps env0 in
let poly =
@@ -132,28 +132,36 @@ let typecheck_params_and_fields finite def id poly pl t ps nots fs =
match Evd.is_sort_variable sigma s' with
| Some l ->
let sigma = Evd.make_flexible_variable sigma ~algebraic:true l in
- sigma, s, s', true
+ (sigma, template), (s, s')
| None ->
- sigma, s, s', false
- else sigma, s, s', false)
+ (sigma, false), (s, s')
+ else (sigma, false), (s, s'))
| _ -> user_err ?loc:(constr_loc t) (str"Sort expected."))
| None ->
let uvarkind = Evd.univ_flexible_alg in
let sigma, s = Evd.new_sort_variable uvarkind sigma in
- sigma, EConstr.mkSort s, s, true
+ (sigma, template), (EConstr.mkSort s, s)
in
- let arity = EConstr.it_mkProd_or_LetIn typ newps in
- let env_ar = EConstr.push_rel_context newps (EConstr.push_rel (LocalAssum (Name id,arity)) env0) in
+ let (sigma, template), typs = List.fold_left_map fold (sigma, true) records in
+ let arities = List.map (fun (typ, _) -> EConstr.it_mkProd_or_LetIn typ newps) typs in
+ let fold accu (id, _, _, _) arity = EConstr.push_rel (LocalAssum (Name id,arity)) accu in
+ let env_ar = EConstr.push_rel_context newps (List.fold_left2 fold env0 records arities) in
let assums = List.filter is_local_assum newps in
- let params = List.map (RelDecl.get_name %> Name.get_id) assums in
- let ty = Inductive (params,(finite != Declarations.BiFinite)) in
- let impls_env = compute_internalization_env env0 sigma ~impls:impls_env ty [id] [arity] [imps] in
- let env2,sigma,impls,newfs,data =
- interp_fields_evars env_ar sigma impls_env nots (binders_of_decls fs)
+ let impls_env =
+ let params = List.map (RelDecl.get_name %> Name.get_id) assums in
+ let ty = Inductive (params, (finite != Declarations.BiFinite)) in
+ let ids = List.map (fun (id, _, _, _) -> id) records in
+ let imps = List.map (fun _ -> imps) arities in
+ compute_internalization_env env0 sigma ~impls:impls_env ty ids arities imps
in
+ let fold sigma (_, _, nots, fs) arity =
+ let _, sigma, impls, newfs, _ = interp_fields_evars env_ar sigma impls_env nots (binders_of_decls fs) in
+ (sigma, (impls, newfs))
+ in
+ let (sigma, data) = List.fold_left2_map fold sigma records arities in
let sigma =
Pretyping.solve_remaining_evars Pretyping.all_and_fail_flags env_ar sigma (Evd.from_env env_ar) in
- let sigma, typ =
+ let fold sigma (typ, sort) (_, newfs) =
let _, univ = compute_constructor_level sigma env_ar newfs in
if not def && (Sorts.is_prop sort ||
(Sorts.is_set sort && is_impredicative_set env0)) then
@@ -164,20 +172,24 @@ let typecheck_params_and_fields finite def id poly pl t ps nots fs =
Option.cata (Evd.is_flexible_level sigma) false (Evd.is_sort_variable sigma sort) then
(* We can assume that the level in aritysort is not constrained
and clear it, if it is flexible *)
- Evd.set_eq_sort env_ar sigma (Prop Pos) sort,
- EConstr.mkSort (Sorts.sort_of_univ univ)
+ Evd.set_eq_sort env_ar sigma Set sort, EConstr.mkSort (Sorts.sort_of_univ univ)
else sigma, typ
in
+ let (sigma, typs) = List.fold_left2_map fold sigma typs data in
let sigma = Evd.minimize_universes sigma in
- let newfs = List.map (EConstr.to_rel_decl sigma) newfs in
let newps = List.map (EConstr.to_rel_decl sigma) newps in
- let typ = EConstr.to_constr sigma typ in
- let ce t = Pretyping.check_evars env0 (Evd.from_env env0) sigma (EConstr.of_constr t) in
let univs = Evd.check_univ_decl ~poly sigma decl in
let ubinders = Evd.universe_binders sigma in
- List.iter (iter_constr ce) (List.rev newps);
+ let ce t = Pretyping.check_evars env0 (Evd.from_env env0) sigma (EConstr.of_constr t) in
+ let () = List.iter (iter_constr ce) (List.rev newps) in
+ let map (impls, newfs) typ =
+ let newfs = List.map (EConstr.to_rel_decl sigma) newfs in
+ let typ = EConstr.to_constr sigma typ in
List.iter (iter_constr ce) (List.rev newfs);
- ubinders, univs, typ, template, imps, newps, impls, newfs
+ (typ, impls, newfs)
+ in
+ let ans = List.map2 map data typs in
+ ubinders, univs, template, newps, imps, ans
let degenerate_decl decl =
let id = match RelDecl.get_name decl with
@@ -261,9 +273,10 @@ let subst_projection fid l c =
raise (NotDefinable (MissingProj (fid,List.rev !bad_projs)));
c''
-let instantiate_possibly_recursive_type indu paramdecls fields =
+let instantiate_possibly_recursive_type ind u ntypes paramdecls fields =
let subst = List.map_i (fun i _ -> mkRel i) 1 paramdecls in
- Termops.substl_rel_context (subst@[mkIndU indu]) fields
+ let subst' = List.init ntypes (fun i -> mkIndU ((ind, ntypes - i - 1), u)) in
+ Termops.substl_rel_context (subst @ subst') fields
let warn_non_primitive_record =
CWarnings.create ~name:"non-primitive-record" ~category:"record"
@@ -281,12 +294,11 @@ let declare_projections indsp ctx ?(kind=StructureComponent) binder_name coers u
| Monomorphic_const_entry ctx -> Univ.Instance.empty
in
let paramdecls = Inductive.inductive_paramdecls (mib, u) in
- let indu = indsp, u in
let r = mkIndU (indsp,u) in
let rp = applist (r, Context.Rel.to_extended_list mkRel 0 paramdecls) in
let paramargs = Context.Rel.to_extended_list mkRel 1 paramdecls in (*def in [[params;x:rp]]*)
let x = Name binder_name in
- let fields = instantiate_possibly_recursive_type indu paramdecls fields in
+ let fields = instantiate_possibly_recursive_type (fst indsp) u mib.mind_ntypes paramdecls fields in
let lifted_fields = Termops.lift_rel_context 1 fields in
let primitive =
if !primitive_flag then
@@ -374,12 +386,9 @@ let declare_projections indsp ctx ?(kind=StructureComponent) binder_name coers u
open Typeclasses
-let declare_structure finite ubinders univs id idbuild paramimpls params arity template
- fieldimpls fields ?(kind=StructureComponent) ?name is_coe coers =
- let nparams = List.length params and nfields = List.length fields in
- let args = Context.Rel.to_extended_list mkRel nfields params in
- let ind = applist (mkRel (1+nparams+nfields), args) in
- let type_constructor = it_mkProd_or_LetIn ind fields in
+
+let declare_structure finite ubinders univs paramimpls params template ?(kind=StructureComponent) ?name record_data =
+ let nparams = List.length params in
let template, ctx =
match univs with
| Monomorphic_ind_entry ctx ->
@@ -389,37 +398,51 @@ let declare_structure finite ubinders univs id idbuild paramimpls params arity t
| Cumulative_ind_entry cumi ->
false, Polymorphic_const_entry (Univ.CumulativityInfo.univ_context cumi)
in
- let binder_name =
+ let binder_name =
match name with
- | None -> Id.of_string (Unicode.lowercase_first_char (Id.to_string id))
+ | None ->
+ let map (id, _, _, _, _, _, _) =
+ Id.of_string (Unicode.lowercase_first_char (Id.to_string id))
+ in
+ Array.map_of_list map record_data
| Some n -> n
in
- let mie_ind =
+ let ntypes = List.length record_data in
+ let mk_block i (id, idbuild, arity, _, fields, _, _) =
+ let nfields = List.length fields in
+ let args = Context.Rel.to_extended_list mkRel nfields params in
+ let ind = applist (mkRel (ntypes - i + nparams + nfields), args) in
+ let type_constructor = it_mkProd_or_LetIn ind fields in
{ mind_entry_typename = id;
mind_entry_arity = arity;
mind_entry_template = template;
mind_entry_consnames = [idbuild];
mind_entry_lc = [type_constructor] }
in
+ let blocks = List.mapi mk_block record_data in
let mie =
{ mind_entry_params = List.map degenerate_decl params;
- mind_entry_record = Some (if !primitive_flag then Some [|binder_name|] else None);
+ mind_entry_record = Some (if !primitive_flag then Some binder_name else None);
mind_entry_finite = finite;
- mind_entry_inds = [mie_ind];
+ mind_entry_inds = blocks;
mind_entry_private = None;
mind_entry_universes = univs;
}
in
let mie = InferCumulativity.infer_inductive (Global.env ()) mie in
- let kn = ComInductive.declare_mutual_inductive_with_eliminations mie ubinders [(paramimpls,[])] in
- let rsp = (kn,0) in (* This is ind path of idstruc *)
- let cstr = (rsp,1) in
- let kinds,sp_projs = declare_projections rsp ctx ~kind binder_name coers ubinders fieldimpls fields in
- let build = ConstructRef cstr in
- let poly = match ctx with | Polymorphic_const_entry _ -> true | Monomorphic_const_entry _ -> false in
- let () = if is_coe then Class.try_add_new_coercion build ~local:false poly in
- Recordops.declare_structure(rsp,cstr,List.rev kinds,List.rev sp_projs);
- rsp
+ let impls = List.map (fun _ -> paramimpls, []) record_data in
+ let kn = ComInductive.declare_mutual_inductive_with_eliminations mie ubinders impls in
+ let map i (_, _, _, fieldimpls, fields, is_coe, coers) =
+ let rsp = (kn, i) in (* This is ind path of idstruc *)
+ let cstr = (rsp, 1) in
+ let kinds,sp_projs = declare_projections rsp ctx ~kind binder_name.(i) coers ubinders fieldimpls fields in
+ let build = ConstructRef cstr in
+ let poly = match ctx with | Polymorphic_const_entry _ -> true | Monomorphic_const_entry _ -> false in
+ let () = if is_coe then Class.try_add_new_coercion build ~local:false poly in
+ let () = Recordops.declare_structure(rsp,cstr,List.rev kinds,List.rev sp_projs) in
+ rsp
+ in
+ List.mapi map record_data
let implicits_of_context ctx =
List.map_i (fun i name ->
@@ -431,22 +454,22 @@ let implicits_of_context ctx =
1 (List.rev (Anonymous :: (List.map RelDecl.get_name ctx)))
let declare_class finite def cum ubinders univs id idbuild paramimpls params arity
- template fieldimpls fields ?(kind=StructureComponent) is_coe coers priorities =
+ template fieldimpls fields ?(kind=StructureComponent) coers priorities =
let fieldimpls =
(* Make the class implicit in the projections, and the params if applicable. *)
let len = List.length params in
let impls = implicits_of_context params in
List.map (fun x -> impls @ Impargs.lift_implicits (succ len) x) fieldimpls
in
- let binder_name = Namegen.next_ident_away (snd id) (Termops.vars_of_env (Global.env())) in
- let impl, projs =
+ let binder_name = Namegen.next_ident_away id (Termops.vars_of_env (Global.env())) in
+ let data =
match fields with
| [LocalAssum (Name proj_name, field) | LocalDef (Name proj_name, _, field)] when def ->
let class_body = it_mkLambda_or_LetIn field params in
let class_type = it_mkProd_or_LetIn arity params in
let class_entry =
Declare.definition_entry ~types:class_type ~univs class_body in
- let cst = Declare.declare_constant (snd id)
+ let cst = Declare.declare_constant id
(DefinitionEntry class_entry, IsDefinition Definition)
in
let cstu = (cst, match univs with
@@ -473,7 +496,7 @@ let declare_class finite def cum ubinders univs id idbuild paramimpls params ari
| Some b -> Some ((if b then Backward else Forward), List.hd priorities)
| None -> None
in
- cref, [Name proj_name, sub, Some proj_cst]
+ [cref, [Name proj_name, sub, Some proj_cst]]
| _ ->
let univs =
match univs with
@@ -485,18 +508,21 @@ let declare_class finite def cum ubinders univs id idbuild paramimpls params ari
| Monomorphic_const_entry univs ->
Monomorphic_ind_entry univs
in
- let ind = declare_structure Declarations.BiFinite ubinders univs (snd id) idbuild paramimpls
- params arity template fieldimpls fields
- ~kind:Method ~name:binder_name false (List.map (fun _ -> false) fields)
- in
+ let record_data = [id, idbuild, arity, fieldimpls, fields, false, List.map (fun _ -> false) fields] in
+ let inds = declare_structure Declarations.BiFinite ubinders univs paramimpls
+ params template ~kind:Method ~name:[|binder_name|] record_data
+ in
let coers = List.map2 (fun coe pri ->
Option.map (fun b ->
if b then Backward, pri else Forward, pri) coe)
coers priorities
in
- let l = List.map3 (fun decl b y -> RelDecl.get_name decl, b, y)
- (List.rev fields) coers (Recordops.lookup_projections ind)
- in IndRef ind, l
+ let map ind =
+ let l = List.map3 (fun decl b y -> RelDecl.get_name decl, b, y)
+ (List.rev fields) coers (Recordops.lookup_projections ind)
+ in IndRef ind, l
+ in
+ List.map map inds
in
let ctx_context =
List.map (fun decl ->
@@ -517,16 +543,19 @@ let declare_class finite def cum ubinders univs id idbuild paramimpls params ari
| Monomorphic_const_entry _ ->
Univ.AUContext.empty, ctx_context, fields
in
- let k =
- { cl_univs = univs;
- cl_impl = impl;
- cl_strict = !typeclasses_strict;
- cl_unique = !typeclasses_unique;
- cl_context = ctx_context;
- cl_props = fields;
- cl_projs = projs }
- in
+ let map (impl, projs) =
+ let k =
+ { cl_univs = univs;
+ cl_impl = impl;
+ cl_strict = !typeclasses_strict;
+ cl_unique = !typeclasses_unique;
+ cl_context = ctx_context;
+ cl_props = fields;
+ cl_projs = projs }
+ in
add_class k; impl
+ in
+ List.map map data
let add_constant_class cst =
@@ -577,39 +606,72 @@ let declare_existing_class g =
open Vernacexpr
-(* [fs] corresponds to fields and [ps] to parameters; [coers] is a
- list telling if the corresponding fields must me declared as coercions
- or subinstances. *)
-let definition_structure (kind,cum,poly,finite,(is_coe,({CAst.loc;v=idstruc},pl)),ps,cfs,idbuild,s) =
- let cfs,notations = List.split cfs in
- let cfs,priorities = List.split cfs in
- let coers,fs = List.split cfs in
- let extract_name acc = function
+let check_unique_names records =
+ let extract_name acc (((_, bnd), _), _) = match bnd with
Vernacexpr.AssumExpr({CAst.v=Name id},_) -> id::acc
| Vernacexpr.DefExpr ({CAst.v=Name id},_,_) -> id::acc
| _ -> acc in
- let allnames = idstruc::(List.fold_left extract_name [] fs) in
- let () = match List.duplicates Id.equal allnames with
+ let allnames =
+ List.fold_left (fun acc (_, id, _, _, cfs, _, _) ->
+ id.CAst.v :: (List.fold_left extract_name acc cfs)) [] records
+ in
+ match List.duplicates Id.equal allnames with
| [] -> ()
| id :: _ -> user_err (str "Two objects have the same name" ++ spc () ++ quote (Id.print id))
- in
+
+let check_priorities kind records =
let isnot_class = match kind with Class false -> false | _ -> true in
- if isnot_class && List.exists (fun opt -> not (Option.is_empty opt)) priorities then
- user_err Pp.(str "Priorities only allowed for type class substructures");
- (* Now, younger decl in params and fields is on top *)
- let pl, univs, arity, template, implpars, params, implfs, fields =
+ let has_priority (_, _, _, _, cfs, _, _) =
+ List.exists (fun ((_, pri), _) -> not (Option.is_empty pri)) cfs
+ in
+ if isnot_class && List.exists has_priority records then
+ user_err Pp.(str "Priorities only allowed for type class substructures")
+
+let extract_record_data records =
+ let map (is_coe, id, _, _, cfs, idbuild, s) =
+ let fs = List.map (fun (((_, f), _), _) -> f) cfs in
+ id.CAst.v, s, List.map snd cfs, fs
+ in
+ let data = List.map map records in
+ let pss = List.map (fun (_, _, _, ps, _, _, _) -> ps) records in
+ let ps = match pss with
+ | [] -> CErrors.anomaly (str "Empty record block")
+ | ps :: rem ->
+ let eq_local_binders bl1 bl2 = List.equal local_binder_eq bl1 bl2 in
+ let () =
+ if not (List.for_all (eq_local_binders ps) rem) then
+ user_err (str "Parameters should be syntactically the \
+ same for each inductive type.")
+ in
+ ps
+ in
+ (** FIXME: Same issue as #7754 *)
+ let _, _, pl, _, _, _, _ = List.hd records in
+ pl, ps, data
+
+(* [fs] corresponds to fields and [ps] to parameters; [coers] is a
+ list telling if the corresponding fields must me declared as coercions
+ or subinstances. *)
+let definition_structure kind cum poly finite records =
+ let () = check_unique_names records in
+ let () = check_priorities kind records in
+ let pl, ps, data = extract_record_data records in
+ let pl, univs, template, params, implpars, data =
States.with_state_protection (fun () ->
- typecheck_params_and_fields finite (kind = Class true) idstruc poly pl s ps notations fs) () in
+ typecheck_params_and_fields finite (kind = Class true) poly pl ps data) () in
match kind with
| Class def ->
- let priorities = List.map (fun id -> {hint_priority = id; hint_pattern = None}) priorities in
- declare_class finite def cum pl univs (loc,idstruc) idbuild
- implpars params arity template implfs fields is_coe coers priorities
+ let (_, id, _, _, cfs, idbuild, _), (arity, implfs, fields) = match records, data with
+ | [r], [d] -> r, d
+ | _, _ -> CErrors.user_err (str "Mutual definitional classes are not handled")
+ in
+ let priorities = List.map (fun ((_, id), _) -> {hint_priority = id; hint_pattern = None}) cfs in
+ let coers = List.map (fun (((coe, _), _), _) -> coe) cfs in
+ declare_class finite def cum pl univs id.CAst.v idbuild
+ implpars params arity template implfs fields coers priorities
| _ ->
- let implfs = List.map
- (fun impls -> implpars @ Impargs.lift_implicits
- (succ (List.length params)) impls) implfs
- in
+ let map impls = implpars @ Impargs.lift_implicits (succ (List.length params)) impls in
+ let data = List.map (fun (arity, implfs, fields) -> (arity, List.map map implfs, fields)) data in
let univs =
match univs with
| Polymorphic_const_entry univs ->
@@ -620,7 +682,11 @@ let definition_structure (kind,cum,poly,finite,(is_coe,({CAst.loc;v=idstruc},pl)
| Monomorphic_const_entry univs ->
Monomorphic_ind_entry univs
in
- let ind = declare_structure finite pl univs idstruc
- idbuild implpars params arity template implfs
- fields is_coe (List.map (fun coe -> not (Option.is_empty coe)) coers) in
- IndRef ind
+ let map (arity, implfs, fields) (is_coe, id, _, _, cfs, idbuild, _) =
+ let coers = List.map (fun (((coe, _), _), _) -> coe) cfs in
+ let coe = List.map (fun coe -> not (Option.is_empty coe)) coers in
+ id.CAst.v, idbuild, arity, implfs, fields, is_coe, coe
+ in
+ let data = List.map2 map data records in
+ let inds = declare_structure finite pl univs implpars params template data in
+ List.map (fun ind -> IndRef ind) inds
diff --git a/vernac/record.mli b/vernac/record.mli
index b2c039f0b..7cf7da1e2 100644
--- a/vernac/record.mli
+++ b/vernac/record.mli
@@ -26,9 +26,14 @@ val declare_projections :
(Name.t * bool) list * Constant.t option list
val definition_structure :
- inductive_kind * Decl_kinds.cumulative_inductive_flag * Decl_kinds.polymorphic *
- Declarations.recursivity_kind * ident_decl with_coercion * local_binder_expr list *
+ inductive_kind -> Decl_kinds.cumulative_inductive_flag -> Decl_kinds.polymorphic ->
+ Declarations.recursivity_kind ->
+ (coercion_flag *
+ Names.lident *
+ universe_decl_expr option *
+ local_binder_expr list *
(local_decl_expr with_instance with_priority with_notation) list *
- Id.t * constr_expr option -> GlobRef.t
+ Id.t * constr_expr option) list ->
+ GlobRef.t list
val declare_existing_class : GlobRef.t -> unit
diff --git a/vernac/vernacentries.ml b/vernac/vernacentries.ml
index 479482095..43c974846 100644
--- a/vernac/vernacentries.ml
+++ b/vernac/vernacentries.ml
@@ -539,25 +539,36 @@ let should_treat_as_cumulative cum poly =
else user_err Pp.(str "The NonCumulative prefix can only be used in a polymorphic context.")
| None -> poly && Flags.is_polymorphic_inductive_cumulativity ()
-let vernac_record cum k poly finite struc binders sort nameopt cfs =
+let vernac_record cum k poly finite records =
let is_cumulative = should_treat_as_cumulative cum poly in
- let const = match nameopt with
- | None -> add_prefix "Build_" (fst (snd struc)).v
- | Some ({v=id} as lid) ->
- Dumpglob.dump_definition lid false "constr"; id in
- if Dumpglob.dump () then (
- Dumpglob.dump_definition (fst (snd struc)) false "rec";
- List.iter (fun (((_, x), _), _) ->
- match x with
- | Vernacexpr.AssumExpr ({loc;v=Name id}, _) -> Dumpglob.dump_definition (make ?loc id) false "proj"
- | _ -> ()) cfs);
- ignore(Record.definition_structure (k,is_cumulative,poly,finite,struc,binders,cfs,const,sort))
+ let map ((coe, (id, pl)), binders, sort, nameopt, cfs) =
+ let const = match nameopt with
+ | None -> add_prefix "Build_" id.v
+ | Some lid ->
+ let () = Dumpglob.dump_definition lid false "constr" in
+ lid.v
+ in
+ let () =
+ if Dumpglob.dump () then
+ let () = Dumpglob.dump_definition id false "rec" in
+ let iter (((_, x), _), _) = match x with
+ | Vernacexpr.AssumExpr ({loc;v=Name id}, _) ->
+ Dumpglob.dump_definition (make ?loc id) false "proj"
+ | _ -> ()
+ in
+ List.iter iter cfs
+ in
+ coe, id, pl, binders, cfs, const, sort
+ in
+ let records = List.map map records in
+ ignore(Record.definition_structure k is_cumulative poly finite records)
(** When [poly] is true the type is declared polymorphic. When [lo] is true,
then the type is declared private (as per the [Private] keyword). [finite]
indicates whether the type is inductive, co-inductive or
neither. *)
let vernac_inductive ~atts cum lo finite indl =
+ let open Pp in
if Dumpglob.dump () then
List.iter (fun (((coe,(lid,_)), _, _, _, cstrs), _) ->
match cstrs with
@@ -567,35 +578,85 @@ let vernac_inductive ~atts cum lo finite indl =
Dumpglob.dump_definition lid false "constr") cstrs
| _ -> () (* dumping is done by vernac_record (called below) *) )
indl;
+ let is_record = function
+ | ((_ , _ , _ , _, RecordDecl _), _) -> true
+ | _ -> false
+ in
+ let is_constructor = function
+ | ((_ , _ , _ , _, Constructors _), _) -> true
+ | _ -> false
+ in
+ let is_defclass = match indl with
+ | [ ( id , bl , c , Class _, Constructors [l]), [] ] -> Some (id, bl, c, l)
+ | _ -> None
+ in
+ if Option.has_some is_defclass then
+ (** Definitional class case *)
+ let (id, bl, c, l) = Option.get is_defclass in
+ let (coe, (lid, ce)) = l in
+ let coe' = if coe then Some true else None in
+ let f = (((coe', AssumExpr ((make ?loc:lid.loc @@ Name lid.v), ce)), None), []) in
+ vernac_record cum (Class true) atts.polymorphic finite [id, bl, c, None, [f]]
+ else if List.for_all is_record indl then
+ (** Mutual record case *)
+ let check_kind ((_, _, _, kind, _), _) = match kind with
+ | Variant ->
+ user_err (str "The Variant keyword does not support syntax { ... }.")
+ | Record | Structure | Class _ | Inductive_kw | CoInductive -> ()
+ in
+ let () = List.iter check_kind indl in
+ let check_where ((_, _, _, _, _), wh) = match wh with
+ | [] -> ()
+ | _ :: _ ->
+ user_err (str "where clause not supported for records")
+ in
+ let () = List.iter check_where indl in
+ let unpack ((id, bl, c, _, decl), _) = match decl with
+ | RecordDecl (oc, fs) ->
+ (id, bl, c, oc, fs)
+ | Constructors _ -> assert false (** ruled out above *)
+ in
+ let ((_, _, _, kind, _), _) = List.hd indl in
+ let kind = match kind with Class _ -> Class false | _ -> kind in
+ let recordl = List.map unpack indl in
+ vernac_record cum kind atts.polymorphic finite recordl
+ else if List.for_all is_constructor indl then
+ (** Mutual inductive case *)
+ let check_kind ((_, _, _, kind, _), _) = match kind with
+ | (Record | Structure) ->
+ user_err (str "The Record keyword is for types defined using the syntax { ... }.")
+ | Class _ ->
+ user_err (str "Inductive classes not supported")
+ | Variant | Inductive_kw | CoInductive -> ()
+ in
+ let () = List.iter check_kind indl in
+ let check_name ((na, _, _, _, _), _) = match na with
+ | (true, _) ->
+ user_err (str "Variant types do not handle the \"> Name\" \
+ syntax, which is reserved for records. Use the \":>\" \
+ syntax on constructors instead.")
+ | _ -> ()
+ in
+ let () = List.iter check_name indl in
+ let unpack (((_, id) , bl, c, _, decl), ntn) = match decl with
+ | Constructors l -> (id, bl, c, l), ntn
+ | RecordDecl _ -> assert false (* ruled out above *)
+ in
+ let indl = List.map unpack indl in
+ let is_cumulative = should_treat_as_cumulative cum atts.polymorphic in
+ ComInductive.do_mutual_inductive indl is_cumulative atts.polymorphic lo finite
+ else
+ user_err (str "Mixed record-inductive definitions are not allowed")
+(*
+
match indl with
- | [ ( _ , _ , _ ,(Record|Structure), Constructors _ ),_ ] ->
- user_err Pp.(str "The Record keyword is for types defined using the syntax { ... }.")
- | [ (_ , _ , _ ,Variant, RecordDecl _),_ ] ->
- user_err Pp.(str "The Variant keyword does not support syntax { ... }.")
- | [ ( id , bl , c , b, RecordDecl (oc,fs) ), [] ] ->
- vernac_record cum (match b with Class _ -> Class false | _ -> b)
- atts.polymorphic finite id bl c oc fs
| [ ( id , bl , c , Class _, Constructors [l]), [] ] ->
let f =
let (coe, ({loc;v=id}, ce)) = l in
let coe' = if coe then Some true else None in
(((coe', AssumExpr ((make ?loc @@ Name id), ce)), None), [])
- in vernac_record cum (Class true) atts.polymorphic finite id bl c None [f]
- | [ ( _ , _, _, Class _, Constructors _), [] ] ->
- user_err Pp.(str "Inductive classes not supported")
- | [ ( id , bl , c , Class _, _), _ :: _ ] ->
- user_err Pp.(str "where clause not supported for classes")
- | [ ( _ , _ , _ , _, RecordDecl _ ) , _ ] ->
- user_err Pp.(str "where clause not supported for (co)inductive records")
- | _ -> let unpack = function
- | ( (false, id) , bl , c , _ , Constructors l ) , ntn -> ( id , bl , c , l ) , ntn
- | ( (true,_),_,_,_,Constructors _),_ ->
- user_err Pp.(str "Variant types do not handle the \"> Name\" syntax, which is reserved for records. Use the \":>\" syntax on constructors instead.")
- | _ -> user_err Pp.(str "Cannot handle mutually (co)inductive records.")
- in
- let indl = List.map unpack indl in
- let is_cumulative = should_treat_as_cumulative cum atts.polymorphic in
- ComInductive.do_mutual_inductive indl is_cumulative atts.polymorphic lo finite
+ in vernac_record cum (Class true) atts.polymorphic finite [id, bl, c, None, [f]]
+ *)
let vernac_fixpoint ~atts discharge l =
let local = enforce_locality_exp atts.locality discharge in