aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
-rw-r--r--API/API.mli8
-rwxr-xr-xdev/ci/ci-vst.sh4
-rw-r--r--dev/doc/changes.md4
-rw-r--r--doc/refman/Classes.tex9
-rw-r--r--library/declaremods.ml34
-rw-r--r--library/lib.ml58
-rw-r--r--library/libnames.ml19
-rw-r--r--library/libnames.mli20
-rw-r--r--library/nametab.ml10
-rw-r--r--pretyping/retyping.ml50
-rw-r--r--pretyping/retyping.mli5
-rw-r--r--printing/prettyp.ml20
-rw-r--r--printing/printmod.ml6
-rw-r--r--tactics/equality.ml2
-rw-r--r--test-suite/bugs/closed/3125.v27
-rw-r--r--test-suite/bugs/closed/3559.v1
-rw-r--r--test-suite/bugs/closed/HoTT_coq_064.v1
-rw-r--r--test-suite/success/Typeclasses.v17
-rw-r--r--test-suite/success/bteauto.v1
-rw-r--r--theories/Compat/Coq87.v3
-rw-r--r--vernac/command.ml21
-rw-r--r--vernac/vernacentries.ml4
22 files changed, 219 insertions, 105 deletions
diff --git a/API/API.mli b/API/API.mli
index 321a0b808..df5a9b131 100644
--- a/API/API.mli
+++ b/API/API.mli
@@ -2082,7 +2082,11 @@ sig
val basename : full_path -> Names.Id.t
type object_name = full_path * Names.KerName.t
- type object_prefix = Names.DirPath.t * (Names.ModPath.t * Names.DirPath.t)
+ type object_prefix = {
+ obj_dir : DirPath.t;
+ obj_mp : ModPath.t;
+ obj_sec : DirPath.t;
+ }
module Dirset : Set.S with type elt = DirPath.t
module Dirmap : Map.ExtS with type key = DirPath.t and module Set := Dirset
@@ -3717,7 +3721,7 @@ end
module Retyping : (* reconstruct the type of a term knowing that it was already typechecked *)
sig
val get_type_of : ?polyprop:bool -> ?lax:bool -> Environ.env -> Evd.evar_map -> EConstr.constr -> EConstr.types
- val get_sort_family_of : ?polyprop:bool -> Environ.env -> Evd.evar_map -> EConstr.types -> Sorts.family
+ val get_sort_family_of : ?truncation_style:bool -> ?polyprop:bool -> Environ.env -> Evd.evar_map -> EConstr.types -> Sorts.family
val expand_projection : Environ.env -> Evd.evar_map -> Names.Projection.t -> EConstr.constr -> EConstr.constr list -> EConstr.constr
val get_sort_of :
?polyprop:bool -> Environ.env -> Evd.evar_map -> EConstr.types -> Sorts.t
diff --git a/dev/ci/ci-vst.sh b/dev/ci/ci-vst.sh
index 5bfc408e9..5760fbafb 100755
--- a/dev/ci/ci-vst.sh
+++ b/dev/ci/ci-vst.sh
@@ -8,6 +8,6 @@ VST_CI_DIR=${CI_BUILD_DIR}/VST
# opam install -j ${NJOBS} -y menhir
git_checkout ${VST_CI_BRANCH} ${VST_CI_GITURL} ${VST_CI_DIR}
-# Targets are: msl veric floyd
+# Targets are: msl veric floyd progs , we remove progs to save time
# Patch to avoid the upper version limit
-( cd ${VST_CI_DIR} && make IGNORECOQVERSION=true )
+( cd ${VST_CI_DIR} && make IGNORECOQVERSION=true .loadpath version.vo msl veric floyd )
diff --git a/dev/doc/changes.md b/dev/doc/changes.md
index 707adce30..c69be4f4d 100644
--- a/dev/doc/changes.md
+++ b/dev/doc/changes.md
@@ -46,9 +46,9 @@ We changed the type of the following functions:
- `Global.body_of_constant`: same as above.
-We renamed the following datatypes:
+We have changed the representation of the following types:
-- `Pp.std_ppcmds` -> `Pp.t`
+- `Lib.object_prefix` is now a record instead of a nested tuple.
Some tactics and related functions now support static configurability, e.g.:
diff --git a/doc/refman/Classes.tex b/doc/refman/Classes.tex
index 22c75b4fc..cab673999 100644
--- a/doc/refman/Classes.tex
+++ b/doc/refman/Classes.tex
@@ -462,11 +462,18 @@ abbreviate a type, like {\tt relation A := A -> A -> Prop}.
This is equivalent to {\tt Hint Transparent,Opaque} {\ident} {\tt: typeclass\_instances}.
+\subsection{\tt Set Typeclasses Axioms Are Instances}
+\optindex{Typeclasses Axioms Are Instances}
+
+This option (off by default since 8.8) automatically declares axioms
+whose type is a typeclass at declaration time as instances of that
+class.
+
\subsection{\tt Set Typeclasses Dependency Order}
\optindex{Typeclasses Dependency Order}
This option (on by default since 8.6) respects the dependency order between
-subgoals, meaning that subgoals which are depended on by other subgoals
+subgoals, meaning that subgoals which are depended on by other subgoals
come first, while the non-dependent subgoals were put before the
dependent ones previously (Coq v8.5 and below). This can result in quite
different performance behaviors of proof search.
diff --git a/library/declaremods.ml b/library/declaremods.ml
index db83dafef..41e00a41c 100644
--- a/library/declaremods.ml
+++ b/library/declaremods.ml
@@ -180,16 +180,16 @@ let compute_visibility exists i =
(** Iterate some function [iter_objects] on all components of a module *)
-let do_module exists iter_objects i dir mp sobjs kobjs =
- let prefix = (dir,(mp,DirPath.empty)) in
+let do_module exists iter_objects i obj_dir obj_mp sobjs kobjs =
+ let prefix = { obj_dir ; obj_mp; obj_sec = DirPath.empty } in
let dirinfo = DirModule prefix in
- consistency_checks exists dir dirinfo;
- Nametab.push_dir (compute_visibility exists i) dir dirinfo;
- ModSubstObjs.set mp sobjs;
+ consistency_checks exists obj_dir dirinfo;
+ Nametab.push_dir (compute_visibility exists i) obj_dir dirinfo;
+ ModSubstObjs.set obj_mp sobjs;
(* If we're not a functor, let's iter on the internal components *)
if sobjs_no_functor sobjs then begin
let objs = expand_sobjs sobjs in
- ModObjs.set mp (prefix,objs,kobjs);
+ ModObjs.set obj_mp (prefix,objs,kobjs);
iter_objects (i+1) prefix objs;
iter_objects (i+1) prefix kobjs
end
@@ -222,20 +222,20 @@ let cache_keep _ = anomaly (Pp.str "This module should not be cached!")
let load_keep i ((sp,kn),kobjs) =
(* Invariant : seg isn't empty *)
- let dir = dir_of_sp sp and mp = mp_of_kn kn in
- let prefix = (dir,(mp,DirPath.empty)) in
+ let obj_dir = dir_of_sp sp and obj_mp = mp_of_kn kn in
+ let prefix = { obj_dir ; obj_mp; obj_sec = DirPath.empty } in
let prefix',sobjs,kobjs0 =
- try ModObjs.get mp
+ try ModObjs.get obj_mp
with Not_found -> assert false (* a substobjs should already be loaded *)
in
assert (eq_op prefix' prefix);
assert (List.is_empty kobjs0);
- ModObjs.set mp (prefix,sobjs,kobjs);
+ ModObjs.set obj_mp (prefix,sobjs,kobjs);
Lib.load_objects i prefix kobjs
let open_keep i ((sp,kn),kobjs) =
- let dir = dir_of_sp sp and mp = mp_of_kn kn in
- let prefix = (dir,(mp,DirPath.empty)) in
+ let obj_dir = dir_of_sp sp and obj_mp = mp_of_kn kn in
+ let prefix = { obj_dir; obj_mp; obj_sec = DirPath.empty } in
Lib.open_objects i prefix kobjs
let in_modkeep : Lib.lib_objects -> obj =
@@ -284,9 +284,9 @@ let (in_modtype : substitutive_objects -> obj),
(** {6 Declaration of substitutive objects for Include} *)
let do_include do_load do_open i ((sp,kn),aobjs) =
- let dir = Libnames.dirpath sp in
- let mp = KerName.modpath kn in
- let prefix = (dir,(mp,DirPath.empty)) in
+ let obj_dir = Libnames.dirpath sp in
+ let obj_mp = KerName.modpath kn in
+ let prefix = { obj_dir; obj_mp; obj_sec = DirPath.empty } in
let o = expand_aobjs aobjs in
if do_load then Lib.load_objects i prefix o;
if do_open then Lib.open_objects i prefix o
@@ -577,7 +577,7 @@ let start_module interp_modast export id args res fs =
in
openmod_info := { cur_typ = res_entry_o; cur_typs = subtyps };
let prefix = Lib.start_module export id mp fs in
- Nametab.push_dir (Nametab.Until 1) (fst prefix) (DirOpenModule prefix);
+ Nametab.push_dir (Nametab.Until 1) (prefix.obj_dir) (DirOpenModule prefix);
mp
let end_module () =
@@ -684,7 +684,7 @@ let start_modtype interp_modast id args mtys fs =
let sub_mty_l = build_subtypes interp_modast env mp arg_entries_r mtys in
openmodtype_info := sub_mty_l;
let prefix = Lib.start_modtype id mp fs in
- Nametab.push_dir (Nametab.Until 1) (fst prefix) (DirOpenModtype prefix);
+ Nametab.push_dir (Nametab.Until 1) (prefix.obj_dir) (DirOpenModtype prefix);
mp
let end_modtype () =
diff --git a/library/lib.ml b/library/lib.ml
index 3dbb16c7b..499e2ae21 100644
--- a/library/lib.ml
+++ b/library/lib.ml
@@ -93,12 +93,16 @@ let segment_of_objects prefix =
sections, but on the contrary there are many constructions of section
paths based on the library path. *)
-let initial_prefix = default_library,(Names.ModPath.initial,Names.DirPath.empty)
+let initial_prefix = {
+ obj_dir = default_library;
+ obj_mp = ModPath.initial;
+ obj_sec = DirPath.empty;
+}
type lib_state = {
- comp_name : Names.DirPath.t option;
+ comp_name : DirPath.t option;
lib_stk : library_segment;
- path_prefix : Names.DirPath.t * (Names.ModPath.t * Names.DirPath.t);
+ path_prefix : object_prefix;
}
let initial_lib_state = {
@@ -115,10 +119,9 @@ let library_dp () =
(* [path_prefix] is a pair of absolute dirpath and a pair of current
module path and relative section path *)
-let cwd () = fst !lib_state.path_prefix
-let current_prefix () = snd !lib_state.path_prefix
-let current_mp () = fst (snd !lib_state.path_prefix)
-let current_sections () = snd (snd !lib_state.path_prefix)
+let cwd () = !lib_state.path_prefix.obj_dir
+let current_mp () = !lib_state.path_prefix.obj_mp
+let current_sections () = !lib_state.path_prefix.obj_sec
let sections_depth () = List.length (Names.DirPath.repr (current_sections ()))
let sections_are_opened () = not (Names.DirPath.is_empty (current_sections ()))
@@ -136,7 +139,7 @@ let make_path_except_section id =
Libnames.make_path (cwd_except_section ()) id
let make_kn id =
- let mp,dir = current_prefix () in
+ let mp, dir = current_mp (), current_sections () in
Names.KerName.make mp dir (Names.Label.of_id id)
let make_oname id = Libnames.make_oname !lib_state.path_prefix id
@@ -152,8 +155,11 @@ let recalc_path_prefix () =
lib_state := { !lib_state with path_prefix = recalc !lib_state.lib_stk }
let pop_path_prefix () =
- let dir,(mp,sec) = !lib_state.path_prefix in
- lib_state := { !lib_state with path_prefix = pop_dirpath dir, (mp, pop_dirpath sec)}
+ let op = !lib_state.path_prefix in
+ lib_state := { !lib_state
+ with path_prefix = { op with obj_dir = pop_dirpath op.obj_dir;
+ obj_sec = pop_dirpath op.obj_sec;
+ } }
let find_entry_p p =
let rec find = function
@@ -226,7 +232,7 @@ let add_anonymous_entry node =
add_entry (make_oname (anonymous_id ())) node
let add_leaf id obj =
- if Names.ModPath.equal (current_mp ()) Names.ModPath.initial then
+ if ModPath.equal (current_mp ()) ModPath.initial then
user_err Pp.(str "No session module started (use -top dir)");
let oname = make_oname id in
cache_object (oname,obj);
@@ -278,8 +284,8 @@ let current_mod_id () =
let start_mod is_type export id mp fs =
- let dir = add_dirpath_suffix (cwd ()) id in
- let prefix = dir,(mp,Names.DirPath.empty) in
+ let dir = add_dirpath_suffix (!lib_state.path_prefix.obj_dir) id in
+ let prefix = { obj_dir = dir; obj_mp = mp; obj_sec = Names.DirPath.empty } in
let exists =
if is_type then Nametab.exists_cci (make_path id)
else Nametab.exists_module dir
@@ -328,10 +334,10 @@ let contents_after sp = let (after,_,_) = split_lib sp in after
let start_compilation s mp =
if !lib_state.comp_name != None then
user_err Pp.(str "compilation unit is already started");
- if not (Names.DirPath.is_empty (current_sections ())) then
+ if not (Names.DirPath.is_empty (!lib_state.path_prefix.obj_sec)) then
user_err Pp.(str "some sections are already opened");
- let prefix = s, (mp, Names.DirPath.empty) in
- let () = add_anonymous_entry (CompilingLibrary prefix) in
+ let prefix = Libnames.{ obj_dir = s; obj_mp = mp; obj_sec = DirPath.empty } in
+ add_anonymous_entry (CompilingLibrary prefix);
lib_state := { !lib_state with comp_name = Some s;
path_prefix = prefix }
@@ -522,15 +528,15 @@ let is_in_section ref =
(*************)
(* Sections. *)
let open_section id =
- let olddir,(mp,oldsec) = !lib_state.path_prefix in
- let dir = add_dirpath_suffix olddir id in
- let prefix = dir, (mp, add_dirpath_suffix oldsec id) in
- if Nametab.exists_section dir then
+ let opp = !lib_state.path_prefix in
+ let obj_dir = add_dirpath_suffix opp.obj_dir id in
+ let prefix = { obj_dir; obj_mp = opp.obj_mp; obj_sec = add_dirpath_suffix opp.obj_sec id } in
+ if Nametab.exists_section obj_dir then
user_err ~hdr:"open_section" (Id.print id ++ str " already exists.");
let fs = Summary.freeze_summaries ~marshallable:`No in
add_entry (make_oname id) (OpenedSection (prefix, fs));
(*Pushed for the lifetime of the section: removed by unfrozing the summary*)
- Nametab.push_dir (Nametab.Until 1) dir (DirOpenSection prefix);
+ Nametab.push_dir (Nametab.Until 1) obj_dir (DirOpenSection prefix);
lib_state := { !lib_state with path_prefix = prefix };
add_section ()
@@ -556,7 +562,7 @@ let close_section () =
in
let (secdecls,mark,before) = split_lib_at_opening oname in
lib_state := { !lib_state with lib_stk = before };
- let full_olddir = fst !lib_state.path_prefix in
+ let full_olddir = !lib_state.path_prefix.obj_dir in
pop_path_prefix ();
add_entry oname (ClosedSection (List.rev (mark::secdecls)));
let newdecls = List.map discharge_item secdecls in
@@ -596,10 +602,10 @@ let init () =
(* Misc *)
let mp_of_global = function
- |VarRef id -> current_mp ()
- |ConstRef cst -> Names.Constant.modpath cst
- |IndRef ind -> Names.ind_modpath ind
- |ConstructRef constr -> Names.constr_modpath constr
+ | VarRef id -> !lib_state.path_prefix.obj_mp
+ | ConstRef cst -> Names.Constant.modpath cst
+ | IndRef ind -> Names.ind_modpath ind
+ | ConstructRef constr -> Names.constr_modpath constr
let rec dp_of_mp = function
|Names.MPfile dp -> dp
diff --git a/library/libnames.ml b/library/libnames.ml
index 81878e72f..a471d8396 100644
--- a/library/libnames.ml
+++ b/library/libnames.ml
@@ -156,10 +156,15 @@ let qualid_of_dirpath dir =
type object_name = full_path * KerName.t
-type object_prefix = DirPath.t * (ModPath.t * DirPath.t)
+type object_prefix = {
+ obj_dir : DirPath.t;
+ obj_mp : ModPath.t;
+ obj_sec : DirPath.t;
+}
-let make_oname (dirpath,(mp,dir)) id =
- make_path dirpath id, KerName.make mp dir (Label.of_id id)
+(* let make_oname (dirpath,(mp,dir)) id = *)
+let make_oname { obj_dir; obj_mp; obj_sec } id =
+ make_path obj_dir id, KerName.make obj_mp obj_sec (Label.of_id id)
(* to this type are mapped DirPath.t's in the nametab *)
type global_dir_reference =
@@ -170,10 +175,10 @@ type global_dir_reference =
| DirClosedSection of DirPath.t
(* this won't last long I hope! *)
-let eq_op (d1, (mp1, p1)) (d2, (mp2, p2)) =
- DirPath.equal d1 d2 &&
- DirPath.equal p1 p2 &&
- ModPath.equal mp1 mp2
+let eq_op op1 op2 =
+ DirPath.equal op1.obj_dir op2.obj_dir &&
+ DirPath.equal op1.obj_sec op2.obj_sec &&
+ ModPath.equal op1.obj_mp op2.obj_mp
let eq_global_dir_reference r1 r2 = match r1, r2 with
| DirOpenModule op1, DirOpenModule op2 -> eq_op op1 op2
diff --git a/library/libnames.mli b/library/libnames.mli
index ed01163ee..71f542240 100644
--- a/library/libnames.mli
+++ b/library/libnames.mli
@@ -94,7 +94,25 @@ val qualid_of_ident : Id.t -> qualid
type object_name = full_path * KerName.t
-type object_prefix = DirPath.t * (ModPath.t * DirPath.t)
+(** Object prefix morally contains the "prefix" naming of an object to
+ be stored by [library], where [obj_dir] is the "absolute" path,
+ [obj_mp] is the current "module" prefix and [obj_sec] is the
+ "section" prefix.
+
+ Thus, for an object living inside [Module A. Section B.] the
+ prefix would be:
+
+ [ { obj_dir = "A.B"; obj_mp = "A"; obj_sec = "B" } ]
+
+ Note that both [obj_dir] and [obj_sec] are "paths" that is to say,
+ as opposed to [obj_mp] which is a single module name.
+
+ *)
+type object_prefix = {
+ obj_dir : DirPath.t;
+ obj_mp : ModPath.t;
+ obj_sec : DirPath.t;
+}
val eq_op : object_prefix -> object_prefix -> bool
diff --git a/library/nametab.ml b/library/nametab.ml
index 0ec4a37cd..222c4cedc 100644
--- a/library/nametab.ml
+++ b/library/nametab.ml
@@ -359,8 +359,8 @@ let push_modtype vis sp kn =
let push_dir vis dir dir_ref =
the_dirtab := DirTab.push vis dir dir_ref !the_dirtab;
match dir_ref with
- DirModule (_,(mp,_)) -> the_modrevtab := MPmap.add mp dir !the_modrevtab
- | _ -> ()
+ | DirModule { obj_mp; _ } -> the_modrevtab := MPmap.add obj_mp dir !the_modrevtab
+ | _ -> ()
(* Locate functions *******************************************************)
@@ -386,17 +386,17 @@ let locate_dir qid = DirTab.locate qid !the_dirtab
let locate_module qid =
match locate_dir qid with
- | DirModule (_,(mp,_)) -> mp
+ | DirModule { obj_mp ; _} -> obj_mp
| _ -> raise Not_found
let full_name_module qid =
match locate_dir qid with
- | DirModule (dir,_) -> dir
+ | DirModule { obj_dir ; _} -> obj_dir
| _ -> raise Not_found
let locate_section qid =
match locate_dir qid with
- | DirOpenSection (dir, _)
+ | DirOpenSection { obj_dir; _ } -> obj_dir
| DirClosedSection dir -> dir
| _ -> raise Not_found
diff --git a/pretyping/retyping.ml b/pretyping/retyping.ml
index 5dd6879d3..f8f086fad 100644
--- a/pretyping/retyping.ml
+++ b/pretyping/retyping.ml
@@ -166,23 +166,6 @@ let retype ?(polyprop=true) sigma =
| Lambda _ | Fix _ | Construct _ -> retype_error NotAType
| _ -> decomp_sort env sigma (type_of env t)
- and sort_family_of env t =
- match EConstr.kind sigma t with
- | Cast (c,_, s) when isSort sigma s -> Sorts.family (destSort sigma s)
- | Sort _ -> InType
- | Prod (name,t,c2) ->
- let s2 = sort_family_of (push_rel (LocalAssum (name,t)) env) c2 in
- if not (is_impredicative_set env) &&
- s2 == InSet && sort_family_of env t == InType then InType else s2
- | App(f,args) when is_template_polymorphic env sigma f ->
- let t = type_of_global_reference_knowing_parameters env f args in
- Sorts.family (sort_of_atomic_type env sigma t args)
- | App(f,args) ->
- Sorts.family (sort_of_atomic_type env sigma (type_of env f) args)
- | Lambda _ | Fix _ | Construct _ -> retype_error NotAType
- | _ ->
- Sorts.family (decomp_sort env sigma (type_of env t))
-
and type_of_global_reference_knowing_parameters env c args =
let argtyps =
Array.map (fun c -> lazy (EConstr.to_constr sigma (type_of env c))) args in
@@ -198,15 +181,34 @@ let retype ?(polyprop=true) sigma =
EConstr.of_constr (type_of_constructor env (cstr, u))
| _ -> assert false
- in type_of, sort_of, sort_family_of,
- type_of_global_reference_knowing_parameters
+ in type_of, sort_of, type_of_global_reference_knowing_parameters
+
+let get_sort_family_of ?(truncation_style=false) ?(polyprop=true) env sigma t =
+ let type_of,_,type_of_global_reference_knowing_parameters = retype ~polyprop sigma in
+ let rec sort_family_of env t =
+ match EConstr.kind sigma t with
+ | Cast (c,_, s) when isSort sigma s -> Sorts.family (destSort sigma s)
+ | Sort _ -> InType
+ | Prod (name,t,c2) ->
+ let s2 = sort_family_of (push_rel (LocalAssum (name,t)) env) c2 in
+ if not (is_impredicative_set env) &&
+ s2 == InSet && sort_family_of env t == InType then InType else s2
+ | App(f,args) when is_template_polymorphic env sigma f ->
+ if truncation_style then InType else
+ let t = type_of_global_reference_knowing_parameters env f args in
+ Sorts.family (sort_of_atomic_type env sigma t args)
+ | App(f,args) ->
+ Sorts.family (sort_of_atomic_type env sigma (type_of env f) args)
+ | Lambda _ | Fix _ | Construct _ -> retype_error NotAType
+ | Ind _ when truncation_style && is_template_polymorphic env sigma t -> InType
+ | _ ->
+ Sorts.family (decomp_sort env sigma (type_of env t))
+ in sort_family_of env t
let get_sort_of ?(polyprop=true) env sigma t =
- let _,f,_,_ = retype ~polyprop sigma in anomaly_on_error (f env) t
-let get_sort_family_of ?(polyprop=true) env sigma c =
- let _,_,f,_ = retype ~polyprop sigma in anomaly_on_error (f env) c
+ let _,f,_ = retype ~polyprop sigma in anomaly_on_error (f env) t
let type_of_global_reference_knowing_parameters env sigma c args =
- let _,_,_,f = retype sigma in anomaly_on_error (f env c) args
+ let _,_,f = retype sigma in anomaly_on_error (f env c) args
let type_of_global_reference_knowing_conclusion env sigma c conclty =
match EConstr.kind sigma c with
@@ -232,7 +234,7 @@ let type_of_global_reference_knowing_conclusion env sigma c conclty =
(* get_type_of polyprop lax env sigma c *)
let get_type_of ?(polyprop=true) ?(lax=false) env sigma c =
- let f,_,_,_ = retype ~polyprop sigma in
+ let f,_,_ = retype ~polyprop sigma in
if lax then f env c else anomaly_on_error (f env) c
(* Makes an unsafe judgment from a constr *)
diff --git a/pretyping/retyping.mli b/pretyping/retyping.mli
index af86df499..6fdde9046 100644
--- a/pretyping/retyping.mli
+++ b/pretyping/retyping.mli
@@ -31,8 +31,11 @@ val get_type_of :
val get_sort_of :
?polyprop:bool -> env -> evar_map -> types -> Sorts.t
+(* When [truncation_style] is [true], tells if the type has been explicitly
+ truncated to Prop or (impredicative) Set; in particular, singleton type and
+ small inductive types, which have all eliminations to Type, are in Type *)
val get_sort_family_of :
- ?polyprop:bool -> env -> evar_map -> types -> Sorts.family
+ ?truncation_style:bool -> ?polyprop:bool -> env -> evar_map -> types -> Sorts.family
(** Makes an unsafe judgment from a constr *)
val get_judgment_of : env -> evar_map -> constr -> unsafe_judgment
diff --git a/printing/prettyp.ml b/printing/prettyp.ml
index 602b7a496..1eb2c31c8 100644
--- a/printing/prettyp.ml
+++ b/printing/prettyp.ml
@@ -360,10 +360,10 @@ let pr_located_qualid = function
str "Notation" ++ spc () ++ pr_path (Nametab.path_of_syndef kn)
| Dir dir ->
let s,dir = match dir with
- | DirOpenModule (dir,_) -> "Open Module", dir
- | DirOpenModtype (dir,_) -> "Open Module Type", dir
- | DirOpenSection (dir,_) -> "Open Section", dir
- | DirModule (dir,_) -> "Module", dir
+ | DirOpenModule { obj_dir ; _ } -> "Open Module", obj_dir
+ | DirOpenModtype { obj_dir ; _ } -> "Open Module Type", obj_dir
+ | DirOpenSection { obj_dir ; _ } -> "Open Section", obj_dir
+ | DirModule { obj_dir ; _ } -> "Module", obj_dir
| DirClosedSection dir -> "Closed Section", dir
in
str s ++ spc () ++ DirPath.print dir
@@ -410,7 +410,7 @@ let locate_term qid =
let locate_module qid =
let all = Nametab.locate_extended_all_dir qid in
let map dir = match dir with
- | DirModule (_, (mp, _)) -> Some (Dir dir, Nametab.shortest_qualid_of_module mp)
+ | DirModule { obj_mp ; _ } -> Some (Dir dir, Nametab.shortest_qualid_of_module obj_mp)
| DirOpenModule _ -> Some (Dir dir, qid)
| _ -> None
in
@@ -649,8 +649,8 @@ let gallina_print_library_entry env sigma with_values ent =
Some (str " >>>>>>> Section " ++ pr_name oname)
| (oname,Lib.ClosedSection _) ->
Some (str " >>>>>>> Closed Section " ++ pr_name oname)
- | (_,Lib.CompilingLibrary (dir,_)) ->
- Some (str " >>>>>>> Library " ++ DirPath.print dir)
+ | (_,Lib.CompilingLibrary { obj_dir; _ }) ->
+ Some (str " >>>>>>> Library " ++ DirPath.print obj_dir)
| (oname,Lib.OpenedModule _) ->
Some (str " >>>>>>> Module " ++ pr_name oname)
| (oname,Lib.ClosedModule _) ->
@@ -779,8 +779,8 @@ let read_sec_context r =
with Not_found ->
user_err ?loc ~hdr:"read_sec_context" (str "Unknown section.") in
let rec get_cxt in_cxt = function
- | (_,Lib.OpenedSection ((dir',_),_) as hd)::rest ->
- if DirPath.equal dir dir' then (hd::in_cxt) else get_cxt (hd::in_cxt) rest
+ | (_,Lib.OpenedSection ({obj_dir;_},_) as hd)::rest ->
+ if DirPath.equal dir obj_dir then (hd::in_cxt) else get_cxt (hd::in_cxt) rest
| (_,Lib.ClosedSection _)::rest ->
user_err Pp.(str "Cannot print the contents of a closed section.")
(* LEM: Actually, we could if we wanted to. *)
@@ -811,7 +811,7 @@ let print_any_name env sigma na udecl =
| Term (ConstructRef ((sp,_),_)) -> print_inductive sp udecl
| Term (VarRef sp) -> print_section_variable env sigma sp
| Syntactic kn -> print_syntactic_def env kn
- | Dir (DirModule(dirpath,(mp,_))) -> print_module (printable_body dirpath) mp
+ | Dir (DirModule { obj_dir; obj_mp; _ } ) -> print_module (printable_body obj_dir) obj_mp
| Dir _ -> mt ()
| ModuleType mp -> print_modtype mp
| Other (obj, info) -> info.print obj
diff --git a/printing/printmod.ml b/printing/printmod.ml
index c34543bbd..05292b06b 100644
--- a/printing/printmod.ml
+++ b/printing/printmod.ml
@@ -245,10 +245,10 @@ let print_kn locals kn =
with
Not_found -> print_modpath locals kn
-let nametab_register_dir mp =
+let nametab_register_dir obj_mp =
let id = mk_fake_top () in
- let dir = DirPath.make [id] in
- Nametab.push_dir (Nametab.Until 1) dir (DirModule (dir,(mp,DirPath.empty)))
+ let obj_dir = DirPath.make [id] in
+ Nametab.push_dir (Nametab.Until 1) obj_dir (DirModule { obj_dir; obj_mp; obj_sec = DirPath.empty })
(** Nota: the [global_reference] we register in the nametab below
might differ from internal ones, since we cannot recreate here
diff --git a/tactics/equality.ml b/tactics/equality.ml
index c36ad980e..0d6263246 100644
--- a/tactics/equality.ml
+++ b/tactics/equality.ml
@@ -739,7 +739,7 @@ let keep_proof_equalities = function
let find_positions env sigma ~keep_proofs ~no_discr t1 t2 =
let project env sorts posn t1 t2 =
let ty1 = get_type_of env sigma t1 in
- let s = get_sort_family_of env sigma ty1 in
+ let s = get_sort_family_of ~truncation_style:true env sigma ty1 in
if Sorts.List.mem s sorts
then [(List.rev posn,t1,t2)] else []
in
diff --git a/test-suite/bugs/closed/3125.v b/test-suite/bugs/closed/3125.v
new file mode 100644
index 000000000..797146174
--- /dev/null
+++ b/test-suite/bugs/closed/3125.v
@@ -0,0 +1,27 @@
+(* Not considering singleton template-polymorphic inductive types as
+ propositions for injection/inversion *)
+
+(* This is also #4560 and #6273 *)
+
+Inductive foo := foo_1.
+
+Goal forall (a b : foo), Some a = Some b -> a = b.
+Proof.
+ intros a b H.
+ inversion H.
+ reflexivity.
+Qed.
+
+(* Check that Prop is not concerned *)
+
+Inductive bar : Prop := bar_1.
+
+Goal
+ forall (a b : bar),
+ Some a = Some b ->
+ a = b.
+Proof.
+ intros a b H.
+ inversion H.
+ Fail reflexivity.
+Abort.
diff --git a/test-suite/bugs/closed/3559.v b/test-suite/bugs/closed/3559.v
index da12b6868..5210b2703 100644
--- a/test-suite/bugs/closed/3559.v
+++ b/test-suite/bugs/closed/3559.v
@@ -65,6 +65,7 @@ Axiom path_iff_hprop_uncurried : forall `{IsHProp A, IsHProp B}, (A <-> B) -> A
= B.
Inductive V : Type@{U'} := | set {A : Type@{U}} (f : A -> V) : V.
Axiom is0trunc_V : IsTrunc (trunc_S (trunc_S minus_two)) V.
+Existing Instance is0trunc_V.
Axiom bisimulation : V@{U' U} -> V@{U' U} -> hProp@{U'}.
Axiom bisimulation_refl : forall (v : V), bisimulation v v.
Axiom bisimulation_eq : forall (u v : V), bisimulation u v -> u = v.
diff --git a/test-suite/bugs/closed/HoTT_coq_064.v b/test-suite/bugs/closed/HoTT_coq_064.v
index b4c745375..d02a5f120 100644
--- a/test-suite/bugs/closed/HoTT_coq_064.v
+++ b/test-suite/bugs/closed/HoTT_coq_064.v
@@ -178,6 +178,7 @@ Definition IsColimit `{Funext} C D (F : Functor D C)
Generalizable All Variables.
Axiom fs : Funext.
+Existing Instance fs.
Section bar.
diff --git a/test-suite/success/Typeclasses.v b/test-suite/success/Typeclasses.v
index 6b1f0315b..cd6eac35c 100644
--- a/test-suite/success/Typeclasses.v
+++ b/test-suite/success/Typeclasses.v
@@ -240,3 +240,20 @@ Module IterativeDeepening.
Qed.
End IterativeDeepening.
+
+Module AxiomsAreInstances.
+ Set Typeclasses Axioms Are Instances.
+ Class TestClass1 := {}.
+ Axiom testax1 : TestClass1.
+ Definition testdef1 : TestClass1 := _.
+
+ Unset Typeclasses Axioms Are Instances.
+ Class TestClass2 := {}.
+ Axiom testax2 : TestClass2.
+ Fail Definition testdef2 : TestClass2 := _.
+
+ (* we didn't break typeclasses *)
+ Existing Instance testax2.
+ Definition testdef2 : TestClass2 := _.
+
+End AxiomsAreInstances.
diff --git a/test-suite/success/bteauto.v b/test-suite/success/bteauto.v
index 3178c6fc1..730b367d6 100644
--- a/test-suite/success/bteauto.v
+++ b/test-suite/success/bteauto.v
@@ -55,6 +55,7 @@ Module Backtracking.
Axiom A : Type.
Existing Class A.
Axioms a b c d e: A.
+ Existing Instances a b c d e.
Ltac get_value H := eval cbv delta [H] in H.
diff --git a/theories/Compat/Coq87.v b/theories/Compat/Coq87.v
index ef1737bf8..aeef9595d 100644
--- a/theories/Compat/Coq87.v
+++ b/theories/Compat/Coq87.v
@@ -15,3 +15,6 @@
and breaks at least fiat-crypto. *)
Declare ML Module "omega_plugin".
Unset Omega UseLocalDefs.
+
+
+Set Typeclasses Axioms Are Instances.
diff --git a/vernac/command.ml b/vernac/command.ml
index 373e5a1be..01c7f149b 100644
--- a/vernac/command.ml
+++ b/vernac/command.ml
@@ -174,6 +174,24 @@ let do_definition ident k univdecl bl red_option c ctypopt hook =
(* 2| Variable/Hypothesis/Parameter/Axiom declarations *)
+let axiom_into_instance = ref false
+
+let _ =
+ let open Goptions in
+ declare_bool_option
+ { optdepr = false;
+ optname = "automatically declare axioms whose type is a typeclass as instances";
+ optkey = ["Typeclasses";"Axioms";"Are";"Instances"];
+ optread = (fun _ -> !axiom_into_instance);
+ optwrite = (:=) axiom_into_instance; }
+
+let should_axiom_into_instance = function
+ | Discharge ->
+ (* The typeclass behaviour of Variable and Context doesn't depend
+ on section status *)
+ true
+ | Global | Local -> !axiom_into_instance
+
let declare_assumption is_coe (local,p,kind) (c,ctx) pl imps impl nl (_,ident) =
match local with
| Discharge when Lib.sections_are_opened () ->
@@ -195,6 +213,7 @@ match local with
(r,Univ.Instance.empty,true)
| Global | Local | Discharge ->
+ let do_instance = should_axiom_into_instance local in
let local = DeclareDef.get_locality ident ~kind:"axiom" local in
let inl = match nl with
| NoInline -> None
@@ -207,7 +226,7 @@ match local with
let () = maybe_declare_manual_implicits false gr imps in
let () = Universes.register_universe_binders gr pl in
let () = assumption_message ident in
- let () = Typeclasses.declare_instance None false gr in
+ let () = if do_instance then Typeclasses.declare_instance None false gr in
let () = if is_coe then Class.try_add_new_coercion gr ~local p in
let inst = match ctx with
| Polymorphic_const_entry ctx -> Univ.UContext.instance ctx
diff --git a/vernac/vernacentries.ml b/vernac/vernacentries.ml
index 1d9cb65f1..63f358a9d 100644
--- a/vernac/vernacentries.ml
+++ b/vernac/vernacentries.ml
@@ -186,8 +186,8 @@ let print_module r =
try
let globdir = Nametab.locate_dir qid in
match globdir with
- DirModule (dirpath,(mp,_)) ->
- Feedback.msg_notice (Printmod.print_module (Printmod.printable_body dirpath) mp)
+ DirModule { obj_dir; obj_mp; _ } ->
+ Feedback.msg_notice (Printmod.print_module (Printmod.printable_body obj_dir) obj_mp)
| _ -> raise Not_found
with
Not_found -> Feedback.msg_error (str"Unknown Module " ++ pr_qualid qid)