aboutsummaryrefslogtreecommitdiff
diff options
context:
space:
mode:
authorGaëtan Gilbert2017-10-31 17:04:02 +0100
committerGaëtan Gilbert2019-03-14 13:27:38 +0100
commit23f84f37c674a07e925925b7e0d50d7ee8414093 (patch)
tree7e470de5769c994d8df37c44fed12cf299d5b194
parent75508769762372043387c67a9abe94e8f940e80a (diff)
Add relevance marks on binders.
Kernel should be mostly correct, higher levels do random stuff at times.
-rw-r--r--checker/checkInductive.ml11
-rw-r--r--checker/checker.ml1
-rw-r--r--checker/values.ml22
-rw-r--r--clib/cArray.ml17
-rw-r--r--clib/cArray.mli2
-rw-r--r--clib/cList.ml7
-rw-r--r--clib/cList.mli3
-rw-r--r--dev/top_printers.ml9
-rw-r--r--doc/plugin_tutorial/tuto3/src/construction_game.ml5
-rw-r--r--doc/plugin_tutorial/tuto3/src/tuto_tactic.ml10
-rw-r--r--engine/eConstr.ml9
-rw-r--r--engine/eConstr.mli27
-rw-r--r--engine/evarutil.ml7
-rw-r--r--engine/namegen.ml33
-rw-r--r--engine/namegen.mli12
-rw-r--r--engine/nameops.ml6
-rw-r--r--engine/nameops.mli3
-rw-r--r--engine/proofview.ml4
-rw-r--r--engine/termops.ml31
-rw-r--r--engine/termops.mli12
-rw-r--r--interp/constrintern.ml13
-rw-r--r--interp/declare.ml2
-rw-r--r--interp/discharge.ml2
-rw-r--r--interp/impargs.ml5
-rw-r--r--interp/implicit_quantifiers.ml5
-rw-r--r--kernel/cClosure.ml12
-rw-r--r--kernel/cClosure.mli8
-rw-r--r--kernel/clambda.ml37
-rw-r--r--kernel/clambda.mli10
-rw-r--r--kernel/constr.ml63
-rw-r--r--kernel/constr.mli22
-rw-r--r--kernel/context.ml113
-rw-r--r--kernel/context.mli47
-rw-r--r--kernel/cooking.ml9
-rw-r--r--kernel/cooking.mli1
-rw-r--r--kernel/declarations.ml5
-rw-r--r--kernel/declareops.ml21
-rw-r--r--kernel/declareops.mli2
-rw-r--r--kernel/environ.ml2
-rw-r--r--kernel/indTyping.ml3
-rw-r--r--kernel/indtypes.ml43
-rw-r--r--kernel/inductive.ml32
-rw-r--r--kernel/inductive.mli4
-rw-r--r--kernel/kernel.mllib1
-rw-r--r--kernel/nativecode.ml9
-rw-r--r--kernel/nativelambda.ml21
-rw-r--r--kernel/nativelambda.mli14
-rw-r--r--kernel/reduction.ml12
-rw-r--r--kernel/retypeops.ml109
-rw-r--r--kernel/retypeops.mli26
-rw-r--r--kernel/safe_typing.ml16
-rw-r--r--kernel/sorts.ml19
-rw-r--r--kernel/sorts.mli10
-rw-r--r--kernel/subtyping.ml5
-rw-r--r--kernel/term.ml18
-rw-r--r--kernel/term.mli31
-rw-r--r--kernel/term_typing.ml11
-rw-r--r--kernel/term_typing.mli4
-rw-r--r--kernel/type_errors.ml19
-rw-r--r--kernel/type_errors.mli19
-rw-r--r--kernel/typeops.ml52
-rw-r--r--kernel/typeops.mli24
-rw-r--r--plugins/cc/ccalgo.ml20
-rw-r--r--plugins/cc/cctac.ml15
-rw-r--r--plugins/derive/derive.ml3
-rw-r--r--plugins/extraction/extract_env.ml2
-rw-r--r--plugins/extraction/extraction.ml55
-rw-r--r--plugins/extraction/table.ml4
-rw-r--r--plugins/firstorder/instances.ml4
-rw-r--r--plugins/firstorder/rules.ml6
-rw-r--r--plugins/firstorder/unify.ml4
-rw-r--r--plugins/funind/functional_principles_proofs.ml26
-rw-r--r--plugins/funind/functional_principles_types.ml51
-rw-r--r--plugins/funind/glob_term_to_relation.ml73
-rw-r--r--plugins/funind/indfun.ml8
-rw-r--r--plugins/funind/indfun_common.mli6
-rw-r--r--plugins/funind/invfun.ml37
-rw-r--r--plugins/funind/recdef.ml59
-rw-r--r--plugins/ltac/evar_tactics.ml6
-rw-r--r--plugins/ltac/extratactics.mlg3
-rw-r--r--plugins/ltac/pptactic.ml2
-rw-r--r--plugins/ltac/rewrite.ml66
-rw-r--r--plugins/ltac/tactic_matching.ml5
-rw-r--r--plugins/ltac/tauto.ml4
-rw-r--r--plugins/micromega/coq_micromega.ml11
-rw-r--r--plugins/omega/coq_omega.ml51
-rw-r--r--plugins/rtauto/refl_tauto.ml5
-rw-r--r--plugins/rtauto/refl_tauto.mli2
-rw-r--r--plugins/ssr/ssrcommon.ml67
-rw-r--r--plugins/ssr/ssrcommon.mli2
-rw-r--r--plugins/ssr/ssrelim.ml9
-rw-r--r--plugins/ssr/ssrequality.ml17
-rw-r--r--plugins/ssr/ssrfwd.ml17
-rw-r--r--plugins/ssr/ssripats.ml15
-rw-r--r--plugins/ssr/ssrtacticals.ml9
-rw-r--r--plugins/ssr/ssrview.ml3
-rw-r--r--plugins/ssrmatching/ssrmatching.ml7
-rw-r--r--pretyping/arguments_renaming.ml3
-rw-r--r--pretyping/cases.ml115
-rw-r--r--pretyping/cbv.ml14
-rw-r--r--pretyping/cbv.mli2
-rw-r--r--pretyping/classops.ml2
-rw-r--r--pretyping/coercion.ml44
-rw-r--r--pretyping/constr_matching.ml24
-rw-r--r--pretyping/detyping.ml51
-rw-r--r--pretyping/evarconv.ml32
-rw-r--r--pretyping/evardefine.ml29
-rw-r--r--pretyping/evardefine.mli2
-rw-r--r--pretyping/evarsolve.ml15
-rw-r--r--pretyping/find_subterm.ml2
-rw-r--r--pretyping/globEnv.ml2
-rw-r--r--pretyping/globEnv.mli2
-rw-r--r--pretyping/indrec.ml143
-rw-r--r--pretyping/inductiveops.ml30
-rw-r--r--pretyping/inductiveops.mli6
-rw-r--r--pretyping/nativenorm.ml51
-rw-r--r--pretyping/patternops.ml20
-rw-r--r--pretyping/pretype_errors.ml4
-rw-r--r--pretyping/pretype_errors.mli6
-rw-r--r--pretyping/pretyping.ml139
-rw-r--r--pretyping/pretyping.mllib4
-rw-r--r--pretyping/reductionops.ml35
-rw-r--r--pretyping/reductionops.mli6
-rw-r--r--pretyping/retyping.ml44
-rw-r--r--pretyping/retyping.mli5
-rw-r--r--pretyping/tacred.ml31
-rw-r--r--pretyping/typing.ml49
-rw-r--r--pretyping/typing.mli4
-rw-r--r--pretyping/unification.ml47
-rw-r--r--pretyping/vnorm.ml40
-rw-r--r--printing/prettyp.ml4
-rw-r--r--printing/printer.ml5
-rw-r--r--printing/printmod.ml5
-rw-r--r--printing/proof_diffs.ml6
-rw-r--r--proofs/clenv.ml9
-rw-r--r--proofs/proof_global.ml5
-rw-r--r--proofs/tacmach.ml2
-rw-r--r--proofs/tacmach.mli2
-rw-r--r--tactics/btermdn.ml2
-rw-r--r--tactics/contradiction.ml14
-rw-r--r--tactics/elimschemes.ml13
-rw-r--r--tactics/elimschemes.mli3
-rw-r--r--tactics/eqdecide.ml8
-rw-r--r--tactics/eqschemes.ml100
-rw-r--r--tactics/equality.ml23
-rw-r--r--tactics/hints.ml5
-rw-r--r--tactics/hipattern.ml8
-rw-r--r--tactics/hipattern.mli2
-rw-r--r--tactics/inv.ml3
-rw-r--r--tactics/leminv.ml18
-rw-r--r--tactics/tactics.ml126
-rw-r--r--tactics/term_dnet.ml2
-rw-r--r--test-suite/misc/poly-capture-global-univs/src/evilImpl.ml4
-rw-r--r--vernac/assumptions.ml4
-rw-r--r--vernac/auto_ind_decl.ml139
-rw-r--r--vernac/class.ml5
-rw-r--r--vernac/classes.ml4
-rw-r--r--vernac/comAssumption.ml4
-rw-r--r--vernac/comFixpoint.ml35
-rw-r--r--vernac/comFixpoint.mli4
-rw-r--r--vernac/comInductive.ml23
-rw-r--r--vernac/comProgramFixpoint.ml30
-rw-r--r--vernac/himsg.ml27
-rw-r--r--vernac/indschemes.ml9
-rw-r--r--vernac/lemmas.ml2
-rw-r--r--vernac/obligations.ml18
-rw-r--r--vernac/record.ml49
-rw-r--r--vernac/vernacentries.ml4
168 files changed, 2129 insertions, 1386 deletions
diff --git a/checker/checkInductive.ml b/checker/checkInductive.ml
index b681fb876e..0eacc24626 100644
--- a/checker/checkInductive.ml
+++ b/checker/checkInductive.ml
@@ -25,7 +25,7 @@ let to_entry (mb:mutual_inductive_body) : Entries.mutual_inductive_entry =
let nparams = List.length mb.mind_params_ctxt in (* include letins *)
let mind_entry_record = match mb.mind_record with
| NotRecord -> None | FakeRecord -> Some None
- | PrimRecord data -> Some (Some (Array.map pi1 data))
+ | PrimRecord data -> Some (Some (Array.map (fun (x,_,_,_) -> x) data))
in
let mind_entry_universes = match mb.mind_universes with
| Monomorphic univs -> Monomorphic_entry univs
@@ -95,8 +95,8 @@ let eq_in_context (ctx1, t1) (ctx2, t2) =
let check_packet env mind ind
{ mind_typename; mind_arity_ctxt; mind_arity; mind_consnames; mind_user_lc;
mind_nrealargs; mind_nrealdecls; mind_kelim; mind_nf_lc;
- mind_consnrealargs; mind_consnrealdecls; mind_recargs; mind_nb_constant;
- mind_nb_args; mind_reloc_tbl } =
+ mind_consnrealargs; mind_consnrealdecls; mind_recargs; mind_relevant;
+ mind_nb_constant; mind_nb_args; mind_reloc_tbl } =
let check = check mind in
ignore mind_typename; (* passed through *)
@@ -117,6 +117,8 @@ let check_packet env mind ind
check "mind_recargs" (Rtree.equal eq_recarg ind.mind_recargs mind_recargs);
+ check "mind_relevant" (Sorts.relevance_equal ind.mind_relevant mind_relevant);
+
check "mind_nb_args" Int.(equal ind.mind_nb_args mind_nb_args);
check "mind_nb_constant" Int.(equal ind.mind_nb_constant mind_nb_constant);
check "mind_reloc_tbl" (eq_reloc_tbl ind.mind_reloc_tbl mind_reloc_tbl);
@@ -128,7 +130,8 @@ let check_same_record r1 r2 = match r1, r2 with
| PrimRecord r1, PrimRecord r2 ->
(* The kernel doesn't care about the names, we just need to check
that the saved types are correct. *)
- Array.for_all2 (fun (_,_,tys1) (_,_,tys2) ->
+ Array.for_all2 (fun (_,_,r1,tys1) (_,_,r2,tys2) ->
+ Array.equal Sorts.relevance_equal r1 r2 &&
Array.equal Constr.equal tys1 tys2)
r1 r2
| (NotRecord | FakeRecord | PrimRecord _), _ -> false
diff --git a/checker/checker.ml b/checker/checker.ml
index 205a3984d5..9be88ee31e 100644
--- a/checker/checker.ml
+++ b/checker/checker.ml
@@ -298,6 +298,7 @@ let explain_exn = function
| IllTypedRecBody _ -> str"IllTypedRecBody"
| UnsatisfiedConstraints _ -> str"UnsatisfiedConstraints"
| DisallowedSProp -> str"DisallowedSProp"
+ | BadRelevance -> str"BadRelevance"
| UndeclaredUniverse _ -> str"UndeclaredUniverse"))
| InductiveError e ->
diff --git a/checker/values.ml b/checker/values.ml
index f2b961ef56..5cbf0ff298 100644
--- a/checker/values.ml
+++ b/checker/values.ml
@@ -119,6 +119,9 @@ let v_context_set = v_tuple "universe_context_set" [|v_hset v_level;v_cstrs|]
let v_sort = v_sum "sort" 3 (*SProp, Prop, Set*) [|[|v_univ(*Type*)|]|]
let v_sortfam = v_enum "sorts_family" 4
+let v_relevance = v_sum "relevance" 2 [||]
+let v_binder_annot x = v_tuple "binder_annot" [|x;v_relevance|]
+
let v_puniverses v = v_tuple "punivs" [|v;v_instance|]
let v_boollist = List v_bool
@@ -126,7 +129,7 @@ let v_boollist = List v_bool
let v_caseinfo =
let v_cstyle = v_enum "case_style" 5 in
let v_cprint = v_tuple "case_printing" [|v_boollist;Array v_boollist;v_cstyle|] in
- v_tuple "case_info" [|v_ind;Int;Array Int;Array Int;v_cprint|]
+ v_tuple "case_info" [|v_ind;Int;Array Int;Array Int;v_relevance;v_cprint|]
let v_cast = v_enum "cast_kind" 4
@@ -141,9 +144,9 @@ let rec v_constr =
[|Fail "Evar"|]; (* Evar *)
[|v_sort|]; (* Sort *)
[|v_constr;v_cast;v_constr|]; (* Cast *)
- [|v_name;v_constr;v_constr|]; (* Prod *)
- [|v_name;v_constr;v_constr|]; (* Lambda *)
- [|v_name;v_constr;v_constr;v_constr|]; (* LetIn *)
+ [|v_binder_annot v_name;v_constr;v_constr|]; (* Prod *)
+ [|v_binder_annot v_name;v_constr;v_constr|]; (* Lambda *)
+ [|v_binder_annot v_name;v_constr;v_constr;v_constr|]; (* LetIn *)
[|v_constr;Array v_constr|]; (* App *)
[|v_puniverses v_cst|]; (* Const *)
[|v_puniverses v_ind|]; (* Ind *)
@@ -156,12 +159,13 @@ let rec v_constr =
|])
and v_prec = Tuple ("prec_declaration",
- [|Array v_name; Array v_constr; Array v_constr|])
+ [|Array (v_binder_annot v_name); Array v_constr; Array v_constr|])
and v_fix = Tuple ("pfixpoint", [|Tuple ("fix2",[|Array Int;Int|]);v_prec|])
and v_cofix = Tuple ("pcofixpoint",[|Int;v_prec|])
-let v_rdecl = v_sum "rel_declaration" 0 [| [|v_name; v_constr|]; (* LocalAssum *)
- [|v_name; v_constr; v_constr|] |] (* LocalDef *)
+let v_rdecl = v_sum "rel_declaration" 0
+ [| [|v_binder_annot v_name; v_constr|]; (* LocalAssum *)
+ [|v_binder_annot v_name; v_constr; v_constr|] |] (* LocalDef *)
let v_rctxt = List v_rdecl
let v_section_ctxt = v_enum "emptylist" 1
@@ -231,6 +235,7 @@ let v_cb = v_tuple "constant_body"
[|v_section_ctxt;
v_cst_def;
v_constr;
+ v_relevance;
Any;
v_univs;
Opt v_context_set;
@@ -265,6 +270,7 @@ let v_one_ind = v_tuple "one_inductive_body"
Array Int;
Array Int;
v_wfp;
+ v_relevance;
Int;
Int;
Any|]
@@ -273,7 +279,7 @@ let v_finite = v_enum "recursivity_kind" 3
let v_record_info =
v_sum "record_info" 2
- [| [| Array (v_tuple "record" [| v_id; Array v_id; Array v_constr |]) |] |]
+ [| [| Array (v_tuple "record" [| v_id; Array v_id; Array v_relevance; Array v_constr |]) |] |]
let v_ind_pack = v_tuple "mutual_inductive_body"
[|Array v_one_ind;
diff --git a/clib/cArray.ml b/clib/cArray.ml
index e0a1859184..145a32cf45 100644
--- a/clib/cArray.ml
+++ b/clib/cArray.ml
@@ -52,6 +52,8 @@ sig
val map2_i : (int -> 'a -> 'b -> 'c) -> 'a array -> 'b array -> 'c array
val map3 :
('a -> 'b -> 'c -> 'd) -> 'a array -> 'b array -> 'c array -> 'd array
+ val map3_i :
+ (int -> 'a -> 'b -> 'c -> 'd) -> 'a array -> 'b array -> 'c array -> 'd array
val map_left : ('a -> 'b) -> 'a array -> 'b array
val iter2_i : (int -> 'a -> 'b -> unit) -> 'a array -> 'b array -> unit
val fold_left_map : ('a -> 'b -> 'a * 'c) -> 'a -> 'b array -> 'a * 'c array
@@ -358,6 +360,21 @@ let map3 f v1 v2 v3 =
res
end
+let map3_i f v1 v2 v3 =
+ let len1 = Array.length v1 in
+ let len2 = Array.length v2 in
+ let len3 = Array.length v3 in
+ let () = if not (Int.equal len1 len2 && Int.equal len1 len3) then invalid_arg "Array.map3_i" in
+ if Int.equal len1 0 then
+ [| |]
+ else begin
+ let res = Array.make len1 (f 0 (uget v1 0) (uget v2 0) (uget v3 0)) in
+ for i = 1 to pred len1 do
+ Array.unsafe_set res i (f i (uget v1 i) (uget v2 i) (uget v3 i))
+ done;
+ res
+ end
+
let map_left f a = (* Ocaml does not guarantee Array.map is LR *)
let l = Array.length a in (* (even if so), then we rewrite it *)
if Int.equal l 0 then [||] else begin
diff --git a/clib/cArray.mli b/clib/cArray.mli
index 21479d2b45..8aca31ad33 100644
--- a/clib/cArray.mli
+++ b/clib/cArray.mli
@@ -83,6 +83,8 @@ sig
val map2_i : (int -> 'a -> 'b -> 'c) -> 'a array -> 'b array -> 'c array
val map3 :
('a -> 'b -> 'c -> 'd) -> 'a array -> 'b array -> 'c array -> 'd array
+ val map3_i :
+ (int -> 'a -> 'b -> 'c -> 'd) -> 'a array -> 'b array -> 'c array -> 'd array
val map_left : ('a -> 'b) -> 'a array -> 'b array
(** As [map] but guaranteed to be left-to-right. *)
diff --git a/clib/cList.ml b/clib/cList.ml
index 524945ef23..aa01f6e5b5 100644
--- a/clib/cList.ml
+++ b/clib/cList.ml
@@ -98,6 +98,7 @@ sig
val split : ('a * 'b) list -> 'a list * 'b list
val combine : 'a list -> 'b list -> ('a * 'b) list
val split3 : ('a * 'b * 'c) list -> 'a list * 'b list * 'c list
+ val split4 : ('a * 'b * 'c * 'd) list -> 'a list * 'b list * 'c list * 'd list
val combine3 : 'a list -> 'b list -> 'c list -> ('a * 'b * 'c) list
val add_set : 'a eq -> 'a -> 'a list -> 'a list
val eq_set : 'a eq -> 'a list -> 'a list -> bool
@@ -846,6 +847,12 @@ let split3 = function
split3_loop cp cq cr l;
(cast cp, cast cq, cast cr)
+(** XXX TODO tailrec *)
+let rec split4 = function
+ | [] -> ([], [], [], [])
+ | (a,b,c,d)::l ->
+ let (ra, rb, rc, rd) = split4 l in (a::ra, b::rb, c::rc, d::rd)
+
let rec combine3_loop p l1 l2 l3 = match l1, l2, l3 with
| [], [], [] -> ()
| x :: l1, y :: l2, z :: l3 ->
diff --git a/clib/cList.mli b/clib/cList.mli
index 8582e6cd65..a2fe0b759a 100644
--- a/clib/cList.mli
+++ b/clib/cList.mli
@@ -308,6 +308,9 @@ sig
val split3 : ('a * 'b * 'c) list -> 'a list * 'b list * 'c list
(** Like [split] but for triples *)
+ val split4 : ('a * 'b * 'c * 'd) list -> 'a list * 'b list * 'c list * 'd list
+ (** Like [split] but for quads *)
+
val combine3 : 'a list -> 'b list -> 'c list -> ('a * 'b * 'c) list
(** Like [combine] but for triples *)
diff --git a/dev/top_printers.ml b/dev/top_printers.ml
index 5ece892aa2..0fbb0634a5 100644
--- a/dev/top_printers.ml
+++ b/dev/top_printers.ml
@@ -20,6 +20,7 @@ open Univ
open Environ
open Printer
open Constr
+open Context
open Genarg
open Clenv
@@ -316,7 +317,7 @@ let constr_display csr =
Array.fold_right (fun x i -> level_display x; (string_of_int !cnt)^(if not(i="")
then (" "^i) else "")) (Instance.to_array l) ""
- and name_display = function
+ and name_display x = match x.binder_name with
| Name id -> "Name("^(Id.to_string id)^")"
| Anonymous -> "Anonymous"
@@ -336,13 +337,13 @@ let print_pure_constr csr =
| Cast (c,_, t) -> open_hovbox 1;
print_string "("; (term_display c); print_cut();
print_string "::"; (term_display t); print_string ")"; close_box()
- | Prod (Name(id),t,c) ->
+ | Prod ({binder_name=Name(id)},t,c) ->
open_hovbox 1;
print_string"("; print_string (Id.to_string id);
print_string ":"; box_display t;
print_string ")"; print_cut();
box_display c; close_box()
- | Prod (Anonymous,t,c) ->
+ | Prod ({binder_name=Anonymous},t,c) ->
print_string"("; box_display t; print_cut(); print_string "->";
box_display c; print_string ")";
| Lambda (na,t,c) ->
@@ -437,7 +438,7 @@ let print_pure_constr csr =
| Type u -> open_hbox();
print_string "Type("; pp (pr_uni u); print_string ")"; close_box()
- and name_display = function
+ and name_display x = match x.binder_name with
| Name id -> print_string (Id.to_string id)
| Anonymous -> print_string "_"
(* Remove the top names for library and Scratch to avoid long names *)
diff --git a/doc/plugin_tutorial/tuto3/src/construction_game.ml b/doc/plugin_tutorial/tuto3/src/construction_game.ml
index 9d9f894e18..663113d012 100644
--- a/doc/plugin_tutorial/tuto3/src/construction_game.ml
+++ b/doc/plugin_tutorial/tuto3/src/construction_game.ml
@@ -1,4 +1,5 @@
open Pp
+open Context
let find_reference = Coqlib.find_reference [@ocaml.warning "-3"]
@@ -32,7 +33,7 @@ let dangling_identity env evd =
let evd, arg_type = Evarutil.new_evar env evd type_type in
(* Notice the use of a De Bruijn index for the inner occurrence of the
bound variable. *)
- evd, EConstr.mkLambda(Names.Name (Names.Id.of_string "x"), arg_type,
+ evd, EConstr.mkLambda(nameR (Names.Id.of_string "x"), arg_type,
EConstr.mkRel 1)
let dangling_identity2 env evd =
@@ -40,7 +41,7 @@ let dangling_identity2 env evd =
is meant to be a type. *)
let evd, (arg_type, type_type) =
Evarutil.new_type_evar env evd Evd.univ_rigid in
- evd, EConstr.mkLambda(Names.Name (Names.Id.of_string "x"), arg_type,
+ evd, EConstr.mkLambda(nameR (Names.Id.of_string "x"), arg_type,
EConstr.mkRel 1)
let example_sort_app_lambda () =
diff --git a/doc/plugin_tutorial/tuto3/src/tuto_tactic.ml b/doc/plugin_tutorial/tuto3/src/tuto_tactic.ml
index 8f2c387d09..2d541087ce 100644
--- a/doc/plugin_tutorial/tuto3/src/tuto_tactic.ml
+++ b/doc/plugin_tutorial/tuto3/src/tuto_tactic.ml
@@ -116,11 +116,11 @@ let repackage i h_hyps_id = Goal.enter begin fun gl ->
mkApp (c_U (), [| ty2; mkVar h_hyps_id|]) |]) in
Refine.refine ~typecheck:true begin fun evd ->
let evd, new_goal = Evarutil.new_evar env evd
- (mkProd (Names.Name.Anonymous,
- mkApp(c_H (), [| new_packed_type |]),
- Vars.lift 1 concl)) in
- evd, mkApp (new_goal,
- [|mkApp(c_M (), [|new_packed_type; new_packed_value |]) |])
+ (mkArrowR (mkApp(c_H (), [| new_packed_type |]))
+ (Vars.lift 1 concl))
+ in
+ evd, mkApp (new_goal,
+ [|mkApp(c_M (), [|new_packed_type; new_packed_value |]) |])
end
end
diff --git a/engine/eConstr.ml b/engine/eConstr.ml
index 521d6b707d..981f9454e4 100644
--- a/engine/eConstr.ml
+++ b/engine/eConstr.ml
@@ -73,7 +73,8 @@ let mkCase (ci, c, r, p) = of_kind (Case (ci, c, r, p))
let mkFix f = of_kind (Fix f)
let mkCoFix f = of_kind (CoFix f)
let mkProj (p, c) = of_kind (Proj (p, c))
-let mkArrow t1 t2 = of_kind (Prod (Anonymous, t1, t2))
+let mkArrow t1 r t2 = of_kind (Prod (make_annot Anonymous r, t1, t2))
+let mkArrowR t1 t2 = mkArrow t1 Sorts.Relevant t2
let mkInt i = of_kind (Int i)
let mkRef (gr,u) = let open GlobRef in match gr with
@@ -668,9 +669,9 @@ let mkLambda_or_LetIn decl c =
| LocalAssum (na,t) -> mkLambda (na, t, c)
| LocalDef (na,b,t) -> mkLetIn (na, b, t, c)
-let mkNamedProd id typ c = mkProd (Name id, typ, Vars.subst_var id c)
-let mkNamedLambda id typ c = mkLambda (Name id, typ, Vars.subst_var id c)
-let mkNamedLetIn id c1 t c2 = mkLetIn (Name id, c1, t, Vars.subst_var id c2)
+let mkNamedProd id typ c = mkProd (map_annot Name.mk_name id, typ, Vars.subst_var id.binder_name c)
+let mkNamedLambda id typ c = mkLambda (map_annot Name.mk_name id, typ, Vars.subst_var id.binder_name c)
+let mkNamedLetIn id c1 t c2 = mkLetIn (map_annot Name.mk_name id, c1, t, Vars.subst_var id.binder_name c2)
let mkNamedProd_or_LetIn decl c =
let open Context.Named.Declaration in
diff --git a/engine/eConstr.mli b/engine/eConstr.mli
index 6a144be6b3..25ceffbd04 100644
--- a/engine/eConstr.mli
+++ b/engine/eConstr.mli
@@ -109,9 +109,9 @@ val mkProp : t
val mkSet : t
val mkType : Univ.Universe.t -> t
val mkCast : t * cast_kind * t -> t
-val mkProd : Name.t * t * t -> t
-val mkLambda : Name.t * t * t -> t
-val mkLetIn : Name.t * t * t * t -> t
+val mkProd : Name.t Context.binder_annot * t * t -> t
+val mkLambda : Name.t Context.binder_annot * t * t -> t
+val mkLetIn : Name.t Context.binder_annot * t * t * t -> t
val mkApp : t * t array -> t
val mkConst : Constant.t -> t
val mkConstU : Constant.t * EInstance.t -> t
@@ -124,7 +124,8 @@ val mkConstructUi : (inductive * EInstance.t) * int -> t
val mkCase : case_info * t * t * t array -> t
val mkFix : (t, t) pfixpoint -> t
val mkCoFix : (t, t) pcofixpoint -> t
-val mkArrow : t -> t -> t
+val mkArrow : t -> Sorts.relevance -> t -> t
+val mkArrowR : t -> t -> t
val mkInt : Uint63.t -> t
val mkRef : GlobRef.t * EInstance.t -> t
@@ -139,9 +140,9 @@ val mkLambda_or_LetIn : rel_declaration -> t -> t
val it_mkProd_or_LetIn : t -> rel_context -> t
val it_mkLambda_or_LetIn : t -> rel_context -> t
-val mkNamedLambda : Id.t -> types -> constr -> constr
-val mkNamedLetIn : Id.t -> constr -> types -> constr -> constr
-val mkNamedProd : Id.t -> types -> types -> types
+val mkNamedLambda : Id.t Context.binder_annot -> types -> constr -> constr
+val mkNamedLetIn : Id.t Context.binder_annot -> constr -> types -> constr -> constr
+val mkNamedProd : Id.t Context.binder_annot -> types -> types -> types
val mkNamedLambda_or_LetIn : named_declaration -> types -> types
val mkNamedProd_or_LetIn : named_declaration -> types -> types
@@ -179,9 +180,9 @@ val destMeta : Evd.evar_map -> t -> metavariable
val destVar : Evd.evar_map -> t -> Id.t
val destSort : Evd.evar_map -> t -> ESorts.t
val destCast : Evd.evar_map -> t -> t * cast_kind * t
-val destProd : Evd.evar_map -> t -> Name.t * types * types
-val destLambda : Evd.evar_map -> t -> Name.t * types * t
-val destLetIn : Evd.evar_map -> t -> Name.t * t * types * t
+val destProd : Evd.evar_map -> t -> Name.t Context.binder_annot * types * types
+val destLambda : Evd.evar_map -> t -> Name.t Context.binder_annot * types * t
+val destLetIn : Evd.evar_map -> t -> Name.t Context.binder_annot * t * types * t
val destApp : Evd.evar_map -> t -> t * t array
val destConst : Evd.evar_map -> t -> Constant.t * EInstance.t
val destEvar : Evd.evar_map -> t -> t pexistential
@@ -197,7 +198,7 @@ val destRef : Evd.evar_map -> t -> GlobRef.t * EInstance.t
val decompose_app : Evd.evar_map -> t -> t * t list
(** Pops lambda abstractions until there are no more, skipping casts. *)
-val decompose_lam : Evd.evar_map -> t -> (Name.t * t) list * t
+val decompose_lam : Evd.evar_map -> t -> (Name.t Context.binder_annot * t) list * t
(** Pops lambda abstractions and letins until there are no more, skipping casts. *)
val decompose_lam_assum : Evd.evar_map -> t -> rel_context * t
@@ -213,10 +214,10 @@ val decompose_lam_n_assum : Evd.evar_map -> int -> t -> rel_context * t
@raise UserError if the term doesn't have enough lambdas/letins. *)
val decompose_lam_n_decls : Evd.evar_map -> int -> t -> rel_context * t
-val compose_lam : (Name.t * t) list -> t -> t
+val compose_lam : (Name.t Context.binder_annot * t) list -> t -> t
val to_lambda : Evd.evar_map -> int -> t -> t
-val decompose_prod : Evd.evar_map -> t -> (Name.t * t) list * t
+val decompose_prod : Evd.evar_map -> t -> (Name.t Context.binder_annot * t) list * t
val decompose_prod_assum : Evd.evar_map -> t -> rel_context * t
val decompose_prod_n_assum : Evd.evar_map -> int -> t -> rel_context * t
diff --git a/engine/evarutil.ml b/engine/evarutil.ml
index 5c2b480db4..96beb72a56 100644
--- a/engine/evarutil.ml
+++ b/engine/evarutil.ml
@@ -11,6 +11,7 @@
open CErrors
open Util
open Names
+open Context
open Constr
open Environ
open Evd
@@ -781,13 +782,13 @@ let cached_evar_of_hyp cache sigma decl accu = match cache with
in
NamedDecl.fold_constr fold decl accu
| Some cache ->
- let id = NamedDecl.get_id decl in
+ let id = NamedDecl.get_annot decl in
let r =
- try Id.Map.find id cache.cache
+ try Id.Map.find id.binder_name cache.cache
with Not_found ->
(* Dummy value *)
let r = ref (NamedDecl.LocalAssum (id, EConstr.mkProp), Evar.Set.empty) in
- let () = cache.cache <- Id.Map.add id r cache.cache in
+ let () = cache.cache <- Id.Map.add id.binder_name r cache.cache in
r
in
let (decl', evs) = !r in
diff --git a/engine/namegen.ml b/engine/namegen.ml
index 46fdf08e10..10ece55a63 100644
--- a/engine/namegen.ml
+++ b/engine/namegen.ml
@@ -18,6 +18,7 @@ open Util
open Names
open Term
open Constr
+open Context
open Environ
open EConstr
open Vars
@@ -117,7 +118,7 @@ let head_name sigma c = (* Find the head constant of a constr if any *)
| Const _ | Ind _ | Construct _ | Var _ as c ->
Some (Nametab.basename_of_global (global_of_constr c))
| Fix ((_,i),(lna,_,_)) | CoFix (i,(lna,_,_)) ->
- Some (match lna.(i) with Name id -> id | _ -> assert false)
+ Some (match lna.(i).binder_name with Name id -> id | _ -> assert false)
| Sort _ | Rel _ | Meta _|Evar _|Case (_, _, _, _) | Int _ -> None
in
hdrec c
@@ -155,12 +156,12 @@ let hdchar env sigma c =
| Rel n ->
(if n<=k then "p" (* the initial term is flexible product/function *)
else
- try match lookup_rel (n-k) env with
- | LocalAssum (Name id,_) | LocalDef (Name id,_,_) -> lowercase_first_char id
- | LocalAssum (Anonymous,t) | LocalDef (Anonymous,_,t) -> hdrec 0 (lift (n-k) t)
+ try match let d = lookup_rel (n-k) env in get_name d, get_type d with
+ | Name id, _ -> lowercase_first_char id
+ | Anonymous, t -> hdrec 0 (lift (n-k) t)
with Not_found -> "y")
| Fix ((_,i),(lna,_,_)) | CoFix (i,(lna,_,_)) ->
- let id = match lna.(i) with Name id -> id | _ -> assert false in
+ let id = match lna.(i).binder_name with Name id -> id | _ -> assert false in
lowercase_first_char id
| Evar _ (* We could do better... *)
| Meta _ | Case (_, _, _, _) -> "y"
@@ -176,18 +177,20 @@ let named_hd env sigma a = function
| Anonymous -> Name (Id.of_string (hdchar env sigma a))
| x -> x
-let mkProd_name env sigma (n,a,b) = mkProd (named_hd env sigma a n, a, b)
-let mkLambda_name env sigma (n,a,b) = mkLambda (named_hd env sigma a n, a, b)
+let mkProd_name env sigma (n,a,b) = mkProd (map_annot (named_hd env sigma a) n, a, b)
+let mkLambda_name env sigma (n,a,b) = mkLambda (map_annot (named_hd env sigma a) n, a, b)
let lambda_name = mkLambda_name
let prod_name = mkProd_name
-let prod_create env sigma (a,b) = mkProd (named_hd env sigma a Anonymous, a, b)
-let lambda_create env sigma (a,b) = mkLambda (named_hd env sigma a Anonymous, a, b)
+let prod_create env sigma (r,a,b) =
+ mkProd (make_annot (named_hd env sigma a Anonymous) r, a, b)
+let lambda_create env sigma (r,a,b) =
+ mkLambda (make_annot (named_hd env sigma a Anonymous) r, a, b)
let name_assumption env sigma = function
- | LocalAssum (na,t) -> LocalAssum (named_hd env sigma t na, t)
- | LocalDef (na,c,t) -> LocalDef (named_hd env sigma c na, c, t)
+ | LocalAssum (na,t) -> LocalAssum (map_annot (named_hd env sigma t) na, t)
+ | LocalDef (na,c,t) -> LocalDef (map_annot (named_hd env sigma c) na, c, t)
let name_context env sigma hyps =
snd
@@ -457,13 +460,13 @@ let rename_bound_vars_as_displayed sigma avoid env c =
| Prod (na,c1,c2) ->
let na',avoid' =
compute_displayed_name_in sigma
- (RenamingElsewhereFor (env,c2)) avoid na c2 in
- mkProd (na', c1, rename avoid' (na' :: env) c2)
+ (RenamingElsewhereFor (env,c2)) avoid na.binder_name c2 in
+ mkProd ({na with binder_name=na'}, c1, rename avoid' (na' :: env) c2)
| LetIn (na,c1,t,c2) ->
let na',avoid' =
compute_displayed_let_name_in sigma
- (RenamingElsewhereFor (env,c2)) avoid na c2 in
- mkLetIn (na',c1,t, rename avoid' (na' :: env) c2)
+ (RenamingElsewhereFor (env,c2)) avoid na.binder_name c2 in
+ mkLetIn ({na with binder_name=na'},c1,t, rename avoid' (na' :: env) c2)
| Cast (c,k,t) -> mkCast (rename avoid env c, k,t)
| _ -> c
in
diff --git a/engine/namegen.mli b/engine/namegen.mli
index 3722cbed24..240fd8fa81 100644
--- a/engine/namegen.mli
+++ b/engine/namegen.mli
@@ -44,15 +44,15 @@ val id_of_name_using_hdchar : env -> evar_map -> types -> Name.t -> Id.t
val named_hd : env -> evar_map -> types -> Name.t -> Name.t
val head_name : evar_map -> types -> Id.t option
-val mkProd_name : env -> evar_map -> Name.t * types * types -> types
-val mkLambda_name : env -> evar_map -> Name.t * types * constr -> constr
+val mkProd_name : env -> evar_map -> Name.t Context.binder_annot * types * types -> types
+val mkLambda_name : env -> evar_map -> Name.t Context.binder_annot * types * constr -> constr
(** Deprecated synonyms of [mkProd_name] and [mkLambda_name] *)
-val prod_name : env -> evar_map -> Name.t * types * types -> types
-val lambda_name : env -> evar_map -> Name.t * types * constr -> constr
+val prod_name : env -> evar_map -> Name.t Context.binder_annot * types * types -> types
+val lambda_name : env -> evar_map -> Name.t Context.binder_annot * types * constr -> constr
-val prod_create : env -> evar_map -> types * types -> constr
-val lambda_create : env -> evar_map -> types * constr -> constr
+val prod_create : env -> evar_map -> Sorts.relevance * types * types -> constr
+val lambda_create : env -> evar_map -> Sorts.relevance * types * constr -> constr
val name_assumption : env -> evar_map -> rel_declaration -> rel_declaration
val name_context : env -> evar_map -> rel_context -> rel_context
diff --git a/engine/nameops.ml b/engine/nameops.ml
index 15e201347c..2047772cfe 100644
--- a/engine/nameops.ml
+++ b/engine/nameops.ml
@@ -132,6 +132,7 @@ sig
val fold_right_map : (Id.t -> 'a -> Id.t * 'a) -> Name.t -> 'a -> Name.t * 'a
val get_id : t -> Id.t
val pick : t -> t -> t
+ val pick_annot : t Context.binder_annot -> t Context.binder_annot -> t Context.binder_annot
val cons : t -> Id.t list -> Id.t list
val to_option : Name.t -> Id.t option
@@ -176,6 +177,11 @@ struct
| Name _ -> na1
| Anonymous -> na2
+ let pick_annot na1 na2 =
+ let open Context in
+ match na1.binder_name with
+ | Name _ -> na1 | Anonymous -> na2
+
let cons na l =
match na with
| Anonymous -> l
diff --git a/engine/nameops.mli b/engine/nameops.mli
index a5308904f5..0e75fed045 100644
--- a/engine/nameops.mli
+++ b/engine/nameops.mli
@@ -84,6 +84,9 @@ module Name : sig
(** [pick na na'] returns [Anonymous] if both names are [Anonymous].
Pick one of [na] or [na'] otherwise. *)
+ val pick_annot : Name.t Context.binder_annot -> Name.t Context.binder_annot ->
+ Name.t Context.binder_annot
+
val cons : Name.t -> Id.t list -> Id.t list
(** [cons na l] returns [id::l] if [na] is [Name id] and [l] otherwise. *)
diff --git a/engine/proofview.ml b/engine/proofview.ml
index a725444e81..2d693e0259 100644
--- a/engine/proofview.ml
+++ b/engine/proofview.ml
@@ -876,9 +876,9 @@ module Progress = struct
let eq_named_declaration d1 d2 =
match d1, d2 with
| LocalAssum (i1,t1), LocalAssum (i2,t2) ->
- Names.Id.equal i1 i2 && eq_constr sigma1 sigma2 t1 t2
+ Context.eq_annot Names.Id.equal i1 i2 && eq_constr sigma1 sigma2 t1 t2
| LocalDef (i1,c1,t1), LocalDef (i2,c2,t2) ->
- Names.Id.equal i1 i2 && eq_constr sigma1 sigma2 c1 c2
+ Context.eq_annot Names.Id.equal i1 i2 && eq_constr sigma1 sigma2 c1 c2
&& eq_constr sigma1 sigma2 t1 t2
| _ ->
false
diff --git a/engine/termops.ml b/engine/termops.ml
index 9c4b64b60d..8e12c9be88 100644
--- a/engine/termops.ml
+++ b/engine/termops.ml
@@ -15,6 +15,7 @@ open Names
open Nameops
open Term
open Constr
+open Context
open Vars
open Environ
@@ -115,8 +116,8 @@ let pr_decl env sigma (decl,ok) =
let open NamedDecl in
let print_constr = print_kconstr in
match decl with
- | LocalAssum (id,_) -> if ok then Id.print id else (str "{" ++ Id.print id ++ str "}")
- | LocalDef (id,c,_) -> str (if ok then "(" else "{") ++ Id.print id ++ str ":=" ++
+ | LocalAssum ({binder_name=id},_) -> if ok then Id.print id else (str "{" ++ Id.print id ++ str "}")
+ | LocalDef ({binder_name=id},c,_) -> str (if ok then "(" else "{") ++ Id.print id ++ str ":=" ++
print_constr env sigma c ++ str (if ok then ")" else "}")
let pr_evar_source env sigma = function
@@ -248,8 +249,8 @@ let pr_evar_universe_context ctx =
let print_env_short env sigma =
let print_constr = print_kconstr in
let pr_rel_decl = function
- | RelDecl.LocalAssum (n,_) -> Name.print n
- | RelDecl.LocalDef (n,b,_) -> str "(" ++ Name.print n ++ str " := "
+ | RelDecl.LocalAssum (n,_) -> Name.print n.binder_name
+ | RelDecl.LocalDef (n,b,_) -> str "(" ++ Name.print n.binder_name ++ str " := "
++ print_constr env sigma (EConstr.of_constr b) ++ str ")"
in
let pr_named_decl = NamedDecl.to_rel_decl %> pr_rel_decl in
@@ -459,9 +460,10 @@ let push_named_rec_types (lna,typarray,_) env =
let ctxt =
Array.map2_i
(fun i na t ->
- match na with
- | Name id -> LocalAssum (id, lift i t)
- | Anonymous -> anomaly (Pp.str "Fix declarations must be named."))
+ let id = map_annot (function
+ | Name id -> id
+ | Anonymous -> anomaly (Pp.str "Fix declarations must be named.")) na
+ in LocalAssum (id, lift i t))
lna typarray in
Array.fold_left
(fun e assum -> push_named assum e) env ctxt
@@ -469,14 +471,11 @@ let push_named_rec_types (lna,typarray,_) env =
let lookup_rel_id id sign =
let open RelDecl in
let rec lookrec n = function
- | [] ->
- raise Not_found
- | (LocalAssum (Anonymous, _) | LocalDef (Anonymous,_,_)) :: l ->
- lookrec (n + 1) l
- | LocalAssum (Name id', t) :: l ->
- if Names.Id.equal id' id then (n,None,t) else lookrec (n + 1) l
- | LocalDef (Name id', b, t) :: l ->
- if Names.Id.equal id' id then (n,Some b,t) else lookrec (n + 1) l
+ | [] -> raise Not_found
+ | decl :: l ->
+ if Names.Name.equal (Name id) (get_name decl)
+ then (n, get_value decl, get_type decl)
+ else lookrec (n+1) l
in
lookrec 1 sign
@@ -1353,7 +1352,7 @@ let compact_named_context sign =
let clear_named_body id env =
let open NamedDecl in
let aux _ = function
- | LocalDef (id',c,t) when Id.equal id id' -> push_named (LocalAssum (id,t))
+ | LocalDef (id',c,t) when Id.equal id id'.binder_name -> push_named (LocalAssum (id',t))
| d -> push_named d in
fold_named_context aux env ~init:(reset_context env)
diff --git a/engine/termops.mli b/engine/termops.mli
index dea59e9efc..1dd9941c5e 100644
--- a/engine/termops.mli
+++ b/engine/termops.mli
@@ -23,9 +23,9 @@ val pr_fix : ('a -> Pp.t) -> ('a, 'a) pfixpoint -> Pp.t
[@@ocaml.deprecated "Use [Constr.debug_print_fix]"]
(** about contexts *)
-val push_rel_assum : Name.t * types -> env -> env
-val push_rels_assum : (Name.t * Constr.types) list -> env -> env
-val push_named_rec_types : Name.t array * Constr.types array * 'a -> env -> env
+val push_rel_assum : Name.t Context.binder_annot * types -> env -> env
+val push_rels_assum : (Name.t Context.binder_annot * Constr.types) list -> env -> env
+val push_named_rec_types : Name.t Context.binder_annot array * Constr.types array * 'a -> env -> env
val lookup_rel_id : Id.t -> ('c, 't) Context.Rel.pt -> int * 'c option * 't
(** Associates the contents of an identifier in a [rel_context]. Raise
@@ -40,8 +40,8 @@ val rel_list : int -> int -> constr list
(** iterators/destructors on terms *)
val mkProd_or_LetIn : rel_declaration -> types -> types
val mkProd_wo_LetIn : rel_declaration -> types -> types
-val it_mkProd : types -> (Name.t * types) list -> types
-val it_mkLambda : constr -> (Name.t * types) list -> constr
+val it_mkProd : types -> (Name.t Context.binder_annot * types) list -> types
+val it_mkLambda : constr -> (Name.t Context.binder_annot * types) list -> constr
val it_mkProd_or_LetIn : types -> rel_context -> types
val it_mkProd_wo_LetIn : types -> rel_context -> types
val it_mkLambda_or_LetIn : Constr.constr -> Constr.rel_context -> Constr.constr
@@ -246,7 +246,7 @@ val add_vname : Id.Set.t -> Name.t -> Id.Set.t
(** other signature iterators *)
val process_rel_context : (rel_declaration -> env -> env) -> env -> env
-val assums_of_rel_context : ('c, 't) Context.Rel.pt -> (Name.t * 't) list
+val assums_of_rel_context : ('c, 't) Context.Rel.pt -> (Name.t Context.binder_annot * 't) list
val lift_rel_context : int -> Constr.rel_context -> Constr.rel_context
val substl_rel_context : Constr.constr list -> Constr.rel_context -> Constr.rel_context
val smash_rel_context : Constr.rel_context -> Constr.rel_context (** expand lets in context *)
diff --git a/interp/constrintern.ml b/interp/constrintern.ml
index b4c570d444..5ede9d6a99 100644
--- a/interp/constrintern.ml
+++ b/interp/constrintern.ml
@@ -16,6 +16,7 @@ open Names
open Nameops
open Namegen
open Constr
+open Context
open Libnames
open Globnames
open Impargs
@@ -2183,7 +2184,7 @@ let internalize globalenv env pattern_mode (_, ntnvars as lvar) c =
(add_name match_acc CAst.(make ?loc x)) (CAst.make ?loc x::var_acc)
| _ ->
let fresh =
- Namegen.next_name_away_with_default_using_types "iV" cano_name forbidden_names (EConstr.of_constr ty) in
+ Namegen.next_name_away_with_default_using_types "iV" cano_name.binder_name forbidden_names (EConstr.of_constr ty) in
canonize_args t tt (Id.Set.add fresh forbidden_names)
((fresh,c)::match_acc) ((CAst.make ?loc:(cases_pattern_loc c) @@ Name fresh)::var_acc)
end
@@ -2432,9 +2433,10 @@ let interp_glob_context_evars ?(program_mode=false) env sigma k bl =
in
let sigma, t = understand_tcc ~flags env sigma ~expected_type:IsType t' in
match b with
- None ->
- let d = LocalAssum (na,t) in
- let impls =
+ None ->
+ let r = Retyping.relevance_of_type env sigma t in
+ let d = LocalAssum (make_annot na r,t) in
+ let impls =
if k == Implicit then
let na = match na with Name n -> Some n | Anonymous -> None in
(ExplByPos (n, na), (true, true, true)) :: impls
@@ -2443,7 +2445,8 @@ let interp_glob_context_evars ?(program_mode=false) env sigma k bl =
(push_rel d env, sigma, d::params, succ n, impls)
| Some b ->
let sigma, c = understand_tcc ~flags env sigma ~expected_type:(OfType t) b in
- let d = LocalDef (na, c, t) in
+ let r = Retyping.relevance_of_type env sigma t in
+ let d = LocalDef (make_annot na r, c, t) in
(push_rel d env, sigma, d::params, n, impls))
(env,sigma,[],k+1,[]) (List.rev bl)
in sigma, ((env, par), impls)
diff --git a/interp/declare.ml b/interp/declare.ml
index 4371b15c82..08a6ac5f7b 100644
--- a/interp/declare.ml
+++ b/interp/declare.ml
@@ -370,7 +370,7 @@ let declare_projections univs mind =
let mib = Environ.lookup_mind mind env in
match mib.mind_record with
| PrimRecord info ->
- let iter_ind i (_, labs, _) =
+ let iter_ind i (_, labs, _, _) =
let ind = (mind, i) in
let projs = Inductiveops.compute_projections env ind in
Array.iter2_i (declare_one_projection univs ind ~proj_npars:mib.mind_nparams) labs projs
diff --git a/interp/discharge.ml b/interp/discharge.ml
index 524bf9a002..1efd13adb1 100644
--- a/interp/discharge.ml
+++ b/interp/discharge.ml
@@ -103,7 +103,7 @@ let process_inductive info modlist mib =
let (params',inds') = abstract_inductive section_decls' nparamdecls inds in
let record = match mib.mind_record with
| PrimRecord info ->
- Some (Some (Array.map pi1 info))
+ Some (Some (Array.map (fun (x,_,_,_) -> x) info))
| FakeRecord -> Some None
| NotRecord -> None
in
diff --git a/interp/impargs.ml b/interp/impargs.ml
index 0f9bff7f1d..d83a0ce918 100644
--- a/interp/impargs.ml
+++ b/interp/impargs.ml
@@ -243,7 +243,7 @@ let compute_implicits_names_gen all env sigma t =
let t = whd_all env sigma t in
match kind sigma t with
| Prod (na,a,b) ->
- let na',avoid' = find_displayed_name_in sigma all avoid na (names,b) in
+ let na',avoid' = find_displayed_name_in sigma all avoid na.Context.binder_name (names,b) in
aux (push_rel (LocalAssum (na,a)) env) avoid' (na'::names) b
| _ -> List.rev names
in aux env Id.Set.empty [] t
@@ -445,7 +445,8 @@ let compute_mib_implicits flags kn =
(fun i mip ->
(* No need to care about constraints here *)
let ty, _ = Typeops.type_of_global_in_context env (IndRef (kn,i)) in
- Context.Rel.Declaration.LocalAssum (Name mip.mind_typename, ty))
+ let r = Inductive.relevance_of_inductive env (kn,i) in
+ Context.Rel.Declaration.LocalAssum (Context.make_annot (Name mip.mind_typename) r, ty))
mib.mind_packets) in
let env_ar = Environ.push_rel_context ar env in
let imps_one_inductive i mip =
diff --git a/interp/implicit_quantifiers.ml b/interp/implicit_quantifiers.ml
index 4f3037b1fc..854651e7b7 100644
--- a/interp/implicit_quantifiers.ml
+++ b/interp/implicit_quantifiers.ml
@@ -10,6 +10,7 @@
(*i*)
open Names
+open Context
open Decl_kinds
open CErrors
open Util
@@ -175,10 +176,10 @@ let combine_params avoid fn applied needed =
match app, need with
[], [] -> List.rev ids, avoid
- | app, (_, (LocalAssum (Name id, _) | LocalDef (Name id, _, _))) :: need when Id.List.mem_assoc id named ->
+ | app, (_, (LocalAssum ({binder_name=Name id}, _) | LocalDef ({binder_name=Name id}, _, _))) :: need when Id.List.mem_assoc id named ->
aux (Id.List.assoc id named :: ids) avoid app need
- | (x, None) :: app, (None, (LocalAssum (Name id, _) | LocalDef (Name id, _, _))) :: need ->
+ | (x, None) :: app, (None, (LocalAssum ({binder_name=Name id}, _) | LocalDef ({binder_name=Name id}, _, _))) :: need ->
aux (x :: ids) avoid app need
| _, (Some cl, _ as d) :: need ->
diff --git a/kernel/cClosure.ml b/kernel/cClosure.ml
index 5fec55fea1..a29f3c6833 100644
--- a/kernel/cClosure.ml
+++ b/kernel/cClosure.ml
@@ -98,7 +98,7 @@ module type RedFlagsSig = sig
val red_projection : reds -> Projection.t -> bool
end
-module RedFlags = (struct
+module RedFlags : RedFlagsSig = struct
(* [r_const=(true,cl)] means all constants but those in [cl] *)
(* [r_const=(false,cl)] means only those in [cl] *)
@@ -195,7 +195,7 @@ module RedFlags = (struct
if Projection.unfolded p then true
else red_set red (fCONST (Projection.constant p))
-end : RedFlagsSig)
+end
open RedFlags
@@ -300,9 +300,9 @@ and fterm =
| FFix of fixpoint * fconstr subs
| FCoFix of cofixpoint * fconstr subs
| FCaseT of case_info * constr * fconstr * constr array * fconstr subs (* predicate and branches are closures *)
- | FLambda of int * (Name.t * constr) list * constr * fconstr subs
- | FProd of Name.t * fconstr * constr * fconstr subs
- | FLetIn of Name.t * fconstr * fconstr * constr * fconstr subs
+ | FLambda of int * (Name.t Context.binder_annot * constr) list * constr * fconstr subs
+ | FProd of Name.t Context.binder_annot * fconstr * constr * fconstr subs
+ | FLetIn of Name.t Context.binder_annot * fconstr * fconstr * constr * fconstr subs
| FEvar of existential * fconstr subs
| FInt of Uint63.t
| FLIFT of int * fconstr
@@ -865,7 +865,7 @@ let rec knh info m stk =
| FLOCKED -> assert false
| FApp(a,b) -> knh info a (append_stack b (zupdate info m stk))
| FCaseT(ci,p,t,br,e) -> knh info t (ZcaseT(ci,p,br,e)::zupdate info m stk)
- | FFix(((ri,n),(_,_,_)),_) ->
+ | FFix(((ri,n),_),_) ->
(match get_nth_arg m ri.(n) stk with
(Some(pars,arg),stk') -> knh info arg (Zfix(m,pars)::stk')
| (None, stk') -> (m,stk'))
diff --git a/kernel/cClosure.mli b/kernel/cClosure.mli
index bd04677374..2852d48a5d 100644
--- a/kernel/cClosure.mli
+++ b/kernel/cClosure.mli
@@ -114,9 +114,9 @@ type fterm =
| FFix of fixpoint * fconstr subs
| FCoFix of cofixpoint * fconstr subs
| FCaseT of case_info * constr * fconstr * constr array * fconstr subs (* predicate and branches are closures *)
- | FLambda of int * (Name.t * constr) list * constr * fconstr subs
- | FProd of Name.t * fconstr * constr * fconstr subs
- | FLetIn of Name.t * fconstr * fconstr * constr * fconstr subs
+ | FLambda of int * (Name.t Context.binder_annot * constr) list * constr * fconstr subs
+ | FProd of Name.t Context.binder_annot * fconstr * constr * fconstr subs
+ | FLetIn of Name.t Context.binder_annot * fconstr * fconstr * constr * fconstr subs
| FEvar of existential * fconstr subs
| FInt of Uint63.t
| FLIFT of int * fconstr
@@ -165,7 +165,7 @@ val mk_red : fterm -> fconstr
val fterm_of : fconstr -> fterm
val term_of_fconstr : fconstr -> constr
val destFLambda :
- (fconstr subs -> constr -> fconstr) -> fconstr -> Name.t * fconstr * fconstr
+ (fconstr subs -> constr -> fconstr) -> fconstr -> Name.t Context.binder_annot * fconstr * fconstr
(** Global and local constant cache *)
type clos_infos
diff --git a/kernel/clambda.ml b/kernel/clambda.ml
index d2147884ba..a764cca354 100644
--- a/kernel/clambda.ml
+++ b/kernel/clambda.ml
@@ -15,8 +15,8 @@ type lambda =
| Lvar of Id.t
| Levar of Evar.t * lambda array
| Lprod of lambda * lambda
- | Llam of Name.t array * lambda
- | Llet of Name.t * lambda * lambda
+ | Llam of Name.t Context.binder_annot array * lambda
+ | Llet of Name.t Context.binder_annot * lambda * lambda
| Lapp of lambda * lambda array
| Lconst of pconstant
| Lprim of pconstant option * CPrimitives.t * lambda array
@@ -38,15 +38,17 @@ type lambda =
stored in [extra_branches]. *)
and lam_branches =
{ constant_branches : lambda array;
- nonconstant_branches : (Name.t array * lambda) array }
+ nonconstant_branches : (Name.t Context.binder_annot array * lambda) array }
(* extra_branches : (name array * lambda) array } *)
-and fix_decl = Name.t array * lambda array * lambda array
+and fix_decl = Name.t Context.binder_annot array * lambda array * lambda array
(** Printing **)
+let pr_annot x = Name.print x.Context.binder_name
+
let pp_names ids =
- prlist_with_sep (fun _ -> brk(1,1)) Name.print (Array.to_list ids)
+ prlist_with_sep (fun _ -> brk(1,1)) pr_annot (Array.to_list ids)
let pp_rel name n =
Name.print name ++ str "##" ++ int n
@@ -80,7 +82,7 @@ let rec pp_lam lam =
str ")")
| Llet(id,def,body) -> hov 0
(str "let " ++
- Name.print id ++
+ pr_annot id ++
str ":=" ++
pp_lam def ++
str " in" ++
@@ -120,7 +122,7 @@ let rec pp_lam lam =
v 0
(prlist_with_sep spc
(fun (na,i,ty,bd) ->
- Name.print na ++ str"/" ++ int i ++ str":" ++
+ pr_annot na ++ str"/" ++ int i ++ str":" ++
pp_lam ty ++ cut() ++ str":=" ++
pp_lam bd) (Array.to_list fixl)) ++
str"}")
@@ -132,7 +134,7 @@ let rec pp_lam lam =
v 0
(prlist_with_sep spc
(fun (na,ty,bd) ->
- Name.print na ++ str":" ++ pp_lam ty ++
+ pr_annot na ++ str":" ++ pp_lam ty ++
cut() ++ str":=" ++ pp_lam bd) (Array.to_list fixl)) ++
str"}")
| Lmakeblock(tag, args) ->
@@ -394,8 +396,8 @@ and reduce_lapp substf lids body substa largs =
Llet(id, a, body)
| [], [] -> simplify substf body
| _::_, _ ->
- Llam(Array.of_list lids, simplify (liftn (List.length lids) substf) body)
- | [], _::_ -> simplify_app substf body substa (Array.of_list largs)
+ Llam(Array.of_list lids, simplify (liftn (List.length lids) substf) body)
+ | [], _ -> simplify_app substf body substa (Array.of_list largs)
@@ -512,7 +514,8 @@ let make_args start _end =
(* Translation of constructors *)
let expand_constructor tag nparams arity =
- let ids = Array.make (nparams + arity) Anonymous in
+ let anon = Context.make_annot Anonymous Sorts.Relevant in (* TODO relevance *)
+ let ids = Array.make (nparams + arity) anon in
if arity = 0 then mkLlam ids (Lint tag)
else
let args = make_args arity 1 in
@@ -554,7 +557,8 @@ let prim kn p args =
Lprim(Some kn, p, args)
let expand_prim kn op arity =
- let ids = Array.make arity Anonymous in
+ (* primitives are always Relevant *)
+ let ids = Array.make arity Context.anonR in
let args = make_args arity 1 in
Llam(ids, prim kn op args)
@@ -629,7 +633,7 @@ struct
construct_tbl = Hashtbl.create 111
}
- let push_rel env id = Vect.push env.name_rel id
+ let push_rel env id = Vect.push env.name_rel id.Context.binder_name
let push_rels env ids =
Array.iter (push_rel env) ids
@@ -679,7 +683,7 @@ let rec lambda_of_constr env c =
Renv.push_rel env id;
let lc = lambda_of_constr env codom in
Renv.pop env;
- Lprod(ld, Llam([|id|], lc))
+ Lprod(ld, Llam([|id|], lc))
| Lambda _ ->
let params, body = decompose_lam c in
@@ -726,7 +730,8 @@ let rec lambda_of_constr env c =
match b with
| Llam(ids, body) when Array.length ids = arity -> (ids, body)
| _ ->
- let ids = Array.make arity Anonymous in
+ let anon = Context.make_annot Anonymous Sorts.Relevant in (* TODO relevance *)
+ let ids = Array.make arity anon in
let args = make_args arity 1 in
let ll = lam_lift arity b in
(ids, mkLapp ll args)
@@ -801,7 +806,7 @@ let optimize_lambda lam =
let lambda_of_constr ~optimize genv c =
let env = Renv.make genv in
- let ids = List.rev_map Context.Rel.Declaration.get_name (rel_context genv) in
+ let ids = List.rev_map Context.Rel.Declaration.get_annot (rel_context genv) in
Renv.push_rels env (Array.of_list ids);
let lam = lambda_of_constr env c in
let lam = if optimize then optimize_lambda lam else lam in
diff --git a/kernel/clambda.mli b/kernel/clambda.mli
index 4d921fd45e..1476bb6e45 100644
--- a/kernel/clambda.mli
+++ b/kernel/clambda.mli
@@ -8,8 +8,8 @@ type lambda =
| Lvar of Id.t
| Levar of Evar.t * lambda array
| Lprod of lambda * lambda
- | Llam of Name.t array * lambda
- | Llet of Name.t * lambda * lambda
+ | Llam of Name.t Context.binder_annot array * lambda
+ | Llet of Name.t Context.binder_annot * lambda * lambda
| Lapp of lambda * lambda array
| Lconst of pconstant
| Lprim of pconstant option * CPrimitives.t * lambda array
@@ -28,15 +28,15 @@ type lambda =
and lam_branches =
{ constant_branches : lambda array;
- nonconstant_branches : (Name.t array * lambda) array }
+ nonconstant_branches : (Name.t Context.binder_annot array * lambda) array }
-and fix_decl = Name.t array * lambda array * lambda array
+and fix_decl = Name.t Context.binder_annot array * lambda array * lambda array
exception TooLargeInductive of Pp.t
val lambda_of_constr : optimize:bool -> env -> Constr.t -> lambda
-val decompose_Llam : lambda -> Name.t array * lambda
+val decompose_Llam : lambda -> Name.t Context.binder_annot array * lambda
val get_alias : env -> Constant.t -> Constant.t
diff --git a/kernel/constr.ml b/kernel/constr.ml
index 57ece320cf..11958c9108 100644
--- a/kernel/constr.ml
+++ b/kernel/constr.ml
@@ -28,6 +28,7 @@
open Util
open Names
open Univ
+open Context
type existential_key = Evar.t
type metavariable = int
@@ -60,6 +61,7 @@ type case_info =
in addition to the parameters of the related inductive type
NOTE: "lets" are therefore excluded from the count
NOTE: parameters of the inductive type are also excluded from the count *)
+ ci_relevance : Sorts.relevance;
ci_pp_info : case_printing (* not interpreted by the kernel *)
}
@@ -71,7 +73,7 @@ type case_info =
the same order (i.e. last argument first) *)
type 'constr pexistential = existential_key * 'constr array
type ('constr, 'types) prec_declaration =
- Name.t array * 'types array * 'constr array
+ Name.t binder_annot array * 'types array * 'constr array
type ('constr, 'types) pfixpoint =
(int array * int) * ('constr, 'types) prec_declaration
type ('constr, 'types) pcofixpoint =
@@ -90,9 +92,9 @@ type ('constr, 'types, 'sort, 'univs) kind_of_term =
| Evar of 'constr pexistential
| Sort of 'sort
| Cast of 'constr * cast_kind * 'types
- | Prod of Name.t * 'types * 'types
- | Lambda of Name.t * 'types * 'constr
- | LetIn of Name.t * 'constr * 'types * 'constr
+ | Prod of Name.t binder_annot * 'types * 'types
+ | Lambda of Name.t binder_annot * 'types * 'constr
+ | LetIn of Name.t binder_annot * 'constr * 'types * 'constr
| App of 'constr * 'constr array
| Const of (Constant.t * 'univs)
| Ind of (inductive * 'univs)
@@ -127,13 +129,15 @@ let rels =
let mkRel n = if 0<n && n<=16 then rels.(n-1) else Rel n
(* Construct a type *)
+let mkSProp = Sort Sorts.sprop
let mkProp = Sort Sorts.prop
let mkSet = Sort Sorts.set
let mkType u = Sort (Sorts.sort_of_univ u)
let mkSort = function
+ | Sorts.SProp -> mkSProp
| Sorts.Prop -> mkProp (* Easy sharing *)
| Sorts.Set -> mkSet
- | s -> Sort s
+ | Sorts.Type _ as s -> Sort s
(* Constructs the term t1::t2, i.e. the term t1 casted with the type t2 *)
(* (that means t2 is declared as the type of t1) *)
@@ -1181,16 +1185,16 @@ let hashcons (sh_sort,sh_ci,sh_construct,sh_ind,sh_con,sh_na,sh_id) =
| Prod (na,t,c) ->
let t, ht = sh_rec t
and c, hc = sh_rec c in
- (Prod (sh_na na, t, c), combinesmall 4 (combine3 (Name.hash na) ht hc))
+ (Prod (sh_na na, t, c), combinesmall 4 (combine3 (hash_annot Name.hash na) ht hc))
| Lambda (na,t,c) ->
let t, ht = sh_rec t
and c, hc = sh_rec c in
- (Lambda (sh_na na, t, c), combinesmall 5 (combine3 (Name.hash na) ht hc))
+ (Lambda (sh_na na, t, c), combinesmall 5 (combine3 (hash_annot Name.hash na) ht hc))
| LetIn (na,b,t,c) ->
let b, hb = sh_rec b in
let t, ht = sh_rec t in
let c, hc = sh_rec c in
- (LetIn (sh_na na, b, t, c), combinesmall 6 (combine4 (Name.hash na) hb ht hc))
+ (LetIn (sh_na na, b, t, c), combinesmall 6 (combine4 (hash_annot Name.hash na) hb ht hc))
| App (c,l) ->
let c, hc = sh_rec c in
let l, hl = hash_term_array l in
@@ -1214,24 +1218,24 @@ let hashcons (sh_sort,sh_ci,sh_construct,sh_ind,sh_con,sh_na,sh_id) =
let p, hp = sh_rec p
and c, hc = sh_rec c in
let bl,hbl = hash_term_array bl in
- let hbl = combine (combine hc hp) hbl in
+ let hbl = combine (combine hc hp) hbl in
(Case (sh_ci ci, p, c, bl), combinesmall 12 hbl)
| Fix (ln,(lna,tl,bl)) ->
- let bl,hbl = hash_term_array bl in
+ let bl,hbl = hash_term_array bl in
let tl,htl = hash_term_array tl in
let () = Array.iteri (fun i x -> Array.unsafe_set lna i (sh_na x)) lna in
- let fold accu na = combine (Name.hash na) accu in
+ let fold accu na = combine (hash_annot Name.hash na) accu in
let hna = Array.fold_left fold 0 lna in
let h = combine3 hna hbl htl in
- (Fix (ln,(lna,tl,bl)), combinesmall 13 h)
+ (Fix (ln,(lna,tl,bl)), combinesmall 13 h)
| CoFix(ln,(lna,tl,bl)) ->
- let bl,hbl = hash_term_array bl in
+ let bl,hbl = hash_term_array bl in
let tl,htl = hash_term_array tl in
let () = Array.iteri (fun i x -> Array.unsafe_set lna i (sh_na x)) lna in
- let fold accu na = combine (Name.hash na) accu in
+ let fold accu na = combine (hash_annot Name.hash na) accu in
let hna = Array.fold_left fold 0 lna in
let h = combine3 hna hbl htl in
- (CoFix (ln,(lna,tl,bl)), combinesmall 14 h)
+ (CoFix (ln,(lna,tl,bl)), combinesmall 14 h)
| Meta n ->
(t, combinesmall 15 n)
| Rel n ->
@@ -1322,6 +1326,7 @@ struct
info1.style == info2.style
let eq ci ci' =
ci.ci_ind == ci'.ci_ind &&
+ ci.ci_relevance == ci'.ci_relevance &&
Int.equal ci.ci_npar ci'.ci_npar &&
Array.equal Int.equal ci.ci_cstr_ndecls ci'.ci_cstr_ndecls && (* we use [Array.equal] on purpose *)
Array.equal Int.equal ci.ci_cstr_nargs ci'.ci_cstr_nargs && (* we use [Array.equal] on purpose *)
@@ -1345,7 +1350,7 @@ struct
let h3 = Array.fold_left combine 0 ci.ci_cstr_ndecls in
let h4 = Array.fold_left combine 0 ci.ci_cstr_nargs in
let h5 = hash_pp_info ci.ci_pp_info in
- combine5 h1 h2 h3 h4 h5
+ combinesmall (Sorts.relevance_hash ci.ci_relevance) (combine5 h1 h2 h3 h4 h5)
end
module Hcaseinfo = Hashcons.Make(CaseinfoHash)
@@ -1354,6 +1359,18 @@ let case_info_hash = CaseinfoHash.hash
let hcons_caseinfo = Hashcons.simple_hcons Hcaseinfo.generate Hcaseinfo.hcons hcons_ind
+module Hannotinfo = struct
+ type t = Name.t binder_annot
+ type u = Name.t -> Name.t
+ let hash = hash_annot Name.hash
+ let eq = eq_annot (fun na1 na2 -> na1 == na2)
+ let hashcons h {binder_name=na;binder_relevance} =
+ {binder_name=h na;binder_relevance}
+ end
+module Hannot = Hashcons.Make(Hannotinfo)
+
+let hcons_annot = Hashcons.simple_hcons Hannot.generate Hannot.hcons Name.hcons
+
let hcons =
hashcons
(Sorts.hcons,
@@ -1361,7 +1378,7 @@ let hcons =
hcons_construct,
hcons_ind,
hcons_con,
- Name.hcons,
+ hcons_annot,
Id.hcons)
(* let hcons_types = hcons_constr *)
@@ -1377,7 +1394,7 @@ type compacted_context = compacted_declaration list
let debug_print_fix pr_constr ((t,i),(lna,tl,bl)) =
let open Pp in
- let fixl = Array.mapi (fun i na -> (na,t.(i),tl.(i),bl.(i))) lna in
+ let fixl = Array.mapi (fun i na -> (na.binder_name,t.(i),tl.(i),bl.(i))) lna in
hov 1
(str"fix " ++ int i ++ spc() ++ str"{" ++
v 0 (prlist_with_sep spc (fun (na,i,ty,bd) ->
@@ -1399,17 +1416,17 @@ let rec debug_print c =
| Cast (c,_, t) -> hov 1
(str"(" ++ debug_print c ++ cut() ++
str":" ++ debug_print t ++ str")")
- | Prod (Name(id),t,c) -> hov 1
+ | Prod ({binder_name=Name id;_},t,c) -> hov 1
(str"forall " ++ Id.print id ++ str":" ++ debug_print t ++ str"," ++
spc() ++ debug_print c)
- | Prod (Anonymous,t,c) -> hov 0
+ | Prod ({binder_name=Anonymous;_},t,c) -> hov 0
(str"(" ++ debug_print t ++ str " ->" ++ spc() ++
debug_print c ++ str")")
| Lambda (na,t,c) -> hov 1
- (str"fun " ++ Name.print na ++ str":" ++
+ (str"fun " ++ Name.print na.binder_name ++ str":" ++
debug_print t ++ str" =>" ++ spc() ++ debug_print c)
| LetIn (na,b,t,c) -> hov 0
- (str"let " ++ Name.print na ++ str":=" ++ debug_print b ++
+ (str"let " ++ Name.print na.binder_name ++ str":=" ++ debug_print b ++
str":" ++ brk(1,2) ++ debug_print t ++ cut() ++
debug_print c)
| App (c,l) -> hov 1
@@ -1434,7 +1451,7 @@ let rec debug_print c =
hov 1
(str"cofix " ++ int i ++ spc() ++ str"{" ++
v 0 (prlist_with_sep spc (fun (na,ty,bd) ->
- Name.print na ++ str":" ++ debug_print ty ++
+ Name.print na.binder_name ++ str":" ++ debug_print ty ++
cut() ++ str":=" ++ debug_print bd) (Array.to_list fixl)) ++
str"}")
| Int i -> str"Int("++str (Uint63.to_string i) ++ str")"
diff --git a/kernel/constr.mli b/kernel/constr.mli
index fdc3296a6a..7fc57cdb8a 100644
--- a/kernel/constr.mli
+++ b/kernel/constr.mli
@@ -45,6 +45,7 @@ type case_info =
in addition to the parameters of the related inductive type
NOTE: "lets" are therefore excluded from the count
NOTE: parameters of the inductive type are also excluded from the count *)
+ ci_relevance : Sorts.relevance; (* relevance of the predicate (not of the inductive!) *)
ci_pp_info : case_printing (* not interpreted by the kernel *)
}
@@ -84,6 +85,7 @@ val mkEvar : existential -> constr
(** Construct a sort *)
val mkSort : Sorts.t -> types
+val mkSProp : types
val mkProp : types
val mkSet : types
val mkType : Univ.Universe.t -> types
@@ -97,13 +99,13 @@ type cast_kind = VMcast | NATIVEcast | DEFAULTcast | REVERTcast
val mkCast : constr * cast_kind * constr -> constr
(** Constructs the product [(x:t1)t2] *)
-val mkProd : Name.t * types * types -> types
+val mkProd : Name.t Context.binder_annot * types * types -> types
(** Constructs the abstraction \[x:t{_ 1}\]t{_ 2} *)
-val mkLambda : Name.t * types * constr -> constr
+val mkLambda : Name.t Context.binder_annot * types * constr -> constr
(** Constructs the product [let x = t1 : t2 in t3] *)
-val mkLetIn : Name.t * constr * types * constr -> constr
+val mkLetIn : Name.t Context.binder_annot * constr * types * constr -> constr
(** [mkApp (f, [|t1; ...; tN|]] constructs the application
{%html:(f t<sub>1</sub> ... t<sub>n</sub>)%}
@@ -160,7 +162,7 @@ val mkCase : case_info * constr * constr * constr array -> constr
where the length of the {% $ %}j{% $ %}th context is {% $ %}ij{% $ %}.
*)
type ('constr, 'types) prec_declaration =
- Name.t array * 'types array * 'constr array
+ Name.t Context.binder_annot array * 'types array * 'constr array
type ('constr, 'types) pfixpoint =
(int array * int) * ('constr, 'types) prec_declaration
(* The array of [int]'s tells for each component of the array of
@@ -213,9 +215,9 @@ type ('constr, 'types, 'sort, 'univs) kind_of_term =
| Evar of 'constr pexistential
| Sort of 'sort
| Cast of 'constr * cast_kind * 'types
- | Prod of Name.t * 'types * 'types (** Concrete syntax ["forall A:B,C"] is represented as [Prod (A,B,C)]. *)
- | Lambda of Name.t * 'types * 'constr (** Concrete syntax ["fun A:B => C"] is represented as [Lambda (A,B,C)]. *)
- | LetIn of Name.t * 'constr * 'types * 'constr (** Concrete syntax ["let A:C := B in D"] is represented as [LetIn (A,B,C,D)]. *)
+ | Prod of Name.t Context.binder_annot * 'types * 'types (** Concrete syntax ["forall A:B,C"] is represented as [Prod (A,B,C)]. *)
+ | Lambda of Name.t Context.binder_annot * 'types * 'constr (** Concrete syntax ["fun A:B => C"] is represented as [Lambda (A,B,C)]. *)
+ | LetIn of Name.t Context.binder_annot * 'constr * 'types * 'constr (** Concrete syntax ["let A:C := B in D"] is represented as [LetIn (A,B,C,D)]. *)
| App of 'constr * 'constr array (** Concrete syntax ["(F P1 P2 ... Pn)"] is represented as [App (F, [|P1; P2; ...; Pn|])].
The {!mkApp} constructor also enforces the following invariant:
@@ -297,13 +299,13 @@ val destSort : constr -> Sorts.t
val destCast : constr -> constr * cast_kind * constr
(** Destructs the product {% $ %}(x:t_1)t_2{% $ %} *)
-val destProd : types -> Name.t * types * types
+val destProd : types -> Name.t Context.binder_annot * types * types
(** Destructs the abstraction {% $ %}[x:t_1]t_2{% $ %} *)
-val destLambda : constr -> Name.t * types * constr
+val destLambda : constr -> Name.t Context.binder_annot * types * constr
(** Destructs the let {% $ %}[x:=b:t_1]t_2{% $ %} *)
-val destLetIn : constr -> Name.t * constr * types * constr
+val destLetIn : constr -> Name.t Context.binder_annot * constr * types * constr
(** Destructs an application *)
val destApp : constr -> constr * constr array
diff --git a/kernel/context.ml b/kernel/context.ml
index 1cc6e79485..290e85294b 100644
--- a/kernel/context.ml
+++ b/kernel/context.ml
@@ -31,6 +31,27 @@
open Util
open Names
+type 'a binder_annot = { binder_name : 'a; binder_relevance : Sorts.relevance }
+
+let eq_annot eq {binder_name=na1;binder_relevance=r1} {binder_name=na2;binder_relevance=r2} =
+ eq na1 na2 && Sorts.relevance_equal r1 r2
+
+let hash_annot h {binder_name=n;binder_relevance=r} =
+ Hashset.Combine.combinesmall (Sorts.relevance_hash r) (h n)
+
+let map_annot f {binder_name=na;binder_relevance} =
+ {binder_name=f na;binder_relevance}
+
+let make_annot x r = {binder_name=x;binder_relevance=r}
+
+let binder_name x = x.binder_name
+let binder_relevance x = x.binder_relevance
+
+let annotR x = make_annot x Sorts.Relevant
+
+let nameR x = annotR (Name x)
+let anonR = annotR Anonymous
+
(** Representation of contexts that can capture anonymous as well as non-anonymous variables.
Individual declarations are then designated by de Bruijn indexes. *)
module Rel =
@@ -40,13 +61,14 @@ struct
struct
(* local declaration *)
type ('constr, 'types) pt =
- | LocalAssum of Name.t * 'types (** name, type *)
- | LocalDef of Name.t * 'constr * 'types (** name, value, type *)
+ | LocalAssum of Name.t binder_annot * 'types (** name, type *)
+ | LocalDef of Name.t binder_annot * 'constr * 'types (** name, value, type *)
+
+ let get_annot = function
+ | LocalAssum (na,_) | LocalDef (na,_,_) -> na
(** Return the name bound by a given declaration. *)
- let get_name = function
- | LocalAssum (na,_)
- | LocalDef (na,_,_) -> na
+ let get_name x = (get_annot x).binder_name
(** Return [Some value] for local-declarations and [None] for local-assumptions. *)
let get_value = function
@@ -57,11 +79,13 @@ struct
let get_type = function
| LocalAssum (_,ty)
| LocalDef (_,_,ty) -> ty
-
+
+ let get_relevance x = (get_annot x).binder_relevance
+
(** Set the name that is bound by a given declaration. *)
let set_name na = function
- | LocalAssum (_,ty) -> LocalAssum (na, ty)
- | LocalDef (_,v,ty) -> LocalDef (na, v, ty)
+ | LocalAssum (x,ty) -> LocalAssum ({x with binder_name=na}, ty)
+ | LocalDef (x,v,ty) -> LocalDef ({x with binder_name=na}, v, ty)
(** Set the type of the bound variable in a given declaration. *)
let set_type ty = function
@@ -92,20 +116,17 @@ struct
let equal eq decl1 decl2 =
match decl1, decl2 with
| LocalAssum (n1,ty1), LocalAssum (n2, ty2) ->
- Name.equal n1 n2 && eq ty1 ty2
+ eq_annot Name.equal n1 n2 && eq ty1 ty2
| LocalDef (n1,v1,ty1), LocalDef (n2,v2,ty2) ->
- Name.equal n1 n2 && eq v1 v2 && eq ty1 ty2
+ eq_annot Name.equal n1 n2 && eq v1 v2 && eq ty1 ty2
| _ ->
false
(** Map the name bound by a given declaration. *)
- let map_name f = function
- | LocalAssum (na, ty) as decl ->
- let na' = f na in
- if na == na' then decl else LocalAssum (na', ty)
- | LocalDef (na, v, ty) as decl ->
- let na' = f na in
- if na == na' then decl else LocalDef (na', v, ty)
+ let map_name f x =
+ let na = get_name x in
+ let na' = f na in
+ if na == na' then x else set_name na' x
(** For local assumptions, this function returns the original local assumptions.
For local definitions, this function maps the value in the local definition. *)
@@ -120,7 +141,7 @@ struct
| LocalAssum (na, ty) as decl ->
let ty' = f ty in
if ty == ty' then decl else LocalAssum (na, ty')
- | LocalDef (na, v, ty) as decl ->
+ | LocalDef (na, v, ty) as decl ->
let ty' = f ty in
if ty == ty' then decl else LocalDef (na, v, ty')
@@ -250,13 +271,14 @@ struct
struct
(** local declaration *)
type ('constr, 'types) pt =
- | LocalAssum of Id.t * 'types (** identifier, type *)
- | LocalDef of Id.t * 'constr * 'types (** identifier, value, type *)
+ | LocalAssum of Id.t binder_annot * 'types (** identifier, type *)
+ | LocalDef of Id.t binder_annot * 'constr * 'types (** identifier, value, type *)
+
+ let get_annot = function
+ | LocalAssum (na,_) | LocalDef (na,_,_) -> na
(** Return the identifier bound by a given declaration. *)
- let get_id = function
- | LocalAssum (id,_) -> id
- | LocalDef (id,_,_) -> id
+ let get_id x = (get_annot x).binder_name
(** Return [Some value] for local-declarations and [None] for local-assumptions. *)
let get_value = function
@@ -268,10 +290,14 @@ struct
| LocalAssum (_,ty)
| LocalDef (_,_,ty) -> ty
+ let get_relevance x = (get_annot x).binder_relevance
+
(** Set the identifier that is bound by a given declaration. *)
- let set_id id = function
- | LocalAssum (_,ty) -> LocalAssum (id, ty)
- | LocalDef (_, v, ty) -> LocalDef (id, v, ty)
+ let set_id id =
+ let set x = {x with binder_name = id} in
+ function
+ | LocalAssum (x,ty) -> LocalAssum (set x, ty)
+ | LocalDef (x, v, ty) -> LocalDef (set x, v, ty)
(** Set the type of the bound variable in a given declaration. *)
let set_type ty = function
@@ -302,20 +328,17 @@ struct
let equal eq decl1 decl2 =
match decl1, decl2 with
| LocalAssum (id1, ty1), LocalAssum (id2, ty2) ->
- Id.equal id1 id2 && eq ty1 ty2
+ eq_annot Id.equal id1 id2 && eq ty1 ty2
| LocalDef (id1, v1, ty1), LocalDef (id2, v2, ty2) ->
- Id.equal id1 id2 && eq v1 v2 && eq ty1 ty2
+ eq_annot Id.equal id1 id2 && eq v1 v2 && eq ty1 ty2
| _ ->
false
(** Map the identifier bound by a given declaration. *)
- let map_id f = function
- | LocalAssum (id, ty) as decl ->
- let id' = f id in
- if id == id' then decl else LocalAssum (id', ty)
- | LocalDef (id, v, ty) as decl ->
- let id' = f id in
- if id == id' then decl else LocalDef (id', v, ty)
+ let map_id f x =
+ let id = get_id x in
+ let id' = f id in
+ if id == id' then x else set_id id' x
(** For local assumptions, this function returns the original local assumptions.
For local definitions, this function maps the value in the local definition. *)
@@ -369,15 +392,17 @@ struct
let of_rel_decl f = function
| Rel.Declaration.LocalAssum (na,t) ->
- LocalAssum (f na, t)
+ LocalAssum (map_annot f na, t)
| Rel.Declaration.LocalDef (na,v,t) ->
- LocalDef (f na, v, t)
-
- let to_rel_decl = function
+ LocalDef (map_annot f na, v, t)
+
+ let to_rel_decl =
+ let name x = {binder_name=Name x.binder_name;binder_relevance=x.binder_relevance} in
+ function
| LocalAssum (id,t) ->
- Rel.Declaration.LocalAssum (Name id, t)
+ Rel.Declaration.LocalAssum (name id, t)
| LocalDef (id,v,t) ->
- Rel.Declaration.LocalDef (Name id,v,t)
+ Rel.Declaration.LocalDef (name id,v,t)
end
(** Named-context is represented as a list of declarations.
@@ -430,7 +455,7 @@ struct
gives [Var id1, Var id3]. All [idj] are supposed distinct. *)
let to_instance mk l =
let filter = function
- | Declaration.LocalAssum (id, _) -> Some (mk id)
+ | Declaration.LocalAssum (id, _) -> Some (mk id.binder_name)
| _ -> None
in
List.map_filter filter l
@@ -441,8 +466,8 @@ module Compacted =
module Declaration =
struct
type ('constr, 'types) pt =
- | LocalAssum of Id.t list * 'types
- | LocalDef of Id.t list * 'constr * 'types
+ | LocalAssum of Id.t binder_annot list * 'types
+ | LocalDef of Id.t binder_annot list * 'constr * 'types
let map_constr f = function
| LocalAssum (ids, ty) as decl ->
diff --git a/kernel/context.mli b/kernel/context.mli
index 8acae73680..7b67e54ba4 100644
--- a/kernel/context.mli
+++ b/kernel/context.mli
@@ -24,6 +24,27 @@
open Names
+type 'a binder_annot = { binder_name : 'a; binder_relevance : Sorts.relevance }
+val eq_annot : ('a -> 'a -> bool) -> 'a binder_annot -> 'a binder_annot -> bool
+
+val hash_annot : ('a -> int) -> 'a binder_annot -> int
+
+val map_annot : ('a -> 'b) -> 'a binder_annot -> 'b binder_annot
+
+val make_annot : 'a -> Sorts.relevance -> 'a binder_annot
+
+val binder_name : 'a binder_annot -> 'a
+val binder_relevance : 'a binder_annot -> Sorts.relevance
+
+val annotR : 'a -> 'a binder_annot
+(** Always Relevant *)
+
+val nameR : Id.t -> Name.t binder_annot
+(** Relevant + Name *)
+
+val anonR : Name.t binder_annot
+(** Relevant + Anonymous *)
+
(** Representation of contexts that can capture anonymous as well as non-anonymous variables.
Individual declarations are then designated by de Bruijn indexes. *)
module Rel :
@@ -32,8 +53,10 @@ sig
sig
(* local declaration *)
type ('constr, 'types) pt =
- | LocalAssum of Name.t * 'types (** name, type *)
- | LocalDef of Name.t * 'constr * 'types (** name, value, type *)
+ | LocalAssum of Name.t binder_annot * 'types (** name, type *)
+ | LocalDef of Name.t binder_annot * 'constr * 'types (** name, value, type *)
+
+ val get_annot : _ pt -> Name.t binder_annot
(** Return the name bound by a given declaration. *)
val get_name : ('c, 't) pt -> Name.t
@@ -44,6 +67,8 @@ sig
(** Return the type of the name bound by a given declaration. *)
val get_type : ('c, 't) pt -> 't
+ val get_relevance : ('c, 't) pt -> Sorts.relevance
+
(** Set the name that is bound by a given declaration. *)
val set_name : Name.t -> ('c, 't) pt -> ('c, 't) pt
@@ -87,7 +112,7 @@ sig
(** Reduce all terms in a given declaration to a single value. *)
val fold_constr : ('c -> 'a -> 'a) -> ('c, 'c) pt -> 'a -> 'a
- val to_tuple : ('c, 't) pt -> Name.t * 'c option * 't
+ val to_tuple : ('c, 't) pt -> Name.t binder_annot * 'c option * 't
(** Turn [LocalDef] into [LocalAssum], identity otherwise. *)
val drop_body : ('c, 't) pt -> ('c, 't) pt
@@ -156,8 +181,10 @@ sig
module Declaration :
sig
type ('constr, 'types) pt =
- | LocalAssum of Id.t * 'types (** identifier, type *)
- | LocalDef of Id.t * 'constr * 'types (** identifier, value, type *)
+ | LocalAssum of Id.t binder_annot * 'types (** identifier, type *)
+ | LocalDef of Id.t binder_annot * 'constr * 'types (** identifier, value, type *)
+
+ val get_annot : _ pt -> Id.t binder_annot
(** Return the identifier bound by a given declaration. *)
val get_id : ('c, 't) pt -> Id.t
@@ -168,6 +195,8 @@ sig
(** Return the type of the name bound by a given declaration. *)
val get_type : ('c, 't) pt -> 't
+ val get_relevance : ('c, 't) pt -> Sorts.relevance
+
(** Set the identifier that is bound by a given declaration. *)
val set_id : Id.t -> ('c, 't) pt -> ('c, 't) pt
@@ -208,8 +237,8 @@ sig
(** Reduce all terms in a given declaration to a single value. *)
val fold_constr : ('c -> 'a -> 'a) -> ('c, 'c) pt -> 'a -> 'a
- val to_tuple : ('c, 't) pt -> Id.t * 'c option * 't
- val of_tuple : Id.t * 'c option * 't -> ('c, 't) pt
+ val to_tuple : ('c, 't) pt -> Id.t binder_annot * 'c option * 't
+ val of_tuple : Id.t binder_annot * 'c option * 't -> ('c, 't) pt
(** Turn [LocalDef] into [LocalAssum], identity otherwise. *)
val drop_body : ('c, 't) pt -> ('c, 't) pt
@@ -276,8 +305,8 @@ sig
module Declaration :
sig
type ('constr, 'types) pt =
- | LocalAssum of Id.t list * 'types
- | LocalDef of Id.t list * 'constr * 'types
+ | LocalAssum of Id.t binder_annot list * 'types
+ | LocalDef of Id.t binder_annot list * 'constr * 'types
val map_constr : ('c -> 'c) -> ('c, 'c) pt -> ('c, 'c) pt
val of_named_decl : ('c, 't) Named.Declaration.pt -> ('c, 't) pt
diff --git a/kernel/cooking.ml b/kernel/cooking.ml
index 22de9bfad5..9b974c4ecc 100644
--- a/kernel/cooking.ml
+++ b/kernel/cooking.ml
@@ -21,6 +21,7 @@ open Term
open Constr
open Declarations
open Univ
+open Context
module NamedDecl = Context.Named.Declaration
module RelDecl = Context.Rel.Declaration
@@ -134,12 +135,12 @@ let abstract_context hyps =
| NamedDecl.LocalDef (id, b, t) ->
let b = Vars.subst_vars subst b in
let t = Vars.subst_vars subst t in
- id, RelDecl.LocalDef (Name id, b, t)
+ id, RelDecl.LocalDef (map_annot Name.mk_name id, b, t)
| NamedDecl.LocalAssum (id, t) ->
let t = Vars.subst_vars subst t in
- id, RelDecl.LocalAssum (Name id, t)
+ id, RelDecl.LocalAssum (map_annot Name.mk_name id, t)
in
- (decl :: ctx, id :: subst)
+ (decl :: ctx, id.binder_name :: subst)
in
Context.Named.fold_outside fold hyps ~init:([], [])
@@ -159,6 +160,7 @@ type result = {
cook_type : types;
cook_universes : universes;
cook_private_univs : Univ.ContextSet.t option;
+ cook_relevance : Sorts.relevance;
cook_inline : inline;
cook_context : Constr.named_context option;
}
@@ -241,6 +243,7 @@ let cook_constant ~hcons { from = cb; info } =
cook_type = typ;
cook_universes = univs;
cook_private_univs = private_univs;
+ cook_relevance = cb.const_relevance;
cook_inline = cb.const_inline_code;
cook_context = Some const_hyps;
}
diff --git a/kernel/cooking.mli b/kernel/cooking.mli
index 89b5c60ad5..b0f143c47d 100644
--- a/kernel/cooking.mli
+++ b/kernel/cooking.mli
@@ -22,6 +22,7 @@ type result = {
cook_type : types;
cook_universes : universes;
cook_private_univs : Univ.ContextSet.t option;
+ cook_relevance : Sorts.relevance;
cook_inline : inline;
cook_context : Constr.named_context option;
}
diff --git a/kernel/declarations.ml b/kernel/declarations.ml
index 567850645e..14cfba187e 100644
--- a/kernel/declarations.ml
+++ b/kernel/declarations.ml
@@ -91,6 +91,7 @@ type constant_body = {
const_hyps : Constr.named_context; (** New: younger hyp at top *)
const_body : Constr.t Mod_subst.substituted constant_def;
const_type : types;
+ const_relevance : Sorts.relevance;
const_body_code : Cemitcodes.to_patch_substituted option;
const_universes : universes;
const_private_poly_univs : Univ.ContextSet.t option;
@@ -133,7 +134,7 @@ v}
type record_info =
| NotRecord
| FakeRecord
-| PrimRecord of (Id.t * Label.t array * types array) array
+| PrimRecord of (Id.t * Label.t array * Sorts.relevance array * types array) array
type regular_inductive_arity = {
mind_user_arity : types;
@@ -176,6 +177,8 @@ type one_inductive_body = {
mind_recargs : wf_paths; (** Signature of recursive arguments in the constructors *)
+ mind_relevant : Sorts.relevance;
+
(** {8 Datas for bytecode compilation } *)
mind_nb_constant : int; (** number of constant constructor *)
diff --git a/kernel/declareops.ml b/kernel/declareops.ml
index d56502a095..da5217eb33 100644
--- a/kernel/declareops.ml
+++ b/kernel/declareops.ml
@@ -114,6 +114,7 @@ let subst_const_body sub cb =
Option.map (Cemitcodes.subst_to_patch_subst sub) cb.const_body_code;
const_universes = cb.const_universes;
const_private_poly_univs = cb.const_private_poly_univs;
+ const_relevance = cb.const_relevance;
const_inline_code = cb.const_inline_code;
const_typing_flags = cb.const_typing_flags }
@@ -222,6 +223,7 @@ let subst_mind_packet sub mbp =
mind_nrealdecls = mbp.mind_nrealdecls;
mind_kelim = mbp.mind_kelim;
mind_recargs = subst_wf_paths sub mbp.mind_recargs (*wf_paths*);
+ mind_relevant = mbp.mind_relevant;
mind_nb_constant = mbp.mind_nb_constant;
mind_nb_args = mbp.mind_nb_args;
mind_reloc_tbl = mbp.mind_reloc_tbl }
@@ -230,10 +232,10 @@ let subst_mind_record sub r = match r with
| NotRecord -> NotRecord
| FakeRecord -> FakeRecord
| PrimRecord infos ->
- let map (id, ps, pb as info) =
+ let map (id, ps, rs, pb as info) =
let pb' = Array.Smart.map (subst_mps sub) pb in
if pb' == pb then info
- else (id, ps, pb')
+ else (id, ps, rs, pb')
in
let infos' = Array.Smart.map map infos in
if infos' == infos then r else PrimRecord infos'
@@ -269,21 +271,32 @@ let inductive_make_projection ind mib ~proj_arg =
match mib.mind_record with
| NotRecord | FakeRecord -> None
| PrimRecord infos ->
+ let _, labs, _, _ = infos.(snd ind) in
Some (Names.Projection.Repr.make ind
~proj_npars:mib.mind_nparams
~proj_arg
- (pi2 infos.(snd ind)).(proj_arg))
+ labs.(proj_arg))
let inductive_make_projections ind mib =
match mib.mind_record with
| NotRecord | FakeRecord -> None
| PrimRecord infos ->
+ let _, labs, _, _ = infos.(snd ind) in
let projs = Array.mapi (fun proj_arg lab ->
Names.Projection.Repr.make ind ~proj_npars:mib.mind_nparams ~proj_arg lab)
- (pi2 infos.(snd ind))
+ labs
in
Some projs
+let relevance_of_projection_repr mib p =
+ let _mind,i = Names.Projection.Repr.inductive p in
+ match mib.mind_record with
+ | NotRecord | FakeRecord ->
+ CErrors.anomaly ~label:"relevance_of_projection" Pp.(str "not a projection")
+ | PrimRecord infos ->
+ let _,_,rs,_ = infos.(i) in
+ rs.(Names.Projection.Repr.arg p)
+
(** {6 Hash-consing of inductive declarations } *)
let hcons_regular_ind_arity a =
diff --git a/kernel/declareops.mli b/kernel/declareops.mli
index 23a44433b3..54a853fc81 100644
--- a/kernel/declareops.mli
+++ b/kernel/declareops.mli
@@ -70,6 +70,8 @@ val inductive_make_projection : Names.inductive -> mutual_inductive_body -> proj
val inductive_make_projections : Names.inductive -> mutual_inductive_body ->
Names.Projection.Repr.t array option
+val relevance_of_projection_repr : mutual_inductive_body -> Names.Projection.Repr.t -> Sorts.relevance
+
(** {6 Kernel flags} *)
(** A default, safe set of flags for kernel type-checking *)
diff --git a/kernel/environ.ml b/kernel/environ.ml
index 5f6d5f3d0e..97c9f8654a 100644
--- a/kernel/environ.ml
+++ b/kernel/environ.ml
@@ -548,7 +548,7 @@ let lookup_projection p env =
match mib.mind_record with
| NotRecord | FakeRecord -> anomaly ~label:"lookup_projection" Pp.(str "not a projection")
| PrimRecord infos ->
- let _,_,typs = infos.(i) in
+ let _,_,_,typs = infos.(i) in
typs.(Projection.arg p)
let get_projection env ind ~proj_arg =
diff --git a/kernel/indTyping.ml b/kernel/indTyping.ml
index 0d63c8feba..5e294c9729 100644
--- a/kernel/indTyping.ml
+++ b/kernel/indTyping.ml
@@ -161,7 +161,8 @@ let check_arity env_params env_ar ind =
full_arity is used as argument or subject to cast, an upper
universe will be generated *)
let arity = it_mkProd_or_LetIn arity (Environ.rel_context env_params) in
- push_rel (LocalAssum (Name ind.mind_entry_typename, arity)) env_ar,
+ let x = Context.make_annot (Name ind.mind_entry_typename) (Sorts.relevance_of_sort ind_sort) in
+ push_rel (LocalAssum (x, arity)) env_ar,
(arity, indices, univ_info)
let check_constructor_univs env_ar_par univ_info (args,_) =
diff --git a/kernel/indtypes.ml b/kernel/indtypes.ml
index 457c17907e..545c0fe7b7 100644
--- a/kernel/indtypes.ml
+++ b/kernel/indtypes.ml
@@ -173,7 +173,9 @@ let ienv_push_inductive (env, n, ntypes, ra_env) ((mi,u),lrecparams) =
let specif = (lookup_mind_specif env mi, u) in
let ty = type_of_inductive env specif in
let env' =
- let decl = LocalAssum (Anonymous, hnf_prod_applist env ty lrecparams) in
+ let r = (snd (fst specif)).mind_relevant in
+ let anon = Context.make_annot Anonymous r in
+ let decl = LocalAssum (anon, hnf_prod_applist env ty lrecparams) in
push_rel decl env in
let ra_env' =
(Imbr mi,(Rtree.mk_rec_calls 1).(0)) ::
@@ -186,8 +188,8 @@ let rec ienv_decompose_prod (env,_,_,_ as ienv) n c =
if Int.equal n 0 then (ienv,c) else
let c' = whd_all env c in
match kind c' with
- Prod(na,a,b) ->
- let ienv' = ienv_push_var ienv (na,a,mk_norec) in
+ Prod(na,a,b) ->
+ let ienv' = ienv_push_var ienv (na,a,mk_norec) in
ienv_decompose_prod ienv' (n-1) b
| _ -> assert false
@@ -215,7 +217,7 @@ let check_positivity_one ~chkpos recursive (env,_,ntypes,_ as ienv) paramsctxt (
let rec check_pos (env, n, ntypes, ra_env as ienv) nmr c =
let x,largs = decompose_app (whd_all env c) in
match kind x with
- | Prod (na,b,d) ->
+ | Prod (na,b,d) ->
let () = assert (List.is_empty largs) in
(** If one of the inductives of the mutually inductive
block occurs in the left-hand side of a product, then
@@ -433,7 +435,7 @@ let compute_projections (kn, i as ind) mib =
mkRel 1 :: List.map (lift 1) subst in
subst
in
- let projections decl (i, j, labs, pbs, letsubst) =
+ let projections decl (i, j, labs, rs, pbs, letsubst) =
match decl with
| LocalDef (_na,c,_t) ->
(* From [params, field1,..,fieldj |- c(params,field1,..,fieldj)]
@@ -445,10 +447,11 @@ let compute_projections (kn, i as ind) mib =
(* From [params-wo-let, x:I |- subst:(params, x:I, field1,..,fieldj)]
to [params-wo-let, x:I |- subst:(params, x:I, field1,..,fieldj+1)] *)
let letsubst = c2 :: letsubst in
- (i, j+1, labs, pbs, letsubst)
+ (i, j+1, labs, rs, pbs, letsubst)
| LocalAssum (na,t) ->
- match na with
+ match na.Context.binder_name with
| Name id ->
+ let r = na.Context.binder_relevance in
let lab = Label.of_id id in
let kn = Projection.Repr.make ind ~proj_npars:mib.mind_nparams ~proj_arg:i lab in
(* from [params, field1,..,fieldj |- t(params,field1,..,fieldj)]
@@ -460,14 +463,15 @@ let compute_projections (kn, i as ind) mib =
(* from [params, x:I, field1,..,fieldj |- t(field1,..,fieldj)]
to [params, x:I |- t(proj1 x,..,projj x)] *)
let fterm = mkProj (Projection.make kn false, mkRel 1) in
- (i + 1, j + 1, lab :: labs, projty :: pbs, fterm :: letsubst)
+ (i + 1, j + 1, lab :: labs, r :: rs, projty :: pbs, fterm :: letsubst)
| Anonymous -> raise UndefinableExpansion
in
- let (_, _, labs, pbs, _letsubst) =
- List.fold_right projections ctx (0, 1, [], [], paramsletsubst)
+ let (_, _, labs, rs, pbs, _letsubst) =
+ List.fold_right projections ctx (0, 1, [], [], [], paramsletsubst)
in
- Array.of_list (List.rev labs),
- Array.of_list (List.rev pbs)
+ Array.of_list (List.rev labs),
+ Array.of_list (List.rev rs),
+ Array.of_list (List.rev pbs)
let build_inductive env names prv univs variance paramsctxt kn isrecord isfinite inds nmr recargs =
let ntypes = Array.length inds in
@@ -483,7 +487,11 @@ let build_inductive env names prv univs variance paramsctxt kn isrecord isfinite
splayed_lc in
let consnrealargs =
Array.map (fun (d,_) -> Context.Rel.nhyps d)
- splayed_lc in
+ splayed_lc in
+ let mind_relevant = match arity with
+ | RegularArity { mind_sort;_ } -> Sorts.relevance_of_sort mind_sort
+ | TemplateArity _ -> Sorts.Relevant
+ in
(* Assigning VM tags to constructors *)
let nconst, nblock = ref 0, ref 0 in
let transf num =
@@ -510,8 +518,9 @@ let build_inductive env names prv univs variance paramsctxt kn isrecord isfinite
mind_consnrealargs = consnrealargs;
mind_user_lc = lc;
mind_nf_lc = nf_lc;
- mind_recargs = recarg;
- mind_nb_constant = !nconst;
+ mind_recargs = recarg;
+ mind_relevant;
+ mind_nb_constant = !nconst;
mind_nb_args = !nblock;
mind_reloc_tbl = rtbl;
} in
@@ -546,8 +555,8 @@ let build_inductive env names prv univs variance paramsctxt kn isrecord isfinite
(** The elimination criterion ensures that all projections can be defined. *)
if Array.for_all is_record packets then
let map i id =
- let labs, projs = compute_projections (kn, i) mib in
- (id, labs, projs)
+ let labs, rs, projs = compute_projections (kn, i) mib in
+ (id, labs, rs, projs)
in
try PrimRecord (Array.mapi map rid)
with UndefinableExpansion -> FakeRecord
diff --git a/kernel/inductive.ml b/kernel/inductive.ml
index 66cd4cba9e..8b8295c64b 100644
--- a/kernel/inductive.ml
+++ b/kernel/inductive.ml
@@ -194,7 +194,11 @@ let instantiate_universes env ctx ar argsorts =
(* Type of an inductive type *)
-let type_of_inductive_gen ?(polyprop=true) env ((_mib,mip),u) paramtyps =
+let relevance_of_inductive env ind =
+ let _, mip = lookup_mind_specif env ind in
+ mip.mind_relevant
+
+let type_of_inductive_gen ?(polyprop=true) env ((_,mip),u) paramtyps =
match mip.mind_arity with
| RegularArity a -> subst_instance_constr u a.mind_user_arity
| TemplateArity ar ->
@@ -301,12 +305,12 @@ let build_dependent_inductive ind (_,mip) params =
@ Context.Rel.to_extended_list mkRel 0 realargs)
(* This exception is local *)
-exception LocalArity of (Sorts.family * Sorts.family * arity_error) option
+exception LocalArity of (Sorts.family list * Sorts.family * Sorts.family * arity_error) option
let check_allowed_sort ksort specif =
if not (CList.exists (Sorts.family_equal ksort) (elim_sorts specif)) then
let s = inductive_sort_family (snd specif) in
- raise (LocalArity (Some(ksort,s,error_elim_explain ksort s)))
+ raise (LocalArity (Some(elim_sorts specif, ksort,s,error_elim_explain ksort s)))
let is_correct_arity env c pj ind specif params =
let arsign,_ = get_instantiated_arity ind specif params in
@@ -320,7 +324,7 @@ let is_correct_arity env c pj ind specif params =
srec (push_rel (LocalAssum (na1,a1)) env) t ar'
(* The last Prod domain is the type of the scrutinee *)
| Prod (na1,a1,a2), [] -> (* whnf of t was not needed here! *)
- let env' = push_rel (LocalAssum (na1,a1)) env in
+ let env' = push_rel (LocalAssum (na1,a1)) env in
let ksort = match kind (whd_all env' a2) with
| Sort s -> Sorts.family s
| _ -> raise (LocalArity None) in
@@ -336,7 +340,7 @@ let is_correct_arity env c pj ind specif params =
in
try srec env pj.uj_type (List.rev arsign)
with LocalArity kinds ->
- error_elim_arity env ind (elim_sorts specif) c pj kinds
+ error_elim_arity env ind c pj kinds
(************************************************************************)
@@ -379,13 +383,14 @@ let type_case_branches env (pind,largs) pj c =
(************************************************************************)
(* Checking the case annotation is relevant *)
-let check_case_info env (indsp,u) ci =
+let check_case_info env (indsp,u) r ci =
let (mib,mip as spec) = lookup_mind_specif env indsp in
if
not (eq_ind indsp ci.ci_ind) ||
not (Int.equal mib.mind_nparams ci.ci_npar) ||
not (Array.equal Int.equal mip.mind_consnrealdecls ci.ci_cstr_ndecls) ||
not (Array.equal Int.equal mip.mind_consnrealargs ci.ci_cstr_nargs) ||
+ not (ci.ci_relevance == r) ||
is_primitive_record spec
then raise (TypeError(env,WrongCaseInfo((indsp,u),ci)))
@@ -574,7 +579,9 @@ let ienv_push_inductive (env, ra_env) ((mind,u),lpar) =
let mib = Environ.lookup_mind mind env in
let ntypes = mib.mind_ntypes in
let push_ind specif env =
- let decl = LocalAssum (Anonymous, hnf_prod_applist env (type_of_inductive env ((mib,specif),u)) lpar) in
+ let r = specif.mind_relevant in
+ let anon = Context.make_annot Anonymous r in
+ let decl = LocalAssum (anon, hnf_prod_applist env (type_of_inductive env ((mib,specif),u)) lpar) in
push_rel decl env
in
let env = Array.fold_right push_ind mib.mind_packets env in
@@ -595,7 +602,8 @@ let rec ienv_decompose_prod (env,_ as ienv) n c =
let dummy_univ = Level.(make (UGlobal.make (DirPath.make [Id.of_string "implicit"]) 0))
let dummy_implicit_sort = mkType (Universe.make dummy_univ)
let lambda_implicit_lift n a =
- let lambda_implicit a = mkLambda (Anonymous, dummy_implicit_sort, a) in
+ let anon = Context.make_annot Anonymous Sorts.Relevant in
+ let lambda_implicit a = mkLambda (anon, dummy_implicit_sort, a) in
iterate lambda_implicit n (lift n a)
(* This removes global parameters of the inductive types in lc (for
@@ -1021,7 +1029,7 @@ let check_one_fix renv recpos trees def =
check_rec_call (assign_var_spec renv (1,recArgsDecrArg)) [] body
else
match kind body with
- | Lambda (x,a,b) ->
+ | Lambda (x,a,b) ->
check_rec_call renv [] a;
let renv' = push_var_renv renv (x,a) in
check_nested_fix_body renv' (decr-1) recArgsDecrArg b
@@ -1054,7 +1062,7 @@ let inductive_of_mutfix env ((nvect,bodynum),(names,types,bodies as recdef)) =
match kind (whd_all env def) with
| Lambda (x,a,b) ->
if noccur_with_meta n nbfix a then
- let env' = push_rel (LocalAssum (x,a)) env in
+ let env' = push_rel (LocalAssum (x,a)) env in
if Int.equal n (k + 1) then
(* get the inductive type of the fixpoint *)
let (mind, _) =
@@ -1111,7 +1119,7 @@ let rec codomain_is_coind env c =
let b = whd_all env c in
match kind b with
| Prod (x,a,b) ->
- codomain_is_coind (push_rel (LocalAssum (x,a)) env) b
+ codomain_is_coind (push_rel (LocalAssum (x,a)) env) b
| _ ->
(try find_coinductive env b
with Not_found ->
@@ -1149,7 +1157,7 @@ let check_one_cofix env nbfix def deftype =
| _ -> anomaly_ill_typed ()
in process_args_of_constr (realargs, lra)
- | Lambda (x,a,b) ->
+ | Lambda (x,a,b) ->
let () = assert (List.is_empty args) in
if noccur_with_meta n nbfix a then
let env' = push_rel (LocalAssum (x,a)) env in
diff --git a/kernel/inductive.mli b/kernel/inductive.mli
index ad35c16c22..997a620742 100644
--- a/kernel/inductive.mli
+++ b/kernel/inductive.mli
@@ -45,6 +45,8 @@ val constrained_type_of_inductive : env -> mind_specif puniverses -> types const
val constrained_type_of_inductive_knowing_parameters :
env -> mind_specif puniverses -> types Lazy.t array -> types constrained
+val relevance_of_inductive : env -> inductive -> Sorts.relevance
+
val type_of_inductive : env -> mind_specif puniverses -> types
val type_of_inductive_knowing_parameters :
@@ -93,7 +95,7 @@ val inductive_sort_family : one_inductive_body -> Sorts.family
(** Check a [case_info] actually correspond to a Case expression on the
given inductive type. *)
-val check_case_info : env -> pinductive -> case_info -> unit
+val check_case_info : env -> pinductive -> Sorts.relevance -> case_info -> unit
(** {6 Guard conditions for fix and cofix-points. } *)
diff --git a/kernel/kernel.mllib b/kernel/kernel.mllib
index 5108744bde..59c1d5890f 100644
--- a/kernel/kernel.mllib
+++ b/kernel/kernel.mllib
@@ -26,6 +26,7 @@ Conv_oracle
Environ
Primred
CClosure
+Retypeops
Reduction
Clambda
Nativelambda
diff --git a/kernel/nativecode.ml b/kernel/nativecode.ml
index df60899b95..b1d177e76d 100644
--- a/kernel/nativecode.ml
+++ b/kernel/nativecode.ml
@@ -11,6 +11,7 @@
open CErrors
open Names
open Constr
+open Context
open Declarations
open Util
open Nativevalues
@@ -763,7 +764,7 @@ let empty_env univ () =
}
let push_rel env id =
- let local = fresh_lname id in
+ let local = fresh_lname id.binder_name in
local, { env with
env_rel = MLlocal local :: env.env_rel;
env_bound = env.env_bound + 1
@@ -772,7 +773,7 @@ let push_rel env id =
let push_rels env ids =
let lnames, env_rel =
Array.fold_left (fun (names,env_rel) id ->
- let local = fresh_lname id in
+ let local = fresh_lname id.binder_name in
(local::names, MLlocal local::env_rel)) ([],env.env_rel) ids in
Array.of_list (List.rev lnames), { env with
env_rel = env_rel;
@@ -1945,7 +1946,7 @@ let compile_mind mb mind stack =
let tbl = ob.mind_reloc_tbl in
(* Building info *)
let ci = { ci_ind = ind; ci_npar = nparams;
- ci_cstr_nargs = [|0|];
+ ci_cstr_nargs = [|0|]; ci_relevance = ob.mind_relevant;
ci_cstr_ndecls = [||] (*FIXME*);
ci_pp_info = { ind_tags = []; cstr_tags = [||] (*FIXME*); style = RegularStyle } } in
let asw = { asw_ind = ind; asw_prefix = ""; asw_ci = ci;
@@ -1968,7 +1969,7 @@ let compile_mind mb mind stack =
let projs = match mb.mind_record with
| NotRecord | FakeRecord -> []
| PrimRecord info ->
- let _, _, pbs = info.(i) in
+ let _, _, _, pbs = info.(i) in
Array.fold_left_i add_proj [] pbs
in
projs @ constructors @ gtype :: accu :: stack
diff --git a/kernel/nativelambda.ml b/kernel/nativelambda.ml
index 0869f94042..ec3a7b893d 100644
--- a/kernel/nativelambda.ml
+++ b/kernel/nativelambda.ml
@@ -26,9 +26,9 @@ type lambda =
| Lmeta of metavariable * lambda (* type *)
| Levar of Evar.t * lambda array (* arguments *)
| Lprod of lambda * lambda
- | Llam of Name.t array * lambda
- | Lrec of Name.t * lambda
- | Llet of Name.t * lambda * lambda
+ | Llam of Name.t Context.binder_annot array * lambda
+ | Lrec of Name.t Context.binder_annot * lambda
+ | Llet of Name.t Context.binder_annot * lambda * lambda
| Lapp of lambda * lambda array
| Lconst of prefix * pconstant
| Lproj of prefix * inductive * int (* prefix, inductive, index starting from 0 *)
@@ -51,9 +51,9 @@ type lambda =
| Llazy
| Lforce
-and lam_branches = (constructor * Name.t array * lambda) array
+and lam_branches = (constructor * Name.t Context.binder_annot array * lambda) array
-and fix_decl = Name.t array * lambda array * lambda array
+and fix_decl = Name.t Context.binder_annot array * lambda array * lambda array
type evars =
{ evars_val : existential -> constr option;
@@ -362,7 +362,8 @@ let prim env kn p args =
Lprim(prefix, kn, p, args)
let expand_prim env kn op arity =
- let ids = Array.make arity Anonymous in
+ (* primitives are always Relevant *)
+ let ids = Array.make arity Context.anonR in
let args = make_args arity 1 in
Llam(ids, prim env kn op args)
@@ -395,7 +396,7 @@ module Cache =
let get_construct_info cache env c : constructor_info =
try ConstrTable.find cache c
- with Not_found ->
+ with Not_found ->
let ((mind,j), i) = c in
let oib = lookup_mind mind env in
let oip = oib.mind_packets.(j) in
@@ -518,8 +519,10 @@ let rec lambda_of_constr cache env sigma c =
else
match b with
| Llam(ids, body) when Int.equal (Array.length ids) arity -> (cn, ids, body)
- | _ ->
- let ids = Array.make arity Anonymous in
+ | _ ->
+ (** TODO relevance *)
+ let anon = Context.make_annot Anonymous Sorts.Relevant in
+ let ids = Array.make arity anon in
let args = make_args arity 1 in
let ll = lam_lift arity b in
(cn, ids, mkLapp ll args) in
diff --git a/kernel/nativelambda.mli b/kernel/nativelambda.mli
index eb06522a33..b0de257a27 100644
--- a/kernel/nativelambda.mli
+++ b/kernel/nativelambda.mli
@@ -21,9 +21,9 @@ type lambda =
| Lmeta of metavariable * lambda (* type *)
| Levar of Evar.t * lambda array (* arguments *)
| Lprod of lambda * lambda
- | Llam of Name.t array * lambda
- | Lrec of Name.t * lambda
- | Llet of Name.t * lambda * lambda
+ | Llam of Name.t Context.binder_annot array * lambda
+ | Lrec of Name.t Context.binder_annot * lambda
+ | Llet of Name.t Context.binder_annot * lambda * lambda
| Lapp of lambda * lambda array
| Lconst of prefix * pconstant
| Lproj of prefix * inductive * int (* prefix, inductive, index starting from 0 *)
@@ -45,9 +45,9 @@ type lambda =
| Llazy
| Lforce
-and lam_branches = (constructor * Name.t array * lambda) array
+and lam_branches = (constructor * Name.t Context.binder_annot array * lambda) array
-and fix_decl = Name.t array * lambda array * lambda array
+and fix_decl = Name.t Context.binder_annot array * lambda array * lambda array
type evars =
{ evars_val : existential -> constr option;
@@ -55,8 +55,8 @@ type evars =
val empty_evars : evars
-val decompose_Llam : lambda -> Name.t array * lambda
-val decompose_Llam_Llet : lambda -> (Name.t * lambda option) array * lambda
+val decompose_Llam : lambda -> Name.t Context.binder_annot array * lambda
+val decompose_Llam_Llet : lambda -> (Name.t Context.binder_annot * lambda option) array * lambda
val is_lazy : constr -> bool
val mk_lazy : lambda -> lambda
diff --git a/kernel/reduction.ml b/kernel/reduction.ml
index d526b25e5b..101f02323f 100644
--- a/kernel/reduction.ml
+++ b/kernel/reduction.ml
@@ -446,7 +446,7 @@ and eqappr cv_pb l2r infos (lft1,st1) (lft2,st2) cuniv =
(* Inconsistency: we tolerate that v1, v2 contain shift and update but
we throw them away *)
if not (is_empty_stack v1 && is_empty_stack v2) then
- anomaly (Pp.str "conversion was given ill-typed terms (FLambda).");
+ anomaly (Pp.str "conversion was given ill-typed terms (FLambda).");
let (_,ty1,bd1) = destFLambda mk_clos hd1 in
let (_,ty2,bd2) = destFLambda mk_clos hd2 in
let el1 = el_stack lft1 v1 in
@@ -470,7 +470,7 @@ and eqappr cv_pb l2r infos (lft1,st1) (lft2,st2) cuniv =
| _ ->
anomaly (Pp.str "conversion was given unreduced term (FLambda).")
in
- let (_,_ty1,bd1) = destFLambda mk_clos hd1 in
+ let (_,_,bd1) = destFLambda mk_clos hd1 in
eqappr CONV l2r infos
(el_lift lft1, (bd1, [])) (el_lift lft2, (hd2, eta_expand_stack v2)) cuniv
| (_, FLambda _) ->
@@ -479,10 +479,10 @@ and eqappr cv_pb l2r infos (lft1,st1) (lft2,st2) cuniv =
| _ ->
anomaly (Pp.str "conversion was given unreduced term (FLambda).")
in
- let (_,_ty2,bd2) = destFLambda mk_clos hd2 in
+ let (_,_,bd2) = destFLambda mk_clos hd2 in
eqappr CONV l2r infos
(el_lift lft1, (hd1, eta_expand_stack v1)) (el_lift lft2, (bd2, [])) cuniv
-
+
(* only one constant, defined var or defined rel *)
| (FFlex fl1, c2) ->
begin match unfold_ref_with_args infos.cnv_inf infos.lft_tab fl1 v1 with
@@ -569,7 +569,7 @@ and eqappr cv_pb l2r infos (lft1,st1) (lft2,st2) cuniv =
with Not_found -> raise NotConvertible)
| (FFix (((op1, i1),(_,tys1,cl1)),e1), FFix(((op2, i2),(_,tys2,cl2)),e2)) ->
- if Int.equal i1 i2 && Array.equal Int.equal op1 op2
+ if Int.equal i1 i2 && Array.equal Int.equal op1 op2
then
let n = Array.length cl1 in
let fty1 = Array.map (mk_clos e1) tys1 in
@@ -896,7 +896,7 @@ let dest_prod env =
let t = whd_all env c in
match kind t with
| Prod (n,a,c0) ->
- let d = LocalAssum (n,a) in
+ let d = LocalAssum (n,a) in
decrec (push_rel d env) (Context.Rel.add d m) c0
| _ -> m,t
in
diff --git a/kernel/retypeops.ml b/kernel/retypeops.ml
new file mode 100644
index 0000000000..dc1aa20310
--- /dev/null
+++ b/kernel/retypeops.ml
@@ -0,0 +1,109 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open Util
+open Names
+open Constr
+open Declarations
+open Environ
+open Context
+
+module RelDecl = Context.Rel.Declaration
+
+let relevance_of_rel env n =
+ let decl = lookup_rel n env in
+ RelDecl.get_relevance decl
+
+let relevance_of_var env x =
+ let decl = lookup_named x env in
+ Context.Named.Declaration.get_relevance decl
+
+let relevance_of_constant env c =
+ let decl = lookup_constant c env in
+ decl.const_relevance
+
+let relevance_of_constructor env ((mi,i),_) =
+ let decl = lookup_mind mi env in
+ let packet = decl.mind_packets.(i) in
+ packet.mind_relevant
+
+let relevance_of_projection env p =
+ let mind = Projection.mind p in
+ let mib = lookup_mind mind env in
+ Declareops.relevance_of_projection_repr mib (Projection.repr p)
+
+let rec relevance_of_rel_extra env extra n =
+ match extra with
+ | [] -> relevance_of_rel env n
+ | r :: _ when Int.equal n 1 -> r
+ | _ :: extra -> relevance_of_rel_extra env extra (n-1)
+
+let relevance_of_flex env extra lft = function
+ | ConstKey (c,_) -> relevance_of_constant env c
+ | VarKey x -> relevance_of_var env x
+ | RelKey p -> relevance_of_rel_extra env extra (Esubst.reloc_rel p lft)
+
+let rec relevance_of_fterm env extra lft f =
+ let open CClosure in
+ match fterm_of f with
+ | FRel n -> relevance_of_rel_extra env extra (Esubst.reloc_rel n lft)
+ | FAtom c -> relevance_of_term_extra env extra lft (Esubst.subs_id 0) c
+ | FFlex key -> relevance_of_flex env extra lft key
+ | FInd _ | FProd _ -> Sorts.Relevant (* types are always relevant *)
+ | FConstruct (c,_) -> relevance_of_constructor env c
+ | FApp (f, _) -> relevance_of_fterm env extra lft f
+ | FProj (p, _) -> relevance_of_projection env p
+ | FFix (((_,i),(lna,_,_)), _) -> (lna.(i)).binder_relevance
+ | FCoFix ((i,(lna,_,_)), _) -> (lna.(i)).binder_relevance
+ | FCaseT (ci, _, _, _, _) -> ci.ci_relevance
+ | FLambda (len, tys, bdy, e) ->
+ let extra = List.rev_append (List.map (fun (x,_) -> binder_relevance x) tys) extra in
+ let lft = Esubst.el_liftn len lft in
+ relevance_of_term_extra env extra lft e bdy
+ | FLetIn (x, _, _, bdy, e) ->
+ relevance_of_term_extra env (x.binder_relevance :: extra)
+ (Esubst.el_lift lft) (Esubst.subs_lift e) bdy
+ | FInt _ -> Sorts.Relevant
+ | FLIFT (k, f) -> relevance_of_fterm env extra (Esubst.el_shft k lft) f
+ | FCLOS (c, e) -> relevance_of_term_extra env extra lft e c
+
+ | FEvar (_, _) -> Sorts.Relevant (* let's assume evars are relevant for now *)
+ | FLOCKED -> assert false
+
+and relevance_of_term_extra env extra lft subs c =
+ match kind c with
+ | Rel n ->
+ (match Esubst.expand_rel n subs with
+ | Inl (k, f) -> relevance_of_fterm env extra (Esubst.el_liftn k lft) f
+ | Inr (n, _) -> relevance_of_rel_extra env extra (Esubst.reloc_rel n lft))
+ | Var x -> relevance_of_var env x
+ | Sort _ | Ind _ | Prod _ -> Sorts.Relevant (* types are always relevant *)
+ | Cast (c, _, _) -> relevance_of_term_extra env extra lft subs c
+ | Lambda ({binder_relevance=r;_}, _, bdy) ->
+ relevance_of_term_extra env (r::extra) (Esubst.el_lift lft) (Esubst.subs_lift subs) bdy
+ | LetIn ({binder_relevance=r;_}, _, _, bdy) ->
+ relevance_of_term_extra env (r::extra) (Esubst.el_lift lft) (Esubst.subs_lift subs) bdy
+ | App (c, _) -> relevance_of_term_extra env extra lft subs c
+ | Const (c,_) -> relevance_of_constant env c
+ | Construct (c,_) -> relevance_of_constructor env c
+ | Case (ci, _, _, _) -> ci.ci_relevance
+ | Fix ((_,i),(lna,_,_)) -> (lna.(i)).binder_relevance
+ | CoFix (i,(lna,_,_)) -> (lna.(i)).binder_relevance
+ | Proj (p, _) -> relevance_of_projection env p
+ | Int _ -> Sorts.Relevant
+
+ | Meta _ | Evar _ -> Sorts.Relevant (* let's assume metas and evars are relevant for now *)
+
+let relevance_of_fterm env extra lft c =
+ if Environ.sprop_allowed env then relevance_of_fterm env extra lft c
+ else Sorts.Relevant
+
+let relevance_of_term env c =
+ if Environ.sprop_allowed env
+ then relevance_of_term_extra env [] Esubst.el_id (Esubst.subs_id 0) c
+ else Sorts.Relevant
diff --git a/kernel/retypeops.mli b/kernel/retypeops.mli
new file mode 100644
index 0000000000..f30c541c3f
--- /dev/null
+++ b/kernel/retypeops.mli
@@ -0,0 +1,26 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(** We can take advantage of non-cumulativity of SProp to avoid fully
+ retyping terms when we just want to know if they inhabit some
+ proof-irrelevant type. *)
+
+val relevance_of_term : Environ.env -> Constr.constr -> Sorts.relevance
+
+val relevance_of_fterm : Environ.env -> Sorts.relevance list ->
+ Esubst.lift -> CClosure.fconstr ->
+ Sorts.relevance
+
+
+(** Helpers *)
+open Names
+val relevance_of_rel_extra : Environ.env -> Sorts.relevance list -> int -> Sorts.relevance
+val relevance_of_var : Environ.env -> Id.t -> Sorts.relevance
+val relevance_of_constant : Environ.env -> Constant.t -> Sorts.relevance
+val relevance_of_constructor : Environ.env -> constructor -> Sorts.relevance
+val relevance_of_projection : Environ.env -> Projection.t -> Sorts.relevance
diff --git a/kernel/safe_typing.ml b/kernel/safe_typing.ml
index badd8d320f..edb1d0a02e 100644
--- a/kernel/safe_typing.ml
+++ b/kernel/safe_typing.ml
@@ -441,14 +441,16 @@ let safe_push_named d env =
let push_named_def (id,de) senv =
- let c, typ = Term_typing.translate_local_def senv.env id de in
- let env'' = safe_push_named (LocalDef (id, c, typ)) senv.env in
+ let c, r, typ = Term_typing.translate_local_def senv.env id de in
+ let x = Context.make_annot id r in
+ let env'' = safe_push_named (LocalDef (x, c, typ)) senv.env in
{ senv with env = env'' }
let push_named_assum ((id,t,poly),ctx) senv =
let senv' = push_context_set poly ctx senv in
- let t = Term_typing.translate_local_assum senv'.env t in
- let env'' = safe_push_named (LocalAssum (id,t)) senv'.env in
+ let t, r = Term_typing.translate_local_assum senv'.env t in
+ let x = Context.make_annot id r in
+ let env'' = safe_push_named (LocalAssum (x,t)) senv'.env in
{senv' with env=env''}
@@ -607,7 +609,7 @@ let inline_side_effects env body side_eff =
if List.is_empty side_eff then (body, Univ.ContextSet.empty, sigs)
else
(** Second step: compute the lifts and substitutions to apply *)
- let cname c = Name (Label.to_id (Constant.label c)) in
+ let cname c r = Context.make_annot (Name (Label.to_id (Constant.label c))) r in
let fold (subst, var, ctx, args) (c, cb, b) =
let (b, opaque) = match cb.const_body, b with
| Def b, _ -> (Mod_subst.force_constr b, false)
@@ -620,7 +622,7 @@ let inline_side_effects env body side_eff =
let ty = cb.const_type in
let subst = Cmap_env.add c (Inr var) subst in
let ctx = Univ.ContextSet.union ctx univs in
- (subst, var + 1, ctx, (cname c, b, ty, opaque) :: args)
+ (subst, var + 1, ctx, (cname c cb.const_relevance, b, ty, opaque) :: args)
| Polymorphic _ ->
(** Inline the term to emulate universe polymorphism *)
let subst = Cmap_env.add c (Inl b) subst in
@@ -1243,7 +1245,7 @@ let check_register_ind ind r env =
check_if (Constr.is_Type d) s;
check_if
(Constr.equal
- (mkProd (Anonymous,mkRel 1, mkApp (mkRel 3,[|mkRel 2|])))
+ (mkProd (Context.anonR,mkRel 1, mkApp (mkRel 3,[|mkRel 2|])))
cd)
s in
check_name 0 "C0";
diff --git a/kernel/sorts.ml b/kernel/sorts.ml
index 1e4e4e00b4..09c98ca1bc 100644
--- a/kernel/sorts.ml
+++ b/kernel/sorts.ml
@@ -128,6 +128,25 @@ module Hsorts =
let hcons = Hashcons.simple_hcons Hsorts.generate Hsorts.hcons hcons_univ
+(** On binders: is this variable proof relevant *)
+type relevance = Relevant | Irrelevant
+
+let relevance_equal r1 r2 = match r1,r2 with
+ | Relevant, Relevant | Irrelevant, Irrelevant -> true
+ | (Relevant | Irrelevant), _ -> false
+
+let relevance_of_sort_family = function
+ | InSProp -> Irrelevant
+ | _ -> Relevant
+
+let relevance_hash = function
+ | Relevant -> 0
+ | Irrelevant -> 1
+
+let relevance_of_sort = function
+ | SProp -> Irrelevant
+ | _ -> Relevant
+
let debug_print = function
| SProp -> Pp.(str "SProp")
| Prop -> Pp.(str "Prop")
diff --git a/kernel/sorts.mli b/kernel/sorts.mli
index 65078c2a62..c49728b146 100644
--- a/kernel/sorts.mli
+++ b/kernel/sorts.mli
@@ -48,6 +48,16 @@ val sort_of_univ : Univ.Universe.t -> t
val super : t -> t
+(** On binders: is this variable proof relevant *)
+type relevance = Relevant | Irrelevant
+
+val relevance_hash : relevance -> int
+
+val relevance_equal : relevance -> relevance -> bool
+
+val relevance_of_sort : t -> relevance
+val relevance_of_sort_family : family -> relevance
+
val debug_print : t -> Pp.t
val pr_sort_family : family -> Pp.t
diff --git a/kernel/subtyping.ml b/kernel/subtyping.ml
index dea72e8b59..1857ea3329 100644
--- a/kernel/subtyping.ml
+++ b/kernel/subtyping.ml
@@ -23,6 +23,7 @@ open Declareops
open Reduction
open Inductive
open Modops
+open Context
open Mod_subst
(*i*)
@@ -190,8 +191,8 @@ let check_inductive cst env mp1 l info1 mp2 mib2 spec2 subst1 subst2 reso1 reso2
check (fun mib -> mib.mind_record <> NotRecord) (==) (fun x -> RecordFieldExpected x);
if mib1.mind_record <> NotRecord then begin
let rec names_prod_letin t = match kind t with
- | Prod(n,_,t) -> n::(names_prod_letin t)
- | LetIn(n,_,_,t) -> n::(names_prod_letin t)
+ | Prod(n,_,t) -> n.binder_name::(names_prod_letin t)
+ | LetIn(n,_,_,t) -> n.binder_name::(names_prod_letin t)
| Cast(t,_,_) -> names_prod_letin t
| _ -> []
in
diff --git a/kernel/term.ml b/kernel/term.ml
index e67c2130d5..f09c45715f 100644
--- a/kernel/term.ml
+++ b/kernel/term.ml
@@ -14,6 +14,7 @@ open CErrors
open Names
open Vars
open Constr
+open Context
(* Deprecated *)
type sorts_family = Sorts.family = InSProp | InProp | InSet | InType
@@ -32,9 +33,11 @@ type sorts = Sorts.t = private
(* Other term constructors *)
(***************************)
-let mkNamedProd id typ c = mkProd (Name id, typ, subst_var id c)
-let mkNamedLambda id typ c = mkLambda (Name id, typ, subst_var id c)
-let mkNamedLetIn id c1 t c2 = mkLetIn (Name id, c1, t, subst_var id c2)
+let name_annot = map_annot Name.mk_name
+
+let mkNamedProd id typ c = mkProd (name_annot id, typ, subst_var id.binder_name c)
+let mkNamedLambda id typ c = mkLambda (name_annot id, typ, subst_var id.binder_name c)
+let mkNamedLetIn id c1 t c2 = mkLetIn (name_annot id, c1, t, subst_var id.binder_name c2)
(* Constructs either [(x:t)c] or [[x=b:t]c] *)
let mkProd_or_LetIn decl c =
@@ -60,10 +63,11 @@ let mkNamedProd_wo_LetIn decl c =
let open Context.Named.Declaration in
match decl with
| LocalAssum (id,t) -> mkNamedProd id t c
- | LocalDef (id,b,_t) -> subst1 b (subst_var id c)
+ | LocalDef (id,b,_) -> subst1 b (subst_var id.binder_name c)
(* non-dependent product t1 -> t2 *)
-let mkArrow t1 t2 = mkProd (Anonymous, t1, t2)
+let mkArrow t1 r t2 = mkProd (make_annot Anonymous r, t1, t2)
+let mkArrowR t1 t2 = mkArrow t1 Sorts.Relevant t2
(* Constructs either [[x:t]c] or [[x=b:t]c] *)
let mkLambda_or_LetIn decl c =
@@ -366,8 +370,8 @@ let rec isArity c =
type ('constr, 'types) kind_of_type =
| SortType of Sorts.t
| CastType of 'types * 'types
- | ProdType of Name.t * 'types * 'types
- | LetInType of Name.t * 'constr * 'types * 'types
+ | ProdType of Name.t Context.binder_annot * 'types * 'types
+ | LetInType of Name.t Context.binder_annot * 'constr * 'types * 'types
| AtomicType of 'constr * 'constr array
let kind_of_type t = match kind t with
diff --git a/kernel/term.mli b/kernel/term.mli
index 7fa817bada..4265324693 100644
--- a/kernel/term.mli
+++ b/kernel/term.mli
@@ -17,12 +17,15 @@ open Constr
[forall (_:t1), t2]. Beware [t_2] is NOT lifted.
Eg: in context [A:Prop], [A->A] is built by [(mkArrow (mkRel 1) (mkRel 2))]
*)
-val mkArrow : types -> types -> constr
+val mkArrow : types -> Sorts.relevance -> types -> constr
+
+val mkArrowR : types -> types -> constr
+(** For an always-relevant domain *)
(** Named version of the functions from [Term]. *)
-val mkNamedLambda : Id.t -> types -> constr -> constr
-val mkNamedLetIn : Id.t -> constr -> types -> constr -> constr
-val mkNamedProd : Id.t -> types -> types -> types
+val mkNamedLambda : Id.t Context.binder_annot -> types -> constr -> constr
+val mkNamedLetIn : Id.t Context.binder_annot -> constr -> types -> constr -> constr
+val mkNamedProd : Id.t Context.binder_annot -> types -> types -> types
(** Constructs either [(x:t)c] or [[x=b:t]c] *)
val mkProd_or_LetIn : Constr.rel_declaration -> types -> types
@@ -45,24 +48,24 @@ val appvectc : constr -> constr array -> constr
(** [prodn n l b] = [forall (x_1:T_1)...(x_n:T_n), b]
where [l] is [(x_n,T_n)...(x_1,T_1)...]. *)
-val prodn : int -> (Name.t * constr) list -> constr -> constr
+val prodn : int -> (Name.t Context.binder_annot * constr) list -> constr -> constr
(** [compose_prod l b]
@return [forall (x_1:T_1)...(x_n:T_n), b]
where [l] is [(x_n,T_n)...(x_1,T_1)].
Inverse of [decompose_prod]. *)
-val compose_prod : (Name.t * constr) list -> constr -> constr
+val compose_prod : (Name.t Context.binder_annot * constr) list -> constr -> constr
(** [lamn n l b]
@return [fun (x_1:T_1)...(x_n:T_n) => b]
where [l] is [(x_n,T_n)...(x_1,T_1)...]. *)
-val lamn : int -> (Name.t * constr) list -> constr -> constr
+val lamn : int -> (Name.t Context.binder_annot * constr) list -> constr -> constr
(** [compose_lam l b]
@return [fun (x_1:T_1)...(x_n:T_n) => b]
where [l] is [(x_n,T_n)...(x_1,T_1)].
Inverse of [it_destLam] *)
-val compose_lam : (Name.t * constr) list -> constr -> constr
+val compose_lam : (Name.t Context.binder_annot * constr) list -> constr -> constr
(** [to_lambda n l]
@return [fun (x_1:T_1)...(x_n:T_n) => T]
@@ -107,22 +110,22 @@ val prod_applist_assum : int -> types -> constr list -> types
(** Transforms a product term {% $ %}(x_1:T_1)..(x_n:T_n)T{% $ %} into the pair
{% $ %}([(x_n,T_n);...;(x_1,T_1)],T){% $ %}, where {% $ %}T{% $ %} is not a product. *)
-val decompose_prod : constr -> (Name.t*constr) list * constr
+val decompose_prod : constr -> (Name.t Context.binder_annot * constr) list * constr
(** Transforms a lambda term {% $ %}[x_1:T_1]..[x_n:T_n]T{% $ %} into the pair
{% $ %}([(x_n,T_n);...;(x_1,T_1)],T){% $ %}, where {% $ %}T{% $ %} is not a lambda. *)
-val decompose_lam : constr -> (Name.t*constr) list * constr
+val decompose_lam : constr -> (Name.t Context.binder_annot * constr) list * constr
(** Given a positive integer n, decompose a product term
{% $ %}(x_1:T_1)..(x_n:T_n)T{% $ %}
into the pair {% $ %}([(xn,Tn);...;(x1,T1)],T){% $ %}.
Raise a user error if not enough products. *)
-val decompose_prod_n : int -> constr -> (Name.t * constr) list * constr
+val decompose_prod_n : int -> constr -> (Name.t Context.binder_annot * constr) list * constr
(** Given a positive integer {% $ %}n{% $ %}, decompose a lambda term
{% $ %}[x_1:T_1]..[x_n:T_n]T{% $ %} into the pair {% $ %}([(x_n,T_n);...;(x_1,T_1)],T){% $ %}.
Raise a user error if not enough lambdas. *)
-val decompose_lam_n : int -> constr -> (Name.t * constr) list * constr
+val decompose_lam_n : int -> constr -> (Name.t Context.binder_annot * constr) list * constr
(** Extract the premisses and the conclusion of a term of the form
"(xi:Ti) ... (xj:=cj:Tj) ..., T" where T is not a product nor a let *)
@@ -183,8 +186,8 @@ val isArity : types -> bool
type ('constr, 'types) kind_of_type =
| SortType of Sorts.t
| CastType of 'types * 'types
- | ProdType of Name.t * 'types * 'types
- | LetInType of Name.t * 'constr * 'types * 'types
+ | ProdType of Name.t Context.binder_annot * 'types * 'types
+ | LetInType of Name.t Context.binder_annot * 'constr * 'types * 'types
| AtomicType of 'constr * 'constr array
val kind_of_type : types -> (constr, types) kind_of_type
diff --git a/kernel/term_typing.ml b/kernel/term_typing.ml
index 929f1c13a3..af96cfdb4f 100644
--- a/kernel/term_typing.ml
+++ b/kernel/term_typing.ml
@@ -74,13 +74,14 @@ let infer_declaration (type a) ~(trust : a trust) env (dcl : a constant_entry) =
in
let j = infer env t in
let usubst, univs = Declareops.abstract_universes uctx in
- let c = Typeops.assumption_of_judgment env j in
+ let c, r = Typeops.assumption_of_judgment env j in
let t = Constr.hcons (Vars.subst_univs_level_constr usubst c) in
{
Cooking.cook_body = Undef nl;
cook_type = t;
cook_universes = univs;
cook_private_univs = None;
+ cook_relevance = r;
cook_inline = false;
cook_context = ctx;
}
@@ -110,7 +111,8 @@ let infer_declaration (type a) ~(trust : a trust) env (dcl : a constant_entry) =
cook_universes = Monomorphic uctxt;
cook_private_univs = None;
cook_inline = false;
- cook_context = None
+ cook_context = None;
+ cook_relevance = Sorts.Relevant;
}
(** Definition [c] is opaque (Qed), non polymorphic and with a specified type,
@@ -159,6 +161,7 @@ the polymorphic case
cook_type = typ;
cook_universes = Monomorphic univs;
cook_private_univs = None;
+ cook_relevance = Sorts.relevance_of_sort tyj.utj_type;
cook_inline = c.const_entry_inline_code;
cook_context = c.const_entry_secctx;
}
@@ -214,6 +217,7 @@ the polymorphic case
cook_type = typ;
cook_universes = univs;
cook_private_univs = private_univs;
+ cook_relevance = Retypeops.relevance_of_term env body;
cook_inline = c.const_entry_inline_code;
cook_context = c.const_entry_secctx;
}
@@ -309,6 +313,7 @@ let build_constant_declaration _kn env result =
const_body_code = tps;
const_universes = univs;
const_private_poly_univs = result.cook_private_univs;
+ const_relevance = result.cook_relevance;
const_inline_code = result.cook_inline;
const_typing_flags = Environ.typing_flags env }
@@ -366,7 +371,7 @@ let translate_local_def env _id centry =
p
| Undef _ | Primitive _ -> assert false
in
- c, typ
+ c, decl.cook_relevance, typ
(* Insertion of inductive types. *)
diff --git a/kernel/term_typing.mli b/kernel/term_typing.mli
index faf434c142..d34c28138e 100644
--- a/kernel/term_typing.mli
+++ b/kernel/term_typing.mli
@@ -27,9 +27,9 @@ type _ trust =
| SideEffects : 'a effect_handler -> 'a trust
val translate_local_def : env -> Id.t -> section_def_entry ->
- constr * types
+ constr * Sorts.relevance * types
-val translate_local_assum : env -> types -> types
+val translate_local_assum : env -> types -> types * Sorts.relevance
val translate_constant :
'a trust -> env -> Constant.t -> 'a constant_entry ->
diff --git a/kernel/type_errors.ml b/kernel/type_errors.ml
index 814ef8bdfc..9168c32f0e 100644
--- a/kernel/type_errors.ml
+++ b/kernel/type_errors.ml
@@ -47,8 +47,8 @@ type ('constr, 'types) ptype_error =
| NotAType of ('constr, 'types) punsafe_judgment
| BadAssumption of ('constr, 'types) punsafe_judgment
| ReferenceVariables of Id.t * 'constr
- | ElimArity of pinductive * Sorts.family list * 'constr * ('constr, 'types) punsafe_judgment
- * (Sorts.family * Sorts.family * arity_error) option
+ | ElimArity of pinductive * 'constr * ('constr, 'types) punsafe_judgment
+ * (Sorts.family list * Sorts.family * Sorts.family * arity_error) option
| CaseNotInductive of ('constr, 'types) punsafe_judgment
| WrongCaseInfo of pinductive * case_info
| NumberBranches of ('constr, 'types) punsafe_judgment * int
@@ -59,12 +59,13 @@ type ('constr, 'types) ptype_error =
| CantApplyBadType of
(int * 'constr * 'constr) * ('constr, 'types) punsafe_judgment * ('constr, 'types) punsafe_judgment array
| CantApplyNonFunctional of ('constr, 'types) punsafe_judgment * ('constr, 'types) punsafe_judgment array
- | IllFormedRecBody of 'constr pguard_error * Name.t array * int * env * ('constr, 'types) punsafe_judgment array
+ | IllFormedRecBody of 'constr pguard_error * Name.t Context.binder_annot array * int * env * ('constr, 'types) punsafe_judgment array
| IllTypedRecBody of
- int * Name.t array * ('constr, 'types) punsafe_judgment array * 'types array
+ int * Name.t Context.binder_annot array * ('constr, 'types) punsafe_judgment array * 'types array
| UnsatisfiedConstraints of Univ.Constraint.t
| UndeclaredUniverse of Univ.Level.t
| DisallowedSProp
+ | BadRelevance
type type_error = (constr, types) ptype_error
@@ -103,8 +104,8 @@ let error_assumption env j =
let error_reference_variables env id c =
raise (TypeError (env, ReferenceVariables (id,c)))
-let error_elim_arity env ind aritylst c pj okinds =
- raise (TypeError (env, ElimArity (ind,aritylst,c,pj,okinds)))
+let error_elim_arity env ind c pj okinds =
+ raise (TypeError (env, ElimArity (ind,c,pj,okinds)))
let error_case_not_inductive env j =
raise (TypeError (env, CaseNotInductive j))
@@ -153,6 +154,9 @@ let error_undeclared_universe env l =
let error_disallowed_sprop env =
raise (TypeError (env, DisallowedSProp))
+let error_bad_relevance env =
+ raise (TypeError (env, BadRelevance))
+
let map_pguard_error f = function
| NotEnoughAbstractionInFixBody -> NotEnoughAbstractionInFixBody
| RecursionNotOnInductiveType c -> RecursionNotOnInductiveType (f c)
@@ -176,7 +180,7 @@ let map_ptype_error f = function
| NotAType j -> NotAType (on_judgment f j)
| BadAssumption j -> BadAssumption (on_judgment f j)
| ReferenceVariables (id, c) -> ReferenceVariables (id, f c)
-| ElimArity (pi, dl, c, j, ar) -> ElimArity (pi, dl, f c, on_judgment f j, ar)
+| ElimArity (pi, c, j, ar) -> ElimArity (pi, f c, on_judgment f j, ar)
| CaseNotInductive j -> CaseNotInductive (on_judgment f j)
| WrongCaseInfo (pi, ci) -> WrongCaseInfo (pi, ci)
| NumberBranches (j, n) -> NumberBranches (on_judgment f j, n)
@@ -194,3 +198,4 @@ let map_ptype_error f = function
| UnsatisfiedConstraints g -> UnsatisfiedConstraints g
| UndeclaredUniverse l -> UndeclaredUniverse l
| DisallowedSProp -> DisallowedSProp
+| BadRelevance -> BadRelevance
diff --git a/kernel/type_errors.mli b/kernel/type_errors.mli
index e208c57e0a..41a5f19e78 100644
--- a/kernel/type_errors.mli
+++ b/kernel/type_errors.mli
@@ -48,8 +48,8 @@ type ('constr, 'types) ptype_error =
| NotAType of ('constr, 'types) punsafe_judgment
| BadAssumption of ('constr, 'types) punsafe_judgment
| ReferenceVariables of Id.t * 'constr
- | ElimArity of pinductive * Sorts.family list * 'constr * ('constr, 'types) punsafe_judgment
- * (Sorts.family * Sorts.family * arity_error) option
+ | ElimArity of pinductive * 'constr * ('constr, 'types) punsafe_judgment
+ * (Sorts.family list * Sorts.family * Sorts.family * arity_error) option
| CaseNotInductive of ('constr, 'types) punsafe_judgment
| WrongCaseInfo of pinductive * case_info
| NumberBranches of ('constr, 'types) punsafe_judgment * int
@@ -60,12 +60,13 @@ type ('constr, 'types) ptype_error =
| CantApplyBadType of
(int * 'constr * 'constr) * ('constr, 'types) punsafe_judgment * ('constr, 'types) punsafe_judgment array
| CantApplyNonFunctional of ('constr, 'types) punsafe_judgment * ('constr, 'types) punsafe_judgment array
- | IllFormedRecBody of 'constr pguard_error * Name.t array * int * env * ('constr, 'types) punsafe_judgment array
+ | IllFormedRecBody of 'constr pguard_error * Name.t Context.binder_annot array * int * env * ('constr, 'types) punsafe_judgment array
| IllTypedRecBody of
- int * Name.t array * ('constr, 'types) punsafe_judgment array * 'types array
+ int * Name.t Context.binder_annot array * ('constr, 'types) punsafe_judgment array * 'types array
| UnsatisfiedConstraints of Univ.Constraint.t
| UndeclaredUniverse of Univ.Level.t
| DisallowedSProp
+ | BadRelevance
type type_error = (constr, types) ptype_error
@@ -101,8 +102,8 @@ val error_assumption : env -> unsafe_judgment -> 'a
val error_reference_variables : env -> Id.t -> constr -> 'a
val error_elim_arity :
- env -> pinductive -> Sorts.family list -> constr -> unsafe_judgment ->
- (Sorts.family * Sorts.family * arity_error) option -> 'a
+ env -> pinductive -> constr -> unsafe_judgment ->
+ (Sorts.family list * Sorts.family * Sorts.family * arity_error) option -> 'a
val error_case_not_inductive : env -> unsafe_judgment -> 'a
@@ -124,10 +125,10 @@ val error_cant_apply_bad_type :
unsafe_judgment -> unsafe_judgment array -> 'a
val error_ill_formed_rec_body :
- env -> guard_error -> Name.t array -> int -> env -> unsafe_judgment array -> 'a
+ env -> guard_error -> Name.t Context.binder_annot array -> int -> env -> unsafe_judgment array -> 'a
val error_ill_typed_rec_body :
- env -> int -> Name.t array -> unsafe_judgment array -> types array -> 'a
+ env -> int -> Name.t Context.binder_annot array -> unsafe_judgment array -> types array -> 'a
val error_elim_explain : Sorts.family -> Sorts.family -> arity_error
@@ -137,5 +138,7 @@ val error_undeclared_universe : env -> Univ.Level.t -> 'a
val error_disallowed_sprop : env -> 'a
+val error_bad_relevance : env -> 'a
+
val map_pguard_error : ('c -> 'd) -> 'c pguard_error -> 'd pguard_error
val map_ptype_error : ('c -> 'd) -> ('c, 'c) ptype_error -> ('d, 'd) ptype_error
diff --git a/kernel/typeops.ml b/kernel/typeops.ml
index 1232ab654e..337b66e8e7 100644
--- a/kernel/typeops.ml
+++ b/kernel/typeops.ml
@@ -12,6 +12,7 @@ open CErrors
open Util
open Names
open Univ
+open Sorts
open Term
open Constr
open Vars
@@ -47,11 +48,19 @@ let check_type env c t =
(* This should be a type intended to be assumed. The error message is
not as useful as for [type_judgment]. *)
-let check_assumption env t ty =
- try let _ = check_type env t ty in t
+let infer_assumption env t ty =
+ try
+ let s = check_type env t ty in
+ t, (match s with Sorts.SProp -> Irrelevant | _ -> Relevant)
with TypeError _ ->
error_assumption env (make_judge t ty)
+let check_assumption env x t ty =
+ let r = x.Context.binder_relevance in
+ let t, r' = infer_assumption env t ty in
+ if not (r == r') then error_bad_relevance env;
+ t
+
(************************************************)
(* Incremental typing rules: builds a typing judgment given the *)
(* judgments for the subterms. *)
@@ -220,7 +229,7 @@ let type_of_prim env t =
in
let rec nary_int63_op arity ty =
if Int.equal arity 0 then ty
- else Constr.mkProd(Name (Id.of_string "x"), int_ty, nary_int63_op (arity-1) ty)
+ else Constr.mkProd(Context.nameR (Id.of_string "x"), int_ty, nary_int63_op (arity-1) ty)
in
let return_ty =
let open CPrimitives in
@@ -377,7 +386,9 @@ let type_of_case env ci p pt c ct _lf lft =
let (pind, _ as indspec) =
try find_rectype env ct
with Not_found -> error_case_not_inductive env (make_judge c ct) in
- let () = check_case_info env pind ci in
+ let _, sp = try dest_arity env pt
+ with NotArity -> error_elim_arity env pind c (make_judge p pt) None in
+ let () = check_case_info env pind (Sorts.relevance_of_sort sp) ci in
let (bty,rslty) =
type_case_branches env indspec (make_judge p pt) c in
let () = check_branch_types env pind c ct lft bty in
@@ -456,6 +467,10 @@ let constr_of_global_in_context env r =
(************************************************************************)
(************************************************************************)
+let check_binder_annot env s x =
+ let r = x.Context.binder_relevance in
+ if not (Sorts.relevance_of_sort s == r) then error_bad_relevance env
+
(* The typing machine. *)
(* ATTENTION : faudra faire le typage du contexte des Const,
Ind et Constructsi un jour cela devient des constructions
@@ -499,20 +514,23 @@ let rec execute env cstr =
type_of_apply env f ft args argst
| Lambda (name,c1,c2) ->
- let _ = execute_is_type env c1 in
+ let s = execute_is_type env c1 in
+ check_binder_annot env s name;
let env1 = push_rel (LocalAssum (name,c1)) env in
let c2t = execute env1 c2 in
type_of_abstraction env name c1 c2t
| Prod (name,c1,c2) ->
let vars = execute_is_type env c1 in
+ check_binder_annot env vars name;
let env1 = push_rel (LocalAssum (name,c1)) env in
let vars' = execute_is_type env1 c2 in
type_of_product env name vars vars'
| LetIn (name,c1,c2,c3) ->
let c1t = execute env c1 in
- let _c2s = execute_is_type env c2 in
+ let c2s = execute_is_type env c2 in
+ check_binder_annot env c2s name;
let () = check_cast env c1 c1t DEFAULTcast c2 in
let env1 = push_rel (LocalDef (name,c1,c2)) env in
let c3t = execute env1 c3 in
@@ -563,7 +581,7 @@ and execute_is_type env constr =
and execute_recdef env (names,lar,vdef) i =
let lart = execute_array env lar in
- let lara = Array.map2 (check_assumption env) lar lart in
+ let lara = Array.map3 (check_assumption env) names lar lart in
let env1 = push_rec_types (names,lara,vdef) env in
let vdeft = execute_array env1 vdef in
let () = check_fixpoint env1 names lara vdef vdeft in
@@ -591,7 +609,7 @@ let infer =
else (fun b c -> infer b c)
let assumption_of_judgment env {uj_val=c; uj_type=t} =
- check_assumption env c t
+ infer_assumption env c t
let type_judgment env {uj_val=c; uj_type=t} =
let s = check_type env c t in
@@ -644,17 +662,17 @@ let judge_of_apply env funj argjv =
let args, argtys = dest_judgev argjv in
make_judge (mkApp (funj.uj_val, args)) (type_of_apply env funj.uj_val funj.uj_type args argtys)
-let judge_of_abstraction env x varj bodyj =
- make_judge (mkLambda (x, varj.utj_val, bodyj.uj_val))
- (type_of_abstraction env x varj.utj_val bodyj.uj_type)
+(* let judge_of_abstraction env x varj bodyj = *)
+(* make_judge (mkLambda (x, varj.utj_val, bodyj.uj_val)) *)
+(* (type_of_abstraction env x varj.utj_val bodyj.uj_type) *)
-let judge_of_product env x varj outj =
- make_judge (mkProd (x, varj.utj_val, outj.utj_val))
- (mkSort (sort_of_product env varj.utj_type outj.utj_type))
+(* let judge_of_product env x varj outj = *)
+(* make_judge (mkProd (x, varj.utj_val, outj.utj_val)) *)
+(* (mkSort (sort_of_product env varj.utj_type outj.utj_type)) *)
-let judge_of_letin _env name defj typj j =
- make_judge (mkLetIn (name, defj.uj_val, typj.utj_val, j.uj_val))
- (subst1 defj.uj_val j.uj_type)
+(* let judge_of_letin env name defj typj j = *)
+(* make_judge (mkLetIn (name, defj.uj_val, typj.utj_val, j.uj_val)) *)
+(* (subst1 defj.uj_val j.uj_type) *)
let judge_of_cast env cj k tj =
let () = check_cast env cj.uj_val cj.uj_type k tj.utj_val in
diff --git a/kernel/typeops.mli b/kernel/typeops.mli
index 52c261c5e8..e57d6622c9 100644
--- a/kernel/typeops.mli
+++ b/kernel/typeops.mli
@@ -34,9 +34,11 @@ val check_context :
(** If [j] is the judgement {% $ %}c:t{% $ %}, then [assumption_of_judgement env j]
returns the type {% $ %}c{% $ %}, checking that {% $ %}t{% $ %} is a sort. *)
-val assumption_of_judgment : env -> unsafe_judgment -> types
+val assumption_of_judgment : env -> unsafe_judgment -> types * Sorts.relevance
val type_judgment : env -> unsafe_judgment -> unsafe_type_judgment
+val check_binder_annot : env -> Sorts.t -> Name.t Context.binder_annot -> unit
+
(** {6 Type of sorts. } *)
val type1 : types
val type_of_sort : Sorts.t -> types
@@ -65,21 +67,21 @@ val judge_of_apply :
-> unsafe_judgment
(** {6 Type of an abstraction. } *)
-val judge_of_abstraction :
- env -> Name.t -> unsafe_type_judgment -> unsafe_judgment
- -> unsafe_judgment
+(* val judge_of_abstraction : *)
+(* env -> Name.t -> unsafe_type_judgment -> unsafe_judgment *)
+(* -> unsafe_judgment *)
(** {6 Type of a product. } *)
val sort_of_product : env -> Sorts.t -> Sorts.t -> Sorts.t
-val type_of_product : env -> Name.t -> Sorts.t -> Sorts.t -> types
-val judge_of_product :
- env -> Name.t -> unsafe_type_judgment -> unsafe_type_judgment
- -> unsafe_judgment
+val type_of_product : env -> Name.t Context.binder_annot -> Sorts.t -> Sorts.t -> types
+(* val judge_of_product : *)
+(* env -> Name.t -> unsafe_type_judgment -> unsafe_type_judgment *)
+(* -> unsafe_judgment *)
(** s Type of a let in. *)
-val judge_of_letin :
- env -> Name.t -> unsafe_judgment -> unsafe_type_judgment -> unsafe_judgment
- -> unsafe_judgment
+(* val judge_of_letin : *)
+(* env -> Name.t -> unsafe_judgment -> unsafe_type_judgment -> unsafe_judgment *)
+(* -> unsafe_judgment *)
(** {6 Type of a cast. } *)
val judge_of_cast :
diff --git a/plugins/cc/ccalgo.ml b/plugins/cc/ccalgo.ml
index 575d964158..23cdae7883 100644
--- a/plugins/cc/ccalgo.ml
+++ b/plugins/cc/ccalgo.ml
@@ -17,6 +17,7 @@ open Pp
open Names
open Sorts
open Constr
+open Context
open Vars
open Goptions
open Tacmach
@@ -421,11 +422,11 @@ let new_representative typ =
let _A_ = Name (Id.of_string "A")
let _B_ = Name (Id.of_string "A")
-let _body_ = mkProd(Anonymous,mkRel 2,mkRel 2)
+let _body_ = mkProd(make_annot Anonymous Sorts.Relevant,mkRel 2,mkRel 2)
let cc_product s1 s2 =
- mkLambda(_A_,mkSort(s1),
- mkLambda(_B_,mkSort(s2),_body_))
+ mkLambda(make_annot _A_ Sorts.Relevant,mkSort(s1),
+ mkLambda(make_annot _B_ Sorts.Relevant,mkSort(s2),_body_))
let rec constr_of_term = function
Symb s-> s
@@ -452,11 +453,11 @@ let rec canonize_name sigma c =
let canon_mind = MutInd.make1 (MutInd.canonical kn) in
mkConstructU (((canon_mind,i),j),u)
| Prod (na,t,ct) ->
- mkProd (na,func t, func ct)
+ mkProd (na,func t, func ct)
| Lambda (na,t,ct) ->
- mkLambda (na, func t,func ct)
+ mkLambda (na, func t,func ct)
| LetIn (na,b,t,ct) ->
- mkLetIn (na, func b,func t,func ct)
+ mkLetIn (na, func b,func t,func ct)
| App (ct,l) ->
mkApp (func ct,Array.Smart.map func l)
| Proj(p,c) ->
@@ -806,7 +807,8 @@ let __eps__ = Id.of_string "_eps_"
let new_state_var typ state =
let ids = Environ.ids_of_named_context_val (Environ.named_context_val state.env) in
let id = Namegen.next_ident_away __eps__ ids in
- state.env<- EConstr.push_named (Context.Named.Declaration.LocalAssum (id,typ)) state.env;
+ let r = Sorts.Relevant in (* TODO relevance *)
+ state.env<- EConstr.push_named (Context.Named.Declaration.LocalAssum (make_annot id r,typ)) state.env;
id
let complete_one_class state i=
@@ -814,9 +816,9 @@ let complete_one_class state i=
Partial pac ->
let rec app t typ n =
if n<=0 then t else
- let _,etyp,rest= destProd typ in
+ let _,etyp,rest= destProd typ in
let id = new_state_var (EConstr.of_constr etyp) state in
- app (Appli(t,Eps id)) (substl [mkVar id] rest) (n-1) in
+ app (Appli(t,Eps id)) (substl [mkVar id] rest) (n-1) in
let _c = Typing.unsafe_type_of state.env state.sigma
(EConstr.of_constr (constr_of_term (term state.uf pac.cnode))) in
let _c = EConstr.Unsafe.to_constr _c in
diff --git a/plugins/cc/cctac.ml b/plugins/cc/cctac.ml
index 055d36747d..5778acce0a 100644
--- a/plugins/cc/cctac.ml
+++ b/plugins/cc/cctac.ml
@@ -15,6 +15,7 @@ open Names
open Inductiveops
open Declarations
open Constr
+open Context
open EConstr
open Vars
open Tactics
@@ -151,19 +152,19 @@ let patterns_of_constr env sigma nrels term=
let rec quantified_atom_of_constr env sigma nrels term =
match EConstr.kind sigma (whd_delta env sigma term) with
- Prod (id,atom,ff) ->
+ Prod (id,atom,ff) ->
if is_global sigma (Lazy.force _False) ff then
let patts=patterns_of_constr env sigma nrels atom in
`Nrule patts
else
- quantified_atom_of_constr (EConstr.push_rel (RelDecl.LocalAssum (id,atom)) env) sigma (succ nrels) ff
+ quantified_atom_of_constr (EConstr.push_rel (RelDecl.LocalAssum (id,atom)) env) sigma (succ nrels) ff
| _ ->
let patts=patterns_of_constr env sigma nrels term in
`Rule patts
let litteral_of_constr env sigma term=
match EConstr.kind sigma (whd_delta env sigma term) with
- | Prod (id,atom,ff) ->
+ | Prod (id,atom,ff) ->
if is_global sigma (Lazy.force _False) ff then
match (atom_of_constr env sigma atom) with
`Eq(t,a,b) -> `Neq(t,a,b)
@@ -171,7 +172,7 @@ let litteral_of_constr env sigma term=
else
begin
try
- quantified_atom_of_constr (EConstr.push_rel (RelDecl.LocalAssum (id,atom)) env) sigma 1 ff
+ quantified_atom_of_constr (EConstr.push_rel (RelDecl.LocalAssum (id,atom)) env) sigma 1 ff
with Not_found ->
`Other (decompose_term env sigma term)
end
@@ -233,7 +234,7 @@ let build_projection intype (cstr:pconstructor) special default gls=
let sigma = project gls in
let body=Equality.build_selector (pf_env gls) sigma ci (mkRel 1) intype special default in
let id=pf_get_new_id (Id.of_string "t") gls in
- sigma, mkLambda(Name id,intype,body)
+ sigma, mkLambda(make_annot (Name id) Sorts.Relevant,intype,body)
(* generate an adhoc tactic following the proof tree *)
@@ -318,7 +319,7 @@ let rec proof_tac p : unit Proofview.tactic =
refresh_universes (type_of tx1) (fun typx ->
refresh_universes (type_of (mkApp (tf1,[|tx1|]))) (fun typfx ->
let id = Tacmach.New.pf_get_new_id (Id.of_string "f") gl in
- let appx1 = mkLambda(Name id,typf,mkApp(mkRel 1,[|tx1|])) in
+ let appx1 = mkLambda(make_annot (Name id) Sorts.Relevant,typf,mkApp(mkRel 1,[|tx1|])) in
let lemma1 = app_global_with_holes _f_equal [|typf;typfx;appx1;tf1;tf2|] 1 in
let lemma2 = app_global_with_holes _f_equal [|typx;typfx;tf2;tx1;tx2|] 1 in
let prf =
@@ -377,7 +378,7 @@ let convert_to_goal_tac c t1 t2 p =
let neweq= app_global _eq [|sort;tt1;tt2|] in
let e = Tacmach.New.pf_get_new_id (Id.of_string "e") gl in
let x = Tacmach.New.pf_get_new_id (Id.of_string "X") gl in
- let identity=mkLambda (Name x,sort,mkRel 1) in
+ let identity=mkLambda (make_annot (Name x) Sorts.Relevant,sort,mkRel 1) in
let endt = app_global _eq_rect [|sort;tt1;identity;c;tt2;mkVar e|] in
Tacticals.New.tclTHENS (neweq (assert_before (Name e)))
[proof_tac p; endt refine_exact_check]
diff --git a/plugins/derive/derive.ml b/plugins/derive/derive.ml
index d06a241969..afdbfa1999 100644
--- a/plugins/derive/derive.ml
+++ b/plugins/derive/derive.ml
@@ -9,6 +9,7 @@
(************************************************************************)
open Constr
+open Context
open Context.Named.Declaration
let map_const_entry_body (f:constr->constr) (x:Safe_typing.private_constants Entries.const_entry_body)
@@ -39,7 +40,7 @@ let start_deriving f suchthat lemma =
TCons ( env , sigma , f_type , (fun sigma ef ->
let f_type = EConstr.Unsafe.to_constr f_type in
let ef = EConstr.Unsafe.to_constr ef in
- let env' = Environ.push_named (LocalDef (f, ef, f_type)) env in
+ let env' = Environ.push_named (LocalDef (annotR f, ef, f_type)) env in
let sigma, suchthat = Constrintern.interp_type_evars ~program_mode:false env' sigma suchthat in
TCons ( env' , sigma , suchthat , (fun sigma _ ->
TNil sigma))))))
diff --git a/plugins/extraction/extract_env.ml b/plugins/extraction/extract_env.ml
index b59e3b608c..0fa9be21c9 100644
--- a/plugins/extraction/extract_env.ml
+++ b/plugins/extraction/extract_env.ml
@@ -150,7 +150,7 @@ let check_fix env sg cb i =
| Undef _ | OpaqueDef _ | Primitive _ -> raise Impossible
let prec_declaration_equal sg (na1, ca1, ta1) (na2, ca2, ta2) =
- Array.equal Name.equal na1 na2 &&
+ Array.equal (Context.eq_annot Name.equal) na1 na2 &&
Array.equal (EConstr.eq_constr sg) ca1 ca2 &&
Array.equal (EConstr.eq_constr sg) ta1 ta2
diff --git a/plugins/extraction/extraction.ml b/plugins/extraction/extraction.ml
index 877ea9a537..d1c4a912f2 100644
--- a/plugins/extraction/extraction.ml
+++ b/plugins/extraction/extraction.ml
@@ -13,6 +13,7 @@ open Util
open Names
open Term
open Constr
+open Context
open Declarations
open Declareops
open Environ
@@ -184,7 +185,7 @@ let rec type_sign_vl env sg c =
| Prod (n,t,d) ->
let s,vl = type_sign_vl (push_rel_assum (n,t) env) sg d in
if not (is_info_scheme env sg t) then Kill Kprop::s, vl
- else Keep::s, (make_typvar n vl) :: vl
+ else Keep::s, (make_typvar n.binder_name vl) :: vl
| _ -> [],[]
let rec nb_default_params env sg c =
@@ -264,14 +265,14 @@ let rec extract_type env sg db j c args =
(* We just accumulate the arguments. *)
extract_type env sg db j d (Array.to_list args' @ args)
| Lambda (_,_,d) ->
- (match args with
+ (match args with
| [] -> assert false (* A lambda cannot be a type. *)
| a :: args -> extract_type env sg db j (EConstr.Vars.subst1 a d) args)
| Prod (n,t,d) ->
- assert (List.is_empty args);
- let env' = push_rel_assum (n,t) env in
+ assert (List.is_empty args);
+ let env' = push_rel_assum (n,t) env in
(match flag_of_type env sg t with
- | (Info, Default) ->
+ | (Info, Default) ->
(* Standard case: two [extract_type] ... *)
let mld = extract_type env' sg (0::db) j d [] in
(match expand env mld with
@@ -296,7 +297,7 @@ let rec extract_type env sg db j c args =
(match EConstr.lookup_rel n env with
| LocalDef (_,t,_) ->
extract_type env sg db j (EConstr.Vars.lift n t) args
- | _ ->
+ | _ ->
(* Asks [db] a translation for [n]. *)
if n > List.length db then Tunknown
else let n' = List.nth db (n-1) in
@@ -497,8 +498,8 @@ and extract_really_ind env kn mib =
(* Now we're sure it's a record. *)
(* First, we find its field names. *)
let rec names_prod t = match Constr.kind t with
- | Prod(n,_,t) -> n::(names_prod t)
- | LetIn(_,_,_,t) -> names_prod t
+ | Prod(n,_,t) -> n::(names_prod t)
+ | LetIn(_,_,_,t) -> names_prod t
| Cast(t,_,_) -> names_prod t
| _ -> []
in
@@ -511,9 +512,9 @@ and extract_really_ind env kn mib =
| [],[] -> []
| _::l, typ::typs when isTdummy (expand env typ) ->
select_fields l typs
- | Anonymous::l, typ::typs ->
+ | {binder_name=Anonymous}::l, typ::typs ->
None :: (select_fields l typs)
- | Name id::l, typ::typs ->
+ | {binder_name=Name id}::l, typ::typs ->
let knp = Constant.make2 mp (Label.of_id id) in
(* Is it safe to use [id] for projections [foo.id] ? *)
if List.for_all ((==) Keep) (type2signature env typ)
@@ -556,8 +557,8 @@ and extract_really_ind env kn mib =
and extract_type_cons env sg db dbmap c i =
match EConstr.kind sg (whd_all env sg c) with
| Prod (n,t,d) ->
- let env' = push_rel_assum (n,t) env in
- let db' = (try Int.Map.find i dbmap with Not_found -> 0) :: db in
+ let env' = push_rel_assum (n,t) env in
+ let db' = (try Int.Map.find i dbmap with Not_found -> 0) :: db in
let l = extract_type_cons env' sg db' dbmap d (i+1) in
(extract_type env sg db 0 t []) :: l
| _ -> []
@@ -620,17 +621,18 @@ let rec extract_term env sg mle mlt c args =
| App (f,a) ->
extract_term env sg mle mlt f (Array.to_list a @ args)
| Lambda (n, t, d) ->
- let id = id_of_name n in
+ let id = map_annot id_of_name n in
+ let idna = map_annot Name.mk_name id in
(match args with
| a :: l ->
(* We make as many [LetIn] as possible. *)
let l' = List.map (EConstr.Vars.lift 1) l in
- let d' = EConstr.mkLetIn (Name id,a,t,applistc d l') in
+ let d' = EConstr.mkLetIn (idna,a,t,applistc d l') in
extract_term env sg mle mlt d' []
- | [] ->
- let env' = push_rel_assum (Name id, t) env in
+ | [] ->
+ let env' = push_rel_assum (idna, t) env in
let id, a =
- try check_default env sg t; Id id, new_meta()
+ try check_default env sg t; Id id.binder_name, new_meta()
with NotDefault d -> Dummy, Tdummy d
in
let b = new_meta () in
@@ -639,9 +641,9 @@ let rec extract_term env sg mle mlt c args =
let d' = extract_term env' sg (Mlenv.push_type mle a) b d [] in
put_magic_if magic (MLlam (id, d')))
| LetIn (n, c1, t1, c2) ->
- let id = id_of_name n in
- let env' = EConstr.push_rel (LocalDef (Name id, c1, t1)) env in
- (* We directly push the args inside the [LetIn].
+ let id = map_annot id_of_name n in
+ let env' = EConstr.push_rel (LocalDef (map_annot Name.mk_name id, c1, t1)) env in
+ (* We directly push the args inside the [LetIn].
TODO: the opt_let_app flag is supposed to prevent that *)
let args' = List.map (EConstr.Vars.lift 1) args in
(try
@@ -654,7 +656,7 @@ let rec extract_term env sg mle mlt c args =
then Mlenv.push_gen mle a
else Mlenv.push_type mle a
in
- MLletin (Id id, c1', extract_term env' sg mle' mlt c2 args')
+ MLletin (Id id.binder_name, c1', extract_term env' sg mle' mlt c2 args')
with NotDefault d ->
let mle' = Mlenv.push_std_type mle (Tdummy d) in
ast_pop (extract_term env' sg mle' mlt c2 args'))
@@ -918,7 +920,7 @@ and extract_fix env sg mle i (fi,ti,ci as recd) mlt =
metas.(i) <- mlt;
let mle = Array.fold_left Mlenv.push_type mle metas in
let ei = Array.map2 (extract_maybe_term env sg mle) metas ci in
- MLfix (i, Array.map id_of_name fi, ei)
+ MLfix (i, Array.map (fun x -> id_of_name x.binder_name) fi, ei)
(*S ML declarations. *)
@@ -994,7 +996,7 @@ let extract_std_constant env sg kn body typ =
(* The initial ML environment. *)
let mle = List.fold_left Mlenv.push_std_type Mlenv.empty l in
(* The lambdas names. *)
- let ids = List.map (fun (n,_) -> Id (id_of_name n)) rels in
+ let ids = List.map (fun (n,_) -> Id (id_of_name n.binder_name)) rels in
(* The according Coq environment. *)
let env = push_rels_assum rels env in
(* The real extraction: *)
@@ -1060,12 +1062,15 @@ let fake_match_projection env p =
ci_npar = mib.mind_nparams;
ci_cstr_ndecls = mip.mind_consnrealdecls;
ci_cstr_nargs = mip.mind_consnrealargs;
+ ci_relevance = Declareops.relevance_of_projection_repr mib p;
ci_pp_info;
}
in
let x = match mib.mind_record with
| NotRecord | FakeRecord -> assert false
- | PrimRecord info -> Name (pi1 info.(snd ind))
+ | PrimRecord info ->
+ let x, _, _, _ = info.(snd ind) in
+ make_annot (Name x) mip.mind_relevant
in
let indty = mkApp (indu, Context.Rel.to_extended_vect mkRel 0 paramslet) in
let rec fold arg j subst = function
@@ -1073,7 +1078,7 @@ let fake_match_projection env p =
| LocalAssum (na,ty) :: rem ->
let ty = Vars.substl subst (liftn 1 j ty) in
if arg != proj_arg then
- let lab = match na with Name id -> Label.of_id id | Anonymous -> assert false in
+ let lab = match na.binder_name with Name id -> Label.of_id id | Anonymous -> assert false in
let kn = Projection.Repr.make ind ~proj_npars:mib.mind_nparams ~proj_arg:arg lab in
fold (arg+1) (j+1) (mkProj (Projection.make kn false, mkRel 1)::subst) rem
else
diff --git a/plugins/extraction/table.ml b/plugins/extraction/table.ml
index 2058837b8e..399a77c596 100644
--- a/plugins/extraction/table.ml
+++ b/plugins/extraction/table.ml
@@ -449,11 +449,11 @@ let argnames_of_global r =
let typ, _ = Typeops.type_of_global_in_context env r in
let rels,_ =
decompose_prod (Reduction.whd_all env typ) in
- List.rev_map fst rels
+ List.rev_map (fun x -> Context.binder_name (fst x)) rels
let msg_of_implicit = function
| Kimplicit (r,i) ->
- let name = match List.nth (argnames_of_global r) (i-1) with
+ let name = match (List.nth (argnames_of_global r) (i-1)) with
| Anonymous -> ""
| Name id -> "(" ^ Id.to_string id ^ ") "
in
diff --git a/plugins/firstorder/instances.ml b/plugins/firstorder/instances.ml
index 286021d68e..1c9ab2e3bd 100644
--- a/plugins/firstorder/instances.ml
+++ b/plugins/firstorder/instances.ml
@@ -107,7 +107,7 @@ let mk_open_instance env evmap id idc m t =
(* since we know we will get a product,
reduction is not too expensive *)
let (nam,_,_)=destProd evmap (whd_all env evmap typ) in
- match nam with
+ match nam.Context.binder_name with
Name id -> id
| Anonymous -> dummy_bvid in
let revt=substl (List.init m (fun i->mkRel (m-i))) t in
@@ -115,7 +115,7 @@ let mk_open_instance env evmap id idc m t =
if Int.equal n 0 then evmap, decls else
let nid=(fresh_id_in_env avoid var_id env) in
let (evmap, (c, _)) = Evarutil.new_type_evar env evmap Evd.univ_flexible in
- let decl = LocalAssum (Name nid, c) in
+ let decl = LocalAssum (Context.make_annot (Name nid) Sorts.Relevant, c) in
aux (n-1) (Id.Set.add nid avoid) (EConstr.push_rel decl env) evmap (decl::decls) in
let evmap, decls = aux m Id.Set.empty env evmap [] in
(evmap, decls, revt)
diff --git a/plugins/firstorder/rules.ml b/plugins/firstorder/rules.ml
index 832a98b7f8..7f06ab6777 100644
--- a/plugins/firstorder/rules.ml
+++ b/plugins/firstorder/rules.ml
@@ -163,9 +163,9 @@ let ll_ind_tac (ind,u as indu) largs backtrack id continue seq =
let ll_arrow_tac a b c backtrack id continue seq=
let open EConstr in
let open Vars in
- let cc=mkProd(Anonymous,a,(lift 1 b)) in
- let d idc = mkLambda (Anonymous,b,
- mkApp (idc, [|mkLambda (Anonymous,(lift 1 a),(mkRel 2))|])) in
+ let cc=mkProd(Context.make_annot Anonymous Sorts.Relevant,a,(lift 1 b)) in
+ let d idc = mkLambda (Context.make_annot Anonymous Sorts.Relevant,b,
+ mkApp (idc, [|mkLambda (Context.make_annot Anonymous Sorts.Relevant,(lift 1 a),(mkRel 2))|])) in
tclORELSE
(tclTHENS (cut c)
[tclTHENLIST
diff --git a/plugins/firstorder/unify.ml b/plugins/firstorder/unify.ml
index d63fe9d799..0c958ddee3 100644
--- a/plugins/firstorder/unify.ml
+++ b/plugins/firstorder/unify.ml
@@ -65,13 +65,13 @@ let unif evd t1 t2=
bind i t else raise (UFAIL(nt1,nt2))
| Cast(_,_,_),_->Queue.add (strip_outer_cast evd nt1,nt2) bige
| _,Cast(_,_,_)->Queue.add (nt1,strip_outer_cast evd nt2) bige
- | (Prod(_,a,b),Prod(_,c,d))|(Lambda(_,a,b),Lambda(_,c,d))->
+ | (Prod(_,a,b),Prod(_,c,d))|(Lambda(_,a,b),Lambda(_,c,d))->
Queue.add (a,c) bige;Queue.add (pop b,pop d) bige
| Case (_,pa,ca,va),Case (_,pb,cb,vb)->
Queue.add (pa,pb) bige;
Queue.add (ca,cb) bige;
let l=Array.length va in
- if not (Int.equal l (Array.length vb)) then
+ if not (Int.equal l (Array.length vb)) then
raise (UFAIL (nt1,nt2))
else
for i=0 to l-1 do
diff --git a/plugins/funind/functional_principles_proofs.ml b/plugins/funind/functional_principles_proofs.ml
index 6fd2f7c2bc..34283c49c3 100644
--- a/plugins/funind/functional_principles_proofs.ml
+++ b/plugins/funind/functional_principles_proofs.ml
@@ -2,6 +2,7 @@ open Printer
open CErrors
open Util
open Constr
+open Context
open EConstr
open Vars
open Namegen
@@ -302,7 +303,7 @@ let change_eq env sigma hyp_id (context:rel_context) x t end_of_type =
in
let old_context_length = List.length context + 1 in
let witness_fun =
- mkLetIn(Anonymous,make_refl_eq constructor t1_typ (fst t1),t,
+ mkLetIn(make_annot Anonymous Sorts.Relevant,make_refl_eq constructor t1_typ (fst t1),t,
mkApp(mkVar hyp_id,Array.init old_context_length (fun i -> mkRel (old_context_length - i)))
)
in
@@ -312,7 +313,8 @@ let change_eq env sigma hyp_id (context:rel_context) x t end_of_type =
try
let witness = Int.Map.find i sub in
if is_local_def decl then anomaly (Pp.str "can not redefine a rel!");
- (pop end_of_type,ctxt_size,mkLetIn (RelDecl.get_name decl, witness, RelDecl.get_type decl, witness_fun))
+ (pop end_of_type,ctxt_size,mkLetIn (RelDecl.get_annot decl,
+ witness, RelDecl.get_type decl, witness_fun))
with Not_found ->
(mkProd_or_LetIn decl end_of_type, ctxt_size + 1, mkLambda_or_LetIn decl witness_fun)
)
@@ -428,7 +430,7 @@ let clean_hyp_with_heq ptes_infos eq_hyps hyp_id env sigma =
else if isProd sigma type_of_hyp
then
begin
- let (x,t_x,t') = destProd sigma type_of_hyp in
+ let (x,t_x,t') = destProd sigma type_of_hyp in
let actual_real_type_of_hyp = it_mkProd_or_LetIn t' context in
if is_property sigma ptes_infos t_x actual_real_type_of_hyp then
begin
@@ -541,7 +543,7 @@ let clean_hyp_with_heq ptes_infos eq_hyps hyp_id env sigma =
(scan_type new_context new_t')
with NoChange ->
(* Last thing todo : push the rel in the context and continue *)
- scan_type (LocalAssum (x,t_x) :: context) t'
+ scan_type (LocalAssum (x,t_x) :: context) t'
end
end
else
@@ -610,7 +612,7 @@ let treat_new_case ptes_infos nb_prod continue_tac term dyn_infos =
anomaly (Pp.str "cannot compute new term value.")
in
let fun_body =
- mkLambda(Anonymous,
+ mkLambda(make_annot Anonymous Sorts.Relevant,
pf_unsafe_type_of g' term,
Termops.replace_term (project g') term (mkRel 1) dyn_infos.info
)
@@ -736,7 +738,7 @@ let build_proof
g
in
build_proof do_finalize_t {dyn_infos with info = t} g
- | Lambda(n,t,b) ->
+ | Lambda(n,t,b) ->
begin
match EConstr.kind sigma (pf_concl g) with
| Prod _ ->
@@ -1149,7 +1151,7 @@ let prove_princ_for_struct (evd:Evd.evar_map ref) interactive_proof fun_num fnam
let fix_offset = List.length princ_params in
let ptes_to_fix,infos =
match EConstr.kind (project g) fbody_with_full_params with
- | Fix((idxs,i),(names,typess,bodies)) ->
+ | Fix((idxs,i),(names,typess,bodies)) ->
let bodies_with_all_params =
Array.map
(fun body ->
@@ -1164,7 +1166,7 @@ let prove_princ_for_struct (evd:Evd.evar_map ref) interactive_proof fun_num fnam
(fun i types ->
let types = prod_applist (project g) types (List.rev_map var_of_decl princ_params) in
{ idx = idxs.(i) - fix_offset;
- name = Nameops.Name.get_id (fresh_id names.(i));
+ name = Nameops.Name.get_id (fresh_id names.(i).binder_name);
types = types;
offset = fix_offset;
nb_realargs =
@@ -1195,9 +1197,9 @@ let prove_princ_for_struct (evd:Evd.evar_map ref) interactive_proof fun_num fnam
applist(body,List.rev_map var_of_decl full_params))
in
match EConstr.kind (project g) body_with_full_params with
- | Fix((_,num),(_,_,bs)) ->
+ | Fix((_,num),(_,_,bs)) ->
Reductionops.nf_betaiota (pf_env g) (project g)
- (
+ (
(applist
(substl
(List.rev
@@ -1514,7 +1516,7 @@ let is_valid_hypothesis sigma predicates_name =
let rec is_valid_hypothesis typ =
is_pte typ ||
match EConstr.kind sigma typ with
- | Prod(_,pte,typ') -> is_pte pte && is_valid_hypothesis typ'
+ | Prod(_,pte,typ') -> is_pte pte && is_valid_hypothesis typ'
| _ -> false
in
is_valid_hypothesis
@@ -1565,7 +1567,7 @@ let prove_principle_for_gen
in
let rec_arg_id =
match List.rev post_rec_arg with
- | (LocalAssum (Name id,_) | LocalDef (Name id,_,_)) :: _ -> id
+ | (LocalAssum ({binder_name=Name id},_) | LocalDef ({binder_name=Name id},_,_)) :: _ -> id
| _ -> assert false
in
(* observe (str "rec_arg_id := " ++ pr_lconstr (mkVar rec_arg_id)); *)
diff --git a/plugins/funind/functional_principles_types.ml b/plugins/funind/functional_principles_types.ml
index ca09cad1f3..1217ba0eba 100644
--- a/plugins/funind/functional_principles_types.ml
+++ b/plugins/funind/functional_principles_types.ml
@@ -14,6 +14,7 @@ open Term
open Sorts
open Util
open Constr
+open Context
open Vars
open Namegen
open Names
@@ -72,7 +73,7 @@ let compute_new_princ_type_from_rel rel_to_fun sorts princ_type =
then List.tl args
else args
in
- Context.Named.Declaration.LocalAssum (Nameops.Name.get_id (Context.Rel.Declaration.get_name decl),
+ Context.Named.Declaration.LocalAssum (map_annot Nameops.Name.get_id (Context.Rel.Declaration.get_annot decl),
Term.compose_prod real_args (mkSort new_sort))
in
let new_predicates =
@@ -137,14 +138,14 @@ let compute_new_princ_type_from_rel rel_to_fun sorts princ_type =
| Rel n ->
begin
try match Environ.lookup_rel n env with
- | LocalAssum (_,t) | LocalDef (_,_,t) when is_dom t -> raise Toberemoved
+ | LocalAssum (_,t) | LocalDef (_,_,t) when is_dom t -> raise Toberemoved
| _ -> pre_princ,[]
with Not_found -> assert false
end
- | Prod(x,t,b) ->
- compute_new_princ_type_for_binder remove mkProd env x t b
- | Lambda(x,t,b) ->
- compute_new_princ_type_for_binder remove mkLambda env x t b
+ | Prod(x,t,b) ->
+ compute_new_princ_type_for_binder remove mkProd env x t b
+ | Lambda(x,t,b) ->
+ compute_new_princ_type_for_binder remove mkLambda env x t b
| Ind _ | Construct _ when is_dom pre_princ -> raise Toberemoved
| App(f,args) when is_dom f ->
let var_to_be_removed = destRel (Array.last args) in
@@ -164,8 +165,8 @@ let compute_new_princ_type_from_rel rel_to_fun sorts princ_type =
let new_f,binders_to_remove_from_f = compute_new_princ_type remove env f in
applistc new_f new_args,
list_union_eq Constr.equal binders_to_remove_from_f binders_to_remove
- | LetIn(x,v,t,b) ->
- compute_new_princ_type_for_letin remove env x v t b
+ | LetIn(x,v,t,b) ->
+ compute_new_princ_type_for_letin remove env x v t b
| _ -> pre_princ,[]
in
(* let _ = match Constr.kind pre_princ with *)
@@ -181,14 +182,14 @@ let compute_new_princ_type_from_rel rel_to_fun sorts princ_type =
begin
try
let new_t,binders_to_remove_from_t = compute_new_princ_type remove env t in
- let new_x : Name.t = get_name (Termops.ids_of_context env) x in
- let new_env = Environ.push_rel (LocalAssum (x,t)) env in
+ let new_x = map_annot (get_name (Termops.ids_of_context env)) x in
+ let new_env = Environ.push_rel (LocalAssum (x,t)) env in
let new_b,binders_to_remove_from_b = compute_new_princ_type remove new_env b in
if List.exists (Constr.equal (mkRel 1)) binders_to_remove_from_b
then (pop new_b), filter_map (Constr.equal (mkRel 1)) pop binders_to_remove_from_b
else
(
- bind_fun(new_x,new_t,new_b),
+ bind_fun(new_x,new_t,new_b),
list_union_eq
Constr.equal
binders_to_remove_from_t
@@ -210,14 +211,14 @@ let compute_new_princ_type_from_rel rel_to_fun sorts princ_type =
try
let new_t,binders_to_remove_from_t = compute_new_princ_type remove env t in
let new_v,binders_to_remove_from_v = compute_new_princ_type remove env v in
- let new_x : Name.t = get_name (Termops.ids_of_context env) x in
- let new_env = Environ.push_rel (LocalDef (x,v,t)) env in
+ let new_x = map_annot (get_name (Termops.ids_of_context env)) x in
+ let new_env = Environ.push_rel (LocalDef (x,v,t)) env in
let new_b,binders_to_remove_from_b = compute_new_princ_type remove new_env b in
if List.exists (Constr.equal (mkRel 1)) binders_to_remove_from_b
then (pop new_b),filter_map (Constr.equal (mkRel 1)) pop binders_to_remove_from_b
else
(
- mkLetIn(new_x,new_v,new_t,new_b),
+ mkLetIn(new_x,new_v,new_t,new_b),
list_union_eq
Constr.equal
(list_union_eq Constr.equal binders_to_remove_from_t binders_to_remove_from_v)
@@ -250,8 +251,11 @@ let compute_new_princ_type_from_rel rel_to_fun sorts princ_type =
in
it_mkProd_or_LetIn
(it_mkProd_or_LetIn
- pre_res (List.map (function Context.Named.Declaration.LocalAssum (id,b) -> LocalAssum (Name (Hashtbl.find tbl id), b)
- | Context.Named.Declaration.LocalDef (id,t,b) -> LocalDef (Name (Hashtbl.find tbl id), t, b))
+ pre_res (List.map (function
+ | Context.Named.Declaration.LocalAssum (id,b) ->
+ LocalAssum (map_annot (fun id -> Name.mk_name (Hashtbl.find tbl id)) id, b)
+ | Context.Named.Declaration.LocalDef (id,t,b) ->
+ LocalDef (map_annot (fun id -> Name.mk_name (Hashtbl.find tbl id)) id, t, b))
new_predicates)
)
(List.map (fun d -> Termops.map_rel_decl EConstr.Unsafe.to_constr d) princ_type_info.params)
@@ -264,7 +268,7 @@ let change_property_sort evd toSort princ princName =
let princ_info = compute_elim_sig evd princ in
let change_sort_in_predicate decl =
LocalAssum
- (get_name decl,
+ (get_annot decl,
let args,ty = decompose_prod (EConstr.Unsafe.to_constr (get_type decl)) in
let s = destSort ty in
Global.add_constraints (Univ.enforce_leq (univ_of_sort toSort) (univ_of_sort s) Univ.Constraint.empty);
@@ -414,7 +418,7 @@ let get_funs_constant mp =
| Fix((_,(na,_,_))) ->
Array.mapi
(fun i na ->
- match na with
+ match na.binder_name with
| Name id ->
let const = Constant.make2 mp (Label.of_id id) in
const,i
@@ -451,7 +455,8 @@ let get_funs_constant mp =
let first_params = List.hd l_params in
List.iter
(fun params ->
- if not (List.equal (fun (n1, c1) (n2, c2) -> Name.equal n1 n2 && Constr.equal c1 c2) first_params params)
+ if not (List.equal (fun (n1, c1) (n2, c2) ->
+ eq_annot Name.equal n1 n2 && Constr.equal c1 c2) first_params params)
then user_err Pp.(str "Not a mutal recursive block")
)
l_params
@@ -461,7 +466,7 @@ let get_funs_constant mp =
try
let extract_info is_first body =
match Constr.kind body with
- | Fix((idxs,_),(na,ta,ca)) -> (idxs,na,ta,ca)
+ | Fix((idxs,_),(na,ta,ca)) -> (idxs,na,ta,ca)
| _ ->
if is_first && Int.equal (List.length l_bodies) 1
then raise Not_Rec
@@ -469,9 +474,9 @@ let get_funs_constant mp =
in
let first_infos = extract_info true (List.hd l_bodies) in
let check body = (* Hope this is correct *)
- let eq_infos (ia1, na1, ta1, ca1) (ia2, na2, ta2, ca2) =
- Array.equal Int.equal ia1 ia2 && Array.equal Name.equal na1 na2 &&
- Array.equal Constr.equal ta1 ta2 && Array.equal Constr.equal ca1 ca2
+ let eq_infos (ia1, na1, ta1, ca1) (ia2, na2, ta2, ca2) =
+ Array.equal Int.equal ia1 ia2 && Array.equal (eq_annot Name.equal) na1 na2 &&
+ Array.equal Constr.equal ta1 ta2 && Array.equal Constr.equal ca1 ca2
in
if not (eq_infos first_infos (extract_info false body))
then user_err Pp.(str "Not a mutal recursive block")
diff --git a/plugins/funind/glob_term_to_relation.ml b/plugins/funind/glob_term_to_relation.ml
index ba0a3bbb5c..8611dcaf83 100644
--- a/plugins/funind/glob_term_to_relation.ml
+++ b/plugins/funind/glob_term_to_relation.ml
@@ -2,6 +2,7 @@ open Printer
open Pp
open Names
open Constr
+open Context
open Vars
open Glob_term
open Glob_ops
@@ -343,12 +344,13 @@ let raw_push_named (na,raw_value,raw_typ) env =
match na with
| Anonymous -> env
| Name id ->
- let typ,_ = Pretyping.understand env (Evd.from_env env) ~expected_type:Pretyping.IsType raw_typ in
+ let typ,_ = Pretyping.understand env (Evd.from_env env) ~expected_type:Pretyping.IsType raw_typ in
+ let na = make_annot id Sorts.Relevant in (* TODO relevance *)
(match raw_value with
| None ->
- EConstr.push_named (NamedDecl.LocalAssum (id,typ)) env
+ EConstr.push_named (NamedDecl.LocalAssum (na,typ)) env
| Some value ->
- EConstr.push_named (NamedDecl.LocalDef (id, value, typ)) env)
+ EConstr.push_named (NamedDecl.LocalDef (na, value, typ)) env)
let add_pat_variables pat typ env : Environ.env =
@@ -356,7 +358,7 @@ let add_pat_variables pat typ env : Environ.env =
observe (str "new rel env := " ++ Printer.pr_rel_context_of env (Evd.from_env env));
match DAst.get pat with
- | PatVar na -> Environ.push_rel (RelDecl.LocalAssum (na,typ)) env
+ | PatVar na -> Environ.push_rel (RelDecl.LocalAssum (make_annot na Sorts.Relevant,typ)) env
| PatCstr(c,patl,na) ->
let Inductiveops.IndType(indf,indargs) =
try Inductiveops.find_rectype env (Evd.from_env env) (EConstr.of_constr typ)
@@ -375,16 +377,18 @@ let add_pat_variables pat typ env : Environ.env =
let open Context.Rel.Declaration in
let sigma, _ = Pfedit.get_current_context () in
match decl with
- | LocalAssum (Anonymous,_) | LocalDef (Anonymous,_,_) -> assert false
- | LocalAssum (Name id, t) ->
+ | LocalAssum ({binder_name=Anonymous},_) | LocalDef ({binder_name=Anonymous},_,_) -> assert false
+ | LocalAssum ({binder_name=Name id} as na, t) ->
+ let na = {na with binder_name=id} in
let new_t = substl ctxt t in
observe (str "for variable " ++ Ppconstr.pr_id id ++ fnl () ++
str "old type := " ++ Printer.pr_lconstr_env env sigma t ++ fnl () ++
str "new type := " ++ Printer.pr_lconstr_env env sigma new_t ++ fnl ()
);
let open Context.Named.Declaration in
- (Environ.push_named (LocalAssum (id,new_t)) env,mkVar id::ctxt)
- | LocalDef (Name id, v, t) ->
+ (Environ.push_named (LocalAssum (na,new_t)) env,mkVar id::ctxt)
+ | LocalDef ({binder_name=Name id} as na, v, t) ->
+ let na = {na with binder_name=id} in
let new_t = substl ctxt t in
let new_v = substl ctxt v in
observe (str "for variable " ++ Ppconstr.pr_id id ++ fnl () ++
@@ -394,7 +398,7 @@ let add_pat_variables pat typ env : Environ.env =
str "new value := " ++ Printer.pr_lconstr_env env sigma new_v ++ fnl ()
);
let open Context.Named.Declaration in
- (Environ.push_named (LocalDef (id,new_v,new_t)) env,mkVar id::ctxt)
+ (Environ.push_named (LocalDef (na,new_v,new_t)) env,mkVar id::ctxt)
)
(Environ.rel_context new_env)
~init:(env,[])
@@ -626,11 +630,12 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return =
let v_res = build_entry_lc env funnames avoid v in
let v_as_constr,ctx = Pretyping.understand env (Evd.from_env env) v in
let v_type = Typing.unsafe_type_of env (Evd.from_env env) v_as_constr in
+ let v_r = Sorts.Relevant in (* TODO relevance *)
let new_env =
match n with
Anonymous -> env
- | Name id -> EConstr.push_named (NamedDecl.LocalDef (id,v_as_constr,v_type)) env
- in
+ | Name id -> EConstr.push_named (NamedDecl.LocalDef (make_annot id v_r,v_as_constr,v_type)) env
+ in
let b_res = build_entry_lc new_env funnames avoid b in
combine_results (combine_letin n) v_res b_res
| GCases(_,_,el,brl) ->
@@ -939,9 +944,10 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
let new_t =
mkGApp(mkGVar(mk_rel_id this_relname),List.tl args'@[res_rt])
in
- let t',ctx = Pretyping.understand env (Evd.from_env env) new_t in
- let new_env = EConstr.push_rel (LocalAssum (n,t')) env in
- let new_b,id_to_exclude =
+ let t',ctx = Pretyping.understand env (Evd.from_env env) new_t in
+ let r = Sorts.Relevant in (* TODO relevance *)
+ let new_env = EConstr.push_rel (LocalAssum (make_annot n r,t')) env in
+ let new_b,id_to_exclude =
rebuild_cons new_env
nb_args relname
args new_crossed_types
@@ -974,9 +980,10 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
let new_args = List.map (replace_var_by_term id rt) args in
let subst_b =
if is_in_b then b else replace_var_by_term id rt b
- in
- let new_env = EConstr.push_rel (LocalAssum (n,t')) env in
- let new_b,id_to_exclude =
+ in
+ let r = Sorts.Relevant in (* TODO relevance *)
+ let new_env = EConstr.push_rel (LocalAssum (make_annot n r,t')) env in
+ let new_b,id_to_exclude =
rebuild_cons
new_env
nb_args relname
@@ -1057,8 +1064,9 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
in
let new_env =
let t',ctx = Pretyping.understand env (Evd.from_env env) eq' in
- EConstr.push_rel (LocalAssum (n,t')) env
- in
+ let r = Sorts.Relevant in (* TODO relevance *)
+ EConstr.push_rel (LocalAssum (make_annot n r,t')) env
+ in
let new_b,id_to_exclude =
rebuild_cons
new_env
@@ -1095,8 +1103,9 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
with Continue ->
observe (str "computing new type for prod : " ++ pr_glob_constr_env env rt);
let t',ctx = Pretyping.understand env (Evd.from_env env) t in
- let new_env = EConstr.push_rel (LocalAssum (n,t')) env in
- let new_b,id_to_exclude =
+ let r = Sorts.Relevant in (* TODO relevance *)
+ let new_env = EConstr.push_rel (LocalAssum (make_annot n r,t')) env in
+ let new_b,id_to_exclude =
rebuild_cons new_env
nb_args relname
args new_crossed_types
@@ -1111,8 +1120,9 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
| _ ->
observe (str "computing new type for prod : " ++ pr_glob_constr_env env rt);
let t',ctx = Pretyping.understand env (Evd.from_env env) t in
- let new_env = EConstr.push_rel (LocalAssum (n,t')) env in
- let new_b,id_to_exclude =
+ let r = Sorts.Relevant in (* TODO relevance *)
+ let new_env = EConstr.push_rel (LocalAssum (make_annot n r,t')) env in
+ let new_b,id_to_exclude =
rebuild_cons new_env
nb_args relname
args new_crossed_types
@@ -1132,8 +1142,9 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
let t',ctx = Pretyping.understand env (Evd.from_env env) t in
match n with
| Name id ->
- let new_env = EConstr.push_rel (LocalAssum (n,t')) env in
- let new_b,id_to_exclude =
+ let r = Sorts.Relevant in (* TODO relevance *)
+ let new_env = EConstr.push_rel (LocalAssum (make_annot n r,t')) env in
+ let new_b,id_to_exclude =
rebuild_cons new_env
nb_args relname
(args@[mkGVar id])new_crossed_types
@@ -1158,7 +1169,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
let type_t' = Typing.unsafe_type_of env evd t' in
let t' = EConstr.Unsafe.to_constr t' in
let type_t' = EConstr.Unsafe.to_constr type_t' in
- let new_env = Environ.push_rel (LocalDef (n,t',type_t')) env in
+ let new_env = Environ.push_rel (LocalDef (make_annot n Sorts.Relevant,t',type_t')) env in
let new_b,id_to_exclude =
rebuild_cons new_env
nb_args relname
@@ -1182,8 +1193,9 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
depth t
in
let t',ctx = Pretyping.understand env (Evd.from_env env) new_t in
- let new_env = EConstr.push_rel (LocalAssum (na,t')) env in
- let new_b,id_to_exclude =
+ let r = Sorts.Relevant in (* TODO relevance *)
+ let new_env = EConstr.push_rel (LocalAssum (make_annot na r,t')) env in
+ let new_b,id_to_exclude =
rebuild_cons new_env
nb_args relname
args (t::crossed_types)
@@ -1320,7 +1332,7 @@ let do_build_inductive
let evd,t = Typing.type_of env evd (EConstr.mkConstU (c, u)) in
let t = EConstr.Unsafe.to_constr t in
evd,
- Environ.push_named (LocalAssum (id,t))
+ Environ.push_named (LocalAssum (make_annot id Sorts.Relevant,t))
env
)
funnames
@@ -1364,7 +1376,8 @@ let do_build_inductive
Util.Array.fold_left2 (fun env rel_name rel_ar ->
let rex = fst (with_full_print (Constrintern.interp_constr env evd) rel_ar) in
let rex = EConstr.Unsafe.to_constr rex in
- Environ.push_named (LocalAssum (rel_name,rex)) env) env relnames rel_arities
+ let r = Sorts.Relevant in (* TODO relevance *)
+ Environ.push_named (LocalAssum (make_annot rel_name r,rex)) env) env relnames rel_arities
in
(* and of the real constructors*)
let constr i res =
diff --git a/plugins/funind/indfun.ml b/plugins/funind/indfun.ml
index 98d369aeda..b69ca7080c 100644
--- a/plugins/funind/indfun.ml
+++ b/plugins/funind/indfun.ml
@@ -3,6 +3,7 @@ open Sorts
open Util
open Names
open Constr
+open Context
open EConstr
open Pp
open Indfun_common
@@ -170,7 +171,8 @@ let build_newrecursive
let evd, (_, (_, impls')) = Constrintern.interp_context_evars ~program_mode:false env evd bl in
let impl = Constrintern.compute_internalization_data env0 evd Constrintern.Recursive arity impls' in
let open Context.Named.Declaration in
- (EConstr.push_named (LocalAssum (recname,arity)) env, Id.Map.add recname impl impls))
+ let r = Sorts.Relevant in (* TODO relevance *)
+ (EConstr.push_named (LocalAssum (make_annot recname r,arity)) env, Id.Map.add recname impl impls))
(env0,Constrintern.empty_internalization_env) lnameargsardef in
let recdef =
(* Declare local notations *)
@@ -622,8 +624,8 @@ let rebuild_bl aux bl typ = rebuild_bl aux bl typ
let recompute_binder_list (fixpoint_exprl : (Vernacexpr.fixpoint_expr * Vernacexpr.decl_notation list) list) =
let fixl,ntns = ComFixpoint.extract_fixpoint_components false fixpoint_exprl in
- let ((_,_,typel),_,ctx,_) = ComFixpoint.interp_fixpoint ~cofix:false fixl ntns in
- let constr_expr_typel =
+ let ((_,_,_,typel),_,ctx,_) = ComFixpoint.interp_fixpoint ~cofix:false fixl ntns in
+ let constr_expr_typel =
with_full_print (List.map (fun c -> Constrextern.extern_constr false (Global.env ()) (Evd.from_ctx ctx) (EConstr.of_constr c))) typel in
let fixpoint_exprl_with_new_bl =
List.map2 (fun ((lna,(rec_arg_opt,rec_order),bl,ret_typ,opt_body),notation_list) fix_typ ->
diff --git a/plugins/funind/indfun_common.mli b/plugins/funind/indfun_common.mli
index 102994f61e..4ec3131518 100644
--- a/plugins/funind/indfun_common.mli
+++ b/plugins/funind/indfun_common.mli
@@ -110,9 +110,9 @@ val evaluable_of_global_reference : GlobRef.t -> Names.evaluable_global_referenc
val list_rewrite : bool -> (EConstr.constr*bool) list -> Tacmach.tactic
val decompose_lam_n : Evd.evar_map -> int -> EConstr.t ->
- (Names.Name.t * EConstr.t) list * EConstr.t
-val compose_lam : (Names.Name.t * EConstr.t) list -> EConstr.t -> EConstr.t
-val compose_prod : (Names.Name.t * EConstr.t) list -> EConstr.t -> EConstr.t
+ (Names.Name.t Context.binder_annot * EConstr.t) list * EConstr.t
+val compose_lam : (Names.Name.t Context.binder_annot * EConstr.t) list -> EConstr.t -> EConstr.t
+val compose_prod : (Names.Name.t Context.binder_annot * EConstr.t) list -> EConstr.t -> EConstr.t
type tcc_lemma_value =
| Undefined
diff --git a/plugins/funind/invfun.ml b/plugins/funind/invfun.ml
index 95e2e9f6e5..37dbfec4c9 100644
--- a/plugins/funind/invfun.ml
+++ b/plugins/funind/invfun.ml
@@ -15,6 +15,7 @@ open Util
open Names
open Term
open Constr
+open Context
open EConstr
open Vars
open Pp
@@ -142,12 +143,13 @@ let generate_type evd g_to_f f graph i =
\[\forall (x_1:t_1)\ldots(x_n:t_n), let fv := f x_1\ldots x_n in, forall res, \]
i*)
let pre_ctxt =
- LocalAssum (Name res_id, lift 1 res_type) :: LocalDef (Name fv_id, mkApp (f,args_as_rels), res_type) :: fun_ctxt
+ LocalAssum (make_annot (Name res_id) Sorts.Relevant, lift 1 res_type) ::
+ LocalDef (make_annot (Name fv_id) Sorts.Relevant, mkApp (f,args_as_rels), res_type) :: fun_ctxt
in
(*i and we can return the solution depending on which lemma type we are defining i*)
if g_to_f
- then LocalAssum (Anonymous,graph_applied)::pre_ctxt,(lift 1 res_eq_f_of_args),graph
- else LocalAssum (Anonymous,res_eq_f_of_args)::pre_ctxt,(lift 1 graph_applied),graph
+ then LocalAssum (make_annot Anonymous Sorts.Relevant,graph_applied)::pre_ctxt,(lift 1 res_eq_f_of_args),graph
+ else LocalAssum (make_annot Anonymous Sorts.Relevant,res_eq_f_of_args)::pre_ctxt,(lift 1 graph_applied),graph
(*
@@ -270,10 +272,10 @@ let prove_fun_correct evd funs_constr graphs_constr schemes lemmas_types_infos i
let type_of_hid = pf_unsafe_type_of g (mkVar hid) in
let sigma = project g in
match EConstr.kind sigma type_of_hid with
- | Prod(_,_,t') ->
+ | Prod(_,_,t') ->
begin
match EConstr.kind sigma t' with
- | Prod(_,t'',t''') ->
+ | Prod(_,t'',t''') ->
begin
match EConstr.kind sigma t'',EConstr.kind sigma t''' with
| App(eq,args), App(graph',_)
@@ -358,17 +360,16 @@ let prove_fun_correct evd funs_constr graphs_constr schemes lemmas_types_infos i
(* end of branche proof *)
let lemmas =
Array.map
- (fun ((_,(ctxt,concl))) ->
- match ctxt with
- | [] | [_] | [_;_] -> anomaly (Pp.str "bad context.")
- | hres::res::decl::ctxt ->
- let res = EConstr.it_mkLambda_or_LetIn
- (EConstr.it_mkProd_or_LetIn concl [hres;res])
- (LocalAssum (RelDecl.get_name decl, RelDecl.get_type decl) :: ctxt)
- in
- res
- )
- lemmas_types_infos
+ (fun ((_,(ctxt,concl))) ->
+ match ctxt with
+ | [] | [_] | [_;_] -> anomaly (Pp.str "bad context.")
+ | hres::res::decl::ctxt ->
+ let res = EConstr.it_mkLambda_or_LetIn
+ (EConstr.it_mkProd_or_LetIn concl [hres;res])
+ (LocalAssum (RelDecl.get_annot decl, RelDecl.get_type decl) :: ctxt)
+ in
+ res)
+ lemmas_types_infos
in
let param_names = fst (List.chop princ_infos.nparams args_names) in
let params = List.map mkVar param_names in
@@ -429,7 +430,7 @@ let generalize_dependent_of x hyp g =
let open Context.Named.Declaration in
tclMAP
(function
- | LocalAssum (id,t) when not (Id.equal id hyp) &&
+ | LocalAssum ({binder_name=id},t) when not (Id.equal id hyp) &&
(Termops.occur_var (pf_env g) (project g) x t) -> tclTHEN (Proofview.V82.of_tactic (Tactics.generalize [mkVar id])) (thin [id])
| _ -> tclIDTAC
)
@@ -456,7 +457,7 @@ and intros_with_rewrite_aux : Tacmach.tactic =
let eq_ind = make_eq () in
let sigma = project g in
match EConstr.kind sigma (pf_concl g) with
- | Prod(_,t,t') ->
+ | Prod(_,t,t') ->
begin
match EConstr.kind sigma t with
| App(eq,args) when (EConstr.eq_constr sigma eq eq_ind) ->
diff --git a/plugins/funind/recdef.ml b/plugins/funind/recdef.ml
index 8746c37309..988cae8fbf 100644
--- a/plugins/funind/recdef.ml
+++ b/plugins/funind/recdef.ml
@@ -12,6 +12,7 @@
module CVars = Vars
open Constr
+open Context
open EConstr
open Vars
open Namegen
@@ -182,7 +183,7 @@ let (value_f: Constr.t list -> GlobRef.t -> Constr.t) =
)
in
let context = List.map
- (fun (x, c) -> LocalAssum (Name x, c)) (List.combine rev_x_id_l (List.rev al))
+ (fun (x, c) -> LocalAssum (make_annot (Name x) Sorts.Relevant, c)) (List.combine rev_x_id_l (List.rev al))
in
let env = Environ.push_rel_context context (Global.env ()) in
let glob_body =
@@ -388,9 +389,9 @@ let add_vars sigma forbidden e =
let treat_case forbid_new_ids to_intros finalize_tac nb_lam e infos : tactic =
fun g ->
let rev_context,b = decompose_lam_n (project g) nb_lam e in
- let ids = List.fold_left (fun acc (na,_) ->
+ let ids = List.fold_left (fun acc (na,_) ->
let pre_id =
- match na with
+ match na.binder_name with
| Name x -> x
| Anonymous -> ano_id
in
@@ -433,10 +434,10 @@ let rec travel_aux jinfo continuation_tac (expr_info:constr infos) g =
match EConstr.kind sigma expr_info.info with
| CoFix _ | Fix _ -> user_err Pp.(str "Function cannot treat local fixpoint or cofixpoint")
| Proj _ -> user_err Pp.(str "Function cannot treat projections")
- | LetIn(na,b,t,e) ->
+ | LetIn(na,b,t,e) ->
begin
let new_continuation_tac =
- jinfo.letiN (na,b,t,e) expr_info continuation_tac
+ jinfo.letiN (na.binder_name,b,t,e) expr_info continuation_tac
in
travel jinfo new_continuation_tac
{expr_info with info = b; is_final=false} g
@@ -450,7 +451,7 @@ let rec travel_aux jinfo continuation_tac (expr_info:constr infos) g =
with e when CErrors.noncritical e ->
user_err ~hdr:"Recdef.travel" (str "the term " ++ Printer.pr_leconstr_env (pf_env g) sigma expr_info.info ++ str " can not contain a recursive call to " ++ Id.print expr_info.f_id)
end
- | Lambda(n,t,b) ->
+ | Lambda(n,t,b) ->
begin
try
check_not_nested sigma (expr_info.f_id::expr_info.forbidden_ids) expr_info.info;
@@ -853,8 +854,8 @@ let rec prove_le g =
EConstr.is_global sigma (le ()) c
| _ -> false
in
- let (h,t) = List.find (fun (_,t) -> matching_fun t) (pf_hyps_types g)
- in
+ let (h,t) = List.find (fun (_,t) -> matching_fun t) (pf_hyps_types g) in
+ let h = h.binder_name in
let y =
let _,args = decompose_app sigma t in
List.hd (List.tl args)
@@ -877,10 +878,10 @@ let rec make_rewrite_list expr_info max = function
let sigma = project g in
let t_eq = compute_renamed_type g (mkVar hp) in
let k,def =
- let k_na,_,t = destProd sigma t_eq in
- let _,_,t = destProd sigma t in
- let def_na,_,_ = destProd sigma t in
- Nameops.Name.get_id k_na,Nameops.Name.get_id def_na
+ let k_na,_,t = destProd sigma t_eq in
+ let _,_,t = destProd sigma t in
+ let def_na,_,_ = destProd sigma t in
+ Nameops.Name.get_id k_na.binder_name,Nameops.Name.get_id def_na.binder_name
in
Proofview.V82.of_tactic (general_rewrite_bindings false Locus.AllOccurrences
true (* dep proofs also: *) true
@@ -903,10 +904,10 @@ let make_rewrite expr_info l hp max =
let sigma = project g in
let t_eq = compute_renamed_type g (mkVar hp) in
let k,def =
- let k_na,_,t = destProd sigma t_eq in
- let _,_,t = destProd sigma t in
- let def_na,_,_ = destProd sigma t in
- Nameops.Name.get_id k_na,Nameops.Name.get_id def_na
+ let k_na,_,t = destProd sigma t_eq in
+ let _,_,t = destProd sigma t in
+ let def_na,_,_ = destProd sigma t in
+ Nameops.Name.get_id k_na.binder_name,Nameops.Name.get_id def_na.binder_name
in
observe_tac (str "general_rewrite_bindings")
(Proofview.V82.of_tactic (general_rewrite_bindings false Locus.AllOccurrences
@@ -1054,20 +1055,19 @@ let compute_terminate_type nb_args func =
let right = mkRel 5 in
let delayed_force c = EConstr.Unsafe.to_constr (delayed_force c) in
let equality = mkApp(delayed_force eq, [|lift 5 b; left; right|]) in
- let result = (mkProd ((Name def_id) , lift 4 a_arrow_b, equality)) in
+ let result = (mkProd (make_annot (Name def_id) Sorts.Relevant, lift 4 a_arrow_b, equality)) in
let cond = mkApp(delayed_force lt, [|(mkRel 2); (mkRel 1)|]) in
let nb_iter =
mkApp(delayed_force ex,
[|delayed_force nat;
(mkLambda
- (Name
- p_id,
+ (make_annot (Name p_id) Sorts.Relevant,
delayed_force nat,
- (mkProd (Name k_id, delayed_force nat,
- mkArrow cond result))))|])in
+ (mkProd (make_annot (Name k_id) Sorts.Relevant, delayed_force nat,
+ mkArrow cond Sorts.Relevant result))))|])in
let value = mkApp(constr_of_global (Util.delayed_force coq_sig_ref),
[|b;
- (mkLambda (Name v_id, b, nb_iter))|]) in
+ (mkLambda (make_annot (Name v_id) Sorts.Relevant, b, nb_iter))|]) in
compose_prod rev_args value
@@ -1165,15 +1165,15 @@ let whole_start (concl_tac:tactic) nb_args is_mes func input_type relation rec_a
let func_body = EConstr.of_constr func_body in
let (f_name, _, body1) = destLambda sigma func_body in
let f_id =
- match f_name with
+ match f_name.binder_name with
| Name f_id -> next_ident_away_in_goal f_id ids
| Anonymous -> anomaly (Pp.str "Anonymous function.")
in
let n_names_types,_ = decompose_lam_n sigma nb_args body1 in
let n_ids,ids =
List.fold_left
- (fun (n_ids,ids) (n_name,_) ->
- match n_name with
+ (fun (n_ids,ids) (n_name,_) ->
+ match n_name.binder_name with
| Name id ->
let n_id = next_ident_away_in_goal id ids in
n_id::n_ids,n_id::ids
@@ -1270,12 +1270,12 @@ let is_rec_res id =
let clear_goals sigma =
let rec clear_goal t =
match EConstr.kind sigma t with
- | Prod(Name id as na,t',b) ->
+ | Prod({binder_name=Name id} as na,t',b) ->
let b' = clear_goal b in
if noccurn sigma 1 b' && (is_rec_res id)
then Vars.lift (-1) b'
else if b' == b then t
- else mkProd(na,t',b')
+ else mkProd(na,t',b')
| _ -> EConstr.map sigma clear_goal t
in
List.map clear_goal
@@ -1519,7 +1519,8 @@ let recursive_definition is_mes function_name rec_impls type_of_f r rec_arg_num
let env = Global.env() in
let evd = Evd.from_env env in
let evd, function_type = interp_type_evars ~program_mode:false env evd type_of_f in
- let env = EConstr.push_named (Context.Named.Declaration.LocalAssum (function_name,function_type)) env in
+ let function_r = Sorts.Relevant in (* TODO relevance *)
+ let env = EConstr.push_named (Context.Named.Declaration.LocalAssum (make_annot function_name function_r,function_type)) env in
(* Pp.msgnl (str "function type := " ++ Printer.pr_lconstr function_type); *)
let evd, ty = interp_type_evars ~program_mode:false env evd ~impls:rec_impls eq in
let evd = Evd.minimize_universes evd in
@@ -1537,7 +1538,7 @@ let recursive_definition is_mes function_name rec_impls type_of_f r rec_arg_num
(* Pp.msgnl (str "eq' := " ++ str (string_of_int rec_arg_num)); *)
match Constr.kind eq' with
| App(e,[|_;_;eq_fix|]) ->
- mkLambda (Name function_name,function_type,subst_var function_name (compose_lam res_vars eq_fix))
+ mkLambda (make_annot (Name function_name) Sorts.Relevant,function_type,subst_var function_name (compose_lam res_vars eq_fix))
| _ -> failwith "Recursive Definition (res not eq)"
in
let pre_rec_args,function_type_before_rec_arg = decompose_prod_n (rec_arg_num - 1) function_type in
diff --git a/plugins/ltac/evar_tactics.ml b/plugins/ltac/evar_tactics.ml
index b0277e9cc2..050fdcb608 100644
--- a/plugins/ltac/evar_tactics.ml
+++ b/plugins/ltac/evar_tactics.ml
@@ -11,6 +11,7 @@
open Util
open Names
open Constr
+open Context
open CErrors
open Evar_refiner
open Tacmach
@@ -62,7 +63,7 @@ let instantiate_tac n c ido =
evar_list sigma (EConstr.of_constr (NamedDecl.get_type decl))
| InHypValueOnly ->
(match decl with
- | LocalDef (_,body,_) -> evar_list sigma (EConstr.of_constr body)
+ | LocalDef (_,body,_) -> evar_list sigma (EConstr.of_constr body)
| _ -> user_err Pp.(str "Not a defined hypothesis.")) in
if List.length evl < n then
user_err Pp.(str "Not enough uninstantiated existential variables.");
@@ -108,5 +109,6 @@ let hget_evar n =
if n <= 0 then user_err Pp.(str "Incorrect existential variable index.");
let ev = List.nth evl (n-1) in
let ev_type = EConstr.existential_type sigma ev in
- Tactics.change_concl (mkLetIn (Name.Anonymous,mkEvar ev,ev_type,concl))
+ let r = Retyping.relevance_of_type (Proofview.Goal.env gl) sigma ev_type in
+ Tactics.change_concl (mkLetIn (make_annot Name.Anonymous r,mkEvar ev,ev_type,concl))
end
diff --git a/plugins/ltac/extratactics.mlg b/plugins/ltac/extratactics.mlg
index ffd8b71e5d..0428f08138 100644
--- a/plugins/ltac/extratactics.mlg
+++ b/plugins/ltac/extratactics.mlg
@@ -12,6 +12,7 @@
open Pp
open Constr
+open Context
open Genarg
open Stdarg
open Tacarg
@@ -674,7 +675,7 @@ let hResolve id c occ t =
let sigma = Evd.merge_universe_context sigma ctx in
let t_constr_type = Retyping.get_type_of env sigma t_constr in
Proofview.tclTHEN (Proofview.Unsafe.tclEVARS sigma)
- (change_concl (mkLetIn (Name.Anonymous,t_constr,t_constr_type,concl)))
+ (change_concl (mkLetIn (make_annot Name.Anonymous Sorts.Relevant,t_constr,t_constr_type,concl)))
end
let hResolve_auto id c t =
diff --git a/plugins/ltac/pptactic.ml b/plugins/ltac/pptactic.ml
index 5e3f4df192..e188971f00 100644
--- a/plugins/ltac/pptactic.ml
+++ b/plugins/ltac/pptactic.ml
@@ -1158,7 +1158,7 @@ let pr_goal_selector ~toplevel s =
if n=0 then (List.rev acc, EConstr.of_constr ty) else
match Constr.kind ty with
| Constr.Prod(na,a,b) ->
- strip_ty (([CAst.make na],EConstr.of_constr a)::acc) (n-1) b
+ strip_ty (([CAst.make na.Context.binder_name],EConstr.of_constr a)::acc) (n-1) b
| _ -> user_err Pp.(str "Cannot translate fix tactic: not enough products") in
strip_ty [] n ty
diff --git a/plugins/ltac/rewrite.ml b/plugins/ltac/rewrite.ml
index e78d0f93a4..b1d5c0252f 100644
--- a/plugins/ltac/rewrite.ml
+++ b/plugins/ltac/rewrite.ml
@@ -15,6 +15,7 @@ open Names
open Nameops
open Namegen
open Constr
+open Context
open EConstr
open Vars
open Reduction
@@ -220,23 +221,23 @@ end) = struct
let rec aux env evars ty l =
let t = Reductionops.whd_all env (goalevars evars) ty in
match EConstr.kind (goalevars evars) t, l with
- | Prod (na, ty, b), obj :: cstrs ->
+ | Prod (na, ty, b), obj :: cstrs ->
let b = Reductionops.nf_betaiota env (goalevars evars) b in
- if noccurn (goalevars evars) 1 b (* non-dependent product *) then
+ if noccurn (goalevars evars) 1 b (* non-dependent product *) then
let ty = Reductionops.nf_betaiota env (goalevars evars) ty in
let (evars, b', arg, cstrs) = aux env evars (subst1 mkProp b) cstrs in
let evars, relty = mk_relty evars env ty obj in
let evars, newarg = app_poly env evars respectful [| ty ; b' ; relty ; arg |] in
- evars, mkProd(na, ty, b), newarg, (ty, Some relty) :: cstrs
+ evars, mkProd(na, ty, b), newarg, (ty, Some relty) :: cstrs
else
let (evars, b, arg, cstrs) =
- aux (push_rel (LocalAssum (na, ty)) env) evars b cstrs
+ aux (push_rel (LocalAssum (na, ty)) env) evars b cstrs
in
let ty = Reductionops.nf_betaiota env (goalevars evars) ty in
- let pred = mkLambda (na, ty, b) in
- let liftarg = mkLambda (na, ty, arg) in
- let evars, arg' = app_poly env evars forall_relation [| ty ; pred ; liftarg |] in
- if Option.is_empty obj then evars, mkProd(na, ty, b), arg', (ty, None) :: cstrs
+ let pred = mkLambda (na, ty, b) in
+ let liftarg = mkLambda (na, ty, arg) in
+ let evars, arg' = app_poly env evars forall_relation [| ty ; pred ; liftarg |] in
+ if Option.is_empty obj then evars, mkProd(na, ty, b), arg', (ty, None) :: cstrs
else user_err Pp.(str "build_signature: no constraint can apply on a dependent argument")
| _, obj :: _ -> anomaly ~label:"build_signature" (Pp.str "not enough products.")
| _, [] ->
@@ -253,7 +254,7 @@ end) = struct
let unfold_impl sigma t =
match EConstr.kind sigma t with
| App (arrow, [| a; b |])(* when eq_constr arrow (Lazy.force impl) *) ->
- mkProd (Anonymous, a, lift 1 b)
+ mkProd (make_annot Anonymous Sorts.Relevant, a, lift 1 b)
| _ -> assert false
let unfold_all sigma t =
@@ -279,7 +280,7 @@ end) = struct
(app_poly env evd arrow [| a; b |]), unfold_impl
(* (evd, mkProd (Anonymous, a, b)), (fun x -> x) *)
else if bp then (* Dummy forall *)
- (app_poly env evd coq_all [| a; mkLambda (Anonymous, a, lift 1 b) |]), unfold_forall
+ (app_poly env evd coq_all [| a; mkLambda (make_annot Anonymous Sorts.Relevant, a, lift 1 b) |]), unfold_forall
else (* None in Prop, use arrow *)
(app_poly env evd arrow [| a; b |]), unfold_impl
@@ -308,7 +309,8 @@ end) = struct
app_poly env evd pointwise_relation [| t; lift (-1) car; lift (-1) rel |]
else
app_poly env evd forall_relation
- [| t; mkLambda (n, t, car); mkLambda (n, t, rel) |]
+ [| t; mkLambda (make_annot n Sorts.Relevant, t, car);
+ mkLambda (make_annot n Sorts.Relevant, t, rel) |]
let lift_cstr env evars (args : constr list) c ty cstr =
let start evars env car =
@@ -323,15 +325,15 @@ end) = struct
else
let sigma = goalevars evars in
match EConstr.kind sigma (Reductionops.whd_all env sigma prod) with
- | Prod (na, ty, b) ->
+ | Prod (na, ty, b) ->
if noccurn sigma 1 b then
let b' = lift (-1) b in
let evars, rb = aux evars env b' (pred n) in
app_poly env evars pointwise_relation [| ty; b'; rb |]
else
- let evars, rb = aux evars (push_rel (LocalAssum (na, ty)) env) b (pred n) in
+ let evars, rb = aux evars (push_rel (LocalAssum (na, ty)) env) b (pred n) in
app_poly env evars forall_relation
- [| ty; mkLambda (na, ty, b); mkLambda (na, ty, rb) |]
+ [| ty; mkLambda (na, ty, b); mkLambda (na, ty, rb) |]
| _ -> raise Not_found
in
let rec find env c ty = function
@@ -481,8 +483,9 @@ let rec decompose_app_rel env evd t =
| App (f, [|arg|]) ->
let (f', argl, argr) = decompose_app_rel env evd arg in
let ty = Typing.unsafe_type_of env evd argl in
- let f'' = mkLambda (Name default_dependent_ident, ty,
- mkLambda (Name (Id.of_string "y"), lift 1 ty,
+ let r = Retyping.relevance_of_type env evd ty in
+ let f'' = mkLambda (make_annot (Name default_dependent_ident) r, ty,
+ mkLambda (make_annot (Name (Id.of_string "y")) r, lift 1 ty,
mkApp (lift 2 f, [| mkApp (lift 2 f', [| mkRel 2; mkRel 1 |]) |])))
in (f'', argl, argr)
| App (f, args) ->
@@ -522,7 +525,7 @@ let decompose_applied_relation env sigma (c,l) =
| Some c -> c
| None ->
let ctx,t' = Reductionops.splay_prod env sigma ctype in (* Search for underlying eq *)
- match find_rel (it_mkProd_or_LetIn t' (List.map (fun (n,t) -> LocalAssum (n, t)) ctx)) with
+ match find_rel (it_mkProd_or_LetIn t' (List.map (fun (n,t) -> LocalAssum (n, t)) ctx)) with
| Some c -> c
| None -> user_err Pp.(str "Cannot find an homogeneous relation to rewrite.")
@@ -803,7 +806,7 @@ let resolve_morphism env avoid oldt m ?(fnewt=fun x -> x) args args' (b,cstr) ev
else TypeGlobal.do_subrelation, TypeGlobal.apply_subrelation
in
EConstr.push_named
- (LocalDef (Id.of_string "do_subrelation",
+ (LocalDef (make_annot (Id.of_string "do_subrelation") Sorts.Relevant,
snd (app_poly_sort b env evars dosub [||]),
snd (app_poly_nocheck env evars appsub [||])))
env
@@ -906,7 +909,7 @@ let make_leibniz_proof env c ty r =
let prf =
e_app_poly env evars coq_f_equal
[| r.rew_car; ty;
- mkLambda (Anonymous, r.rew_car, c);
+ mkLambda (make_annot Anonymous Sorts.Relevant, r.rew_car, c);
r.rew_from; r.rew_to; prf |]
in RewPrf (rel, prf)
| RewCast k -> r.rew_prf
@@ -1103,7 +1106,7 @@ let subterm all flags (s : 'a pure_strategy) : 'a pure_strategy =
(* else *)
| Prod (n, dom, codom) ->
- let lam = mkLambda (n, dom, codom) in
+ let lam = mkLambda (n, dom, codom) in
let (evars', app), unfold =
if eq_constr (fst evars) ty mkProp then
(app_poly_sort prop env evars coq_all [| dom; lam |]), TypeGlobal.unfold_all
@@ -1149,9 +1152,9 @@ let subterm all flags (s : 'a pure_strategy) : 'a pure_strategy =
(* | _ -> b') *)
| Lambda (n, t, b) when flags.under_lambdas ->
- let n' = Nameops.Name.map (fun id -> Tactics.fresh_id_in_env unfresh id env) n in
+ let n' = map_annot (Nameops.Name.map (fun id -> Tactics.fresh_id_in_env unfresh id env)) n in
let open Context.Rel.Declaration in
- let env' = EConstr.push_rel (LocalAssum (n', t)) env in
+ let env' = EConstr.push_rel (LocalAssum (n', t)) env in
let bty = Retyping.get_type_of env' (goalevars evars) b in
let unlift = if prop then PropGlobal.unlift_cstr else TypeGlobal.unlift_cstr in
let state, b' = s.strategy { state ; env = env' ; unfresh ;
@@ -1166,15 +1169,15 @@ let subterm all flags (s : 'a pure_strategy) : 'a pure_strategy =
let point = if prop then PropGlobal.pointwise_or_dep_relation else
TypeGlobal.pointwise_or_dep_relation
in
- let evars, rel = point env r.rew_evars n' t r.rew_car rel in
- let prf = mkLambda (n', t, prf) in
+ let evars, rel = point env r.rew_evars n'.binder_name t r.rew_car rel in
+ let prf = mkLambda (n', t, prf) in
{ r with rew_prf = RewPrf (rel, prf); rew_evars = evars }
| x -> r
in
Success { r with
- rew_car = mkProd (n, t, r.rew_car);
- rew_from = mkLambda(n, t, r.rew_from);
- rew_to = mkLambda (n, t, r.rew_to) }
+ rew_car = mkProd (n, t, r.rew_car);
+ rew_from = mkLambda(n, t, r.rew_from);
+ rew_to = mkLambda (n, t, r.rew_to) }
| Fail | Identity -> b'
in state, res
@@ -1516,7 +1519,7 @@ let cl_rewrite_clause_aux ?(abs=None) strat env avoid sigma concl is_hyp : resul
| Some (t, ty) ->
let t = Reductionops.nf_evar evars' t in
let ty = Reductionops.nf_evar evars' ty in
- mkApp (mkLambda (Name (Id.of_string "lemma"), ty, p), [| t |])
+ mkApp (mkLambda (make_annot (Name (Id.of_string "lemma")) Sorts.Relevant, ty, p), [| t |])
in
let proof = match is_hyp with
| None -> term
@@ -1542,7 +1545,8 @@ let assert_replacing id newt tac =
let after, before = List.split_when (NamedDecl.get_id %> Id.equal id) ctx in
let nc = match before with
| [] -> assert false
- | d :: rem -> insert_dependent env sigma (LocalAssum (NamedDecl.get_id d, newt)) [] after @ rem
+ | d :: rem -> insert_dependent env sigma
+ (LocalAssum (make_annot (NamedDecl.get_id d) Sorts.Relevant, newt)) [] after @ rem
in
let env' = Environ.reset_with_named_context (val_of_named_context nc) env in
Refine.refine ~typecheck:true begin fun sigma ->
@@ -1586,7 +1590,7 @@ let cl_rewrite_clause_newtac ?abs ?origsigma ~progress strat clause =
tclTHENFIRST (assert_replacing id newt tac) (beta_hyp id)
| Some id, None ->
Proofview.Unsafe.tclEVARS undef <*>
- convert_hyp_no_check (LocalAssum (id, newt)) <*>
+ convert_hyp_no_check (LocalAssum (make_annot id Sorts.Relevant, newt)) <*>
beta_hyp id
| None, Some p ->
Proofview.Unsafe.tclEVARS undef <*>
@@ -1905,7 +1909,7 @@ let build_morphism_signature env sigma m =
let cstrs =
let rec aux t =
match EConstr.kind sigma t with
- | Prod (na, a, b) ->
+ | Prod (na, a, b) ->
None :: aux b
| _ -> []
in aux t
diff --git a/plugins/ltac/tactic_matching.ml b/plugins/ltac/tactic_matching.ml
index 54924f1644..2b5e496168 100644
--- a/plugins/ltac/tactic_matching.ml
+++ b/plugins/ltac/tactic_matching.ml
@@ -12,6 +12,7 @@
(lazy)match and (lazy)match goal. *)
open Names
+open Context
open Tacexpr
open Context.Named.Declaration
@@ -299,8 +300,8 @@ module PatternMatching (E:StaticEnvironment) = struct
| LocalDef (id,body,hyp) ->
pattern_match_term false bodypat body () <*>
pattern_match_term true typepat hyp () <*>
- put_terms (id_map_try_add_name hypname (EConstr.mkVar id) empty_term_subst) <*>
- return id
+ put_terms (id_map_try_add_name hypname (EConstr.mkVar id.binder_name) empty_term_subst) <*>
+ return id.binder_name
| LocalAssum (id,hyp) -> fail
(** [hyp_match pat hyps] dispatches to
diff --git a/plugins/ltac/tauto.ml b/plugins/ltac/tauto.ml
index 19256e054d..4c65445b89 100644
--- a/plugins/ltac/tauto.ml
+++ b/plugins/ltac/tauto.ml
@@ -142,7 +142,7 @@ let flatten_contravariant_conj _ ist =
~onlybinary:flags.binary_mode typ
with
| Some (_,args) ->
- let newtyp = List.fold_right mkArrow args c in
+ let newtyp = List.fold_right (fun a b -> mkArrow a Sorts.Relevant b) args c in
let intros = tclMAP (fun _ -> intro) args in
let by = tclTHENLIST [intros; apply hyp; split; assumption] in
tclTHENLIST [assert_ ~by newtyp; clear (destVar sigma hyp)]
@@ -173,7 +173,7 @@ let flatten_contravariant_disj _ ist =
typ with
| Some (_,args) ->
let map i arg =
- let typ = mkArrow arg c in
+ let typ = mkArrow arg Sorts.Relevant c in
let ci = Tactics.constructor_tac false None (succ i) Tactypes.NoBindings in
let by = tclTHENLIST [intro; apply hyp; ci; assumption] in
assert_ ~by typ
diff --git a/plugins/micromega/coq_micromega.ml b/plugins/micromega/coq_micromega.ml
index 7adae148bd..ac34faa7da 100644
--- a/plugins/micromega/coq_micromega.ml
+++ b/plugins/micromega/coq_micromega.ml
@@ -23,6 +23,7 @@ open Names
open Goptions
open Mutils
open Constr
+open Context
open Tactypes
(**
@@ -1243,7 +1244,7 @@ let dump_rexpr = lazy
let prodn n env b =
let rec prodrec = function
| (0, env, b) -> b
- | (n, ((v,t)::l), b) -> prodrec (n-1, l, EConstr.mkProd (v,t,b))
+ | (n, ((v,t)::l), b) -> prodrec (n-1, l, EConstr.mkProd (make_annot v Sorts.Relevant,t,b))
| _ -> assert false
in
prodrec (n,env,b)
@@ -1293,8 +1294,8 @@ let make_goal_of_formula sigma dexpr form =
| FF -> Lazy.force coq_False
| C(x,y) -> EConstr.mkApp(Lazy.force coq_and,[|xdump pi xi x ; xdump pi xi y|])
| D(x,y) -> EConstr.mkApp(Lazy.force coq_or,[| xdump pi xi x ; xdump pi xi y|])
- | I(x,_,y) -> EConstr.mkArrow (xdump pi xi x) (xdump (pi+1) (xi+1) y)
- | N(x) -> EConstr.mkArrow (xdump pi xi x) (Lazy.force coq_False)
+ | I(x,_,y) -> EConstr.mkArrow (xdump pi xi x) Sorts.Relevant (xdump (pi+1) (xi+1) y)
+ | N(x) -> EConstr.mkArrow (xdump pi xi x) Sorts.Relevant (Lazy.force coq_False)
| A(x,_,_) -> dump_cstr xi x
| X(t) -> let idx = Env.get_rank props sigma t in
EConstr.mkRel (pi+idx) in
@@ -1327,7 +1328,7 @@ let make_goal_of_formula sigma dexpr form =
| (e::l) ->
let (name,expr,typ) = e in
xset (EConstr.mkNamedLetIn
- (Names.Id.of_string name)
+ (make_annot (Names.Id.of_string name) Sorts.Relevant)
expr typ acc) l in
xset concl l
@@ -1614,7 +1615,7 @@ let abstract_formula hyps f =
| I(f1,hyp,f2) ->
(match xabs f1 , hyp, xabs f2 with
| X a1 , Some _ , af2 -> af2
- | X a1 , None , X a2 -> X (EConstr.mkArrow a1 a2)
+ | X a1 , None , X a2 -> X (EConstr.mkArrow a1 Sorts.Relevant a2)
| af1 , _ , af2 -> I(af1,hyp,af2)
)
| FF -> FF
diff --git a/plugins/omega/coq_omega.ml b/plugins/omega/coq_omega.ml
index dff25b3a42..4802608fda 100644
--- a/plugins/omega/coq_omega.ml
+++ b/plugins/omega/coq_omega.ml
@@ -19,6 +19,7 @@ open CErrors
open Util
open Names
open Constr
+open Context
open Nameops
open EConstr
open Tacticals.New
@@ -431,8 +432,8 @@ let destructurate_prop sigma t =
| Ind (isp,_), args ->
Kapp (Other (string_of_path (path_of_global (IndRef isp))),args)
| Var id,[] -> Kvar id
- | Prod (Anonymous,typ,body), [] -> Kimp(typ,body)
- | Prod (Name _,_,_),[] -> CErrors.user_err Pp.(str "Omega: Not a quantifier-free goal")
+ | Prod ({binder_name=Anonymous},typ,body), [] -> Kimp(typ,body)
+ | Prod ({binder_name=Name _},_,_),[] -> CErrors.user_err Pp.(str "Omega: Not a quantifier-free goal")
| _ -> Kufo
let nf = Tacred.simpl
@@ -499,13 +500,13 @@ let context sigma operation path (t : constr) =
| (p, Fix ((_,n as ln),(tys,lna,v))) ->
let l = Array.length v in
let v' = Array.copy v in
- v'.(n)<- loop (Pervasives.(+) i l) p v.(n); (mkFix (ln,(tys,lna,v')))
+ v'.(n)<- loop (Pervasives.(+) i l) p v.(n); (mkFix (ln,(tys,lna,v')))
| ((P_TYPE :: p), Prod (n,t,c)) ->
- (mkProd (n,loop i p t,c))
+ (mkProd (n,loop i p t,c))
| ((P_TYPE :: p), Lambda (n,t,c)) ->
- (mkLambda (n,loop i p t,c))
+ (mkLambda (n,loop i p t,c))
| ((P_TYPE :: p), LetIn (n,b,t,c)) ->
- (mkLetIn (n,b,loop i p t,c))
+ (mkLetIn (n,b,loop i p t,c))
| (p, _) ->
failwith ("abstract_path " ^ string_of_int(List.length p))
in
@@ -528,7 +529,7 @@ let occurrence sigma path (t : constr) =
let abstract_path sigma typ path t =
let term_occur = ref (mkRel 0) in
let abstract = context sigma (fun i t -> term_occur:= t; mkRel i) path t in
- mkLambda (Name (Id.of_string "x"), typ, abstract), !term_occur
+ mkLambda (make_annot (Name (Id.of_string "x")) Sorts.Relevant, typ, abstract), !term_occur
let focused_simpl path =
let open Tacmach.New in
@@ -604,10 +605,10 @@ let clever_rewrite_base_poly typ p result theorem =
let t =
applist
(mkLambda
- (Name (Id.of_string "P"),
- mkArrow typ mkProp,
+ (make_annot (Name (Id.of_string "P")) Sorts.Relevant,
+ mkArrow typ Sorts.Relevant mkProp,
mkLambda
- (Name (Id.of_string "H"),
+ (make_annot (Name (Id.of_string "H")) Sorts.Relevant,
applist (mkRel 1,[result]),
mkApp (Lazy.force coq_eq_ind_r,
[| typ; result; mkRel 2; mkRel 1; occ; theorem |]))),
@@ -1264,7 +1265,7 @@ let replay_history tactic_normalisation =
mkApp (Lazy.force coq_ex, [|
Lazy.force coq_Z;
mkLambda
- (Name vid,
+ (make_annot (Name vid) Sorts.Relevant,
Lazy.force coq_Z,
mk_eq (mkRel 1) eq1) |])
in
@@ -1725,11 +1726,11 @@ let destructure_hyps =
try
match destructurate_type env sigma typ with
| Kapp(Nat,_) | Kapp(Z,_) ->
- let hid = fresh_id Id.Set.empty (add_suffix i "_eqn") gl in
- let hty = mk_gen_eq typ (mkVar i) body in
+ let hid = fresh_id Id.Set.empty (add_suffix i.binder_name "_eqn") gl in
+ let hty = mk_gen_eq typ (mkVar i.binder_name) body in
tclTHEN
(assert_by (Name hid) hty reflexivity)
- (loop (LocalAssum (hid, hty) :: lit))
+ (loop (LocalAssum (make_annot hid Sorts.Relevant, hty) :: lit))
| _ -> loop lit
with e when catchable_exception e -> loop lit
end
@@ -1742,18 +1743,20 @@ let destructure_hyps =
| Kapp(Or,[t1;t2]) ->
(tclTHENS
(elim_id i)
- [ onClearedName i (fun i -> (loop (LocalAssum (i,t1)::lit)));
- onClearedName i (fun i -> (loop (LocalAssum (i,t2)::lit))) ])
+ [ onClearedName i (fun i -> (loop (LocalAssum (make_annot i Sorts.Relevant,t1)::lit)));
+ onClearedName i (fun i -> (loop (LocalAssum (make_annot i Sorts.Relevant,t2)::lit))) ])
| Kapp(And,[t1;t2]) ->
tclTHEN
(elim_id i)
(onClearedName2 i (fun i1 i2 ->
- loop (LocalAssum (i1,t1) :: LocalAssum (i2,t2) :: lit)))
+ loop (LocalAssum (make_annot i1 Sorts.Relevant,t1) ::
+ LocalAssum (make_annot i2 Sorts.Relevant,t2) :: lit)))
| Kapp(Iff,[t1;t2]) ->
tclTHEN
(elim_id i)
(onClearedName2 i (fun i1 i2 ->
- loop (LocalAssum (i1,mkArrow t1 t2) :: LocalAssum (i2,mkArrow t2 t1) :: lit)))
+ loop (LocalAssum (make_annot i1 Sorts.Relevant,mkArrow t1 Sorts.Relevant t2) ::
+ LocalAssum (make_annot i2 Sorts.Relevant,mkArrow t2 Sorts.Relevant t1) :: lit)))
| Kimp(t1,t2) ->
(* t1 and t2 might be in Type rather than Prop.
For t1, the decidability check will ensure being Prop. *)
@@ -1764,7 +1767,7 @@ let destructure_hyps =
(generalize_tac [mkApp (Lazy.force coq_imp_simp,
[| t1; t2; d1; mkVar i|])]);
(onClearedName i (fun i ->
- (loop (LocalAssum (i,mk_or (mk_not t1) t2) :: lit))))
+ (loop (LocalAssum (make_annot i Sorts.Relevant,mk_or (mk_not t1) t2) :: lit))))
]
else
loop lit
@@ -1775,7 +1778,7 @@ let destructure_hyps =
(generalize_tac
[mkApp (Lazy.force coq_not_or,[| t1; t2; mkVar i |])]);
(onClearedName i (fun i ->
- (loop (LocalAssum (i,mk_and (mk_not t1) (mk_not t2)) :: lit))))
+ (loop (LocalAssum (make_annot i Sorts.Relevant,mk_and (mk_not t1) (mk_not t2)) :: lit))))
]
| Kapp(And,[t1;t2]) ->
let d1 = decidability t1 in
@@ -1784,7 +1787,7 @@ let destructure_hyps =
[mkApp (Lazy.force coq_not_and,
[| t1; t2; d1; mkVar i |])]);
(onClearedName i (fun i ->
- (loop (LocalAssum (i,mk_or (mk_not t1) (mk_not t2)) :: lit))))
+ (loop (LocalAssum (make_annot i Sorts.Relevant,mk_or (mk_not t1) (mk_not t2)) :: lit))))
]
| Kapp(Iff,[t1;t2]) ->
let d1 = decidability t1 in
@@ -1794,7 +1797,7 @@ let destructure_hyps =
[mkApp (Lazy.force coq_not_iff,
[| t1; t2; d1; d2; mkVar i |])]);
(onClearedName i (fun i ->
- (loop (LocalAssum (i, mk_or (mk_and t1 (mk_not t2))
+ (loop (LocalAssum (make_annot i Sorts.Relevant,mk_or (mk_and t1 (mk_not t2))
(mk_and (mk_not t1) t2)) :: lit))))
]
| Kimp(t1,t2) ->
@@ -1806,14 +1809,14 @@ let destructure_hyps =
[mkApp (Lazy.force coq_not_imp,
[| t1; t2; d1; mkVar i |])]);
(onClearedName i (fun i ->
- (loop (LocalAssum (i,mk_and t1 (mk_not t2)) :: lit))))
+ (loop (LocalAssum (make_annot i Sorts.Relevant,mk_and t1 (mk_not t2)) :: lit))))
]
| Kapp(Not,[t]) ->
let d = decidability t in
tclTHENLIST [
(generalize_tac
[mkApp (Lazy.force coq_not_not, [| t; d; mkVar i |])]);
- (onClearedName i (fun i -> (loop (LocalAssum (i,t) :: lit))))
+ (onClearedName i (fun i -> (loop (LocalAssum (make_annot i Sorts.Relevant,t) :: lit))))
]
| Kapp(op,[t1;t2]) ->
(try
diff --git a/plugins/rtauto/refl_tauto.ml b/plugins/rtauto/refl_tauto.ml
index a6b6c57ff9..89528fe357 100644
--- a/plugins/rtauto/refl_tauto.ml
+++ b/plugins/rtauto/refl_tauto.ml
@@ -16,6 +16,7 @@ open CErrors
open Util
open Term
open Constr
+open Context
open Proof_search
open Context.Named.Declaration
@@ -127,7 +128,7 @@ let rec make_hyps env sigma atom_env lenv = function
| LocalAssum (id,typ)::rest ->
let hrec=
make_hyps env sigma atom_env (typ::lenv) rest in
- if List.exists (fun c -> Termops.local_occur_var Evd.empty (* FIXME *) id c) lenv ||
+ if List.exists (fun c -> Termops.local_occur_var Evd.empty (* FIXME *) id.binder_name c) lenv ||
(Retyping.get_sort_family_of env sigma typ != InProp)
then
hrec
@@ -291,7 +292,7 @@ let rtauto_tac =
build_form formula;
build_proof [] 0 prf|]) in
let term=
- applistc main (List.rev_map (fun (id,_) -> mkVar id) hyps) in
+ applistc main (List.rev_map (fun (id,_) -> mkVar id.binder_name) hyps) in
let build_end_time=System.get_time () in
let () = if !verbose then
begin
diff --git a/plugins/rtauto/refl_tauto.mli b/plugins/rtauto/refl_tauto.mli
index 49b5ee5ac7..3de0ba44df 100644
--- a/plugins/rtauto/refl_tauto.mli
+++ b/plugins/rtauto/refl_tauto.mli
@@ -23,6 +23,6 @@ val make_hyps
-> atom_env
-> EConstr.types list
-> EConstr.named_context
- -> (Names.Id.t * Proof_search.form) list
+ -> (Names.Id.t Context.binder_annot * Proof_search.form) list
val rtauto_tac : unit Proofview.tactic
diff --git a/plugins/ssr/ssrcommon.ml b/plugins/ssr/ssrcommon.ml
index 0961edb6cb..f6f92dcb3f 100644
--- a/plugins/ssr/ssrcommon.ml
+++ b/plugins/ssr/ssrcommon.ml
@@ -15,6 +15,7 @@ open Names
open Evd
open Term
open Constr
+open Context
open Termops
open Printer
open Locusops
@@ -429,15 +430,16 @@ let convert_concl t = Tactics.convert_concl t DEFAULTcast
let rename_hd_prod orig_name_ref gl =
match EConstr.kind (project gl) (pf_concl gl) with
- | Prod(_,src,tgt) ->
- Proofview.V82.of_tactic (convert_concl_no_check (EConstr.mkProd (!orig_name_ref,src,tgt))) gl
+ | Prod(x,src,tgt) ->
+ let x = {x with binder_name = !orig_name_ref} in
+ Proofview.V82.of_tactic (convert_concl_no_check (EConstr.mkProd (x,src,tgt))) gl
| _ -> CErrors.anomaly (str "gentac creates no product")
(* Reduction that preserves the Prod/Let spine of the "in" tactical. *)
let inc_safe n = if n = 0 then n else n + 1
let rec safe_depth s c = match EConstr.kind s c with
-| LetIn (Name x, _, _, c') when is_discharged_id x -> safe_depth s c' + 1
+| LetIn ({binder_name=Name x}, _, _, c') when is_discharged_id x -> safe_depth s c' + 1
| LetIn (_, _, _, c') | Prod (_, _, c') -> inc_safe (safe_depth s c')
| _ -> 0
@@ -529,7 +531,7 @@ let pf_abs_evars2 gl rigid (sigma, c0) =
let concl = EConstr.Unsafe.to_constr evi.evar_concl in
let dc = EConstr.Unsafe.to_named_context (CList.firstn n (evar_filtered_context evi)) in
let abs_dc c = function
- | NamedDecl.LocalDef (x,b,t) -> mkNamedLetIn x b t (mkArrow t c)
+ | NamedDecl.LocalDef (x,b,t) -> mkNamedLetIn x b t (mkArrow t x.binder_relevance c)
| NamedDecl.LocalAssum (x,t) -> mkNamedProd x t c in
let t = Context.Named.fold_inside abs_dc ~init:concl dc in
nf_evar sigma t in
@@ -552,7 +554,7 @@ let pf_abs_evars2 gl rigid (sigma, c0) =
| _ -> Constr.map_with_binders ((+) 1) get i c in
let rec loop c i = function
| (_, (n, t)) :: evl ->
- loop (mkLambda (mk_evar_name n, get (i - 1) t, c)) (i - 1) evl
+ loop (mkLambda (make_annot (mk_evar_name n) Sorts.Relevant, get (i - 1) t, c)) (i - 1) evl
| [] -> c in
List.length evlist, EConstr.of_constr (loop (get 1 c0) 1 evlist), List.map fst evlist, ucst
@@ -590,7 +592,7 @@ let pf_abs_evars_pirrel gl (sigma, c0) =
let concl = EConstr.Unsafe.to_constr evi.evar_concl in
let dc = EConstr.Unsafe.to_named_context (CList.firstn n (evar_filtered_context evi)) in
let abs_dc c = function
- | NamedDecl.LocalDef (x,b,t) -> mkNamedLetIn x b t (mkArrow t c)
+ | NamedDecl.LocalDef (x,b,t) -> mkNamedLetIn x b t (mkArrow t x.binder_relevance c)
| NamedDecl.LocalAssum (x,t) -> mkNamedProd x t c in
let t = Context.Named.fold_inside abs_dc ~init:concl dc in
nf_evar sigma0 (nf_evar sigma t) in
@@ -646,7 +648,7 @@ let pf_abs_evars_pirrel gl (sigma, c0) =
| (_, (n, t, _)) :: evl ->
let t = get evlist (i - 1) t in
let n = Name (Id.of_string (ssr_anon_hyp ^ string_of_int n)) in
- loopP evlist (mkProd (n, t, c)) (i - 1) evl
+ loopP evlist (mkProd (make_annot n Sorts.Relevant, t, c)) (i - 1) evl
| [] -> c in
let rec loop c i = function
| (_, (n, t, _)) :: evl ->
@@ -658,7 +660,7 @@ let pf_abs_evars_pirrel gl (sigma, c0) =
List.map (fun (k,_) -> mkRel (fst (lookup k i evlist)))
(List.rev t_evplist) in
let c = if extra_args = [] then c else app extra_args 1 c in
- loop (mkLambda (mk_evar_name n, t, c)) (i - 1) evl
+ loop (mkLambda (make_annot (mk_evar_name n) Sorts.Relevant, t, c)) (i - 1) evl
| [] -> c in
let res = loop (get evlist 1 c0) 1 evlist in
pp(lazy(str"res= " ++ pr_constr res));
@@ -710,13 +712,13 @@ let pf_abs_cterm gl n c0 =
| _ -> [], strip i c in
let rec strip_evars i c = match Constr.kind c with
| Lambda (x, t1, c1) when i < n ->
- let na = nb_evar_deps x in
+ let na = nb_evar_deps x.binder_name in
let dl, t2 = strip_ndeps (i + na) i t1 in
let na' = List.length dl in
eva.(i) <- Array.of_list (na - na' :: dl);
let x' =
if na' = 0 then Name (pf_type_id gl (EConstr.of_constr t2)) else mk_evar_name na' in
- mkLambda (x', t2, strip_evars (i + 1) c1)
+ mkLambda ({x with binder_name=x'}, t2, strip_evars (i + 1) c1)
(* if noccurn 1 c2 then lift (-1) c2 else
mkLambda (Name (pf_type_id gl t2), t2, c2) *)
| _ -> strip i c in
@@ -740,8 +742,9 @@ let rec constr_name sigma c = match EConstr.kind sigma c with
let pf_mkprod gl c ?(name=constr_name (project gl) c) cl =
let gl, t = pfe_type_of gl c in
- if name <> Anonymous || EConstr.Vars.noccurn (project gl) 1 cl then gl, EConstr.mkProd (name, t, cl) else
- gl, EConstr.mkProd (Name (pf_type_id gl t), t, cl)
+ let r = pf_apply Retyping.relevance_of_term gl c in
+ if name <> Anonymous || EConstr.Vars.noccurn (project gl) 1 cl then gl, EConstr.mkProd (make_annot name r, t, cl) else
+ gl, EConstr.mkProd (make_annot (Name (pf_type_id gl t)) r, t, cl)
let pf_abs_prod name gl c cl = pf_mkprod gl c ~name (Termops.subst_term (project gl) c cl)
@@ -783,13 +786,17 @@ let mkRefl t c gl =
let discharge_hyp (id', (id, mode)) gl =
let cl' = Vars.subst_var id (pf_concl gl) in
- match pf_get_hyp gl id, mode with
- | NamedDecl.LocalAssum (_, t), _ | NamedDecl.LocalDef (_, _, t), "(" ->
- Proofview.V82.of_tactic (Tactics.apply_type ~typecheck:true (EConstr.of_constr (mkProd (Name id', t, cl')))
+ let decl = pf_get_hyp gl id in
+ match decl, mode with
+ | NamedDecl.LocalAssum _, _ | NamedDecl.LocalDef _, "(" ->
+ let id' = {(NamedDecl.get_annot decl) with binder_name = Name id'} in
+ Proofview.V82.of_tactic (Tactics.apply_type ~typecheck:true
+ (EConstr.of_constr (mkProd (id', NamedDecl.get_type decl, cl')))
[EConstr.of_constr (mkVar id)]) gl
| NamedDecl.LocalDef (_, v, t), _ ->
+ let id' = {(NamedDecl.get_annot decl) with binder_name = Name id'} in
Proofview.V82.of_tactic
- (convert_concl (EConstr.of_constr (mkLetIn (Name id', v, t, cl')))) gl
+ (convert_concl (EConstr.of_constr (mkLetIn (id', v, t, cl')))) gl
(* wildcard names *)
let clear_wilds wilds gl =
@@ -983,7 +990,7 @@ let applyn ~with_evars ?beta ?(with_shelve=false) n t gl =
let rec loop sigma bo args = function (* saturate with metas *)
| 0 -> EConstr.mkApp (t, Array.of_list (List.rev args)), re_sig si sigma
| n -> match EConstr.kind sigma bo with
- | Lambda (_, ty, bo) ->
+ | Lambda (_, ty, bo) ->
if not (EConstr.Vars.closed0 sigma ty) then
raise dependent_apply_error;
let m = Evarutil.new_meta () in
@@ -1019,7 +1026,7 @@ let () = CLexer.set_keyword_state frozen_lexer ;;
let rec fst_prod red tac = Proofview.Goal.enter begin fun gl ->
let concl = Proofview.Goal.concl gl in
match EConstr.kind (Proofview.Goal.sigma gl) concl with
- | Prod (id,_,tgt) | LetIn(id,_,_,tgt) -> tac id
+ | Prod (id,_,tgt) | LetIn(id,_,_,tgt) -> tac id.binder_name
| _ -> if red then Tacticals.New.tclZEROMSG (str"No product even after head-reduction.")
else Tacticals.New.tclTHEN Tactics.hnf_in_concl (fst_prod true tac)
end
@@ -1122,14 +1129,15 @@ let pf_interp_gen_aux gl to_ind ((oclr, occ), t) =
errorstrm (str "@ can be used with variables only")
else match Tacmach.pf_get_hyp gl (EConstr.destVar sigma c) with
| NamedDecl.LocalAssum _ -> errorstrm (str "@ can be used with let-ins only")
- | NamedDecl.LocalDef (name, b, ty) -> true, pat, EConstr.mkLetIn (Name name,b,ty,cl),c,clr,ucst,gl
+ | NamedDecl.LocalDef (name, b, ty) -> true, pat, EConstr.mkLetIn (map_annot Name.mk_name name,b,ty,cl),c,clr,ucst,gl
else let gl, ccl = pf_mkprod gl c cl in false, pat, ccl, c, clr,ucst,gl
else if to_ind && occ = None then
let nv, p, _, ucst' = pf_abs_evars gl (fst pat, c) in
let ucst = UState.union ucst ucst' in
if nv = 0 then anomaly "occur_existential but no evars" else
let gl, pty = pfe_type_of gl p in
- false, pat, EConstr.mkProd (constr_name (project gl) c, pty, Tacmach.pf_concl gl), p, clr,ucst,gl
+ let rp = pf_apply Retyping.relevance_of_term gl p in
+ false, pat, EConstr.mkProd (make_annot (constr_name (project gl) c) rp, pty, Tacmach.pf_concl gl), p, clr,ucst,gl
else CErrors.user_err ?loc:(loc_of_cpattern t) (str "generalized term didn't match")
let apply_type x xs = Proofview.V82.of_tactic (Tactics.apply_type ~typecheck:true x xs)
@@ -1235,7 +1243,7 @@ let abs_wgen keep_let f gen (gl,args,c) =
(EConstr.Vars.subst_var x c)
| _, Some ((x, _), None) ->
let x = hoi_id x in
- gl, EConstr.mkVar x :: args, EConstr.mkProd (Name (f x),Tacmach.pf_get_hyp_typ gl x, EConstr.Vars.subst_var x c)
+ gl, EConstr.mkVar x :: args, EConstr.mkProd (make_annot (Name (f x)) Sorts.Relevant,Tacmach.pf_get_hyp_typ gl x, EConstr.Vars.subst_var x c)
| _, Some ((x, "@"), Some p) ->
let x = hoi_id x in
let cp = interp_cpattern gl p None in
@@ -1247,7 +1255,7 @@ let abs_wgen keep_let f gen (gl,args,c) =
evar_closed t p;
let ut = red_product_skip_id env sigma t in
let gl, ty = pfe_type_of gl t in
- pf_merge_uc ucst gl, args, EConstr.mkLetIn(Name (f x), ut, ty, c)
+ pf_merge_uc ucst gl, args, EConstr.mkLetIn(make_annot (Name (f x)) Sorts.Relevant, ut, ty, c)
| _, Some ((x, _), Some p) ->
let x = hoi_id x in
let cp = interp_cpattern gl p None in
@@ -1258,7 +1266,7 @@ let abs_wgen keep_let f gen (gl,args,c) =
let t = EConstr.of_constr t in
evar_closed t p;
let gl, ty = pfe_type_of gl t in
- pf_merge_uc ucst gl, t :: args, EConstr.mkProd(Name (f x), ty, c)
+ pf_merge_uc ucst gl, t :: args, EConstr.mkProd(make_annot (Name (f x)) Sorts.Relevant, ty, c)
| _ -> gl, args, c
let clr_of_wgen gen clrs = match gen with
@@ -1321,8 +1329,8 @@ let unsafe_intro env decl b =
end
let set_decl_id id = let open Context in function
- | Rel.Declaration.LocalAssum(name,ty) -> Named.Declaration.LocalAssum(id,ty)
- | Rel.Declaration.LocalDef(name,ty,t) -> Named.Declaration.LocalDef(id,ty,t)
+ | Rel.Declaration.LocalAssum(name,ty) -> Named.Declaration.LocalAssum({name with binder_name=id},ty)
+ | Rel.Declaration.LocalDef(name,ty,t) -> Named.Declaration.LocalDef({name with binder_name=id},ty,t)
let rec decompose_assum env sigma orig_goal =
let open Context in
@@ -1400,8 +1408,8 @@ let tclRENAME_HD_PROD name = Goal.enter begin fun gl ->
let concl = Goal.concl gl in
let sigma = Goal.sigma gl in
match EConstr.kind sigma concl with
- | Prod(_,src,tgt) ->
- convert_concl_no_check EConstr.(mkProd (name,src,tgt))
+ | Prod(x,src,tgt) ->
+ convert_concl_no_check EConstr.(mkProd ({x with binder_name = name},src,tgt))
| _ -> CErrors.anomaly (Pp.str "rename_hd_prod: no head product")
end
@@ -1429,11 +1437,12 @@ let tacMKPROD c ?name cl =
tacCONSTR_NAME ?name c >>= fun name ->
Goal.enter_one ~__LOC__ begin fun g ->
let sigma, env = Goal.sigma g, Goal.env g in
+ let r = Sorts.Relevant in (* TODO relevance *)
if name <> Names.Name.Anonymous || EConstr.Vars.noccurn sigma 1 cl
- then tclUNIT (EConstr.mkProd (name, t, cl))
+ then tclUNIT (EConstr.mkProd (make_annot name r, t, cl))
else
let name = Names.Id.of_string (Namegen.hdchar env sigma t) in
- tclUNIT (EConstr.mkProd (Names.Name.Name name, t, cl))
+ tclUNIT (EConstr.mkProd (make_annot (Name.Name name) r, t, cl))
end
let tacINTERP_CPATTERN cp =
diff --git a/plugins/ssr/ssrcommon.mli b/plugins/ssr/ssrcommon.mli
index e642b5e788..f0e7b7e6a5 100644
--- a/plugins/ssr/ssrcommon.mli
+++ b/plugins/ssr/ssrcommon.mli
@@ -155,7 +155,7 @@ val pf_e_type_of :
val splay_open_constr :
Goal.goal Evd.sigma ->
evar_map * EConstr.t ->
- (Names.Name.t * EConstr.t) list * EConstr.t
+ (Names.Name.t Context.binder_annot * EConstr.t) list * EConstr.t
val isAppInd : Environ.env -> Evd.evar_map -> EConstr.types -> bool
val mk_term : ssrtermkind -> constr_expr -> ssrterm
diff --git a/plugins/ssr/ssrelim.ml b/plugins/ssr/ssrelim.ml
index 7216849948..82a88678f0 100644
--- a/plugins/ssr/ssrelim.ml
+++ b/plugins/ssr/ssrelim.ml
@@ -15,6 +15,7 @@ open Names
open Printer
open Term
open Constr
+open Context
open Termops
open Tactypes
open Tacmach
@@ -364,14 +365,14 @@ let ssrelim ?(is_case=false) deps what ?elim eqid elim_intro_tac =
let gl, eq = get_eq_type gl in
let gen_eq_tac, gl =
let refl = EConstr.mkApp (eq, [|t; c; c|]) in
- let new_concl = EConstr.mkArrow refl (EConstr.Vars.lift 1 (pf_concl orig_gl)) in
+ let new_concl = EConstr.mkArrow refl Sorts.Relevant (EConstr.Vars.lift 1 (pf_concl orig_gl)) in
let new_concl = fire_subst gl new_concl in
let erefl, gl = mkRefl t c gl in
let erefl = fire_subst gl erefl in
apply_type new_concl [erefl], gl in
let rel = k + if c_is_head_p then 1 else 0 in
let src, gl = mkProt EConstr.mkProp EConstr.(mkApp (eq,[|t; c; mkRel rel|])) gl in
- let concl = EConstr.mkArrow src (EConstr.Vars.lift 1 concl) in
+ let concl = EConstr.mkArrow src Sorts.Relevant (EConstr.Vars.lift 1 concl) in
let clr = if deps <> [] then clr else [] in
concl, gen_eq_tac, clr, gl
| _ -> concl, Tacticals.tclIDTAC, clr, gl in
@@ -446,7 +447,7 @@ let injecteq_id = mk_internal_id "injection equation"
let revtoptac n0 gl =
let n = pf_nb_prod gl - n0 in
let dc, cl = EConstr.decompose_prod_n_assum (project gl) n (pf_concl gl) in
- let dc' = dc @ [Context.Rel.Declaration.LocalAssum(Name rev_id, EConstr.it_mkProd_or_LetIn cl (List.rev dc))] in
+ let dc' = dc @ [Context.Rel.Declaration.LocalAssum(make_annot (Name rev_id) Sorts.Relevant, EConstr.it_mkProd_or_LetIn cl (List.rev dc))] in
let f = EConstr.it_mkLambda_or_LetIn (mkEtaApp (EConstr.mkRel (n + 1)) (-n) 1) dc' in
Refiner.refiner ~check:true EConstr.Unsafe.(to_constr (EConstr.mkApp (f, [|Evarutil.mk_new_meta ()|]))) gl
@@ -486,7 +487,7 @@ let perform_injection c gl =
CErrors.user_err (Pp.str "can't decompose a quantified equality") else
let cl = pf_concl gl in let n = List.length dc in
let c_eq = mkEtaApp c n 2 in
- let cl1 = EConstr.mkLambda EConstr.(Anonymous, mkArrow eqt cl, mkApp (mkRel 1, [|c_eq|])) in
+ let cl1 = EConstr.mkLambda EConstr.(make_annot Anonymous Sorts.Relevant, mkArrow eqt Sorts.Relevant cl, mkApp (mkRel 1, [|c_eq|])) in
let id = injecteq_id in
let id_with_ebind = (EConstr.mkVar id, NoBindings) in
let injtac = Tacticals.tclTHEN (introid id) (injectidl2rtac id id_with_ebind) in
diff --git a/plugins/ssr/ssrequality.ml b/plugins/ssr/ssrequality.ml
index 64e023c68a..18461c0533 100644
--- a/plugins/ssr/ssrequality.ml
+++ b/plugins/ssr/ssrequality.ml
@@ -15,6 +15,7 @@ open Util
open Names
open Term
open Constr
+open Context
open Vars
open Locus
open Printer
@@ -136,7 +137,7 @@ let newssrcongrtac arg ist gl =
(fun ty -> congrtac (arg, Detyping.detype Detyping.Now false Id.Set.empty (pf_env gl) (project gl) ty) ist)
(fun () ->
let lhs, gl' = mk_evar gl EConstr.mkProp in let rhs, gl' = mk_evar gl' EConstr.mkProp in
- let arrow = EConstr.mkArrow lhs (EConstr.Vars.lift 1 rhs) in
+ let arrow = EConstr.mkArrow lhs Sorts.Relevant (EConstr.Vars.lift 1 rhs) in
tclMATCH_GOAL (arrow, gl') (fun gl' -> [|fs gl' lhs;fs gl' rhs|])
(fun lr -> tclTHEN (Proofview.V82.of_tactic (Tactics.apply (ssr_congr lr))) (congrtac (arg, mkRType) ist))
(fun _ _ -> errorstrm Pp.(str"Conclusion is not an equality nor an arrow")))
@@ -335,7 +336,7 @@ let pirrel_rewrite pred rdx rdx_ty new_rdx dir (sigma, c) c_ty gl =
let (sigma, ev) = Evarutil.new_evar env sigma (beta (EConstr.Vars.subst1 new_rdx pred)) in
(sigma, ev)
in
- let pred = EConstr.mkNamedLambda pattern_id rdx_ty pred in
+ let pred = EConstr.mkNamedLambda (make_annot pattern_id Sorts.Relevant) rdx_ty pred in
let elim, gl =
let ((kn, i) as ind, _), unfolded_c_ty = pf_reduce_to_quantified_ind gl c_ty in
let sort = elimination_sort_of_goal gl in
@@ -362,7 +363,7 @@ let pirrel_rewrite pred rdx rdx_ty new_rdx dir (sigma, c) c_ty gl =
let names = let rec aux t = function 0 -> [] | n ->
let t = Reductionops.whd_all env sigma t in
match EConstr.kind_of_type sigma t with
- | ProdType (name, _, t) -> name :: aux t (n-1)
+ | ProdType (name, _, t) -> name.binder_name :: aux t (n-1)
| _ -> assert false in aux hd_ty (Array.length args) in
hd_ty, Util.List.map_filter (fun (t, name) ->
let evs = Evar.Set.elements (Evarutil.undefined_evars_of_term sigma t) in
@@ -403,7 +404,7 @@ let rwcltac cl rdx dir sr gl =
let new_rdx = if dir = L2R then a.(2) else a.(1) in
pirrel_rewrite cl rdx rdxt new_rdx dir (sigma,c) c_ty, tclIDTAC, gl
| _ ->
- let cl' = EConstr.mkApp (EConstr.mkNamedLambda pattern_id rdxt cl, [|rdx|]) in
+ let cl' = EConstr.mkApp (EConstr.mkNamedLambda (make_annot pattern_id Sorts.Relevant) rdxt cl, [|rdx|]) in
let sigma, _ = Typing.type_of env sigma cl' in
let gl = pf_merge_uc_of sigma gl in
Proofview.V82.of_tactic (convert_concl cl'), rewritetac dir r', gl
@@ -413,8 +414,8 @@ let rwcltac cl rdx dir sr gl =
try EConstr.destCast (project gl) r2 with _ ->
errorstrm Pp.(str "no cast from " ++ pr_constr_pat (EConstr.Unsafe.to_constr (snd sr))
++ str " to " ++ pr_econstr_env (pf_env gl) (project gl) r2) in
- let cl' = EConstr.mkNamedProd rule_id (EConstr.it_mkProd_or_LetIn r3t dc) (EConstr.Vars.lift 1 cl) in
- let cl'' = EConstr.mkNamedProd pattern_id rdxt cl' in
+ let cl' = EConstr.mkNamedProd (make_annot rule_id Sorts.Relevant) (EConstr.it_mkProd_or_LetIn r3t dc) (EConstr.Vars.lift 1 cl) in
+ let cl'' = EConstr.mkNamedProd (make_annot pattern_id Sorts.Relevant) rdxt cl' in
let itacs = [introid pattern_id; introid rule_id] in
let cltac = Proofview.V82.of_tactic (Tactics.clear [pattern_id; rule_id]) in
let rwtacs = [rewritetac dir (EConstr.mkVar rule_id); cltac] in
@@ -426,7 +427,9 @@ let rwcltac cl rdx dir sr gl =
if occur_existential (project gl) (Tacmach.pf_concl gl)
then errorstrm Pp.(str "Rewriting impacts evars")
else errorstrm Pp.(str "Dependent type error in rewrite of "
- ++ pr_constr_env (pf_env gl) (project gl) (Term.mkNamedLambda pattern_id (EConstr.Unsafe.to_constr rdxt) (EConstr.Unsafe.to_constr cl)))
+ ++ pr_constr_env (pf_env gl) (project gl)
+ (Term.mkNamedLambda (make_annot pattern_id Sorts.Relevant)
+ (EConstr.Unsafe.to_constr rdxt) (EConstr.Unsafe.to_constr cl)))
in
tclTHEN cvtac' rwtac gl
diff --git a/plugins/ssr/ssrfwd.ml b/plugins/ssr/ssrfwd.ml
index 8c1363020a..9ea35b8694 100644
--- a/plugins/ssr/ssrfwd.ml
+++ b/plugins/ssr/ssrfwd.ml
@@ -13,6 +13,7 @@
open Pp
open Names
open Constr
+open Context
open Tacmach
open Ssrmatching_plugin.Ssrmatching
@@ -54,7 +55,7 @@ let ssrsettac id ((_, (pat, pty)), (_, occ)) gl =
let c, (gl, cty) = match EConstr.kind sigma c with
| Cast(t, DEFAULTcast, ty) -> t, (gl, ty)
| _ -> c, pfe_type_of gl c in
- let cl' = EConstr.mkLetIn (Name id, c, cty, cl) in
+ let cl' = EConstr.mkLetIn (make_annot (Name id) Sorts.Relevant, c, cty, cl) in
Tacticals.tclTHEN (Proofview.V82.of_tactic (convert_concl cl')) (introid id) gl
open Util
@@ -162,7 +163,7 @@ let havetac ist
let assert_is_conv gl =
try Proofview.V82.of_tactic (convert_concl (EConstr.it_mkProd_or_LetIn concl ctx)) gl
with _ -> errorstrm (str "Given proof term is not of type " ++
- pr_econstr_env (pf_env gl) (project gl) (EConstr.mkArrow (EConstr.mkVar (Id.of_string "_")) concl)) in
+ pr_econstr_env (pf_env gl) (project gl) (EConstr.mkArrow (EConstr.mkVar (Id.of_string "_")) Sorts.Relevant concl)) in
gl, ty, Tacticals.tclTHEN assert_is_conv (Proofview.V82.of_tactic (Tactics.apply t)), id, itac_c
| FwdHave, false, false ->
let skols = List.flatten (List.map (function
@@ -190,10 +191,10 @@ let havetac ist
Proofview.V82.of_tactic (unfold [abstract; abstract_key]) gl))
| _,true,true ->
let _, ty, uc = interp_ty gl fixtc cty in let gl = pf_merge_uc uc gl in
- gl, EConstr.mkArrow ty concl, hint, itac, clr
+ gl, EConstr.mkArrow ty Sorts.Relevant concl, hint, itac, clr
| _,false,true ->
let _, ty, uc = interp_ty gl fixtc cty in let gl = pf_merge_uc uc gl in
- gl, EConstr.mkArrow ty concl, hint, id, itac_c
+ gl, EConstr.mkArrow ty Sorts.Relevant concl, hint, id, itac_c
| _, false, false ->
let n, cty, uc = interp_ty gl fixtc cty in let gl = pf_merge_uc uc gl in
gl, cty, Tacticals.tclTHEN (binderstac n) hint, id, Tacticals.tclTHEN itac_c simpltac
@@ -233,7 +234,7 @@ let wlogtac ist (((clr0, pats),_),_) (gens, ((_, ct))) hint suff ghave gl =
let gens = List.filter (function _, Some _ -> true | _ -> false) gens in
let concl = pf_concl gl in
let c = EConstr.mkProp in
- let c = if cut_implies_goal then EConstr.mkArrow c concl else c in
+ let c = if cut_implies_goal then EConstr.mkArrow c Sorts.Relevant concl else c in
let gl, args, c = List.fold_right mkabs gens (gl,[],c) in
let env, _ =
List.fold_left (fun (env, c) _ ->
@@ -245,10 +246,10 @@ let wlogtac ist (((clr0, pats),_),_) (gens, ((_, ct))) hint suff ghave gl =
let fake_gl = {Evd.it = k; Evd.sigma = sigma} in
let _, ct, _, uc = pf_interp_ty ist fake_gl ct in
let rec var2rel c g s = match EConstr.kind sigma c, g with
- | Prod(Anonymous,_,c), [] -> EConstr.mkProd(Anonymous, EConstr.Vars.subst_vars s ct, c)
+ | Prod({binder_name=Anonymous} as x,_,c), [] -> EConstr.mkProd(x, EConstr.Vars.subst_vars s ct, c)
| Sort _, [] -> EConstr.Vars.subst_vars s ct
- | LetIn(Name id as n,b,ty,c), _::g -> EConstr.mkLetIn (n,b,ty,var2rel c g (id::s))
- | Prod(Name id as n,ty,c), _::g -> EConstr.mkProd (n,ty,var2rel c g (id::s))
+ | LetIn({binder_name=Name id} as n,b,ty,c), _::g -> EConstr.mkLetIn (n,b,ty,var2rel c g (id::s))
+ | Prod({binder_name=Name id} as n,ty,c), _::g -> EConstr.mkProd (n,ty,var2rel c g (id::s))
| _ -> CErrors.anomaly(str"SSR: wlog: var2rel: " ++ pr_econstr_env env sigma c) in
let c = var2rel c gens [] in
let rec pired c = function
diff --git a/plugins/ssr/ssripats.ml b/plugins/ssr/ssripats.ml
index a8dfd69240..e9fe1f3e48 100644
--- a/plugins/ssr/ssripats.ml
+++ b/plugins/ssr/ssripats.ml
@@ -13,6 +13,7 @@ open Ssrmatching_plugin
open Util
open Names
open Constr
+open Context
open Proofview
open Proofview.Notations
@@ -393,12 +394,12 @@ let tcltclMK_ABSTRACT_VAR id = Goal.enter begin fun gl ->
let sigma, m = Evarutil.new_evar env sigma abstract_ty in
sigma, (m, abstract_ty) in
let sigma, kont =
- let rd = Context.Rel.Declaration.LocalAssum (Name id, abstract_ty) in
+ let rd = Context.Rel.Declaration.LocalAssum (make_annot (Name id) Sorts.Relevant, abstract_ty) in
let sigma, ev = Evarutil.new_evar (EConstr.push_rel rd env) sigma concl in
sigma, ev
in
let term =
- EConstr.(mkApp (mkLambda(Name id,abstract_ty,kont),[|abstract_proof|])) in
+ EConstr.(mkApp (mkLambda(make_annot (Name id) Sorts.Relevant,abstract_ty,kont),[|abstract_proof|])) in
let sigma, _ = Typing.type_of env sigma term in
sigma, term
end in
@@ -608,7 +609,7 @@ let with_defective maintac deps clr = Goal.enter begin fun g ->
let sigma, concl = Goal.(sigma g, concl g) in
let top_id =
match EConstr.kind_of_type sigma concl with
- | Term.ProdType (Name id, _, _)
+ | Term.ProdType ({binder_name=Name id}, _, _)
when Ssrcommon.is_discharged_id id -> id
| _ -> Ssrcommon.top_id in
let top_gen = Ssrequality.mkclr clr, Ssrmatching.cpattern_of_id top_id in
@@ -683,7 +684,7 @@ let elim_intro_tac ipats ?seed what eqid ssrelim is_rec clr =
let name = Ssrcommon.mk_anon_id "K" (Tacmach.New.pf_ids_of_hyps g) in
let new_concl =
- mkProd (Name name, case_ty, mkArrow refl (Vars.lift 2 concl)) in
+ mkProd (make_annot (Name name) Sorts.Relevant, case_ty, mkArrow refl Sorts.Relevant (Vars.lift 2 concl)) in
let erefl, sigma = mkCoqRefl case_ty case env sigma in
Proofview.Unsafe.tclEVARS sigma <*>
Tactics.apply_type ~typecheck:true new_concl [case;erefl]
@@ -707,7 +708,7 @@ let mkEq dir cl c t n env sigma =
eqargs.(Ssrequality.dir_org dir) <- mkRel n;
let eq, sigma = mkCoqEq env sigma in
let refl, sigma = mkCoqRefl t c env sigma in
- mkArrow (mkApp (eq, eqargs)) (Vars.lift 1 cl), refl, sigma
+ mkArrow (mkApp (eq, eqargs)) Sorts.Relevant (Vars.lift 1 cl), refl, sigma
(** in [tac/v: last gens..] the first (last to be run) generalization is
"special" in that is it also the main argument of [tac] and is eventually
@@ -743,7 +744,7 @@ let tclLAST_GEN ~to_ind ((oclr, occ), t) conclusion = tclINDEPENDENTL begin
Ssrcommon.errorstrm Pp.(str "@ can be used with let-ins only")
| Context.Named.Declaration.LocalDef (name, b, ty) ->
Unsafe.tclEVARS sigma <*>
- tclUNIT (true, EConstr.mkLetIn (Name name,b,ty,cl), c, clr)
+ tclUNIT (true, EConstr.mkLetIn (map_annot Name.mk_name name,b,ty,cl), c, clr)
else
Unsafe.tclEVARS sigma <*>
Ssrcommon.tacMKPROD c cl >>= fun ccl ->
@@ -757,7 +758,7 @@ let tclLAST_GEN ~to_ind ((oclr, occ), t) conclusion = tclINDEPENDENTL begin
Unsafe.tclEVARS sigma <*>
Ssrcommon.tacTYPEOF p >>= fun pty ->
(* TODO: check bug: cl0 no lift? *)
- let ccl = EConstr.mkProd (Ssrcommon.constr_name sigma c, pty, cl0) in
+ let ccl = EConstr.mkProd (make_annot (Ssrcommon.constr_name sigma c) Sorts.Relevant, pty, cl0) in
tclUNIT (false, ccl, p, clr)
else
Ssrcommon.errorstrm Pp.(str "generalized term didn't match")
diff --git a/plugins/ssr/ssrtacticals.ml b/plugins/ssr/ssrtacticals.ml
index f12f9fac0f..bbe7bde78b 100644
--- a/plugins/ssr/ssrtacticals.ml
+++ b/plugins/ssr/ssrtacticals.ml
@@ -12,6 +12,7 @@
open Names
open Constr
+open Context
open Termops
open Tacmach
@@ -102,10 +103,10 @@ let endclausestac id_map clseq gl_id cl0 gl =
forced && ids = [] && (not hide_goal || dc' = [] && c_hidden) in
let rec unmark c = match EConstr.kind (project gl) c with
| Var id when hidden_clseq clseq && id = gl_id -> cl0
- | Prod (Name id, t, c') when List.mem_assoc id id_map ->
- EConstr.mkProd (Name (orig_id id), unmark t, unmark c')
- | LetIn (Name id, v, t, c') when List.mem_assoc id id_map ->
- EConstr.mkLetIn (Name (orig_id id), unmark v, unmark t, unmark c')
+ | Prod ({binder_name=Name id} as na, t, c') when List.mem_assoc id id_map ->
+ EConstr.mkProd ({na with binder_name=Name (orig_id id)}, unmark t, unmark c')
+ | LetIn ({binder_name=Name id} as na, v, t, c') when List.mem_assoc id id_map ->
+ EConstr.mkLetIn ({na with binder_name=Name (orig_id id)}, unmark v, unmark t, unmark c')
| _ -> EConstr.map (project gl) unmark c in
let utac hyp =
Proofview.V82.of_tactic
diff --git a/plugins/ssr/ssrview.ml b/plugins/ssr/ssrview.ml
index 2794696017..537fd7d7b4 100644
--- a/plugins/ssr/ssrview.ml
+++ b/plugins/ssr/ssrview.ml
@@ -10,6 +10,7 @@
open Util
open Names
+open Context
open Ltac_plugin
@@ -95,7 +96,7 @@ let vsBOOTSTRAP = Goal.enter_one ~__LOC__ begin fun gl ->
let concl = Goal.concl gl in
let id = (* We keep the orig name for checks in "in" tcl *)
match EConstr.kind_of_type (Goal.sigma gl) concl with
- | Term.ProdType(Name.Name id, _, _)
+ | Term.ProdType({binder_name=Name.Name id}, _, _)
when Ssrcommon.is_discharged_id id -> id
| _ -> mk_anon_id "view_subject" (Tacmach.New.pf_ids_of_hyps gl) in
let view = EConstr.mkVar id in
diff --git a/plugins/ssrmatching/ssrmatching.ml b/plugins/ssrmatching/ssrmatching.ml
index fb99b87108..b83a6a34cb 100644
--- a/plugins/ssrmatching/ssrmatching.ml
+++ b/plugins/ssrmatching/ssrmatching.ml
@@ -16,6 +16,7 @@ open Pp
open Genarg
open Stdarg
open Term
+open Context
module CoqConstr = Constr
open CoqConstr
open Vars
@@ -383,7 +384,7 @@ let evars_for_FO ~hack env sigma0 (ise0:evar_map) c0 =
| Context.Named.Declaration.LocalDef (x, b, t) ->
d, mkNamedLetIn x (put b) (put t) c
| Context.Named.Declaration.LocalAssum (x, t) ->
- mkVar x :: d, mkNamedProd x (put t) c in
+ mkVar x.binder_name :: d, mkNamedProd x (put t) c in
let a, t =
Context.Named.fold_inside abs_dc
~init:([], (put @@ EConstr.Unsafe.to_constr evi.evar_concl))
@@ -548,7 +549,7 @@ let match_upats_FO upats env sigma0 ise orig_c =
if skip || not (closed0 c') then () else try
let _ = match u.up_k with
| KpatFlex ->
- let kludge v = mkLambda (Anonymous, mkProp, v) in
+ let kludge v = mkLambda (make_annot Anonymous Sorts.Relevant, mkProp, v) in
unif_FO env ise (kludge u.up_FO) (kludge c')
| KpatLet ->
let kludge vla =
@@ -1286,7 +1287,7 @@ let ssrpatterntac _ist arg gl =
let t = EConstr.of_constr t in
let concl_x = EConstr.of_constr concl_x in
let gl, tty = pf_type_of gl t in
- let concl = EConstr.mkLetIn (Name (Id.of_string "selected"), t, tty, concl_x) in
+ let concl = EConstr.mkLetIn (make_annot (Name (Id.of_string "selected")) Sorts.Relevant, t, tty, concl_x) in
Proofview.V82.of_tactic (convert_concl concl DEFAULTcast) gl
(* Register "ssrpattern" tactic *)
diff --git a/pretyping/arguments_renaming.ml b/pretyping/arguments_renaming.ml
index 08df9a2460..3b3de33d8e 100644
--- a/pretyping/arguments_renaming.ml
+++ b/pretyping/arguments_renaming.ml
@@ -13,6 +13,7 @@ open Names
open Globnames
open Term
open Constr
+open Context
open Environ
open Util
open Libobject
@@ -72,7 +73,7 @@ let arguments_names r = GlobRef.Map.find r !name_table
let rename_type ty ref =
let name_override old_name override =
match override with
- | Name _ as x -> x
+ | Name _ as x -> {old_name with binder_name=x}
| Anonymous -> old_name in
let rec rename_type_aux c = function
| [] -> c
diff --git a/pretyping/cases.ml b/pretyping/cases.ml
index 069ba9572a..e22368d5e5 100644
--- a/pretyping/cases.ml
+++ b/pretyping/cases.ml
@@ -16,6 +16,7 @@ open Util
open Names
open Nameops
open Constr
+open Context
open Termops
open Environ
open EConstr
@@ -472,7 +473,8 @@ let push_current_pattern ~program_mode sigma (cur,ty) eqn =
let hypnaming = if program_mode then ProgramNaming else KeepUserNameAndRenameExistingButSectionNames in
match eqn.patterns with
| pat::pats ->
- let _,rhs_env = push_rel ~hypnaming sigma (LocalDef (alias_of_pat pat,cur,ty)) eqn.rhs.rhs_env in
+ let r = Sorts.Relevant in (* TODO relevance *)
+ let _,rhs_env = push_rel ~hypnaming sigma (LocalDef (make_annot (alias_of_pat pat) r,cur,ty)) eqn.rhs.rhs_env in
{ eqn with
rhs = { eqn.rhs with rhs_env = rhs_env };
patterns = pats }
@@ -762,7 +764,10 @@ let get_names avoid env sigma sign eqns =
(fun (l,avoid) d na ->
let na =
merge_name
- (fun (LocalAssum (na,t) | LocalDef (na,_,t)) -> Name (next_name_away (named_hd env sigma t na) avoid))
+ (fun decl ->
+ let na = get_name decl in
+ let t = get_type decl in
+ Name (next_name_away (named_hd env sigma t na) avoid))
d na
in
(na::l,Id.Set.add (Name.get_id na) avoid))
@@ -782,10 +787,10 @@ let recover_and_adjust_alias_names (_,avoid) names sign =
let rec aux = function
| [],[] ->
[]
- | x::names, LocalAssum (_,t)::sign ->
- (x, LocalAssum (alias_of_pat x,t)) :: aux (names,sign)
+ | x::names, LocalAssum (x',t)::sign ->
+ (x, LocalAssum ({x' with binder_name=alias_of_pat x},t)) :: aux (names,sign)
| names, (LocalDef (na,_,_) as decl)::sign ->
- (DAst.make @@ PatVar na, decl) :: aux (names,sign)
+ (DAst.make @@ PatVar na.binder_name, decl) :: aux (names,sign)
| _ -> assert false
in
List.split (aux (names,sign))
@@ -1247,7 +1252,7 @@ let rec generalize_problem names sigma pb = function
let pb',deps = generalize_problem names sigma pb l in
let d = map_constr (lift i) (lookup_rel i !!(pb.env)) in
begin match d with
- | LocalDef (Anonymous,_,_) -> pb', deps
+ | LocalDef ({binder_name=Anonymous},_,_) -> pb', deps
| _ ->
(* for better rendering *)
let d = RelDecl.map_type (fun c -> whd_betaiota sigma c) d in
@@ -1436,16 +1441,15 @@ let compile ~program_mode sigma pb =
brvals pb.tomatch pb.pred deps cstrs in
let brvals = Array.map (fun (sign,body) ->
it_mkLambda_or_LetIn body sign) brvals in
- let (pred,typ) =
+ let (pred,typ) =
find_predicate pb.caseloc pb.env sigma
- pred current indt (names,dep) tomatch in
- let ci = make_case_info !!(pb.env) (fst mind) pb.casestyle in
+ pred current indt (names,dep) tomatch
+ in
+ let rci = Typing.check_allowed_sort !!(pb.env) sigma mind current pred in
+ let ci = make_case_info !!(pb.env) (fst mind) rci pb.casestyle in
let pred = nf_betaiota !!(pb.env) sigma pred in
- let case =
- make_case_or_project !!(pb.env) sigma indf ci pred current brvals
- in
+ let case = make_case_or_project !!(pb.env) sigma indf ci pred current brvals in
let sigma, _ = Typing.type_of !!(pb.env) sigma pred in
- Typing.check_allowed_sort !!(pb.env) sigma mind current pred;
sigma, { uj_val = applist (case, inst);
uj_type = prod_applist sigma typ inst }
@@ -1460,7 +1464,7 @@ let compile ~program_mode sigma pb =
let hypnaming = if program_mode then ProgramNaming else KeepUserNameAndRenameExistingButSectionNames in
let pb =
{ pb with
- env = snd (push_rel ~hypnaming sigma (LocalDef (na,current,ty)) env);
+ env = snd (push_rel ~hypnaming sigma (LocalDef (annotR na,current,ty)) env);
tomatch = tomatch;
pred = lift_predicate 1 pred tomatch;
history = pop_history pb.history;
@@ -1511,7 +1515,8 @@ let compile ~program_mode sigma pb =
and compile_alias initial sigma pb (na,orig,(expanded,expanded_typ)) rest =
let hypnaming = if program_mode then ProgramNaming else KeepUserNameAndRenameExistingButSectionNames in
let f c t =
- let alias = LocalDef (na,c,t) in
+ let r = Retyping.relevance_of_type !!(pb.env) sigma t in
+ let alias = LocalDef (make_annot na r,c,t) in
let pb =
{ pb with
env = snd (push_rel ~hypnaming sigma alias pb.env);
@@ -1524,7 +1529,7 @@ let compile ~program_mode sigma pb =
if isRel sigma c || isVar sigma c || count_occurrences sigma (mkRel 1) j.uj_val <= 1 then
subst1 c j.uj_val
else
- mkLetIn (na,c,t,j.uj_val);
+ mkLetIn (make_annot na r,c,t,j.uj_val);
uj_type = subst1 c j.uj_type } in
(* spiwack: when an alias appears on a deep branch, its non-expanded
form is automatically a variable of the same name. We avoid
@@ -1812,7 +1817,7 @@ let build_inversion_problem ~program_mode loc env sigma tms t =
List.rev_append patl patl',acc_sign,acc
| (t, NotInd (bo,typ)) :: tms ->
let pat,acc = make_patvar t acc in
- let d = LocalAssum (alias_of_pat pat,typ) in
+ let d = LocalAssum (annotR (alias_of_pat pat),typ) in
let patl,acc_sign,acc = aux (n+1) (snd (push_rel ~hypnaming:KeepUserNameAndRenameExistingButSectionNames sigma d env)) (d::acc_sign) tms acc in
pat::patl,acc_sign,acc in
let avoid0 = GlobEnv.vars_of_env env in
@@ -1913,9 +1918,11 @@ let extract_arity_signature ?(dolift=true) env0 tomatchl tmsign =
match tm with
| NotInd (bo,typ) ->
(match t with
- | None -> let sign = match bo with
- | None -> [LocalAssum (na, lift n typ)]
- | Some b -> [LocalDef (na, lift n b, lift n typ)] in sign
+ | None ->
+ let r = Sorts.Relevant in (* TODO relevance *)
+ let sign = match bo with
+ | None -> [LocalAssum (make_annot na r, lift n typ)]
+ | Some b -> [LocalDef (make_annot na r, lift n b, lift n typ)] in sign
| Some {CAst.loc} ->
user_err ?loc
(str"Unexpected type annotation for a term of non inductive type."))
@@ -1923,7 +1930,7 @@ let extract_arity_signature ?(dolift=true) env0 tomatchl tmsign =
let indf' = if dolift then lift_inductive_family n indf else indf in
let ((ind,u),_) = dest_ind_family indf' in
let nrealargs_ctxt = inductive_nrealdecls_env env0 ind in
- let arsign = fst (get_arity env0 indf') in
+ let arsign, inds = get_arity env0 indf' in
let arsign = List.map (fun d -> map_rel_decl EConstr.of_constr d) arsign in
let realnal =
match t with
@@ -1935,8 +1942,9 @@ let extract_arity_signature ?(dolift=true) env0 tomatchl tmsign =
List.rev realnal
| None ->
List.make nrealargs_ctxt Anonymous in
+ let r = Sorts.relevance_of_sort_family inds in
let t = EConstr.of_constr (build_dependent_inductive env0 indf') in
- LocalAssum (na, t) :: List.map2 RelDecl.set_name realnal arsign in
+ LocalAssum (make_annot na r, t) :: List.map2 RelDecl.set_name realnal arsign in
let rec buildrec n = function
| [],[] -> []
| (_,tm)::ltm, (_,x)::tmsign ->
@@ -2143,9 +2151,10 @@ let constr_of_pat env sigma arsign pat avoid =
| Anonymous ->
let previd, id = prime avoid (Name (Id.of_string "wildcard")) in
Name id, Id.Set.add id avoid
- in
- (sigma, (DAst.make ?loc @@ PatVar name), [LocalAssum (name, ty)] @ realargs, mkRel 1, ty,
- (List.map (fun x -> mkRel 1) realargs), 1, avoid)
+ in
+ let r = Sorts.Relevant in (* TODO relevance *)
+ (sigma, (DAst.make ?loc @@ PatVar name), [LocalAssum (make_annot name r, ty)] @ realargs, mkRel 1, ty,
+ (List.map (fun x -> mkRel 1) realargs), 1, avoid)
| PatCstr (((_, i) as cstr),args,alias) ->
let cind = inductive_of_constructor cstr in
let IndType (indf, _) =
@@ -2184,9 +2193,11 @@ let constr_of_pat env sigma arsign pat avoid =
match alias with
Anonymous ->
sigma, pat', sign, app, apptype, realargs, n, avoid
- | Name id ->
- let sign = LocalAssum (alias, lift m ty) :: sign in
- let avoid = Id.Set.add id avoid in
+ | Name id ->
+ let _, inds = get_arity env indf in
+ let r = Sorts.relevance_of_sort_family inds in
+ let sign = LocalAssum (make_annot alias r, lift m ty) :: sign in
+ let avoid = Id.Set.add id avoid in
let sigma, sign, i, avoid =
try
let env = EConstr.push_rel_context sign env in
@@ -2196,11 +2207,12 @@ let constr_of_pat env sigma arsign pat avoid =
(mkRel 1) (* alias *)
(lift 1 app) (* aliased term *)
in
- let neq = eq_id avoid id in
- sigma, LocalDef (Name neq, mkRel 0, eq_t) :: sign, 2, Id.Set.add neq avoid
+ let neq = eq_id avoid id in
+ (* if we ever allow using a SProp-typed coq_eq_ind this relevance will be wrong *)
+ sigma, LocalDef (nameR neq, mkRel 0, eq_t) :: sign, 2, Id.Set.add neq avoid
with Evarconv.UnableToUnify _ -> sigma, sign, 1, avoid
in
- (* Mark the equality as a hole *)
+ (* Mark the equality as a hole *)
sigma, pat', sign, lift i app, lift i apptype, realargs, n + i, avoid
in
let sigma, pat', sign, patc, patty, args, z, avoid = typ env sigma (RelDecl.get_type (List.hd arsign), List.tl arsign) pat avoid in
@@ -2222,18 +2234,18 @@ match EConstr.kind sigma t with
let rels_of_patsign sigma =
List.map (fun decl ->
match decl with
- | LocalDef (na,t',t) when is_topvar sigma t' -> LocalAssum (na,t)
+ | LocalDef (na,t',t) when is_topvar sigma t' -> LocalAssum (na,t)
| _ -> decl)
let vars_of_ctx sigma ctx =
let _, y =
List.fold_right (fun decl (prev, vars) ->
match decl with
- | LocalDef (na,t',t) when is_topvar sigma t' ->
+ | LocalDef (na,t',t) when is_topvar sigma t' ->
prev,
(DAst.make @@ GApp (
(DAst.make @@ GRef (delayed_force coq_eq_refl_ref, None)),
- [hole na; DAst.make @@ GVar prev])) :: vars
+ [hole na.binder_name; DAst.make @@ GVar prev])) :: vars
| _ ->
match RelDecl.get_name decl with
Anonymous -> invalid_arg "vars_of_ctx"
@@ -2343,12 +2355,13 @@ let constrs_of_pats typing_fun env sigma eqns tomatchs sign neqs arity =
let args = List.rev args in
substl args (liftn signlen (succ nargs) arity)
in
- let rhs_rels', tycon =
+ let r = Sorts.Relevant in (* TODO relevance *)
+ let rhs_rels', tycon =
let neqs_rels, arity =
match ineqs with
| None -> [], arity
| Some ineqs ->
- [LocalAssum (Anonymous, ineqs)], lift 1 arity
+ [LocalAssum (make_annot Anonymous r, ineqs)], lift 1 arity
in
let eqs_rels, arity = decompose_prod_n_assum sigma neqs arity in
eqs_rels @ neqs_rels @ rhs_rels', arity
@@ -2359,7 +2372,7 @@ let constrs_of_pats typing_fun env sigma eqns tomatchs sign neqs arity =
and btype = it_mkProd_or_LetIn j.uj_type rhs_rels' in
let sigma, _btype = Typing.type_of !!env sigma bbody in
let branch_name = Id.of_string ("program_branch_" ^ (string_of_int !i)) in
- let branch_decl = LocalDef (Name branch_name, lift !i bbody, lift !i btype) in
+ let branch_decl = LocalDef (make_annot (Name branch_name) r, lift !i bbody, lift !i btype) in
let branch =
let bref = DAst.make @@ GVar branch_name in
match vars_of_ctx sigma rhs_rels with
@@ -2407,9 +2420,10 @@ let abstract_tomatch env sigma tomatchs tycon =
| _ ->
let tycon = Option.map
(fun t -> subst_term sigma (lift 1 c) (lift 1 t)) tycon in
- let name = next_ident_away (Id.of_string "filtered_var") names in
+ let name = next_ident_away (Id.of_string "filtered_var") names in
+ let r = Sorts.Relevant in (* TODO relevance *)
(mkRel 1, lift_tomatch_type (succ lenctx) t) :: lift_ctx 1 prev,
- LocalDef (Name name, lift lenctx c, lift lenctx $ type_of_tomatch t) :: ctx,
+ LocalDef (make_annot (Name name) r, lift lenctx c, lift lenctx $ type_of_tomatch t) :: ctx,
Id.Set.add name names, tycon)
([], [], Id.Set.empty, tycon) tomatchs
in List.rev prev, ctx, tycon
@@ -2471,8 +2485,8 @@ let build_dependent_signature env sigma avoid tomatchs arsign =
make_prime avoid name
in
(sigma, env, succ nargeqs,
- (LocalAssum (Name (eq_id avoid previd), eq)) :: argeqs,
- refl_arg :: refl_args,
+ (LocalAssum (make_annot (Name (eq_id avoid previd)) Sorts.Relevant, eq)) :: argeqs,
+ refl_arg :: refl_args,
pred slift,
RelDecl.set_name (Name id) decl :: argsign'))
(sigma, env, neqs, [], [], slift, []) args argsign
@@ -2486,8 +2500,8 @@ let build_dependent_signature env sigma avoid tomatchs arsign =
in
let sigma, refl_eq = mk_JMeq_refl sigma ty tm in
let previd, id = make_prime avoid appn in
- (sigma, (LocalAssum (Name (eq_id avoid previd), eq) :: argeqs) :: eqs,
- succ nargeqs,
+ (sigma, (LocalAssum (make_annot (Name (eq_id avoid previd)) Sorts.Relevant, eq) :: argeqs) :: eqs,
+ succ nargeqs,
refl_eq :: refl_args,
pred slift,
((RelDecl.set_name (Name id) app_decl :: argsign') :: arsigns))
@@ -2503,8 +2517,9 @@ let build_dependent_signature env sigma avoid tomatchs arsign =
(mkRel slift) (lift nar tm)
in
let sigma, refl = mk_eq_refl sigma tomatch_ty tm in
+ let na = make_annot (Name (eq_id avoid previd)) Sorts.Relevant in
(sigma,
- [LocalAssum (Name (eq_id avoid previd), eq)] :: eqs, succ neqs,
+ [LocalAssum (na, eq)] :: eqs, succ neqs,
refl :: refl_args,
pred slift, (arsign' :: []) :: arsigns))
(sigma, [], 0, [], nar, []) tomatchs arsign
@@ -2580,11 +2595,12 @@ let compile_program_cases ?loc style (typing_function, sigma) tycon env
(* We push the initial terms to match and push their alias to rhs' envs *)
(* names of aliases will be recovered from patterns (hence Anonymous here) *)
- let out_tmt na = function NotInd (None,t) -> LocalAssum (na,t)
- | NotInd (Some b, t) -> LocalDef (na,b,t)
- | IsInd (typ,_,_) -> LocalAssum (na,typ) in
+ (* TODO relevance *)
+ let out_tmt na = function NotInd (None,t) -> LocalAssum (make_annot na Sorts.Relevant,t)
+ | NotInd (Some b, t) -> LocalDef (make_annot na Sorts.Relevant,b,t)
+ | IsInd (typ,_,_) -> LocalAssum (make_annot na Sorts.Relevant,typ) in
let typs = List.map2 (fun (na,_) (tm,tmt) -> (tm,out_tmt na tmt)) nal tomatchs in
-
+
let typs =
List.map (fun (c,d) -> (c,extract_inductive_data !!env sigma d,d)) typs in
@@ -2654,10 +2670,11 @@ let compile_cases ?loc ~program_mode style (typing_fun, sigma) tycon env (predop
(* names of aliases will be recovered from patterns (hence Anonymous *)
(* here) *)
+ (* TODO relevance *)
let out_tmt na = function NotInd (None,t) -> LocalAssum (na,t)
| NotInd (Some b,t) -> LocalDef (na,b,t)
| IsInd (typ,_,_) -> LocalAssum (na,typ) in
- let typs = List.map2 (fun (na,_) (tm,tmt) -> (tm,out_tmt na tmt)) nal tomatchs in
+ let typs = List.map2 (fun (na,_) (tm,tmt) -> (tm,out_tmt (make_annot na Sorts.Relevant) tmt)) nal tomatchs in
let typs =
List.map (fun (c,d) -> (c,extract_inductive_data !!env sigma d,d)) typs in
diff --git a/pretyping/cbv.ml b/pretyping/cbv.ml
index e27fc536eb..c9f18d89be 100644
--- a/pretyping/cbv.ml
+++ b/pretyping/cbv.ml
@@ -48,7 +48,7 @@ type cbv_value =
| VAL of int * constr
| STACK of int * cbv_value * cbv_stack
| CBN of constr * cbv_value subs
- | LAM of int * (Name.t * constr) list * constr * cbv_value subs
+ | LAM of int * (Name.t Context.binder_annot * constr) list * constr * cbv_value subs
| FIXP of fixpoint * cbv_value subs * cbv_value array
| COFIXP of cofixpoint * cbv_value subs * cbv_value array
| CONSTR of constructor Univ.puniverses * cbv_value array
@@ -281,11 +281,11 @@ and reify_value = function (* reduction under binders *)
apply_env env @@
List.fold_left (fun c (n,t) ->
mkLambda (n, t, c)) b ctxt
- | FIXP ((lij,(names,lty,bds)),env,args) ->
- let fix = mkFix (lij, (names, lty, bds)) in
+ | FIXP ((lij,fix),env,args) ->
+ let fix = mkFix (lij, fix) in
mkApp (apply_env env fix, Array.map reify_value args)
- | COFIXP ((j,(names,lty,bds)),env,args) ->
- let cofix = mkCoFix (j, (names,lty,bds)) in
+ | COFIXP ((j,cofix),env,args) ->
+ let cofix = mkCoFix (j, cofix) in
mkApp (apply_env env cofix, Array.map reify_value args)
| CONSTR (c,args) ->
mkApp(mkConstructU c, Array.map reify_value args)
@@ -550,7 +550,7 @@ and cbv_norm_value info = function (* reduction under binders *)
| FIXP ((lij,(names,lty,bds)),env,args) ->
mkApp
(mkFix (lij,
- (names,
+ (names,
Array.map (cbv_norm_term info env) lty,
Array.map (cbv_norm_term info
(subs_liftn (Array.length lty) env)) bds)),
@@ -558,7 +558,7 @@ and cbv_norm_value info = function (* reduction under binders *)
| COFIXP ((j,(names,lty,bds)),env,args) ->
mkApp
(mkCoFix (j,
- (names,Array.map (cbv_norm_term info env) lty,
+ (names,Array.map (cbv_norm_term info env) lty,
Array.map (cbv_norm_term info
(subs_liftn (Array.length lty) env)) bds)),
Array.map (cbv_norm_value info) args)
diff --git a/pretyping/cbv.mli b/pretyping/cbv.mli
index 0a1e771921..d6c2ad146e 100644
--- a/pretyping/cbv.mli
+++ b/pretyping/cbv.mli
@@ -32,7 +32,7 @@ type cbv_value =
| VAL of int * constr
| STACK of int * cbv_value * cbv_stack
| CBN of constr * cbv_value subs
- | LAM of int * (Name.t * constr) list * constr * cbv_value subs
+ | LAM of int * (Name.t Context.binder_annot * constr) list * constr * cbv_value subs
| FIXP of fixpoint * cbv_value subs * cbv_value array
| COFIXP of cofixpoint * cbv_value subs * cbv_value array
| CONSTR of constructor Univ.puniverses * cbv_value array
diff --git a/pretyping/classops.ml b/pretyping/classops.ml
index 306a76e35e..54a1aa9aa0 100644
--- a/pretyping/classops.ml
+++ b/pretyping/classops.ml
@@ -179,7 +179,7 @@ let find_class_type sigma t =
| Proj (p, c) when not (Projection.unfolded p) ->
CL_PROJ (Projection.repr p), EInstance.empty, (c :: args)
| Ind (ind_sp,u) -> CL_IND ind_sp, u, args
- | Prod (_,_,_) -> CL_FUN, EInstance.empty, []
+ | Prod _ -> CL_FUN, EInstance.empty, []
| Sort _ -> CL_SORT, EInstance.empty, []
| _ -> raise Not_found
diff --git a/pretyping/coercion.ml b/pretyping/coercion.ml
index aa3fbc4577..82411ba2ef 100644
--- a/pretyping/coercion.ml
+++ b/pretyping/coercion.ml
@@ -21,6 +21,7 @@ open Util
open Names
open Term
open Constr
+open Context
open Environ
open EConstr
open Vars
@@ -175,23 +176,23 @@ and coerce ?loc env evdref (x : EConstr.constr) (y : EConstr.constr)
try evdref := unify_leq_delay env !evdref hdx hdy;
let (n, eqT), restT = dest_prod typ in
let (n', eqT'), restT' = dest_prod typ' in
- aux (hdx :: tele) (subst1 hdx restT) (subst1 hdy restT') (succ i) co
+ aux (hdx :: tele) (subst1 hdx restT) (subst1 hdy restT') (succ i) co
with UnableToUnify _ ->
- let (n, eqT), restT = dest_prod typ in
+ let (n, eqT), restT = dest_prod typ in
let (n', eqT'), restT' = dest_prod typ' in
let () =
try evdref := unify_leq_delay env !evdref eqT eqT'
with UnableToUnify _ -> raise NoSubtacCoercion
- in
+ in
(* Disallow equalities on arities *)
if Reductionops.is_arity env !evdref eqT then raise NoSubtacCoercion;
let restargs = lift_args 1
(List.rev (Array.to_list (Array.sub l (succ i) (len - (succ i)))))
in
let args = List.rev (restargs @ mkRel 1 :: List.map (lift 1) tele) in
- let pred = mkLambda (n, eqT, applist (lift 1 c, args)) in
+ let pred = mkLambda (n, eqT, applist (lift 1 c, args)) in
let eq = papp evdref coq_eq_ind [| eqT; hdx; hdy |] in
- let evar = make_existential ?loc n env evdref eq in
+ let evar = make_existential ?loc n.binder_name env evdref eq in
let eq_app x = papp evdref coq_eq_rect
[| eqT; hdx; pred; x; hdy; evar|]
in
@@ -216,9 +217,12 @@ and coerce ?loc env evdref (x : EConstr.constr) (y : EConstr.constr)
| _ -> subco ())
| Prod (name, a, b), Prod (name', a', b') ->
let name' =
- Name (Namegen.next_ident_away Namegen.default_dependent_ident (Termops.vars_of_env env))
- in
- let env' = push_rel (LocalAssum (name', a')) env in
+ {name' with
+ binder_name =
+ Name (Namegen.next_ident_away
+ Namegen.default_dependent_ident (Termops.vars_of_env env))}
+ in
+ let env' = push_rel (LocalAssum (name', a')) env in
let c1 = coerce_unify env' (lift 1 a') (lift 1 a) in
(* env, x : a' |- c1 : lift 1 a' > lift 1 a *)
let coec1 = app_opt env' evdref c1 (mkRel 1) in
@@ -230,7 +234,7 @@ and coerce ?loc env evdref (x : EConstr.constr) (y : EConstr.constr)
| _, _ ->
Some
(fun f ->
- mkLambda (name', a',
+ mkLambda (name', a',
app_opt env' evdref c2
(mkApp (lift 1 f, [| coec1 |])))))
@@ -253,11 +257,11 @@ and coerce ?loc env evdref (x : EConstr.constr) (y : EConstr.constr)
let c1 = coerce_unify env a a' in
let remove_head a c =
match EConstr.kind !evdref c with
- | Lambda (n, t, t') -> c, t'
+ | Lambda (n, t, t') -> c, t'
| Evar (k, args) ->
let (evs, t) = Evardefine.define_evar_as_lambda env !evdref (k,args) in
evdref := evs;
- let (n, dom, rng) = destLambda !evdref t in
+ let (n, dom, rng) = destLambda !evdref t in
if isEvar !evdref dom then
let (domk, args) = destEvar !evdref dom in
evdref := define domk a !evdref;
@@ -265,8 +269,12 @@ and coerce ?loc env evdref (x : EConstr.constr) (y : EConstr.constr)
t, rng
| _ -> raise NoSubtacCoercion
in
- let (pb, b), (pb', b') = remove_head a pb, remove_head a' pb' in
- let env' = push_rel (LocalAssum (Name Namegen.default_dependent_ident, a)) env in
+ let (pb, b), (pb', b') = remove_head a pb, remove_head a' pb' in
+ let ra = Retyping.relevance_of_type env !evdref a in
+ let env' = push_rel
+ (LocalAssum (make_annot (Name Namegen.default_dependent_ident) ra, a))
+ env
+ in
let c2 = coerce_unify env' b b' in
match c1, c2 with
| None, None -> None
@@ -396,7 +404,7 @@ let apply_coercion env sigma p hj typ_cl =
let inh_app_fun_core ~program_mode env evd j =
let t = whd_all env evd j.uj_type in
match EConstr.kind evd t with
- | Prod (_,_,_) -> (evd,j)
+ | Prod _ -> (evd,j)
| Evar ev ->
let (evd',t) = Evardefine.define_evar_as_product env evd ev in
(evd',{ uj_val = j.uj_val; uj_type = t })
@@ -505,11 +513,11 @@ let rec inh_conv_coerce_to_fail ?loc env evd ?(flags=default_flags_of env) rigid
(* has type forall (x:u1), u2 (with v' recursively obtained) *)
(* Note: we retype the term because template polymorphism may have *)
(* weakened its type *)
- let name = match name with
+ let name = map_annot (function
| Anonymous -> Name Namegen.default_dependent_ident
- | _ -> name in
+ | na -> na) name in
let open Context.Rel.Declaration in
- let env1 = push_rel (LocalAssum (name,u1)) env in
+ let env1 = push_rel (LocalAssum (name,u1)) env in
let (evd', v1) =
inh_conv_coerce_to_fail ?loc env1 evd rigidonly
(Some (mkRel 1)) (lift 1 u1) (lift 1 t1) in
@@ -519,7 +527,7 @@ let rec inh_conv_coerce_to_fail ?loc env evd ?(flags=default_flags_of env) rigid
| None -> subst_term evd' v1 t2
| Some v2 -> Retyping.get_type_of env1 evd' v2 in
let (evd'',v2') = inh_conv_coerce_to_fail ?loc env1 evd' rigidonly v2 t2 u2 in
- (evd'', Option.map (fun v2' -> mkLambda (name, u1, v2')) v2')
+ (evd'', Option.map (fun v2' -> mkLambda (name, u1, v2')) v2')
| _ -> raise (NoCoercionNoUnifier (best_failed_evd,e))
(* Look for cj' obtained from cj by inserting coercions, s.t. cj'.typ = t *)
diff --git a/pretyping/constr_matching.ml b/pretyping/constr_matching.ml
index 94257fedd7..bc083ed9d9 100644
--- a/pretyping/constr_matching.ml
+++ b/pretyping/constr_matching.ml
@@ -14,6 +14,7 @@ open CErrors
open Util
open Names
open Constr
+open Context
open Globnames
open Termops
open Term
@@ -70,7 +71,7 @@ let constrain sigma n (ids, m) ((names,seen as names_seen), terms as subst) =
(names_seen, Id.Map.add n (ids, m) terms)
let add_binders na1 na2 binding_vars ((names,seen), terms as subst) =
- match na1, na2 with
+ match na1, na2.binder_name with
| Name id1, Name id2 when Id.Set.mem id1 binding_vars ->
if Id.Map.mem id1 names then
let () = Glob_ops.warn_variable_collision id1 in
@@ -94,7 +95,7 @@ let rec build_lambda sigma vars ctx m = match vars with
let (na, t, suf) = match suf with
| [] -> assert false
| (_, id, t) :: suf ->
- (Name id, t, suf)
+ (map_annot Name.mk_name id, t, suf)
in
(* Check that the abstraction is legal by generating a transitive closure of
its dependencies. *)
@@ -178,11 +179,12 @@ let make_renaming ids = function
| _ -> dummy_constr
let push_binder na1 na2 t ctx =
- let id2 = match na2 with
- | Name id2 -> id2
- | Anonymous ->
- let avoid = Id.Set.of_list (List.map pi2 ctx) in
- Namegen.next_ident_away Namegen.default_non_dependent_ident avoid in
+ let id2 = map_annot (function
+ | Name id2 -> id2
+ | Anonymous ->
+ let avoid = Id.Set.of_list (List.map (fun (_,id,_) -> id.binder_name) ctx) in
+ Namegen.next_ident_away Namegen.default_non_dependent_ident avoid) na2
+ in
(na1, id2, t) :: ctx
(* This is an optimization of the main pattern-matching which shares
@@ -341,19 +343,19 @@ let matches_core env sigma allow_bound_rels
sorec ctx env subst c1 c2
| PProd (na1,c1,d1), Prod(na2,c2,d2) ->
- sorec (push_binder na1 na2 c2 ctx) (EConstr.push_rel (LocalAssum (na2,c2)) env)
+ sorec (push_binder na1 na2 c2 ctx) (EConstr.push_rel (LocalAssum (na2,c2)) env)
(add_binders na1 na2 binding_vars (sorec ctx env subst c1 c2)) d1 d2
| PLambda (na1,c1,d1), Lambda(na2,c2,d2) ->
- sorec (push_binder na1 na2 c2 ctx) (EConstr.push_rel (LocalAssum (na2,c2)) env)
+ sorec (push_binder na1 na2 c2 ctx) (EConstr.push_rel (LocalAssum (na2,c2)) env)
(add_binders na1 na2 binding_vars (sorec ctx env subst c1 c2)) d1 d2
| PLetIn (na1,c1,Some t1,d1), LetIn(na2,c2,t2,d2) ->
- sorec (push_binder na1 na2 t2 ctx) (EConstr.push_rel (LocalDef (na2,c2,t2)) env)
+ sorec (push_binder na1 na2 t2 ctx) (EConstr.push_rel (LocalDef (na2,c2,t2)) env)
(add_binders na1 na2 binding_vars (sorec ctx env (sorec ctx env subst c1 c2) t1 t2)) d1 d2
| PLetIn (na1,c1,None,d1), LetIn(na2,c2,t2,d2) ->
- sorec (push_binder na1 na2 t2 ctx) (EConstr.push_rel (LocalDef (na2,c2,t2)) env)
+ sorec (push_binder na1 na2 t2 ctx) (EConstr.push_rel (LocalDef (na2,c2,t2)) env)
(add_binders na1 na2 binding_vars (sorec ctx env subst c1 c2)) d1 d2
| PIf (a1,b1,b1'), Case (ci,_,a2,[|b2;b2'|]) ->
diff --git a/pretyping/detyping.ml b/pretyping/detyping.ml
index 48e4b5414e..ac7c3d30d5 100644
--- a/pretyping/detyping.ml
+++ b/pretyping/detyping.ml
@@ -13,6 +13,7 @@ open CErrors
open Util
open Names
open Constr
+open Context
open Term
open EConstr
open Vars
@@ -40,9 +41,12 @@ let print_evar_arguments = ref false
let add_name na b t (nenv, env) =
let open Context.Rel.Declaration in
+ (* Is this just a dummy? Be careful, printing doesn't always give us
+ a correct env. *)
+ let r = Sorts.Relevant in
add_name na nenv, push_rel (match b with
- | None -> LocalAssum (na,t)
- | Some b -> LocalDef (na,b,t)
+ | None -> LocalAssum (make_annot na r,t)
+ | Some b -> LocalDef (make_annot na r,b,t)
)
env
@@ -202,11 +206,11 @@ let computable sigma p k =
let lookup_name_as_displayed env sigma t s =
let rec lookup avoid n c = match EConstr.kind sigma c with
| Prod (name,_,c') ->
- (match compute_displayed_name_in sigma RenamingForGoal avoid name c' with
+ (match compute_displayed_name_in sigma RenamingForGoal avoid name.binder_name c' with
| (Name id,avoid') -> if Id.equal id s then Some n else lookup avoid' (n+1) c'
| (Anonymous,avoid') -> lookup avoid' (n+1) (pop c'))
| LetIn (name,_,_,c') ->
- (match compute_displayed_name_in sigma RenamingForGoal avoid name c' with
+ (match compute_displayed_name_in sigma RenamingForGoal avoid name.binder_name c' with
| (Name id,avoid') -> if Id.equal id s then Some n else lookup avoid' (n+1) c'
| (Anonymous,avoid') -> lookup avoid' (n+1) (pop c'))
| Cast (c,_,_) -> lookup avoid n c
@@ -216,7 +220,7 @@ let lookup_name_as_displayed env sigma t s =
let lookup_index_as_renamed env sigma t n =
let rec lookup n d c = match EConstr.kind sigma c with
| Prod (name,_,c') ->
- (match compute_displayed_name_in sigma RenamingForGoal Id.Set.empty name c' with
+ (match compute_displayed_name_in sigma RenamingForGoal Id.Set.empty name.binder_name c' with
(Name _,_) -> lookup n (d+1) c'
| (Anonymous,_) ->
if Int.equal n 0 then
@@ -226,7 +230,7 @@ let lookup_index_as_renamed env sigma t n =
else
lookup (n-1) (d+1) c')
| LetIn (name,_,_,c') ->
- (match compute_displayed_name_in sigma RenamingForGoal Id.Set.empty name c' with
+ (match compute_displayed_name_in sigma RenamingForGoal Id.Set.empty name.binder_name c' with
| (Name _,_) -> lookup n (d+1) c'
| (Anonymous,_) ->
if Int.equal n 0 then
@@ -342,9 +346,9 @@ let rec decomp_branch tags nal b (avoid,env as e) sigma c =
| b::tags ->
let na,c,f,body,t =
match EConstr.kind sigma (strip_outer_cast sigma c), b with
- | Lambda (na,t,c),false -> na,c,compute_displayed_let_name_in,None,Some t
- | LetIn (na,b,t,c),true ->
- na,c,compute_displayed_name_in,Some b,Some t
+ | Lambda (na,t,c),false -> na.binder_name,c,compute_displayed_let_name_in,None,Some t
+ | LetIn (na,b,t,c),true ->
+ na.binder_name,c,compute_displayed_name_in,Some b,Some t
| _, false ->
Name default_dependent_ident,(applist (lift 1 c, [mkRel 1])),
compute_displayed_name_in,None,None
@@ -490,19 +494,16 @@ let rec share_names detype n l avoid env sigma c t =
match EConstr.kind sigma c, EConstr.kind sigma t with
(* factorize even when not necessary to have better presentation *)
| Lambda (na,t,c), Prod (na',t',c') ->
- let na = match (na,na') with
- Name _, _ -> na
- | _, Name _ -> na'
- | _ -> na in
+ let na = Nameops.Name.pick_annot na na' in
let t' = detype avoid env sigma t in
- let id = next_name_away na avoid in
+ let id = next_name_away na.binder_name avoid in
let avoid = Id.Set.add id avoid and env = add_name (Name id) None t env in
share_names detype (n-1) ((Name id,Explicit,None,t')::l) avoid env sigma c c'
(* May occur for fix built interactively *)
| LetIn (na,b,t',c), _ when n > 0 ->
let t'' = detype avoid env sigma t' in
let b' = detype avoid env sigma b in
- let id = next_name_away na avoid in
+ let id = next_name_away na.binder_name avoid in
let avoid = Id.Set. add id avoid and env = add_name (Name id) (Some b) t' env in
share_names detype n ((Name id,Explicit,Some b',t'')::l) avoid env sigma c (lift 1 t)
(* Only if built with the f/n notation or w/o let-expansion in types *)
@@ -511,7 +512,7 @@ let rec share_names detype n l avoid env sigma c t =
(* If it is an open proof: we cheat and eta-expand *)
| _, Prod (na',t',c') when n > 0 ->
let t'' = detype avoid env sigma t' in
- let id = next_name_away na' avoid in
+ let id = next_name_away na'.binder_name avoid in
let avoid = Id.Set.add id avoid and env = add_name (Name id) None t' env in
let appc = mkApp (lift 1 c,[|mkRel 1|]) in
share_names detype (n-1) ((Name id,Explicit,None,t'')::l) avoid env sigma appc c'
@@ -549,7 +550,7 @@ let detype_fix detype avoid env sigma (vn,_ as nvn) (names,tys,bodies) =
let def_avoid, def_env, lfi =
Array.fold_left2
(fun (avoid, env, l) na ty ->
- let id = next_name_away na avoid in
+ let id = next_name_away na.binder_name avoid in
(Id.Set.add id avoid, add_name (Name id) None ty env, id::l))
(avoid, env, []) names tys in
let n = Array.length tys in
@@ -565,7 +566,7 @@ let detype_cofix detype avoid env sigma n (names,tys,bodies) =
let def_avoid, def_env, lfi =
Array.fold_left2
(fun (avoid, env, l) na ty ->
- let id = next_name_away na avoid in
+ let id = next_name_away na.binder_name avoid in
(Id.Set.add id avoid, add_name (Name id) None ty env, id::l))
(avoid, env, []) names tys in
let ntys = Array.length tys in
@@ -703,9 +704,9 @@ and detype_r d flags avoid env sigma t =
match decl with
| LocalDef _ -> true
| LocalAssum (id,_) ->
- try let n = List.index Name.equal (Name id) (fst env) in
+ try let n = List.index Name.equal (Name id.binder_name) (fst env) in
isRelN sigma n c
- with Not_found -> isVarId sigma id c
+ with Not_found -> isVarId sigma id.binder_name c
in
let id,l =
try
@@ -766,11 +767,11 @@ and detype_eqn d (lax,isgoal as flags) avoid env sigma constr construct_nargs br
[DAst.make @@ PatCstr(constr, List.rev patlist,Anonymous)],
detype d flags avoid env sigma b)
| Lambda (x,t,b), false::l ->
- let pat,new_avoid,new_env,new_ids = make_pat x avoid env b None t ids in
+ let pat,new_avoid,new_env,new_ids = make_pat x.binder_name avoid env b None t ids in
buildrec new_ids (pat::patlist) new_avoid new_env l b
| LetIn (x,b,t,b'), true::l ->
- let pat,new_avoid,new_env,new_ids = make_pat x avoid env b' (Some b) t ids in
+ let pat,new_avoid,new_env,new_ids = make_pat x.binder_name avoid env b' (Some b) t ids in
buildrec new_ids (pat::patlist) new_avoid new_env l b'
| Cast (c,_,_), l -> (* Oui, il y a parfois des cast *)
@@ -792,7 +793,7 @@ and detype_eqn d (lax,isgoal as flags) avoid env sigma constr construct_nargs br
in
buildrec Id.Set.empty [] avoid env construct_nargs branch
-and detype_binder d (lax,isgoal as flags) bk avoid env sigma na body ty c =
+and detype_binder d (lax,isgoal as flags) bk avoid env sigma {binder_name=na} body ty c =
let flag = if isgoal then RenamingForGoal else RenamingElsewhereFor (fst env,c) in
let na',avoid' = match bk with
| BLetIn -> compute_displayed_let_name_in sigma flag avoid na c
@@ -828,7 +829,7 @@ let detype_rel_context d ?(lax=false) where avoid env sigma sign =
(RenamingElsewhereFor (fst env,c)) avoid na c in
let b = match decl with
| LocalAssum _ -> None
- | LocalDef (_,b,_) -> Some b
+ | LocalDef (_,b,_) -> Some b
in
let b' = Option.map (detype d (lax,false) avoid env sigma) b in
let t' = detype d (lax,false) avoid env sigma t in
@@ -865,7 +866,7 @@ let detype_closed_glob ?lax isgoal avoid env sigma t =
(* spiwack: I'm not sure it is the right thing to do,
but I'm computing the detyping environment like
[Printer.pr_constr_under_binders_env] does. *)
- let assums = List.map (fun id -> LocalAssum (Name id,(* dummy *) mkProp)) b in
+ let assums = List.map (fun id -> LocalAssum (make_annot (Name id) Sorts.Relevant,(* dummy *) mkProp)) b in
let env = push_rel_context assums env in
DAst.get (detype Now ?lax isgoal avoid env sigma c)
(* if [id] is bound to a [closed_glob_constr]. *)
diff --git a/pretyping/evarconv.ml b/pretyping/evarconv.ml
index 8e273fb4a8..28a97bb91a 100644
--- a/pretyping/evarconv.ml
+++ b/pretyping/evarconv.ml
@@ -15,6 +15,7 @@ open Constr
open Termops
open Environ
open EConstr
+open Context
open Vars
open Reduction
open Reductionops
@@ -78,8 +79,8 @@ let impossible_default_case env =
let coq_unit_judge =
let open Environ in
let make_judge c t = make_judge (EConstr.of_constr c) (EConstr.of_constr t) in
- let na1 = Name (Id.of_string "A") in
- let na2 = Name (Id.of_string "H") in
+ let na1 = make_annot (Name (Id.of_string "A")) Sorts.Relevant in
+ let na2 = make_annot (Name (Id.of_string "H")) Sorts.Relevant in
fun env ->
match impossible_default_case env with
| Some (id, type_of_id, ctx) ->
@@ -87,7 +88,7 @@ let coq_unit_judge =
| None ->
(* In case the constants id/ID are not defined *)
Environ.make_judge (mkLambda (na1,mkProp,mkLambda(na2,mkRel 1,mkRel 1)))
- (mkProd (na1,mkProp,mkArrow (mkRel 1) (mkRel 2))),
+ (mkProd (na1,mkProp,mkArrow (mkRel 1) Sorts.Relevant (mkRel 2))),
Univ.ContextSet.empty
let unfold_projection env evd ts p c =
@@ -251,8 +252,8 @@ let check_conv_record env sigma (t1,sk1) (t2,sk2) =
let canon_s,sk2_effective =
try
match EConstr.kind sigma t2 with
- Prod (_,a,b) -> (* assert (l2=[]); *)
- let _, a, b = destProd sigma t2 in
+ Prod (_,a,b) -> (* assert (l2=[]); *)
+ let _, a, b = destProd sigma t2 in
if noccurn sigma 1 b then
lookup_canonical_conversion (proji, Prod_cs),
(Stack.append_app [|a;pop b|] Stack.empty)
@@ -815,10 +816,10 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) flags env evd pbty
(fun i ->
let b = nf_evar i b1 in
let t = nf_evar i t1 in
- let na = Nameops.Name.pick na1 na2 in
+ let na = Nameops.Name.pick_annot na1 na2 in
evar_conv_x flags (push_rel (RelDecl.LocalDef (na,b,t)) env) i pbty c'1 c'2);
(fun i -> exact_ise_stack2 env i (evar_conv_x flags) sk1 sk2)]
- and f2 i =
+ and f2 i =
let out1 = whd_betaiota_deltazeta_for_iota_state flags.open_ts env i csts1 (v1,sk1)
and out2 = whd_betaiota_deltazeta_for_iota_state flags.open_ts env i csts2 (v2,sk2)
in evar_eqappr_x flags env i pbty out1 out2
@@ -930,7 +931,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) flags env evd pbty
[(fun i -> evar_conv_x flags env i CONV c1 c2);
(fun i ->
let c = nf_evar i c1 in
- let na = Nameops.Name.pick na1 na2 in
+ let na = Nameops.Name.pick_annot na1 na2 in
evar_conv_x flags (push_rel (RelDecl.LocalAssum (na,c)) env) i CONV c'1 c'2);
(* When in modulo_betaiota = false case, lambda's are not reduced *)
(fun i -> exact_ise_stack2 env i (evar_conv_x flags) sk1 sk2)]
@@ -988,12 +989,12 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) flags env evd pbty
UnifFailure (evd,UnifUnivInconsistency p)
| e when CErrors.noncritical e -> UnifFailure (evd,NotSameHead))
- | Prod (n1,c1,c'1), Prod (n2,c2,c'2) when app_empty ->
+ | Prod (n1,c1,c'1), Prod (n2,c2,c'2) when app_empty ->
ise_and evd
[(fun i -> evar_conv_x flags env i CONV c1 c2);
(fun i ->
let c = nf_evar i c1 in
- let na = Nameops.Name.pick n1 n2 in
+ let na = Nameops.Name.pick_annot n1 n2 in
evar_conv_x flags (push_rel (RelDecl.LocalAssum (na,c)) env) i pbty c'1 c'2)]
| Rel x1, Rel x2 ->
@@ -1027,7 +1028,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) flags env evd pbty
| _, Construct u ->
eta_constructor flags env evd sk2 u sk1 term1
- | Fix ((li1, i1),(_,tys1,bds1 as recdef1)), Fix ((li2, i2),(_,tys2,bds2)) -> (* Partially applied fixs *)
+ | Fix ((li1, i1),(_,tys1,bds1 as recdef1)), Fix ((li2, i2),(_,tys2,bds2)) -> (* Partially applied fixs *)
if Int.equal i1 i2 && Array.equal Int.equal li1 li2 then
ise_and evd [
(fun i -> ise_array2 i (fun i' -> evar_conv_x flags env i' CONV) tys1 tys2);
@@ -1035,7 +1036,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) flags env evd pbty
(fun i -> exact_ise_stack2 env i (evar_conv_x flags) sk1 sk2)]
else UnifFailure (evd, NotSameHead)
- | CoFix (i1,(_,tys1,bds1 as recdef1)), CoFix (i2,(_,tys2,bds2)) ->
+ | CoFix (i1,(_,tys1,bds1 as recdef1)), CoFix (i2,(_,tys2,bds2)) ->
if Int.equal i1 i2 then
ise_and evd
[(fun i -> ise_array2 i
@@ -1352,7 +1353,7 @@ let second_order_matching flags env_rhs evd (evk,args) (test,argoccs) rhs =
make_subst (ctxt',l,occsl)
end
| decl'::ctxt', c::l, occs::occsl ->
- let id = NamedDecl.get_id decl' in
+ let id = NamedDecl.get_annot decl' in
let t = NamedDecl.get_type decl' in
let evs = ref [] in
let c = nf_evar evd c in
@@ -1369,7 +1370,7 @@ let second_order_matching flags env_rhs evd (evk,args) (test,argoccs) rhs =
let c = nf_evar evd c in
if !debug_ho_unification then
Feedback.msg_debug Pp.(str"set holes for: " ++
- prc env_rhs evd (mkVar id) ++ spc () ++
+ prc env_rhs evd (mkVar id.binder_name) ++ spc () ++
prc env_rhs evd c ++ str" in " ++
prc env_rhs evd rhs);
let occ = ref 1 in
@@ -1381,7 +1382,7 @@ let second_order_matching flags env_rhs evd (evk,args) (test,argoccs) rhs =
incr occ;
match occs with
| AtOccurrences occs ->
- if Locusops.is_selected oc occs then evd, mkVar id
+ if Locusops.is_selected oc occs then evd, mkVar id.binder_name
else evd, inst
| Unspecified prefer_abstraction ->
let evd, evty = set_holes env_rhs evd cty subst in
@@ -1436,6 +1437,7 @@ let second_order_matching flags env_rhs evd (evk,args) (test,argoccs) rhs =
let rec abstract_free_holes evd = function
| (id,idty,c,cty,evsref,_,_)::l ->
+ let id = id.binder_name in
let c = nf_evar evd c in
if !debug_ho_unification then
Feedback.msg_debug Pp.(str"abstracting: " ++
diff --git a/pretyping/evardefine.ml b/pretyping/evardefine.ml
index 799a81cf4b..a51cb22c20 100644
--- a/pretyping/evardefine.ml
+++ b/pretyping/evardefine.ml
@@ -13,6 +13,7 @@ open Util
open Pp
open Names
open Constr
+open Context
open Termops
open EConstr
open Vars
@@ -84,8 +85,9 @@ let define_pure_evar_as_product env evd evk =
let evd1,(dom,u1) =
new_type_evar evenv evd univ_flexible_alg ~src ~filter:(evar_filter evi)
in
+ let rdom = Sorts.Relevant in (* TODO relevance *)
let evd2,rng =
- let newenv = push_named (LocalAssum (id, dom)) evenv in
+ let newenv = push_named (LocalAssum (make_annot id rdom, dom)) evenv in
let src = subterm_source evk ~where:Codomain evksrc in
let filter = Filter.extend 1 (evar_filter evi) in
if Environ.is_impredicative_sort env (ESorts.kind evd1 s) then
@@ -100,7 +102,7 @@ let define_pure_evar_as_product env evd evk =
let evd3 = Evd.set_leq_sort evenv evd3 (Sorts.sort_of_univ prods) (ESorts.kind evd1 s) in
evd3, rng
in
- let prod = mkProd (Name id, dom, subst_var id rng) in
+ let prod = mkProd (make_annot (Name id) rdom, dom, subst_var id rng) in
let evd3 = Evd.define evk prod evd2 in
evd3,prod
@@ -135,13 +137,15 @@ let define_pure_evar_as_lambda env evd evk =
| _ -> error_not_product env evd typ in
let avoid = Environ.ids_of_named_context_val evi.evar_hyps in
let id =
- next_name_away_with_default_using_types "x" na avoid (Reductionops.whd_evar evd dom) in
+ map_annot (fun na -> next_name_away_with_default_using_types "x" na avoid
+ (Reductionops.whd_evar evd dom)) na
+ in
let newenv = push_named (LocalAssum (id, dom)) evenv in
let filter = Filter.extend 1 (evar_filter evi) in
let src = subterm_source evk ~where:Body (evar_source evk evd1) in
let abstract_arguments = Abstraction.abstract_last evi.evar_abstract_arguments in
- let evd2,body = new_evar newenv evd1 ~src (subst1 (mkVar id) rng) ~filter ~abstract_arguments in
- let lam = mkLambda (Name id, dom, subst_var id body) in
+ let evd2,body = new_evar newenv evd1 ~src (subst1 (mkVar id.binder_name) rng) ~filter ~abstract_arguments in
+ let lam = mkLambda (map_annot Name.mk_name id, dom, subst_var id.binder_name body) in
Evd.define evk lam evd2, lam
let define_evar_as_lambda env evd (evk,args) =
@@ -180,21 +184,22 @@ let split_tycon ?loc env evd tycon =
let rec real_split evd c =
let t = Reductionops.whd_all env evd c in
match EConstr.kind evd t with
- | Prod (na,dom,rng) -> evd, (na, dom, rng)
+ | Prod (na,dom,rng) -> evd, (na, dom, rng)
| Evar ev (* ev is undefined because of whd_all *) ->
let (evd',prod) = define_evar_as_product env evd ev in
- let (_,dom,rng) = destProd evd prod in
- evd',(Anonymous, dom, rng)
+ let (na,dom,rng) = destProd evd prod in
+ let anon = {na with binder_name = Anonymous} in
+ evd',(anon, dom, rng)
| App (c,args) when isEvar evd c ->
- let (evd',lam) = define_evar_as_lambda env evd (destEvar evd c) in
+ let (evd',lam) = define_evar_as_lambda env evd (destEvar evd c) in
real_split evd' (mkApp (lam,args))
| _ -> error_not_product ?loc env evd c
in
match tycon with
- | None -> evd,(Anonymous,None,None)
+ | None -> evd,(make_annot Anonymous Relevant,None,None)
| Some c ->
- let evd', (n, dom, rng) = real_split evd c in
- evd', (n, mk_tycon dom, mk_tycon rng)
+ let evd', (n, dom, rng) = real_split evd c in
+ evd', (n, mk_tycon dom, mk_tycon rng)
let valcon_of_tycon x = x
let lift_tycon n = Option.map (lift n)
diff --git a/pretyping/evardefine.mli b/pretyping/evardefine.mli
index 7b5ced49bf..8ff113196b 100644
--- a/pretyping/evardefine.mli
+++ b/pretyping/evardefine.mli
@@ -33,7 +33,7 @@ val evar_absorb_arguments : env -> evar_map -> existential -> constr list ->
val split_tycon :
?loc:Loc.t -> env -> evar_map -> type_constraint ->
- evar_map * (Name.t * type_constraint * type_constraint)
+ evar_map * (Name.t Context.binder_annot * type_constraint * type_constraint)
val valcon_of_tycon : type_constraint -> val_constraint
val lift_tycon : int -> type_constraint -> type_constraint
diff --git a/pretyping/evarsolve.ml b/pretyping/evarsolve.ml
index e5f2207333..a4a078bfa0 100644
--- a/pretyping/evarsolve.ml
+++ b/pretyping/evarsolve.ml
@@ -12,6 +12,7 @@ open Sorts
open Util
open CErrors
open Names
+open Context
open Constr
open Environ
open Termops
@@ -193,9 +194,9 @@ let recheck_applications unify flags env evdref t =
let rec aux i ty =
if i < Array.length argsty then
match EConstr.kind !evdref (whd_all env !evdref ty) with
- | Prod (na, dom, codom) ->
+ | Prod (na, dom, codom) ->
(match unify flags TypeUnification env !evdref Reduction.CUMUL argsty.(i) dom with
- | Success evd -> evdref := evd;
+ | Success evd -> evdref := evd;
aux (succ i) (subst1 args.(i) codom)
| UnifFailure (evd, reason) ->
Pretype_errors.error_cannot_unify env evd ~reason (argsty.(i), dom))
@@ -353,7 +354,7 @@ let compute_rel_aliases var_aliases rels sigma =
(fun decl (n,aliases) ->
(n-1,
match decl with
- | LocalDef (_,t,u) ->
+ | LocalDef (_,t,u) ->
(match EConstr.kind sigma t with
| Var id' ->
let aliases_of_n =
@@ -640,7 +641,7 @@ let make_projectable_subst aliases sigma evi args =
List.fold_right_i
(fun i decl (args,all,cstrs,revmap) ->
match decl,args with
- | LocalAssum (id,c), a::rest ->
+ | LocalAssum ({binder_name=id},c), a::rest ->
let revmap = Id.Map.add id i revmap in
let cstrs =
let a',args = decompose_app_vect sigma a in
@@ -651,7 +652,7 @@ let make_projectable_subst aliases sigma evi args =
| _ -> cstrs in
let all = Int.Map.add i [a,normalize_alias_opt sigma aliases a,id] all in
(rest,all,cstrs,revmap)
- | LocalDef (id,c,_), a::rest ->
+ | LocalDef ({binder_name=id},c,_), a::rest ->
let revmap = Id.Map.add id i revmap in
(match EConstr.kind sigma c with
| Var id' ->
@@ -727,7 +728,7 @@ let materialize_evar define_fun env evd k (evk1,args1) ty_in_env =
let (sign2,filter2,inst2_in_env,inst2_in_sign,_,evd,_) =
List.fold_right (fun d (sign,filter,inst_in_env,inst_in_sign,env,evd,avoid) ->
let LocalAssum (na,t_in_env) | LocalDef (na,_,t_in_env) = d in
- let id = next_name_away na avoid in
+ let id = map_annot (fun na -> next_name_away na avoid) na in
let evd,t_in_sign =
let s = Retyping.get_sort_of env evd t_in_env in
let evd,ty_t_in_sign = refresh_universes
@@ -743,7 +744,7 @@ let materialize_evar define_fun env evd k (evk1,args1) ty_in_env =
(push_named_context_val d' sign, Filter.extend 1 filter,
(mkRel 1)::(List.map (lift 1) inst_in_env),
(mkRel 1)::(List.map (lift 1) inst_in_sign),
- push_rel d env,evd,Id.Set.add id avoid))
+ push_rel d env,evd,Id.Set.add id.binder_name avoid))
rel_sign
(sign1,filter1,Array.to_list args1,inst_in_sign,env1,evd,avoid)
in
diff --git a/pretyping/find_subterm.ml b/pretyping/find_subterm.ml
index d150f8e1cb..7019cdf046 100644
--- a/pretyping/find_subterm.ml
+++ b/pretyping/find_subterm.ml
@@ -70,7 +70,7 @@ let map_named_declaration_with_hyploc f hyploc acc decl =
in
match decl,hyploc with
| LocalAssum (id,_), InHypValueOnly ->
- error_occurrences_error (IncorrectInValueOccurrence id)
+ error_occurrences_error (IncorrectInValueOccurrence id.Context.binder_name)
| LocalAssum (id,typ), _ ->
let acc,typ = f acc typ in acc, LocalAssum (id,typ)
| LocalDef (id,body,typ), InHypTypeOnly ->
diff --git a/pretyping/globEnv.ml b/pretyping/globEnv.ml
index d6b204561e..cd82b1993b 100644
--- a/pretyping/globEnv.ml
+++ b/pretyping/globEnv.ml
@@ -92,7 +92,7 @@ let push_rec_types ~hypnaming sigma (lna,typarray) env =
let open Context.Rel.Declaration in
let ctxt = Array.map2_i (fun i na t -> Context.Rel.Declaration.LocalAssum (na, lift i t)) lna typarray in
let env,ctx = Array.fold_left_map (fun e assum -> let (d,e) = push_rel sigma assum e ~hypnaming in (e,d)) env ctxt in
- Array.map get_name ctx, env
+ Array.map get_annot ctx, env
let new_evar env sigma ?src ?naming typ =
let open Context.Named.Declaration in
diff --git a/pretyping/globEnv.mli b/pretyping/globEnv.mli
index 63f72e60bd..65ae495135 100644
--- a/pretyping/globEnv.mli
+++ b/pretyping/globEnv.mli
@@ -50,7 +50,7 @@ val vars_of_env : t -> Id.Set.t
val push_rel : hypnaming:naming_mode -> evar_map -> rel_declaration -> t -> rel_declaration * t
val push_rel_context : hypnaming:naming_mode -> ?force_names:bool -> evar_map -> rel_context -> t -> rel_context * t
-val push_rec_types : hypnaming:naming_mode -> evar_map -> Name.t array * constr array -> t -> Name.t array * t
+val push_rec_types : hypnaming:naming_mode -> evar_map -> Name.t Context.binder_annot array * constr array -> t -> Name.t Context.binder_annot array * t
(** Declare an evar using renaming information *)
diff --git a/pretyping/indrec.ml b/pretyping/indrec.ml
index a9b419f03f..4f940fa16a 100644
--- a/pretyping/indrec.ml
+++ b/pretyping/indrec.ml
@@ -21,6 +21,7 @@ open Globnames
open Nameops
open Term
open Constr
+open Context
open Vars
open Namegen
open Declarations
@@ -43,8 +44,8 @@ exception RecursionSchemeError of env * recursion_scheme_error
let named_hd env t na = named_hd env (Evd.from_env env) (EConstr.of_constr t) na
let name_assumption env = function
-| LocalAssum (na,t) -> LocalAssum (named_hd env t na, t)
-| LocalDef (na,c,t) -> LocalDef (named_hd env c na, c, t)
+| LocalAssum (na,t) -> LocalAssum (map_annot (named_hd env t) na, t)
+| LocalDef (na,c,t) -> LocalDef (map_annot (named_hd env c) na, c, t)
let mkLambda_or_LetIn_name env d b = mkLambda_or_LetIn (name_assumption env d) b
let mkProd_or_LetIn_name env d b = mkProd_or_LetIn (name_assumption env d) b
@@ -54,7 +55,7 @@ let it_mkProd_or_LetIn_name env b l = List.fold_left (fun c d -> mkProd_or_LetIn
let it_mkLambda_or_LetIn_name env b l = List.fold_left (fun c d -> mkLambda_or_LetIn_name env d c) b l
let make_prod_dep dep env = if dep then mkProd_name env else mkProd
-let mkLambda_string s t c = mkLambda (Name (Id.of_string s), t, c)
+let mkLambda_string s r t c = mkLambda (make_annot (Name (Id.of_string s)) r, t, c)
(*******************************************)
@@ -79,6 +80,7 @@ let mis_make_case_com dep env sigma (ind, u as pind) (mib,mip as specif) kind =
let indf = make_ind_family(pind, Context.Rel.to_extended_list mkRel 0 lnamespar) in
let constrs = get_constructors env indf in
let projs = get_projections env ind in
+ let relevance = Sorts.relevance_of_sort_family kind in
let () = if Option.is_empty projs then check_privacy_block mib in
let () =
@@ -98,11 +100,13 @@ let mis_make_case_com dep env sigma (ind, u as pind) (mib,mip as specif) kind =
let nbprod = k+1 in
let indf' = lift_inductive_family nbprod indf in
- let arsign,_ = get_arity env indf' in
+ let arsign,sort = get_arity env indf' in
+ let r = Sorts.relevance_of_sort_family sort in
let depind = build_dependent_inductive env indf' in
- let deparsign = LocalAssum (Anonymous,depind)::arsign in
+ let deparsign = LocalAssum (make_annot Anonymous r,depind)::arsign in
- let ci = make_case_info env (fst pind) RegularStyle in
+ let rci = relevance in
+ let ci = make_case_info env (fst pind) rci RegularStyle in
let pbody =
appvect
(mkRel (ndepar + nbprod),
@@ -111,7 +115,7 @@ let mis_make_case_com dep env sigma (ind, u as pind) (mib,mip as specif) kind =
let p =
it_mkLambda_or_LetIn_name env'
((if dep then mkLambda_name env' else mkLambda)
- (Anonymous,depind,pbody))
+ (make_annot Anonymous r,depind,pbody))
arsign
in
let obj =
@@ -132,16 +136,16 @@ let mis_make_case_com dep env sigma (ind, u as pind) (mib,mip as specif) kind =
else
let cs = lift_constructor (k+1) constrs.(k) in
let t = build_branch_type env sigma dep (mkRel (k+1)) cs in
- mkLambda_string "f" t
- (add_branch (push_rel (LocalAssum (Anonymous, t)) env) (k+1))
+ mkLambda_string "f" relevance t
+ (add_branch (push_rel (LocalAssum (make_annot Anonymous relevance, t)) env) (k+1))
in
let (sigma, s) = Evd.fresh_sort_in_family ~rigid:Evd.univ_flexible_alg sigma kind in
let typP = make_arity env' sigma dep indf s in
let typP = EConstr.Unsafe.to_constr typP in
let c =
it_mkLambda_or_LetIn_name env
- (mkLambda_string "P" typP
- (add_branch (push_rel (LocalAssum (Anonymous,typP)) env') 0)) lnamespar
+ (mkLambda_string "P" Sorts.Relevant typP
+ (add_branch (push_rel (LocalAssum (make_annot Anonymous Sorts.Relevant,typP)) env') 0)) lnamespar
in
(sigma, c)
@@ -171,12 +175,12 @@ let type_rec_branch is_rec dep env sigma (vargs,depPvect,decP) tyi cs recargs =
let p' = EConstr.Unsafe.to_constr p' in
let largs = List.map EConstr.Unsafe.to_constr largs in
match kind p' with
- | Prod (n,t,c) ->
- let d = LocalAssum (n,t) in
- make_prod env (n,t,prec (push_rel d env) (i+1) (d::sign) c)
- | LetIn (n,b,t,c) when List.is_empty largs ->
- let d = LocalDef (n,b,t) in
- mkLetIn (n,b,t,prec (push_rel d env) (i+1) (d::sign) c)
+ | Prod (n,t,c) ->
+ let d = LocalAssum (n,t) in
+ make_prod env (n,t,prec (push_rel d env) (i+1) (d::sign) c)
+ | LetIn (n,b,t,c) when List.is_empty largs ->
+ let d = LocalDef (n,b,t) in
+ mkLetIn (n,b,t,prec (push_rel d env) (i+1) (d::sign) c)
| Ind (_,_) ->
let realargs = List.skipn nparams largs in
let base = applist (lift i pk,realargs) in
@@ -208,23 +212,24 @@ let type_rec_branch is_rec dep env sigma (vargs,depPvect,decP) tyi cs recargs =
(match optionpos with
| None ->
make_prod env
- (n,t,
- process_constr (push_rel (LocalAssum (n,t)) env) (i+1) c_0 rest
+ (n,t,
+ process_constr (push_rel (LocalAssum (n,t)) env) (i+1) c_0 rest
(nhyps-1) (i::li))
| Some(dep',p) ->
let nP = lift (i+1+decP) p in
let env' = push_rel (LocalAssum (n,t)) env in
- let t_0 = process_pos env' dep' nP (lift 1 t) in
+ let t_0 = process_pos env' dep' nP (lift 1 t) in
+ let r_0 = Retyping.relevance_of_type env' sigma (EConstr.of_constr t_0) in
make_prod_dep (dep || dep') env
(n,t,
- mkArrow t_0
+ mkArrow t_0 r_0
(process_constr
- (push_rel (LocalAssum (Anonymous,t_0)) env')
+ (push_rel (LocalAssum (make_annot Anonymous n.binder_relevance,t_0)) env')
(i+2) (lift 1 c_0) rest (nhyps-1) (i::li))))
| LetIn (n,b,t,c_0) ->
- mkLetIn (n,b,t,
+ mkLetIn (n,b,t,
process_constr
- (push_rel (LocalDef (n,b,t)) env)
+ (push_rel (LocalDef (n,b,t)) env)
(i+1) c_0 recargs (nhyps-1) li)
| _ -> assert false
else
@@ -250,12 +255,12 @@ let make_rec_branch_arg env sigma (nparrec,fvect,decF) f cstr recargs =
let p' = EConstr.Unsafe.to_constr p' in
let largs = List.map EConstr.Unsafe.to_constr largs in
match kind p' with
- | Prod (n,t,c) ->
- let d = LocalAssum (n,t) in
- mkLambda_name env (n,t,prec (push_rel d env) (i+1) (d::hyps) c)
- | LetIn (n,b,t,c) when List.is_empty largs ->
- let d = LocalDef (n,b,t) in
- mkLetIn (n,b,t,prec (push_rel d env) (i+1) (d::hyps) c)
+ | Prod (n,t,c) ->
+ let d = LocalAssum (n,t) in
+ mkLambda_name env (n,t,prec (push_rel d env) (i+1) (d::hyps) c)
+ | LetIn (n,b,t,c) when List.is_empty largs ->
+ let d = LocalDef (n,b,t) in
+ mkLetIn (n,b,t,prec (push_rel d env) (i+1) (d::hyps) c)
| Ind _ ->
let realargs = List.skipn nparrec largs
and arg = appvect (mkRel (i+1), Context.Rel.to_extended_vect mkRel 0 hyps) in
@@ -280,7 +285,7 @@ let make_rec_branch_arg env sigma (nparrec,fvect,decF) f cstr recargs =
(match optionpos with
| None ->
mkLambda_name env
- (n,t,process_constr (push_rel d env) (i+1)
+ (n,t,process_constr (push_rel d env) (i+1)
(EConstr.Unsafe.to_constr (whd_beta Evd.empty (EConstr.of_constr (applist (lift 1 f, [(mkRel 1)])))))
(cprest,rest))
| Some(_,f_0) ->
@@ -288,12 +293,12 @@ let make_rec_branch_arg env sigma (nparrec,fvect,decF) f cstr recargs =
let env' = push_rel d env in
let arg = process_pos env' nF (lift 1 t) in
mkLambda_name env
- (n,t,process_constr env' (i+1)
+ (n,t,process_constr env' (i+1)
(EConstr.Unsafe.to_constr (whd_beta Evd.empty (EConstr.of_constr (applist (lift 1 f, [(mkRel 1); arg])))))
(cprest,rest)))
| (LocalDef (n,c,t) as d)::cprest, rest ->
mkLetIn
- (n,c,t,
+ (n,c,t,
process_constr (push_rel d env) (i+1) (lift 1 f)
(cprest,rest))
| [],[] -> f
@@ -329,25 +334,26 @@ let mis_make_indrec env sigma ?(force_mutual=false) listdepkind mib u =
let recargpar = recargparn [] (nparams-nparrec) in
let make_one_rec p =
let makefix nbconstruct =
- let rec mrec i ln ltyp ldef = function
- | ((indi,u),mibi,mipi,dep,_)::rest ->
- let tyi = snd indi in
- let nctyi =
- Array.length mipi.mind_consnames in (* nb constructeurs du type*)
+ let rec mrec i ln lrelevance ltyp ldef = function
+ | ((indi,u),mibi,mipi,dep,target_sort)::rest ->
+ let tyi = snd indi in
+ let nctyi =
+ Array.length mipi.mind_consnames in (* nb constructeurs du type*)
- (* arity in the context of the fixpoint, i.e.
+ (* arity in the context of the fixpoint, i.e.
P1..P_nrec f1..f_nbconstruct *)
- let args = Context.Rel.to_extended_list mkRel (nrec+nbconstruct) lnamesparrec in
- let indf = make_ind_family((indi,u),args) in
+ let args = Context.Rel.to_extended_list mkRel (nrec+nbconstruct) lnamesparrec in
+ let indf = make_ind_family((indi,u),args) in
- let arsign,_ = get_arity env indf in
- let depind = build_dependent_inductive env indf in
- let deparsign = LocalAssum (Anonymous,depind)::arsign in
+ let arsign,s = get_arity env indf in
+ let r = Sorts.relevance_of_sort_family s in
+ let depind = build_dependent_inductive env indf in
+ let deparsign = LocalAssum (make_annot Anonymous r,depind)::arsign in
- let nonrecpar = Context.Rel.length lnonparrec in
- let larsign = Context.Rel.length deparsign in
- let ndepar = larsign - nonrecpar in
- let dect = larsign+nrec+nbconstruct in
+ let nonrecpar = Context.Rel.length lnonparrec in
+ let larsign = Context.Rel.length deparsign in
+ let ndepar = larsign - nonrecpar in
+ let dect = larsign+nrec+nbconstruct in
(* constructors in context of the Cases expr, i.e.
P1..P_nrec f1..f_nbconstruct F_1..F_nrec a_1..a_nar x:I *)
@@ -375,9 +381,10 @@ let mis_make_indrec env sigma ?(force_mutual=false) listdepkind mib u =
(* Predicate in the context of the case *)
- let depind' = build_dependent_inductive env indf' in
- let arsign',_ = get_arity env indf' in
- let deparsign' = LocalAssum (Anonymous,depind')::arsign' in
+ let depind' = build_dependent_inductive env indf' in
+ let arsign',s = get_arity env indf' in
+ let r = Sorts.relevance_of_sort_family s in
+ let deparsign' = LocalAssum (make_annot Anonymous r,depind')::arsign' in
let pargs =
let nrpar = Context.Rel.to_extended_list mkRel (2*ndepar) lnonparrec
@@ -388,13 +395,15 @@ let mis_make_indrec env sigma ?(force_mutual=false) listdepkind mib u =
in
(* body of i-th component of the mutual fixpoint *)
+ let target_relevance = Sorts.relevance_of_sort_family target_sort in
let deftyi =
- let ci = make_case_info env indi RegularStyle in
+ let rci = target_relevance in
+ let ci = make_case_info env indi rci RegularStyle in
let concl = applist (mkRel (dect+j+ndepar),pargs) in
let pred =
it_mkLambda_or_LetIn_name env
((if dep then mkLambda_name env else mkLambda)
- (Anonymous,depind',concl))
+ (make_annot Anonymous r,depind',concl))
arsign'
in
let obj =
@@ -416,20 +425,21 @@ let mis_make_indrec env sigma ?(force_mutual=false) listdepkind mib u =
in it_mkProd_or_LetIn_name env
concl
deparsign
- in
- mrec (i+nctyi) (Context.Rel.nhyps arsign ::ln) (typtyi::ltyp)
+ in
+ mrec (i+nctyi) (Context.Rel.nhyps arsign ::ln) (target_relevance::lrelevance) (typtyi::ltyp)
(deftyi::ldef) rest
| [] ->
let fixn = Array.of_list (List.rev ln) in
let fixtyi = Array.of_list (List.rev ltyp) in
let fixdef = Array.of_list (List.rev ldef) in
- let names = Array.make nrec (Name(Id.of_string "F")) in
- mkFix ((fixn,p),(names,fixtyi,fixdef))
+ let lrelevance = CArray.rev_of_list lrelevance in
+ let names = Array.map (fun r -> make_annot (Name(Id.of_string "F")) r) lrelevance in
+ mkFix ((fixn,p),(names,fixtyi,fixdef))
in
- mrec 0 [] [] []
+ mrec 0 [] [] [] []
in
let rec make_branch env i = function
- | ((indi,u),mibi,mipi,dep,_)::rest ->
+ | ((indi,u),mibi,mipi,dep,sfam)::rest ->
let tyi = snd indi in
let nconstr = Array.length mipi.mind_consnames in
let rec onerec env j =
@@ -443,9 +453,10 @@ let mis_make_indrec env sigma ?(force_mutual=false) listdepkind mib u =
let p_0 =
type_rec_branch
true dep env !evdref (vargs,depPvec,i+j) tyi cs recarg
- in
- mkLambda_string "f" p_0
- (onerec (push_rel (LocalAssum (Anonymous,p_0)) env) (j+1))
+ in
+ let r_0 = Sorts.relevance_of_sort_family sfam in
+ mkLambda_string "f" r_0 p_0
+ (onerec (push_rel (LocalAssum (make_annot Anonymous r_0,p_0)) env) (j+1))
in onerec env 0
| [] ->
makefix i listdepkind
@@ -458,9 +469,9 @@ let mis_make_indrec env sigma ?(force_mutual=false) listdepkind mib u =
evdref := sigma; res
in
let typP = make_arity env !evdref dep indf s in
- let typP = EConstr.Unsafe.to_constr typP in
- mkLambda_string "P" typP
- (put_arity (push_rel (LocalAssum (Anonymous,typP)) env) (i+1) rest)
+ let typP = EConstr.Unsafe.to_constr typP in
+ mkLambda_string "P" Sorts.Relevant typP
+ (put_arity (push_rel (LocalAssum (anonR,typP)) env) (i+1) rest)
| [] ->
make_branch env 0 listdepkind
in
@@ -530,7 +541,7 @@ let weaken_sort_scheme env evd set sort npars term ty =
mkLambda (n, t', mkApp(term,Termops.rel_vect 0 (npars+1)))
else
let c',term' = drec (np-1) c in
- mkProd (n, t, c'), mkLambda (n, t, term')
+ mkProd (n, t, c'), mkLambda (n, t, term')
| LetIn (n,b,t,c) -> let c',term' = drec np c in
mkLetIn (n,b,t,c'), mkLetIn (n,b,t,term')
| _ -> anomaly ~label:"weaken_sort_scheme" (Pp.str "wrong elimination type.")
diff --git a/pretyping/inductiveops.ml b/pretyping/inductiveops.ml
index d937456bcb..204618034e 100644
--- a/pretyping/inductiveops.ml
+++ b/pretyping/inductiveops.ml
@@ -15,6 +15,7 @@ open Univ
open Term
open Constr
open Vars
+open Context
open Termops
open Declarations
open Declareops
@@ -60,6 +61,8 @@ let lift_inductive_family n = liftn_inductive_family n 1
let substnl_ind_family l n = map_ind_family (substnl l n)
+let relevance_of_inductive_family env ((ind,_),_ : inductive_family) =
+ Inductive.relevance_of_inductive env ind
type inductive_type = IndType of inductive_family * EConstr.constr list
@@ -75,6 +78,9 @@ let lift_inductive_type n = liftn_inductive_type n 1
let substnl_ind_type l n = map_inductive_type (EConstr.Vars.substnl l n)
+let relevance_of_inductive_type env (IndType (indf, _)) =
+ relevance_of_inductive_family env indf
+
let mkAppliedInd (IndType ((ind,params), realargs)) =
let open EConstr in
let ind = on_snd EInstance.make ind in
@@ -276,7 +282,7 @@ let has_dependent_elim mib =
| NotRecord | FakeRecord -> true
(* Annotation for cases *)
-let make_case_info env ind style =
+let make_case_info env ind r style =
let (mib,mip) = Inductive.lookup_mind_specif env ind in
let ind_tags =
Context.Rel.to_tags (List.firstn mip.mind_nrealdecls mip.mind_arity_ctxt) in
@@ -289,6 +295,7 @@ let make_case_info env ind style =
ci_npar = mib.mind_nparams;
ci_cstr_ndecls = mip.mind_consnrealdecls;
ci_cstr_nargs = mip.mind_consnrealargs;
+ ci_relevance = r;
ci_pp_info = print_info }
(*s Useful functions *)
@@ -419,12 +426,14 @@ let build_dependent_inductive env ((ind, params) as indf) =
(* builds the arity of an elimination predicate in sort [s] *)
let make_arity_signature env sigma dep indf =
- let (arsign,_) = get_arity env indf in
+ let (arsign,s) = get_arity env indf in
+ let r = Sorts.relevance_of_sort_family s in
+ let anon = make_annot Anonymous r in
let arsign = List.map (fun d -> Termops.map_rel_decl EConstr.of_constr d) arsign in
if dep then
(* We need names everywhere *)
Namegen.name_context env sigma
- ((LocalAssum (Anonymous,EConstr.of_constr (build_dependent_inductive env indf)))::arsign)
+ ((LocalAssum (anon,EConstr.of_constr (build_dependent_inductive env indf)))::arsign)
(* Costly: would be better to name once for all at definition time *)
else
(* No need to enforce names *)
@@ -457,7 +466,9 @@ let compute_projections env (kn, i as ind) =
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)))
+ | PrimRecord info ->
+ let id, _, _, _ = info.(i) in
+ make_annot (Name id) mib.mind_packets.(i).mind_relevant
in
let pkt = mib.mind_packets.(i) in
let { mind_nparams = nparamargs; mind_params_ctxt = params } = mib in
@@ -491,7 +502,7 @@ let compute_projections env (kn, i as ind) =
let subst = c1 :: subst in
(proj_arg, j+1, pbs, subst)
| LocalAssum (na,t) ->
- match na with
+ match na.binder_name with
| Name id ->
let lab = Label.of_id id in
let kn = Projection.Repr.make ind ~proj_npars:mib.mind_nparams ~proj_arg lab in
@@ -601,7 +612,7 @@ let is_predicate_explicitly_dep env sigma pred arsign =
From Coq > 8.2, using or not the effective dependency of
the predicate is parametrable! *)
- begin match na with
+ begin match na.binder_name with
| Anonymous -> false
| Name _ -> true
end
@@ -643,9 +654,10 @@ let type_case_branches_with_names env sigma indspec p c =
(* Type of Case predicates *)
let arity_of_case_predicate env (ind,params) dep k =
- let arsign,_ = get_arity env (ind,params) in
+ let arsign,s = get_arity env (ind,params) in
+ let r = Sorts.relevance_of_sort_family s in
let mind = build_dependent_inductive env (ind,params) in
- let concl = if dep then mkArrow mind (mkSort k) else mkSort k in
+ let concl = if dep then mkArrow mind r (mkSort k) else mkSort k in
Term.it_mkProd_or_LetIn concl arsign
(***********************************************)
@@ -720,7 +732,7 @@ let control_only_guard env sigma c =
match kind c with
| CoFix (_,(_,_,_) as cofix) ->
Inductive.check_cofix e cofix
- | Fix (_,(_,_,_) as fix) ->
+ | Fix fix ->
Inductive.check_fix e fix
| _ -> ()
in
diff --git a/pretyping/inductiveops.mli b/pretyping/inductiveops.mli
index 5a4257e175..c74bbfe98b 100644
--- a/pretyping/inductiveops.mli
+++ b/pretyping/inductiveops.mli
@@ -38,6 +38,8 @@ val lift_inductive_family : int -> inductive_family -> inductive_family
val substnl_ind_family :
constr list -> int -> inductive_family -> inductive_family
+val relevance_of_inductive_family : env -> inductive_family -> Sorts.relevance
+
(** An inductive type with its parameters and real arguments *)
type inductive_type = IndType of inductive_family * EConstr.constr list
val make_ind_type : inductive_family * EConstr.constr list -> inductive_type
@@ -47,6 +49,8 @@ val liftn_inductive_type : int -> int -> inductive_type -> inductive_type
val lift_inductive_type : int -> inductive_type -> inductive_type
val substnl_ind_type : EConstr.constr list -> int -> inductive_type -> inductive_type
+val relevance_of_inductive_type : env -> inductive_type -> Sorts.relevance
+
val mkAppliedInd : inductive_type -> EConstr.constr
val mis_is_recursive_subset : int list -> wf_paths -> bool
val mis_is_recursive :
@@ -176,7 +180,7 @@ val type_case_branches_with_names :
env -> evar_map -> pinductive * EConstr.constr list -> constr -> constr -> types array * types
(** Annotation for cases *)
-val make_case_info : env -> inductive -> case_style -> case_info
+val make_case_info : env -> inductive -> Sorts.relevance -> case_style -> case_info
(** Make a case or substitute projections if the inductive type is a record
with primitive projections.
diff --git a/pretyping/nativenorm.ml b/pretyping/nativenorm.ml
index 77ae09ee6f..0b2d760ca8 100644
--- a/pretyping/nativenorm.ml
+++ b/pretyping/nativenorm.ml
@@ -10,6 +10,7 @@
open CErrors
open Term
open Constr
+open Context
open Vars
open Environ
open Reduction
@@ -89,10 +90,12 @@ let invert_tag cst tag reloc_tbl =
with Find_at j -> (j+1)
let decompose_prod env t =
- let (name,dom,codom as res) = destProd (whd_all env t) in
- match name with
- | Anonymous -> (Name (Id.of_string "x"),dom,codom)
- | _ -> res
+ let (name,dom,codom) = destProd (whd_all env t) in
+ let name = map_annot (function
+ | Anonymous -> Name (Id.of_string "x")
+ | na -> na) name
+ in
+ (name,dom,codom)
let app_type env c =
let t = whd_all env c in
@@ -194,7 +197,7 @@ let rec nf_val env sigma v typ =
| Vaccu accu -> nf_accu env sigma accu
| Vfun f ->
let lvl = nb_rel env in
- let name,dom,codom =
+ let name,dom,codom =
try decompose_prod env typ
with DestKO ->
CErrors.anomaly
@@ -275,11 +278,13 @@ and nf_atom env sigma atom =
| Asort s -> mkSort s
| Avar id -> mkVar id
| Aprod(n,dom,codom) ->
- let dom = nf_type env sigma dom in
- let vn = mk_rel_accu (nb_rel env) in
- let env = push_rel (LocalAssum (n,dom)) env in
- let codom = nf_type env sigma (codom vn) in
- mkProd(n,dom,codom)
+ let dom, sdom = nf_type_sort env sigma dom in
+ let rdom = Sorts.relevance_of_sort sdom in
+ let n = make_annot n rdom in
+ let vn = mk_rel_accu (nb_rel env) in
+ let env = push_rel (LocalAssum (n,dom)) env in
+ let codom = nf_type env sigma (codom vn) in
+ mkProd(n,dom,codom)
| Ameta (mv,_) -> mkMeta mv
| Aproj (p, c) ->
let c = nf_accu env sigma c in
@@ -325,28 +330,34 @@ and nf_atom_type env sigma atom =
let ci = ans.asw_ci in
mkCase(ci, p, a, branchs), tcase
| Afix(tt,ft,rp,s) ->
- let tt = Array.map (fun t -> nf_type env sigma t) tt in
- let name = Array.map (fun _ -> (Name (Id.of_string "Ffix"))) tt in
+ let tt = Array.map (fun t -> nf_type_sort env sigma t) tt in
+ let tt = Array.map fst tt and rt = Array.map snd tt in
+ let name = Name (Id.of_string "Ffix") in
+ let names = Array.map (fun s -> make_annot name (Sorts.relevance_of_sort s)) rt in
let lvl = nb_rel env in
let nbfix = Array.length ft in
let fargs = mk_rels_accu lvl (Array.length ft) in
- (* Third argument of the tuple is ignored by push_rec_types *)
- let env = push_rec_types (name,tt,[||]) env in
+ (* Body argument of the tuple is ignored by push_rec_types *)
+ let env = push_rec_types (names,tt,[||]) env in
(* We lift here because the types of arguments (in tt) will be evaluated
in an environment where the fixpoints have been pushed *)
let norm_body i v = nf_val env sigma (napply v fargs) (lift nbfix tt.(i)) in
let ft = Array.mapi norm_body ft in
- mkFix((rp,s),(name,tt,ft)), tt.(s)
+ mkFix((rp,s),(names,tt,ft)), tt.(s)
| Acofix(tt,ft,s,_) | Acofixe(tt,ft,s,_) ->
- let tt = Array.map (nf_type env sigma) tt in
- let name = Array.map (fun _ -> (Name (Id.of_string "Fcofix"))) tt in
+ let tt = Array.map (fun t -> nf_type_sort env sigma t) tt in
+ let tt = Array.map fst tt and rt = Array.map snd tt in
+ let name = Name (Id.of_string "Fcofix") in
let lvl = nb_rel env in
+ let names = Array.map (fun s -> make_annot name (Sorts.relevance_of_sort s)) rt in
let fargs = mk_rels_accu lvl (Array.length ft) in
- let env = push_rec_types (name,tt,[||]) env in
+ let env = push_rec_types (names,tt,[||]) env in
let ft = Array.mapi (fun i v -> nf_val env sigma (napply v fargs) tt.(i)) ft in
- mkCoFix(s,(name,tt,ft)), tt.(s)
+ mkCoFix(s,(names,tt,ft)), tt.(s)
| Aprod(n,dom,codom) ->
let dom,s1 = nf_type_sort env sigma dom in
+ let r1 = Sorts.relevance_of_sort s1 in
+ let n = make_annot n r1 in
let vn = mk_rel_accu (nb_rel env) in
let env = push_rel (LocalAssum (n,dom)) env in
let codom,s2 = nf_type_sort env sigma (codom vn) in
@@ -390,6 +401,8 @@ and nf_predicate env sigma ind mip params v pT =
let rargs = Array.init n (fun i -> mkRel (n-i)) in
let params = if Int.equal n 0 then params else Array.map (lift n) params in
let dom = mkApp(mkIndU ind,Array.append params rargs) in
+ let r = Inductive.relevance_of_inductive env (fst ind) in
+ let name = make_annot name r in
let body = nf_type (push_rel (LocalAssum (name,dom)) env) sigma vb in
mkLambda(name,dom,body)
| _ -> nf_type env sigma v
diff --git a/pretyping/patternops.ml b/pretyping/patternops.ml
index 305ebf3dd5..13034d078a 100644
--- a/pretyping/patternops.ml
+++ b/pretyping/patternops.ml
@@ -15,6 +15,7 @@ open Globnames
open Nameops
open Term
open Constr
+open Context
open Glob_term
open Pp
open Mod_subst
@@ -158,12 +159,15 @@ let pattern_of_constr env sigma t =
| 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),
- pattern_of_constr (push_rel (LocalDef (na,c,t)) env) b)
- | Prod (na,c,b) -> PProd (na,pattern_of_constr env c,
- pattern_of_constr (push_rel (LocalAssum (na, c)) env) b)
- | Lambda (na,c,b) -> PLambda (na,pattern_of_constr env c,
- pattern_of_constr (push_rel (LocalAssum (na, c)) env) b)
+ | LetIn (na,c,t,b) -> PLetIn (na.binder_name,
+ pattern_of_constr env c,Some (pattern_of_constr env t),
+ pattern_of_constr (push_rel (LocalDef (na,c,t)) env) b)
+ | Prod (na,c,b) -> PProd (na.binder_name,
+ pattern_of_constr env c,
+ pattern_of_constr (push_rel (LocalAssum (na, c)) env) b)
+ | Lambda (na,c,b) -> PLambda (na.binder_name,
+ pattern_of_constr env c,
+ pattern_of_constr (push_rel (LocalAssum (na, c)) env) b)
| App (f,a) ->
(match
match kind f with
@@ -207,12 +211,12 @@ let pattern_of_constr env sigma t =
| Fix (lni,(lna,tl,bl)) ->
let push env na2 c2 = push_rel (LocalAssum (na2,c2)) env in
let env' = Array.fold_left2 push env lna tl in
- PFix (lni,(lna,Array.map (pattern_of_constr env) tl,
+ PFix (lni,(Array.map binder_name lna,Array.map (pattern_of_constr env) tl,
Array.map (pattern_of_constr env') bl))
| CoFix (ln,(lna,tl,bl)) ->
let push env na2 c2 = push_rel (LocalAssum (na2,c2)) env in
let env' = Array.fold_left2 push env lna tl in
- PCoFix (ln,(lna,Array.map (pattern_of_constr env) tl,
+ PCoFix (ln,(Array.map binder_name lna,Array.map (pattern_of_constr env) tl,
Array.map (pattern_of_constr env') bl))
| Int i -> PInt i in
pattern_of_constr env t
diff --git a/pretyping/pretype_errors.ml b/pretyping/pretype_errors.ml
index 1a51ea4db7..35a7036af4 100644
--- a/pretyping/pretype_errors.ml
+++ b/pretyping/pretype_errors.ml
@@ -108,9 +108,9 @@ let error_ill_typed_rec_body ?loc env sigma i na jl tys =
raise_type_error ?loc
(env, sigma, IllTypedRecBody (i, na, jl, tys))
-let error_elim_arity ?loc env sigma pi s c j a =
+let error_elim_arity ?loc env sigma pi c j a =
raise_type_error ?loc
- (env, sigma, ElimArity (pi, s, c, j, a))
+ (env, sigma, ElimArity (pi, c, j, a))
let error_not_a_type ?loc env sigma j =
raise_type_error ?loc (env, sigma, NotAType j)
diff --git a/pretyping/pretype_errors.mli b/pretyping/pretype_errors.mli
index 82926ba280..a9e2b0ea8f 100644
--- a/pretyping/pretype_errors.mli
+++ b/pretyping/pretype_errors.mli
@@ -102,12 +102,12 @@ val error_number_branches :
val error_ill_typed_rec_body :
?loc:Loc.t -> env -> Evd.evar_map ->
- int -> Name.t array -> unsafe_judgment array -> types array -> 'b
+ int -> Name.t Context.binder_annot array -> unsafe_judgment array -> types array -> 'b
val error_elim_arity :
?loc:Loc.t -> env -> Evd.evar_map ->
- pinductive -> Sorts.family list -> constr ->
- unsafe_judgment -> (Sorts.family * Sorts.family * arity_error) option -> 'b
+ pinductive -> constr ->
+ unsafe_judgment -> (Sorts.family list * Sorts.family * Sorts.family * arity_error) option -> 'b
val error_not_a_type :
?loc:Loc.t -> env -> Evd.evar_map -> unsafe_judgment -> 'b
diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml
index 62dd50d934..8e9a2e114b 100644
--- a/pretyping/pretyping.ml
+++ b/pretyping/pretyping.ml
@@ -29,6 +29,7 @@ open Util
open Names
open Evd
open Constr
+open Context
open Termops
open Environ
open EConstr
@@ -475,8 +476,8 @@ let mark_obligation_evar sigma k evc =
let rec pretype ~program_mode k0 resolve_tc (tycon : type_constraint) (env : GlobEnv.t) (sigma : evar_map) t =
let inh_conv_coerce_to_tycon ?loc = inh_conv_coerce_to_tycon ?loc ~program_mode resolve_tc in
- let pretype_type = pretype_type k0 resolve_tc in
- let pretype = pretype k0 resolve_tc in
+ let pretype_type = pretype_type ~program_mode k0 resolve_tc in
+ let pretype = pretype ~program_mode k0 resolve_tc in
let open Context.Rel.Declaration in
let loc = t.CAst.loc in
match DAst.get t with
@@ -485,7 +486,7 @@ let rec pretype ~program_mode k0 resolve_tc (tycon : type_constraint) (env : Glo
inh_conv_coerce_to_tycon ?loc env sigma t_ref tycon
| GVar id ->
- let sigma, t_id = pretype_id (fun e r t -> pretype ~program_mode tycon e r t) k0 loc env sigma id in
+ let sigma, t_id = pretype_id (fun e r t -> pretype tycon e r t) k0 loc env sigma id in
inh_conv_coerce_to_tycon ?loc env sigma t_id tycon
| GEvar (id, inst) ->
@@ -537,21 +538,23 @@ let rec pretype ~program_mode k0 resolve_tc (tycon : type_constraint) (env : Glo
let rec type_bl env sigma ctxt = function
| [] -> sigma, ctxt
| (na,bk,None,ty)::bl ->
- let sigma, ty' = pretype_type ~program_mode empty_valcon env sigma ty in
- let dcl = LocalAssum (na, ty'.utj_val) in
+ let sigma, ty' = pretype_type empty_valcon env sigma ty in
+ let rty' = Sorts.relevance_of_sort ty'.utj_type in
+ let dcl = LocalAssum (make_annot na rty', ty'.utj_val) in
let dcl', env = push_rel ~hypnaming sigma dcl env in
type_bl env sigma (Context.Rel.add dcl' ctxt) bl
| (na,bk,Some bd,ty)::bl ->
- let sigma, ty' = pretype_type ~program_mode empty_valcon env sigma ty in
- let sigma, bd' = pretype ~program_mode (mk_tycon ty'.utj_val) env sigma bd in
- let dcl = LocalDef (na, bd'.uj_val, ty'.utj_val) in
+ let sigma, ty' = pretype_type empty_valcon env sigma ty in
+ let rty' = Sorts.relevance_of_sort ty'.utj_type in
+ let sigma, bd' = pretype (mk_tycon ty'.utj_val) env sigma bd in
+ let dcl = LocalDef (make_annot na rty', bd'.uj_val, ty'.utj_val) in
let dcl', env = push_rel ~hypnaming sigma dcl env in
type_bl env sigma (Context.Rel.add dcl' ctxt) bl in
let sigma, ctxtv = Array.fold_left_map (fun sigma -> type_bl env sigma Context.Rel.empty) sigma bl in
let sigma, larj =
Array.fold_left2_map
(fun sigma e ar ->
- pretype_type ~program_mode empty_valcon (snd (push_rel_context ~hypnaming sigma e env)) sigma ar)
+ pretype_type empty_valcon (snd (push_rel_context ~hypnaming sigma e env)) sigma ar)
sigma ctxtv lar in
let lara = Array.map (fun a -> a.utj_val) larj in
let ftys = Array.map2 (fun e a -> it_mkProd_or_LetIn a e) ctxtv lara in
@@ -570,6 +573,10 @@ let rec pretype ~program_mode k0 resolve_tc (tycon : type_constraint) (env : Glo
end
| None -> sigma
in
+ let names = Array.map2 (fun na t ->
+ make_annot na (Retyping.relevance_of_type !!(env) sigma t))
+ names ftys
+ in
(* Note: bodies are not used by push_rec_types, so [||] is safe *)
let names,newenv = push_rec_types ~hypnaming sigma (names,ftys) env in
let sigma, vdefj =
@@ -581,7 +588,7 @@ let rec pretype ~program_mode k0 resolve_tc (tycon : type_constraint) (env : Glo
decompose_prod_n_assum sigma (Context.Rel.length ctxt)
(lift nbfix ftys.(i)) in
let ctxt,nenv = push_rel_context ~hypnaming sigma ctxt newenv in
- let sigma, j = pretype ~program_mode (mk_tycon ty) nenv sigma def in
+ let sigma, j = pretype (mk_tycon ty) nenv sigma def in
sigma, { uj_val = it_mkLambda_or_LetIn j.uj_val ctxt;
uj_type = it_mkProd_or_LetIn j.uj_type ctxt })
sigma ctxtv vdef in
@@ -604,10 +611,10 @@ let rec pretype ~program_mode k0 resolve_tc (tycon : type_constraint) (env : Glo
| None -> List.map_i (fun i _ -> i) 0 ctxtv.(i))
vn)
in
- let fixdecls = (names,ftys,fdefs) in
+ let fixdecls = (names,ftys,fdefs) in
let indexes = esearch_guard ?loc !!env sigma possible_indexes fixdecls in
make_judge (mkFix ((indexes,i),fixdecls)) ftys.(i)
- | GCoFix i ->
+ | GCoFix i ->
let fixdecls = (names,ftys,fdefs) in
let cofix = (i, fixdecls) in
(try check_cofix !!env (i, nf_fix sigma fixdecls)
@@ -624,7 +631,7 @@ let rec pretype ~program_mode k0 resolve_tc (tycon : type_constraint) (env : Glo
inh_conv_coerce_to_tycon ?loc env sigma j tycon
| GApp (f,args) ->
- let sigma, fj = pretype ~program_mode empty_tycon env sigma f in
+ let sigma, fj = pretype empty_tycon env sigma f in
let floc = loc_of_glob_constr f in
let length = List.length args in
let candargs =
@@ -667,7 +674,7 @@ let rec pretype ~program_mode k0 resolve_tc (tycon : type_constraint) (env : Glo
match EConstr.kind sigma resty with
| Prod (na,c1,c2) ->
let tycon = Some c1 in
- let sigma, hj = pretype ~program_mode tycon env sigma c in
+ let sigma, hj = pretype tycon env sigma c in
let sigma, candargs, ujval =
match candargs with
| [] -> sigma, [], j_val hj
@@ -679,12 +686,12 @@ let rec pretype ~program_mode k0 resolve_tc (tycon : type_constraint) (env : Glo
sigma, args, nf_evar sigma (j_val hj)
end
in
- let sigma, ujval = adjust_evar_source sigma na ujval in
- let value, typ = app_f n (j_val resj) ujval, subst1 ujval c2 in
+ let sigma, ujval = adjust_evar_source sigma na.binder_name ujval in
+ let value, typ = app_f n (j_val resj) ujval, subst1 ujval c2 in
let j = { uj_val = value; uj_type = typ } in
apply_rec env sigma (n+1) j candargs rest
| _ ->
- let sigma, hj = pretype ~program_mode empty_tycon env sigma c in
+ let sigma, hj = pretype empty_tycon env sigma c in
error_cant_apply_not_functional
?loc:(Loc.merge_opt floc argloc) !!env sigma resj [|hj|]
in
@@ -714,26 +721,28 @@ let rec pretype ~program_mode k0 resolve_tc (tycon : type_constraint) (env : Glo
in
let sigma, (name',dom,rng) = split_tycon ?loc !!env sigma tycon' in
let dom_valcon = valcon_of_tycon dom in
- let sigma, j = pretype_type ~program_mode dom_valcon env sigma c1 in
+ let sigma, j = pretype_type dom_valcon env sigma c1 in
+ let name = {binder_name=name; binder_relevance=Sorts.relevance_of_sort j.utj_type} in
let var = LocalAssum (name, j.utj_val) in
let hypnaming = if program_mode then ProgramNaming else KeepUserNameAndRenameExistingButSectionNames in
let var',env' = push_rel ~hypnaming sigma var env in
- let sigma, j' = pretype ~program_mode rng env' sigma c2 in
+ let sigma, j' = pretype rng env' sigma c2 in
let name = get_name var' in
- let resj = judge_of_abstraction !!env (orelse_name name name') j j' in
+ let resj = judge_of_abstraction !!env (orelse_name name name'.binder_name) j j' in
inh_conv_coerce_to_tycon ?loc env sigma resj tycon
| GProd(name,bk,c1,c2) ->
- let sigma, j = pretype_type ~program_mode empty_valcon env sigma c1 in
+ let sigma, j = pretype_type empty_valcon env sigma c1 in
let hypnaming = if program_mode then ProgramNaming else KeepUserNameAndRenameExistingButSectionNames in
let sigma, name, j' = match name with
| Anonymous ->
- let sigma, j = pretype_type ~program_mode empty_valcon env sigma c2 in
+ let sigma, j = pretype_type empty_valcon env sigma c2 in
sigma, name, { j with utj_val = lift 1 j.utj_val }
| Name _ ->
- let var = LocalAssum (name, j.utj_val) in
+ let r = Sorts.relevance_of_sort j.utj_type in
+ let var = LocalAssum (make_annot name r, j.utj_val) in
let var, env' = push_rel ~hypnaming sigma var env in
- let sigma, c2_j = pretype_type ~program_mode empty_valcon env' sigma c2 in
+ let sigma, c2_j = pretype_type empty_valcon env' sigma c2 in
sigma, get_name var, c2_j
in
let resj =
@@ -749,24 +758,25 @@ let rec pretype ~program_mode k0 resolve_tc (tycon : type_constraint) (env : Glo
let sigma, tycon1 =
match t with
| Some t ->
- let sigma, t_j = pretype_type ~program_mode empty_valcon env sigma t in
+ let sigma, t_j = pretype_type empty_valcon env sigma t in
sigma, mk_tycon t_j.utj_val
| None ->
sigma, empty_tycon in
- let sigma, j = pretype ~program_mode tycon1 env sigma c1 in
+ let sigma, j = pretype tycon1 env sigma c1 in
let sigma, t = Evarsolve.refresh_universes
~onlyalg:true ~status:Evd.univ_flexible (Some false) !!env sigma j.uj_type in
- let var = LocalDef (name, j.uj_val, t) in
+ let r = Retyping.relevance_of_term !!env sigma j.uj_val in
+ let var = LocalDef (make_annot name r, j.uj_val, t) in
let tycon = lift_tycon 1 tycon in
let hypnaming = if program_mode then ProgramNaming else KeepUserNameAndRenameExistingButSectionNames in
let var, env = push_rel ~hypnaming sigma var env in
- let sigma, j' = pretype ~program_mode tycon env sigma c2 in
+ let sigma, j' = pretype tycon env sigma c2 in
let name = get_name var in
- sigma, { uj_val = mkLetIn (name, j.uj_val, t, j'.uj_val) ;
+ sigma, { uj_val = mkLetIn (make_annot name r, j.uj_val, t, j'.uj_val) ;
uj_type = subst1 j.uj_val j'.uj_type }
| GLetTuple (nal,(na,po),c,d) ->
- let sigma, cj = pretype ~program_mode empty_tycon env sigma c in
+ let sigma, cj = pretype empty_tycon env sigma c in
let (IndType (indf,realargs)) =
try find_rectype !!env sigma cj.uj_type
with Not_found ->
@@ -790,10 +800,11 @@ let rec pretype ~program_mode k0 resolve_tc (tycon : type_constraint) (env : Glo
| Some ps ->
let rec aux n k names l =
match names, l with
- | na :: names, (LocalAssum (_,t) :: l) ->
+ | na :: names, (LocalAssum (na', t) :: l) ->
let t = EConstr.of_constr t in
let proj = Projection.make ps.(cs.cs_nargs - k) true in
- LocalDef (na, lift (cs.cs_nargs - n) (mkProj (proj, cj.uj_val)), t)
+ LocalDef ({na' with binder_name = na},
+ lift (cs.cs_nargs - n) (mkProj (proj, cj.uj_val)), t)
:: aux (n+1) (k + 1) names l
| na :: names, (decl :: l) ->
set_name na decl :: aux (n+1) k names l
@@ -803,27 +814,27 @@ let rec pretype ~program_mode k0 resolve_tc (tycon : type_constraint) (env : Glo
let fsign = Context.Rel.map (whd_betaiota sigma) fsign in
let hypnaming = if program_mode then ProgramNaming else KeepUserNameAndRenameExistingButSectionNames in
let fsign,env_f = push_rel_context ~hypnaming sigma fsign env in
- let obj ind p v f =
+ let obj ind rci p v f =
if not record then
- let f = it_mkLambda_or_LetIn f fsign in
- let ci = make_case_info !!env (fst ind) LetStyle in
- mkCase (ci, p, cj.uj_val,[|f|])
+ let f = it_mkLambda_or_LetIn f fsign in
+ let ci = make_case_info !!env (fst ind) rci LetStyle in
+ mkCase (ci, p, cj.uj_val,[|f|])
else it_mkLambda_or_LetIn f fsign
in
(* Make dependencies from arity signature impossible *)
- let arsgn =
- let arsgn,_ = get_arity !!env indf in
- List.map (set_name Anonymous) arsgn
+ let arsgn, indr =
+ let arsgn,s = get_arity !!env indf in
+ List.map (set_name Anonymous) arsgn, Sorts.relevance_of_sort_family s
in
let indt = build_dependent_inductive !!env indf in
- let psign = LocalAssum (na, indt) :: arsgn in (* For locating names in [po] *)
+ let psign = LocalAssum (make_annot na indr, indt) :: arsgn in (* For locating names in [po] *)
let psign = List.map (fun d -> map_rel_decl EConstr.of_constr d) psign in
let predenv = Cases.make_return_predicate_ltac_lvar env sigma na c cj.uj_val in
let nar = List.length arsgn in
let psign',env_p = push_rel_context ~hypnaming ~force_names:true sigma psign predenv in
(match po with
| Some p ->
- let sigma, pj = pretype_type ~program_mode empty_valcon env_p sigma p in
+ let sigma, pj = pretype_type empty_valcon env_p sigma p in
let ccl = nf_evar sigma pj.utj_val in
let p = it_mkLambda_or_LetIn ccl psign' in
let inst =
@@ -831,17 +842,17 @@ let rec pretype ~program_mode k0 resolve_tc (tycon : type_constraint) (env : Glo
@[EConstr.of_constr (build_dependent_constructor cs)] in
let lp = lift cs.cs_nargs p in
let fty = hnf_lam_applist !!env sigma lp inst in
- let sigma, fj = pretype ~program_mode (mk_tycon fty) env_f sigma d in
+ let sigma, fj = pretype (mk_tycon fty) env_f sigma d in
let v =
let ind,_ = dest_ind_family indf in
- Typing.check_allowed_sort !!env sigma ind cj.uj_val p;
- obj ind p cj.uj_val fj.uj_val
- in
+ let rci = Typing.check_allowed_sort !!env sigma ind cj.uj_val p in
+ obj ind rci p cj.uj_val fj.uj_val
+ in
sigma, { uj_val = v; uj_type = (substl (realargs@[cj.uj_val]) ccl) }
| None ->
let tycon = lift_tycon cs.cs_nargs tycon in
- let sigma, fj = pretype ~program_mode tycon env_f sigma d in
+ let sigma, fj = pretype tycon env_f sigma d in
let ccl = nf_evar sigma fj.uj_type in
let ccl =
if noccur_between sigma 1 cs.cs_nargs ccl then
@@ -853,12 +864,12 @@ let rec pretype ~program_mode k0 resolve_tc (tycon : type_constraint) (env : Glo
let p = it_mkLambda_or_LetIn (lift (nar+1) ccl) psign' in
let v =
let ind,_ = dest_ind_family indf in
- Typing.check_allowed_sort !!env sigma ind cj.uj_val p;
- obj ind p cj.uj_val fj.uj_val
+ let rci = Typing.check_allowed_sort !!env sigma ind cj.uj_val p in
+ obj ind rci p cj.uj_val fj.uj_val
in sigma, { uj_val = v; uj_type = ccl })
| GIf (c,(na,po),b1,b2) ->
- let sigma, cj = pretype ~program_mode empty_tycon env sigma c in
+ let sigma, cj = pretype empty_tycon env sigma c in
let (IndType (indf,realargs)) =
try find_rectype !!env sigma cj.uj_type
with Not_found ->
@@ -869,21 +880,21 @@ let rec pretype ~program_mode k0 resolve_tc (tycon : type_constraint) (env : Glo
user_err ?loc
(str "If is only for inductive types with two constructors.");
- let arsgn =
- let arsgn,_ = get_arity !!env indf in
+ let arsgn, indr =
+ let arsgn,s = get_arity !!env indf in
(* Make dependencies from arity signature impossible *)
- List.map (set_name Anonymous) arsgn
+ List.map (set_name Anonymous) arsgn, Sorts.relevance_of_sort_family s
in
let nar = List.length arsgn in
let indt = build_dependent_inductive !!env indf in
- let psign = LocalAssum (na, indt) :: arsgn in (* For locating names in [po] *)
+ let psign = LocalAssum (make_annot na indr, indt) :: arsgn in (* For locating names in [po] *)
let psign = List.map (fun d -> map_rel_decl EConstr.of_constr d) psign in
let predenv = Cases.make_return_predicate_ltac_lvar env sigma na c cj.uj_val in
let hypnaming = if program_mode then ProgramNaming else KeepUserNameAndRenameExistingButSectionNames in
let psign,env_p = push_rel_context ~hypnaming sigma psign predenv in
let sigma, pred, p = match po with
| Some p ->
- let sigma, pj = pretype_type ~program_mode empty_valcon env_p sigma p in
+ let sigma, pj = pretype_type empty_valcon env_p sigma p in
let ccl = nf_evar sigma pj.utj_val in
let pred = it_mkLambda_or_LetIn ccl psign in
let typ = lift (- nar) (beta_applist sigma (pred,[cj.uj_val])) in
@@ -906,38 +917,38 @@ let rec pretype ~program_mode k0 resolve_tc (tycon : type_constraint) (env : Glo
List.map (set_name Anonymous) cs_args
in
let _,env_c = push_rel_context ~hypnaming sigma csgn env in
- let sigma, bj = pretype ~program_mode (mk_tycon pi) env_c sigma b in
+ let sigma, bj = pretype (mk_tycon pi) env_c sigma b in
sigma, it_mkLambda_or_LetIn bj.uj_val cs_args in
let sigma, b1 = f sigma cstrs.(0) b1 in
let sigma, b2 = f sigma cstrs.(1) b2 in
let v =
let ind,_ = dest_ind_family indf in
- let ci = make_case_info !!env (fst ind) IfStyle in
let pred = nf_evar sigma pred in
- Typing.check_allowed_sort !!env sigma ind cj.uj_val pred;
- mkCase (ci, pred, cj.uj_val, [|b1;b2|])
+ let rci = Typing.check_allowed_sort !!env sigma ind cj.uj_val pred in
+ let ci = make_case_info !!env (fst ind) rci IfStyle in
+ mkCase (ci, pred, cj.uj_val, [|b1;b2|])
in
let cj = { uj_val = v; uj_type = p } in
inh_conv_coerce_to_tycon ?loc env sigma cj tycon
| GCases (sty,po,tml,eqns) ->
- Cases.compile_cases ?loc ~program_mode sty (pretype ~program_mode, sigma) tycon env (po,tml,eqns)
+ Cases.compile_cases ?loc ~program_mode sty (pretype, sigma) tycon env (po,tml,eqns)
| GCast (c,k) ->
let sigma, cj =
match k with
| CastCoerce ->
- let sigma, cj = pretype ~program_mode empty_tycon env sigma c in
+ let sigma, cj = pretype empty_tycon env sigma c in
Coercion.inh_coerce_to_base ?loc ~program_mode !!env sigma cj
| CastConv t | CastVM t | CastNative t ->
let k = (match k with CastVM _ -> VMcast | CastNative _ -> NATIVEcast | _ -> DEFAULTcast) in
- let sigma, tj = pretype_type ~program_mode empty_valcon env sigma t in
+ let sigma, tj = pretype_type empty_valcon env sigma t in
let sigma, tval = Evarsolve.refresh_universes
~onlyalg:true ~status:Evd.univ_flexible (Some false) !!env sigma tj.utj_val in
let tval = nf_evar sigma tval in
let (sigma, cj), tval = match k with
| VMcast ->
- let sigma, cj = pretype ~program_mode empty_tycon env sigma c in
+ let sigma, cj = pretype empty_tycon env sigma c in
let cty = nf_evar sigma cj.uj_type and tval = nf_evar sigma tval in
if not (occur_existential sigma cty || occur_existential sigma tval) then
match Reductionops.vm_infer_conv !!env sigma cty tval with
@@ -948,7 +959,7 @@ let rec pretype ~program_mode k0 resolve_tc (tycon : type_constraint) (env : Glo
else user_err ?loc (str "Cannot check cast with vm: " ++
str "unresolved arguments remain.")
| NATIVEcast ->
- let sigma, cj = pretype ~program_mode empty_tycon env sigma c in
+ let sigma, cj = pretype empty_tycon env sigma c in
let cty = nf_evar sigma cj.uj_type and tval = nf_evar sigma tval in
begin
match Nativenorm.native_infer_conv !!env sigma cty tval with
@@ -958,7 +969,7 @@ let rec pretype ~program_mode k0 resolve_tc (tycon : type_constraint) (env : Glo
(ConversionFailed (!!env,cty,tval))
end
| _ ->
- pretype ~program_mode (mk_tycon tval) env sigma c, tval
+ pretype (mk_tycon tval) env sigma c, tval
in
let v = mkCast (cj.uj_val, k, tval) in
sigma, { uj_val = v; uj_type = tval }
diff --git a/pretyping/pretyping.mllib b/pretyping/pretyping.mllib
index d0359b43f4..34a6cecc95 100644
--- a/pretyping/pretyping.mllib
+++ b/pretyping/pretyping.mllib
@@ -5,10 +5,10 @@ Pretype_errors
Reductionops
Inductiveops
InferCumulativity
-Vnorm
Arguments_renaming
-Nativenorm
Retyping
+Vnorm
+Nativenorm
Cbv
Find_subterm
Evardefine
diff --git a/pretyping/reductionops.ml b/pretyping/reductionops.ml
index 98ca329117..71fbfe8716 100644
--- a/pretyping/reductionops.ml
+++ b/pretyping/reductionops.ml
@@ -12,6 +12,7 @@ open CErrors
open Util
open Names
open Constr
+open Context
open Termops
open Univ
open Evd
@@ -479,10 +480,10 @@ struct
| App (i,a,j) ->
let le = j - i + 1 in
App (0,Array.map f (Array.sub a i le), le-1)
- | Case (info,ty,br,alt) -> Case (info, f ty, Array.map f br, alt)
- | Fix ((r,(na,ty,bo)),arg,alt) ->
- Fix ((r,(na,Array.map f ty, Array.map f bo)),map f arg,alt)
- | Cst (cst,curr,remains,params,alt) ->
+ | Case (info,ty,br,alt) -> Case (info, f ty, Array.map f br, alt)
+ | Fix ((r,(na,ty,bo)),arg,alt) ->
+ Fix ((r,(na,Array.map f ty, Array.map f bo)),map f arg,alt)
+ | Cst (cst,curr,remains,params,alt) ->
Cst (cst,curr,remains,map f params,alt)
| Primitive (p,c,args,kargs,cst_l) ->
Primitive(p,c, map f args, kargs, cst_l)
@@ -775,7 +776,7 @@ let contract_cofix ?env sigma ?reference (bodynum,(names,types,bodies as typedbo
| Some e ->
match reference with
| None -> bd
- | Some r -> magicaly_constant_of_fixbody e sigma r bd names.(ind) in
+ | Some r -> magicaly_constant_of_fixbody e sigma r bd names.(ind).binder_name in
let closure = List.init nbodies make_Fi in
substl closure bodies.(bodynum)
@@ -817,7 +818,7 @@ let contract_fix ?env sigma ?reference ((recindices,bodynum),(names,types,bodies
| Some e ->
match reference with
| None -> bd
- | Some r -> magicaly_constant_of_fixbody e sigma r bd names.(ind) in
+ | Some r -> magicaly_constant_of_fixbody e sigma r bd names.(ind).binder_name in
let closure = List.init nbodies make_Fi in
substl closure bodies.(bodynum)
@@ -1062,7 +1063,7 @@ let rec whd_state_gen ?csts ~refold ~tactic_mode flags env sigma =
| Some _ when CClosure.RedFlags.red_set flags CClosure.RedFlags.fBETA ->
apply_subst (fun _ -> whrec) [] sigma refold cst_l x stack
| None when CClosure.RedFlags.red_set flags CClosure.RedFlags.fETA ->
- let env' = push_rel (LocalAssum (na, t)) env in
+ let env' = push_rel (LocalAssum (na, t)) env in
let whrec' = whd_state_gen ~refold ~tactic_mode flags env' sigma in
(match EConstr.kind sigma (Stack.zip ~refold sigma (fst (whrec' (c, Stack.empty)))) with
| App (f,cl) ->
@@ -1520,7 +1521,9 @@ let plain_instance sigma s c =
match EConstr.kind sigma g with
| App _ ->
let l' = Array.Fun1.Smart.map lift 1 l' in
- mkLetIn (Name default_plain_instance_ident,g,t,mkApp(mkRel 1, l'))
+ let r = Sorts.Relevant in (* TODO fix relevance *)
+ let na = make_annot (Name default_plain_instance_ident) r in
+ mkLetIn (na,g,t,mkApp(mkRel 1, l'))
| _ -> mkApp (g,l')
with Not_found -> mkApp (f,l'))
| _ -> mkApp (irec n f,l'))
@@ -1623,11 +1626,11 @@ let splay_prod_assum env sigma =
let t = whd_allnolet env sigma c in
match EConstr.kind sigma t with
| Prod (x,t,c) ->
- prodec_rec (push_rel (LocalAssum (x,t)) env)
- (Context.Rel.add (LocalAssum (x,t)) l) c
+ prodec_rec (push_rel (LocalAssum (x,t)) env)
+ (Context.Rel.add (LocalAssum (x,t)) l) c
| LetIn (x,b,t,c) ->
- prodec_rec (push_rel (LocalDef (x,b,t)) env)
- (Context.Rel.add (LocalDef (x,b,t)) l) c
+ prodec_rec (push_rel (LocalDef (x,b,t)) env)
+ (Context.Rel.add (LocalDef (x,b,t)) l) c
| Cast (c,_,_) -> prodec_rec env l c
| _ ->
let t' = whd_all env sigma t in
@@ -1648,8 +1651,8 @@ let splay_prod_n env sigma n =
let rec decrec env m ln c = if Int.equal m 0 then (ln,c) else
match EConstr.kind sigma (whd_all env sigma c) with
| Prod (n,a,c0) ->
- decrec (push_rel (LocalAssum (n,a)) env)
- (m-1) (Context.Rel.add (LocalAssum (n,a)) ln) c0
+ decrec (push_rel (LocalAssum (n,a)) env)
+ (m-1) (Context.Rel.add (LocalAssum (n,a)) ln) c0
| _ -> invalid_arg "splay_prod_n"
in
decrec env n Context.Rel.empty
@@ -1658,8 +1661,8 @@ let splay_lam_n env sigma n =
let rec decrec env m ln c = if Int.equal m 0 then (ln,c) else
match EConstr.kind sigma (whd_all env sigma c) with
| Lambda (n,a,c0) ->
- decrec (push_rel (LocalAssum (n,a)) env)
- (m-1) (Context.Rel.add (LocalAssum (n,a)) ln) c0
+ decrec (push_rel (LocalAssum (n,a)) env)
+ (m-1) (Context.Rel.add (LocalAssum (n,a)) ln) c0
| _ -> invalid_arg "splay_lam_n"
in
decrec env n Context.Rel.empty
diff --git a/pretyping/reductionops.mli b/pretyping/reductionops.mli
index fae0b23b83..5938d9b367 100644
--- a/pretyping/reductionops.mli
+++ b/pretyping/reductionops.mli
@@ -235,9 +235,9 @@ val hnf_lam_app : env -> evar_map -> constr -> constr -> constr
val hnf_lam_appvect : env -> evar_map -> constr -> constr array -> constr
val hnf_lam_applist : env -> evar_map -> constr -> constr list -> constr
-val splay_prod : env -> evar_map -> constr -> (Name.t * constr) list * constr
-val splay_lam : env -> evar_map -> constr -> (Name.t * constr) list * constr
-val splay_arity : env -> evar_map -> constr -> (Name.t * constr) list * ESorts.t
+val splay_prod : env -> evar_map -> constr -> (Name.t Context.binder_annot * constr) list * constr
+val splay_lam : env -> evar_map -> constr -> (Name.t Context.binder_annot * constr) list * constr
+val splay_arity : env -> evar_map -> constr -> (Name.t Context.binder_annot * constr) list * ESorts.t
val sort_of_arity : env -> evar_map -> constr -> ESorts.t
val splay_prod_n : env -> evar_map -> int -> constr -> rel_context * constr
val splay_lam_n : env -> evar_map -> int -> constr -> rel_context * constr
diff --git a/pretyping/retyping.ml b/pretyping/retyping.ml
index 1ca9e6590d..20120f4182 100644
--- a/pretyping/retyping.ml
+++ b/pretyping/retyping.ml
@@ -13,6 +13,7 @@ open CErrors
open Util
open Term
open Constr
+open Context
open Inductive
open Inductiveops
open Names
@@ -79,7 +80,8 @@ let rec subst_type env sigma typ = function
let sort_of_atomic_type env sigma ft args =
let rec concl_of_arity env n ar args =
match EConstr.kind sigma (whd_all env sigma ar), args with
- | Prod (na, t, b), h::l -> concl_of_arity (push_rel (LocalDef (na, lift n h, t)) env) (n + 1) b l
+ | Prod (na, t, b), h::l ->
+ concl_of_arity (push_rel (LocalDef (na, lift n h, t)) env) (n + 1) b l
| Sort s, [] -> ESorts.kind sigma s
| _ -> retype_error NotASort
in concl_of_arity env 0 ft (Array.to_list args)
@@ -188,7 +190,7 @@ let get_sort_family_of ?(truncation_style=false) ?(polyprop=true) env sigma t =
| 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
+ 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 Termops.is_template_polymorphic_ind env sigma f ->
@@ -256,3 +258,41 @@ let expand_projection env sigma pr c args =
in
mkApp (mkConstU (Projection.constant pr,u),
Array.of_list (ind_args @ (c :: args)))
+
+let relevance_of_term env sigma c =
+ if Environ.sprop_allowed env then
+ let rec aux rels c =
+ match kind sigma c with
+ | Rel n -> Retypeops.relevance_of_rel_extra env rels n
+ | Var x -> Retypeops.relevance_of_var env x
+ | Sort _ -> Sorts.Relevant
+ | Cast (c, _, _) -> aux rels c
+ | Prod ({binder_relevance=r}, _, codom) ->
+ aux (r::rels) codom
+ | Lambda ({binder_relevance=r}, _, bdy) ->
+ aux (r::rels) bdy
+ | LetIn ({binder_relevance=r}, _, _, bdy) ->
+ aux (r::rels) bdy
+ | App (c, _) -> aux rels c
+ | Const (c,_) -> Retypeops.relevance_of_constant env c
+ | Ind _ -> Sorts.Relevant
+ | Construct (c,_) -> Retypeops.relevance_of_constructor env c
+ | Case (ci, _, _, _) -> ci.ci_relevance
+ | Fix ((_,i),(lna,_,_)) -> (lna.(i)).binder_relevance
+ | CoFix (i,(lna,_,_)) -> (lna.(i)).binder_relevance
+ | Proj (p, _) -> Retypeops.relevance_of_projection env p
+ | Int _ -> Sorts.Relevant
+
+ | Meta _ | Evar _ -> Sorts.Relevant
+
+ in
+ aux [] c
+ else Sorts.Relevant
+
+let relevance_of_type env sigma t =
+ let s = get_sort_family_of env sigma t in
+ Sorts.relevance_of_sort_family s
+
+let relevance_of_sort s = Sorts.relevance_of_sort (EConstr.Unsafe.to_sorts s)
+
+let relevance_of_sort_family f = Sorts.relevance_of_sort_family f
diff --git a/pretyping/retyping.mli b/pretyping/retyping.mli
index 2aff0c7775..252bfb1a84 100644
--- a/pretyping/retyping.mli
+++ b/pretyping/retyping.mli
@@ -53,3 +53,8 @@ val sorts_of_context : env -> evar_map -> rel_context -> Sorts.t list
val expand_projection : env -> evar_map -> Names.Projection.t -> constr -> constr list -> constr
val print_retype_error : retype_error -> Pp.t
+
+val relevance_of_term : env -> evar_map -> constr -> Sorts.relevance
+val relevance_of_type : env -> evar_map -> types -> Sorts.relevance
+val relevance_of_sort : ESorts.t -> Sorts.relevance
+val relevance_of_sort_family : Sorts.family -> Sorts.relevance
diff --git a/pretyping/tacred.ml b/pretyping/tacred.ml
index 5db571433a..bcc20a41b4 100644
--- a/pretyping/tacred.ml
+++ b/pretyping/tacred.ml
@@ -13,6 +13,7 @@ open CErrors
open Util
open Names
open Constr
+open Context
open Libnames
open Globnames
open Termops
@@ -229,7 +230,8 @@ let check_fix_reversibility sigma labs args ((lv,i),(_,tys,bds)) =
(* Heuristic to look if global names are associated to other
components of a mutual fixpoint *)
-let invert_name labs l na0 env sigma ref = function
+let invert_name labs l {binder_name=na0} env sigma ref na =
+ match na.binder_name with
| Name id ->
let minfxargs = List.length l in
begin match na0 with
@@ -249,7 +251,7 @@ let invert_name labs l na0 env sigma ref = function
| Some c ->
let labs',ccl = decompose_lam sigma c in
let _, l' = whd_betalet_stack sigma ccl in
- let labs' = List.map snd labs' in
+ let labs' = List.map snd labs' in
(* ppedrot: there used to be generic equality on terms here *)
let eq_constr c1 c2 = EConstr.eq_constr sigma c1 c2 in
if List.equal eq_constr labs' labs &&
@@ -269,7 +271,7 @@ let compute_consteval_direct env sigma ref =
match EConstr.kind sigma c' with
| Lambda (id,t,g) when List.is_empty l && not onlyproj ->
let open Context.Rel.Declaration in
- srec (push_rel (LocalAssum (id,t)) env) (n+1) (t::labs) onlyproj g
+ srec (push_rel (LocalAssum (id,t)) env) (n+1) (t::labs) onlyproj g
| Fix fix when not onlyproj ->
(try check_fix_reversibility sigma labs l fix
with Elimconst -> NotAnElimination)
@@ -289,7 +291,7 @@ let compute_consteval_mutual_fix env sigma ref =
match EConstr.kind sigma c' with
| Lambda (na,t,g) when List.is_empty l ->
let open Context.Rel.Declaration in
- srec (push_rel (LocalAssum (na,t)) env) (minarg+1) (t::labs) ref g
+ srec (push_rel (LocalAssum (na,t)) env) (minarg+1) (t::labs) ref g
| Fix ((lv,i),(names,_,_)) ->
(* Last known constant wrapping Fix is ref = [labs](Fix l) *)
(match compute_consteval_direct env sigma ref with
@@ -374,7 +376,8 @@ let make_elim_fun (names,(nbfix,lv,n)) u largs =
List.fold_left_i (fun q (* j = n+1-q *) c (ij,tij) ->
let subst = List.map (Vars.lift (-q)) (List.firstn (n-ij) la) in
let tij' = Vars.substl (List.rev subst) tij in
- mkLambda (x,tij',c)) 1 body (List.rev lv)
+ let x = make_annot x Sorts.Relevant in (* TODO relevance *)
+ mkLambda (x,tij',c)) 1 body (List.rev lv)
in Some (minargs,g)
(* [f] is convertible to [Fix(recindices,bodynum),bodyvect)]:
@@ -384,7 +387,8 @@ let dummy = mkProp
let vfx = Id.of_string "_expanded_fix_"
let vfun = Id.of_string "_eliminator_function_"
let venv = let open Context.Named.Declaration in
- val_of_named_context [LocalAssum (vfx, dummy); LocalAssum (vfun, dummy)]
+ val_of_named_context [LocalAssum (make_annot vfx Sorts.Relevant, dummy);
+ LocalAssum (make_annot vfun Sorts.Relevant, dummy)]
(* Mark every occurrence of substituted vars (associated to a function)
as a problem variable: an evar that can be instantiated either by
@@ -513,7 +517,7 @@ let reduce_mind_case_use_function func env sigma mia =
let minargs = List.length mia.mcargs in
fun i ->
if Int.equal i bodynum then Some (minargs,func)
- else match names.(i) with
+ else match names.(i).binder_name with
| Anonymous -> None
| Name id ->
(* In case of a call to another component of a block of
@@ -627,12 +631,12 @@ let whd_nothing_for_iota env sigma s =
| Rel n ->
let open Context.Rel.Declaration in
(match lookup_rel n env with
- | LocalDef (_,body,_) -> whrec (lift n body, stack)
+ | LocalDef (_,body,_) -> whrec (lift n body, stack)
| _ -> s)
| Var id ->
let open Context.Named.Declaration in
(match lookup_named id env with
- | LocalDef (_,body,_) -> whrec (body, stack)
+ | LocalDef (_,body,_) -> whrec (body, stack)
| _ -> s)
| Evar ev -> s
| Meta ev ->
@@ -838,10 +842,10 @@ let try_red_product env sigma c =
| Cast (c,_,_) -> redrec env c
| Prod (x,a,b) ->
let open Context.Rel.Declaration in
- mkProd (x, a, redrec (push_rel (LocalAssum (x, a)) env) b)
+ mkProd (x, a, redrec (push_rel (LocalAssum (x, a)) env) b)
| LetIn (x,a,b,t) -> redrec env (Vars.subst1 a t)
| Case (ci,p,d,lf) -> simpfun (mkCase (ci,p,redrec env d,lf))
- | Proj (p, c) ->
+ | Proj (p, c) ->
let c' =
match EConstr.kind sigma c with
| Construct _ -> c
@@ -1150,6 +1154,7 @@ let compute = cbv_betadeltaiota
let abstract_scheme env sigma (locc,a) (c, sigma) =
let ta = Retyping.get_type_of env sigma a in
let na = named_hd env sigma ta Anonymous in
+ let na = make_annot na Sorts.Relevant in (* TODO relevance *)
if occur_meta sigma ta then user_err Pp.(str "Cannot find a type for the generalisation.");
if occur_meta sigma a then
mkLambda (na,ta,c), sigma
@@ -1192,7 +1197,7 @@ let reduce_to_ind_gen allow_product env sigma t =
| Prod (n,ty,t') ->
let open Context.Rel.Declaration in
if allow_product then
- elimrec (push_rel (LocalAssum (n,ty)) env) t' ((LocalAssum (n,ty))::l)
+ elimrec (push_rel (LocalAssum (n,ty)) env) t' ((LocalAssum (n,ty))::l)
else
user_err (str"Not an inductive definition.")
| _ ->
@@ -1270,7 +1275,7 @@ let reduce_to_ref_gen allow_product env sigma ref t =
| Prod (n,ty,t') ->
if allow_product then
let open Context.Rel.Declaration in
- elimrec (push_rel (LocalAssum (n,t)) env) t' ((LocalAssum (n,ty))::l)
+ elimrec (push_rel (LocalAssum (n,ty)) env) t' ((LocalAssum (n,ty))::l)
else
error_cannot_recognize ref
| _ ->
diff --git a/pretyping/typing.ml b/pretyping/typing.ml
index ed020e243d..b660973cdd 100644
--- a/pretyping/typing.ml
+++ b/pretyping/typing.ml
@@ -15,6 +15,7 @@ open CErrors
open Util
open Term
open Constr
+open Context
open Environ
open EConstr
open Vars
@@ -122,7 +123,7 @@ let max_sort l =
let is_correct_arity env sigma c pj ind specif params =
let arsign = make_arity_signature env sigma true (make_ind_family (ind,params)) in
let allowed_sorts = elim_sorts specif in
- let error () = Pretype_errors.error_elim_arity env sigma ind allowed_sorts c pj None in
+ let error () = Pretype_errors.error_elim_arity env sigma ind c pj None in
let rec srec env sigma pt ar =
let pt' = whd_all env sigma pt in
match EConstr.kind sigma pt', ar with
@@ -136,11 +137,11 @@ let is_correct_arity env sigma c pj ind specif params =
let s = ESorts.kind sigma s in
if not (Sorts.List.mem (Sorts.family s) allowed_sorts)
then error ()
- else sigma
+ else sigma, s
| Evar (ev,_), [] ->
let sigma, s = Evd.fresh_sort_in_family sigma (max_sort allowed_sorts) in
let sigma = Evd.define ev (mkSort s) sigma in
- sigma
+ sigma, s
| _, (LocalDef _ as d)::ar' ->
srec (push_rel d env) sigma (lift 1 pt') ar'
| _ ->
@@ -165,20 +166,20 @@ let type_case_branches env sigma (ind,largs) pj c =
let (params,realargs) = List.chop nparams largs in
let p = pj.uj_val in
let params = List.map EConstr.Unsafe.to_constr params in
- let sigma = is_correct_arity env sigma c pj ind specif params in
+ let sigma, ps = is_correct_arity env sigma c pj ind specif params in
let lc = build_branches_type ind specif params (EConstr.to_constr ~abort_on_undefined_evars:false sigma p) in
let lc = Array.map EConstr.of_constr lc in
let n = (snd specif).Declarations.mind_nrealdecls in
let ty = whd_betaiota sigma (lambda_applist_assum sigma (n+1) p (realargs@[c])) in
- sigma, (lc, ty)
+ sigma, (lc, ty, Sorts.relevance_of_sort ps)
let judge_of_case env sigma ci pj cj lfj =
let ((ind, u), spec) =
try find_mrectype env sigma cj.uj_type
with Not_found -> error_case_not_inductive env sigma cj in
let indspec = ((ind, EInstance.kind sigma u), spec) in
- let _ = check_case_info env (fst indspec) ci in
- let sigma, (bty,rslty) = type_case_branches env sigma indspec pj cj.uj_val in
+ let sigma, (bty,rslty,rci) = type_case_branches env sigma indspec pj cj.uj_val in
+ let () = check_case_info env (fst indspec) rci ci in
let sigma = check_branch_types env sigma (fst indspec) cj (lfj,bty) in
sigma, { uj_val = mkCase (ci, pj.uj_val, cj.uj_val, Array.map j_val lfj);
uj_type = rslty }
@@ -203,11 +204,13 @@ let check_allowed_sort env sigma ind c p =
let _, s = splay_prod env sigma pj.uj_type in
let ksort = match EConstr.kind sigma s with
| Sort s -> Sorts.family (ESorts.kind sigma s)
- | _ -> error_elim_arity env sigma ind sorts c pj None in
+ | _ -> error_elim_arity env sigma ind c pj None in
if not (List.exists ((==) ksort) sorts) then
let s = inductive_sort_family (snd specif) in
- error_elim_arity env sigma ind sorts c pj
- (Some(ksort,s,Type_errors.error_elim_explain ksort s))
+ error_elim_arity env sigma ind c pj
+ (Some(sorts,ksort,s,Type_errors.error_elim_explain ksort s))
+ else
+ Sorts.relevance_of_sort_family ksort
let judge_of_cast env sigma cj k tj =
let expected_type = tj.utj_val in
@@ -266,16 +269,19 @@ let judge_of_projection env sigma p cj =
uj_type = ty}
let judge_of_abstraction env name var j =
- { uj_val = mkLambda (name, var.utj_val, j.uj_val);
- uj_type = mkProd (name, var.utj_val, j.uj_type) }
+ let r = Sorts.relevance_of_sort var.utj_type in
+ { uj_val = mkLambda (make_annot name r, var.utj_val, j.uj_val);
+ uj_type = mkProd (make_annot name r, var.utj_val, j.uj_type) }
let judge_of_product env name t1 t2 =
+ let r = Sorts.relevance_of_sort t1.utj_type in
let s = sort_of_product env t1.utj_type t2.utj_type in
- { uj_val = mkProd (name, t1.utj_val, t2.utj_val);
+ { uj_val = mkProd (make_annot name r, t1.utj_val, t2.utj_val);
uj_type = mkSort s }
let judge_of_letin env name defj typj j =
- { uj_val = mkLetIn (name, defj.uj_val, typj.utj_val, j.uj_val) ;
+ let r = Sorts.relevance_of_sort typj.utj_type in
+ { uj_val = mkLetIn (make_annot name r, defj.uj_val, typj.utj_val, j.uj_val) ;
uj_type = subst1 defj.uj_val j.uj_type }
let check_hyps_inclusion env sigma f x hyps =
@@ -353,7 +359,7 @@ let rec execute env sigma cstr =
| Fix ((vn,i as vni),recdef) ->
let sigma, (_,tys,_ as recdef') = execute_recdef env sigma recdef in
- let fix = (vni,recdef') in
+ let fix = (vni,recdef') in
check_fix env sigma fix;
sigma, make_judge (mkFix fix) tys.(i)
@@ -391,26 +397,29 @@ let rec execute env sigma cstr =
| Lambda (name,c1,c2) ->
let sigma, j = execute env sigma c1 in
let sigma, var = type_judgment env sigma j in
- let env1 = push_rel (LocalAssum (name, var.utj_val)) env in
+ let () = check_binder_annot env var.utj_type name in
+ let env1 = push_rel (LocalAssum (name, var.utj_val)) env in
let sigma, j' = execute env1 sigma c2 in
- sigma, judge_of_abstraction env1 name var j'
+ sigma, judge_of_abstraction env1 name.binder_name var j'
| Prod (name,c1,c2) ->
let sigma, j = execute env sigma c1 in
let sigma, varj = type_judgment env sigma j in
- let env1 = push_rel (LocalAssum (name, varj.utj_val)) env in
+ let () = check_binder_annot env varj.utj_type name in
+ let env1 = push_rel (LocalAssum (name, varj.utj_val)) env in
let sigma, j' = execute env1 sigma c2 in
let sigma, varj' = type_judgment env1 sigma j' in
- sigma, judge_of_product env name varj varj'
+ sigma, judge_of_product env name.binder_name varj varj'
| LetIn (name,c1,c2,c3) ->
let sigma, j1 = execute env sigma c1 in
let sigma, j2 = execute env sigma c2 in
let sigma, j2 = type_judgment env sigma j2 in
let sigma, _ = judge_of_cast env sigma j1 DEFAULTcast j2 in
+ let () = check_binder_annot env j2.utj_type name in
let env1 = push_rel (LocalDef (name, j1.uj_val, j2.utj_val)) env in
let sigma, j3 = execute env1 sigma c3 in
- sigma, judge_of_letin env name j1 j2 j3
+ sigma, judge_of_letin env name.binder_name j1 j2 j3
| Cast (c,k,t) ->
let sigma, cj = execute env sigma c in
diff --git a/pretyping/typing.mli b/pretyping/typing.mli
index 9bd369a155..f68820429b 100644
--- a/pretyping/typing.mli
+++ b/pretyping/typing.mli
@@ -39,12 +39,12 @@ val solve_evars : env -> evar_map -> constr -> evar_map * constr
(** Raise an error message if incorrect elimination for this inductive
(first constr is term to match, second is return predicate) *)
val check_allowed_sort : env -> evar_map -> pinductive -> constr -> constr ->
- unit
+ Sorts.relevance
(** Raise an error message if bodies have types not unifiable with the
expected ones *)
val check_type_fixpoint : ?loc:Loc.t -> env -> evar_map ->
- Names.Name.t array -> types array -> unsafe_judgment array -> evar_map
+ Names.Name.t Context.binder_annot array -> types array -> unsafe_judgment array -> evar_map
val judge_of_sprop : unsafe_judgment
val judge_of_prop : unsafe_judgment
diff --git a/pretyping/unification.ml b/pretyping/unification.ml
index 3de8c381d0..9ba51dcfa9 100644
--- a/pretyping/unification.ml
+++ b/pretyping/unification.ml
@@ -13,6 +13,7 @@ open Pp
open Util
open Names
open Constr
+open Context
open Termops
open Environ
open EConstr
@@ -103,13 +104,13 @@ let occur_meta_evd sigma mv c =
let abstract_scheme env evd c l lname_typ =
let mkLambda_name env (n,a,b) =
- mkLambda (named_hd env evd a n, a, b)
+ mkLambda (map_annot (named_hd env evd a) n, a, b)
in
List.fold_left2
(fun (t,evd) (locc,a) decl ->
- let na = RelDecl.get_name decl in
+ let na = RelDecl.get_annot decl in
let ta = RelDecl.get_type decl in
- let na = match EConstr.kind evd a with Var id -> Name id | _ -> na in
+ let na = match EConstr.kind evd a with Var id -> {na with binder_name=Name id} | _ -> na in
(* [occur_meta ta] test removed for support of eelim/ecase but consequences
are unclear...
if occur_meta ta then error "cannot find a type for the generalisation"
@@ -117,7 +118,7 @@ let abstract_scheme env evd c l lname_typ =
if occur_meta evd a then mkLambda_name env (na,ta,t), evd
else
let t', evd' = Find_subterm.subst_closed_term_occ env evd locc a t in
- mkLambda_name env (na,ta,t'), evd')
+ mkLambda_name env (na,ta,t'), evd')
(c,evd)
(List.rev l)
lname_typ
@@ -561,8 +562,8 @@ let is_rigid_head sigma flags t =
| Ind (i,u) -> true
| Construct _ | Int _ -> true
| Fix _ | CoFix _ -> true
- | Rel _ | Var _ | Meta _ | Evar _ | Sort _ | Cast (_, _, _) | Prod (_, _, _)
- | Lambda (_, _, _) | LetIn (_, _, _, _) | App (_, _) | Case (_, _, _, _)
+ | Rel _ | Var _ | Meta _ | Evar _ | Sort _ | Cast (_, _, _) | Prod _
+ | Lambda _ | LetIn _ | App (_, _) | Case (_, _, _, _)
| Proj (_, _) -> false (* Why aren't Prod, Sort rigid heads ? *)
let force_eqs c =
@@ -662,7 +663,7 @@ let is_eta_constructor_app env sigma ts f l1 term =
let mib = lookup_mind (fst ind) env in
(match mib.Declarations.mind_record with
| PrimRecord info when mib.Declarations.mind_finite == Declarations.BiFinite &&
- let (_, projs, _) = info.(i) in
+ let (_, projs, _, _) = info.(i) in
Array.length projs == Array.length l1 - mib.Declarations.mind_nparams ->
(* Check that the other term is neutral *)
is_neutral env sigma ts term
@@ -782,14 +783,14 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst : subst0) conv_at_top e
with e when CErrors.noncritical e ->
error_cannot_unify curenv sigma (m,n))
- | Lambda (na,t1,c1), Lambda (_,t2,c2) ->
- unirec_rec (push (na,t1) curenvnb) CONV {opt with at_top = true}
+ | Lambda (na,t1,c1), Lambda (__,t2,c2) ->
+ unirec_rec (push (na,t1) curenvnb) CONV {opt with at_top = true}
(unirec_rec curenvnb CONV {opt with at_top = true; with_types = false} substn t1 t2) c1 c2
- | Prod (na,t1,c1), Prod (_,t2,c2) ->
- unirec_rec (push (na,t1) curenvnb) pb {opt with at_top = true}
+ | Prod (na,t1,c1), Prod (_,t2,c2) ->
+ unirec_rec (push (na,t1) curenvnb) pb {opt with at_top = true}
(unirec_rec curenvnb CONV {opt with at_top = true; with_types = false} substn t1 t2) c1 c2
- | LetIn (_,a,_,c), _ -> unirec_rec curenvnb pb opt substn (subst1 a c) cN
- | _, LetIn (_,a,_,c) -> unirec_rec curenvnb pb opt substn cM (subst1 a c)
+ | LetIn (_,a,_,c), _ -> unirec_rec curenvnb pb opt substn (subst1 a c) cN
+ | _, LetIn (_,a,_,c) -> unirec_rec curenvnb pb opt substn cM (subst1 a c)
(* Fast path for projections. *)
| Proj (p1,c1), Proj (p2,c2) when Constant.equal
@@ -800,11 +801,11 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst : subst0) conv_at_top e
unify_not_same_head curenvnb pb opt substn cM cN)
(* eta-expansion *)
- | Lambda (na,t1,c1), _ when flags.modulo_eta ->
- unirec_rec (push (na,t1) curenvnb) CONV {opt with at_top = true} substn
+ | Lambda (na,t1,c1), _ when flags.modulo_eta ->
+ unirec_rec (push (na,t1) curenvnb) CONV {opt with at_top = true} substn
c1 (mkApp (lift 1 cN,[|mkRel 1|]))
- | _, Lambda (na,t2,c2) when flags.modulo_eta ->
- unirec_rec (push (na,t2) curenvnb) CONV {opt with at_top = true} substn
+ | _, Lambda (na,t2,c2) when flags.modulo_eta ->
+ unirec_rec (push (na,t2) curenvnb) CONV {opt with at_top = true} substn
(mkApp (lift 1 cM,[|mkRel 1|])) c2
(* For records *)
@@ -1775,7 +1776,7 @@ let w_unify_to_subterm env evd ?(flags=default_unify_flags ()) (op,cl) =
matchrec c
with ex when precatchable_exception ex ->
iter_fail matchrec lf)
- | LetIn(_,c1,_,c2) ->
+ | LetIn(_,c1,_,c2) ->
(try
matchrec c1
with ex when precatchable_exception ex ->
@@ -1783,13 +1784,13 @@ let w_unify_to_subterm env evd ?(flags=default_unify_flags ()) (op,cl) =
| Proj (p,c) -> matchrec c
- | Fix(_,(_,types,terms)) ->
+ | Fix(_,(_,types,terms)) ->
(try
iter_fail matchrec types
with ex when precatchable_exception ex ->
iter_fail matchrec terms)
- | CoFix(_,(_,types,terms)) ->
+ | CoFix(_,(_,types,terms)) ->
(try
iter_fail matchrec types
with ex when precatchable_exception ex ->
@@ -1860,13 +1861,13 @@ let w_unify_to_subterm_all env evd ?(flags=default_unify_flags ()) (op,cl) =
| Proj (p,c) -> matchrec c
- | LetIn(_,c1,_,c2) ->
+ | LetIn(_,c1,_,c2) ->
bind (matchrec c1) (matchrec c2)
- | Fix(_,(_,types,terms)) ->
+ | Fix(_,(_,types,terms)) ->
bind (bind_iter matchrec types) (bind_iter matchrec terms)
- | CoFix(_,(_,types,terms)) ->
+ | CoFix(_,(_,types,terms)) ->
bind (bind_iter matchrec types) (bind_iter matchrec terms)
| Prod (_,t,c) ->
diff --git a/pretyping/vnorm.ml b/pretyping/vnorm.ml
index ff528bd2cf..62e9e477f7 100644
--- a/pretyping/vnorm.ml
+++ b/pretyping/vnorm.ml
@@ -13,6 +13,7 @@ open Names
open Declarations
open Term
open Constr
+open Context
open Vars
open Environ
open Inductive
@@ -31,10 +32,12 @@ module NamedDecl = Context.Named.Declaration
let crazy_type = mkSet
let decompose_prod env t =
- let (name,dom,codom as res) = destProd (whd_all env t) in
- match name with
- | Anonymous -> (Name (Id.of_string "x"), dom, codom)
- | Name _ -> res
+ let (name,dom,codom) = destProd (whd_all env t) in
+ let name = map_annot (function
+ | Anonymous -> Name (Id.of_string "x")
+ | Name _ as na -> na) name
+ in
+ (name,dom,codom)
exception Find_at of int
@@ -138,6 +141,8 @@ and nf_whd env sigma whd typ =
let dom = nf_vtype env sigma (dom p) in
let name = Name (Id.of_string "x") in
let vc = reduce_fun (nb_rel env) (codom p) in
+ let r = Retyping.relevance_of_type env sigma (EConstr.of_constr dom) in
+ let name = make_annot name r in
let codom = nf_vtype (push_rel (LocalAssum (name,dom)) env) sigma vc in
mkProd(name,dom,codom)
| Vfun f -> nf_fun env sigma f typ
@@ -307,6 +312,8 @@ and nf_predicate env sigma ind mip params v pT =
let rargs = Array.init n (fun i -> mkRel (n-i)) in
let params = if Int.equal n 0 then params else Array.map (lift n) params in
let dom = mkApp(mkIndU ind,Array.append params rargs) in
+ let r = Inductive.relevance_of_inductive env (fst ind) in
+ let name = make_annot name r in
let body = nf_vtype (push_rel (LocalAssum (name,dom)) env) sigma vb in
mkLambda(name,dom,body)
| _ -> assert false
@@ -317,7 +324,7 @@ and nf_args env sigma vargs ?from:(f=0) t =
let args =
Array.init len
(fun i ->
- let _,dom,codom = decompose_prod env !t in
+ let _,dom,codom = decompose_prod env !t in
let c = nf_val env sigma (arg vargs (f+i)) dom in
t := subst1 c codom; c) in
!t,args
@@ -328,7 +335,7 @@ and nf_bargs env sigma b ofs t =
let args =
Array.init len
(fun i ->
- let _,dom,codom = decompose_prod env !t in
+ let _,dom,codom = decompose_prod env !t in
let c = nf_val env sigma (bfield b (i+ofs)) dom in
t := subst1 c codom; c) in
args
@@ -353,14 +360,17 @@ and nf_fix env sigma f =
let vb, vt = reduce_fix k f in
let ndef = Array.length vt in
let ft = Array.map (fun v -> nf_val env sigma v crazy_type) vt in
- let name = Array.init ndef (fun _ -> (Name (Id.of_string "Ffix"))) in
- (* Third argument of the tuple is ignored by push_rec_types *)
- let env = push_rec_types (name,ft,ft) env in
+ let name = Name (Id.of_string "Ffix") in
+ let names = Array.map (fun t ->
+ make_annot name @@
+ Retyping.relevance_of_type env sigma (EConstr.of_constr t)) ft in
+ (* Body argument of the tuple is ignored by push_rec_types *)
+ let env = push_rec_types (names,ft,ft) env in
(* We lift here because the types of arguments (in tt) will be evaluated
in an environment where the fixpoints have been pushed *)
let norm_vb v t = nf_fun env sigma v (lift ndef t) in
let fb = Util.Array.map2 norm_vb vb ft in
- mkFix ((rec_args,init),(name,ft,fb))
+ mkFix ((rec_args,init),(names,ft,fb))
and nf_fix_app env sigma f vargs =
let fd = nf_fix env sigma f in
@@ -373,12 +383,14 @@ and nf_cofix env sigma cf =
let init = current_cofix cf in
let k = nb_rel env in
let vb,vt = reduce_cofix k cf in
- let ndef = Array.length vt in
let cft = Array.map (fun v -> nf_val env sigma v crazy_type) vt in
- let name = Array.init ndef (fun _ -> (Name (Id.of_string "Fcofix"))) in
- let env = push_rec_types (name,cft,cft) env in
+ let name = Name (Id.of_string "Fcofix") in
+ let names = Array.map (fun t ->
+ make_annot name @@
+ Retyping.relevance_of_type env sigma (EConstr.of_constr t)) cft in
+ let env = push_rec_types (names,cft,cft) env in
let cfb = Util.Array.map2 (fun v t -> nf_val env sigma v t) vb cft in
- mkCoFix (init,(name,cft,cfb))
+ mkCoFix (init,(names,cft,cfb))
let cbv_vm env sigma c t =
if Termops.occur_meta sigma c then
diff --git a/printing/prettyp.ml b/printing/prettyp.ml
index 797b6faa08..8bf86e9ef6 100644
--- a/printing/prettyp.ml
+++ b/printing/prettyp.ml
@@ -505,9 +505,9 @@ let gallina_print_named_decl env sigma =
let open Context.Named.Declaration in
function
| LocalAssum (id, typ) ->
- print_named_assum env sigma (Id.to_string id) typ
+ print_named_assum env sigma (Id.to_string id.Context.binder_name) typ
| LocalDef (id, body, typ) ->
- print_named_def env sigma (Id.to_string id) body typ
+ print_named_def env sigma (Id.to_string id.Context.binder_name) body typ
let assumptions_for_print lna =
List.fold_right (fun na env -> add_name na env) lna empty_names_context
diff --git a/printing/printer.ml b/printing/printer.ml
index bc936975c2..fa55a28cb3 100644
--- a/printing/printer.ml
+++ b/printing/printer.ml
@@ -13,6 +13,7 @@ open CErrors
open Util
open Names
open Constr
+open Context
open Environ
open Globnames
open Evd
@@ -100,7 +101,7 @@ let pr_constr_under_binders_env_gen pr env sigma (ids,c) =
(* Warning: clashes can occur with variables of same name in env but *)
(* we also need to preserve the actual names of the patterns *)
(* So what to do? *)
- let assums = List.map (fun id -> (Name id,(* dummy *) mkProp)) ids in
+ let assums = List.map (fun id -> (make_annot (Name id) Sorts.Relevant,(* dummy *) mkProp)) ids in
pr (Termops.push_rels_assum assums env) sigma c
let pr_constr_under_binders_env = pr_constr_under_binders_env_gen pr_econstr_env
@@ -290,7 +291,7 @@ let pr_compacted_decl env sigma decl =
let pb = if isCast c then surround pb else pb in
ids, (str" := " ++ pb ++ cut ()), typ
in
- let pids = prlist_with_sep pr_comma pr_id ids in
+ let pids = prlist_with_sep pr_comma (fun id -> pr_id id.binder_name) ids in
let pt = pr_ltype_env env sigma typ in
let ptyp = (str" : " ++ pt) in
hov 0 (pids ++ pbody ++ ptyp)
diff --git a/printing/printmod.ml b/printing/printmod.ml
index 3438063f76..f4986652b3 100644
--- a/printing/printmod.ml
+++ b/printing/printmod.ml
@@ -10,6 +10,7 @@
open Util
open Constr
+open Context
open Pp
open Names
open Environ
@@ -132,10 +133,10 @@ let get_fields =
let rec prodec_rec l subst c =
match kind c with
| Prod (na,t,c) ->
- let id = match na with Name id -> id | Anonymous -> Id.of_string "_" in
+ let id = match na.binder_name with Name id -> id | Anonymous -> Id.of_string "_" in
prodec_rec ((id,true,Vars.substl subst t)::l) (mkVar id::subst) c
| LetIn (na,b,_,c) ->
- let id = match na with Name id -> id | Anonymous -> Id.of_string "_" in
+ let id = match na.binder_name with Name id -> id | Anonymous -> Id.of_string "_" in
prodec_rec ((id,false,Vars.substl subst b)::l) (mkVar id::subst) c
| _ -> List.rev l
in
diff --git a/printing/proof_diffs.ml b/printing/proof_diffs.ml
index 878e9f477b..5aa7b3c7bd 100644
--- a/printing/proof_diffs.ml
+++ b/printing/proof_diffs.ml
@@ -204,14 +204,14 @@ let diff_hyps o_line_idents o_map n_line_idents n_map =
List.rev !rv;;
-type 'a hyp = (Names.Id.t list * 'a option * 'a)
+type 'a hyp = (Names.Id.t Context.binder_annot list * 'a option * 'a)
type 'a reified_goal = { name: string; ty: 'a; hyps: 'a hyp list; env : Environ.env; sigma: Evd.evar_map }
(* XXX: Port to proofview, one day. *)
(* open Proofview *)
module CDC = Context.Compacted.Declaration
-let to_tuple : Constr.compacted_declaration -> (Names.Id.t list * 'pc option * 'pc) =
+let to_tuple : Constr.compacted_declaration -> (Names.Id.t Context.binder_annot list * 'pc option * 'pc) =
let open CDC in function
| LocalAssum(idl, tm) -> (idl, None, EConstr.of_constr tm)
| LocalDef(idl,tdef,tm) -> (idl, Some (EConstr.of_constr tdef), EConstr.of_constr tm);;
@@ -283,7 +283,7 @@ let goal_info goal sigma =
let build_hyp_info env sigma hyp =
let (names, body, ty) = hyp in
let open Pp in
- let idents = List.map (fun x -> Names.Id.to_string x) names in
+ let idents = List.map (fun x -> Names.Id.to_string x.Context.binder_name) names in
line_idents := idents :: !line_idents;
let mid = match body with
diff --git a/proofs/clenv.ml b/proofs/clenv.ml
index 9540d3de44..2d2113b636 100644
--- a/proofs/clenv.ml
+++ b/proofs/clenv.ml
@@ -15,6 +15,7 @@ open Names
open Nameops
open Termops
open Constr
+open Context
open Namegen
open Environ
open Evd
@@ -69,7 +70,7 @@ let clenv_push_prod cl =
| Prod (na,t,u) ->
let mv = new_meta () in
let dep = not (noccurn (cl_sigma cl) 1 u) in
- let na' = if dep then na else Anonymous in
+ let na' = if dep then na.binder_name else Anonymous in
let e' = meta_declare mv t ~name:na' cl.evd in
let concl = if dep then subst1 (mkMeta mv) u else u in
let def = applist (cl.templval.rebus,[mkMeta mv]) in
@@ -103,7 +104,7 @@ let clenv_environments evd bound t =
| (n, Prod (na,t1,t2)) ->
let mv = new_meta () in
let dep = not (noccurn evd 1 t2) in
- let na' = if dep then na else Anonymous in
+ let na' = if dep then na.binder_name else Anonymous in
let e' = meta_declare mv t1 ~name:na' e in
clrec (e', (mkMeta mv)::metas) (Option.map ((+) (-1)) n)
(if dep then (subst1 (mkMeta mv) t2) else t2)
@@ -277,7 +278,7 @@ let adjust_meta_source evd mv = function
| loc,Evar_kinds.VarInstance id ->
let rec match_name c l =
match EConstr.kind evd c, l with
- | Lambda (Name id,_,c), a::l when EConstr.eq_constr evd a (mkMeta mv) -> Some id
+ | Lambda ({binder_name=Name id},_,c), a::l when EConstr.eq_constr evd a (mkMeta mv) -> Some id
| Lambda (_,_,c), a::l -> match_name c l
| _ -> None in
(* This is very ad hoc code so that an evar inherits the name of the binder
@@ -623,7 +624,7 @@ let make_evar_clause env sigma ?len t =
hole_type = t1;
hole_deps = dep;
(* We fix it later *)
- hole_name = na;
+ hole_name = na.binder_name;
} in
let t2 = if dep then subst1 ev t2 else t2 in
clrec (sigma, hole :: holes) inst (pred n) t2
diff --git a/proofs/proof_global.ml b/proofs/proof_global.ml
index a47fa78f4d..6174b75a96 100644
--- a/proofs/proof_global.ml
+++ b/proofs/proof_global.ml
@@ -19,6 +19,7 @@
open Util
open Pp
open Names
+open Context
module NamedDecl = Context.Named.Declaration
@@ -198,10 +199,10 @@ let set_used_variables l =
let vars_of = Environ.global_vars_set in
let aux env entry (ctx, all_safe as orig) =
match entry with
- | LocalAssum (x,_) ->
+ | LocalAssum ({binder_name=x},_) ->
if Id.Set.mem x all_safe then orig
else (ctx, all_safe)
- | LocalDef (x,bo, ty) as decl ->
+ | LocalDef ({binder_name=x},bo, ty) as decl ->
if Id.Set.mem x all_safe then orig else
let vars = Id.Set.union (vars_of env bo) (vars_of env ty) in
if Id.Set.subset vars all_safe
diff --git a/proofs/tacmach.ml b/proofs/tacmach.ml
index df90354717..8196f5e198 100644
--- a/proofs/tacmach.ml
+++ b/proofs/tacmach.ml
@@ -172,7 +172,7 @@ module New = struct
let env = Proofview.Goal.env gl in
let sign = Environ.named_context env in
List.map (function LocalAssum (id,x)
- | LocalDef (id,_,x) -> id, EConstr.of_constr x)
+ | LocalDef (id,_,x) -> id.Context.binder_name, EConstr.of_constr x)
sign
let pf_last_hyp gl =
diff --git a/proofs/tacmach.mli b/proofs/tacmach.mli
index 213ed7bfda..1454140dd7 100644
--- a/proofs/tacmach.mli
+++ b/proofs/tacmach.mli
@@ -29,7 +29,7 @@ val pf_concl : Goal.goal sigma -> types
val pf_env : Goal.goal sigma -> env
val pf_hyps : Goal.goal sigma -> named_context
(*i val pf_untyped_hyps : Goal.goal sigma -> (Id.t * constr) list i*)
-val pf_hyps_types : Goal.goal sigma -> (Id.t * types) list
+val pf_hyps_types : Goal.goal sigma -> (Id.t Context.binder_annot * types) list
val pf_nth_hyp_id : Goal.goal sigma -> int -> Id.t
val pf_last_hyp : Goal.goal sigma -> named_declaration
val pf_ids_of_hyps : Goal.goal sigma -> Id.t list
diff --git a/tactics/btermdn.ml b/tactics/btermdn.ml
index 2f2bd8d2bc..06246ef584 100644
--- a/tactics/btermdn.ml
+++ b/tactics/btermdn.ml
@@ -77,7 +77,7 @@ let constr_val_discr_st sigma ts t =
| Construct (cstr_sp,u) -> Label(GRLabel (ConstructRef cstr_sp),l)
| Var id when not (TransparentState.is_transparent_variable ts id) -> Label(GRLabel (VarRef id),l)
| Prod (n, d, c) -> Label(ProdLabel, [d; c])
- | Lambda (n, d, c) ->
+ | Lambda (n, d, c) ->
if List.is_empty l then
Label(LambdaLabel, [d; c] @ l)
else Everything
diff --git a/tactics/contradiction.ml b/tactics/contradiction.ml
index bd95a62532..3ff2e3852d 100644
--- a/tactics/contradiction.ml
+++ b/tactics/contradiction.ml
@@ -9,6 +9,7 @@
(************************************************************************)
open Constr
+open Context
open EConstr
open Hipattern
open Tactics
@@ -19,10 +20,10 @@ module NamedDecl = Context.Named.Declaration
(* Absurd *)
-let mk_absurd_proof coq_not t =
+let mk_absurd_proof coq_not r t =
let id = Namegen.default_dependent_ident in
- mkLambda (Names.Name id,mkApp(coq_not,[|t|]),
- mkLambda (Names.Name id,t,mkApp (mkRel 2,[|mkRel 1|])))
+ mkLambda (make_annot (Names.Name id) Sorts.Relevant,mkApp(coq_not,[|t|]),
+ mkLambda (make_annot (Names.Name id) r,t,mkApp (mkRel 2,[|mkRel 1|])))
let absurd c =
Proofview.Goal.enter begin fun gl ->
@@ -31,12 +32,13 @@ let absurd c =
let j = Retyping.get_judgment_of env sigma c in
let sigma, j = Coercion.inh_coerce_to_sort env sigma j in
let t = j.Environ.utj_val in
+ let r = Sorts.relevance_of_sort j.Environ.utj_type in
Proofview.Unsafe.tclEVARS sigma <*>
Tacticals.New.pf_constr_of_global (Coqlib.(lib_ref "core.not.type")) >>= fun coqnot ->
Tacticals.New.pf_constr_of_global (Coqlib.(lib_ref "core.False.type")) >>= fun coqfalse ->
Tacticals.New.tclTHENLIST [
elim_type coqfalse;
- Simple.apply (mk_absurd_proof coqnot t)
+ Simple.apply (mk_absurd_proof coqnot r t)
]
end
@@ -68,9 +70,9 @@ let contradiction_context =
if is_empty_type sigma typ then
simplest_elim (mkVar id)
else match EConstr.kind sigma typ with
- | Prod (na,t,u) when is_empty_type sigma u ->
+ | Prod (na,t,u) when is_empty_type sigma u ->
let is_unit_or_eq = match_with_unit_or_eq_type sigma t in
- Tacticals.New.tclORELSE
+ Tacticals.New.tclORELSE
(match is_unit_or_eq with
| Some _ ->
let hd,args = decompose_app sigma t in
diff --git a/tactics/elimschemes.ml b/tactics/elimschemes.ml
index 3b69d9922d..1fae4c3d9d 100644
--- a/tactics/elimschemes.ml
+++ b/tactics/elimschemes.ml
@@ -88,14 +88,27 @@ let ind_scheme_kind_from_type =
declare_individual_scheme_object "_ind_nodep"
(optimize_non_type_induction_scheme rec_scheme_kind_from_type false InProp)
+let sind_scheme_kind_from_type =
+ declare_individual_scheme_object "_sind_nodep"
+ (fun _ x -> build_induction_scheme_in_type false InSProp x, Safe_typing.empty_private_constants)
+
let ind_dep_scheme_kind_from_type =
declare_individual_scheme_object "_ind" ~aux:"_ind_from_type"
(optimize_non_type_induction_scheme rec_dep_scheme_kind_from_type true InProp)
+let sind_dep_scheme_kind_from_type =
+ declare_individual_scheme_object "_sind" ~aux:"_sind_from_type"
+ (fun _ x -> build_induction_scheme_in_type true InSProp x, Safe_typing.empty_private_constants)
+
let ind_scheme_kind_from_prop =
declare_individual_scheme_object "_ind" ~aux:"_ind_from_prop"
(optimize_non_type_induction_scheme rec_scheme_kind_from_prop false InProp)
+let sind_scheme_kind_from_prop =
+ declare_individual_scheme_object "_sind" ~aux:"_sind_from_prop"
+ (fun _ x -> build_induction_scheme_in_type false InSProp x, Safe_typing.empty_private_constants)
+
+
(* Case analysis *)
let build_case_analysis_scheme_in_type dep sort ind =
diff --git a/tactics/elimschemes.mli b/tactics/elimschemes.mli
index ece4124b8b..4472792449 100644
--- a/tactics/elimschemes.mli
+++ b/tactics/elimschemes.mli
@@ -22,11 +22,14 @@ val optimize_non_type_induction_scheme :
val rect_scheme_kind_from_prop : individual scheme_kind
val ind_scheme_kind_from_prop : individual scheme_kind
+val sind_scheme_kind_from_prop : individual scheme_kind
val rec_scheme_kind_from_prop : individual scheme_kind
val rect_scheme_kind_from_type : individual scheme_kind
val rect_dep_scheme_kind_from_type : individual scheme_kind
val ind_scheme_kind_from_type : individual scheme_kind
val ind_dep_scheme_kind_from_type : individual scheme_kind
+val sind_scheme_kind_from_type : individual scheme_kind
+val sind_dep_scheme_kind_from_type : individual scheme_kind
val rec_scheme_kind_from_type : individual scheme_kind
val rec_dep_scheme_kind_from_type : individual scheme_kind
diff --git a/tactics/eqdecide.ml b/tactics/eqdecide.ml
index 6388aa2c33..e75a61d0c6 100644
--- a/tactics/eqdecide.ml
+++ b/tactics/eqdecide.ml
@@ -18,6 +18,7 @@ open Util
open Names
open Namegen
open Constr
+open Context
open EConstr
open Declarations
open Tactics
@@ -74,7 +75,8 @@ let generalize_right mk typ c1 c2 =
let env = Proofview.Goal.env gl in
Refine.refine ~typecheck:false begin fun sigma ->
let na = Name (next_name_away_with_default "x" Anonymous (Termops.vars_of_env env)) in
- let newconcl = mkProd (na, typ, mk typ c1 (mkRel 1)) in
+ let r = Retyping.relevance_of_type env sigma typ in
+ let newconcl = mkProd (make_annot na r, typ, mk typ c1 (mkRel 1)) in
let (sigma, x) = Evarutil.new_evar env sigma ~principal:true newconcl in
(sigma, mkApp (x, [|c2|]))
end
@@ -123,8 +125,8 @@ let mkGenDecideEqGoal rectype ops g =
let hypnames = pf_ids_set_of_hyps g in
let xname = next_ident_away idx hypnames
and yname = next_ident_away idy hypnames in
- (mkNamedProd xname rectype
- (mkNamedProd yname rectype
+ (mkNamedProd (make_annot xname Sorts.Relevant) rectype
+ (mkNamedProd (make_annot yname Sorts.Relevant) rectype
(mkDecideEqGoal true ops
rectype (mkVar xname) (mkVar yname))))
diff --git a/tactics/eqschemes.ml b/tactics/eqschemes.ml
index 554fe3cbed..073d66e4aa 100644
--- a/tactics/eqschemes.ml
+++ b/tactics/eqschemes.ml
@@ -51,6 +51,7 @@ open Util
open Names
open Term
open Constr
+open Context
open Vars
open Declarations
open Environ
@@ -80,8 +81,8 @@ let build_dependent_inductive ind (mib,mip) =
let named_hd env t na = named_hd env (Evd.from_env env) (EConstr.of_constr t) na
let name_assumption env = function
-| LocalAssum (na,t) -> LocalAssum (named_hd env t na, t)
-| LocalDef (na,c,t) -> LocalDef (named_hd env c na, c, t)
+| LocalAssum (na,t) -> LocalAssum (map_annot (named_hd env t) na, t)
+| LocalDef (na,c,t) -> LocalDef (map_annot (named_hd env c) na, c, t)
let name_context env hyps =
snd
@@ -202,11 +203,14 @@ let build_sym_scheme env ind =
get_sym_eq_data env indu in
let cstr n =
mkApp (mkConstructUi(indu,1),Context.Rel.to_extended_vect mkRel n mib.mind_params_ctxt) in
- let varH = fresh env (default_id_of_sort (snd (mind_arity mip))) in
+ let inds = snd (mind_arity mip) in
+ let varH = fresh env (default_id_of_sort inds) in
let applied_ind = build_dependent_inductive indu specif in
+ let indr = Sorts.relevance_of_sort_family inds in
let realsign_ind =
- name_context env ((LocalAssum (Name varH,applied_ind))::realsign) in
- let ci = make_case_info (Global.env()) ind RegularStyle in
+ name_context env ((LocalAssum (make_annot (Name varH) indr,applied_ind))::realsign) in
+ let rci = Sorts.Relevant in (* TODO relevance *)
+ let ci = make_case_info (Global.env()) ind rci RegularStyle in
let c =
(my_it_mkLambda_or_LetIn paramsctxt
(my_it_mkLambda_or_LetIn_name realsign_ind
@@ -256,7 +260,9 @@ let build_sym_involutive_scheme env ind =
let eq,eqrefl,ctx = get_coq_eq ctx in
let sym, ctx, eff = const_of_scheme sym_scheme_kind env ind ctx in
let cstr n = mkApp (mkConstructUi (indu,1),Context.Rel.to_extended_vect mkRel n paramsctxt) in
- let varH = fresh env (default_id_of_sort (snd (mind_arity mip))) in
+ let inds = snd (mind_arity mip) in
+ let indr = Sorts.relevance_of_sort_family inds in
+ let varH = fresh env (default_id_of_sort inds) in
let applied_ind = build_dependent_inductive indu specif in
let applied_ind_C =
mkApp
@@ -264,8 +270,9 @@ let build_sym_involutive_scheme env ind =
(Context.Rel.to_extended_vect mkRel (nrealargs+1) mib.mind_params_ctxt)
(rel_vect (nrealargs+1) nrealargs)) in
let realsign_ind =
- name_context env ((LocalAssum (Name varH,applied_ind))::realsign) in
- let ci = make_case_info (Global.env()) ind RegularStyle in
+ name_context env ((LocalAssum (make_annot (Name varH) indr,applied_ind))::realsign) in
+ let rci = Sorts.Relevant in (* TODO relevance *)
+ let ci = make_case_info (Global.env()) ind rci RegularStyle in
let c =
(my_it_mkLambda_or_LetIn paramsctxt
(my_it_mkLambda_or_LetIn_name realsign_ind
@@ -368,7 +375,9 @@ let build_l2r_rew_scheme dep env ind kind =
mkApp (mkConstructUi(indu,1),
Array.concat [Context.Rel.to_extended_vect mkRel n paramsctxt1;
rel_vect p nrealargs]) in
- let varH = fresh env (default_id_of_sort (snd (mind_arity mip))) in
+ let inds = snd (mind_arity mip) in
+ let indr = Sorts.relevance_of_sort_family inds in
+ let varH = fresh env (default_id_of_sort inds) in
let varHC = fresh env (Id.of_string "HC") in
let varP = fresh env (Id.of_string "P") in
let applied_ind = build_dependent_inductive indu specif in
@@ -384,9 +393,9 @@ let build_l2r_rew_scheme dep env ind kind =
rel_vect 0 nrealargs]) in
let realsign_P = lift_rel_context nrealargs realsign in
let realsign_ind_P =
- name_context env ((LocalAssum (Name varH,applied_ind_P))::realsign_P) in
+ name_context env ((LocalAssum (make_annot (Name varH) indr,applied_ind_P))::realsign_P) in
let realsign_ind_G =
- name_context env ((LocalAssum (Name varH,applied_ind_G))::
+ name_context env ((LocalAssum (make_annot (Name varH) indr,applied_ind_G))::
lift_rel_context (nrealargs+3) realsign) in
let applied_sym_C n =
mkApp(sym,
@@ -400,8 +409,9 @@ let build_l2r_rew_scheme dep env ind kind =
let s, ctx' = UnivGen.fresh_sort_in_family kind in
let ctx = Univ.ContextSet.union ctx ctx' in
let s = mkSort s in
- let ci = make_case_info (Global.env()) ind RegularStyle in
- let cieq = make_case_info (Global.env()) (fst (destInd eq)) RegularStyle in
+ let rci = Sorts.Relevant in (* TODO relevance *)
+ let ci = make_case_info (Global.env()) ind rci RegularStyle in
+ let cieq = make_case_info (Global.env()) (fst (destInd eq)) rci RegularStyle in
let applied_PC =
mkApp (mkVar varP,Array.append (Context.Rel.to_extended_vect mkRel 1 realsign)
(if dep then [|cstr (2*nrealargs+1) 1|] else [||])) in
@@ -429,14 +439,14 @@ let build_l2r_rew_scheme dep env ind kind =
let c =
(my_it_mkLambda_or_LetIn paramsctxt
(my_it_mkLambda_or_LetIn_name realsign
- (mkNamedLambda varP
+ (mkNamedLambda (make_annot varP indr)
(my_it_mkProd_or_LetIn (if dep then realsign_ind_P else realsign_P) s)
- (mkNamedLambda varHC applied_PC
- (mkNamedLambda varH (lift 2 applied_ind)
+ (mkNamedLambda (make_annot varHC indr) applied_PC
+ (mkNamedLambda (make_annot varH indr) (lift 2 applied_ind)
(if dep then (* we need a coercion *)
mkCase (cieq,
- mkLambda (Name varH,lift 3 applied_ind,
- mkLambda (Anonymous,
+ mkLambda (make_annot (Name varH) indr,lift 3 applied_ind,
+ mkLambda (make_annot Anonymous indr,
mkApp (eq,[|lift 4 applied_ind;applied_sym_sym;mkRel 1|]),
applied_PR)),
mkApp (sym_involutive,
@@ -481,7 +491,9 @@ let build_l2r_forward_rew_scheme dep env ind kind =
mkApp (mkConstructUi(indu,1),
Array.concat [Context.Rel.to_extended_vect mkRel n paramsctxt1;
rel_vect p nrealargs]) in
- let varH = fresh env (default_id_of_sort (snd (mind_arity mip))) in
+ let inds = snd (mind_arity mip) in
+ let indr = Sorts.relevance_of_sort_family inds in
+ let varH = fresh env (default_id_of_sort inds) in
let varHC = fresh env (Id.of_string "HC") in
let varP = fresh env (Id.of_string "P") in
let applied_ind = build_dependent_inductive indu specif in
@@ -497,13 +509,14 @@ let build_l2r_forward_rew_scheme dep env ind kind =
rel_vect (2*nrealargs+1) nrealargs]) in
let realsign_P n = lift_rel_context (nrealargs*n+n) realsign in
let realsign_ind =
- name_context env ((LocalAssum (Name varH,applied_ind))::realsign) in
+ name_context env ((LocalAssum (make_annot (Name varH) indr,applied_ind))::realsign) in
let realsign_ind_P n aP =
- name_context env ((LocalAssum (Name varH,aP))::realsign_P n) in
+ name_context env ((LocalAssum (make_annot (Name varH) indr,aP))::realsign_P n) in
let s, ctx' = UnivGen.fresh_sort_in_family kind in
let ctx = Univ.ContextSet.union ctx ctx' in
let s = mkSort s in
- let ci = make_case_info (Global.env()) ind RegularStyle in
+ let rci = Sorts.Relevant in
+ let ci = make_case_info (Global.env()) ind rci RegularStyle in
let applied_PC =
mkApp (mkVar varP,Array.append
(rel_vect (nrealargs*2+3) nrealargs)
@@ -519,19 +532,19 @@ let build_l2r_forward_rew_scheme dep env ind kind =
let c =
(my_it_mkLambda_or_LetIn paramsctxt
(my_it_mkLambda_or_LetIn_name realsign
- (mkNamedLambda varH applied_ind
+ (mkNamedLambda (make_annot varH indr) applied_ind
(mkCase (ci,
my_it_mkLambda_or_LetIn_name
(lift_rel_context (nrealargs+1) realsign_ind)
- (mkNamedProd varP
+ (mkNamedProd (make_annot varP indr)
(my_it_mkProd_or_LetIn
(if dep then realsign_ind_P 2 applied_ind_P else realsign_P 2) s)
- (mkNamedProd varHC applied_PC applied_PG)),
+ (mkNamedProd (make_annot varHC indr) applied_PC applied_PG)),
(mkVar varH),
- [|mkNamedLambda varP
+ [|mkNamedLambda (make_annot varP indr)
(my_it_mkProd_or_LetIn
(if dep then realsign_ind_P 1 applied_ind_P' else realsign_P 2) s)
- (mkNamedLambda varHC applied_PC'
+ (mkNamedLambda (make_annot varHC indr) applied_PC'
(mkVar varHC))|])))))
in c, UState.of_context_set ctx
@@ -572,16 +585,19 @@ let build_r2l_forward_rew_scheme dep env ind kind =
let cstr n =
mkApp (mkConstructUi(indu,1),Context.Rel.to_extended_vect mkRel n mib.mind_params_ctxt) in
let constrargs_cstr = constrargs@[cstr 0] in
- let varH = fresh env (default_id_of_sort (snd (mind_arity mip))) in
+ let inds = snd (mind_arity mip) in
+ let indr = Sorts.relevance_of_sort_family inds in
+ let varH = fresh env (default_id_of_sort inds) in
let varHC = fresh env (Id.of_string "HC") in
let varP = fresh env (Id.of_string "P") in
let applied_ind = build_dependent_inductive indu specif in
let realsign_ind =
- name_context env ((LocalAssum (Name varH,applied_ind))::realsign) in
+ name_context env ((LocalAssum (make_annot (Name varH) indr,applied_ind))::realsign) in
let s, ctx' = UnivGen.fresh_sort_in_family kind in
let ctx = Univ.ContextSet.union ctx ctx' in
let s = mkSort s in
- let ci = make_case_info (Global.env()) ind RegularStyle in
+ let rci = Sorts.Relevant in (* TODO relevance *)
+ let ci = make_case_info (Global.env()) ind rci RegularStyle in
let applied_PC =
applist (mkVar varP,if dep then constrargs_cstr else constrargs) in
let applied_PG =
@@ -591,18 +607,18 @@ let build_r2l_forward_rew_scheme dep env ind kind =
let c =
(my_it_mkLambda_or_LetIn paramsctxt
(my_it_mkLambda_or_LetIn_name realsign_ind
- (mkNamedLambda varP
+ (mkNamedLambda (make_annot varP indr)
(my_it_mkProd_or_LetIn (lift_rel_context (nrealargs+1)
(if dep then realsign_ind else realsign)) s)
- (mkNamedLambda varHC (lift 1 applied_PG)
+ (mkNamedLambda (make_annot varHC indr) (lift 1 applied_PG)
(mkApp
(mkCase (ci,
my_it_mkLambda_or_LetIn_name
(lift_rel_context (nrealargs+3) realsign_ind)
- (mkArrow applied_PG (lift (2*nrealargs+5) applied_PC)),
+ (mkArrow applied_PG indr (lift (2*nrealargs+5) applied_PC)),
mkRel 3 (* varH *),
[|mkLambda
- (Name varHC,
+ (make_annot (Name varHC) indr,
lift (nrealargs+3) applied_PC,
mkRel 1)|]),
[|mkVar varHC|]))))))
@@ -775,7 +791,10 @@ let build_congr env (eq,refl,ctx) ind =
if List.exists is_local_def realsign then
error "Inductive equalities with local definitions in arity not supported.";
let env_with_arity = push_rel_context arityctxt env in
- let ty = RelDecl.get_type (lookup_rel (mip.mind_nrealargs - i + 1) env_with_arity) in
+ let ty, tyr =
+ let decl = lookup_rel (mip.mind_nrealargs - i + 1) env_with_arity in
+ RelDecl.get_type decl, RelDecl.get_relevance decl
+ in
let constrsign,ccl = mip.mind_nf_lc.(0) in
let _,constrargs = decompose_app ccl in
if not (Int.equal (Context.Rel.length constrsign) (Context.Rel.length mib.mind_params_ctxt)) then
@@ -784,15 +803,16 @@ let build_congr env (eq,refl,ctx) ind =
let varB = fresh env (Id.of_string "B") in
let varH = fresh env (Id.of_string "H") in
let varf = fresh env (Id.of_string "f") in
- let ci = make_case_info (Global.env()) ind RegularStyle in
+ let rci = Sorts.Relevant in (* TODO relevance *)
+ let ci = make_case_info (Global.env()) ind rci RegularStyle in
let uni, ctx = Univ.extend_in_context_set (UnivGen.new_global_univ ()) ctx in
let ctx = (fst ctx, Univ.enforce_leq uni (univ_of_eq env eq) (snd ctx)) in
let c =
my_it_mkLambda_or_LetIn paramsctxt
- (mkNamedLambda varB (mkType uni)
- (mkNamedLambda varf (mkArrow (lift 1 ty) (mkVar varB))
+ (mkNamedLambda (make_annot varB Sorts.Relevant) (mkType uni)
+ (mkNamedLambda (make_annot varf Sorts.Relevant) (mkArrow (lift 1 ty) tyr (mkVar varB))
(my_it_mkLambda_or_LetIn_name (lift_rel_context 2 realsign)
- (mkNamedLambda varH
+ (mkNamedLambda (make_annot varH Sorts.Relevant)
(applist
(mkIndU indu,
Context.Rel.to_extended_list mkRel (mip.mind_nrealargs+2) paramsctxt @
@@ -801,7 +821,7 @@ let build_congr env (eq,refl,ctx) ind =
my_it_mkLambda_or_LetIn_name
(lift_rel_context (mip.mind_nrealargs+3) realsign)
(mkLambda
- (Anonymous,
+ (make_annot Anonymous Sorts.Relevant,
applist
(mkIndU indu,
Context.Rel.to_extended_list mkRel (2*mip.mind_nrealdecls+3)
diff --git a/tactics/equality.ml b/tactics/equality.ml
index 4a1bb37fa4..88ce9868af 100644
--- a/tactics/equality.ml
+++ b/tactics/equality.ml
@@ -16,6 +16,7 @@ open Names
open Nameops
open Term
open Constr
+open Context
open Termops
open EConstr
open Vars
@@ -887,7 +888,8 @@ let descend_then env sigma head dirn =
let brl =
List.map build_branch
(List.interval 1 (Array.length mip.mind_consnames)) in
- let ci = make_case_info env ind RegularStyle in
+ let rci = Sorts.Relevant in (* TODO relevance *)
+ let ci = make_case_info env ind rci RegularStyle in
Inductiveops.make_case_or_project env sigma indf ci p head (Array.of_list brl)))
(* Now we need to construct the discriminator, given a discriminable
@@ -932,7 +934,8 @@ let build_selector env sigma dirn c ind special default =
it_mkLambda_or_LetIn endpt args in
let brl =
List.map build_branch(List.interval 1 (Array.length mip.mind_consnames)) in
- let ci = make_case_info env ind RegularStyle in
+ let rci = Sorts.Relevant in (* TODO relevance *)
+ let ci = make_case_info env ind rci RegularStyle in
let ans = Inductiveops.make_case_or_project env sigma indf ci p c (Array.of_list brl) in
ans
@@ -997,7 +1000,7 @@ let discrimination_pf e (t,t1,t2) discriminator lbeq =
Proofview.tclEFFECTS eff <*>
pf_constr_of_global eq_elim >>= fun eq_elim ->
Proofview.tclUNIT
- (applist (eq_elim, [t;t1;mkNamedLambda e t discriminator;i;t2]), absurd_term)
+ (applist (eq_elim, [t;t1;mkNamedLambda (make_annot e Sorts.Relevant) t discriminator;i;t2]), absurd_term)
let eq_baseid = Id.of_string "e"
@@ -1015,7 +1018,7 @@ let discr_positions env sigma (lbeq,eqn,(t,t1,t2)) eq_clause cpath dirn =
build_coq_True () >>= fun true_0 ->
build_coq_False () >>= fun false_0 ->
let e = next_ident_away eq_baseid (vars_of_env env) in
- let e_env = push_named (Context.Named.Declaration.LocalAssum (e,t)) env in
+ let e_env = push_named (Context.Named.Declaration.LocalAssum (make_annot e Sorts.Relevant,t)) env in
let discriminator =
try
Proofview.tclUNIT
@@ -1025,7 +1028,7 @@ let discr_positions env sigma (lbeq,eqn,(t,t1,t2)) eq_clause cpath dirn =
in
discriminator >>= fun discriminator ->
discrimination_pf e (t,t1,t2) discriminator lbeq >>= fun (pf, absurd_term) ->
- let pf_ty = mkArrow eqn absurd_term in
+ let pf_ty = mkArrow eqn Sorts.Relevant absurd_term in
let absurd_clause = apply_on_clause (pf,pf_ty) eq_clause in
let pf = Clenvtac.clenv_value_cast_meta absurd_clause in
tclTHENS (assert_after Anonymous absurd_term)
@@ -1114,7 +1117,7 @@ let make_tuple env sigma (rterm,rty) lind =
assert (not (noccurn sigma lind rty));
let sigdata = find_sigma_data env (get_sort_of env sigma rty) in
let sigma, a = type_of ~refresh:true env sigma (mkRel lind) in
- let na = Context.Rel.Declaration.get_name (lookup_rel lind env) in
+ let na = Context.Rel.Declaration.get_annot (lookup_rel lind env) in
(* We move [lind] to [1] and lift other rels > [lind] by 1 *)
let rty = lift (1-lind) (liftn lind (lind+1) rty) in
(* Now [lind] is [mkRel 1] and we abstract on (na:a) *)
@@ -1374,13 +1377,13 @@ let simplify_args env sigma t =
let inject_at_positions env sigma l2r (eq,_,(t,t1,t2)) eq_clause posns tac =
let e = next_ident_away eq_baseid (vars_of_env env) in
- let e_env = push_named (LocalAssum (e,t)) env in
+ let e_env = push_named (LocalAssum (make_annot e Sorts.Relevant,t)) env in
let evdref = ref sigma in
let filter (cpath, t1', t2') =
try
(* arbitrarily take t1' as the injector default value *)
let sigma, (injbody,resty) = build_injector e_env !evdref t1' (mkVar e) cpath in
- let injfun = mkNamedLambda e t injbody in
+ let injfun = mkNamedLambda (make_annot e Sorts.Relevant) t injbody in
let sigma,congr = Evd.fresh_global env sigma eq.congr in
let pf = applist(congr,[t;resty;injfun;t1;t2]) in
let sigma, pf_typ = Typing.type_of env sigma pf in
@@ -1565,9 +1568,9 @@ let subst_tuple_term env sigma dep_pair1 dep_pair2 b =
(* We build the expected goal *)
let abst_B =
List.fold_right
- (fun (e,t) body -> lambda_create env sigma (t,subst_term sigma e body)) e1_list b in
+ (fun (e,t) body -> lambda_create env sigma (Sorts.Relevant,t,subst_term sigma e body)) e1_list b in
let pred_body = beta_applist sigma (abst_B,proj_list) in
- let body = mkApp (lambda_create env sigma (typ,pred_body),[|dep_pair1|]) in
+ let body = mkApp (lambda_create env sigma (Sorts.Relevant,typ,pred_body),[|dep_pair1|]) in
let expected_goal = beta_applist sigma (abst_B,List.map fst e2_list) in
(* Simulate now the normalisation treatment made by Logic.mk_refgoals *)
let expected_goal = nf_betaiota env sigma expected_goal in
diff --git a/tactics/hints.ml b/tactics/hints.ml
index c1f6365f5d..a04a9f9db9 100644
--- a/tactics/hints.ml
+++ b/tactics/hints.ml
@@ -13,6 +13,7 @@ open Util
open CErrors
open Names
open Constr
+open Context
open Evd
open EConstr
open Vars
@@ -1275,7 +1276,7 @@ let prepare_hint check (poly,local) env init (sigma,c) =
let id = next_ident_away_from default_prepare_hint_ident (fun id -> Id.Set.mem id !vars) in
vars := Id.Set.add id !vars;
subst := (evar,mkVar id)::!subst;
- mkNamedLambda id t (iter (replace_term sigma evar (mkVar id) c)) in
+ mkNamedLambda (make_annot id Sorts.Relevant) t (iter (replace_term sigma evar (mkVar id) c)) in
let c' = iter c in
let env = Global.env () in
let empty_sigma = Evd.from_env env in
@@ -1305,7 +1306,7 @@ let project_hint ~poly pri l2r r =
let sigma, p = Evd.fresh_global env sigma p in
let c = Reductionops.whd_beta sigma (mkApp (c, Context.Rel.to_extended_vect mkRel 0 sign)) in
let c = it_mkLambda_or_LetIn
- (mkApp (p,[|mkArrow a (lift 1 b);mkArrow b (lift 1 a);c|])) sign in
+ (mkApp (p,[|mkArrow a Sorts.Relevant (lift 1 b);mkArrow b Sorts.Relevant (lift 1 a);c|])) sign in
let id =
Nameops.add_suffix (Nametab.basename_of_global gr) ("_proj_" ^ (if l2r then "l2r" else "r2l"))
in
diff --git a/tactics/hipattern.ml b/tactics/hipattern.ml
index 395b4928ce..08131f6309 100644
--- a/tactics/hipattern.ml
+++ b/tactics/hipattern.ml
@@ -182,10 +182,10 @@ let match_with_disjunction ?(strict=false) ?(onlybinary=false) sigma t =
let car = constructors_nrealargs ind in
let (mib,mip) = Global.lookup_inductive ind in
if Array.for_all (fun ar -> Int.equal ar 1) car
- && not (mis_is_recursive (ind,mib,mip))
- && (Int.equal mip.mind_nrealargs 0)
+ && not (mis_is_recursive (ind,mib,mip))
+ && (Int.equal mip.mind_nrealargs 0)
then
- if strict then
+ if strict then
if test_strict_disjunction (mib, mip) then
Some (hdapp,args)
else
@@ -196,7 +196,7 @@ let match_with_disjunction ?(strict=false) ?(onlybinary=false) sigma t =
pi2 (destProd sigma (prod_applist sigma ar args))
in
let cargs = Array.map map mip.mind_nf_lc in
- Some (hdapp,Array.to_list cargs)
+ Some (hdapp,Array.to_list cargs)
else
None
| _ -> None in
diff --git a/tactics/hipattern.mli b/tactics/hipattern.mli
index f04cda1232..741f6713e3 100644
--- a/tactics/hipattern.mli
+++ b/tactics/hipattern.mli
@@ -86,7 +86,7 @@ val is_equality_type : testing_function
val match_with_nottype : Environ.env -> (constr * constr) matching_function
val is_nottype : Environ.env -> testing_function
-val match_with_forall_term : (Name.t * constr * constr) matching_function
+val match_with_forall_term : (Name.t Context.binder_annot * constr * constr) matching_function
val is_forall_term : testing_function
val match_with_imp_term : (constr * constr) matching_function
diff --git a/tactics/inv.ml b/tactics/inv.ml
index 2ae37ab471..776148d4cf 100644
--- a/tactics/inv.ml
+++ b/tactics/inv.ml
@@ -15,6 +15,7 @@ open Names
open Term
open Termops
open Constr
+open Context
open EConstr
open Vars
open Namegen
@@ -131,7 +132,7 @@ let make_inv_predicate env evd indf realargs id status concl =
let eq_term = eqdata.Coqlib.eq in
let eq = evd_comb1 (Evd.fresh_global env) evd eq_term in
let eqn = applist (eq,[eqnty;lhs;rhs]) in
- let eqns = (Anonymous, lift n eqn) :: eqns in
+ let eqns = (make_annot Anonymous Sorts.Relevant, lift n eqn) :: eqns in
let refl_term = eqdata.Coqlib.refl in
let refl_term = evd_comb1 (Evd.fresh_global env) evd refl_term in
let refl = mkApp (refl_term, [|eqnty; rhs|]) in
diff --git a/tactics/leminv.ml b/tactics/leminv.ml
index 335f3c74ff..4aa4d13e1e 100644
--- a/tactics/leminv.ml
+++ b/tactics/leminv.ml
@@ -15,6 +15,7 @@ open Names
open Termops
open Environ
open Constr
+open Context
open EConstr
open Vars
open Namegen
@@ -120,13 +121,13 @@ let max_prefix_sign lid sign =
let rec add_prods_sign env sigma t =
match EConstr.kind sigma (whd_all env sigma t) with
| Prod (na,c1,b) ->
- let id = id_of_name_using_hdchar env sigma t na in
+ let id = id_of_name_using_hdchar env sigma t na.binder_name in
let b'= subst1 (mkVar id) b in
- add_prods_sign (push_named (LocalAssum (id,c1)) env) sigma b'
+ add_prods_sign (push_named (LocalAssum ({na with binder_name=id},c1)) env) sigma b'
| LetIn (na,c1,t1,b) ->
- let id = id_of_name_using_hdchar env sigma t na in
+ let id = id_of_name_using_hdchar env sigma t na.binder_name in
let b'= subst1 (mkVar id) b in
- add_prods_sign (push_named (LocalDef (id,c1,t1)) env) sigma b'
+ add_prods_sign (push_named (LocalDef ({na with binder_name=id},c1,t1)) env) sigma b'
| _ -> (env,t)
(* [dep_option] indicates whether the inversion lemma is dependent or not.
@@ -149,9 +150,10 @@ let compute_first_inversion_scheme env sigma ind sort dep_option =
let pty,goal =
if dep_option then
let pty = make_arity env sigma true indf sort in
+ let r = relevance_of_inductive_type env ind in
let goal =
mkProd
- (Anonymous, mkAppliedInd ind, applist(mkVar p,realargs@[mkRel 1]))
+ (make_annot Anonymous r, mkAppliedInd ind, applist(mkVar p,realargs@[mkRel 1]))
in
pty,goal
else
@@ -169,11 +171,11 @@ let compute_first_inversion_scheme env sigma ind sort dep_option =
env ~init:([],[])
in
let pty = it_mkNamedProd_or_LetIn (mkSort sort) ownsign in
- let goal = mkArrow i (applist(mkVar p, List.rev revargs)) in
+ let goal = mkArrow i Sorts.Relevant (applist(mkVar p, List.rev revargs)) in
(pty,goal)
in
let npty = nf_all env sigma pty in
- let extenv = push_named (LocalAssum (p,npty)) env in
+ let extenv = push_named (LocalAssum (make_annot p Sorts.Relevant,npty)) env in
extenv, goal
(* [inversion_scheme sign I]
@@ -225,7 +227,7 @@ let inversion_scheme ~name ~poly env sigma t sort dep_option inv_op =
let h = next_ident_away (Id.of_string "H") !avoid in
let ty,inst = Evarutil.generalize_evar_over_rels sigma (e,args) in
avoid := Id.Set.add h !avoid;
- ownSign := Context.Named.add (LocalAssum (h,ty)) !ownSign;
+ ownSign := Context.Named.add (LocalAssum (make_annot h Sorts.Relevant,ty)) !ownSign;
applist (mkVar h, inst)
| _ -> EConstr.map sigma fill_holes c
in
diff --git a/tactics/tactics.ml b/tactics/tactics.ml
index ee9d117d4d..b8308dc49b 100644
--- a/tactics/tactics.ml
+++ b/tactics/tactics.ml
@@ -14,6 +14,7 @@ open Util
open Names
open Nameops
open Constr
+open Context
open Termops
open Environ
open EConstr
@@ -137,8 +138,8 @@ let introduction id =
in
let open Context.Named.Declaration in
match EConstr.kind sigma concl with
- | Prod (_, t, b) -> unsafe_intro env (LocalAssum (id, t)) b
- | LetIn (_, c, t, b) -> unsafe_intro env (LocalDef (id, c, t)) b
+ | Prod (id0, t, b) -> unsafe_intro env (LocalAssum ({id0 with binder_name=id}, t)) b
+ | LetIn (id0, c, t, b) -> unsafe_intro env (LocalDef ({id0 with binder_name=id}, c, t)) b
| _ -> raise (RefinerError (env, sigma, IntroNeedsProduct))
end
@@ -366,8 +367,8 @@ let default_id env sigma decl =
match decl with
| LocalAssum (name,t) ->
let dft = default_id_of_sort (Retyping.get_sort_of env sigma t) in
- id_of_name_with_default dft name
- | LocalDef (name,b,_) -> id_of_name_using_hdchar env sigma b name
+ id_of_name_with_default dft name.binder_name
+ | LocalDef (name,b,_) -> id_of_name_using_hdchar env sigma b name.binder_name
(* Non primitive introduction tactics are treated by intro_then_gen
There is possibly renaming, with possibly names to avoid and
@@ -437,16 +438,17 @@ let internal_cut_gen ?(check=true) dir replace id t =
let sigma = Tacmach.New.project gl in
let concl = Proofview.Goal.concl gl in
let sign = named_context_val env in
+ let r = Retyping.relevance_of_type env sigma t in
let sign',t,concl,sigma =
if replace then
let nexthyp = get_next_hyp_position env sigma id (named_context_of_val sign) in
let sigma,sign',t,concl = clear_hyps2 env sigma (Id.Set.singleton id) sign t concl in
- let sign' = insert_decl_in_named_context env sigma (LocalAssum (id,t)) nexthyp sign' in
+ let sign' = insert_decl_in_named_context env sigma (LocalAssum (make_annot id r,t)) nexthyp sign' in
sign',t,concl,sigma
else
(if check && mem_named_context_val id sign then
user_err (str "Variable " ++ Id.print id ++ str " is already declared.");
- push_named_context_val (LocalAssum (id,t)) sign,t,concl,sigma) in
+ push_named_context_val (LocalAssum (make_annot id r,t)) sign,t,concl,sigma) in
let nf_t = nf_betaiota env sigma t in
Proofview.tclTHEN
(Proofview.Unsafe.tclEVARS sigma)
@@ -460,7 +462,7 @@ let internal_cut_gen ?(check=true) dir replace id t =
let (sigma, ev') = Evarutil.new_evar_from_context sign' sigma ~principal:true concl in
let (sigma, ev) = Evarutil.new_evar_from_context sign sigma nf_t in
(sigma,ev,ev') in
- let term = mkLetIn (Name id, ev, t, EConstr.Vars.subst_var id ev') in
+ let term = mkLetIn (make_annot (Name id) r, ev, t, EConstr.Vars.subst_var id ev') in
(sigma, term)
end)
end
@@ -471,7 +473,7 @@ let internal_cut_rev ?(check=true) = internal_cut_gen ~check false
let assert_before_then_gen b naming t tac =
let open Context.Rel.Declaration in
Proofview.Goal.enter begin fun gl ->
- let id = find_name b (LocalAssum (Anonymous,t)) naming gl in
+ let id = find_name b (LocalAssum (make_annot Anonymous Sorts.Relevant,t)) naming gl in
Tacticals.New.tclTHENLAST
(internal_cut b id t)
(tac id)
@@ -486,7 +488,7 @@ let assert_before_replacing id = assert_before_gen true (NamingMustBe (CAst.make
let assert_after_then_gen b naming t tac =
let open Context.Rel.Declaration in
Proofview.Goal.enter begin fun gl ->
- let id = find_name b (LocalAssum (Anonymous,t)) naming gl in
+ let id = find_name b (LocalAssum (make_annot Anonymous Sorts.Relevant,t)) naming gl in
Tacticals.New.tclTHENFIRST
(internal_cut_rev b id t)
(tac id)
@@ -542,7 +544,7 @@ let mutual_fix f n rest j = Proofview.Goal.enter begin fun gl ->
if mem_named_context_val f sign then
user_err ~hdr:"Logic.prim_refiner"
(str "Name " ++ Id.print f ++ str " already used in the environment");
- mk_sign (push_named_context_val (LocalAssum (f, ar)) sign) oth
+ mk_sign (push_named_context_val (LocalAssum (make_annot f Sorts.Relevant, ar)) sign) oth
in
let nenv = reset_with_named_context (mk_sign (named_context_val env) all) env in
Refine.refine ~typecheck:false begin fun sigma ->
@@ -550,7 +552,8 @@ let mutual_fix f n rest j = Proofview.Goal.enter begin fun gl ->
let ids = List.map pi1 all in
let evs = List.map (Vars.subst_vars (List.rev ids)) evs in
let indxs = Array.of_list (List.map (fun n -> n-1) (List.map pi2 all)) in
- let funnames = Array.of_list (List.map (fun i -> Name i) ids) in
+ (* TODO relevance *)
+ let funnames = Array.of_list (List.map (fun i -> make_annot (Name i) Sorts.Relevant) ids) in
let typarray = Array.of_list (List.map pi3 all) in
let bodies = Array.of_list evs in
let oterm = mkFix ((indxs,0),(funnames,typarray,bodies)) in
@@ -586,14 +589,15 @@ let mutual_cofix f others j = Proofview.Goal.enter begin fun gl ->
let open Context.Named.Declaration in
if mem_named_context_val f sign then
error "Name already used in the environment.";
- mk_sign (push_named_context_val (LocalAssum (f, ar)) sign) oth
+ mk_sign (push_named_context_val (LocalAssum (make_annot f Sorts.Relevant, ar)) sign) oth
in
let nenv = reset_with_named_context (mk_sign (named_context_val env) all) env in
Refine.refine ~typecheck:false begin fun sigma ->
let (ids, types) = List.split all in
let (sigma, evs) = mk_holes nenv sigma types in
let evs = List.map (Vars.subst_vars (List.rev ids)) evs in
- let funnames = Array.of_list (List.map (fun i -> Name i) ids) in
+ (* TODO relevance *)
+ let funnames = Array.of_list (List.map (fun i -> make_annot (Name i) Sorts.Relevant) ids) in
let typarray = Array.of_list types in
let bodies = Array.of_list evs in
let oterm = mkCoFix (0, (funnames, typarray, bodies)) in
@@ -616,7 +620,7 @@ let pf_reduce_decl redfun where decl gl =
match decl with
| LocalAssum (id,ty) ->
if where == InHypValueOnly then
- user_err (Id.print id ++ str " has no value.");
+ user_err (Id.print id.binder_name ++ str " has no value.");
LocalAssum (id,redfun' ty)
| LocalDef (id,b,ty) ->
let b' = if where != InHypTypeOnly then redfun' b else b in
@@ -717,7 +721,7 @@ let pf_e_reduce_decl redfun where decl gl =
match decl with
| LocalAssum (id,ty) ->
if where == InHypValueOnly then
- user_err (Id.print id ++ str " has no value.");
+ user_err (Id.print id.binder_name ++ str " has no value.");
let (sigma, ty') = redfun sigma ty in
(sigma, LocalAssum (id, ty'))
| LocalDef (id,b,ty) ->
@@ -760,7 +764,7 @@ let e_pf_change_decl (redfun : bool -> e_reduction_function) where decl env sigm
match decl with
| LocalAssum (id,ty) ->
if where == InHypValueOnly then
- user_err (Id.print id ++ str " has no value.");
+ user_err (Id.print id.binder_name ++ str " has no value.");
let (sigma, ty') = redfun false env sigma ty in
(sigma, LocalAssum (id, ty'))
| LocalDef (id,b,ty) ->
@@ -1238,27 +1242,29 @@ let cut c =
let env = Proofview.Goal.env gl in
let sigma = Tacmach.New.project gl in
let concl = Proofview.Goal.concl gl in
- let is_sort =
+ let relevance =
try
- (* Backward compat: ensure that [c] is well-typed. *)
+ (* Backward compat: ensure that [c] is well-typed. Plus we
+ need to know the relevance *)
let typ = Typing.unsafe_type_of env sigma c in
let typ = whd_all env sigma typ in
match EConstr.kind sigma typ with
- | Sort _ -> true
- | _ -> false
- with e when Pretype_errors.precatchable_exception e -> false
+ | Sort s -> Some (Sorts.relevance_of_sort (ESorts.kind sigma s))
+ | _ -> None
+ with e when Pretype_errors.precatchable_exception e -> None
in
- if is_sort then
+ match relevance with
+ | Some r ->
let id = next_name_away_with_default "H" Anonymous (Tacmach.New.pf_ids_set_of_hyps gl) in
(* Backward compat: normalize [c]. *)
let c = if normalize_cut then local_strong whd_betaiota sigma c else c in
Refine.refine ~typecheck:false begin fun h ->
- let (h, f) = Evarutil.new_evar ~principal:true env h (mkArrow c (Vars.lift 1 concl)) in
+ let (h, f) = Evarutil.new_evar ~principal:true env h (mkArrow c r (Vars.lift 1 concl)) in
let (h, x) = Evarutil.new_evar env h c in
- let f = mkLetIn (Name id, x, c, mkApp (Vars.lift 1 f, [|mkRel 1|])) in
+ let f = mkLetIn (make_annot (Name id) r, x, c, mkApp (Vars.lift 1 f, [|mkRel 1|])) in
(h, f)
end
- else
+ | None ->
Tacticals.New.tclZEROMSG (str "Not a proposition or a type.")
end
@@ -1823,7 +1829,7 @@ let apply_in_once ?(respect_opaque = false) sidecond_first with_delta
let sigma = Tacmach.New.project gl in
let t' = Tacmach.New.pf_get_hyp_typ id gl in
let innerclause = mk_clenv_from_env env sigma (Some 0) (mkVar id,t') in
- let targetid = find_name true (LocalAssum (Anonymous,t')) naming gl in
+ let targetid = find_name true (LocalAssum (make_annot Anonymous Sorts.Relevant,t')) naming gl in
let rec aux idstoclear with_destruct c =
Proofview.Goal.enter begin fun gl ->
let env = Proofview.Goal.env gl in
@@ -1890,7 +1896,7 @@ let cut_and_apply c =
let concl = Proofview.Goal.concl gl in
let env = Tacmach.New.pf_env gl in
Refine.refine ~typecheck:false begin fun sigma ->
- let typ = mkProd (Anonymous, c2, concl) in
+ let typ = mkProd (make_annot Anonymous Sorts.Relevant, c2, concl) in
let (sigma, f) = Evarutil.new_evar env sigma typ in
let (sigma, x) = Evarutil.new_evar env sigma c1 in
(sigma, mkApp (f, [|mkApp (c, [|x|])|]))
@@ -2013,12 +2019,12 @@ let clear_body ids =
let ctx = named_context env in
let map = function
| LocalAssum (id,t) as decl ->
- let () = if List.mem_f Id.equal id ids then
- user_err (str "Hypothesis " ++ Id.print id ++ str " is not a local definition")
+ let () = if List.mem_f Id.equal id.binder_name ids then
+ user_err (str "Hypothesis " ++ Id.print id.binder_name ++ str " is not a local definition")
in
decl
| LocalDef (id,_,t) as decl ->
- if List.mem_f Id.equal id ids then LocalAssum (id, t) else decl
+ if List.mem_f Id.equal id.binder_name ids then LocalAssum (id, t) else decl
in
let ctx = List.map map ctx in
let base_env = reset_context env in
@@ -2624,7 +2630,8 @@ let letin_tac_gen with_eq (id,depdecls,lastlhyp,ccl,c) ty =
let (sigma, refl) = Evd.fresh_global env sigma eqdata.refl in
let eq = applist (eq,args) in
let refl = applist (refl, [t;mkVar id]) in
- let term = mkNamedLetIn id c t (mkLetIn (Name heq, refl, eq, ccl)) in
+ let term = mkNamedLetIn (make_annot id Sorts.Relevant) c t
+ (mkLetIn (make_annot (Name heq) Sorts.Relevant, refl, eq, ccl)) in
let sigma, _ = Typing.type_of env sigma term in
let ans = term,
Tacticals.New.tclTHENLIST
@@ -2634,7 +2641,7 @@ let letin_tac_gen with_eq (id,depdecls,lastlhyp,ccl,c) ty =
in
(sigma, ans)
| None ->
- (sigma, (mkNamedLetIn id c t ccl, Proofview.tclUNIT ()))
+ (sigma, (mkNamedLetIn (make_annot id Sorts.Relevant) c t ccl, Proofview.tclUNIT ()))
in
Tacticals.New.tclTHENLIST
[ Proofview.Unsafe.tclEVARS sigma;
@@ -2669,8 +2676,9 @@ let mk_eq_name env id {CAst.loc;v=ido} =
let mkletin_goal env sigma with_eq dep (id,lastlhyp,ccl,c) ty =
let open Context.Named.Declaration in
let t = match ty with Some t -> t | _ -> typ_of env sigma c in
- let decl = if dep then LocalDef (id,c,t)
- else LocalAssum (id,t)
+ let r = Retyping.relevance_of_type env sigma t in
+ let decl = if dep then LocalDef (make_annot id r,c,t)
+ else LocalAssum (make_annot id r,t)
in
match with_eq with
| Some (lr,heq) ->
@@ -2680,13 +2688,14 @@ let mkletin_goal env sigma with_eq dep (id,lastlhyp,ccl,c) ty =
let (sigma, refl) = Evd.fresh_global env sigma eqdata.refl in
let eq = applist (eq,args) in
let refl = applist (refl, [t;mkVar id]) in
- let newenv = insert_before [LocalAssum (heq,eq); decl] lastlhyp env in
+ let newenv = insert_before [LocalAssum (make_annot heq Sorts.Relevant,eq); decl] lastlhyp env in
let (sigma, x) = new_evar newenv sigma ~principal:true ccl in
- (sigma, mkNamedLetIn id c t (mkNamedLetIn heq refl eq x))
+ (sigma, mkNamedLetIn (make_annot id r) c t
+ (mkNamedLetIn (make_annot heq Sorts.Relevant) refl eq x))
| None ->
let newenv = insert_before [decl] lastlhyp env in
let (sigma, x) = new_evar newenv sigma ~principal:true ccl in
- (sigma, mkNamedLetIn id c t x)
+ (sigma, mkNamedLetIn (make_annot id r) c t x)
let pose_tac na c =
Proofview.Goal.enter begin fun gl ->
@@ -2708,11 +2717,13 @@ let pose_tac na c =
in
Proofview.Unsafe.tclEVARS sigma <*>
Refine.refine ~typecheck:false begin fun sigma ->
+ (* TODO relevance *)
+ let id = make_annot id Sorts.Relevant in
let nhyps = EConstr.push_named_context_val (NamedDecl.LocalDef (id, c, t)) hyps in
let (sigma, ev) = Evarutil.new_pure_evar nhyps sigma concl in
let inst = Array.map_of_list (fun d -> mkVar (get_id d)) (named_context env) in
let body = mkEvar (ev, Array.append [|mkRel 1|] inst) in
- (sigma, mkLetIn (Name id, c, t, body))
+ (sigma, mkLetIn (map_annot Name.mk_name id, c, t, body))
end
end
@@ -2806,9 +2817,10 @@ let generalize_goal_gen env sigma ids i ((occs,c,b),na) t cl =
let newdecls,_ = decompose_prod_n_assum sigma i (subst_term_gen sigma EConstr.eq_constr_nounivs c dummy_prod) in
let cl',sigma' = subst_closed_term_occ env sigma (AtOccs occs) c (it_mkProd_or_LetIn cl newdecls) in
let na = generalized_name env sigma c t ids cl' na in
+ let r = Retyping.relevance_of_type env sigma t in
let decl = match b with
- | None -> LocalAssum (na,t)
- | Some b -> LocalDef (na,b,t)
+ | None -> LocalAssum (make_annot na r,t)
+ | Some b -> LocalDef (make_annot na r,b,t)
in
mkProd_or_LetIn decl cl', sigma'
@@ -2948,8 +2960,8 @@ let specialize (c,lbind) ipat =
(* If the term is lambda then we put a letin to put avoid
interaction between the term and the bindings. *)
let c = match EConstr.kind sigma c with
- | Lambda(_,_,_) ->
- mkLetIn(Name.Anonymous, c, typ_of_c, (mkRel 1))
+ | Lambda _ ->
+ mkLetIn(make_annot Name.Anonymous Sorts.Relevant, c, typ_of_c, (mkRel 1))
| _ -> c in
let clause = make_clenv_binding env sigma (c,typ_of_c) lbind in
let flags = { (default_unify_flags ()) with resolve_evars = true } in
@@ -2973,14 +2985,15 @@ let specialize (c,lbind) ipat =
(* nme has not been resolved, let us re-abstract it. Same
name but type updated by instanciation of other args. *)
let sigma,new_typ_of_t = Typing.type_of clause.env sigma t in
+ let r = Retyping.relevance_of_type env sigma new_typ_of_t in
let liftedargs = List.map liftrel args in
(* lifting rels in the accumulator args *)
let sigma,hd' = rebuild_lambdas sigma lp' (liftedargs@[mkRel 1 ]) hd l' in
(* replace meta variable by the abstracted variable *)
let hd'' = subst_term sigma t hd' in
(* lambda expansion *)
- sigma,mkLambda (nme,new_typ_of_t,hd'')
- | Context.Rel.Declaration.LocalAssum(_,_)::lp' , t::l' ->
+ sigma,mkLambda ({nme with binder_relevance=r},new_typ_of_t,hd'')
+ | Context.Rel.Declaration.LocalAssum _::lp' , t::l' ->
let sigma,hd' = rebuild_lambdas sigma lp' (args@[t]) hd l' in
sigma,hd'
| _ ,_ -> assert false in
@@ -3631,15 +3644,18 @@ let make_abstract_generalize env id typ concl dep ctx body c eqs args refls =
let homogeneous = Reductionops.is_conv env sigma ty typ in
let sigma, (eq, refl) =
mk_term_eq homogeneous (push_rel_context ctx env) sigma ty (mkRel 1) typ (mkVar id) in
- sigma, mkProd (Anonymous, eq, lift 1 concl), [| refl |]
+ sigma, mkProd (make_annot Anonymous Sorts.Relevant, eq, lift 1 concl), [| refl |]
else sigma, concl, [||]
in
(* Abstract by equalities *)
let eqs = lift_togethern 1 eqs in (* lift together and past genarg *)
- let abseqs = it_mkProd_or_LetIn (lift eqslen abshypeq) (List.map (fun x -> LocalAssum (Anonymous, x)) eqs) in
+ let abseqs = it_mkProd_or_LetIn (lift eqslen abshypeq)
+ (List.map (fun x -> LocalAssum (make_annot Anonymous Sorts.Relevant, x)) eqs)
+ in
+ let r = Sorts.Relevant in (* TODO relevance *)
let decl = match body with
- | None -> LocalAssum (Name id, c)
- | Some body -> LocalDef (Name id, body, c)
+ | None -> LocalAssum (make_annot (Name id) r, c)
+ | Some body -> LocalDef (make_annot (Name id) r, body, c)
in
(* Abstract by the "generalized" hypothesis. *)
let genarg = mkProd_or_LetIn decl abseqs in
@@ -3714,10 +3730,10 @@ let abstract_args gl generalize_vars dep id defined f args =
eqs are not lifted w.r.t. each other yet. (* will be needed when going to dependent indexes *)
*)
let aux (prod, ctx, ctxenv, c, args, eqs, refls, nongenvars, vars, env) arg =
- let name, ty, arity =
+ let name, ty_relevance, ty, arity =
let rel, c = Reductionops.splay_prod_n env !sigma 1 prod in
let decl = List.hd rel in
- RelDecl.get_name decl, RelDecl.get_type decl, c
+ RelDecl.get_name decl, RelDecl.get_relevance decl, RelDecl.get_type decl, c
in
let argty = Tacmach.New.pf_unsafe_type_of gl arg in
let sigma', ty = Evarsolve.refresh_universes (Some true) env !sigma ty in
@@ -3731,7 +3747,7 @@ let abstract_args gl generalize_vars dep id defined f args =
Id.Set.add id nongenvars, Id.Set.remove id vars, env)
| _ ->
let name = get_id name in
- let decl = LocalAssum (Name name, ty) in
+ let decl = LocalAssum (make_annot (Name name) ty_relevance, ty) in
let ctx = decl :: ctx in
let c' = mkApp (lift 1 c, [|mkRel 1|]) in
let args = arg :: args in
@@ -3869,7 +3885,7 @@ let specialize_eqs id =
else
let sigma, e = Evarutil.new_evar (push_rel_context ctx env) !evars t in
evars := sigma;
- aux false (LocalDef (na,e,t) :: ctx) (mkApp (lift 1 acc, [| mkRel 1 |])) b)
+ aux false (LocalDef (na,e,t) :: ctx) (mkApp (lift 1 acc, [| mkRel 1 |])) b)
| t -> acc, in_eqs, ctx, ty
in
let acc, worked, ctx, ty = aux false [] (mkVar id) ty in
@@ -3917,7 +3933,7 @@ let decompose_paramspred_branch_args sigma elimt =
| Prod(nme,tpe,elimt') ->
let hd_tpe,_ = decompose_app sigma (snd (decompose_prod_assum sigma tpe)) in
if not (occur_rel sigma 1 elimt') && isRel sigma hd_tpe
- then cut_noccur elimt' (LocalAssum (nme,tpe)::acc2)
+ then cut_noccur elimt' (LocalAssum (nme,tpe)::acc2)
else let acc3,ccl = decompose_prod_assum sigma elimt in acc2 , acc3 , ccl
| App(_, _) | Rel _ -> acc2 , [] , elimt
| _ -> error_ind_scheme "" in
@@ -3999,8 +4015,8 @@ let compute_elim_sig sigma ?elimc elimt =
(* 3- Look at last arg: is it the indarg? *)
ignore (
match List.hd args_indargs with
- | LocalDef (hiname,_,hi) -> error_ind_scheme ""
- | LocalAssum (hiname,hi) ->
+ | LocalDef (hiname,_,hi) -> error_ind_scheme ""
+ | LocalAssum (hiname,hi) ->
let hi_ind, hi_args = decompose_app sigma hi in
let hi_is_ind = (* hi est d'un type globalisable *)
match EConstr.kind sigma hi_ind with
diff --git a/tactics/term_dnet.ml b/tactics/term_dnet.ml
index e8a66f1889..2831aec9f6 100644
--- a/tactics/term_dnet.ml
+++ b/tactics/term_dnet.ml
@@ -316,7 +316,7 @@ struct
Term(DCoFix(i,Array.map pat_of_constr ta,Array.map pat_of_constr ca))
| Cast (c,_,_) -> pat_of_constr c
| Lambda (_,t,c) -> Term(DLambda (pat_of_constr t, pat_of_constr c))
- | (Prod (_,_,_) | LetIn(_,_,_,_)) ->
+ | (Prod _ | LetIn _) ->
let (ctx,c) = ctx_of_constr (Term DNil) c in Term (DCtx (ctx,c))
| App (f,ca) ->
Array.fold_left (fun c a -> Term (DApp (c,a)))
diff --git a/test-suite/misc/poly-capture-global-univs/src/evilImpl.ml b/test-suite/misc/poly-capture-global-univs/src/evilImpl.ml
index f5043db099..adabb7a0a0 100644
--- a/test-suite/misc/poly-capture-global-univs/src/evilImpl.ml
+++ b/test-suite/misc/poly-capture-global-univs/src/evilImpl.ml
@@ -16,7 +16,7 @@ let evil t f =
let fe = Declare.definition_entry
~univs:(Polymorphic_entry ([|Anonymous|], UContext.make (Instance.of_array [|u|],Constraint.empty)))
- ~types:(Term.mkArrow tc tu)
- (mkLambda (Name.Name (Id.of_string "x"), tc, mkRel 1))
+ ~types:(Term.mkArrowR tc tu)
+ (mkLambda (Context.nameR (Id.of_string "x"), tc, mkRel 1))
in
ignore (Declare.declare_constant f (DefinitionEntry fe, k))
diff --git a/vernac/assumptions.ml b/vernac/assumptions.ml
index b5cc74b594..3773bc9c88 100644
--- a/vernac/assumptions.ml
+++ b/vernac/assumptions.ml
@@ -21,6 +21,7 @@ open CErrors
open Util
open Names
open Constr
+open Context
open Declarations
open Mod_subst
open Globnames
@@ -238,8 +239,9 @@ and traverse_inductive (curr, data, ax2ty) mind obj =
Array.fold_left (fun accu oib ->
let pspecif = Univ.in_punivs (mib, oib) in
let ind_type = Inductive.type_of_inductive global_env pspecif in
+ let indr = oib.mind_relevant in
let ind_name = Name oib.mind_typename in
- Context.Rel.add (Context.Rel.Declaration.LocalAssum (ind_name, ind_type)) accu)
+ Context.Rel.add (Context.Rel.Declaration.LocalAssum (make_annot ind_name indr, ind_type)) accu)
Context.Rel.empty mib.mind_packets
in
(* For each inductive, collects references in their arity and in the type
diff --git a/vernac/auto_ind_decl.ml b/vernac/auto_ind_decl.ml
index 868a6ed3e9..528829f3a5 100644
--- a/vernac/auto_ind_decl.ml
+++ b/vernac/auto_ind_decl.ml
@@ -16,6 +16,7 @@ open Util
open Pp
open Term
open Constr
+open Context
open Vars
open Termops
open Declarations
@@ -144,7 +145,7 @@ let build_beq_scheme mode kn =
in
(* construct the "fun A B ... N, eqA eqB eqC ... N => fixpoint" part *)
let create_input c =
- let myArrow u v = mkArrow u (lift 1 v)
+ let myArrow u v = mkArrow u Sorts.Relevant (lift 1 v)
and eqName = function
| Name s -> Id.of_string ("eq_"^(Id.to_string s))
| Anonymous -> Id.of_string "eq_A"
@@ -161,14 +162,16 @@ let build_beq_scheme mode kn =
( fun a b decl -> (* mkLambda(n,b,a) ) *)
(* here I leave the Naming thingy so that the type of
the function is more readable for the user *)
- mkNamedLambda (eqName (RelDecl.get_name decl)) b a )
+ mkNamedLambda (map_annot eqName (RelDecl.get_annot decl)) b a )
c (List.rev eqs_typ) lnamesparrec
in
List.fold_left (fun a decl ->(* mkLambda(n,t,a)) eq_input rel_list *)
- (* Same here , hoping the auto renaming will do something good ;) *)
- mkNamedLambda
- (match RelDecl.get_name decl with Name s -> s | Anonymous -> Id.of_string "A")
- (RelDecl.get_type decl) a) eq_input lnamesparrec
+ (* Same here , hoping the auto renaming will do something good ;) *)
+ let x = map_annot
+ (function Name s -> s | Anonymous -> Id.of_string "A")
+ (RelDecl.get_annot decl)
+ in
+ mkNamedLambda x (RelDecl.get_type decl) a) eq_input lnamesparrec
in
let make_one_eq cur =
let u = Univ.Instance.empty in
@@ -251,8 +254,8 @@ let build_beq_scheme mode kn =
in
(* construct the predicate for the Case part*)
let do_predicate rel_list n =
- List.fold_left (fun a b -> mkLambda(Anonymous,b,a))
- (mkLambda (Anonymous,
+ List.fold_left (fun a b -> mkLambda(make_annot Anonymous Sorts.Relevant,b,a))
+ (mkLambda (make_annot Anonymous Sorts.Relevant,
mkFullInd ind (n+3+(List.length rettyp_l)+nb_ind-1),
(bb ())))
(List.rev rettyp_l) in
@@ -260,7 +263,8 @@ let build_beq_scheme mode kn =
(* do the [| C1 ... => match Y with ... end
...
Cn => match Y with ... end |] part *)
- let ci = make_case_info env (fst ind) MatchStyle in
+ let rci = Sorts.Relevant in (* TODO relevance *)
+ let ci = make_case_info env (fst ind) rci MatchStyle in
let constrs n = get_constructors env (make_ind_family (ind,
Context.Rel.to_extended_list mkRel (n+nb_ind-1) mib.mind_params_ctxt)) in
let constrsi = constrs (3+nparrec) in
@@ -296,32 +300,32 @@ let build_beq_scheme mode kn =
(Array.sub eqs 1 (nb_cstr_args - 1))
)
in
- (List.fold_left (fun a decl -> mkLambda (RelDecl.get_name decl, RelDecl.get_type decl, a)) cc
+ (List.fold_left (fun a decl -> mkLambda (RelDecl.get_annot decl, RelDecl.get_type decl, a)) cc
(constrsj.(j).cs_args)
)
else ar2.(j) <- (List.fold_left (fun a decl ->
- mkLambda (RelDecl.get_name decl, RelDecl.get_type decl, a)) (ff ()) (constrsj.(j).cs_args) )
- done;
+ mkLambda (RelDecl.get_annot decl, RelDecl.get_type decl, a)) (ff ()) (constrsj.(j).cs_args) )
+ done;
- ar.(i) <- (List.fold_left (fun a decl -> mkLambda (RelDecl.get_name decl, RelDecl.get_type decl, a))
+ ar.(i) <- (List.fold_left (fun a decl -> mkLambda (RelDecl.get_annot decl, RelDecl.get_type decl, a))
(mkCase (ci,do_predicate rel_list nb_cstr_args,
mkVar (Id.of_string "Y") ,ar2))
(constrsi.(i).cs_args))
- done;
- mkNamedLambda (Id.of_string "X") (mkFullInd ind (nb_ind-1+1)) (
- mkNamedLambda (Id.of_string "Y") (mkFullInd ind (nb_ind-1+2)) (
+ done;
+ mkNamedLambda (make_annot (Id.of_string "X") Sorts.Relevant) (mkFullInd ind (nb_ind-1+1)) (
+ mkNamedLambda (make_annot (Id.of_string "Y") Sorts.Relevant) (mkFullInd ind (nb_ind-1+2)) (
mkCase (ci, do_predicate rel_list 0,mkVar (Id.of_string "X"),ar))),
!eff
in (* build_beq_scheme *)
- let names = Array.make nb_ind Anonymous and
+ let names = Array.make nb_ind (make_annot Anonymous Sorts.Relevant) and
types = Array.make nb_ind mkSet and
cores = Array.make nb_ind mkSet in
let eff = ref Safe_typing.empty_private_constants in
let u = Univ.Instance.empty in
for i=0 to (nb_ind-1) do
- names.(i) <- Name (Id.of_string (rec_name i));
- types.(i) <- mkArrow (mkFullInd ((kn,i),u) 0)
- (mkArrow (mkFullInd ((kn,i),u) 1) (bb ()));
+ names.(i) <- make_annot (Name (Id.of_string (rec_name i))) Sorts.Relevant;
+ types.(i) <- mkArrow (mkFullInd ((kn,i),u) 0) Sorts.Relevant
+ (mkArrow (mkFullInd ((kn,i),u) 1) Sorts.Relevant (bb ()));
let c, eff' = make_one_eq i in
cores.(i) <- c;
eff := Safe_typing.concat_private eff' !eff
@@ -562,34 +566,39 @@ let compute_bl_goal ind lnamesparrec nparrec =
let x = next_ident_away (Id.of_string "x") avoid and
y = next_ident_away (Id.of_string "y") avoid in
let bl_typ = List.map (fun (s,seq,_,_) ->
- mkNamedProd x (mkVar s) (
- mkNamedProd y (mkVar s) (
+ mkNamedProd (make_annot x Sorts.Relevant) (mkVar s) (
+ mkNamedProd (make_annot y Sorts.Relevant) (mkVar s) (
mkArrow
( mkApp(eq (),[|bb (); mkApp(mkVar seq,[|mkVar x;mkVar y|]);tt () |]))
+ Sorts.Relevant
( mkApp(eq (),[|mkVar s;mkVar x;mkVar y|]))
))
) list_id in
let bl_input = List.fold_left2 ( fun a (s,_,sbl,_) b ->
- mkNamedProd sbl b a
+ mkNamedProd (make_annot sbl Sorts.Relevant) b a
) c (List.rev list_id) (List.rev bl_typ) in
let eqs_typ = List.map (fun (s,_,_,_) ->
- mkProd(Anonymous,mkVar s,mkProd(Anonymous,mkVar s,(bb ())))
+ mkProd(make_annot Anonymous Sorts.Relevant,mkVar s,mkProd(make_annot Anonymous Sorts.Relevant,mkVar s,(bb ())))
) list_id in
let eq_input = List.fold_left2 ( fun a (s,seq,_,_) b ->
- mkNamedProd seq b a
+ mkNamedProd (make_annot seq Sorts.Relevant) b a
) bl_input (List.rev list_id) (List.rev eqs_typ) in
- List.fold_left (fun a decl -> mkNamedProd
- (match RelDecl.get_name decl with Name s -> s | Anonymous -> next_ident_away (Id.of_string "A") avoid)
- (RelDecl.get_type decl) a) eq_input lnamesparrec
+ List.fold_left (fun a decl ->
+ let x = map_annot
+ (function Name s -> s | Anonymous -> next_ident_away (Id.of_string "A") avoid)
+ (RelDecl.get_annot decl)
+ in
+ mkNamedProd x (RelDecl.get_type decl) a) eq_input lnamesparrec
in
let n = next_ident_away (Id.of_string "x") avoid and
m = next_ident_away (Id.of_string "y") avoid in
let u = Univ.Instance.empty in
create_input (
- mkNamedProd n (mkFullInd (ind,u) nparrec) (
- mkNamedProd m (mkFullInd (ind,u) (nparrec+1)) (
+ mkNamedProd (make_annot n Sorts.Relevant) (mkFullInd (ind,u) nparrec) (
+ mkNamedProd (make_annot m Sorts.Relevant) (mkFullInd (ind,u) (nparrec+1)) (
mkArrow
(mkApp(eq (),[|bb ();mkApp(eqI,[|mkVar n;mkVar m|]);tt ()|]))
+ Sorts.Relevant
(mkApp(eq (),[|mkFullInd (ind,u) (nparrec+3);mkVar n;mkVar m|]))
))), eff
@@ -706,34 +715,40 @@ let compute_lb_goal ind lnamesparrec nparrec =
let x = next_ident_away (Id.of_string "x") avoid and
y = next_ident_away (Id.of_string "y") avoid in
let lb_typ = List.map (fun (s,seq,_,_) ->
- mkNamedProd x (mkVar s) (
- mkNamedProd y (mkVar s) (
+ mkNamedProd (make_annot x Sorts.Relevant) (mkVar s) (
+ mkNamedProd (make_annot y Sorts.Relevant) (mkVar s) (
mkArrow
- ( mkApp(eq,[|mkVar s;mkVar x;mkVar y|]))
- ( mkApp(eq,[|bb;mkApp(mkVar seq,[|mkVar x;mkVar y|]);tt|]))
+ ( mkApp(eq,[|mkVar s;mkVar x;mkVar y|]))
+ Sorts.Relevant
+ ( mkApp(eq,[|bb;mkApp(mkVar seq,[|mkVar x;mkVar y|]);tt|]))
))
) list_id in
let lb_input = List.fold_left2 ( fun a (s,_,_,slb) b ->
- mkNamedProd slb b a
+ mkNamedProd (make_annot slb Sorts.Relevant) b a
) c (List.rev list_id) (List.rev lb_typ) in
let eqs_typ = List.map (fun (s,_,_,_) ->
- mkProd(Anonymous,mkVar s,mkProd(Anonymous,mkVar s,bb))
+ mkProd(make_annot Anonymous Sorts.Relevant,mkVar s,
+ mkProd(make_annot Anonymous Sorts.Relevant,mkVar s,bb))
) list_id in
let eq_input = List.fold_left2 ( fun a (s,seq,_,_) b ->
- mkNamedProd seq b a
+ mkNamedProd (make_annot seq Sorts.Relevant) b a
) lb_input (List.rev list_id) (List.rev eqs_typ) in
- List.fold_left (fun a decl -> mkNamedProd
- (match (RelDecl.get_name decl) with Name s -> s | Anonymous -> Id.of_string "A")
- (RelDecl.get_type decl) a) eq_input lnamesparrec
+ List.fold_left (fun a decl ->
+ let x = map_annot
+ (function Name s -> s | Anonymous -> Id.of_string "A")
+ (RelDecl.get_annot decl)
+ in
+ mkNamedProd x (RelDecl.get_type decl) a) eq_input lnamesparrec
in
let n = next_ident_away (Id.of_string "x") avoid and
m = next_ident_away (Id.of_string "y") avoid in
let u = Univ.Instance.empty in
create_input (
- mkNamedProd n (mkFullInd (ind,u) nparrec) (
- mkNamedProd m (mkFullInd (ind,u) (nparrec+1)) (
+ mkNamedProd (make_annot n Sorts.Relevant) (mkFullInd (ind,u) nparrec) (
+ mkNamedProd (make_annot m Sorts.Relevant) (mkFullInd (ind,u) (nparrec+1)) (
mkArrow
(mkApp(eq,[|mkFullInd (ind,u) (nparrec+2);mkVar n;mkVar m|]))
+ Sorts.Relevant
(mkApp(eq,[|bb;mkApp(eqI,[|mkVar n;mkVar m|]);tt|]))
))), eff
@@ -835,45 +850,51 @@ let compute_dec_goal ind lnamesparrec nparrec =
let x = next_ident_away (Id.of_string "x") avoid and
y = next_ident_away (Id.of_string "y") avoid in
let lb_typ = List.map (fun (s,seq,_,_) ->
- mkNamedProd x (mkVar s) (
- mkNamedProd y (mkVar s) (
+ mkNamedProd (make_annot x Sorts.Relevant) (mkVar s) (
+ mkNamedProd (make_annot y Sorts.Relevant) (mkVar s) (
mkArrow
- ( mkApp(eq,[|mkVar s;mkVar x;mkVar y|]))
- ( mkApp(eq,[|bb;mkApp(mkVar seq,[|mkVar x;mkVar y|]);tt|]))
+ ( mkApp(eq,[|mkVar s;mkVar x;mkVar y|]))
+ Sorts.Relevant
+ ( mkApp(eq,[|bb;mkApp(mkVar seq,[|mkVar x;mkVar y|]);tt|]))
))
) list_id in
let bl_typ = List.map (fun (s,seq,_,_) ->
- mkNamedProd x (mkVar s) (
- mkNamedProd y (mkVar s) (
+ mkNamedProd (make_annot x Sorts.Relevant) (mkVar s) (
+ mkNamedProd (make_annot y Sorts.Relevant) (mkVar s) (
mkArrow
- ( mkApp(eq,[|bb;mkApp(mkVar seq,[|mkVar x;mkVar y|]);tt|]))
- ( mkApp(eq,[|mkVar s;mkVar x;mkVar y|]))
+ ( mkApp(eq,[|bb;mkApp(mkVar seq,[|mkVar x;mkVar y|]);tt|]))
+ Sorts.Relevant
+ ( mkApp(eq,[|mkVar s;mkVar x;mkVar y|]))
))
) list_id in
let lb_input = List.fold_left2 ( fun a (s,_,_,slb) b ->
- mkNamedProd slb b a
+ mkNamedProd (make_annot slb Sorts.Relevant) b a
) c (List.rev list_id) (List.rev lb_typ) in
let bl_input = List.fold_left2 ( fun a (s,_,sbl,_) b ->
- mkNamedProd sbl b a
+ mkNamedProd (make_annot sbl Sorts.Relevant) b a
) lb_input (List.rev list_id) (List.rev bl_typ) in
let eqs_typ = List.map (fun (s,_,_,_) ->
- mkProd(Anonymous,mkVar s,mkProd(Anonymous,mkVar s,bb))
+ mkProd(make_annot Anonymous Sorts.Relevant,mkVar s,
+ mkProd(make_annot Anonymous Sorts.Relevant,mkVar s,bb))
) list_id in
let eq_input = List.fold_left2 ( fun a (s,seq,_,_) b ->
- mkNamedProd seq b a
+ mkNamedProd (make_annot seq Sorts.Relevant) b a
) bl_input (List.rev list_id) (List.rev eqs_typ) in
- List.fold_left (fun a decl -> mkNamedProd
- (match RelDecl.get_name decl with Name s -> s | Anonymous -> Id.of_string "A")
- (RelDecl.get_type decl) a) eq_input lnamesparrec
+ List.fold_left (fun a decl ->
+ let x = map_annot
+ (function Name s -> s | Anonymous -> Id.of_string "A")
+ (RelDecl.get_annot decl)
+ in
+ mkNamedProd x (RelDecl.get_type decl) a) eq_input lnamesparrec
in
let n = next_ident_away (Id.of_string "x") avoid and
m = next_ident_away (Id.of_string "y") avoid in
let eqnm = mkApp(eq,[|mkFullInd ind (2*nparrec+2);mkVar n;mkVar m|]) in
create_input (
- mkNamedProd n (mkFullInd ind (2*nparrec)) (
- mkNamedProd m (mkFullInd ind (2*nparrec+1)) (
+ mkNamedProd (make_annot n Sorts.Relevant) (mkFullInd ind (2*nparrec)) (
+ mkNamedProd (make_annot m Sorts.Relevant) (mkFullInd ind (2*nparrec+1)) (
mkApp(sumbool(),[|eqnm;mkApp (UnivGen.constr_of_monomorphic_global @@ Coqlib.lib_ref "core.not.type",[|eqnm|])|])
)
)
diff --git a/vernac/class.ml b/vernac/class.ml
index a6b3242cae..0837beccee 100644
--- a/vernac/class.ml
+++ b/vernac/class.ml
@@ -14,6 +14,7 @@ open Pp
open Names
open Term
open Constr
+open Context
open Vars
open Termops
open Entries
@@ -188,14 +189,14 @@ let build_id_coercion idf_opt source poly =
let lams,t = decompose_lam_assum c in
let val_f =
it_mkLambda_or_LetIn
- (mkLambda (Name Namegen.default_dependent_ident,
+ (mkLambda (make_annot (Name Namegen.default_dependent_ident) Sorts.Relevant,
applistc vs (Context.Rel.to_extended_list mkRel 0 lams),
mkRel 1))
lams
in
let typ_f =
List.fold_left (fun d c -> Term.mkProd_wo_LetIn c d)
- (mkProd (Anonymous, applistc vs (Context.Rel.to_extended_list mkRel 0 lams), lift 1 t))
+ (mkProd (make_annot Anonymous Sorts.Relevant, applistc vs (Context.Rel.to_extended_list mkRel 0 lams), lift 1 t))
lams
in
(* juste pour verification *)
diff --git a/vernac/classes.ml b/vernac/classes.ml
index 4664df3182..1981e24ae4 100644
--- a/vernac/classes.ml
+++ b/vernac/classes.ml
@@ -352,8 +352,8 @@ let named_of_rel_context l =
(fun decl (subst, ctx) ->
let id = match RelDecl.get_name decl with Anonymous -> invalid_arg "named_of_rel_context" | Name id -> id in
let d = match decl with
- | LocalAssum (_,t) -> id, None, substl subst t
- | LocalDef (_,b,t) -> id, Some (substl subst b), substl subst t in
+ | LocalAssum (_,t) -> id, None, substl subst t
+ | LocalDef (_,b,t) -> id, Some (substl subst b), substl subst t in
(mkVar id :: subst, d :: ctx))
l ([], [])
in ctx
diff --git a/vernac/comAssumption.ml b/vernac/comAssumption.ml
index 35d8be5c56..37a33daf8f 100644
--- a/vernac/comAssumption.ml
+++ b/vernac/comAssumption.ml
@@ -13,6 +13,7 @@ open Util
open Vars
open Declare
open Names
+open Context
open Globnames
open Constrexpr_ops
open Constrintern
@@ -148,8 +149,9 @@ let do_assumptions ~program_mode kind nl l =
(* We intepret all declarations in the same evar_map, i.e. as a telescope. *)
let (sigma,_,_),l = List.fold_left_map (fun (sigma,env,ienv) (is_coe,(idl,c)) ->
let sigma,(t,imps) = interp_assumption ~program_mode sigma env ienv c in
+ let r = Retyping.relevance_of_type env sigma t in
let env =
- EConstr.push_named_context (List.map (fun {CAst.v=id} -> LocalAssum (id,t)) idl) env in
+ EConstr.push_named_context (List.map (fun {CAst.v=id} -> LocalAssum (make_annot id r,t)) idl) env in
let ienv = List.fold_right (fun {CAst.v=id} ienv ->
let impls = compute_internalization_data env sigma Variable t imps in
Id.Map.add id impls ienv) idl ienv in
diff --git a/vernac/comFixpoint.ml b/vernac/comFixpoint.ml
index 5229d9e8e8..2f00b41b7c 100644
--- a/vernac/comFixpoint.ml
+++ b/vernac/comFixpoint.ml
@@ -12,6 +12,7 @@ open Pp
open CErrors
open Util
open Constr
+open Context
open Vars
open Termops
open Declare
@@ -126,7 +127,9 @@ let interp_fix_context ~program_mode ~cofix env sigma fix =
sigma, ((env'', ctx' @ ctx), (impl_env',imps @ imps'), annot)
let interp_fix_ccl ~program_mode sigma impls (env,_) fix =
- interp_type_evars_impls ~program_mode ~impls env sigma fix.fix_type
+ let sigma, (c, impl) = interp_type_evars_impls ~program_mode ~impls env sigma fix.fix_type in
+ let r = Retyping.relevance_of_type env sigma c in
+ sigma, (c, r, impl)
let interp_fix_body ~program_mode env_rec sigma impls (_,ctx) fix ccl =
let open EConstr in
@@ -137,9 +140,9 @@ let interp_fix_body ~program_mode env_rec sigma impls (_,ctx) fix ccl =
let build_fix_type (_,ctx) ccl = EConstr.it_mkProd_or_LetIn ccl ctx
-let prepare_recursive_declaration fixnames fixtypes fixdefs =
+let prepare_recursive_declaration fixnames fixrs fixtypes fixdefs =
let defs = List.map (subst_vars (List.rev fixnames)) fixdefs in
- let names = List.map (fun id -> Name id) fixnames in
+ let names = List.map2 (fun id r -> make_annot (Name id) r) fixnames fixrs in
(Array.of_list names, Array.of_list fixtypes, Array.of_list defs)
(* Jump over let-bindings. *)
@@ -158,7 +161,7 @@ let compute_possible_guardness_evidences (ctx,_,recindex) =
List.interval 0 (Context.Rel.nhyps ctx - 1)
type recursive_preentry =
- Id.t list * constr option list * types list
+ Id.t list * Sorts.relevance list * constr option list * types list
(* Wellfounded definition *)
@@ -188,8 +191,8 @@ let interp_recursive ~program_mode ~cofix fixl notations =
on_snd List.split3 @@
List.fold_left_map (fun sigma -> interp_fix_context ~program_mode env sigma ~cofix) sigma fixl in
let fixctximpenvs, fixctximps = List.split fiximppairs in
- let sigma, (fixccls,fixcclimps) =
- on_snd List.split @@
+ let sigma, (fixccls,fixrs,fixcclimps) =
+ on_snd List.split3 @@
List.fold_left3_map (interp_fix_ccl ~program_mode) sigma fixctximpenvs fixctxs fixl in
let fixtypes = List.map2 build_fix_type fixctxs fixccls in
let fixtypes = List.map (fun c -> nf_evar sigma c) fixtypes in
@@ -208,8 +211,8 @@ let interp_recursive ~program_mode ~cofix fixl notations =
Typing.solve_evars env sigma app
with e when CErrors.noncritical e -> sigma, t
in
- sigma, LocalAssum (id,fixprot) :: env'
- else sigma, LocalAssum (id,t) :: env')
+ sigma, LocalAssum (make_annot id Sorts.Relevant,fixprot) :: env'
+ else sigma, LocalAssum (make_annot id Sorts.Relevant,t) :: env')
(sigma,[]) fixnames fixtypes
in
let env_rec = push_named_context rec_sign env in
@@ -232,19 +235,19 @@ let interp_recursive ~program_mode ~cofix fixl notations =
let fixctxs = List.map (fun (_,ctx) -> ctx) fixctxs in
(* Build the fix declaration block *)
- (env,rec_sign,decl,sigma), (fixnames,fixdefs,fixtypes), List.combine3 fixctxs fiximps fixannots
+ (env,rec_sign,decl,sigma), (fixnames,fixrs,fixdefs,fixtypes), List.combine3 fixctxs fiximps fixannots
-let check_recursive isfix env evd (fixnames,fixdefs,_) =
+let check_recursive isfix env evd (fixnames,_,fixdefs,_) =
if List.for_all Option.has_some fixdefs then begin
let fixdefs = List.map Option.get fixdefs in
check_mutuality env evd isfix (List.combine fixnames fixdefs)
end
-let ground_fixpoint env evd (fixnames,fixdefs,fixtypes) =
+let ground_fixpoint env evd (fixnames,fixrs,fixdefs,fixtypes) =
check_evars_are_solved ~program_mode:false env evd;
let fixdefs = List.map (fun c -> Option.map EConstr.(to_constr evd) c) fixdefs in
let fixtypes = List.map EConstr.(to_constr evd) fixtypes in
- Evd.evar_universe_context evd, (fixnames,fixdefs,fixtypes)
+ Evd.evar_universe_context evd, (fixnames,fixrs,fixdefs,fixtypes)
let interp_fixpoint ~cofix l ntns =
let (env,_,pl,evd),fix,info = interp_recursive ~program_mode:false ~cofix l ntns in
@@ -252,7 +255,7 @@ let interp_fixpoint ~cofix l ntns =
let uctx,fix = ground_fixpoint env evd fix in
(fix,pl,uctx,info)
-let declare_fixpoint local poly ((fixnames,fixdefs,fixtypes),pl,ctx,fiximps) indexes ntns =
+let declare_fixpoint local poly ((fixnames,fixrs,fixdefs,fixtypes),pl,ctx,fiximps) indexes ntns =
if List.exists Option.is_empty fixdefs then
(* Some bodies to define by proof *)
let thms =
@@ -267,7 +270,7 @@ let declare_fixpoint local poly ((fixnames,fixdefs,fixtypes),pl,ctx,fiximps) ind
else begin
(* We shortcut the proof process *)
let fixdefs = List.map Option.get fixdefs in
- let fixdecls = prepare_recursive_declaration fixnames fixtypes fixdefs in
+ let fixdecls = prepare_recursive_declaration fixnames fixrs fixtypes fixdefs in
let env = Global.env() in
let indexes = search_guard env indexes fixdecls in
let fiximps = List.map (fun (n,r,p) -> r) fiximps in
@@ -287,7 +290,7 @@ let declare_fixpoint local poly ((fixnames,fixdefs,fixtypes),pl,ctx,fiximps) ind
(* Declare notations *)
List.iter (Metasyntax.add_notation_interpretation (Global.env())) ntns
-let declare_cofixpoint local poly ((fixnames,fixdefs,fixtypes),pl,ctx,fiximps) ntns =
+let declare_cofixpoint local poly ((fixnames,fixrs,fixdefs,fixtypes),pl,ctx,fiximps) ntns =
if List.exists Option.is_empty fixdefs then
(* Some bodies to define by proof *)
let thms =
@@ -302,7 +305,7 @@ let declare_cofixpoint local poly ((fixnames,fixdefs,fixtypes),pl,ctx,fiximps) n
else begin
(* We shortcut the proof process *)
let fixdefs = List.map Option.get fixdefs in
- let fixdecls = prepare_recursive_declaration fixnames fixtypes fixdefs in
+ let fixdecls = prepare_recursive_declaration fixnames fixrs fixtypes fixdefs in
let fixdecls = List.map_i (fun i _ -> mkCoFix (i,fixdecls)) 0 fixnames in
let vars = Vars.universes_of_constr (List.hd fixdecls) in
let fixdecls = List.map Safe_typing.mk_pure_proof fixdecls in
diff --git a/vernac/comFixpoint.mli b/vernac/comFixpoint.mli
index 338dfa5ef5..9bcb53697b 100644
--- a/vernac/comFixpoint.mli
+++ b/vernac/comFixpoint.mli
@@ -51,7 +51,7 @@ val interp_recursive :
(* env / signature / univs / evar_map *)
(Environ.env * EConstr.named_context * UState.universe_decl * Evd.evar_map) *
(* names / defs / types *)
- (Id.t list * EConstr.constr option list * EConstr.types list) *
+ (Id.t list * Sorts.relevance list * EConstr.constr option list * EConstr.types list) *
(* ctx per mutual def / implicits / struct annotations *)
(EConstr.rel_context * Impargs.manual_explicitation list * int option) list
@@ -69,7 +69,7 @@ val extract_cofixpoint_components :
structured_fixpoint_expr list * decl_notation list
type recursive_preentry =
- Id.t list * constr option list * types list
+ Id.t list * Sorts.relevance list * constr option list * types list
val interp_fixpoint :
cofix:bool ->
diff --git a/vernac/comInductive.ml b/vernac/comInductive.ml
index e71946e8b8..8b8307c14a 100644
--- a/vernac/comInductive.ml
+++ b/vernac/comInductive.ml
@@ -13,6 +13,7 @@ open CErrors
open Sorts
open Util
open Constr
+open Context
open Environ
open Declare
open Names
@@ -70,9 +71,9 @@ let rec complete_conclusion a cs = CAst.map_with_loc (fun ?loc -> function
| c -> c
)
-let push_types env idl tl =
- List.fold_left2 (fun env id t -> EConstr.push_rel (LocalAssum (Name id,t)) env)
- env idl tl
+let push_types env idl rl tl =
+ List.fold_left3 (fun env id r t -> EConstr.push_rel (LocalAssum (make_annot (Name id) r,t)) env)
+ env idl rl tl
type structured_one_inductive_expr = {
ind_name : Id.t;
@@ -152,7 +153,7 @@ let interp_ind_arity env sigma ind =
user_err ?loc:(constr_loc ind.ind_arity) (str "Not an arity")
| s ->
let concl = if pseudo_poly then Some s else None in
- sigma, (t, concl, impls)
+ sigma, (t, Retyping.relevance_of_sort s, concl, impls)
let interp_cstrs env sigma impls mldata arity ind =
let cnames,ctyps = List.split ind.ind_lc in
@@ -183,7 +184,7 @@ let sup_list min = List.fold_left Univ.sup min
let extract_level env evd min tys =
let sorts = List.map (fun ty ->
let ctx, concl = Reduction.dest_prod_assum env ty in
- sign_level env evd (LocalAssum (Anonymous, concl) :: ctx)) tys
+ sign_level env evd (LocalAssum (make_annot Anonymous Sorts.Relevant, concl) :: ctx)) tys
in sup_list min sorts
let is_flexible_sort evd u =
@@ -370,15 +371,15 @@ let interp_mutual_inductive_gen env0 ~template udecl (uparamsl,paramsl,indl) not
(* Interpret the arities *)
let sigma, arities = List.fold_left_map (fun sigma -> interp_ind_arity env_params sigma) sigma indl in
+ let arities, relevances, arityconcl, indimpls = List.split4 arities in
- let fullarities = List.map (fun (c, _, _) -> EConstr.it_mkProd_or_LetIn c ctx_params) arities in
- let env_ar = push_types env_uparams indnames fullarities in
+ let fullarities = List.map (fun c -> EConstr.it_mkProd_or_LetIn c ctx_params) arities in
+ let env_ar = push_types env_uparams indnames relevances fullarities in
let env_ar_params = EConstr.push_rel_context ctx_params env_ar in
(* Compute interpretation metadatas *)
- let indimpls = List.map (fun (_, _, impls) -> userimpls @
- lift_implicits (Context.Rel.nhyps ctx_params) impls) arities in
- let arities = List.map pi1 arities and arityconcl = List.map pi2 arities in
+ let indimpls = List.map (fun impls -> userimpls @
+ lift_implicits (Context.Rel.nhyps ctx_params) impls) indimpls in
let impls = compute_internalization_env env_uparams sigma ~impls (Inductive (params,true)) indnames fullarities indimpls in
let ntn_impls = compute_internalization_env env_uparams sigma (Inductive (params,true)) indnames fullarities indimpls in
let mldatas = List.map2 (mk_mltype_data sigma env_params params) arities indnames in
@@ -407,7 +408,7 @@ let interp_mutual_inductive_gen env0 ~template udecl (uparamsl,paramsl,indl) not
let userimpls = useruimpls @ (lift_implicits (Context.Rel.nhyps ctx_uparams) userimpls) in
let indimpls = List.map (fun iimpl -> useruimpls @ (lift_implicits (Context.Rel.nhyps ctx_uparams) iimpl)) indimpls in
let fullarities = List.map (fun c -> EConstr.it_mkProd_or_LetIn c ctx_uparams) fullarities in
- let env_ar = push_types env0 indnames fullarities in
+ let env_ar = push_types env0 indnames relevances fullarities in
let env_ar_params = EConstr.push_rel_context ctx_params env_ar in
(* Try further to solve evars, and instantiate them *)
diff --git a/vernac/comProgramFixpoint.ml b/vernac/comProgramFixpoint.ml
index ae77cf12e5..ad7c65b70c 100644
--- a/vernac/comProgramFixpoint.ml
+++ b/vernac/comProgramFixpoint.ml
@@ -12,6 +12,7 @@ open Pp
open CErrors
open Util
open Constr
+open Context
open Entries
open Vars
open Declare
@@ -41,7 +42,7 @@ let well_founded sigma = init_constant sigma (lib_ref "core.wf.well_founded")
let mkSubset sigma name typ prop =
let open EConstr in
let sigma, app_h = Evarutil.new_global sigma (delayed_force build_sigma).typ in
- sigma, mkApp (app_h, [| typ; mkLambda (name, typ, prop) |])
+ sigma, mkApp (app_h, [| typ; mkLambda (make_annot name Sorts.Relevant, typ, prop) |])
let make_qref s = qualid_of_string s
let lt_ref = make_qref "Init.Peano.lt"
@@ -58,7 +59,7 @@ let rec telescope sigma l =
List.fold_left
(fun (sigma, ty, tys, (k, constr)) decl ->
let t = RelDecl.get_type decl in
- let pred = mkLambda (RelDecl.get_name decl, t, ty) in
+ let pred = mkLambda (RelDecl.get_annot decl, t, ty) in
let sigma, ty = Evarutil.new_global sigma (lib_ref "core.sigT.type") in
let sigma, intro = Evarutil.new_global sigma (lib_ref "core.sigT.intro") in
let sigty = mkApp (ty, [|t; pred|]) in
@@ -73,7 +74,7 @@ let rec telescope sigma l =
let sigma, p2 = Evarutil.new_global sigma (lib_ref "core.sigT.proj2") in
let proj1 = applist (p1, [t; pred; prev]) in
let proj2 = applist (p2, [t; pred; prev]) in
- (sigma, lift 1 proj2, LocalDef (get_name decl, proj1, t) :: subst))
+ (sigma, lift 1 proj2, LocalDef (get_annot decl, proj1, t) :: subst))
(List.rev tys) tl (sigma, mkRel 1, [])
in sigma, ty, (LocalDef (n, last, t) :: subst), constr
@@ -98,7 +99,7 @@ let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation =
let full_arity = it_mkProd_or_LetIn top_arity binders_rel in
let sigma, argtyp, letbinders, make = telescope sigma binders_rel in
let argname = Id.of_string "recarg" in
- let arg = LocalAssum (Name argname, argtyp) in
+ let arg = LocalAssum (make_annot (Name argname) Sorts.Relevant, argtyp) in
let binders = letbinders @ [arg] in
let binders_env = push_rel_context binders_rel env in
let sigma, (rel, _) = interp_constr_evars_impls ~program_mode:true env sigma r in
@@ -135,7 +136,7 @@ let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation =
let argid' = Id.of_string (Id.to_string argname ^ "'") in
let wfarg sigma len =
let sigma, ss_term = mkSubset sigma (Name argid') argtyp (wf_rel_fun (mkRel 1) (mkRel (len + 1))) in
- sigma, LocalAssum (Name argid', ss_term)
+ sigma, LocalAssum (make_annot (Name argid') Sorts.Relevant, ss_term)
in
let sigma, intern_bl =
let sigma, wfa = wfarg sigma 1 in
@@ -143,7 +144,7 @@ let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation =
in
let _intern_env = push_rel_context intern_bl env in
let sigma, proj = Evarutil.new_global sigma (delayed_force build_sigma).Coqlib.proj1 in
- let wfargpred = mkLambda (Name argid', argtyp, wf_rel_fun (mkRel 1) (mkRel 3)) in
+ let wfargpred = mkLambda (make_annot (Name argid') Sorts.Relevant, argtyp, wf_rel_fun (mkRel 1) (mkRel 3)) in
let projection = (* in wfarg :: arg :: before *)
mkApp (proj, [| argtyp ; wfargpred ; mkRel 1 |])
in
@@ -153,22 +154,23 @@ let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation =
now intern_arity is in wfarg :: arg *)
let sigma, wfa = wfarg sigma 1 in
let intern_fun_arity_prod = it_mkProd_or_LetIn intern_arity [wfa] in
- let intern_fun_binder = LocalAssum (Name (add_suffix recname "'"), intern_fun_arity_prod) in
+ let intern_fun_binder = LocalAssum (make_annot (Name (add_suffix recname "'")) Sorts.Relevant,
+ intern_fun_arity_prod) in
let sigma, curry_fun =
- let wfpred = mkLambda (Name argid', argtyp, wf_rel_fun (mkRel 1) (mkRel (2 * len + 4))) in
+ let wfpred = mkLambda (make_annot (Name argid') Sorts.Relevant, argtyp, wf_rel_fun (mkRel 1) (mkRel (2 * len + 4))) in
let sigma, intro = Evarutil.new_global sigma (delayed_force build_sigma).Coqlib.intro in
let arg = mkApp (intro, [| argtyp; wfpred; lift 1 make; mkRel 1 |]) in
let app = mkApp (mkRel (2 * len + 2 (* recproof + orig binders + current binders *)), [| arg |]) in
let rcurry = mkApp (rel, [| measure; lift len measure |]) in
- let lam = LocalAssum (Name (Id.of_string "recproof"), rcurry) in
+ let lam = LocalAssum (make_annot (Name (Id.of_string "recproof")) Sorts.Relevant, rcurry) in
let body = it_mkLambda_or_LetIn app (lam :: binders_rel) in
let ty = it_mkProd_or_LetIn (lift 1 top_arity) (lam :: binders_rel) in
- sigma, LocalDef (Name recname, body, ty)
+ sigma, LocalDef (make_annot (Name recname) Sorts.Relevant, body, ty)
in
let fun_bl = intern_fun_binder :: [arg] in
let lift_lets = lift_rel_context 1 letbinders in
let sigma, intern_body =
- let ctx = LocalAssum (Name recname, get_type curry_fun) :: binders_rel in
+ let ctx = LocalAssum (make_annot (Name recname) Sorts.Relevant, get_type curry_fun) :: binders_rel in
let (r, l, impls, scopes) =
Constrintern.compute_internalization_data env sigma
Constrintern.Recursive full_arity impls
@@ -180,7 +182,7 @@ let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation =
~impls:newimpls body (lift 1 top_arity)
in
let intern_body_lam = it_mkLambda_or_LetIn intern_body (curry_fun :: lift_lets @ fun_bl) in
- let prop = mkLambda (Name argname, argtyp, top_arity_let) in
+ let prop = mkLambda (make_annot (Name argname) Sorts.Relevant, argtyp, top_arity_let) in
(* XXX: Previous code did parallel evdref update, so possible old
weak ordering semantics may bite here. *)
let sigma, def =
@@ -272,7 +274,7 @@ let do_program_recursive local poly fixkind fixl ntns =
(List.length rec_sign) def typ
in (id, def, typ, imps, evars)
in
- let (fixnames,fixdefs,fixtypes) = fix in
+ let (fixnames,fixrs,fixdefs,fixtypes) = fix in
let fiximps = List.map pi2 info in
let fixdefs = List.map out_def fixdefs in
let defs = List.map4 collect_evars fixnames fixdefs fixtypes fiximps in
@@ -281,7 +283,7 @@ let do_program_recursive local poly fixkind fixl ntns =
(* XXX: are we allowed to have evars here? *)
let fixtypes = List.map (EConstr.to_constr ~abort_on_undefined_evars:false evd) fixtypes in
let fixdefs = List.map (EConstr.to_constr ~abort_on_undefined_evars:false evd) fixdefs in
- let fixdecls = Array.of_list (List.map (fun x -> Name x) fixnames),
+ let fixdecls = Array.of_list (List.map2 (fun x r -> make_annot (Name x) r) fixnames fixrs),
Array.of_list fixtypes,
Array.of_list (List.map (subst_vars (List.rev fixnames)) fixdefs)
in
diff --git a/vernac/himsg.ml b/vernac/himsg.ml
index fd5970e8ca..0e0788c470 100644
--- a/vernac/himsg.ml
+++ b/vernac/himsg.ml
@@ -14,6 +14,7 @@ open Names
open Nameops
open Namegen
open Constr
+open Context
open Termops
open Environ
open Pretype_errors
@@ -103,9 +104,9 @@ let canonize_constr sigma c =
let dn = Name.Anonymous in
let rec canonize_binders c =
match EConstr.kind sigma c with
- | Prod (_,t,b) -> mkProd(dn,t,b)
- | Lambda (_,t,b) -> mkLambda(dn,t,b)
- | LetIn (_,u,t,b) -> mkLetIn(dn,u,t,b)
+ | Prod (x,t,b) -> mkProd({x with binder_name=dn},t,b)
+ | Lambda (x,t,b) -> mkLambda({x with binder_name=dn},t,b)
+ | LetIn (x,u,t,b) -> mkLetIn({x with binder_name=dn},u,t,b)
| _ -> EConstr.map sigma canonize_binders c
in
canonize_binders c
@@ -193,13 +194,13 @@ let rec pr_disjunction pr = function
| a::l -> pr a ++ str "," ++ spc () ++ pr_disjunction pr l
| [] -> assert false
-let explain_elim_arity env sigma ind sorts c pj okinds =
+let explain_elim_arity env sigma ind c pj okinds =
let open EConstr in
let env = make_all_name_different env sigma in
let pi = pr_inductive env (fst ind) in
let pc = pr_leconstr_env env sigma c in
let msg = match okinds with
- | Some(kp,ki,explanation) ->
+ | Some(sorts,kp,ki,explanation) ->
let pki = Sorts.pr_sort_family ki in
let pkp = Sorts.pr_sort_family kp in
let explanation = match explanation with
@@ -262,7 +263,7 @@ let explain_ill_formed_branch env sigma c ci actty expty =
let explain_generalization env sigma (name,var) j =
let pe = pr_ne_context_of (str "In environment") env sigma in
let pv = pr_letype_env env sigma var in
- let (pc,pt) = pr_ljudge_env (push_rel_assum (name,var) env) sigma j in
+ let (pc,pt) = pr_ljudge_env (push_rel_assum (make_annot name Sorts.Relevant,var) env) sigma j in
pe ++ str "Cannot generalize" ++ brk(1,1) ++ pv ++ spc () ++
str "over" ++ brk(1,1) ++ pc ++ str "," ++ spc () ++
str "it has type" ++ spc () ++ pt ++
@@ -414,7 +415,7 @@ let explain_not_product env sigma c =
let explain_ill_formed_rec_body env sigma err names i fixenv vdefj =
let pr_lconstr_env env sigma c = pr_leconstr_env env sigma c in
let prt_name i =
- match names.(i) with
+ match names.(i).binder_name with
Name id -> str "Recursive definition of " ++ Id.print id
| Anonymous -> str "The " ++ pr_nth i ++ str " definition" in
@@ -429,7 +430,7 @@ let explain_ill_formed_rec_body env sigma err names i fixenv vdefj =
| RecursionOnIllegalTerm(j,(arg_env, arg),le,lt) ->
let arg_env = make_all_name_different arg_env sigma in
let called =
- match names.(j) with
+ match names.(j).binder_name with
Name id -> Id.print id
| Anonymous -> str "the " ++ pr_nth i ++ str " definition" in
let pr_db x = quote (pr_db env x) in
@@ -449,7 +450,7 @@ let explain_ill_formed_rec_body env sigma err names i fixenv vdefj =
| NotEnoughArgumentsForFixCall j ->
let called =
- match names.(j) with
+ match names.(j).binder_name with
Name id -> Id.print id
| Anonymous -> str "the " ++ pr_nth i ++ str " definition" in
str "Recursive call to " ++ called ++ str " has not enough arguments"
@@ -713,6 +714,9 @@ let explain_undeclared_universe env sigma l =
let explain_disallowed_sprop () =
Pp.(str "SProp not allowed, you need to use -allow-sprop.")
+let explain_bad_relevance env =
+ strbrk "Bad relevance (maybe a bugged tactic)."
+
let explain_type_error env sigma err =
let env = make_all_name_different env sigma in
match err with
@@ -726,8 +730,8 @@ let explain_type_error env sigma err =
explain_bad_assumption env sigma c
| ReferenceVariables (id,c) ->
explain_reference_variables sigma id c
- | ElimArity (ind, aritylst, c, pj, okinds) ->
- explain_elim_arity env sigma ind aritylst c pj okinds
+ | ElimArity (ind, c, pj, okinds) ->
+ explain_elim_arity env sigma ind c pj okinds
| CaseNotInductive cj ->
explain_case_not_inductive env sigma cj
| NumberBranches (cj, n) ->
@@ -755,6 +759,7 @@ let explain_type_error env sigma err =
| UndeclaredUniverse l ->
explain_undeclared_universe env sigma l
| DisallowedSProp -> explain_disallowed_sprop ()
+ | BadRelevance -> explain_bad_relevance env
let pr_position (cl,pos) =
let clpos = match cl with
diff --git a/vernac/indschemes.ml b/vernac/indschemes.ml
index 28ee3d184f..1e733acc59 100644
--- a/vernac/indschemes.ml
+++ b/vernac/indschemes.ml
@@ -228,17 +228,20 @@ let declare_one_case_analysis_scheme ind =
let kinds_from_prop =
[InType,rect_scheme_kind_from_prop;
InProp,ind_scheme_kind_from_prop;
- InSet,rec_scheme_kind_from_prop]
+ InSet,rec_scheme_kind_from_prop;
+ InSProp,sind_scheme_kind_from_prop]
let kinds_from_type =
[InType,rect_dep_scheme_kind_from_type;
InProp,ind_dep_scheme_kind_from_type;
- InSet,rec_dep_scheme_kind_from_type]
+ InSet,rec_dep_scheme_kind_from_type;
+ InSProp,sind_dep_scheme_kind_from_type]
let nondep_kinds_from_type =
[InType,rect_scheme_kind_from_type;
InProp,ind_scheme_kind_from_type;
- InSet,rec_scheme_kind_from_type]
+ InSet,rec_scheme_kind_from_type;
+ InSProp,sind_scheme_kind_from_type]
let declare_one_induction_scheme ind =
let (mib,mip) = Global.lookup_inductive ind in
diff --git a/vernac/lemmas.ml b/vernac/lemmas.ml
index 77f125e878..0d0732cbb4 100644
--- a/vernac/lemmas.ml
+++ b/vernac/lemmas.ml
@@ -330,7 +330,7 @@ let initialize_named_context_for_proof () =
List.fold_right
(fun d signv ->
let id = NamedDecl.get_id d in
- let d = if variable_opacity id then NamedDecl.LocalAssum (id, NamedDecl.get_type d) else d in
+ let d = if variable_opacity id then NamedDecl.drop_body d else d in
Environ.push_named_context_val d signv) sign Environ.empty_named_context_val
let start_proof id ?pl kind sigma ?terminator ?sign ?(compute_guard=[]) ?(hook : declaration_hook option) c =
diff --git a/vernac/obligations.ml b/vernac/obligations.ml
index 38cdfc2d7a..9aca48f529 100644
--- a/vernac/obligations.ml
+++ b/vernac/obligations.ml
@@ -13,6 +13,7 @@ open Declare
open Term
open Constr
+open Context
open Vars
open Names
open Evd
@@ -124,11 +125,11 @@ let etype_of_evar evs hyps concl =
| LocalDef (id,c,_) ->
let c', s'', trans'' = subst_evar_constr evs n mkVar c in
let c' = subst_vars acc 0 c' in
- mkNamedProd_or_LetIn (LocalDef (id, c', t'')) rest,
+ mkNamedProd_or_LetIn (LocalDef (id, c', t'')) rest,
Int.Set.union s'' s',
Id.Set.union trans'' trans'
- | LocalAssum (id,_) ->
- mkNamedProd_or_LetIn (LocalAssum (id, t'')) rest, s', trans')
+ | LocalAssum (id,_) ->
+ mkNamedProd_or_LetIn (LocalAssum (id, t'')) rest, s', trans')
| [] ->
let t', s, trans = subst_evar_constr evs n mkVar concl in
subst_vars acc 0 t', s, trans
@@ -479,7 +480,7 @@ let declare_definition prg =
let rec lam_index n t acc =
match Constr.kind t with
- | Lambda (Name n', _, _) when Id.equal n n' ->
+ | Lambda ({binder_name=Name n'}, _, _) when Id.equal n n' ->
acc
| Lambda (_, _, b) ->
lam_index n b (succ acc)
@@ -508,11 +509,12 @@ let declare_mutual_definition l =
let subs, typ = subst_prog oblsubst x in
let env = Global.env () in
let sigma = Evd.from_ctx x.prg_ctx in
+ let r = Retyping.relevance_of_type env sigma (EConstr.of_constr typ) in
let term = snd (Reductionops.splay_lam_n env sigma len (EConstr.of_constr subs)) in
let typ = snd (Reductionops.splay_prod_n env sigma len (EConstr.of_constr typ)) in
let term = EConstr.to_constr sigma term in
let typ = EConstr.to_constr sigma typ in
- let def = (x.prg_reduce term, x.prg_reduce typ, x.prg_implicits) in
+ let def = (x.prg_reduce term, r, x.prg_reduce typ, x.prg_implicits) in
let oblsubst = List.map (fun (id, (_, c)) -> id, c) oblsubst in
def, oblsubst
in
@@ -522,10 +524,12 @@ let declare_mutual_definition l =
(xdef :: defs, xobls @ obls)) l ([], [])
in
(* let fixdefs = List.map reduce_fix fixdefs in *)
- let fixdefs, fixtypes, fiximps = List.split3 defs in
+ let fixdefs, fixrs, fixtypes, fiximps = List.split4 defs in
let fixkind = Option.get first.prg_fixkind in
let arrrec, recvec = Array.of_list fixtypes, Array.of_list fixdefs in
- let fixdecls = (Array.of_list (List.map (fun x -> Name x.prg_name) l), arrrec, recvec) in
+ let rvec = Array.of_list fixrs in
+ let namevec = Array.of_list (List.map (fun x -> Name x.prg_name) l) in
+ let fixdecls = (Array.map2 make_annot namevec rvec, arrrec, recvec) in
let (local,poly,kind) = first.prg_kind in
let fixnames = first.prg_deps in
let opaque = first.prg_opaque in
diff --git a/vernac/record.ml b/vernac/record.ml
index 9f5658ecbd..6b223f845b 100644
--- a/vernac/record.ml
+++ b/vernac/record.ml
@@ -17,6 +17,7 @@ open Names
open Globnames
open Nameops
open Constr
+open Context
open Vars
open Environ
open Declarations
@@ -66,6 +67,7 @@ let interp_fields_evars env sigma impls_env nots l =
List.fold_left2
(fun (env, sigma, uimpls, params, impls) no ({CAst.loc;v=i}, b, t) ->
let sigma, (t', impl) = interp_type_evars_impls ~program_mode:false env sigma ~impls t in
+ let r = Retyping.relevance_of_type env sigma t' in
let sigma, b' =
Option.cata (fun x -> on_snd (fun x -> Some (fst x)) @@
interp_casted_constr_evars_impls ~program_mode:false env sigma ~impls x t') (sigma,None) b in
@@ -75,8 +77,8 @@ let interp_fields_evars env sigma impls_env nots l =
| Name id -> Id.Map.add id (compute_internalization_data env sigma Constrintern.Method t' impl) impls
in
let d = match b' with
- | None -> LocalAssum (i,t')
- | Some b' -> LocalDef (i,b',t')
+ | None -> LocalAssum (make_annot i r,t')
+ | Some b' -> LocalDef (make_annot i r,b',t')
in
List.iter (Metasyntax.set_notation_for_interpretation env impls) no;
(EConstr.push_rel d env, sigma, impl :: uimpls, d::params, impls))
@@ -144,8 +146,10 @@ let typecheck_params_and_fields finite def poly pl ps records =
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 relevances = List.map (fun (_,s) -> Sorts.relevance_of_sort s) typs in
+ let fold accu (id, _, _, _) arity r =
+ EConstr.push_rel (LocalAssum (make_annot (Name id) r,arity)) accu in
+ let env_ar = EConstr.push_rel_context newps (List.fold_left3 fold env0 records arities relevances) in
let assums = List.filter is_local_assum newps in
let impls_env =
let params = List.map (RelDecl.get_name %> Name.get_id) assums in
@@ -213,12 +217,12 @@ let warning_or_error coe indsp err =
strbrk " not defined.")
| BadTypedProj (fi,ctx,te) ->
match te with
- | ElimArity (_,_,_,_,Some (_,_,NonInformativeToInformative)) ->
+ | ElimArity (_,_,_,Some (_,_,_,NonInformativeToInformative)) ->
(Id.print fi ++
strbrk" cannot be defined because it is informative and " ++
Printer.pr_inductive (Global.env()) indsp ++
strbrk " is not.")
- | ElimArity (_,_,_,_,Some (_,_,StrongEliminationOnNonSmallType)) ->
+ | ElimArity (_,_,_,Some (_,_,_,StrongEliminationOnNonSmallType)) ->
(Id.print fi ++
strbrk" cannot be defined because it is large and " ++
Printer.pr_inductive (Global.env()) indsp ++
@@ -284,7 +288,7 @@ let declare_projections indsp ctx ?(kind=StructureComponent) binder_name coers f
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 x = make_annot (Name binder_name) mip.mind_relevant 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 =
@@ -316,18 +320,19 @@ let declare_projections indsp ctx ?(kind=StructureComponent) binder_name coers f
else
let ccl = subst_projection fid subst ti in
let body = match decl with
- | LocalDef (_,ci,_) -> subst_projection fid subst ci
- | LocalAssum _ ->
+ | LocalDef (_,ci,_) -> subst_projection fid subst ci
+ | LocalAssum ({binder_relevance=rci},_) ->
(* [ccl] is defined in context [params;x:rp] *)
(* [ccl'] is defined in context [params;x:rp;x:rp] *)
let ccl' = liftn 1 2 ccl in
- let p = mkLambda (x, lift 1 rp, ccl') in
+ let p = mkLambda (x, lift 1 rp, ccl') in
let branch = it_mkLambda_or_LetIn (mkRel nfi) lifted_fields in
- let ci = Inductiveops.make_case_info env indsp LetStyle in
- mkCase (ci, p, mkRel 1, [|branch|])
- in
+ let ci = Inductiveops.make_case_info env indsp rci LetStyle in
+ (* Record projections have no is *)
+ mkCase (ci, p, mkRel 1, [|branch|])
+ in
let proj =
- it_mkLambda_or_LetIn (mkLambda (x,rp,body)) paramdecls in
+ it_mkLambda_or_LetIn (mkLambda (x,rp,body)) paramdecls in
let projtyp =
it_mkProd_or_LetIn (mkProd (x,rp,ccl)) paramdecls in
try
@@ -463,7 +468,9 @@ let declare_class def cum ubinders univs id idbuild paramimpls params arity
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 ->
+ | [LocalAssum ({binder_name=Name proj_name} as binder, field)
+ | LocalDef ({binder_name=Name proj_name} as binder, _, field)] when def ->
+ let binder = {binder with binder_name=Name binder_name} in
let class_body = it_mkLambda_or_LetIn field params in
let class_type = it_mkProd_or_LetIn arity params in
let class_entry =
@@ -477,11 +484,11 @@ let declare_class def cum ubinders univs id idbuild paramimpls params arity
in
let cstu = (cst, inst) in
let inst_type = appvectc (mkConstU cstu)
- (Termops.rel_vect 0 (List.length params)) in
+ (Termops.rel_vect 0 (List.length params)) in
let proj_type =
- it_mkProd_or_LetIn (mkProd(Name binder_name, inst_type, lift 1 field)) params in
+ it_mkProd_or_LetIn (mkProd(binder, inst_type, lift 1 field)) params in
let proj_body =
- it_mkLambda_or_LetIn (mkLambda (Name binder_name, inst_type, mkRel 1)) params in
+ it_mkLambda_or_LetIn (mkLambda (binder, inst_type, mkRel 1)) params in
let proj_entry = Declare.definition_entry ~types:proj_type ~univs proj_body in
let proj_cst = Declare.declare_constant proj_name
(DefinitionEntry proj_entry, IsDefinition Definition)
@@ -548,12 +555,13 @@ let declare_class def cum ubinders univs id idbuild paramimpls params arity
let add_constant_class env cst =
let ty, univs = Typeops.type_of_global_in_context env (ConstRef cst) in
+ let r = (Environ.lookup_constant cst env).const_relevance in
let ctx, arity = decompose_prod_assum ty in
let tc =
{ cl_univs = univs;
cl_impl = ConstRef cst;
cl_context = (List.map (const None) ctx, ctx);
- cl_props = [LocalAssum (Anonymous, arity)];
+ cl_props = [LocalAssum (make_annot Anonymous r, arity)];
cl_projs = [];
cl_strict = !typeclasses_strict;
cl_unique = !typeclasses_unique
@@ -570,10 +578,11 @@ let add_inductive_class env ind =
let env = push_rel_context ctx env in
let inst = Univ.make_abstract_instance univs in
let ty = Inductive.type_of_inductive env ((mind, oneind), inst) in
+ let r = Inductive.relevance_of_inductive env ind in
{ cl_univs = univs;
cl_impl = IndRef ind;
cl_context = List.map (const None) ctx, ctx;
- cl_props = [LocalAssum (Anonymous, ty)];
+ cl_props = [LocalAssum (make_annot Anonymous r, ty)];
cl_projs = [];
cl_strict = !typeclasses_strict;
cl_unique = !typeclasses_unique }
diff --git a/vernac/vernacentries.ml b/vernac/vernacentries.ml
index d227834fcf..b9d0a27b39 100644
--- a/vernac/vernacentries.ml
+++ b/vernac/vernacentries.ml
@@ -144,8 +144,8 @@ let make_cases_aux glob_ref =
| [] -> []
| RelDecl.LocalDef _ :: l -> "_" :: rename avoid l
| RelDecl.LocalAssum (n, _)::l ->
- let n' = Namegen.next_name_away_with_default (Id.to_string Namegen.default_dependent_ident) n avoid in
- Id.to_string n' :: rename (Id.Set.add n' avoid) l in
+ let n' = Namegen.next_name_away_with_default (Id.to_string Namegen.default_dependent_ident) n.Context.binder_name avoid in
+ Id.to_string n' :: rename (Id.Set.add n' avoid) l in
let al' = rename Id.Set.empty al in
let consref = ConstructRef (ith_constructor_of_inductive ind (i + 1)) in
(Libnames.string_of_qualid (Nametab.shortest_qualid_of_global Id.Set.empty consref) :: al') :: l)