diff --git a/checker/checkInductive.ml b/checker/checkInductive.ml index 655ac2d0b395..9bca850c2a0a 100644 --- a/checker/checkInductive.ml +++ b/checker/checkInductive.ml @@ -12,6 +12,7 @@ open Declarations open Environ open Names open Univ +open UVars open Util [@@@ocaml.warning "+9+27"] @@ -181,7 +182,7 @@ let check_inductive env mind mb = check "mind_finite" (mb.mind_finite == mind_finite); check "mind_ntypes" Int.(equal mb.mind_ntypes mind_ntypes); check "mind_hyps" (List.is_empty mind_hyps); - check "mind_univ_hyps" (Univ.Instance.is_empty mind_univ_hyps); + check "mind_univ_hyps" (UVars.Instance.is_empty mind_univ_hyps); check "mind_nparams" Int.(equal mb.mind_nparams mind_nparams); check "mind_nparams_rec" (mb.mind_nparams_rec <= mind_nparams_rec); @@ -191,7 +192,7 @@ let check_inductive env mind mb = check "mind_params_ctxt" (Context.Rel.equal Constr.equal mb.mind_params_ctxt mind_params_ctxt); ignore mind_universes; (* Indtypes did the necessary checking *) check "mind_template" (check_template mb.mind_template mind_template); - check "mind_variance" (Option.equal (Array.equal Univ.Variance.equal) + check "mind_variance" (Option.equal (Array.equal UVars.Variance.equal) mb.mind_variance mind_variance); check "mind_sec_variance" (Option.is_empty mind_sec_variance); ignore mind_private; (* passed through Indtypes *) diff --git a/checker/checker.ml b/checker/checker.ml index e1420b3631a3..f6a0e04e3769 100644 --- a/checker/checker.ml +++ b/checker/checker.ml @@ -245,7 +245,7 @@ let explain_exn = function let msg = if CDebug.(get_flag misc) then str "." ++ spc() ++ - UGraph.explain_universe_inconsistency Univ.Level.raw_pr i + UGraph.explain_universe_inconsistency Sorts.QVar.raw_pr Univ.Level.raw_pr i else mt() in hov 0 (str "Error: Universe inconsistency" ++ msg ++ str ".") @@ -288,11 +288,13 @@ let explain_exn = function | CantApplyNonFunctional _ -> str"CantApplyNonFunctional" | IllFormedRecBody _ -> str"IllFormedRecBody" | IllTypedRecBody _ -> str"IllTypedRecBody" + | UnsatisfiedQConstraints _ -> str"UnsatisfiedQConstraints" | UnsatisfiedConstraints _ -> str"UnsatisfiedConstraints" | DisallowedSProp -> str"DisallowedSProp" | BadBinderRelevance _ -> str"BadBinderRelevance" | BadCaseRelevance _ -> str"BadCaseRelevance" | BadInvert -> str"BadInvert" + | UndeclaredQualities _ -> str"UndeclaredQualities" | UndeclaredUniverse _ -> str"UndeclaredUniverse" | BadVariance _ -> str "BadVariance" | UndeclaredUsedVariables _ -> str "UndeclaredUsedVariables" diff --git a/checker/mod_checking.ml b/checker/mod_checking.ml index f441fa054fc7..dc6053a3cb2c 100644 --- a/checker/mod_checking.ml +++ b/checker/mod_checking.ml @@ -41,7 +41,7 @@ let check_constant_declaration env opac kn cb opacify = ones in const_universes should not be needed *) false, env | Polymorphic auctx -> - let ctx = Univ.AbstractContext.repr auctx in + let ctx = UVars.AbstractContext.repr auctx in (* [env] contains De Bruijn universe variables *) let env = push_context ~strict:false ctx env in true, env diff --git a/checker/values.ml b/checker/values.ml index c997c576b9ff..96ba9d0c4444 100644 --- a/checker/values.ml +++ b/checker/values.ml @@ -94,6 +94,12 @@ let v_level = v_tuple "level" [|Int;v_raw_level|] let v_expr = v_tuple "levelexpr" [|v_level;Int|] let v_univ = List v_expr +let v_qvar = v_sum "qvar" 0 [|[|Int|];[|String;Int|]|] + +let v_constant_quality = v_enum "constant_quality" 3 + +let v_quality = v_sum "quality" 0 [|[|v_qvar|];[|v_constant_quality|]|] + let v_cstrs = Annot ("Univ.constraints", @@ -103,16 +109,16 @@ let v_cstrs = let v_variance = v_enum "variance" 3 -let v_instance = Annot ("instance", Array v_level) -let v_abs_context = v_tuple "abstract_universe_context" [|Array v_name; v_cstrs|] +let v_instance = Annot ("instance", v_pair (Array v_quality) (Array v_level)) +let v_abs_context = v_tuple "abstract_universe_context" [|v_pair (Array v_name) (Array v_name); v_cstrs|] let v_context_set = v_tuple "universe_context_set" [|v_hset v_level;v_cstrs|] (** kernel/term *) -let v_sort = v_sum "sort" 3 (*SProp, Prop, Set*) [|[|v_univ(*Type*)|]|] +let v_sort = v_sum "sort" 3 (*SProp, Prop, Set*) [|[|v_univ(*Type*)|];[|v_qvar;v_univ(*QSort*)|]|] let v_sortfam = v_enum "sorts_family" 4 -let v_relevance = v_sum "relevance" 2 [||] +let v_relevance = v_sum "relevance" 2 [|[|v_qvar|]|] let v_binder_annot x = v_tuple "binder_annot" [|x;v_relevance|] let v_puniverses v = v_tuple "punivs" [|v;v_instance|] @@ -126,7 +132,7 @@ let v_caseinfo = let v_cast = v_enum "cast_kind" 3 -let v_proj_repr = v_tuple "projection_repr" [|v_ind;v_bool;Int;Int;v_id|] +let v_proj_repr = v_tuple "projection_repr" [|v_ind;Int;Int;v_id|] let v_proj = v_tuple "projection" [|v_proj_repr; v_bool|] let v_uint63 = @@ -150,7 +156,7 @@ let rec v_constr = [|v_caseinfo;v_instance; Array v_constr; v_case_return; v_case_invert; v_constr; Array v_case_branch|]; (* Case *) [|v_fix|]; (* Fix *) [|v_cofix|]; (* CoFix *) - [|v_proj;v_constr|]; (* Proj *) + [|v_proj;v_relevance;v_constr|]; (* Proj *) [|v_uint63|]; (* Int *) [|Float64|]; (* Float *) [|v_instance;Array v_constr;v_constr;v_constr|] (* Array *) diff --git a/dev/bench/bench.sh b/dev/bench/bench.sh index d08a6aa553ae..109ed1e1fb32 100755 --- a/dev/bench/bench.sh +++ b/dev/bench/bench.sh @@ -63,9 +63,9 @@ check_variable () { : "${old_ocaml_version:=4.09.1}" : "${new_coq_repository:=$CI_REPOSITORY_URL}" : "${old_coq_repository:=$CI_REPOSITORY_URL}" -: "${new_coq_opam_archive_git_uri:=https://github.com/coq/opam-coq-archive.git}" +: "${new_coq_opam_archive_git_uri:=https://github.com/skyskimmer/opam-coq-archive.git}" : "${old_coq_opam_archive_git_uri:=https://github.com/coq/opam-coq-archive.git}" -: "${new_coq_opam_archive_git_branch:=master}" +: "${new_coq_opam_archive_git_branch:=sort-poly}" : "${old_coq_opam_archive_git_branch:=master}" : "${num_of_iterations:=1}" : "${timeout:=3h}" diff --git a/dev/ci/user-overlays/17836-SkySkimmer-sort-poly.sh b/dev/ci/user-overlays/17836-SkySkimmer-sort-poly.sh new file mode 100644 index 000000000000..1a02af4d643a --- /dev/null +++ b/dev/ci/user-overlays/17836-SkySkimmer-sort-poly.sh @@ -0,0 +1,33 @@ +overlay bedrock2 https://github.com/SkySkimmer/bedrock2 sort-poly 17836 + +overlay rewriter https://github.com/SkySkimmer/rewriter sort-poly 17836 + +overlay fiat_crypto https://github.com/SkySkimmer/fiat-crypto sort-poly 17836 + +overlay waterproof https://github.com/SkySkimmer/coq-waterproof sort-poly 17836 + +overlay tactician https://github.com/SkySkimmer/coq-tactician sort-poly 17836 + +overlay serapi https://github.com/SkySkimmer/coq-serapi sort-poly 17836 + +overlay quickchick https://github.com/SkySkimmer/QuickChick sort-poly 17836 + +overlay paramcoq https://github.com/SkySkimmer/paramcoq sort-poly 17836 + +overlay unicoq https://github.com/SkySkimmer/unicoq sort-poly 17836 + +overlay mtac2 https://github.com/SkySkimmer/Mtac2 sort-poly 17836 + +overlay lean_importer https://github.com/SkySkimmer/coq-lean-import sort-poly 17836 + +overlay equations https://github.com/SkySkimmer/Coq-Equations sort-poly 17836 + +overlay elpi https://github.com/SkySkimmer/coq-elpi sort-poly 17836 + +overlay coqhammer https://github.com/SkySkimmer/coqhammer sort-poly 17836 + +overlay coq_dpdgraph https://github.com/SkySkimmer/coq-dpdgraph sort-poly 17836 + +overlay metacoq https://github.com/SkySkimmer/metacoq sort-poly 17836 + +overlay coq_lsp https://github.com/SkySkimmer/coq-lsp sort-poly 17836 diff --git a/dev/doc/changes.md b/dev/doc/changes.md index 83c984ba21c5..6dc3fdfa2e4a 100644 --- a/dev/doc/changes.md +++ b/dev/doc/changes.md @@ -67,7 +67,7 @@ In general, code can be adapted by: ### Universes -- Type `Univ.UContext` now embeds universe user names, generally +- Type `UVars.UContext` now embeds universe user names, generally resulting in more concise code. - Renaming `Univ.Constraint` into `Univ.Constraints` diff --git a/dev/doc/universes.md b/dev/doc/universes.md index 883d635ade83..c80f15f4209b 100644 --- a/dev/doc/universes.md +++ b/dev/doc/universes.md @@ -6,7 +6,7 @@ to the API of Coq. First and foremost, the term language changes, as global references now carry a universe level substitution: ~~~ocaml -type 'a puniverses = 'a * Univ.Instance.t +type 'a puniverses = 'a * UVars.Instance.t type pconstant = constant puniverses type pinductive = inductive puniverses type pconstructor = constructor puniverses @@ -35,7 +35,7 @@ universes etc.. Structures are hashconsed (with a hack to take care of the fact that deserialization breaks sharing). Definitions (constants, inductives) now carry around not only -constraints but also the universes they introduced (a Univ.UContext.t). +constraints but also the universes they introduced (a UVars.UContext.t). There is another kind of contexts `Univ.ContextSet.t`, the latter has a set of universes, while the former has serialized the levels in an array, and is used for polymorphic objects. Both have "reified" diff --git a/dev/include_printers b/dev/include_printers index 938d89f251c5..1164543107c8 100644 --- a/dev/include_printers +++ b/dev/include_printers @@ -14,6 +14,7 @@ #install_printer (* univ context *) ppaucontext;; #install_printer (* univ context future *) ppuniverse_context_future;; #install_printer (* univ context set *) ppuniverse_context_set;; +#install_printer (* qvar set *) ppqvarset;; #install_printer (* univ set *) ppuniverse_set;; #install_printer (* univ instance *) ppuniverse_instance;; #install_printer (* univ subst *) ppuniverse_subst;; diff --git a/dev/top_printers.dbg b/dev/top_printers.dbg index 95b81ba8ce65..9c8e8109c141 100644 --- a/dev/top_printers.dbg +++ b/dev/top_printers.dbg @@ -69,6 +69,7 @@ install_printer Top_printers.ppuni install_printer Top_printers.ppesorts install_printer Top_printers.ppqvar install_printer Top_printers.ppuni_level +install_printer Top_printers.ppqvarset install_printer Top_printers.ppuniverse_set install_printer Top_printers.ppuniverse_instance install_printer Top_printers.ppuniverse_context diff --git a/dev/top_printers.ml b/dev/top_printers.ml index 079d526cc77c..8e6b048b6880 100644 --- a/dev/top_printers.ml +++ b/dev/top_printers.ml @@ -16,6 +16,7 @@ open Pp open Names open Libnames open Univ +open UVars open Environ open Printer open Constr @@ -256,19 +257,21 @@ let pproof p = pp(Proof.pr_proof p) let ppuni u = pp(Universe.raw_pr u) let ppuni_level u = pp (Level.raw_pr u) -let ppqvar q = pp (QVar.pr q) +let ppqvar q = pp (QVar.raw_pr q) let ppesorts s = pp (Sorts.debug_print (Evd.MiniEConstr.ESorts.unsafe_to_sorts s)) -let prlev l = UnivNames.pr_with_global_universes l +let prlev l = UnivNames.pr_level_with_global_universes l +let prqvar q = Sorts.QVar.raw_pr q +let ppqvarset l = pp (hov 1 (str "{" ++ prlist_with_sep spc QVar.raw_pr (QVar.Set.elements l) ++ str "}")) let ppuniverse_set l = pp (Level.Set.pr prlev l) -let ppuniverse_instance l = pp (Instance.pr prlev l) -let ppuniverse_context l = pp (pr_universe_context prlev l) +let ppuniverse_instance l = pp (Instance.pr prqvar prlev l) +let ppuniverse_context l = pp (pr_universe_context prqvar prlev l) let ppuniverse_context_set l = pp (pr_universe_context_set prlev l) let ppuniverse_subst l = pp (UnivSubst.pr_universe_subst Level.raw_pr l) let ppuniverse_opt_subst l = pp (UnivFlex.pr Level.raw_pr l) let ppuniverse_level_subst l = pp (Univ.pr_universe_level_subst Level.raw_pr l) let ppevar_universe_context l = pp (Termops.pr_evar_universe_context l) -let ppconstraints c = pp (pr_constraints Level.raw_pr c) +let ppconstraints c = pp (Constraints.pr Level.raw_pr c) let ppuniverseconstraints c = pp (UnivProblem.Set.pr c) let ppuniverse_context_future c = let ctx = Future.force c in @@ -280,14 +283,16 @@ let ppnamedcontextval e = pp (pr_named_context env sigma (named_context_of_val e)) let ppaucontext auctx = - let nas = AbstractContext.names auctx in - let prlev l = match Level.var_index l with + let qnas, unas = AbstractContext.names auctx in + let prgen pr var_index nas l = match var_index l with | Some n -> (match nas.(n) with - | Anonymous -> prlev l + | Anonymous -> pr l | Name id -> Id.print id) - | None -> prlev l + | None -> pr l in - pp (pr_universe_context prlev (AbstractContext.repr auctx)) + let prqvar l = prgen prqvar Sorts.QVar.var_index qnas l in + let prlev l = prgen prlev Level.var_index unas l in + pp (pr_universe_context prqvar prlev (AbstractContext.repr auctx)) let ppenv e = pp @@ -342,7 +347,7 @@ let constr_display csr = | Construct (((sp,i),j),u) -> "MutConstruct(("^(MutInd.to_string sp)^","^(string_of_int i)^")," ^","^(universes_display u)^(string_of_int j)^")" - | Proj (p, c) -> "Proj("^(Constant.to_string (Projection.constant p))^","^term_display c ^")" + | Proj (p, r, c) -> "Proj("^(Constant.to_string (Projection.constant p))^","^term_display c ^")" | Case (ci,u,pms,(_,p),iv,c,bl) -> "MutCase(,"^(term_display p)^","^(term_display c)^"," ^(array_display (Array.map snd bl))^")" @@ -374,6 +379,9 @@ let constr_display csr = and univ_display u = incr cnt; pp (str "with " ++ int !cnt ++ str" " ++ Universe.raw_pr u ++ fnl ()) + and quality_display q = + incr cnt; pp (str "with " ++ int !cnt ++ str" " ++ Sorts.Quality.raw_pr q ++ fnl ()) + and level_display u = incr cnt; pp (str "with " ++ int !cnt ++ str" " ++ Level.raw_pr u ++ fnl ()) @@ -386,8 +394,15 @@ let constr_display csr = | QSort (q, u) -> univ_display u; Printf.sprintf "QSort(%s, %i)" (Sorts.QVar.to_string q) !cnt and universes_display l = + let qs, us = Instance.to_array l in + let qs = Array.fold_right (fun x i -> + quality_display x; + (string_of_int !cnt)^ + (if not(i="") then (" "^i) else "")) + qs "" + in Array.fold_right (fun x i -> level_display x; (string_of_int !cnt)^(if not(i="") - then (" "^i) else "")) (Instance.to_array l) "" + then (" "^i) else "")) us (if qs = "" then "" else (qs^" | ")) and name_display x = match x.binder_name with | Name id -> "Name("^(Id.to_string id)^")" @@ -440,7 +455,7 @@ let print_pure_constr csr = sp_con_display c; print_string ","; universes_display u; print_string ")" - | Proj (p,c') -> print_string "Proj("; + | Proj (p,_,c') -> print_string "Proj("; sp_con_display (Projection.constant p); print_string ","; box_display c'; @@ -526,7 +541,9 @@ let print_pure_constr csr = and box_display c = open_hovbox 1; term_display c; close_box() and universes_display u = - Array.iter (fun u -> print_space (); pp (Level.raw_pr u)) (Instance.to_array u) + let qs, us = Instance.to_array u in + Array.iter (fun u -> print_space (); pp (Sorts.Quality.raw_pr u)) qs; + Array.iter (fun u -> print_space (); pp (Level.raw_pr u)) us and sort_display = function | SProp -> print_string "SProp" @@ -535,7 +552,7 @@ let print_pure_constr csr = | Type u -> open_hbox(); print_string "Type("; pp (Universe.raw_pr u); print_string ")"; close_box() | QSort (q, u) -> open_hbox(); - print_string "QSort("; pp (QVar.pr q); print_string ", "; pp (Universe.raw_pr u); print_string ")"; close_box() + print_string "QSort("; pp (QVar.raw_pr q); print_string ", "; pp (Universe.raw_pr u); print_string ")"; close_box() and name_display x = match x.binder_name with | Name id -> print_string (Id.to_string id) diff --git a/dev/top_printers.mli b/dev/top_printers.mli index 7a81b39aab70..c97142381d6d 100644 --- a/dev/top_printers.mli +++ b/dev/top_printers.mli @@ -147,10 +147,11 @@ val ppuni_level : Univ.Level.t -> unit (* raw *) val ppqvar : Sorts.QVar.t -> unit val ppesorts : EConstr.ESorts.t -> unit val prlev : Univ.Level.t -> Pp.t (* with global names (does this work?) *) +val ppqvarset : Sorts.QVar.Set.t -> unit val ppuniverse_set : Univ.Level.Set.t -> unit -val ppuniverse_instance : Univ.Instance.t -> unit -val ppuniverse_context : Univ.UContext.t -> unit -val ppaucontext : Univ.AbstractContext.t -> unit +val ppuniverse_instance : UVars.Instance.t -> unit +val ppuniverse_context : UVars.UContext.t -> unit +val ppaucontext : UVars.AbstractContext.t -> unit val ppuniverse_context_set : Univ.ContextSet.t -> unit val ppuniverse_subst : UnivSubst.universe_subst -> unit val ppuniverse_opt_subst : UState.universe_opt_subst -> unit @@ -158,7 +159,7 @@ val ppuniverse_level_subst : Univ.universe_level_subst -> unit val ppevar_universe_context : UState.t -> unit val ppconstraints : Univ.Constraints.t -> unit val ppuniverseconstraints : UnivProblem.Set.t -> unit -val ppuniverse_context_future : Univ.UContext.t Future.computation -> unit +val ppuniverse_context_future : UVars.UContext.t Future.computation -> unit val ppuniverses : UGraph.t -> unit val ppnamedcontextval : Environ.named_context_val -> unit diff --git a/doc/changelog/01-kernel/17836-sort-poly.rst b/doc/changelog/01-kernel/17836-sort-poly.rst new file mode 100644 index 000000000000..d34ac26360ef --- /dev/null +++ b/doc/changelog/01-kernel/17836-sort-poly.rst @@ -0,0 +1,5 @@ +- **Added:** + :ref:`sort-polymorphism` makes it possible to share common constructs + over `Type` `Prop` and `SProp` + (`#17836 `_, + by Gaëtan Gilbert). diff --git a/doc/sphinx/addendum/universe-polymorphism.rst b/doc/sphinx/addendum/universe-polymorphism.rst index 9b36e27acb62..8d791d02a918 100644 --- a/doc/sphinx/addendum/universe-polymorphism.rst +++ b/doc/sphinx/addendum/universe-polymorphism.rst @@ -530,14 +530,15 @@ Explicit Universes universe_name ::= @qualid | Set | Prop - univ_annot ::= @%{ {* @universe_level } %} - universe_level ::= Set + univ_annot ::= @%{ {* @univ_level_or_quality } {? %| {* @univ_level_or_quality } } %} + univ_level_or_quality ::= Set + | SProp | Prop | Type | _ | @qualid - univ_decl ::= @%{ {* @ident } {? + } {? %| {*, @univ_constraint } {? + } } %} - cumul_univ_decl ::= @%{ {* {? {| + | = | * } } @ident } {? + } {? %| {*, @univ_constraint } {? + } } %} + univ_decl ::= @%{ {? {* @ident } %| } {* @ident } {? + } {? %| {*, @univ_constraint } {? + } } %} + cumul_univ_decl ::= @%{ {? {* @ident } %| } {* {? {| + | = | * } } @ident } {? + } {? %| {*, @univ_constraint } {? + } } %} univ_constraint ::= @universe_name {| < | = | <= } @universe_name The syntax has been extended to allow users to explicitly bind names @@ -739,6 +740,69 @@ underscore or by omitting the annotation to a polymorphic definition. About baz. +.. _sort-polymorphism: + +Sort polymorphism +----------------- + +Quantifying over universes does not allow instantiation with `Prop` or `SProp`. For instance + +.. coqtop:: in reset + + Polymorphic Definition type@{u} := Type@{u}. + +.. coqtop:: all + + Fail Check type@{Prop}. + +To be able to instantiate a sort with `Prop` or `SProp`, we must +quantify over :gdef:`sort qualities`. Definitions which quantify over +sort qualities are called :gdef:`sort polymorphic`. + +All sort quality variables must be explicitly bound. + +.. coqtop:: all + + Polymorphic Definition sort@{s | u |} := Type@{s|u}. + +To help the parser, both `|` in the :n:`@univ_decl` are required. + +Sort quality variables of a sort polymorphic definition may be +instantiated by the concrete values `SProp`, `Prop` and `Type` or by a +bound variable. + +Instantiating `s` in `Type@{s|u}` with the impredicative `Prop` or +`SProp` produces `Prop` or `SProp` respectively regardless of the +instantiation fof `u`. + +.. coqtop:: all + + Eval cbv in sort@{Prop|Set}. + Eval cbv in sort@{Type|Set}. + +When no explicit instantiation is provided or `_` is used, a temporary +variable is generated. Temporary sort variables are instantiated with +`Type` if not unified with another quality when universe minimization +runs (typically at the end of a definition). + +:cmd:`Check` and :cmd:`Eval` run minimization so we cannot use them to +witness these temporary variables. + +.. coqtop:: in + + Goal True. + Set Printing Universes. + +.. coqtop:: all abort + + let c := constr:(sort) in idtac c. + +.. note:: + + We recommend you do not name explicitly quantified sort variables + `α` followed by a number as printing will not distinguish between + your bound variables and temporary variables. + .. _universe-polymorphism-in-sections: Universe polymorphism and sections diff --git a/doc/sphinx/language/core/sorts.rst b/doc/sphinx/language/core/sorts.rst index 04aee32f6fcb..7fe731804b8a 100644 --- a/doc/sphinx/language/core/sorts.rst +++ b/doc/sphinx/language/core/sorts.rst @@ -17,7 +17,7 @@ Sorts | SProp | Type | Type @%{ _ %} - | Type @%{ @universe %} + | Type @%{ {? @qualid %| } @universe %} universe ::= max ( {+, @universe_expr } ) | @universe_expr universe_expr ::= @universe_name {? + @natural } @@ -71,31 +71,30 @@ Formally, we call :math:`\Sort` the set of sorts which is defined by: Their properties, such as :math:`\Prop:\Type(1)`, :math:`\Set:\Type(1)`, and :math:`\Type(i):\Type(i+1)`, are described in :ref:`subtyping-rules`. -The user does not have to mention explicitly the index :math:`i` when -referring to the universe :math:`\Type(i)`. One only writes `Type`. The system -itself generates for each instance of `Type` a new index for the +**Algebraic universes** In practice, the Type hierarchy is +implemented using algebraic universes, +which appear in the syntax :n:`Type@{@universe}`. +An :gdef:`algebraic universe` :math:`u` is either a variable, +a successor of an algebraic universe (an expression :math:`u+1`), +an upper bound of algebraic universes (an expression :math:`\max(u_1 ,...,u_n )`), +or the base universe :math:`\Set`. + +A graph of constraints between the universe variables is maintained +globally. To ensure the existence of a mapping of the universes to the +positive integers, the graph of constraints must remain acyclic. +Typing expressions that violate the acyclicity of the graph of +constraints results in a :exn:`Universe inconsistency` error. + +The user does not have to mention explicitly the universe :math:`u` when +referring to the universe `Type@{u}`. One only writes `Type`. The system +itself generates for each instance of `Type` a new variable for the universe and checks that the constraints between these indexes can be solved. From the user point of view we consequently have :math:`\Type:\Type`. We shall make precise in the typing rules the constraints between the indices. - -.. _Implementation-issues: - -**Implementation issues** In practice, the Type hierarchy is -implemented using algebraic universes. -An :gdef:`algebraic universe` :math:`u` is either a variable (a qualified -identifier with a number) or a successor of an algebraic universe (an -expression :math:`u+1`), or an upper bound of algebraic universes (an -expression :math:`\max(u_1 ,...,u_n )`), or the base universe (the expression -:math:`0`) which corresponds, in the arity of template polymorphic inductive -types (see Section -:ref:`well-formed-inductive-definitions`), -to the predicative sort :math:`\Set`. A graph of -constraints between the universe variables is maintained globally. To -ensure the existence of a mapping of the universes to the positive -integers, the graph of constraints must remain acyclic. Typing -expressions that violate the acyclicity of the graph of constraints -results in a :exn:`Universe inconsistency` error. +The syntax :n:`Type@{@qualid | @universe}` is used with +:ref:`polymorphicuniverses` when quantifying over all sorts including +:math:`\Prop` and :math:`\SProp`. .. seealso:: :ref:`printing-universes`, :ref:`explicit-universes`. diff --git a/doc/tools/docgram/common.edit_mlg b/doc/tools/docgram/common.edit_mlg index b9ada86d7f09..ed3dd1b29386 100644 --- a/doc/tools/docgram/common.edit_mlg +++ b/doc/tools/docgram/common.edit_mlg @@ -163,6 +163,7 @@ DELETE: [ | test_variance_ident | test_qualid_with_or_lpar_or_rbrac | test_leftsquarebracket_equal +| test_sort_qvar (* unused *) | constr_comma_sequence' @@ -326,6 +327,12 @@ scope_delimiter: [ | WITH "%" scope ] +sort: [ +| REPLACE "Type" "@{" reference "|" universe "}" +| WITH "Type" "@{" OPT [ qualid "|" ] universe "}" +| DELETE "Type" "@{" universe "}" +] + term100: [ | REPLACE term99 "<:" term200 | WITH term99 "<:" type @@ -624,13 +631,15 @@ inline: [ ] univ_decl: [ -| REPLACE "@{" LIST0 identref [ "+" | ] [ "|" LIST0 univ_constraint SEP "," [ "+" | ] "}" | [ "}" | bar_cbrace ] ] -| WITH "@{" LIST0 identref OPT "+" OPT [ "|" LIST0 univ_constraint SEP "," OPT "+" ] "}" +| REPLACE "@{" test_univ_decl LIST0 identref "|" LIST0 identref [ "+" | ] univ_decl_constraints +| WITH "@{" OPT [ LIST0 identref "|" ] LIST0 identref OPT "+" OPT [ "|" LIST0 univ_constraint SEP "," OPT "+" ] "}" +| DELETE "@{" LIST0 identref [ "+" | ] univ_decl_constraints ] cumul_univ_decl: [ -| REPLACE "@{" LIST0 variance_identref [ "+" | ] [ "|" LIST0 univ_constraint SEP "," [ "+" | ] "}" | [ "}" | bar_cbrace ] ] -| WITH "@{" LIST0 variance_identref OPT "+" OPT [ "|" LIST0 univ_constraint SEP "," OPT "+" ] "}" +| REPLACE "@{" test_cumul_univ_decl LIST0 identref "|" LIST0 variance_identref [ "+" | ] univ_decl_constraints +| WITH "@{" OPT [ LIST0 identref "|" ] LIST0 variance_identref OPT "+" OPT [ "|" LIST0 univ_constraint SEP "," OPT "+" ] "}" +| DELETE "@{" LIST0 variance_identref [ "+" | ] univ_decl_constraints ] of_type: [ diff --git a/doc/tools/docgram/fullGrammar b/doc/tools/docgram/fullGrammar index 62e0f216c1f8..41be5d4a3ec1 100644 --- a/doc/tools/docgram/fullGrammar +++ b/doc/tools/docgram/fullGrammar @@ -27,6 +27,7 @@ sort: [ | "SProp" | "Type" | "Type" "@{" "_" "}" +| "Type" "@{" test_sort_qvar reference "|" universe "}" | "Type" "@{" universe "}" ] @@ -182,12 +183,13 @@ evar_instance: [ ] univ_annot: [ -| "@{" LIST0 universe_level "}" +| "@{" LIST0 univ_level_or_quality OPT [ "|" LIST0 univ_level_or_quality ] "}" | ] -universe_level: [ +univ_level_or_quality: [ | "Set" +| "SProp" | "Prop" | "Type" | "_" @@ -866,8 +868,14 @@ univ_constraint: [ | universe_name [ "<" | "=" | "<=" ] universe_name ] +univ_decl_constraints: [ +| "|" LIST0 univ_constraint SEP "," [ "+" | ] "}" +| [ "}" | bar_cbrace ] +] + univ_decl: [ -| "@{" LIST0 identref [ "+" | ] [ "|" LIST0 univ_constraint SEP "," [ "+" | ] "}" | [ "}" | bar_cbrace ] ] +| "@{" test_univ_decl LIST0 identref "|" LIST0 identref [ "+" | ] univ_decl_constraints +| "@{" LIST0 identref [ "+" | ] univ_decl_constraints ] variance: [ @@ -882,7 +890,8 @@ variance_identref: [ ] cumul_univ_decl: [ -| "@{" LIST0 variance_identref [ "+" | ] [ "|" LIST0 univ_constraint SEP "," [ "+" | ] "}" | [ "}" | bar_cbrace ] ] +| "@{" test_cumul_univ_decl LIST0 identref "|" LIST0 variance_identref [ "+" | ] univ_decl_constraints +| "@{" LIST0 variance_identref [ "+" | ] univ_decl_constraints ] ident_decl: [ diff --git a/doc/tools/docgram/orderedGrammar b/doc/tools/docgram/orderedGrammar index 57883c19f583..12364c190c51 100644 --- a/doc/tools/docgram/orderedGrammar +++ b/doc/tools/docgram/orderedGrammar @@ -280,7 +280,7 @@ sort: [ | "SProp" | "Type" | "Type" "@{" "_" "}" -| "Type" "@{" universe "}" +| "Type" "@{" OPT [ qualid "|" ] universe "}" ] universe: [ @@ -299,11 +299,12 @@ universe_name: [ ] univ_annot: [ -| "@{" LIST0 universe_level "}" +| "@{" LIST0 univ_level_or_quality OPT [ "|" LIST0 univ_level_or_quality ] "}" ] -universe_level: [ +univ_level_or_quality: [ | "Set" +| "SProp" | "Prop" | "Type" | "_" @@ -311,17 +312,22 @@ universe_level: [ ] univ_decl: [ -| "@{" LIST0 ident OPT "+" OPT [ "|" LIST0 univ_constraint SEP "," OPT "+" ] "}" +| "@{" OPT [ LIST0 ident "|" ] LIST0 ident OPT "+" OPT [ "|" LIST0 univ_constraint SEP "," OPT "+" ] "}" ] cumul_univ_decl: [ -| "@{" LIST0 ( OPT [ "+" | "=" | "*" ] ident ) OPT "+" OPT [ "|" LIST0 univ_constraint SEP "," OPT "+" ] "}" +| "@{" OPT [ LIST0 ident "|" ] LIST0 ( OPT [ "+" | "=" | "*" ] ident ) OPT "+" OPT [ "|" LIST0 univ_constraint SEP "," OPT "+" ] "}" ] univ_constraint: [ | universe_name [ "<" | "=" | "<=" ] universe_name ] +univ_decl_constraints: [ +| "|" LIST0 univ_constraint SEP "," [ "+" | ] "}" +| [ "}" | "|}" ] +] + term_fix: [ | "let" "fix" fix_decl "in" term | "fix" fix_decl OPT ( LIST1 ( "with" fix_decl ) "for" ident ) diff --git a/engine/eConstr.ml b/engine/eConstr.ml index e901168731db..7169f59eda56 100644 --- a/engine/eConstr.ml +++ b/engine/eConstr.ml @@ -41,13 +41,15 @@ module ESorts = struct let family sigma s = Sorts.family (kind sigma s) + let quality sigma s = Sorts.quality (kind sigma s) + end module EInstance = struct include Evd.MiniEConstr.EInstance let equal sigma i1 i2 = - Univ.Instance.equal (kind sigma i1) (kind sigma i2) + UVars.Instance.equal (kind sigma i1) (kind sigma i2) end include (Evd.MiniEConstr : module type of Evd.MiniEConstr @@ -97,7 +99,7 @@ let mkConstructUi ((ind,u),i) = of_kind (Construct ((ind,i),u)) let mkCase (ci, u, pms, c, iv, r, p) = of_kind (Case (ci, u, pms, c, iv, 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 mkProj (p, r, c) = of_kind (Proj (p, r, c)) 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) @@ -230,7 +232,7 @@ let destCase sigma c = match kind sigma c with | _ -> raise DestKO let destProj sigma c = match kind sigma c with -| Proj (p, c) -> (p, c) +| Proj (p, r, c) -> (p, r, c) | _ -> raise DestKO let destRef sigma c = let open GlobRef in match kind sigma c with @@ -498,7 +500,7 @@ let iter_with_full_binders env sigma g f n c = let (ci, _, pms, p, iv, c, bl) = annotate_case env sigma (ci, u, pms, p, iv, c, bl) in let f_ctx (ctx, c) = f (List.fold_right g ctx n) c in Array.Fun1.iter f n pms; f_ctx p; iter_invert (f n) iv; f n c; Array.iter f_ctx bl - | Proj (p,c) -> f n c + | Proj (_,_,c) -> f n c | Fix (_,(lna,tl,bl)) -> Array.iter (f n) tl; let n' = Array.fold_left2_i (fun i n na t -> g (LocalAssum (na, lift i t)) n) n lna tl in @@ -572,31 +574,14 @@ let compare_constr sigma cmp c1 c2 = let cmp nargs c1 c2 = cmp c1 c2 in compare_gen kind eq_inst eq_sorts (eq_existential cmp) cmp 0 c1 c2 -let compare_cumulative_instances cv_pb nargs_ok variances u u' cstrs = - let open UnivProblem in - if not nargs_ok then enforce_eq_instances_univs false u u' cstrs - else - let make u = Sorts.sort_of_univ @@ Univ.Universe.make u in - CArray.fold_left3 - (fun cstrs v u u' -> - let open Univ.Variance in - match v with - | Irrelevant -> Set.add (UWeak (u,u')) cstrs - | Covariant -> - (match cv_pb with - | Conversion.CONV -> Set.add (UEq (make u, make u')) cstrs - | Conversion.CUMUL -> Set.add (ULe (make u, make u')) cstrs) - | Invariant -> - Set.add (UEq (make u, make u')) cstrs) - cstrs variances (Univ.Instance.to_array u) (Univ.Instance.to_array u') - let cmp_inductives cv_pb (mind,ind as spec) nargs u1 u2 cstrs = let open UnivProblem in match mind.Declarations.mind_variance with | None -> enforce_eq_instances_univs false u1 u2 cstrs | Some variances -> let num_param_arity = Conversion.inductive_cumulativity_arguments spec in - compare_cumulative_instances cv_pb (Int.equal num_param_arity nargs) variances u1 u2 cstrs + if not (Int.equal num_param_arity nargs) then enforce_eq_instances_univs false u1 u2 cstrs + else compare_cumulative_instances cv_pb variances u1 u2 cstrs let cmp_constructors (mind, ind, cns as spec) nargs u1 u2 cstrs = let open UnivProblem in @@ -607,8 +592,11 @@ let cmp_constructors (mind, ind, cns as spec) nargs u1 u2 cstrs = if not (Int.equal num_cnstr_args nargs) then enforce_eq_instances_univs false u1 u2 cstrs else + let qs1, us1 = UVars.Instance.to_array u1 + and qs2, us2 = UVars.Instance.to_array u2 in + let cstrs = enforce_eq_qualities qs1 qs2 cstrs in Array.fold_left2 (fun cstrs u1 u2 -> UnivProblem.(Set.add (UWeak (u1,u2)) cstrs)) - cstrs (Univ.Instance.to_array u1) (Univ.Instance.to_array u2) + cstrs us1 us2 let eq_universes env sigma cstrs cv_pb refargs l l' = if EInstance.is_empty l then (assert (EInstance.is_empty l'); true) @@ -619,7 +607,7 @@ let eq_universes env sigma cstrs cv_pb refargs l l' = let open UnivProblem in match refargs with | Some (ConstRef c, 1) when Environ.is_array_type env c -> - cstrs := compare_cumulative_instances cv_pb true [|Univ.Variance.Irrelevant|] l l' !cstrs; + cstrs := compare_cumulative_instances cv_pb [|UVars.Variance.Irrelevant|] l l' !cstrs; true | None | Some (ConstRef _, _) -> cstrs := enforce_eq_instances_univs true l l' !cstrs; true @@ -681,8 +669,8 @@ let leq_constr_universes env sigma ?nargs m n = let compare_head_gen_proj env sigma equ eqs eqev eqc' nargs m n = let kind c = kind sigma c in match kind m, kind n with - | Proj (p, c), App (f, args) - | App (f, args), Proj (p, c) -> + | Proj (p, _, c), App (f, args) + | App (f, args), Proj (p, _, c) -> (match kind f with | Const (p', u) when Environ.QConstant.equal env (Projection.constant p) p' -> let npars = Projection.npars p in @@ -714,30 +702,77 @@ let eq_constr_universes_proj env sigma m n = let res = eq_constr' 0 m n in if res then Some !cstrs else None +let add_universes_of_instance sigma (qs,us) u = + let u = EInstance.kind sigma u in + let qs', us' = UVars.Instance.levels u in + let qs = Sorts.Quality.(Set.fold (fun q qs -> match q with + | QVar q -> Sorts.QVar.Set.add q qs + | QConstant _ -> qs) + qs' qs) + in + qs, Univ.Level.Set.union us us' + +let fold_annot_relevance f acc na = + f acc na.Context.binder_relevance + +let fold_case_under_context_relevance f acc (nas,_) = + Array.fold_left (fold_annot_relevance f) acc nas + +let fold_rec_declaration_relevance f acc (nas,_,_) = + Array.fold_left (fold_annot_relevance f) acc nas + +let fold_constr_relevance sigma f acc c = + match kind sigma c with + | Rel _ | Var _ | Meta _ | Evar _ + | Sort _ | Cast _ | App _ + | Const _ | Ind _ | Construct _ | Proj _ + | Int _ | Float _ | Array _ -> acc + + | Prod (na,_,_) | Lambda (na,_,_) | LetIn (na,_,_,_) -> + fold_annot_relevance f acc na + + | Case (ci,_u,_params,ret,_iv,_v,brs) -> + let acc = f acc ci.ci_relevance in + let acc = fold_case_under_context_relevance f acc ret in + let acc = CArray.fold_left (fold_case_under_context_relevance f) acc brs in + acc + + | Fix (_,data) + | CoFix (_,data) -> + fold_rec_declaration_relevance f acc data + +let add_relevance sigma (qs,us as v) r = + let open Sorts in + (* NB this normalizes above_prop to Relevant which makes it disappear *) + match UState.nf_relevance (Evd.evar_universe_context sigma) r with + | Irrelevant | Relevant -> v + | RelevanceVar q -> QVar.Set.add q qs, us + let universes_of_constr sigma c = let open Univ in let rec aux s c = + let s = fold_constr_relevance sigma (add_relevance sigma) s c in match kind sigma c with - | Const (c, u) -> - Level.Set.fold Level.Set.add (Instance.levels (EInstance.kind sigma u)) s - | Ind ((mind,_), u) | Construct (((mind,_),_), u) -> - Level.Set.fold Level.Set.add (Instance.levels (EInstance.kind sigma u)) s - | Sort u -> - let sort = ESorts.kind sigma u in - if Sorts.is_small sort then s - else - Level.Set.fold Level.Set.add (Sorts.levels sort) s + | Const (_, u) | Ind (_, u) | Construct (_,u) -> add_universes_of_instance sigma s u + | Sort u -> begin match ESorts.kind sigma u with + | Type u -> + Util.on_snd (Level.Set.fold Level.Set.add (Universe.levels u)) s + | QSort (q, u) -> + let qs, us = s in + Sorts.QVar.Set.add q qs, Level.Set.union us (Universe.levels u) + | SProp | Prop | Set -> s + end | Evar (k, args) -> let concl = Evd.evar_concl (Evd.find_undefined sigma k) in fold sigma aux (aux s concl) c | Array (u,_,_,_) -> - let s = Level.Set.fold Level.Set.add (Instance.levels (EInstance.kind sigma u)) s in + let s = add_universes_of_instance sigma s u in fold sigma aux s c | Case (_,u,_,_,_,_,_) -> - let s = Level.Set.fold Level.Set.add (Instance.levels (EInstance.kind sigma u)) s in + let s = add_universes_of_instance sigma s u in fold sigma aux s c | _ -> fold sigma aux s c - in aux Level.Set.empty c + in aux (Sorts.QVar.Set.empty,Level.Set.empty) c open Context open Environ diff --git a/engine/eConstr.mli b/engine/eConstr.mli index fcfd73d4fb35..568aa43a49bd 100644 --- a/engine/eConstr.mli +++ b/engine/eConstr.mli @@ -66,16 +66,18 @@ sig val family : Evd.evar_map -> t -> Sorts.family + val quality : Evd.evar_map -> t -> Sorts.Quality.t + end module EInstance : sig type t (** Type of universe instances up-to universe unification. Similar to - [ESorts.t] for [Univ.Instance.t]. *) + [ESorts.t] for [UVars.Instance.t]. *) - val make : Univ.Instance.t -> t - val kind : Evd.evar_map -> t -> Univ.Instance.t + val make : UVars.Instance.t -> t + val kind : Evd.evar_map -> t -> UVars.Instance.t val empty : t val is_empty : t -> bool end @@ -91,7 +93,7 @@ val kind : Evd.evar_map -> t -> (t, t, ESorts.t, EInstance.t) Constr.kind_of_ter (** Same as {!Constr.kind} except that it expands evars and normalizes universes on the fly. *) -val kind_upto : Evd.evar_map -> Constr.t -> (Constr.t, Constr.t, Sorts.t, Univ.Instance.t) Constr.kind_of_term +val kind_upto : Evd.evar_map -> Constr.t -> (Constr.t, Constr.t, Sorts.t, UVars.Instance.t) Constr.kind_of_term val to_constr : ?abort_on_undefined_evars:bool -> Evd.evar_map -> t -> Constr.t (** Returns the evar-normal form of the argument. Note that this @@ -146,7 +148,7 @@ 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 mkConstU : Constant.t * EInstance.t -> t -val mkProj : (Projection.t * t) -> t +val mkProj : (Projection.t * Sorts.relevance * t) -> t val mkIndU : inductive * EInstance.t -> t val mkConstructU : constructor * EInstance.t -> t val mkConstructUi : (inductive * EInstance.t) * int -> t @@ -266,7 +268,7 @@ val destEvar : Evd.evar_map -> t -> t pexistential val destInd : Evd.evar_map -> t -> inductive * EInstance.t val destConstruct : Evd.evar_map -> t -> constructor * EInstance.t val destCase : Evd.evar_map -> t -> case -val destProj : Evd.evar_map -> t -> Projection.t * t +val destProj : Evd.evar_map -> t -> Projection.t * Sorts.relevance * t val destFix : Evd.evar_map -> t -> (t, t) pfixpoint val destCoFix : Evd.evar_map -> t -> (t, t) pcofixpoint @@ -334,7 +336,7 @@ val fold_with_binders : Evd.evar_map -> ('a -> 'a) -> ('a -> 'b -> t -> 'b) -> ' (** Gather the universes transitively used in the term, including in the type of evars appearing in it. *) -val universes_of_constr : Evd.evar_map -> t -> Univ.Level.Set.t +val universes_of_constr : Evd.evar_map -> t -> Sorts.QVar.Set.t * Univ.Level.Set.t (** {6 Substitutions} *) @@ -368,9 +370,9 @@ val noccur_between : Evd.evar_map -> int -> int -> t -> bool val closedn : Evd.evar_map -> int -> t -> bool val closed0 : Evd.evar_map -> t -> bool -val subst_univs_level_constr : Univ.universe_level_subst -> t -> t -val subst_instance_context : Univ.Instance.t -> rel_context -> rel_context -val subst_instance_constr : Univ.Instance.t -> t -> t +val subst_univs_level_constr : UVars.sort_level_subst -> t -> t +val subst_instance_context : UVars.Instance.t -> rel_context -> rel_context +val subst_instance_constr : UVars.Instance.t -> t -> t val subst_of_rel_context_instance : rel_context -> instance -> substl val subst_of_rel_context_instance_list : rel_context -> instance_list -> substl @@ -419,7 +421,7 @@ val identity_subst_val : named_context_val -> t SList.t (* XXX Missing Sigma proxy *) val fresh_global : - ?loc:Loc.t -> ?rigid:Evd.rigid -> ?names:Univ.Instance.t -> Environ.env -> + ?loc:Loc.t -> ?rigid:Evd.rigid -> ?names:UVars.Instance.t -> Environ.env -> Evd.evar_map -> GlobRef.t -> Evd.evar_map * t val is_global : Environ.env -> Evd.evar_map -> GlobRef.t -> t -> bool @@ -478,7 +480,7 @@ sig val to_sorts : ESorts.t -> Sorts.t (** Physical identity. Does not care for normalization. *) - val to_instance : EInstance.t -> Univ.Instance.t + val to_instance : EInstance.t -> UVars.Instance.t (** Physical identity. Does not care for normalization. *) val to_case_invert : case_invert -> Constr.case_invert diff --git a/engine/evarutil.ml b/engine/evarutil.ml index b042164ca523..c9932039e6f1 100644 --- a/engine/evarutil.ml +++ b/engine/evarutil.ml @@ -35,7 +35,7 @@ let finalize ?abort_on_undefined_evars sigma f = let sigma = minimize_universes sigma in let uvars = ref Univ.Level.Set.empty in let v = f (fun c -> - let varsc = EConstr.universes_of_constr sigma c in + let _, varsc = EConstr.universes_of_constr sigma c in let c = EConstr.to_constr ?abort_on_undefined_evars sigma c in uvars := Univ.Level.Set.union !uvars varsc; c) @@ -67,12 +67,9 @@ let whd_evar = EConstr.whd_evar let nf_evar sigma c = let lsubst = Evd.universe_subst sigma in let evar_value ev = Evd.existential_opt_value0 sigma ev in - let level_value l = - UnivSubst.level_subst_of (fun l -> UnivFlex.normalize_univ_variable lsubst l) l - in - let sort_value s = UState.nf_sort (Evd.evar_universe_context sigma) s in - let rel_value r = UState.nf_relevance (Evd.evar_universe_context sigma) r in - EConstr.of_constr @@ UnivSubst.nf_evars_and_universes_opt_subst evar_value level_value sort_value rel_value (EConstr.Unsafe.to_constr c) + let univ_value l = UnivFlex.normalize_univ_variable lsubst l in + let qvar_value q = UState.nf_qvar (Evd.evar_universe_context sigma) q in + EConstr.of_constr @@ UnivSubst.nf_evars_and_universes_opt_subst evar_value qvar_value univ_value (EConstr.Unsafe.to_constr c) let j_nf_evar sigma j = { uj_val = nf_evar sigma j.uj_val; @@ -141,7 +138,7 @@ let head_evar sigma c = | Case (_, _, _, _, _, c, _) -> hrec c | App (c,_) -> hrec c | Cast (c,_,_) -> hrec c - | Proj (p, c) -> hrec c + | Proj (_, _, c) -> hrec c | _ -> raise NoHeadEvar in hrec c @@ -764,14 +761,20 @@ let compare_cumulative_instances cv_pb variances u u' sigma = let open UnivProblem in let cstrs = Univ.Constraints.empty in let soft = Set.empty in + let qs, us = UVars.Instance.to_array u + and qs', us' = UVars.Instance.to_array u' in + let qcstrs = enforce_eq_qualities qs qs' Set.empty in + match Evd.add_universe_constraints sigma qcstrs with + | exception UGraph.UniverseInconsistency p -> Inr p + | sigma -> let cstrs, soft = Array.fold_left3 (fun (cstrs, soft) v u u' -> - let open Univ.Variance in + let open UVars.Variance in match v with | Irrelevant -> cstrs, Set.add (UWeak (u,u')) soft | Covariant when cv_pb == Conversion.CUMUL -> Univ.Constraints.add (u,Univ.Le,u') cstrs, soft | Covariant | Invariant -> Univ.Constraints.add (u,Univ.Eq,u') cstrs, soft) - (cstrs,soft) variances (Univ.Instance.to_array u) (Univ.Instance.to_array u') + (cstrs,soft) variances us us' in match Evd.add_constraints sigma cstrs with | sigma -> @@ -780,11 +783,17 @@ let compare_cumulative_instances cv_pb variances u u' sigma = let compare_constructor_instances evd u u' = let open UnivProblem in - let soft = - Array.fold_left2 (fun cs u u' -> Set.add (UWeak (u,u')) cs) - Set.empty (Univ.Instance.to_array u) (Univ.Instance.to_array u') - in - Evd.add_universe_constraints evd soft + let qs, us = UVars.Instance.to_array u + and qs', us' = UVars.Instance.to_array u' in + let qcstrs = enforce_eq_qualities qs qs' Set.empty in + match Evd.add_universe_constraints evd qcstrs with + | exception UGraph.UniverseInconsistency p -> Inr p + | evd -> + let soft = + Array.fold_left2 (fun cs u u' -> Set.add (UWeak (u,u')) cs) + Set.empty us us' + in + Inl (Evd.add_universe_constraints evd soft) (** [eq_constr_univs_test ~evd ~extended_evd t u] tests equality of [t] and [u] up to existential variable instantiation and diff --git a/engine/evarutil.mli b/engine/evarutil.mli index 276873881b7d..ded6676b8447 100644 --- a/engine/evarutil.mli +++ b/engine/evarutil.mli @@ -170,7 +170,7 @@ val finalize : ?abort_on_undefined_evars:bool -> evar_map -> as an evar [e] only if [e] is uninstantiated in [sigma]. Otherwise the value of [e] in [sigma] is (recursively) used. *) val kind_of_term_upto : evar_map -> Constr.constr -> - (Constr.constr, Constr.types, Sorts.t, Univ.Instance.t) kind_of_term + (Constr.constr, Constr.types, Sorts.t, UVars.Instance.t) kind_of_term (** [eq_constr_univs_test ~evd ~extended_evd t u] tests equality of [t] and [u] up to existential variable instantiation and @@ -189,14 +189,17 @@ val eq_constr_univs_test : constraints such that [u1 cv_pb? u2] according to [variance]. Additionally flexible universes in irrelevant positions are unified if possible. Returns [Inr p] when the former is impossible. *) -val compare_cumulative_instances : Conversion.conv_pb -> Univ.Variance.t array -> - Univ.Instance.t -> Univ.Instance.t -> evar_map -> +val compare_cumulative_instances : Conversion.conv_pb -> UVars.Variance.t array -> + UVars.Instance.t -> UVars.Instance.t -> evar_map -> (evar_map, UGraph.univ_inconsistency) Util.union (** We should only compare constructors at convertible types, so this - is only an opportunity to unify universes. *) + is only an opportunity to unify universes. + + But what about qualities? +*) val compare_constructor_instances : evar_map -> - Univ.Instance.t -> Univ.Instance.t -> evar_map + UVars.Instance.t -> UVars.Instance.t -> (evar_map, UGraph.univ_inconsistency) Util.union (** {6 Unification problems} *) type unification_pb = conv_pb * env * constr * constr diff --git a/engine/evd.ml b/engine/evd.ml index dfe4e337ee95..79db14f37e76 100644 --- a/engine/evd.ml +++ b/engine/evd.ml @@ -10,7 +10,6 @@ open Pp open CErrors -open Sorts open Util open Names open Nameops @@ -913,8 +912,8 @@ let mkLEvar sigma (evk, args) = let evar_handler sigma = let evar_expand ev = existential_expand_value0 sigma ev in let qvar_relevant q = match UState.nf_qvar sigma.universes q with - | QSProp -> false - | QProp | QType | QVar _ -> true + | QConstant QSProp -> false + | QConstant QProp | QConstant QType | QVar _ -> true in let evar_relevant (evk, _) = match find sigma evk with | EvarInfo evi -> @@ -943,6 +942,9 @@ let existential_type0 = existential_type let add_constraints d c = { d with universes = UState.add_constraints d.universes c } +let add_quconstraints d c = + { d with universes = UState.add_quconstraints d.universes c } + let add_universe_constraints d c = { d with universes = UState.add_universe_constraints d.universes c } @@ -1084,6 +1086,8 @@ let evar_universe_context d = d.universes let universe_context_set d = UState.context_set d.universes +let sort_context_set d = UState.sort_context_set d.universes + let to_universe_context evd = UState.context evd.universes let univ_entry ~poly evd = UState.univ_entry ~poly evd.universes @@ -1099,9 +1103,18 @@ let universe_subst evd = let merge_context_set ?loc ?(sideff=false) rigid evd uctx' = {evd with universes = UState.merge ?loc ~sideff rigid evd.universes uctx'} +let merge_sort_context_set ?loc ?(sideff=false) rigid evd ctx' = + {evd with universes = UState.merge_sort_context ?loc ~sideff rigid evd.universes ctx'} + +let merge_sort_variables ?loc ?(sideff=false) evd qs = + { evd with universes = UState.merge_sort_variables ?loc ~sideff evd.universes qs } + let with_context_set ?loc rigid evd (a, uctx) = (merge_context_set ?loc rigid evd uctx, a) +let with_sort_context_set ?loc rigid d (a, ctx) = + (merge_sort_context_set ?loc rigid d ctx, a) + let new_univ_level_variable ?loc ?name rigid evd = let uctx', u = UState.new_univ_variable ?loc rigid name evd.universes in ({evd with universes = uctx'}, u) @@ -1110,8 +1123,12 @@ let new_univ_variable ?loc ?name rigid evd = let uctx', u = UState.new_univ_variable ?loc rigid name evd.universes in ({evd with universes = uctx'}, Univ.Universe.make u) -let new_sort_variable ?loc ?name rigid sigma = - let (sigma, u) = new_univ_variable ?loc rigid ?name sigma in +let new_quality_variable ?loc ?name evd = + let uctx, q = UState.new_sort_variable ?loc ?name evd.universes in + {evd with universes = uctx}, q + +let new_sort_variable ?loc rigid sigma = + let (sigma, u) = new_univ_variable ?loc rigid sigma in let uctx, q = UState.new_sort_variable sigma.universes in ({ sigma with universes = uctx }, Sorts.qsort q u) @@ -1126,22 +1143,22 @@ let make_nonalgebraic_variable evd u = (****************************************) let fresh_sort_in_family ?loc ?(rigid=univ_flexible) evd s = - with_context_set ?loc rigid evd (UnivGen.fresh_sort_in_family s) + with_sort_context_set ?loc rigid evd (UnivGen.fresh_sort_in_family s) let fresh_constant_instance ?loc ?(rigid=univ_flexible) env evd c = - with_context_set ?loc rigid evd (UnivGen.fresh_constant_instance env c) + with_sort_context_set ?loc rigid evd (UnivGen.fresh_constant_instance env c) let fresh_inductive_instance ?loc ?(rigid=univ_flexible) env evd i = - with_context_set ?loc rigid evd (UnivGen.fresh_inductive_instance env i) + with_sort_context_set ?loc rigid evd (UnivGen.fresh_inductive_instance env i) let fresh_constructor_instance ?loc ?(rigid=univ_flexible) env evd c = - with_context_set ?loc rigid evd (UnivGen.fresh_constructor_instance env c) + with_sort_context_set ?loc rigid evd (UnivGen.fresh_constructor_instance env c) let fresh_array_instance ?loc ?(rigid=univ_flexible) env evd = - with_context_set ?loc rigid evd (UnivGen.fresh_array_instance env) + with_sort_context_set ?loc rigid evd (UnivGen.fresh_array_instance env) let fresh_global ?loc ?(rigid=univ_flexible) ?names env evd gr = - with_context_set ?loc rigid evd (UnivGen.fresh_global_instance ?loc ?names env gr) + with_sort_context_set ?loc rigid evd (UnivGen.fresh_global_instance ?loc ?names env gr) let is_flexible_level evd l = let uctx = evd.universes in @@ -1159,33 +1176,11 @@ let universe_rigidity evd l = UnivFlexible (UState.is_algebraic l uctx) else UnivRigid -let normalize_universe evd = - let vars = UState.subst evd.universes in - let normalize = UnivFlex.normalize_universe vars in - normalize - let normalize_universe_instance evd l = - let vars = UState.subst evd.universes in - let normalize = UnivSubst.level_subst_of (UnivFlex.normalize_univ_variable vars) in - UnivSubst.subst_instance normalize l + UState.nf_instance evd.universes l let normalize_sort evars s = - match s with - | SProp | Prop | Set -> s - | Type u -> - let u' = normalize_universe evars u in - if u' == u then s else Sorts.sort_of_univ u' - | QSort (q, u) -> - begin match UState.nf_qvar evars.universes q with - | QProp -> Sorts.prop - | QSProp -> Sorts.sprop - | QType -> - let u' = normalize_universe evars u in - Sorts.sort_of_univ u' - | QVar q' -> - let u' = normalize_universe evars u in - if q' == q && u' == u then s else Sorts.qsort q' u' - end + UState.nf_sort evars.universes s (* FIXME inefficient *) let set_eq_sort env d s1 s2 = @@ -1230,6 +1225,12 @@ let check_leq evd s s' = let check_constraints evd csts = UGraph.check_constraints csts (UState.ugraph evd.universes) +let check_qconstraints evd csts = + UState.check_qconstraints evd.universes csts + +let check_quconstraints evd (qcsts,ucsts) = + check_qconstraints evd qcsts && check_constraints evd ucsts + let fix_undefined_variables evd = { evd with universes = UState.fix_undefined_variables evd.universes } @@ -1249,6 +1250,8 @@ let minimize_universes evd = let universe_of_name evd s = UState.universe_of_name evd.universes s +let quality_of_name evd s = UState.quality_of_name evd.universes s + let universe_binders evd = UState.universe_binders evd.universes let universes evd = UState.ugraph evd.universes @@ -1639,13 +1642,13 @@ module MiniEConstr = struct module EInstance = struct - type t = Univ.Instance.t + type t = UVars.Instance.t let make i = i let kind sigma i = - if Univ.Instance.is_empty i then i + if UVars.Instance.is_empty i then i else normalize_universe_instance sigma i - let empty = Univ.Instance.empty - let is_empty = Univ.Instance.is_empty + let empty = UVars.Instance.empty + let is_empty = UVars.Instance.is_empty let unsafe_to_instance t = t end @@ -1703,12 +1706,11 @@ module MiniEConstr = struct Some (instantiate_evar_array sigma info c args) in let lsubst = universe_subst sigma in - let level_value l = - UnivSubst.level_subst_of (fun l -> UnivFlex.normalize_univ_variable lsubst l) l + let univ_value l = + UnivFlex.normalize_univ_variable lsubst l in - let sort_value s = UState.nf_sort (evar_universe_context sigma) s in - let rel_value r = UState.nf_relevance (evar_universe_context sigma) r in - UnivSubst.nf_evars_and_universes_opt_subst evar_value level_value sort_value rel_value c + let qvar_value q = UState.nf_qvar sigma.universes q in + UnivSubst.nf_evars_and_universes_opt_subst evar_value qvar_value univ_value c let to_constr_gen sigma c = let saw_evar = ref false in @@ -1718,12 +1720,11 @@ module MiniEConstr = struct v in let lsubst = universe_subst sigma in - let level_value l = - UnivSubst.level_subst_of (fun l -> UnivFlex.normalize_univ_variable lsubst l) l + let univ_value l = + UnivFlex.normalize_univ_variable lsubst l in - let sort_value s = UState.nf_sort (evar_universe_context sigma) s in - let rel_value r = UState.nf_relevance (evar_universe_context sigma) r in - let c = UnivSubst.nf_evars_and_universes_opt_subst evar_value level_value sort_value rel_value c in + let qvar_value q = UState.nf_qvar sigma.universes q in + let c = UnivSubst.nf_evars_and_universes_opt_subst evar_value qvar_value univ_value c in let saw_evar = if not !saw_evar then false else let exception SawEvar in diff --git a/engine/evd.mli b/engine/evd.mli index 5cf3ec18cf8f..6b44ff233398 100644 --- a/engine/evd.mli +++ b/engine/evd.mli @@ -271,6 +271,8 @@ val is_undefined : evar_map -> Evar.t-> bool val add_constraints : evar_map -> Univ.Constraints.t -> evar_map (** Add universe constraints in an evar map. *) +val add_quconstraints : evar_map -> Sorts.QUConstraints.t -> evar_map + val undefined_map : evar_map -> undefined evar_info Evar.Map.t (** Access the undefined evar mapping directly. *) @@ -639,11 +641,13 @@ val restrict_universe_context : evar_map -> Univ.Level.Set.t -> evar_map (** Raises Not_found if not a name for a universe in this map. *) val universe_of_name : evar_map -> Id.t -> Univ.Level.t +val quality_of_name : evar_map -> Id.t -> Sorts.QVar.t val universe_binders : evar_map -> UnivNames.universe_binders val new_univ_level_variable : ?loc:Loc.t -> ?name:Id.t -> rigid -> evar_map -> evar_map * Univ.Level.t -val new_sort_variable : ?loc:Loc.t -> ?name:Id.t -> rigid -> evar_map -> evar_map * esorts +val new_quality_variable : ?loc:Loc.t -> ?name:Id.t -> evar_map -> evar_map * Sorts.QVar.t +val new_sort_variable : ?loc:Loc.t -> rigid -> evar_map -> evar_map * esorts val add_global_univ : evar_map -> Univ.Level.t -> evar_map @@ -654,29 +658,32 @@ val make_nonalgebraic_variable : evar_map -> Univ.Level.t -> evar_map val is_flexible_level : evar_map -> Univ.Level.t -> bool -val normalize_universe_instance : evar_map -> Univ.Instance.t -> Univ.Instance.t +val normalize_universe_instance : evar_map -> UVars.Instance.t -> UVars.Instance.t val set_leq_sort : env -> evar_map -> esorts -> esorts -> evar_map val set_eq_sort : env -> evar_map -> esorts -> esorts -> evar_map val set_eq_level : evar_map -> Univ.Level.t -> Univ.Level.t -> evar_map val set_leq_level : evar_map -> Univ.Level.t -> Univ.Level.t -> evar_map val set_eq_instances : ?flex:bool -> - evar_map -> Univ.Instance.t -> Univ.Instance.t -> evar_map + evar_map -> UVars.Instance.t -> UVars.Instance.t -> evar_map val check_eq : evar_map -> esorts -> esorts -> bool val check_leq : evar_map -> esorts -> esorts -> bool val check_constraints : evar_map -> Univ.Constraints.t -> bool +val check_qconstraints : evar_map -> Sorts.QConstraints.t -> bool +val check_quconstraints : evar_map -> Sorts.QUConstraints.t -> bool val evar_universe_context : evar_map -> UState.t val universe_context_set : evar_map -> Univ.ContextSet.t +val sort_context_set : evar_map -> UnivGen.sort_context_set val universe_subst : evar_map -> UnivFlex.t val universes : evar_map -> UGraph.t (** [to_universe_context evm] extracts the local universes and constraints of [evm] and orders the universes the same as [Univ.ContextSet.to_context]. *) -val to_universe_context : evar_map -> Univ.UContext.t +val to_universe_context : evar_map -> UVars.UContext.t val univ_entry : poly:bool -> evar_map -> UState.named_universes_entry @@ -687,8 +694,14 @@ val set_universe_context : evar_map -> UState.t -> evar_map val merge_context_set : ?loc:Loc.t -> ?sideff:bool -> rigid -> evar_map -> Univ.ContextSet.t -> evar_map +val merge_sort_context_set : ?loc:Loc.t -> ?sideff:bool -> rigid -> evar_map -> UnivGen.sort_context_set -> evar_map + +val merge_sort_variables : ?loc:Loc.t -> ?sideff:bool -> evar_map -> Sorts.QVar.Set.t -> evar_map + val with_context_set : ?loc:Loc.t -> rigid -> evar_map -> 'a Univ.in_universe_context_set -> evar_map * 'a +val with_sort_context_set : ?loc:Loc.t -> rigid -> evar_map -> 'a UnivGen.in_sort_context_set -> evar_map * 'a + val nf_univ_variables : evar_map -> evar_map val collapse_sort_variables : evar_map -> evar_map @@ -712,9 +725,9 @@ val fresh_inductive_instance : ?loc:Loc.t -> ?rigid:rigid val fresh_constructor_instance : ?loc:Loc.t -> ?rigid:rigid -> env -> evar_map -> constructor -> evar_map * pconstructor val fresh_array_instance : ?loc:Loc.t -> ?rigid:rigid - -> env -> evar_map -> evar_map * Univ.Instance.t + -> env -> evar_map -> evar_map * UVars.Instance.t -val fresh_global : ?loc:Loc.t -> ?rigid:rigid -> ?names:Univ.Instance.t -> env -> +val fresh_global : ?loc:Loc.t -> ?rigid:rigid -> ?names:UVars.Instance.t -> env -> evar_map -> GlobRef.t -> evar_map * econstr (********************************************************************) @@ -752,17 +765,17 @@ module MiniEConstr : sig module EInstance : sig type t - val make : Univ.Instance.t -> t - val kind : evar_map -> t -> Univ.Instance.t + val make : UVars.Instance.t -> t + val kind : evar_map -> t -> UVars.Instance.t val empty : t val is_empty : t -> bool - val unsafe_to_instance : t -> Univ.Instance.t + val unsafe_to_instance : t -> UVars.Instance.t end type t = econstr val kind : evar_map -> t -> (t, t, ESorts.t, EInstance.t) Constr.kind_of_term - val kind_upto : evar_map -> constr -> (constr, types, Sorts.t, Univ.Instance.t) Constr.kind_of_term + val kind_upto : evar_map -> constr -> (constr, types, Sorts.t, UVars.Instance.t) Constr.kind_of_term val whd_evar : evar_map -> t -> t diff --git a/engine/namegen.ml b/engine/namegen.ml index cd4492974c8b..b5fb31840b4b 100644 --- a/engine/namegen.ml +++ b/engine/namegen.ml @@ -118,7 +118,7 @@ let head_name sigma c = (* Find the head constant of a constr if any *) match EConstr.kind sigma c with | Prod (_,_,c) | Lambda (_,_,c) | LetIn (_,_,_,c) | Cast (c,_,_) | App (c,_) -> hdrec c - | Proj (kn,_) -> Some (Label.to_id (Constant.label (Projection.constant kn))) + | Proj (kn,_,_) -> Some (Label.to_id (Constant.label (Projection.constant kn))) | Const _ | Ind _ | Construct _ | Var _ as c -> Some (Nametab.basename_of_global (global_of_constr c)) | Fix ((_,i),(lna,_,_)) | CoFix (i,(lna,_,_)) -> @@ -151,7 +151,7 @@ let hdchar env sigma c = match EConstr.kind sigma c with | Prod (_,_,c) | Lambda (_,_,c) | LetIn (_,_,_,c) -> hdrec (k+1) c | Cast (c,_,_) | App (c,_) -> hdrec k c - | Proj (kn,_) -> lowercase_first_char (Label.to_id (Constant.label (Projection.constant kn))) + | Proj (kn,_,_) -> lowercase_first_char (Label.to_id (Constant.label (Projection.constant kn))) | Const (kn,_) -> lowercase_first_char (Label.to_id (Constant.label kn)) | Ind (x,_) -> (try lowercase_first_char (Nametab.basename_of_global (GlobRef.IndRef x)) with Not_found when !Flags.in_debugger -> "zz") | Construct (x,_) -> (try lowercase_first_char (Nametab.basename_of_global (GlobRef.ConstructRef x)) with Not_found when !Flags.in_debugger -> "zz") diff --git a/engine/termops.ml b/engine/termops.ml index 4ec037e89aec..f69ed73a883b 100644 --- a/engine/termops.ml +++ b/engine/termops.ml @@ -296,6 +296,9 @@ let has_no_evar sigma = with Exit -> false let pr_evd_level sigma = UState.pr_uctx_level (Evd.evar_universe_context sigma) + +let pr_evd_qvar sigma = UState.pr_uctx_qvar (Evd.evar_universe_context sigma) + let reference_of_level sigma l = UState.qualid_of_level (Evd.evar_universe_context sigma) l let pr_evar_universe_context ctx = @@ -639,10 +642,10 @@ let map_constr_with_binders_left_to_right env sigma g f l c = let a' = f l a in if app' == app && a' == a then c else mkApp (app', [| a' |]) - | Proj (p,b) -> + | Proj (p,r,b) -> let b' = f l b in if b' == b then c - else mkProj (p, b') + else mkProj (p, r, b') | Evar ev -> let ev' = EConstr.map_existential sigma (fun c -> f l c) ev in if ev' == ev then c else mkEvar ev' @@ -706,9 +709,9 @@ let map_constr_with_full_binders env sigma g f l cstr = let c' = f l c in let al' = Array.map (f l) al in if c==c' && Array.for_all2 (==) al al' then cstr else mkApp (c', al') - | Proj (p,c) -> + | Proj (p,r,c) -> let c' = f l c in - if c' == c then cstr else mkProj (p, c') + if c' == c then cstr else mkProj (p, r, c') | Evar ev -> let ev' = EConstr.map_existential sigma (fun c -> f l c) ev in if ev' == ev then cstr else mkEvar ev' @@ -762,7 +765,7 @@ let fold_constr_with_full_binders env sigma g f n acc c = | Lambda (na,t,c) -> f (g (LocalAssum (na,t)) n) (f n acc t) c | LetIn (na,b,t,c) -> f (g (LocalDef (na,b,t)) n) (f n (f n acc b) t) c | App (c,l) -> Array.fold_left (f n) (f n acc c) l - | Proj (_,c) -> f n acc c + | Proj (_,_,c) -> f n acc c | Evar ev -> let args = Evd.expand_existential sigma ev in List.fold_left (fun c -> f n c) acc args @@ -1327,7 +1330,7 @@ let global_app_of_constr sigma c = | Ind (i, u) -> (IndRef i, u), None | Construct (c, u) -> (ConstructRef c, u), None | Var id -> (VarRef id, EConstr.EInstance.empty), None - | Proj (p, c) -> (ConstRef (Projection.constant p), EConstr.EInstance.empty), Some c + | Proj (p, _, c) -> (ConstRef (Projection.constant p), EConstr.EInstance.empty), Some c | _ -> raise Not_found let prod_applist sigma c l = diff --git a/engine/termops.mli b/engine/termops.mli index da70e9b621b0..5acf326945f6 100644 --- a/engine/termops.mli +++ b/engine/termops.mli @@ -292,6 +292,7 @@ val pr_evar_map_filter : ?with_univs:bool -> (Evar.t -> any_evar_info -> bool) - val pr_metaset : Metaset.t -> Pp.t val pr_evar_universe_context : UState.t -> Pp.t val pr_evd_level : evar_map -> Univ.Level.t -> Pp.t +val pr_evd_qvar : evar_map -> Sorts.QVar.t -> Pp.t module Internal : sig diff --git a/engine/uState.ml b/engine/uState.ml index 9670c2d95f79..f0cfe9893042 100644 --- a/engine/uState.ml +++ b/engine/uState.ml @@ -12,70 +12,65 @@ open CErrors open Util open Names open Univ +open Sorts +open UVars module UnivFlex = UnivFlex type universes_entry = | Monomorphic_entry of Univ.ContextSet.t -| Polymorphic_entry of Univ.UContext.t +| Polymorphic_entry of UVars.UContext.t -module UNameMap = Names.Id.Map +module UNameMap = Id.Map type uinfo = { uname : Id.t option; uloc : Loc.t option; } -type quality = QVar of Sorts.QVar.t | QProp | QSProp | QType +open Quality let sort_inconsistency ?explain cst l r = let explain = Option.map (fun p -> UGraph.Other p) explain in raise (UGraph.UniverseInconsistency (cst, l, r, explain)) -let pr_quality = function - | QVar v -> Sorts.QVar.pr v - | QProp -> Pp.str "Prop" - | QSProp -> Pp.str "SProp" - | QType -> Pp.str "Type" - module QState : sig type t - type elt = Sorts.QVar.t + type elt = QVar.t val empty : t - val union : fail:(t -> quality -> quality -> t) -> t -> t -> t - val add : elt -> t -> t - val repr : elt -> t -> quality - val unify_quality : fail:(unit -> t) -> Conversion.conv_pb -> quality -> quality -> t -> t + val union : fail:(t -> Quality.t -> Quality.t -> t) -> t -> t -> t + val add : check_fresh:bool -> named:bool -> elt -> t -> t + val repr : elt -> t -> Quality.t + val unify_quality : fail:(unit -> t) -> Conversion.conv_pb -> Quality.t -> Quality.t -> t -> t val is_above_prop : elt -> t -> bool + val undefined : t -> QVar.Set.t val collapse : t -> t val pr : t -> Pp.t + val of_set : QVar.Set.t -> t end = struct -module QSet = Set.Make(Sorts.QVar) -module QMap = Map.Make(Sorts.QVar) +module QSet = QVar.Set +module QMap = QVar.Map type t = { - qmap : quality option QMap.t; + named : QSet.t; + (** Named variables, may not be set to another *) + qmap : Quality.t option QMap.t; (* TODO: use a persistent union-find structure *) above : QSet.t; (** Set of quality variables known to be either in Prop or Type. If q ∈ above then it must map to None in qmap. *) } -type elt = Sorts.QVar.t - -let empty = { qmap = QMap.empty; above = QSet.empty } +type elt = QVar.t -let quality_eq a b = match a, b with - | QProp, QProp | QSProp, QSProp | QType, QType -> true - | QVar q1, QVar q2 -> Sorts.QVar.equal q1 q2 - | (QVar _ | QProp | QSProp | QType), _ -> false +let empty = { named = QSet.empty; qmap = QMap.empty; above = QSet.empty } let rec repr q m = match QMap.find q m.qmap with | None -> QVar q | Some (QVar q) -> repr q m -| Some (QProp | QSProp | QType as q) -> q +| Some (QConstant _ as q) -> q | exception Not_found -> (* let () = assert !Flags.in_debugger in *) (* FIXME *) QVar q @@ -84,47 +79,60 @@ let is_above_prop q m = QSet.mem q m.above let set q qv m = let q = repr q m in - let q = match q with QVar q -> q | QProp | QSProp | QType -> assert false in - let qv = match qv with QVar qv -> repr qv m | (QSProp | QProp | QType as qv) -> qv in + let q = match q with QVar q -> q | QConstant _ -> assert false in + let qv = match qv with QVar qv -> repr qv m | (QConstant _ as qv) -> qv in match q, qv with | q, QVar qv -> - if Sorts.QVar.equal q qv then Some m + if QVar.equal q qv then Some m + else + if QSet.mem q m.named then None else let above = if QSet.mem q m.above then QSet.add qv (QSet.remove q m.above) else m.above in - Some { qmap = QMap.add q (Some (QVar qv)) m.qmap; above } - | q, (QProp | QSProp | QType as qv) -> - if qv == QSProp && QSet.mem q m.above then None - else Some { qmap = QMap.add q (Some qv) m.qmap; above = QSet.remove q m.above } + Some { named = m.named; qmap = QMap.add q (Some (QVar qv)) m.qmap; above } + | q, (QConstant qc as qv) -> + if qc == QSProp && QSet.mem q m.above then None + else if QSet.mem q m.named then None + else + Some { named = m.named; qmap = QMap.add q (Some qv) m.qmap; above = QSet.remove q m.above } let set_above_prop q m = + (* XXX fail if named? *) let q = repr q m in - let q = match q with QVar q -> q | QProp | QSProp | QType -> assert false in - { qmap = m.qmap; above = QSet.add q m.above } + let q = match q with QVar q -> q | QConstant _ -> assert false in + { named = m.named; qmap = m.qmap; above = QSet.add q m.above } let unify_quality ~fail c q1 q2 local = match q1, q2 with -| QType, QType | QProp, QProp | QSProp, QSProp -> local -| QProp, QVar q when c == Conversion.CUMUL -> +| QConstant QType, QConstant QType +| QConstant QProp, QConstant QProp +| QConstant QSProp, QConstant QSProp -> local +| QConstant QProp, QVar q when c == Conversion.CUMUL -> set_above_prop q local -| QVar q, (QType | QProp | QSProp | QVar _ as qv) -| (QType | QProp | QSProp as qv), QVar q -> +| QVar qv1, QVar qv2 -> begin match set qv1 q2 local with + | Some local -> local + | None -> match set qv2 q1 local with + | Some local -> local + | None -> fail () + end +| QVar q, (QConstant (QType | QProp | QSProp) as qv) +| (QConstant (QType | QProp | QSProp) as qv), QVar q -> begin match set q qv local with | Some local -> local | None -> fail () end -| (QType, (QProp | QSProp)) -> fail () -| (QProp, QType) -> +| (QConstant QType, QConstant (QProp | QSProp)) -> fail () +| (QConstant QProp, QConstant QType) -> begin match c with | CONV -> fail () | CUMUL -> local end -| (QSProp, (QType | QProp)) -> fail () -| (QProp, QSProp) -> fail () +| (QConstant QSProp, QConstant (QType | QProp)) -> fail () +| (QConstant QProp, QConstant QSProp) -> fail () let nf_quality m = function - | QSProp | QProp | QType as q -> q + | QConstant _ as q -> q | QVar q -> repr q m let union ~fail s1 s2 = @@ -134,7 +142,7 @@ let union ~fail s1 s2 = | Some q, None | None, Some q -> Some (Some q) | None, None -> Some None | Some q1, Some q2 -> - let () = if not (quality_eq q1 q2) then extra := (q1,q2) :: !extra in + let () = if not (Quality.equal q1 q2) then extra := (q1,q2) :: !extra in Some (Some q1)) s1.qmap s2.qmap in @@ -145,39 +153,52 @@ let union ~fail s1 s2 = | exception Not_found -> false in let above = QSet.filter filter @@ QSet.union s1.above s2.above in - let s = { qmap; above } in + let s = { named = QSet.union s1.named s2.named; qmap; above } in List.fold_left (fun s (q1,q2) -> let q1 = nf_quality s q1 and q2 = nf_quality s q2 in unify_quality ~fail:(fun () -> fail s q1 q2) CONV q1 q2 s) s extra -let add q m = { qmap = QMap.add q None m.qmap; above = m.above } +let add ~check_fresh ~named q m = + if check_fresh then assert (not (QMap.mem q m.qmap)); + { named = if named then QSet.add q m.named else m.named; + qmap = QMap.add q None m.qmap; + above = m.above } + +let of_set qs = + { named = QSet.empty; qmap = QMap.bind (fun _ -> None) qs; above = QSet.empty } + +(* XXX what about [above]? *) +let undefined m = + let m = QMap.filter (fun _ v -> Option.is_empty v) m.qmap in + QMap.domain m let collapse m = let map q v = match v with - | None -> Some QType + | None -> if QSet.mem q m.named then None else Some (QConstant QType) | Some _ -> v in - { qmap = QMap.mapi map m.qmap; above = QSet.empty } + { named = m.named; qmap = QMap.mapi map m.qmap; above = QSet.empty } let pr { qmap; above } = + (* TODO names *) let open Pp in let prbody u = function | None -> if QSet.mem u above then str " >= Prop" else mt () | Some q -> - let q = pr_quality q in + let q = Quality.raw_pr q in str " := " ++ q in - h (prlist_with_sep fnl (fun (u, v) -> Sorts.QVar.pr u ++ prbody u v) (QMap.bindings qmap)) + h (prlist_with_sep fnl (fun (u, v) -> QVar.raw_pr u ++ prbody u v) (QMap.bindings qmap)) end module UPairSet = UnivMinim.UPairSet -type univ_names = UnivNames.universe_binders * uinfo Level.Map.t +type univ_names = UnivNames.universe_binders * (uinfo QVar.Map.t * uinfo Level.Map.t) (* 2nd part used to check consistency on the fly. *) type t = @@ -195,7 +216,7 @@ type t = } let empty = - { names = UNameMap.empty, Level.Map.empty; + { names = UnivNames.empty_binders, (QVar.Map.empty, Level.Map.empty); local = ContextSet.empty; seff_univs = Level.Set.empty; univ_variables = UnivFlex.empty; @@ -223,14 +244,20 @@ let uname_union s t = | Some _, _ -> l | _, _ -> r) s t +let names_union ((qbind,ubind),(qrev,urev)) ((qbind',ubind'),(qrev',urev')) = + let qbind = uname_union qbind qbind' + and ubind = uname_union ubind ubind' + and qrev = QVar.Map.union (fun _ l _ -> Some l) qrev qrev' + and urev = Level.Map.lunion urev urev' in + ((qbind,ubind),(qrev,urev)) + let union uctx uctx' = if uctx == uctx' then uctx else if is_empty uctx' then uctx else let local = ContextSet.union uctx.local uctx'.local in let seff = Level.Set.union uctx.seff_univs uctx'.seff_univs in - let names = uname_union (fst uctx.names) (fst uctx'.names) in - let names_rev = Level.Map.lunion (snd uctx.names) (snd uctx'.names) in + let names = names_union uctx.names uctx'.names in let newus = Level.Set.diff (ContextSet.levels uctx'.local) (ContextSet.levels uctx.local) in let newus = Level.Set.diff newus (UnivFlex.domain uctx.univ_variables) in @@ -242,9 +269,9 @@ let union uctx uctx' = if UGraph.type_in_type uctx.universes then s else CErrors.user_err Pp.(str "Could not merge universe contexts: could not unify" ++ spc() ++ - pr_quality q1 ++ strbrk " and " ++ pr_quality q2 ++ str ".") + Quality.raw_pr q1 ++ strbrk " and " ++ Quality.raw_pr q2 ++ str ".") in - { names = (names, names_rev); + { names; local = local; seff_univs = seff; univ_variables = @@ -261,18 +288,31 @@ let union uctx uctx' = let context_set uctx = uctx.local +let sort_context_set uctx = + let us, csts = uctx.local in + (QState.undefined uctx.sort_variables, us), csts + let constraints uctx = snd uctx.local -let compute_instance_binders rbinders inst = - let map lvl = - try Name (Option.get (Level.Map.find lvl rbinders).uname) +let compute_instance_binders (qrev,urev) inst = + let qinst, uinst = Instance.to_array inst in + let qmap = function + | QVar q -> + begin try Name (Option.get (QVar.Map.find q qrev).uname) + with Option.IsNone | Not_found -> Anonymous + end + | QConstant _ -> assert false + in + let umap lvl = + try Name (Option.get (Level.Map.find lvl urev).uname) with Option.IsNone | Not_found -> Anonymous in - Array.map map (Instance.to_array inst) + Array.map qmap qinst, Array.map umap uinst let context uctx = let (_, rbinders) = uctx.names in - ContextSet.to_context (compute_instance_binders rbinders) uctx.local + let qvars = QState.undefined uctx.sort_variables in + UContext.of_context_set (compute_instance_binders rbinders) qvars uctx.local type named_universes_entry = universes_entry * UnivNames.universe_binders @@ -283,7 +323,11 @@ let univ_entry ~poly uctx = else Monomorphic_entry (context_set uctx) in entry, binders -let of_context_set local = { empty with local } +let of_context_set ((qs,us),csts) = + let sort_variables = QState.of_set qs in + { empty with + local = (us,csts); + sort_variables;} type universe_opt_subst = UnivFlex.t @@ -295,27 +339,23 @@ let initial_graph uctx = uctx.initial_universes let is_algebraic l uctx = UnivFlex.is_algebraic l uctx.univ_variables -let add_names ?loc s l (names, names_rev) = - if UNameMap.mem s names - then user_err ?loc - Pp.(str "Universe " ++ Names.Id.print s ++ str" already bound."); - (UNameMap.add s l names, Level.Map.add l { uname = Some s; uloc = loc } names_rev) - -let add_loc l loc (names, names_rev) = - match loc with - | None -> (names, names_rev) - | Some _ -> (names, Level.Map.add l { uname = None; uloc = loc } names_rev) - -let of_names (ubind,revubind) = +let of_names (ubind,(revqbind,revubind)) = + let revqbind = QVar.Map.map (fun id -> { uname = Some id; uloc = None }) revqbind in let revubind = Level.Map.map (fun id -> { uname = Some id; uloc = None }) revubind in - {empty with names = (ubind,revubind)} + {empty with names = (ubind,(revqbind,revubind))} let universe_of_name uctx s = - UNameMap.find s (fst uctx.names) + UNameMap.find s (snd (fst uctx.names)) + +let quality_of_name uctx s = + Id.Map.find s (fst (fst uctx.names)) let name_level level id uctx = - assert(not(Names.Id.Map.mem id (fst uctx.names))); - { uctx with names = (Names.Id.Map.add id level (fst uctx.names), Univ.Level.Map.add level { uname = Some id; uloc = None } (snd uctx.names)) } + let ((qbind,ubind),(qrev,urev)) = uctx.names in + assert(not(Id.Map.mem id ubind)); + let ubind = Id.Map.add id level ubind in + let urev = Level.Map.add level { uname = Some id; uloc = None } urev in + { uctx with names = ((qbind,ubind),(qrev,urev)) } let universe_binders uctx = @@ -339,40 +379,38 @@ let level_inconsistency cst l r = let mk u = Sorts.sort_of_univ @@ Universe.make u in raise (UGraph.UniverseInconsistency (cst, mk l, mk r, None)) -let subst_univs_sort normalize qnormalize s = match s with -| Sorts.Set | Sorts.Prop | Sorts.SProp -> s -| Sorts.Type u -> Sorts.sort_of_univ (UnivSubst.subst_univs_universe normalize u) -| Sorts.QSort (q, u) -> - match qnormalize q with - | QSProp -> Sorts.sprop - | QProp -> Sorts.prop - | QType -> Sorts.sort_of_univ (UnivSubst.subst_univs_universe normalize u) - | QVar q -> Sorts.qsort q (UnivSubst.subst_univs_universe normalize u) +let nf_universe uctx u = + UnivSubst.(subst_univs_universe (UnivFlex.normalize_univ_variable uctx.univ_variables)) u + +let nf_level uctx u = + UnivSubst.(level_subst_of (UnivFlex.normalize_univ_variable uctx.univ_variables)) u + +let nf_instance uctx u = Instance.subst_fn (nf_qvar uctx, nf_level uctx) u + +let nf_quality uctx q = Quality.subst (nf_qvar uctx) q let nf_sort uctx s = - let normalize u = UnivFlex.normalize_univ_variable uctx.univ_variables u in + let normalize u = nf_universe uctx u in let qnormalize q = QState.repr q uctx.sort_variables in - subst_univs_sort normalize qnormalize s + Sorts.subst_fn (qnormalize, normalize) s let nf_relevance uctx r = match r with -| Sorts.Relevant | Sorts.Irrelevant -> r -| Sorts.RelevanceVar q -> +| Relevant | Irrelevant -> r +| RelevanceVar q -> match nf_qvar uctx q with - | QSProp -> Sorts.Irrelevant - | QProp | QType -> Sorts.Relevant + | QConstant QSProp -> Sorts.Irrelevant + | QConstant QProp | QConstant QType -> Sorts.Relevant | QVar q' -> + (* XXX currently not used in nf_evars_and_universes_opt_subst + does it matter? *) if QState.is_above_prop q' uctx.sort_variables then Relevant - else if Sorts.QVar.equal q q' then r + else if QVar.equal q q' then r else Sorts.RelevanceVar q' let nf_universes uctx c = let lsubst = uctx.univ_variables in - let level_value l = - UnivSubst.level_subst_of (fun l -> UnivFlex.normalize_univ_variable lsubst l) l - in - let sort_value s = nf_sort uctx s in - let rel_value r = nf_relevance uctx r in - UnivSubst.nf_evars_and_universes_opt_subst (fun _ -> None) level_value sort_value rel_value c + let nf_univ u = UnivFlex.normalize_univ_variable lsubst u in + UnivSubst.nf_evars_and_universes_opt_subst (fun _ -> None) (nf_qvar uctx) nf_univ c type small_universe = USet | UProp | USProp @@ -385,10 +423,10 @@ type sort_classification = | UAlgebraic of Universe.t (* Arbitrary algebraic expression *) let classify s = match s with -| Sorts.Prop -> USmall UProp -| Sorts.SProp -> USmall USProp -| Sorts.Set -> USmall USet -| Sorts.Type u | Sorts.QSort (_, u) -> +| Prop -> USmall UProp +| SProp -> USmall USProp +| Set -> USmall USet +| Type u | QSort (_, u) -> if Universe.is_levels u then match Universe.level u with | None -> UMax (u, Universe.levels u) | Some u -> ULevel u @@ -408,12 +446,6 @@ let add_local cst local = let enforce_leq_up u v local = { local with local_cst = UnivSubst.enforce_leq u (Universe.make v) local.local_cst } -let quality_of_sort = function -| Sorts.Set | Sorts.Type _ -> QType -| Sorts.Prop -> QProp -| Sorts.SProp -> QSProp -| Sorts.QSort (q, _) -> QVar q - let get_constraint = function | Conversion.CONV -> Eq | Conversion.CUMUL -> Le @@ -424,7 +456,7 @@ let unify_quality univs c s1 s2 l = in { l with local_sorts = QState.unify_quality ~fail - c (quality_of_sort s1) (quality_of_sort s2) l.local_sorts; + c (Sorts.quality s1) (Sorts.quality s2) l.local_sorts; } let process_universe_constraints uctx cstrs = @@ -433,11 +465,13 @@ let process_universe_constraints uctx cstrs = let univs = uctx.universes in let vars = ref uctx.univ_variables in let normalize u = UnivFlex.normalize_univ_variable !vars u in + let qnormalize sorts q = QState.repr q sorts in let normalize_sort sorts s = - let qnormalize q = QState.repr q sorts in - subst_univs_sort normalize qnormalize s + Sorts.subst_fn ((qnormalize sorts), subst_univs_universe normalize) s in let nf_constraint sorts = function + | QLeq (a, b) -> QLeq (Quality.subst (qnormalize sorts) a, Quality.subst (qnormalize sorts) b) + | QEq (a, b) -> QEq (Quality.subst (qnormalize sorts) a, Quality.subst (qnormalize sorts) b) | ULub (u, v) -> ULub (level_subst_of normalize u, level_subst_of normalize v) | UWeak (u, v) -> UWeak (level_subst_of normalize u, level_subst_of normalize v) | UEq (u, v) -> UEq (normalize_sort sorts u, normalize_sort sorts v) @@ -446,19 +480,19 @@ let process_universe_constraints uctx cstrs = let is_local l = UnivFlex.mem l !vars in let equalize_small l s local = let ls = match l with - | USProp -> Sorts.sprop - | UProp -> Sorts.prop - | USet -> Sorts.set + | USProp -> sprop + | UProp -> prop + | USet -> set in if UGraph.check_eq_sort univs ls s then local else if is_uset l then match classify s with - | USmall _ -> sort_inconsistency Eq Sorts.set s + | USmall _ -> sort_inconsistency Eq set s | ULevel r -> if is_local r then let () = instantiate_variable r Universe.type0 vars in add_local (Level.set, Eq, r) local else - sort_inconsistency Eq Sorts.set s + sort_inconsistency Eq set s | UMax (u, _)| UAlgebraic u -> if univ_level_mem Level.set u then let inst = univ_level_rem Level.set u u in @@ -494,7 +528,7 @@ let process_universe_constraints uctx cstrs = else if univ_level_mem l ru then enforce_leq_up inst l local - else sort_inconsistency Eq (Sorts.sort_of_univ (Universe.make l)) (Sorts.sort_of_univ ru) + else sort_inconsistency Eq (sort_of_univ (Universe.make l)) (sort_of_univ ru) in let equalize_universes l r local = match classify l, classify r with | USmall l', (USmall _ | ULevel _ | UMax _ | UAlgebraic _) -> @@ -514,6 +548,26 @@ let process_universe_constraints uctx cstrs = let cst = nf_constraint local.local_sorts cst in if UnivProblem.is_trivial cst then local else match cst with + | QEq (a, b) -> + (* TODO sort_inconsistency should be able to handle raw + qualities instead of having to make a dummy sort *) + let mk = function + | QVar q -> qsort q Universe.type0 + | QConstant QSProp -> sprop + | QConstant QProp -> prop + | QConstant QType -> set + in + unify_quality univs CONV (mk a) (mk b) local + | QLeq (a, b) -> + (* TODO sort_inconsistency should be able to handle raw + qualities instead of having to make a dummy sort *) + let mk = function + | QVar q -> qsort q Universe.type0 + | QConstant QSProp -> sprop + | QConstant QProp -> prop + | QConstant QType -> set + in + unify_quality univs CUMUL (mk a) (mk b) local | ULe (l, r) -> let local = unify_quality univs CUMUL l r local in let l = normalize_sort local.local_sorts l in @@ -546,7 +600,7 @@ let process_universe_constraints uctx cstrs = | UMax (_, levels) -> if is_uset r' then let fold l' local = - let l = Sorts.sort_of_univ @@ Universe.make l' in + let l = sort_of_univ @@ Universe.make l' in if Level.is_set l' || is_local l' then equalize_variables false l' Level.set local else sort_inconsistency Le l r @@ -605,27 +659,6 @@ let process_universe_constraints uctx cstrs = let extra = { UnivMinim.above_prop = local.local_above_prop; UnivMinim.weak_constraints = local.local_weak } in !vars, extra, local.local_cst, local.local_sorts -let add_constraints uctx cstrs = - let univs, old_cstrs = uctx.local in - let cstrs' = Constraints.fold (fun (l,d,r) acc -> - let l = Universe.make l and r = Sorts.sort_of_univ @@ Universe.make r in - let cstr' = let open UnivProblem in - match d with - | Lt -> - ULe (Sorts.sort_of_univ @@ Universe.super l, r) - | Le -> ULe (Sorts.sort_of_univ l, r) - | Eq -> UEq (Sorts.sort_of_univ l, r) - in UnivProblem.Set.add cstr' acc) - cstrs UnivProblem.Set.empty - in - let vars, extra, cstrs', sorts = process_universe_constraints uctx cstrs' in - { uctx with - local = (univs, Constraints.union old_cstrs cstrs'); - univ_variables = vars; - universes = UGraph.merge_constraints cstrs' uctx.universes; - sort_variables = sorts; - minim_extra = extra; } - let add_universe_constraints uctx cstrs = let univs, local = uctx.local in let vars, extra, local', sorts = process_universe_constraints uctx cstrs in @@ -636,22 +669,96 @@ let add_universe_constraints uctx cstrs = sort_variables = sorts; minim_extra = extra; } +let problem_of_constraints cstrs = + Constraints.fold (fun (l,d,r) acc -> + let l = Universe.make l and r = sort_of_univ @@ Universe.make r in + let cstr' = let open UnivProblem in + match d with + | Lt -> + ULe (sort_of_univ @@ Universe.super l, r) + | Le -> ULe (sort_of_univ l, r) + | Eq -> UEq (sort_of_univ l, r) + in UnivProblem.Set.add cstr' acc) + cstrs UnivProblem.Set.empty + +let add_constraints uctx cstrs = + let cstrs = problem_of_constraints cstrs in + add_universe_constraints uctx cstrs + +let add_quconstraints uctx (qcstrs,ucstrs) = + let cstrs = problem_of_constraints ucstrs in + let cstrs = QConstraints.fold (fun (l,d,r) cstrs -> + match d with + | Equal -> UnivProblem.Set.add (QEq (l,r)) cstrs + | Leq -> UnivProblem.Set.add (QLeq (l,r)) cstrs) + qcstrs cstrs + in + add_universe_constraints uctx cstrs + +let check_qconstraints uctx csts = + Sorts.QConstraints.for_all (fun (l,k,r) -> + let l = nf_quality uctx l in + let r = nf_quality uctx r in + if Quality.equal l r then true + else match l,k,r with + | _, Equal, _ -> false + | QConstant QProp, Leq, QConstant QType -> true + | QConstant QProp, Leq, QVar q -> QState.is_above_prop q uctx.sort_variables + | _, Leq, _ -> false) + csts + +let check_universe_constraint uctx (c:UnivProblem.t) = + match c with + | QEq (a,b) -> + let a = nf_quality uctx a in + let b = nf_quality uctx b in + Quality.equal a b + | QLeq (a,b) -> + let a = nf_quality uctx a in + let b = nf_quality uctx b in + if Quality.equal a b then true + else begin match a, b with + | QConstant QProp, QConstant QType -> true + | QConstant QProp, QVar q -> QState.is_above_prop q uctx.sort_variables + | _ -> false + end + | ULe (u,v) -> UGraph.check_leq_sort uctx.universes u v + | UEq (u,v) -> UGraph.check_eq_sort uctx.universes u v + | ULub (u,v) -> UGraph.check_eq_level uctx.universes u v + | UWeak _ -> true + +let check_universe_constraints uctx csts = + UnivProblem.Set.for_all (check_universe_constraint uctx) csts + let constrain_variables diff uctx = let local, vars = UnivFlex.constrain_variables diff uctx.univ_variables uctx.local in { uctx with local; univ_variables = vars } let id_of_level uctx l = - try Some (Option.get (Level.Map.find l (snd uctx.names)).uname) + try (Level.Map.find l (snd (snd uctx.names))).uname + with Not_found -> None + +let id_of_qvar uctx l = + try (QVar.Map.find l (fst (snd uctx.names))).uname + with Not_found -> None + +let qualid_of_qvar_names (bind, (qrev,_)) l = + try Some (Libnames.qualid_of_ident (Option.get (QVar.Map.find l qrev).uname)) with Not_found | Option.IsNone -> - None + None (* no global qvars *) -let qualid_of_level_names (map, map_rev) l = - try Some (Libnames.qualid_of_ident (Option.get (Level.Map.find l map_rev).uname)) +let qualid_of_level_names (bind, (_,urev)) l = + try Some (Libnames.qualid_of_ident (Option.get (Level.Map.find l urev).uname)) with Not_found | Option.IsNone -> - UnivNames.qualid_of_level map l + UnivNames.qualid_of_level bind l let qualid_of_level uctx l = qualid_of_level_names uctx.names l +let pr_uctx_qvar_names names l = + match qualid_of_qvar_names names l with + | Some qid -> Libnames.pr_qualid qid + | None -> QVar.raw_pr l + let pr_uctx_level_names names l = match qualid_of_level_names names l with | Some qid -> Libnames.pr_qualid qid @@ -659,54 +766,88 @@ let pr_uctx_level_names names l = let pr_uctx_level uctx l = pr_uctx_level_names uctx.names l -type ('a, 'b) gen_universe_decl = { - univdecl_instance : 'a; (* Declared universes *) +let pr_uctx_qvar uctx l = pr_uctx_qvar_names uctx.names l + +type ('a, 'b, 'c) gen_universe_decl = { + univdecl_qualities : 'a; + univdecl_instance : 'b; (* Declared universes *) univdecl_extensible_instance : bool; (* Can new universes be added *) - univdecl_constraints : 'b; (* Declared constraints *) + univdecl_constraints : 'c; (* Declared constraints *) univdecl_extensible_constraints : bool (* Can new constraints be added *) } type universe_decl = - (Level.t list, Constraints.t) gen_universe_decl + (QVar.t list, Level.t list, Constraints.t) gen_universe_decl let default_univ_decl = - { univdecl_instance = []; + { univdecl_qualities = []; + univdecl_instance = []; univdecl_extensible_instance = true; univdecl_constraints = Constraints.empty; univdecl_extensible_constraints = true } -let pr_error_unbound_universes left names = +let pr_error_unbound_universes quals univs names = let open Pp in - let n = Level.Set.cardinal left in + let nqs = QVar.Set.cardinal quals in + let prqvar q = + let info = QVar.Map.find_opt q (fst (snd names)) in + h (pr_uctx_qvar_names names q ++ (match info with + | None | Some {uloc=None} -> mt () + | Some {uloc=Some loc} -> spc() ++ str"(" ++ Loc.pr loc ++ str")")) + in + let nus = Level.Set.cardinal univs in let prlev u = - let info = Level.Map.find_opt u (snd names) in + let info = Level.Map.find_opt u (snd (snd names)) in h (pr_uctx_level_names names u ++ (match info with | None | Some {uloc=None} -> mt () | Some {uloc=Some loc} -> spc() ++ str"(" ++ Loc.pr loc ++ str")")) in + let ppqs = if nqs > 0 then + str (if nqs = 1 then "Quality" else "Qualities") ++ spc () ++ + prlist_with_sep spc prqvar (QVar.Set.elements quals) + else mt() + in + let ppus = if nus > 0 then + let universe_s = CString.plural nus "universe" in + let universe_s = if nqs = 0 then CString.capitalize_ascii universe_s else universe_s in + str universe_s ++ spc () ++ + prlist_with_sep spc prlev (Level.Set.elements univs) + else mt() + in (hv 0 - (str (CString.plural n "Universe") ++ spc () ++ - (prlist_with_sep spc prlev (Level.Set.elements left)) ++ - spc () ++ str (CString.conjugate_verb_to_be n) ++ str" unbound.")) + (ppqs ++ + (if nqs > 0 && nus > 0 then strbrk " and " else mt()) ++ + ppus ++ + spc () ++ str (CString.conjugate_verb_to_be (nus + nqs)) ++ str" unbound.")) -exception UnboundUnivs of Level.Set.t * univ_names +exception UnboundUnivs of QVar.Set.t * Level.Set.t * univ_names (* Deliberately using no location as the location of the univs doesn't correspond to the failing command. *) -let error_unbound_universes left uctx = raise (UnboundUnivs (left,uctx)) +let error_unbound_universes qs us uctx = raise (UnboundUnivs (qs,us,uctx)) let _ = CErrors.register_handler (function - | UnboundUnivs (left,uctx) -> Some (pr_error_unbound_universes left uctx) + | UnboundUnivs (qs,us,uctx) -> Some (pr_error_unbound_universes qs us uctx) | _ -> None) -let universe_context_inst ~prefix ~extensible levels names = - let left = List.fold_left (fun acc l -> Level.Set.remove l acc) levels prefix in - if not extensible && not (Level.Set.is_empty left) - then error_unbound_universes left names - else - let left = ContextSet.sort_levels (Array.of_list (Level.Set.elements left)) in - let inst = Array.append (Array.of_list prefix) left in - let inst = Instance.of_array inst in - inst +let universe_context_inst decl qvars levels names = + let leftqs = List.fold_left (fun acc l -> QVar.Set.remove l acc) qvars decl.univdecl_qualities in + let leftus = List.fold_left (fun acc l -> Level.Set.remove l acc) levels decl.univdecl_instance in + let () = + let unboundus = if decl.univdecl_extensible_instance then Level.Set.empty else leftus in + if not (QVar.Set.is_empty leftqs && Level.Set.is_empty unboundus) + then error_unbound_universes leftqs unboundus names + in + let leftqs = UContext.sort_qualities + (Array.map_of_list (fun q -> Quality.QVar q) (QVar.Set.elements leftqs)) + in + let leftus = UContext.sort_levels (Array.of_list (Level.Set.elements leftus)) in + let instq = Array.append + (Array.map_of_list (fun q -> Quality.QVar q) decl.univdecl_qualities) + leftqs + in + let instu = Array.append (Array.of_list decl.univdecl_instance) leftus in + let inst = Instance.of_array (instq,instu) in + inst let check_universe_context_set ~prefix levels names = let left = @@ -714,7 +855,7 @@ let check_universe_context_set ~prefix levels names = levels prefix in if not (Level.Set.is_empty left) - then error_unbound_universes left names + then error_unbound_universes QVar.Set.empty left names let check_implication uctx cstrs cstrs' = let gr = initial_graph uctx in @@ -723,9 +864,14 @@ let check_implication uctx cstrs cstrs' = if Constraints.is_empty cstrs' then () else CErrors.user_err Pp.(str "Universe constraints are not implied by the ones declared: " ++ - pr_constraints (pr_uctx_level uctx) cstrs') + Constraints.pr (pr_uctx_level uctx) cstrs') let check_mono_univ_decl uctx decl = + let () = + if not (List.is_empty decl.univdecl_qualities) + || not (QVar.Set.is_empty (QState.undefined uctx.sort_variables)) + then CErrors.user_err Pp.(str "Monomorphic declarations may not have sort variables.") + in let levels, csts = uctx.local in let () = let prefix = decl.univdecl_instance in @@ -741,10 +887,9 @@ let check_mono_univ_decl uctx decl = end let check_poly_univ_decl uctx decl = - let prefix = decl.univdecl_instance in - let extensible = decl.univdecl_extensible_instance in let levels, csts = uctx.local in - let inst = universe_context_inst ~prefix ~extensible levels uctx.names in + let qvars = QState.undefined uctx.sort_variables in + let inst = universe_context_inst decl qvars levels uctx.names in let nas = compute_instance_binders (snd uctx.names) inst in let csts = if decl.univdecl_extensible_constraints then csts else begin @@ -788,8 +933,8 @@ let restrict_universe_context ~lbound (univs, csts) keep = let restrict uctx vars = let vars = Level.Set.union vars uctx.seff_univs in - let vars = Names.Id.Map.fold (fun na l vars -> Level.Set.add l vars) - (fst uctx.names) vars + let vars = Id.Map.fold (fun na l vars -> Level.Set.add l vars) + (snd (fst uctx.names)) vars in let uctx' = restrict_universe_context ~lbound:uctx.universes_lbound uctx.local vars in { uctx with local = uctx' } @@ -808,11 +953,8 @@ let univ_flexible = UnivFlexible false let univ_flexible_alg = UnivFlexible true (** ~sideff indicates that it is ok to redeclare a universe. - ~extend also merges the universe context in the local constraint structures - and not only in the graph. This depends if the - context we merge comes from a side effect that is already inlined - or defined separately. In the later case, there is no extension, - see [emit_side_effects] for example. *) + Also merges the universe context in the local constraint structures + and not only in the graph. *) let merge ?loc ~sideff rigid uctx uctx' = let levels = ContextSet.levels uctx' in let local = ContextSet.append uctx' uctx.local in @@ -824,14 +966,15 @@ let merge ?loc ~sideff rigid uctx uctx' = in let names = let fold u accu = - let modify _ info = match info.uloc with - | None -> { info with uloc = loc } - | Some _ -> info + let update = function + | None -> Some { uname = None; uloc = loc } + | Some info -> match info.uloc with + | None -> Some { info with uloc = loc } + | Some _ -> Some info in - try Level.Map.modify u modify accu - with Not_found -> Level.Map.add u { uname = None; uloc = loc } accu + Level.Map.update u update accu in - (fst uctx.names, Level.Set.fold fold levels (snd uctx.names)) + (fst uctx.names, (fst (snd uctx.names), Level.Set.fold fold levels (snd (snd uctx.names)))) in let initial = declare uctx.initial_universes in let univs = declare uctx.universes in @@ -847,6 +990,31 @@ let merge ?loc ~sideff rigid uctx uctx' = { uctx with names; local; universes; initial_universes = initial } +let merge_sort_variables ?loc ~sideff uctx qvars = + let sort_variables = + QVar.Set.fold (fun qv qstate -> QState.add ~check_fresh:(not sideff) ~named:false qv qstate) + qvars + uctx.sort_variables + in + let names = + let fold u accu = + let update = function + | None -> Some { uname = None; uloc = loc } + | Some info -> match info.uloc with + | None -> Some { info with uloc = loc } + | Some _ -> Some info + in + QVar.Map.update u update accu + in + let qrev = QVar.Set.fold fold qvars (fst (snd uctx.names)) in + (fst uctx.names, (qrev, snd (snd uctx.names))) + in + { uctx with sort_variables; names } + +let merge_sort_context ?loc ~sideff rigid uctx ((qvars,levels),csts) = + let uctx = merge_sort_variables ?loc ~sideff uctx qvars in + merge ?loc ~sideff rigid uctx (levels,csts) + (* Check bug_4363 and bug_6323 when changing this code *) let demote_seff_univs univs uctx = let seff = Level.Set.union uctx.seff_univs univs in @@ -886,6 +1054,30 @@ let update_sigma_univs uctx univs = in merge_seff eunivs eunivs.local +let add_qnames ?loc s l ((qnames,unames), (qnames_rev,unames_rev)) = + if Id.Map.mem s qnames + then user_err ?loc + Pp.(str "Quality " ++ Id.print s ++ str" already bound."); + ((Id.Map.add s l qnames, unames), + (QVar.Map.add l { uname = Some s; uloc = loc } qnames_rev, unames_rev)) + +let add_names ?loc s l ((qnames,unames), (qnames_rev,unames_rev)) = + if UNameMap.mem s unames + then user_err ?loc + Pp.(str "Universe " ++ Id.print s ++ str" already bound."); + ((qnames,UNameMap.add s l unames), + (qnames_rev, Level.Map.add l { uname = Some s; uloc = loc } unames_rev)) + +let add_qloc l loc (names, (qnames_rev,unames_rev) as orig) = + match loc with + | None -> orig + | Some _ -> (names, (QVar.Map.add l { uname = None; uloc = loc } qnames_rev, unames_rev)) + +let add_loc l loc (names, (qnames_rev,unames_rev) as orig) = + match loc with + | None -> orig + | Some _ -> (names, (qnames_rev, Level.Map.add l { uname = None; uloc = loc } unames_rev)) + let add_universe ?loc name strict lbound uctx u = let initial_universes = UGraph.add_universe ~lbound ~strict u uctx.initial_universes in let universes = UGraph.add_universe ~lbound ~strict u uctx.universes in @@ -897,10 +1089,17 @@ let add_universe ?loc name strict lbound uctx u = in { uctx with names; local; initial_universes; universes } -let new_sort_variable uctx = +let new_sort_variable ?loc ?name uctx = let q = UnivGen.new_sort_global () in - let sort_variables = QState.add q uctx.sort_variables in - { uctx with sort_variables }, q + (* don't need to check_fresh as it's guaranteed new *) + let sort_variables = QState.add ~check_fresh:false ~named:(Option.has_some name) + q uctx.sort_variables + in + let names = match name with + | Some n -> add_qnames ?loc n q uctx.names + | None -> add_qloc q loc uctx.names + in + { uctx with sort_variables; names }, q let new_univ_variable ?loc rigid name uctx = let u = UnivGen.fresh_level () in diff --git a/engine/uState.mli b/engine/uState.mli index 1675d73e439a..19adc28ff63a 100644 --- a/engine/uState.mli +++ b/engine/uState.mli @@ -14,12 +14,11 @@ open Names open Univ - -type quality = QVar of Sorts.QVar.t | QProp | QSProp | QType +open Sorts type universes_entry = | Monomorphic_entry of Univ.ContextSet.t -| Polymorphic_entry of Univ.UContext.t +| Polymorphic_entry of UVars.UContext.t exception UniversesDiffer @@ -43,10 +42,10 @@ val from_env : ?binders:lident list -> Environ.env -> t (** Main entry point at the beginning of a declaration declaring the binding names as rigid universes. *) -val of_names : (UnivNames.universe_binders * Id.t Level.Map.t) -> t +val of_names : (UnivNames.universe_binders * UnivNames.rev_binders) -> t (** Main entry point when only names matter, e.g. for printing. *) -val of_context_set : Univ.ContextSet.t -> t +val of_context_set : UnivGen.sort_context_set -> t (** Main entry point when starting from the instance of a global reference, e.g. when building a scheme. *) @@ -62,6 +61,8 @@ val context_set : t -> Univ.ContextSet.t (** The local context of the state, i.e. a set of bound variables together with their associated constraints. *) +val sort_context_set : t -> UnivGen.sort_context_set + type universe_opt_subst = UnivFlex.t (* Reexport because UnivSubst is private *) @@ -84,7 +85,7 @@ val is_algebraic : Level.t -> t -> bool val constraints : t -> Univ.Constraints.t (** Shorthand for {!context_set} composed with {!ContextSet.constraints}. *) -val context : t -> Univ.UContext.t +val context : t -> UVars.UContext.t (** Shorthand for {!context_set} with {!Context_set.to_context}. *) type named_universes_entry = universes_entry * UnivNames.universe_binders @@ -95,13 +96,22 @@ val univ_entry : poly:bool -> t -> named_universes_entry val universe_binders : t -> UnivNames.universe_binders (** Return local names of universes. *) -val nf_qvar : t -> Sorts.QVar.t -> quality +val nf_qvar : t -> QVar.t -> Quality.t (** Returns the normal form of the sort variable. *) +val nf_quality : t -> Quality.t -> Quality.t + +val nf_instance : t -> UVars.Instance.t -> UVars.Instance.t + +val nf_level : t -> Level.t -> Level.t +(** Must not be allowed to be algebraic *) + +val nf_universe : t -> Universe.t -> Universe.t + val nf_sort : t -> Sorts.t -> Sorts.t (** Returns the normal form of the sort. *) -val nf_relevance : t -> Sorts.relevance -> Sorts.relevance +val nf_relevance : t -> relevance -> relevance (** Returns the normal form of the relevance. *) (** {5 Constraints handling} *) @@ -116,8 +126,16 @@ val add_universe_constraints : t -> UnivProblem.Set.t -> t @raise UniversesDiffer when universes differ *) +val check_qconstraints : t -> QConstraints.t -> bool + +val check_universe_constraints : t -> UnivProblem.Set.t -> bool + +val add_quconstraints : t -> QUConstraints.t -> t + (** {5 Names} *) +val quality_of_name : t -> Id.t -> Sorts.QVar.t + val universe_of_name : t -> Id.t -> Univ.Level.t (** Retrieve the universe associated to the name. *) @@ -156,6 +174,9 @@ val univ_flexible : rigid val univ_flexible_alg : rigid val merge : ?loc:Loc.t -> sideff:bool -> rigid -> t -> Univ.ContextSet.t -> t +val merge_sort_variables : ?loc:Loc.t -> sideff:bool -> t -> QVar.Set.t -> t +val merge_sort_context : ?loc:Loc.t -> sideff:bool -> rigid -> t -> UnivGen.sort_context_set -> t + val emit_side_effects : Safe_typing.private_constants -> t -> t val demote_global_univs : Environ.env -> t -> t @@ -168,7 +189,7 @@ val demote_seff_univs : Univ.Level.Set.t -> t -> t globally declared by some side effect. You should be using emit_side_effects instead. *) -val new_sort_variable : t -> t * Sorts.QVar.t +val new_sort_variable : ?loc:Loc.t -> ?name:Id.t -> t -> t * QVar.t (** Declare a new local sort. *) val new_univ_variable : ?loc:Loc.t -> rigid -> Id.t option -> t -> t * Univ.Level.t @@ -197,14 +218,15 @@ val minimize : t -> t val collapse_sort_variables : t -> t -type ('a, 'b) gen_universe_decl = { - univdecl_instance : 'a; (* Declared universes *) +type ('a, 'b, 'c) gen_universe_decl = { + univdecl_qualities : 'a; + univdecl_instance : 'b; (* Declared universes *) univdecl_extensible_instance : bool; (* Can new universes be added *) - univdecl_constraints : 'b; (* Declared constraints *) + univdecl_constraints : 'c; (* Declared constraints *) univdecl_extensible_constraints : bool (* Can new constraints be added *) } type universe_decl = - (Level.t list, Univ.Constraints.t) gen_universe_decl + (QVar.t list, Level.t list, Univ.Constraints.t) gen_universe_decl val default_univ_decl : universe_decl @@ -230,11 +252,14 @@ val update_sigma_univs : t -> UGraph.t -> t (** {5 Pretty-printing} *) val pr_uctx_level : t -> Univ.Level.t -> Pp.t +val pr_uctx_qvar : t -> Sorts.QVar.t -> Pp.t val qualid_of_level : t -> Univ.Level.t -> Libnames.qualid option (** Only looks in the local names, not in the nametab. *) val id_of_level : t -> Univ.Level.t -> Id.t option +val id_of_qvar : t -> Sorts.QVar.t -> Id.t option + val pr_weak : (Univ.Level.t -> Pp.t) -> t -> Pp.t val pr_sort_opt_subst : t -> Pp.t diff --git a/engine/univGen.ml b/engine/univGen.ml index 6bfcca7eb232..18006dedc2d6 100644 --- a/engine/univGen.ml +++ b/engine/univGen.ml @@ -12,19 +12,40 @@ open Sorts open Names open Constr open Univ +open UVars + +type sort_context_set = (Sorts.QVar.Set.t * Univ.Level.Set.t) * Univ.Constraints.t + +type 'a in_sort_context_set = 'a * sort_context_set + +let empty_sort_context = (QVar.Set.empty, Level.Set.empty), Constraints.empty + +let is_empty_sort_context ((qs,us),csts) = + QVar.Set.is_empty qs && Level.Set.is_empty us && Constraints.is_empty csts + +let sort_context_union ((qs,us),csts) ((qs',us'),csts') = + ((QVar.Set.union qs qs', Level.Set.union us us'),Constraints.union csts csts') + +let diff_sort_context ((qs,us),csts) ((qs',us'),csts') = + (QVar.Set.diff qs qs', Level.Set.diff us us'), Constraints.diff csts csts' type univ_length_mismatch = { - actual : int ; - expect : int ; + actual : int * int; + expect : int * int; } (* Due to an OCaml bug ocaml/ocaml#10027 inlining this record will cause compliation with -rectypes to crash. *) exception UniverseLengthMismatch of univ_length_mismatch let () = CErrors.register_handler (function - | UniverseLengthMismatch { actual; expect } -> - Some Pp.(str "Universe instance length is " ++ int actual - ++ str " but should be " ++ int expect ++ str ".") + | UniverseLengthMismatch { actual=(aq,au); expect=(eq,eu) } -> + let ppreal, ppexpected = + if aq = 0 && eq = 0 then Pp.(int au, int eu) + else Pp.(str "(" ++ int aq ++ str " | " ++ int au ++ str ")" + , str "(" ++ int eq ++ str " | " ++ int eu ++ str ")") + in + Some Pp.(str "Universe instance length is " ++ ppreal + ++ str " but should be " ++ ppexpected ++ str".") | _ -> None) (* Generator of levels *) @@ -45,23 +66,31 @@ let new_sort_id = let new_sort_global () = let s = if Flags.async_proofs_is_worker() then !Flags.async_proofs_worker_id else "" in - Sorts.QVar.make s (new_sort_id ()) - -let fresh_instance auctx = - let inst = Array.init (AbstractContext.size auctx) (fun _ -> fresh_level()) in - let ctx = Array.fold_right Level.Set.add inst Level.Set.empty in - let inst = Instance.of_array inst in - inst, (ctx, AbstractContext.instantiate inst auctx) + Sorts.QVar.make_unif s (new_sort_id ()) + +let fresh_instance auctx : _ in_sort_context_set = + let qlen, ulen = AbstractContext.size auctx in + let qinst = Array.init qlen (fun _ -> Sorts.Quality.QVar (new_sort_global())) in + let uinst = Array.init ulen (fun _ -> fresh_level()) in + let qctx = Array.fold_left (fun qctx q -> match q with + | Sorts.Quality.QVar q -> Sorts.QVar.Set.add q qctx + | _ -> assert false) + Sorts.QVar.Set.empty + qinst + in + let uctx = Array.fold_right Level.Set.add uinst Level.Set.empty in + let inst = Instance.of_array (qinst,uinst) in + inst, ((qctx,uctx), AbstractContext.instantiate inst auctx) let existing_instance ?loc auctx inst = let () = - let actual = Array.length (Instance.to_array inst) + let actual = Instance.length inst and expect = AbstractContext.size auctx in - if not (Int.equal actual expect) then + if not (UVars.eq_sizes actual expect) then Loc.raise ?loc (UniverseLengthMismatch { actual; expect }) else () in - inst, (Level.Set.empty, AbstractContext.instantiate inst auctx) + inst, ((Sorts.QVar.Set.empty,Level.Set.empty), AbstractContext.instantiate inst auctx) let fresh_instance_from ?loc ctx = function | Some inst -> existing_instance ?loc ctx inst @@ -103,12 +132,12 @@ let constr_of_monomorphic_global env gr = str " would forget universes.") let fresh_sort_in_family = function - | InSProp -> Sorts.sprop, ContextSet.empty - | InProp -> Sorts.prop, ContextSet.empty - | InSet -> Sorts.set, ContextSet.empty + | InSProp -> Sorts.sprop, empty_sort_context + | InProp -> Sorts.prop, empty_sort_context + | InSet -> Sorts.set, empty_sort_context | InType | InQSort (* Treat as Type *) -> let u = fresh_level () in - sort_of_univ (Univ.Universe.make u), ContextSet.singleton u + sort_of_univ (Univ.Universe.make u), ((QVar.Set.empty,Level.Set.singleton u),Constraints.empty) let new_global_univ () = let u = fresh_level () in @@ -126,3 +155,13 @@ let fresh_universe_context_set_instance ctx = in let cst' = subst_univs_level_constraints subst cst in subst, (univs', cst') + +let fresh_sort_context_instance ((qs,us),csts) = + let usubst, (us, csts) = fresh_universe_context_set_instance (us,csts) in + let qsubst, qs = QVar.Set.fold (fun q (qsubst,qs) -> + let q' = new_sort_global () in + QVar.Map.add q (Sorts.Quality.QVar q') qsubst, QVar.Set.add q' qs) + qs + (QVar.Map.empty, QVar.Set.empty) + in + (qsubst, usubst), ((qs, us), csts) diff --git a/engine/univGen.mli b/engine/univGen.mli index 66197dd46c4f..1fe05a77c1c3 100644 --- a/engine/univGen.mli +++ b/engine/univGen.mli @@ -12,10 +12,11 @@ open Names open Constr open Environ open Univ +open UVars type univ_length_mismatch = { - actual : int ; - expect : int ; + actual : int * int; + expect : int * int; } (* Due to an OCaml bug ocaml/ocaml#10027 inlining this record will cause compliation with -rectypes to crash. *) @@ -32,30 +33,47 @@ val new_global_univ : unit -> Universe.t in_universe_context_set (** Build a fresh instance for a given context, its associated substitution and the instantiated constraints. *) -val fresh_instance : AbstractContext.t -> Instance.t in_universe_context_set +type sort_context_set = (Sorts.QVar.Set.t * Univ.Level.Set.t) * Univ.Constraints.t + +type 'a in_sort_context_set = 'a * sort_context_set + +val sort_context_union : sort_context_set -> sort_context_set -> sort_context_set + +val empty_sort_context : sort_context_set + +val is_empty_sort_context : sort_context_set -> bool + +val diff_sort_context : sort_context_set -> sort_context_set -> sort_context_set + +val fresh_instance : AbstractContext.t -> Instance.t in_sort_context_set val fresh_instance_from : ?loc:Loc.t -> AbstractContext.t -> Instance.t option -> - Instance.t in_universe_context_set + Instance.t in_sort_context_set val fresh_sort_in_family : Sorts.family -> - Sorts.t in_universe_context_set + Sorts.t in_sort_context_set +(** NB: InQSort is treated as InType *) + val fresh_constant_instance : env -> Constant.t -> - pconstant in_universe_context_set + pconstant in_sort_context_set val fresh_inductive_instance : env -> inductive -> - pinductive in_universe_context_set + pinductive in_sort_context_set val fresh_constructor_instance : env -> constructor -> - pconstructor in_universe_context_set + pconstructor in_sort_context_set val fresh_array_instance : env -> - Instance.t in_universe_context_set + Instance.t in_sort_context_set -val fresh_global_instance : ?loc:Loc.t -> ?names:Univ.Instance.t -> env -> GlobRef.t -> - constr in_universe_context_set +val fresh_global_instance : ?loc:Loc.t -> ?names:UVars.Instance.t -> env -> GlobRef.t -> + constr in_sort_context_set (** Get fresh variables for the universe context. Useful to make tactics that manipulate constrs in universe contexts polymorphic. *) val fresh_universe_context_set_instance : ContextSet.t -> universe_level_subst * ContextSet.t +val fresh_sort_context_instance : sort_context_set -> + sort_level_subst * sort_context_set + (** Create a fresh global in the environment argument, without side effects. BEWARE: this raises an error on polymorphic constants/inductives: the constraints should be properly added to an evd. diff --git a/engine/univNames.ml b/engine/univNames.ml index 16fd451fe107..432f8845855d 100644 --- a/engine/univNames.ml +++ b/engine/univNames.ml @@ -10,21 +10,28 @@ open Names open Univ +open Sorts -type universe_binders = Level.t Names.Id.Map.t +type universe_binders = QVar.t Id.Map.t * Level.t Id.Map.t -let empty_binders = Id.Map.empty +type rev_binders = Id.t QVar.Map.t * Id.t Level.Map.t + +let empty_binders = Id.Map.empty, Id.Map.empty + +let empty_rev_binders = QVar.Map.empty, Level.Map.empty type univ_name_list = Names.lname list -let qualid_of_level ctx l = +type full_name_list = lname list * lname list + +let qualid_of_level (_,ctx) l = match Level.name l with | Some qid -> (try Some (Nametab.shortest_qualid_of_universe ctx qid) with Not_found -> None) | None -> None -let pr_with_global_universes ?(binders=empty_binders) l = +let pr_level_with_global_universes ?(binders=empty_binders) l = match qualid_of_level binders l with | Some qid -> Libnames.pr_qualid qid | None -> Level.raw_pr l diff --git a/engine/univNames.mli b/engine/univNames.mli index 85a3ea097ecb..e2d6b9b5fd21 100644 --- a/engine/univNames.mli +++ b/engine/univNames.mli @@ -8,15 +8,23 @@ (* * (see LICENSE file for the text of the license) *) (************************************************************************) +open Names open Univ +open Sorts (** Local universe name <-> level mapping *) -type universe_binders = Univ.Level.t Names.Id.Map.t +type universe_binders = QVar.t Id.Map.t * Level.t Id.Map.t + +type rev_binders = Id.t QVar.Map.t * Id.t Level.Map.t val empty_binders : universe_binders +val empty_rev_binders : rev_binders + type univ_name_list = Names.lname list -val pr_with_global_universes : ?binders:universe_binders -> Level.t -> Pp.t +type full_name_list = lname list * lname list + +val pr_level_with_global_universes : ?binders:universe_binders -> Level.t -> Pp.t val qualid_of_level : universe_binders -> Level.t -> Libnames.qualid option diff --git a/engine/univProblem.ml b/engine/univProblem.ml index 4e27e50e09ed..55eb359c7ac6 100644 --- a/engine/univProblem.ml +++ b/engine/univProblem.ml @@ -11,6 +11,8 @@ open Univ type t = + | QEq of Sorts.Quality.t * Sorts.Quality.t + | QLeq of Sorts.Quality.t * Sorts.Quality.t | ULe of Sorts.t * Sorts.t | UEq of Sorts.t * Sorts.t | ULub of Level.t * Level.t @@ -18,21 +20,17 @@ type t = let is_trivial = function + | QLeq (QConstant QProp, QConstant QType) -> true + | QLeq (a,b) | QEq (a, b) -> Sorts.Quality.equal a b | ULe (u, v) | UEq (u, v) -> Sorts.equal u v | ULub (u, v) | UWeak (u, v) -> Level.equal u v let force = function - | ULe _ | UEq _ | UWeak _ as cst -> cst + | QEq _ | QLeq _ | ULe _ | UEq _ | UWeak _ as cst -> cst | ULub (u,v) -> UEq (Sorts.sort_of_univ @@ Universe.make u, Sorts.sort_of_univ @@ Universe.make v) let check_eq_level g u v = UGraph.check_eq_level g u v -let check g = function - | ULe (u,v) -> UGraph.check_leq_sort g u v - | UEq (u,v) -> UGraph.check_eq_sort g u v - | ULub (u,v) -> check_eq_level g u v - | UWeak _ -> true - module Set = struct module S = Set.Make( struct @@ -40,6 +38,14 @@ module Set = struct let compare x y = match x, y with + | QEq (a, b), QEq (a', b') -> + let i = Sorts.Quality.compare a a' in + if i <> 0 then i + else Sorts.Quality.compare b b' + | QLeq (a, b), QLeq (a', b') -> + let i = Sorts.Quality.compare a a' in + if i <> 0 then i + else Sorts.Quality.compare b b' | ULe (u, v), ULe (u', v') -> let i = Sorts.compare u u' in if Int.equal i 0 then Sorts.compare v v' @@ -54,6 +60,10 @@ module Set = struct if Int.equal i 0 then Level.compare v v' else if Level.equal u v' && Level.equal v u' then 0 else i + | QEq _, _ -> -1 + | _, QEq _ -> 1 + | QLeq _, _ -> -1 + | _, QLeq _ -> 1 | ULe _, _ -> -1 | _, ULe _ -> 1 | UEq _, _ -> -1 @@ -69,6 +79,8 @@ module Set = struct else add cst s let pr_one = let open Pp in function + | QEq (a, b) -> Sorts.Quality.raw_pr a ++ str " = " ++ Sorts.Quality.raw_pr b + | QLeq (a, b) -> Sorts.Quality.raw_pr a ++ str " <= " ++ Sorts.Quality.raw_pr b | ULe (u, v) -> Sorts.debug_print u ++ str " <= " ++ Sorts.debug_print v | UEq (u, v) -> Sorts.debug_print u ++ str " = " ++ Sorts.debug_print v | ULub (u, v) -> Level.raw_pr u ++ str " /\\ " ++ Level.raw_pr v @@ -83,8 +95,6 @@ module Set = struct x == y || equal x y let force s = map force s - - let check g s = for_all (check g) s end type 'a constraint_function = 'a -> 'a -> Set.t -> Set.t @@ -92,10 +102,40 @@ type 'a constraint_function = 'a -> 'a -> Set.t -> Set.t let enforce_eq_instances_univs strict x y c = let mkU u = Sorts.sort_of_univ @@ Universe.make u in let mk u v = if strict then ULub (u, v) else UEq (mkU u, mkU v) in - let ax = Instance.to_array x and ay = Instance.to_array y in - if Array.length ax != Array.length ay then - CErrors.anomaly Pp.(str "Invalid argument: enforce_eq_instances_univs called with" ++ - str " instances of different lengths."); - CArray.fold_right2 - (fun x y -> Set.add (mk x y)) - ax ay c + if not (UVars.eq_sizes (UVars.Instance.length x) (UVars.Instance.length y)) then + CErrors.anomaly Pp.(str "Invalid argument: enforce_eq_instances_univs called with" ++ + str " instances of different lengths."); + let xq, xu = UVars.Instance.to_array x and yq, yu = UVars.Instance.to_array y in + let c = CArray.fold_left2 + (* TODO strict? *) + (fun c x y -> if Sorts.Quality.equal x y then c else Set.add (QEq (x,y)) c) + c xq yq + in + let c = CArray.fold_left2 + (fun c x y -> Set.add (mk x y) c) + c xu yu + in + c + +let enforce_eq_qualities qs qs' cstrs = + CArray.fold_left2 (fun c a b -> + if Sorts.Quality.equal a b then c else Set.add (QEq (a, b)) c) + cstrs qs qs' + +let compare_cumulative_instances cv_pb variances u u' cstrs = + let make u = Sorts.sort_of_univ @@ Univ.Universe.make u in + let qs, us = UVars.Instance.to_array u + and qs', us' = UVars.Instance.to_array u' in + let cstrs = enforce_eq_qualities qs qs' cstrs in + CArray.fold_left3 + (fun cstrs v u u' -> + let open UVars.Variance in + match v with + | Irrelevant -> Set.add (UWeak (u,u')) cstrs + | Covariant -> + (match cv_pb with + | Conversion.CONV -> Set.add (UEq (make u, make u')) cstrs + | Conversion.CUMUL -> Set.add (ULe (make u, make u')) cstrs) + | Invariant -> + Set.add (UEq (make u, make u')) cstrs) + cstrs variances us us' diff --git a/engine/univProblem.mli b/engine/univProblem.mli index 95556f211bb7..5875d43ad7b6 100644 --- a/engine/univProblem.mli +++ b/engine/univProblem.mli @@ -9,6 +9,7 @@ (************************************************************************) open Univ +open UVars (** {6 Constraints for type inference} @@ -20,6 +21,8 @@ open Univ *) type t = + | QEq of Sorts.Quality.t * Sorts.Quality.t + | QLeq of Sorts.Quality.t * Sorts.Quality.t | ULe of Sorts.t * Sorts.t | UEq of Sorts.t * Sorts.t | ULub of Level.t * Level.t @@ -27,8 +30,6 @@ type t = val is_trivial : t -> bool -val check : UGraph.t -> t -> bool - (** Wrapper around the UGraph function to handle Prop *) val check_eq_level : UGraph.t -> Level.t -> Level.t -> bool @@ -39,10 +40,12 @@ module Set : sig (** Replace ULub constraints by UEq *) val force : t -> t - - val check : UGraph.t -> t -> bool end type 'a constraint_function = 'a -> 'a -> Set.t -> Set.t val enforce_eq_instances_univs : bool -> Instance.t constraint_function + +val enforce_eq_qualities : Sorts.Quality.t array constraint_function + +val compare_cumulative_instances : Conversion.conv_pb -> Variance.t array -> Instance.t constraint_function diff --git a/engine/univSubst.ml b/engine/univSubst.ml index ec5d4efd679b..62b2aa398393 100644 --- a/engine/univSubst.ml +++ b/engine/univSubst.ml @@ -12,14 +12,15 @@ open Sorts open Util open Constr open Univ +open UVars type 'a universe_map = 'a Level.Map.t type universe_subst = Universe.t universe_map type universe_subst_fn = Level.t -> Universe.t option type universe_level_subst_fn = Level.t -> Level.t -let subst_instance fn i = - Instance.of_array (Array.Smart.map fn (Instance.to_array i)) +type quality_subst = Quality.t QVar.Map.t +type quality_subst_fn = QVar.t -> Quality.t let subst_univs_universe fn ul = let addn n u = iterate Universe.super n u in @@ -158,7 +159,7 @@ let level_subst_of f = | Some l -> l let subst_univs_fn_puniverses f (c, u as cu) = - let u' = subst_instance f u in + let u' = Instance.subst_fn f u in if u' == u then cu else (c, u') let nf_binder_annot frel na = @@ -167,7 +168,9 @@ let nf_binder_annot frel na = if rel' == na.binder_relevance then na else { binder_name = na.binder_name; binder_relevance = rel' } -let nf_evars_and_universes_opt_subst fevar flevel fsort frel c = +let nf_evars_and_universes_opt_subst fevar fqual funiv c = + let frel = Sorts.relevance_subst_fn fqual in + let flevel = fqual, level_subst_of funiv in let rec aux c = match kind c with | Evar (evk, args) -> @@ -185,10 +188,10 @@ let nf_evars_and_universes_opt_subst fevar flevel fsort frel c = let pu' = subst_univs_fn_puniverses flevel pu in if pu' == pu then c else mkConstructU pu' | Sort s -> - let s' = fsort s in + let s' = Sorts.subst_fn (fqual, subst_univs_universe funiv) s in if s' == s then c else mkSort s' | Case (ci,u,pms,p,iv,t,br) -> - let u' = subst_instance flevel u in + let u' = Instance.subst_fn flevel u in let ci' = let rel' = frel ci.ci_relevance in if rel' == ci.ci_relevance then ci else { ci with ci_relevance = rel' } @@ -201,7 +204,7 @@ let nf_evars_and_universes_opt_subst fevar flevel fsort frel c = if ci' == ci && u' == u && pms' == pms && p' == p && iv' == iv && t' == t && br' == br then c else mkCase (ci', u', pms', p', iv', t', br') | Array (u,elems,def,ty) -> - let u' = subst_instance flevel u in + let u' = Instance.subst_fn flevel u in let elems' = CArray.Smart.map aux elems in let def' = aux def in let ty' = aux ty in @@ -234,6 +237,11 @@ let nf_evars_and_universes_opt_subst fevar flevel fsort frel c = let rc' = aux_rec rc in if rc' == rc then c else mkCoFix (i, rc') + | Proj (p, r, v) -> + let r' = frel r in + let v' = aux v in + if r' == r && v' == v then c + else mkProj (p, r', v') | _ -> Constr.map aux c and aux_rec ((nas, tys, bds) as rc) = let nas' = Array.Smart.map (fun na -> nf_binder_annot frel na) nas in diff --git a/engine/univSubst.mli b/engine/univSubst.mli index a5c29e894dec..3b74122c8140 100644 --- a/engine/univSubst.mli +++ b/engine/univSubst.mli @@ -10,29 +10,31 @@ open Constr open Univ +open Sorts type 'a universe_map = 'a Level.Map.t type universe_subst = Universe.t universe_map type universe_subst_fn = Level.t -> Universe.t option type universe_level_subst_fn = Level.t -> Level.t +type quality_subst = Quality.t QVar.Map.t +type quality_subst_fn = QVar.t -> Quality.t + val level_subst_of : universe_subst_fn -> universe_level_subst_fn (** The resulting function must never be called on a level which would produce an algebraic. *) val subst_univs_constraints : universe_subst_fn -> Constraints.t -> Constraints.t -val subst_instance : universe_level_subst_fn -> Instance.t -> Instance.t val nf_binder_annot : (Sorts.relevance -> Sorts.relevance) -> 'a Context.binder_annot -> 'a Context.binder_annot (** Full universes substitutions into terms *) -val nf_evars_and_universes_opt_subst : - (existential -> constr option) -> - (Level.t -> Level.t) -> - (Sorts.t -> Sorts.t) -> - (Sorts.relevance -> Sorts.relevance) -> - constr -> constr +val nf_evars_and_universes_opt_subst + : (existential -> constr option) + -> quality_subst_fn + -> universe_subst_fn + -> constr -> constr val subst_univs_universe : universe_subst_fn -> Universe.t -> Universe.t diff --git a/interp/constrexpr.mli b/interp/constrexpr.mli index 1c0fede07ef3..d632c1e8dd84 100644 --- a/interp/constrexpr.mli +++ b/interp/constrexpr.mli @@ -20,16 +20,26 @@ type sort_name_expr = | CRawType of Univ.Level.t (** Universes like "foo.1" have no qualid form *) type univ_level_expr = sort_name_expr Glob_term.glob_sort_gen -type sort_expr = (Sorts.QVar.t option * (sort_name_expr * int) list) Glob_term.glob_sort_gen -type instance_expr = univ_level_expr list +type qvar_expr = + | CQVar of qualid + | CQAnon of Loc.t option + | CRawQVar of Sorts.QVar.t + +type quality_expr = + | CQConstant of Sorts.Quality.constant + | CQualVar of qvar_expr + +type sort_expr = (qvar_expr option * (sort_name_expr * int) list) Glob_term.glob_sort_gen + +type instance_expr = quality_expr list * univ_level_expr list (** Constraints don't have anonymous universes *) type univ_constraint_expr = sort_name_expr * Univ.constraint_type * sort_name_expr -type universe_decl_expr = (lident list, univ_constraint_expr list) UState.gen_universe_decl +type universe_decl_expr = (lident list, lident list, univ_constraint_expr list) UState.gen_universe_decl type cumul_univ_decl_expr = - ((lident * Univ.Variance.t option) list, univ_constraint_expr list) UState.gen_universe_decl + (lident list, (lident * UVars.Variance.t option) list, univ_constraint_expr list) UState.gen_universe_decl type ident_decl = lident * universe_decl_expr option type cumul_ident_decl = lident * cumul_univ_decl_expr option diff --git a/interp/constrexpr_ops.ml b/interp/constrexpr_ops.ml index 6394979dea96..a82019651e39 100644 --- a/interp/constrexpr_ops.ml +++ b/interp/constrexpr_ops.ml @@ -29,12 +29,18 @@ let sort_name_expr_eq c1 c2 = match c1, c2 with | CRawType u1, CRawType u2 -> Univ.Level.equal u1 u2 | (CSProp|CProp|CSet|CType _|CRawType _), _ -> false +let qvar_expr_eq c1 c2 = match c1, c2 with + | CQVar q1, CQVar q2 -> Libnames.qualid_eq q1 q2 + | CQAnon _, CQAnon _ -> true + | CRawQVar q1, CRawQVar q2 -> Sorts.QVar.equal q1 q2 + | (CQVar _ | CQAnon _ | CRawQVar _), _ -> false + let univ_level_expr_eq u1 u2 = Glob_ops.glob_sort_gen_eq sort_name_expr_eq u1 u2 let sort_expr_eq u1 u2 = let eq (q1, l1) (q2, l2) = - Option.equal Sorts.QVar.equal q1 q2 && + Option.equal qvar_expr_eq q1 q2 && List.equal (fun (x,m) (y,n) -> sort_name_expr_eq x y && Int.equal m n) l1 l2 in Glob_ops.glob_sort_gen_eq eq u1 u2 diff --git a/interp/constrextern.ml b/interp/constrextern.ml index 3eedfd174e85..15e51e220ca5 100644 --- a/interp/constrextern.ml +++ b/interp/constrextern.ml @@ -849,7 +849,7 @@ let extern_float f scopes = type extern_env = Id.Set.t * UnivNames.universe_binders let extern_env env sigma = vars_of_env env, Evd.universe_binders sigma -let empty_extern_env = Id.Set.empty, Id.Map.empty +let empty_extern_env = Id.Set.empty, UnivNames.empty_binders let extern_glob_sort_name uvars = function | GSProp -> CSProp @@ -862,9 +862,19 @@ let extern_glob_sort_name uvars = function | None -> CRawType u end +let extern_glob_qvar uvars = function + | GLocalQVar {v=Anonymous;loc} -> CQAnon loc + | GLocalQVar {v=Name id; loc} -> CQVar (qualid_of_ident ?loc id) + | GRawQVar q -> CRawQVar q + | GQVar q -> CRawQVar q + +let extern_glob_quality uvars = function + | GQConstant q -> CQConstant q + | GQualVar q -> CQualVar (extern_glob_qvar uvars q) + let extern_glob_sort uvars u = let map (q, l) = - q, List.map (on_fst (extern_glob_sort_name uvars)) l + Option.map (extern_glob_qvar uvars) q, List.map (on_fst (extern_glob_sort_name uvars)) l in map_glob_sort_gen map u @@ -876,8 +886,10 @@ let extern_glob_sort uvars = function | UNamed _ | UAnonymous _ as u -> extern_glob_sort uvars u let extern_instance uvars = function - | Some l when !print_universes -> - Some (List.map (map_glob_sort_gen (extern_glob_sort_name uvars)) l) + | Some (ql,ul) when !print_universes -> + let ql = List.map (extern_glob_quality uvars) ql in + let ul = List.map (map_glob_sort_gen (extern_glob_sort_name uvars)) ul in + Some (ql,ul) | _ -> None let extern_ref (vars,uvars) ref us = diff --git a/interp/constrintern.ml b/interp/constrintern.ml index 66908e93d7f4..c9ea7c558d37 100644 --- a/interp/constrintern.ml +++ b/interp/constrintern.ml @@ -234,9 +234,9 @@ let contract_curly_brackets_pat ntn (l,ll,bl) = (* side effect; don't inline *) (InConstrEntry,!ntn'),(l,ll,bl) -type local_univs = { bound : Univ.Level.t Id.Map.t; unb_univs : bool } +type local_univs = { bound : UnivNames.universe_binders; unb_univs : bool } -let empty_local_univs = { bound = Id.Map.empty; unb_univs = false } +let empty_local_univs = { bound = UnivNames.empty_binders; unb_univs = false } type abstraction_kind = AbsLambda | AbsPi @@ -1104,7 +1104,7 @@ let string_of_ty = function | Variable -> "var" let gvar (loc, id) us = match us with - | None | Some [] -> DAst.make ?loc @@ GVar id + | None | Some ([],[]) -> DAst.make ?loc @@ GVar id | Some _ -> user_err ?loc (str "Variable " ++ Id.print id ++ str " cannot have a universe instance") @@ -1187,7 +1187,7 @@ let intern_sort_name ~local_univs = function | CType qid -> let is_id = qualid_is_ident qid in let local = if not is_id then None - else Id.Map.find_opt (qualid_basename qid) local_univs.bound + else Id.Map.find_opt (qualid_basename qid) (snd local_univs.bound) in match local with | Some u -> GUniv u @@ -1197,18 +1197,41 @@ let intern_sort_name ~local_univs = function if is_id && local_univs.unb_univs then GLocalUniv (CAst.make ?loc:qid.loc (qualid_basename qid)) else - CErrors.user_err Pp.(str "Undeclared universe " ++ pr_qualid qid ++ str".") + CErrors.user_err ?loc:qid.loc Pp.(str "Undeclared universe " ++ pr_qualid qid ++ str".") + +let intern_qvar ~local_univs = function + | CQAnon loc -> GLocalQVar (CAst.make ?loc Anonymous) + | CRawQVar q -> GRawQVar q + | CQVar qid -> + let is_id = qualid_is_ident qid in + let local = if not is_id then None + else Id.Map.find_opt (qualid_basename qid) (fst local_univs.bound) + in + match local with + | Some u -> GQVar u + | None -> + if is_id && local_univs.unb_univs + then GLocalQVar (CAst.make ?loc:qid.loc (Name (qualid_basename qid))) + else + CErrors.user_err ?loc:qid.loc Pp.(str "Undeclared quality " ++ pr_qualid qid ++ str".") + +let intern_quality ~local_univs q = + match q with + | CQConstant q -> GQConstant q + | CQualVar q -> GQualVar (intern_qvar ~local_univs q) let intern_sort ~local_univs s = let map (q, l) = - (* No user-facing syntax for qualities *) - let () = assert (Option.is_empty q) in - None, List.map (on_fst (intern_sort_name ~local_univs)) l + Option.map (intern_qvar ~local_univs) q, List.map (on_fst (intern_sort_name ~local_univs)) l in map_glob_sort_gen map s -let intern_instance ~local_univs us = - Option.map (List.map (map_glob_sort_gen (intern_sort_name ~local_univs))) us +let intern_instance ~local_univs = function + | None -> None + | Some (qs, us) -> + let qs = List.map (intern_quality ~local_univs) qs in + let us = List.map (map_glob_sort_gen (intern_sort_name ~local_univs)) us in + Some (qs, us) let intern_name_alias = function | { CAst.v = CRef(qid,u) } -> @@ -1326,8 +1349,9 @@ let intern_qualid ?(no_secvar=false) qid intern env ntnvars us args = DAst.make ?loc @@ GApp (DAst.make ?loc:loc' @@ GRef (ref, us), arg) | _ -> err () end - | Some [s], GSort (UAnonymous {rigid=UnivRigid}) -> DAst.make ?loc @@ GSort (glob_sort_of_level s) - | Some [_old_level], GSort _new_sort -> + | Some ([],[s]), GSort (UAnonymous {rigid=UnivRigid}) -> + DAst.make ?loc @@ GSort (glob_sort_of_level s) + | Some ([],[_old_level]), GSort _new_sort -> (* TODO: add old_level and new_sort to the error message *) user_err ?loc (str "Cannot change universe level of notation " ++ pr_qualid qid) | Some _, _ -> err () @@ -2907,21 +2931,26 @@ let interp_univ_constraints env evd cstrs = with UGraph.UniverseInconsistency e as exn -> let _, info = Exninfo.capture exn in CErrors.user_err ~info - (UGraph.explain_universe_inconsistency (Termops.pr_evd_level evd) e) + (UGraph.explain_universe_inconsistency (Termops.pr_evd_qvar evd) (Termops.pr_evd_level evd) e) in List.fold_left interp (evd,Univ.Constraints.empty) cstrs let interp_univ_decl env decl = let open UState in - let binders : lident list = decl.univdecl_instance in let evd = Evd.from_env env in + let evd, qualities = List.fold_left_map (fun evd lid -> + Evd.new_quality_variable ?loc:lid.loc ~name:lid.v evd) + evd + decl.univdecl_qualities + in let evd, instance = List.fold_left_map (fun evd lid -> Evd.new_univ_level_variable ?loc:lid.loc univ_rigid ~name:lid.v evd) evd - binders + decl.univdecl_instance in let evd, cstrs = interp_univ_constraints env evd decl.univdecl_constraints in let decl = { + univdecl_qualities = qualities; univdecl_instance = instance; univdecl_extensible_instance = decl.univdecl_extensible_instance; univdecl_constraints = cstrs; @@ -2934,6 +2963,11 @@ let interp_cumul_univ_decl env decl = let binders = List.map fst decl.univdecl_instance in let variances = Array.map_of_list snd decl.univdecl_instance in let evd = Evd.from_env env in + let evd, qualities = List.fold_left_map (fun evd lid -> + Evd.new_quality_variable ?loc:lid.loc ~name:lid.v evd) + evd + decl.univdecl_qualities + in let evd, instance = List.fold_left_map (fun evd lid -> Evd.new_univ_level_variable ?loc:lid.loc univ_rigid ~name:lid.v evd) evd @@ -2941,6 +2975,7 @@ let interp_cumul_univ_decl env decl = in let evd, cstrs = interp_univ_constraints env evd decl.univdecl_constraints in let decl = { + univdecl_qualities = qualities; univdecl_instance = instance; univdecl_extensible_instance = decl.univdecl_extensible_instance; univdecl_constraints = cstrs; diff --git a/interp/constrintern.mli b/interp/constrintern.mli index c1ad7f6a6f46..6eea239a5dac 100644 --- a/interp/constrintern.mli +++ b/interp/constrintern.mli @@ -150,7 +150,7 @@ val intern_reference : qualid -> GlobRef.t option (** Returns None if not a reference or a abbrev not bound to a name *) val intern_name_alias : - constr_expr -> (GlobRef.t * Glob_term.glob_level list option) option + constr_expr -> (GlobRef.t * Glob_term.glob_instance option) option (** Expands abbreviations; raise an error if not existing *) val interp_reference : ltac_sign -> qualid -> glob_constr diff --git a/interp/impargs.ml b/interp/impargs.ml index 094701c57ea5..70cb0370243e 100644 --- a/interp/impargs.ml +++ b/interp/impargs.ml @@ -222,7 +222,7 @@ let add_free_rels_until strict strongly_strict revpat bound env sigma m pos acc | App (f,_) when rig && is_flexible_reference env sigma bound depth f -> if strict then () else iter_with_full_binders env sigma push_lift (frec false) ed c - | Proj (p, _) when rig -> + | Proj (p, _, _) when rig -> if strict then () else iter_with_full_binders env sigma push_lift (frec false) ed c | Case _ when rig -> @@ -241,7 +241,7 @@ let rec is_rigid_head sigma t = match kind sigma t with | Rel _ | Evar _ -> false | Ind _ | Const _ | Var _ | Sort _ -> true | Case (_,_,_,_,_,f,_) -> is_rigid_head sigma f - | Proj (p,c) -> true + | Proj _ -> true | App (f,args) -> (match kind sigma f with | Fix ((fi,i),_) -> is_rigid_head sigma (args.(fi.(i))) diff --git a/interp/modintern.ml b/interp/modintern.ml index 2dd31ee5b43a..f3ccda0beb04 100644 --- a/interp/modintern.ml +++ b/interp/modintern.ml @@ -137,8 +137,8 @@ let interp_with_decl env base kind = function let poly = lookup_polymorphism env base kind fqid in begin match fst (UState.check_univ_decl ~poly ectx udecl) with | UState.Polymorphic_entry ctx -> - let inst, ctx = Univ.abstract_universes ctx in - let c = EConstr.Vars.subst_univs_level_constr (Univ.make_instance_subst inst) c in + let inst, ctx = UVars.abstract_universes ctx in + let c = EConstr.Vars.subst_univs_level_constr (UVars.make_instance_subst inst) c in let c = EConstr.to_constr sigma c in WithDef (fqid,(c, Some ctx)), Univ.ContextSet.empty | UState.Monomorphic_entry ctx -> diff --git a/interp/notation_ops.ml b/interp/notation_ops.ml index de4b2933d17b..22316c33cb66 100644 --- a/interp/notation_ops.ml +++ b/interp/notation_ops.ml @@ -87,13 +87,24 @@ let compare_glob_universe_instances lt strictly_lt us1 us2 = | None, None -> true | Some _, None -> strictly_lt := true; lt | None, Some _ -> false - | Some l1, Some l2 -> - CList.for_all2eq (fun u1 u2 -> + | Some (ql1,ul1), Some (ql2,ul2) -> + let is_anon = function + | GQualVar (GLocalQVar {v=Anonymous}) -> true + | _ -> false + in + CList.for_all2eq (fun q1 q2 -> + match is_anon q1, is_anon q2 with + | true, true -> true + | false, true -> strictly_lt := true; lt + | true, false -> false + | false, false -> glob_quality_eq q1 q2) + ql1 ql2 + && CList.for_all2eq (fun u1 u2 -> match u1, u2 with | UAnonymous {rigid}, UAnonymous {rigid=rigid'} -> eq_rigid rigid rigid' | UNamed _, UAnonymous _ -> strictly_lt := true; lt | UAnonymous _, UNamed _ -> false - | UNamed _, UNamed _ -> glob_level_eq u1 u2) l1 l2 + | UNamed _, UNamed _ -> glob_level_eq u1 u2) ul1 ul2 (* Compute us1 <= us2, as a boolean *) let compare_glob_universe_instances_le us1 us2 = @@ -796,7 +807,7 @@ let rec subst_notation_constr subst bound raw = if ref' == ref then raw else (match t with | None -> NRef (ref',u) | Some t -> - fst (notation_constr_of_constr bound t.Univ.univ_abstracted_value)) + fst (notation_constr_of_constr bound t.UVars.univ_abstracted_value)) | NVar _ -> raw diff --git a/interp/notation_term.mli b/interp/notation_term.mli index fdd48e03871a..887b544c7da4 100644 --- a/interp/notation_term.mli +++ b/interp/notation_term.mli @@ -21,10 +21,10 @@ open Glob_term type notation_constr = (* Part common to [glob_constr] and [cases_pattern] *) - | NRef of GlobRef.t * glob_level list option + | NRef of GlobRef.t * glob_instance option | NVar of Id.t | NApp of notation_constr * notation_constr list - | NProj of (Constant.t * glob_level list option) * notation_constr list * notation_constr + | NProj of (Constant.t * glob_instance option) * notation_constr list * notation_constr | NHole of Evar_kinds.t * Namegen.intro_pattern_naming_expr | NGenarg of Genarg.glob_generic_argument | NList of Id.t * Id.t * notation_constr * notation_constr * (* associativity: *) bool diff --git a/kernel/cClosure.ml b/kernel/cClosure.ml index 9c3f20903c1f..554036240fce 100644 --- a/kernel/cClosure.ml +++ b/kernel/cClosure.ml @@ -71,7 +71,7 @@ type red_state = Ntrl | Cstr | Red let neutr = function Ntrl -> Ntrl | Red | Cstr -> Red let is_red = function Red -> true | Ntrl | Cstr -> false -type table_key = Constant.t Univ.puniverses tableKey +type table_key = Constant.t UVars.puniverses tableKey type evar_repack = Evar.t * constr list -> constr @@ -87,24 +87,24 @@ and fterm = | FInd of pinductive | FConstruct of pconstructor | FApp of fconstr * fconstr array - | FProj of Projection.t * fconstr + | FProj of Projection.t * Sorts.relevance * fconstr | FFix of fixpoint * usubs | FCoFix of cofixpoint * usubs - | FCaseT of case_info * Univ.Instance.t * constr array * case_return * fconstr * case_branch array * usubs (* predicate and branches are closures *) - | FCaseInvert of case_info * Univ.Instance.t * constr array * case_return * finvert * fconstr * case_branch array * usubs + | FCaseT of case_info * UVars.Instance.t * constr array * case_return * fconstr * case_branch array * usubs (* predicate and branches are closures *) + | FCaseInvert of case_info * UVars.Instance.t * constr array * case_return * finvert * fconstr * case_branch array * usubs | FLambda of int * (Name.t Context.binder_annot * constr) list * constr * usubs | FProd of Name.t Context.binder_annot * fconstr * constr * usubs | FLetIn of Name.t Context.binder_annot * fconstr * fconstr * constr * usubs | FEvar of Evar.t * constr list * usubs * evar_repack | FInt of Uint63.t | FFloat of Float64.t - | FArray of Univ.Instance.t * fconstr Parray.t * fconstr + | FArray of UVars.Instance.t * fconstr Parray.t * fconstr | FLIFT of int * fconstr | FCLOS of constr * usubs | FIrrelevant | FLOCKED -and usubs = fconstr subs Univ.puniverses +and usubs = fconstr subs UVars.puniverses and finvert = fconstr array @@ -137,12 +137,14 @@ let info_flags info = info.i_flags let info_env info = info.i_cache.i_env let info_univs info = info.i_cache.i_univs -let push_relevance infos r = - { infos with i_relevances = Range.cons r.binder_relevance infos.i_relevances } +let push_relevance infos x = + { infos with i_relevances = Range.cons x.binder_relevance infos.i_relevances } let push_relevances infos nas = - { infos with i_relevances = Array.fold_left (fun l x -> Range.cons x.binder_relevance l) - infos.i_relevances nas } + { infos with + i_relevances = + Array.fold_left (fun l x -> Range.cons x.binder_relevance l) + infos.i_relevances nas } let set_info_relevances info r = { info with i_relevances = r } @@ -154,8 +156,8 @@ type 'a next_native_args = (CPrimitives.arg_kind * 'a) list type stack_member = | Zapp of fconstr array - | ZcaseT of case_info * Univ.Instance.t * constr array * case_return * case_branch array * usubs - | Zproj of Projection.Repr.t + | ZcaseT of case_info * UVars.Instance.t * constr array * case_return * case_branch array * usubs + | Zproj of Projection.Repr.t * Sorts.relevance | Zfix of fconstr * stack | Zprimitive of CPrimitives.t * pconstant * fconstr list * fconstr next_native_args (* operator, constr def, arguments already seen (in rev order), next arguments *) @@ -239,6 +241,34 @@ let zupdate info m s = Zupdate(m)::s' else s +(* We use empty as a special identity value, if we don't check + subst_instance_instance will raise array out of bounds. *) +let usubst_instance (_,u) u' = + if UVars.Instance.is_empty u then u' + else UVars.subst_instance_instance u u' + +let usubst_punivs (_,u) (v,u' as orig) = + if UVars.Instance.is_empty u then orig + else v, UVars.subst_instance_instance u u' + +let usubst_sort (_,u) s = + if UVars.Instance.is_empty u then s + else UVars.subst_instance_sort u s + +let usubst_relevance (_,u) r = + if UVars.Instance.is_empty u then r + else UVars.subst_instance_relevance u r + +let usubst_binder e x = + let r = x.binder_relevance in + let r' = usubst_relevance e r in + if r == r' then x else { x with binder_relevance = r' } + +let usubst_case_info e ci = + let r = ci.ci_relevance in + let r' = usubst_relevance e r in + if r == r' then ci else { ci with ci_relevance = r' } + let mk_lambda env t = let (rvars,t') = Term.decompose_lambda t in FLambda(List.length rvars, List.rev rvars, t', env) @@ -247,32 +277,14 @@ let usubs_lift (e,u) = subs_lift e, u let usubs_liftn n (e,u) = subs_liftn n e, u +(* t must be a FLambda and binding list cannot be empty *) let destFLambda clos_fun t = match [@ocaml.warning "-4"] t.term with - FLambda(_,[(na,ty)],b,e) -> (na,clos_fun e ty,clos_fun (usubs_lift e) b) - | FLambda(n,(na,ty)::tys,b,e) -> - (na,clos_fun e ty,{mark=t.mark;term=FLambda(n-1,tys,b,usubs_lift e)}) - | _ -> assert false - (* t must be a FLambda and binding list cannot be empty *) - -(* We use empty as a special identity value, if we don't check - subst_instance_instance will raise array out of bounds. *) -let usubst_instance (_,u) u' = - if Univ.Instance.is_empty u then u' - else Univ.subst_instance_instance u u' - -let usubst_punivs (_,u) (v,u' as orig) = - if Univ.Instance.is_empty u then orig - else v, Univ.subst_instance_instance u u' - -let usubst_sort (_,u) s = match s with - | Sorts.Type su -> - if Univ.Instance.is_empty u then s - else Sorts.(sort_of_univ (Univ.subst_instance_universe u su)) - | Sorts.QSort (q, v) -> - if Univ.Instance.is_empty u then s - else Sorts.qsort q (Univ.subst_instance_universe u v) - | Sorts.(SProp | Prop | Set) -> s + | FLambda(_,[(na,ty)],b,e) -> + (usubst_binder e na,clos_fun e ty,clos_fun (usubs_lift e) b) + | FLambda(n,(na,ty)::tys,b,e) -> + (usubst_binder e na,clos_fun e ty,{mark=t.mark;term=FLambda(n-1,tys,b,usubs_lift e)}) + | _ -> assert false (* Optimization: do not enclose variables in a closure. Makes variable access much faster *) @@ -294,7 +306,7 @@ let mk_clos (e:usubs) t = let injectu c u = mk_clos (subs_id 0, u) c -let inject c = injectu c Univ.Instance.empty +let inject c = injectu c UVars.Instance.empty let mk_irrelevant = { mark = Cstr; term = FIrrelevant } @@ -316,7 +328,7 @@ module Table : sig end = struct module Table = Hashtbl.Make(struct type t = table_key - let equal = eq_table_key (eq_pair eq_constant_key Univ.Instance.equal) + let equal = eq_table_key (eq_pair eq_constant_key UVars.Instance.equal) let hash = hash_table_key (fun (c, _) -> Constant.UserOrd.hash c) end) @@ -368,7 +380,7 @@ end = struct raise Not_found | ConstKey (cst,u) -> let cb = lookup_constant cst env in - shortcut_irrelevant info (cb.const_relevance); + shortcut_irrelevant info (UVars.subst_instance_relevance u cb.const_relevance); let ts = RedFlags.red_transparent info.i_flags in if TransparentState.is_transparent_constant ts cst then Def (constant_value_in u cb.const_body) @@ -403,7 +415,9 @@ let mk_clos_vect env v = match v with [|mk_clos env v0; mk_clos env v1; mk_clos env v2; mk_clos env v3|] | v -> Array.Fun1.map mk_clos env v -let rec subst_constr (subst,usubst as e) c = match [@ocaml.warning "-4"] Constr.kind c with +let rec subst_constr (subst,usubst as e) c = + let c = Vars.map_constr_relevance (usubst_relevance e) c in + match [@ocaml.warning "-4"] Constr.kind c with | Rel i -> begin match expand_rel i subst with | Inl (k, lazy v) -> Vars.lift k v @@ -445,6 +459,7 @@ let rec to_constr (lfts, usubst as ulfts) v = let n = Array.length bds in let subs_ty = comp_subs ulfts e in let subs_bd = comp_subs (on_fst (el_liftn n) ulfts) (on_fst (subs_liftn n) e) in + let lna = Array.Fun1.map usubst_binder subs_ty lna in let tys = Array.Fun1.map subst_constr subs_ty tys in let bds = Array.Fun1.map subst_constr subs_bd bds in mkFix (op, (lna, tys, bds)) @@ -455,21 +470,25 @@ let rec to_constr (lfts, usubst as ulfts) v = let n = Array.length bds in let subs_ty = comp_subs ulfts e in let subs_bd = comp_subs (on_fst (el_liftn n) ulfts) (on_fst (subs_liftn n) e) in + let lna = Array.Fun1.map usubst_binder subs_ty lna in let tys = Array.Fun1.map subst_constr subs_ty tys in let bds = Array.Fun1.map subst_constr subs_bd bds in mkCoFix (op, (lna, tys, bds)) | FApp (f,ve) -> mkApp (to_constr ulfts f, Array.Fun1.map to_constr ulfts ve) - | FProj (p,c) -> - mkProj (p,to_constr ulfts c) + | FProj (p,r,c) -> + mkProj (p,usubst_relevance ulfts r,to_constr ulfts c) | FLambda (len, tys, f, e) -> if is_subs_id (fst e) && is_lift_id lfts then subst_instance_constr (usubst_instance ulfts (snd e)) (Term.compose_lam (List.rev tys) f) else let subs = comp_subs ulfts e in - let tys = List.mapi (fun i (na, c) -> na, subst_constr (usubs_liftn i subs) c) tys in + let tys = List.mapi (fun i (na, c) -> + usubst_binder subs na, subst_constr (usubs_liftn i subs) c) + tys + in let f = subst_constr (usubs_liftn len subs) f in Term.compose_lam (List.rev tys) f | FProd (n, t, c, e) -> @@ -477,12 +496,15 @@ let rec to_constr (lfts, usubst as ulfts) v = mkProd (n, to_constr ulfts t, subst_instance_constr (usubst_instance ulfts (snd e)) c) else let subs' = comp_subs ulfts e in - mkProd (n, to_constr ulfts t, subst_constr (usubs_lift subs') c) + mkProd (usubst_binder subs' n, + to_constr ulfts t, + subst_constr (usubs_lift subs') c) | FLetIn (n,b,t,f,e) -> let subs = comp_subs (on_fst el_lift ulfts) (usubs_lift e) in - mkLetIn (n, to_constr ulfts b, - to_constr ulfts t, - subst_constr subs f) + mkLetIn (usubst_binder subs n, + to_constr ulfts b, + to_constr ulfts t, + subst_constr subs f) | FEvar (ev, args, env, repack) -> let subs = comp_subs ulfts env in repack (ev, List.map (fun a -> subst_constr subs a) args) @@ -514,14 +536,18 @@ let rec to_constr (lfts, usubst as ulfts) v = and to_constr_case (lfts,_ as ulfts) ci u pms p iv c ve env = let subs = comp_subs ulfts env in + let ci = usubst_case_info subs ci in if is_subs_id (fst env) && is_lift_id lfts then mkCase (ci, usubst_instance subs u, pms, p, iv, to_constr ulfts c, ve) else let f_ctx (nas, c) = + let nas = Array.map (usubst_binder subs) nas in let c = subst_constr (usubs_liftn (Array.length nas) subs) c in (nas, c) in - mkCase (ci, usubst_instance subs u, Array.map (fun c -> subst_constr subs c) pms, + mkCase (ci, + usubst_instance subs u, + Array.map (fun c -> subst_constr subs c) pms, f_ctx p, iv, to_constr ulfts c, @@ -534,7 +560,7 @@ and comp_subs (el,u) (s,u') = fconstr. When we find a closure whose substitution is the identity, then we directly return the constr to avoid possibly huge reallocation. *) -let term_of_fconstr c = to_constr (el_id, Univ.Instance.empty) c +let term_of_fconstr c = to_constr (el_id, UVars.Instance.empty) c (* fstrong applies unfreeze_fun recursively on the (freeze) term and * yields a term. Assumes that the unfreeze_fun never returns a @@ -551,9 +577,9 @@ let rec zip m stk = let t = FCaseT(ci, u, pms, p, m, br, e) in let mark = (neutr m.mark) in zip {mark; term=t} s - | Zproj p :: s -> + | Zproj (p,r) :: s -> let mark = (neutr m.mark) in - zip {mark; term=FProj(Projection.make p true,m)} s + zip {mark; term=FProj(Projection.make p true,r,m)} s | Zfix(fx,par)::s -> zip fx (par @ append_stack [|m|] s) | Zshift(n)::s -> @@ -813,7 +839,7 @@ let get_branch infos depth ci u pms (ind, c) br e args = @raise Not_found if the inductive is not a primitive record, or if the constructor is partially applied. *) -let eta_expand_ind_stack env ind m s (f, s') = +let eta_expand_ind_stack env (ind,u) m s (f, s') = let open Declarations in let mib = lookup_mind (fst ind) env in (* disallow eta-exp for non-primitive records *) @@ -827,9 +853,9 @@ let eta_expand_ind_stack env ind m s (f, s') = let (depth, args, _s) = strip_update_shift_app m s in (** Try to drop the params, might fail on partially applied constructors. *) let argss = try_drop_parameters depth pars args in - let hstack = Array.map (fun p -> + let hstack = Array.map (fun (p,r) -> { mark = Red; (* right can't be a constructor though *) - term = FProj (Projection.make p true, right) }) + term = FProj (Projection.make p true, UVars.subst_instance_relevance u r, right) }) projs in argss, [Zapp hstack] @@ -873,10 +899,10 @@ let contract_fix_vect fix = in (on_fst (fun env -> mk_subs env 0) env, thisbody) -let unfold_projection info p = +let unfold_projection info p r = if red_projection info.i_flags p then - Some (Zproj (Projection.repr p)) + Some (Zproj (Projection.repr p, r)) else None (************************************************************************) @@ -889,11 +915,11 @@ module FNativeEntries = type elem = fconstr type args = fconstr array type evd = unit - type uinstance = Univ.Instance.t + type uinstance = UVars.Instance.t let mk_construct c = (* All constructors used in primitive functions are relevant *) - { mark = Cstr; term = FConstruct (Univ.in_punivs c) } + { mark = Cstr; term = FConstruct (UVars.in_punivs c) } let get = Array.get @@ -922,7 +948,7 @@ module FNativeEntries = match retro.Retroknowledge.retro_int63 with | Some c -> defined_int := true; - fint := { mark = Ntrl; term = FFlex (ConstKey (Univ.in_punivs c)) } + fint := { mark = Ntrl; term = FFlex (ConstKey (UVars.in_punivs c)) } | None -> defined_int := false let defined_float = ref false @@ -932,7 +958,7 @@ module FNativeEntries = match retro.Retroknowledge.retro_float64 with | Some c -> defined_float := true; - ffloat := { mark = Ntrl; term = FFlex (ConstKey (Univ.in_punivs c)) } + ffloat := { mark = Ntrl; term = FFlex (ConstKey (UVars.in_punivs c)) } | None -> defined_float := false let defined_bool = ref false @@ -983,7 +1009,7 @@ module FNativeEntries = fLt := mk_construct cLt; fGt := mk_construct cGt; let (icmp, _) = cEq in - fcmp := { mark = Ntrl; term = FInd (Univ.in_punivs icmp) } + fcmp := { mark = Ntrl; term = FInd (UVars.in_punivs icmp) } | None -> defined_cmp := false let defined_f_cmp = ref false @@ -1189,8 +1215,9 @@ let rec skip_irrelevant_stack info stk = match stk with | (Zfix _ | Zproj _) :: s -> (* Typing rules ensure that fix / proj over SProp is irrelevant *) skip_irrelevant_stack info s -| ZcaseT (ci, _, _, _, _, _) :: s -> - if is_irrelevant info ci.ci_relevance then skip_irrelevant_stack info s +| ZcaseT (ci, _, _, _, _, e) :: s -> + let r = usubst_relevance e ci.ci_relevance in + if is_irrelevant info r then skip_irrelevant_stack info s else stk | Zprimitive _ :: _ -> assert false (* no irrelevant primitives so far *) | Zupdate m :: s -> @@ -1202,10 +1229,6 @@ let is_irrelevant_constructor infos (ind,_) = match infos.i_cache.i_mode with | Conversion -> Indset_env.mem ind infos.i_cache.i_env.irr_inds | Reduction -> false -let is_irrelevant_projection infos p = match infos.i_cache.i_mode with -| Conversion -> not @@ Projection.Repr.relevant @@ Projection.repr p -| Reduction -> false - (*********************************************************************) (* A machine that inspects the head of a term until it finds an atom or a subterm that may produce a redex (abstraction, @@ -1218,22 +1241,23 @@ let rec knh info m stk = | FLOCKED -> assert false | FApp(a,b) -> knh info a (append_stack b (zupdate info m stk)) | FCaseT(ci,u,pms,p,t,br,e) -> - if is_irrelevant info ci.ci_relevance then + let r = usubst_relevance e ci.ci_relevance in + if is_irrelevant info r then (mk_irrelevant, skip_irrelevant_stack info stk) else knh info t (ZcaseT(ci,u,pms,p,br,e)::zupdate info m stk) - | FFix (((ri, n), (lna, _, _)), _) -> - if is_irrelevant info (lna.(n)).binder_relevance then + | FFix (((ri, n), (lna, _, _)), e) -> + if is_irrelevant info (usubst_relevance e (lna.(n)).binder_relevance) then (mk_irrelevant, skip_irrelevant_stack info stk) else (match get_nth_arg m ri.(n) stk with (Some(pars,arg),stk') -> knh info arg (Zfix(m,pars)::stk') | (None, stk') -> (m,stk')) - | FProj (p,c) -> - if is_irrelevant_projection info p then + | FProj (p,r,c) -> + if is_irrelevant info r then (mk_irrelevant, skip_irrelevant_stack info stk) else - (match unfold_projection info p with + (match unfold_projection info p r with | None -> (m, stk) | Some s -> knh info c (s :: zupdate info m stk)) @@ -1248,35 +1272,42 @@ and knht info e t stk = | App(a,b) -> knht info e a (append_stack (mk_clos_vect e b) stk) | Case(ci,u,pms,p,NoInvert,t,br) -> - if is_irrelevant info ci.ci_relevance then + if is_irrelevant info (usubst_relevance e ci.ci_relevance) then (mk_irrelevant, skip_irrelevant_stack info stk) else knht info e t (ZcaseT(ci, u, pms, p, br, e)::stk) | Case(ci,u,pms,p,CaseInvert{indices},t,br) -> - if is_irrelevant info ci.ci_relevance then + if is_irrelevant info (usubst_relevance e ci.ci_relevance) then (mk_irrelevant, skip_irrelevant_stack info stk) else let term = FCaseInvert (ci, u, pms, p, (Array.map (mk_clos e) indices), mk_clos e t, br, e) in { mark = Red; term }, stk | Fix (((_, n), (lna, _, _)) as fx) -> - if is_irrelevant info (lna.(n)).binder_relevance then + if is_irrelevant info (usubst_relevance e (lna.(n)).binder_relevance) then (mk_irrelevant, skip_irrelevant_stack info stk) else knh info { mark = Cstr; term = FFix (fx, e) } stk | Cast(a,_,_) -> knht info e a stk | Rel n -> knh info (clos_rel (fst e) n) stk - | Proj (p, c) -> knh info { mark = Red; term = FProj (p, mk_clos e c) } stk + | Proj (p, r, c) -> + let r = usubst_relevance e r in + if is_irrelevant info r then + (mk_irrelevant, skip_irrelevant_stack info stk) + else + knh info { mark = Red; term = FProj (p, r, mk_clos e c) } stk | (Ind _|Const _|Construct _|Var _|Meta _ | Sort _ | Int _|Float _) -> (mk_clos e t, stk) - | CoFix cfx -> { mark = Cstr; term = FCoFix (cfx,e) }, stk + | CoFix cfx -> + { mark = Cstr; term = FCoFix (cfx,e) }, stk | Lambda _ -> { mark = Cstr ; term = mk_lambda e t }, stk | Prod (n, t, c) -> - { mark = Ntrl; term = FProd (n, mk_clos e t, c, e) }, stk + { mark = Ntrl; term = FProd (usubst_binder e n, mk_clos e t, c, e) }, stk | LetIn (n,b,t,c) -> - { mark = Red; term = FLetIn (n, mk_clos e b, mk_clos e t, c, e) }, stk + { mark = Red; term = FLetIn (usubst_binder e n, mk_clos e b, mk_clos e t, c, e) }, stk | Evar ev -> begin match info.i_cache.i_sigma.evar_expand ev with | EvarDefined c -> knht info e c stk | EvarUndefined (evk, args) -> + assert (UVars.Instance.is_empty (snd e)); if info.i_cache.i_sigma.evar_relevant ev then let repack = info.i_cache.i_sigma.evar_repack in { mark = Ntrl; term = FEvar (evk, args, e, repack) }, stk @@ -1329,7 +1360,7 @@ let rec knr info tab m stk = let stk' = par @ append_stack [|rarg|] s in let (fxe,fxbd) = contract_fix_vect fx.term in knit info tab fxe fxbd stk' - | (depth, args, Zproj p::s) when use_match -> + | (depth, args, Zproj (p,_)::s) when use_match -> let rargs = drop_parameters depth (Projection.Repr.npars p) args in let rarg = project_nth_arg (Projection.Repr.arg p) rargs in kni info tab rarg s @@ -1339,8 +1370,8 @@ let rec knr info tab m stk = (mk_irrelevant, skip_irrelevant_stack info stk) else (m, stk) - | FCoFix ((i, (lna, _, _)), _) -> - if is_irrelevant info (lna.(i)).binder_relevance then + | FCoFix ((i, (lna, _, _)), e) -> + if is_irrelevant info (usubst_relevance e (lna.(i)).binder_relevance) then (mk_irrelevant, skip_irrelevant_stack info stk) else if red_set info.i_flags fCOFIX then (match strip_update_shift_app m stk with @@ -1468,15 +1499,17 @@ and klt info tab e t = match kind t with | App _ -> assert false end | Lambda (na, u, c) -> + let na' = usubst_binder e na in let u' = klt info tab e u in - let c' = klt (push_relevance info na) tab (usubs_lift e) c in - if u' == u && c' == c then t - else mkLambda (na, u', c') + let c' = klt (push_relevance info na') tab (usubs_lift e) c in + if na' == na && u' == u && c' == c then t + else mkLambda (na', u', c') | Prod (na, u, v) -> + let na' = usubst_binder e na in let u' = klt info tab e u in - let v' = klt (push_relevance info na) tab (usubs_lift e) v in - if u' == u && v' == v then t - else mkProd (na, u', v') + let v' = klt (push_relevance info na') tab (usubs_lift e) v in + if na' == na && u' == u && v' == v then t + else mkProd (na', u', v') | Cast (t, _, _) -> klt info tab e t | Var _ | Const _ | CoFix _ | Fix _ | Evar _ | Case _ | LetIn _ | Proj _ | Array _ -> let share = info.i_cache.i_share in @@ -1492,6 +1525,7 @@ and norm_head info tab m = match m.term with | FLambda(_n,tys,f,e) -> let fold (e, info, ctxt) (na, ty) = + let na = usubst_binder e na in let ty = klt info tab e ty in let info = push_relevance info na in (usubs_lift e, info, (na, ty) :: ctxt) @@ -1500,25 +1534,29 @@ and norm_head info tab m = let bd = klt info tab e' f in List.fold_left (fun b (na,ty) -> mkLambda(na,ty,b)) bd rvtys | FLetIn(na,a,b,f,e) -> + let na = usubst_binder e na in let c = klt (push_relevance info na) tab (usubs_lift e) f in mkLetIn(na, kl info tab a, kl info tab b, c) | FProd(na,dom,rng,e) -> + let na = usubst_binder e na in let rng = klt (push_relevance info na) tab (usubs_lift e) rng in mkProd(na, kl info tab dom, rng) | FCoFix((n,(na,tys,bds)),e) -> + let na = Array.Smart.map (usubst_binder e) na in let infobd = push_relevances info na in let ftys = Array.map (fun ty -> klt info tab e ty) tys in let fbds = Array.map (fun bd -> klt infobd tab (usubs_liftn (Array.length na) e) bd) bds in mkCoFix (n, (na, ftys, fbds)) | FFix((n,(na,tys,bds)),e) -> + let na = Array.Smart.map (usubst_binder e) na in let infobd = push_relevances info na in let ftys = Array.map (fun ty -> klt info tab e ty) tys in let fbds = Array.map (fun bd -> klt infobd tab (usubs_liftn (Array.length na) e) bd) bds in mkFix (n, (na, ftys, fbds)) | FEvar(ev, args, env, repack) -> repack (ev, List.map (fun a -> klt info tab env a) args) - | FProj (p,c) -> - mkProj (p, kl info tab c) + | FProj (p,r,c) -> + mkProj (p, r, kl info tab c) | FArray (u, a, ty) -> let a, def = Parray.to_array a in let a = Array.map (kl info tab) a in @@ -1535,16 +1573,18 @@ and zip_term info tab m stk = match stk with | Zapp args :: s -> zip_term info tab (mkApp(m, Array.map (kl info tab) args)) s | ZcaseT(ci, u, pms, p, br, e) :: s -> - let zip_ctx (nas, c) = + let zip_ctx (nas, c) = + let nas = Array.map (usubst_binder e) nas in let e = usubs_liftn (Array.length nas) e in (nas, klt info tab e c) in + let ci = usubst_case_info e ci in let u = usubst_instance e u in let t = mkCase(ci, u, Array.map (fun c -> klt info tab e c) pms, zip_ctx p, NoInvert, m, Array.map zip_ctx br) in zip_term info tab t s -| Zproj p::s -> - let t = mkProj (Projection.make p true, m) in +| Zproj (p,r)::s -> + let t = mkProj (Projection.make p true, r, m) in zip_term info tab t s | Zfix(fx,par)::s -> let h = mkApp(zip_term info tab (kl info tab fx) par,[|m|]) in diff --git a/kernel/cClosure.mli b/kernel/cClosure.mli index 2bcbbe2cf54e..6ec20e015f8e 100644 --- a/kernel/cClosure.mli +++ b/kernel/cClosure.mli @@ -26,10 +26,12 @@ type finvert type evar_repack -type usubs = fconstr subs Univ.puniverses +type usubs = fconstr subs UVars.puniverses -type table_key = Constant.t Univ.puniverses tableKey +type table_key = Constant.t UVars.puniverses tableKey +(** Relevances (eg in binder_annot or case_info) have NOT been substituted + when there is a usubs field *) type fterm = | FRel of int | FAtom of constr (** Metas and Sorts *) @@ -37,18 +39,18 @@ type fterm = | FInd of pinductive | FConstruct of pconstructor | FApp of fconstr * fconstr array - | FProj of Projection.t * fconstr + | FProj of Projection.t * Sorts.relevance * fconstr | FFix of fixpoint * usubs | FCoFix of cofixpoint * usubs - | FCaseT of case_info * Univ.Instance.t * constr array * case_return * fconstr * case_branch array * usubs (* predicate and branches are closures *) - | FCaseInvert of case_info * Univ.Instance.t * constr array * case_return * finvert * fconstr * case_branch array * usubs + | FCaseT of case_info * UVars.Instance.t * constr array * case_return * fconstr * case_branch array * usubs (* predicate and branches are closures *) + | FCaseInvert of case_info * UVars.Instance.t * constr array * case_return * finvert * fconstr * case_branch array * usubs | FLambda of int * (Name.t Context.binder_annot * constr) list * constr * usubs | FProd of Name.t Context.binder_annot * fconstr * constr * usubs | FLetIn of Name.t Context.binder_annot * fconstr * fconstr * constr * usubs | FEvar of Evar.t * constr list * usubs * evar_repack | FInt of Uint63.t | FFloat of Float64.t - | FArray of Univ.Instance.t * fconstr Parray.t * fconstr + | FArray of UVars.Instance.t * fconstr Parray.t * fconstr | FLIFT of int * fconstr | FCLOS of constr * usubs | FIrrelevant @@ -61,8 +63,8 @@ type 'a next_native_args = (CPrimitives.arg_kind * 'a) list type stack_member = | Zapp of fconstr array - | ZcaseT of case_info * Univ.Instance.t * constr array * case_return * case_branch array * usubs - | Zproj of Projection.Repr.t + | ZcaseT of case_info * UVars.Instance.t * constr array * case_return * case_branch array * usubs + | Zproj of Projection.Repr.t * Sorts.relevance | Zfix of fconstr * stack | Zprimitive of CPrimitives.t * pconstant * fconstr list * fconstr next_native_args (* operator, constr def, arguments already seen (in rev order), next arguments *) @@ -81,7 +83,7 @@ val get_native_args1 : CPrimitives.t -> pconstant -> stack -> val stack_args_size : stack -> int val inductive_subst : Declarations.mutual_inductive_body - -> Univ.Instance.t + -> UVars.Instance.t -> fconstr array -> usubs @@ -90,7 +92,9 @@ val usubs_liftn : int -> usubs -> usubs val usubs_cons : fconstr -> usubs -> usubs (** identity if the first instance is empty *) -val usubst_instance : 'a Univ.puniverses -> Univ.Instance.t -> Univ.Instance.t +val usubst_instance : 'a UVars.puniverses -> UVars.Instance.t -> UVars.Instance.t + +val usubst_binder : _ UVars.puniverses -> 'a Context.binder_annot -> 'a Context.binder_annot (** To lazy reduce a constr, create a [clos_infos] with [create_clos_infos], inject the term to reduce with [inject]; then use @@ -123,7 +127,7 @@ val create_tab : unit -> clos_tab val info_env : clos_infos -> env val info_flags: clos_infos -> reds val info_univs : clos_infos -> UGraph.t -val unfold_projection : clos_infos -> Projection.t -> stack_member option +val unfold_projection : clos_infos -> Projection.t -> Sorts.relevance -> stack_member option val push_relevance : clos_infos -> 'b Context.binder_annot -> clos_infos val push_relevances : clos_infos -> 'b Context.binder_annot array -> clos_infos @@ -131,6 +135,8 @@ val set_info_relevances : clos_infos -> Sorts.relevance Range.t -> clos_infos val info_relevances : clos_infos -> Sorts.relevance Range.t +val is_irrelevant : clos_infos -> Sorts.relevance -> bool + val infos_with_reds : clos_infos -> reds -> clos_infos (** Reduction function *) @@ -162,7 +168,7 @@ val eta_expand_stack : clos_infos -> Name.t Context.binder_annot -> stack -> sta @raise Not_found if the inductive is not a primitive record, or if the constructor is partially applied. *) -val eta_expand_ind_stack : env -> inductive -> fconstr -> stack -> +val eta_expand_ind_stack : env -> pinductive -> fconstr -> stack -> (fconstr * stack) -> stack * stack (** Conversion auxiliary functions to do step by step normalisation *) @@ -199,6 +205,6 @@ val zip : fconstr -> stack -> fconstr val term_of_process : fconstr -> stack -> constr -val to_constr : lift Univ.puniverses -> fconstr -> constr +val to_constr : lift UVars.puniverses -> fconstr -> constr (** End of cbn debug section i*) diff --git a/kernel/cPrimitives.ml b/kernel/cPrimitives.ml index 691e13bd9087..6a9fcf679e23 100644 --- a/kernel/cPrimitives.ml +++ b/kernel/cPrimitives.ml @@ -12,6 +12,7 @@ (* number of primitives is changed. *) open Univ +open UVars type t = | Int63head0 @@ -278,7 +279,7 @@ and ind_or_type = | PITT_param : int -> ind_or_type (* DeBruijn index referring to prenex type quantifiers *) let one_univ = - AbstractContext.make Names.[|Name (Id.of_string "u")|] Constraints.empty + AbstractContext.make ([||],Names.[|Name (Id.of_string "u")|]) Constraints.empty let typ_univs (type a) (t : a prim_type) = match t with | PT_int63 -> AbstractContext.empty @@ -295,7 +296,7 @@ let types = let array_ty = PITT_type (PT_array, - (Instance.of_array [|Level.var 0|], + (Instance.of_array ([||],[|Level.var 0|]), PITT_param 1)) in function diff --git a/kernel/cPrimitives.mli b/kernel/cPrimitives.mli index b1af775bef0e..9e704fa16f6c 100644 --- a/kernel/cPrimitives.mli +++ b/kernel/cPrimitives.mli @@ -98,7 +98,7 @@ val kind : t -> args_red type 'a prim_type = | PT_int63 : unit prim_type | PT_float64 : unit prim_type - | PT_array : (Univ.Instance.t * ind_or_type) prim_type + | PT_array : (UVars.Instance.t * ind_or_type) prim_type and 'a prim_ind = | PIT_bool : unit prim_ind @@ -113,7 +113,7 @@ and ind_or_type = | PITT_type : 'a prim_type * 'a -> ind_or_type | PITT_param : int -> ind_or_type (* DeBruijn index referring to prenex type quantifiers *) -val typ_univs : 'a prim_type -> Univ.AbstractContext.t +val typ_univs : 'a prim_type -> UVars.AbstractContext.t type prim_type_ex = PTE : 'a prim_type -> prim_type_ex @@ -128,7 +128,7 @@ type op_or_type = | OT_type : 'a prim_type -> op_or_type | OT_const of const -val op_or_type_univs : op_or_type -> Univ.AbstractContext.t +val op_or_type_univs : op_or_type -> UVars.AbstractContext.t val prim_ind_to_string : 'a prim_ind -> string @@ -139,7 +139,7 @@ val op_or_type_to_string : op_or_type -> string val parse_op_or_type : ?loc:Loc.t -> string -> op_or_type -val univs : t -> Univ.AbstractContext.t +val univs : t -> UVars.AbstractContext.t val types : t -> Constr.rel_context * ind_or_type list * ind_or_type (** Parameters * Reduction relevant arguments * output type diff --git a/kernel/constant_typing.ml b/kernel/constant_typing.ml index 075e53087b6c..0fdeda52791a 100644 --- a/kernel/constant_typing.ml +++ b/kernel/constant_typing.ml @@ -21,6 +21,7 @@ open Declarations open Environ open Entries open Univ +open UVars module NamedDecl = Context.Named.Declaration @@ -95,17 +96,17 @@ let rec unzip ctx j = unzip ctx { j with uj_val = mkApp (mkLambda (n,ty,j.uj_val),arg) } type typing_context = - TyCtx of Environ.env * unsafe_type_judgment * Id.Set.t * universe_level_subst * universes + TyCtx of Environ.env * unsafe_type_judgment * Id.Set.t * UVars.sort_level_subst * universes let process_universes env = function | Entries.Monomorphic_entry -> - env, Univ.empty_level_subst, Univ.Instance.empty, Monomorphic + env, UVars.empty_sort_subst, UVars.Instance.empty, Monomorphic | Entries.Polymorphic_entry uctx -> (** [ctx] must contain local universes, such that it has no impact on the rest of the graph (up to transitivity). *) let env = Environ.push_context ~strict:false uctx env in - let inst, auctx = Univ.abstract_universes uctx in - let usubst = Univ.make_instance_subst inst in + let inst, auctx = UVars.abstract_universes uctx in + let usubst = UVars.make_instance_subst inst in env, usubst, inst, Polymorphic auctx let check_primitive_type env op_t u t = @@ -172,7 +173,7 @@ let infer_primitive env { prim_entry_type = utyp; prim_entry_content = p; } = let make_univ_hyps = function | None -> Instance.empty - | Some us -> Instance.of_array us + | Some us -> us let infer_parameter ~sec_univs env entry = let env, usubst, _, univs = process_universes env entry.parameter_entry_universes in @@ -189,7 +190,7 @@ let infer_parameter ~sec_univs env entry = const_type = typ; const_body_code = tps; const_universes = univs; - const_relevance = r; + const_relevance = UVars.subst_sort_level_relevance usubst r; const_inline_code = false; const_typing_flags = Environ.typing_flags env; } @@ -254,12 +255,12 @@ let check_delayed (type a) (handle : a effect_handler) tyenv (body : a proof_out let uctx = ContextSet.union uctx uctx' in let env, univs = match univs with | Monomorphic -> - assert (is_empty_level_subst usubst); + assert (UVars.is_empty_sort_subst usubst); push_context_set uctx env, Opaqueproof.PrivateMonomorphic uctx | Polymorphic _ -> assert (Int.equal valid_signatures 0); push_subgraph uctx env, - let private_univs = on_snd (subst_univs_level_constraints usubst) uctx in + let private_univs = on_snd (subst_univs_level_constraints (snd usubst)) uctx in Opaqueproof.PrivatePolymorphic private_univs in (* Note: non-trivial trusted side-effects only in monomorphic case *) diff --git a/kernel/constant_typing.mli b/kernel/constant_typing.mli index 520317386ec5..d5c79f9f829b 100644 --- a/kernel/constant_typing.mli +++ b/kernel/constant_typing.mli @@ -30,11 +30,11 @@ val infer_local_def : env -> Id.t -> section_def_entry -> val infer_local_assum : env -> types -> types * Sorts.relevance val infer_constant : - sec_univs:Univ.Level.t array option -> env -> constant_entry -> + sec_univs:UVars.Instance.t option -> env -> constant_entry -> 'a pconstant_body val infer_opaque : - sec_univs:Univ.Level.t array option -> env -> 'a opaque_entry -> + sec_univs:UVars.Instance.t option -> env -> 'a opaque_entry -> unit pconstant_body * typing_context val check_delayed : 'a effect_handler -> typing_context -> 'a proof_output -> (Constr.t * Univ.ContextSet.t Opaqueproof.delayed_universes) diff --git a/kernel/constr.ml b/kernel/constr.ml index 8604b316ab04..c674aab25b86 100644 --- a/kernel/constr.ml +++ b/kernel/constr.ml @@ -27,7 +27,7 @@ open Util open Names -open Univ +open UVars open Context type existential_key = Evar.t @@ -76,7 +76,7 @@ type ('constr, 'types) pfixpoint = (int array * int) * ('constr, 'types) prec_declaration type ('constr, 'types) pcofixpoint = int * ('constr, 'types) prec_declaration -type 'a puniverses = 'a Univ.puniverses +type 'a puniverses = 'a UVars.puniverses type pconstant = Constant.t puniverses type pinductive = inductive puniverses type pconstructor = constructor puniverses @@ -110,7 +110,7 @@ type ('constr, 'types, 'sort, 'univs) kind_of_term = | Case of case_info * 'univs * 'constr array * 'types pcase_return * 'constr pcase_invert * 'constr * 'constr pcase_branch array | Fix of ('constr, 'types) pfixpoint | CoFix of ('constr, 'types) pcofixpoint - | Proj of Projection.t * 'constr + | Proj of Projection.t * Sorts.relevance * 'constr | Int of Uint63.t | Float of Float64.t | Array of 'univs * 'constr array * 'constr * 'types @@ -203,14 +203,14 @@ let mkLambda (x,t1,t2) = of_kind @@ Lambda (x,t1,t2) let mkLetIn (x,c1,t,c2) = of_kind @@ LetIn (x,c1,t,c2) let map_puniverses f (x,u) = (f x, u) -let in_punivs a = (a, Univ.Instance.empty) +let in_punivs a = (a, UVars.Instance.empty) (* Constructs a constant *) let mkConst c = of_kind @@ Const (in_punivs c) let mkConstU c = of_kind @@ Const c (* Constructs an applied projection *) -let mkProj (p,c) = of_kind @@ Proj (p,c) +let mkProj (p,r,c) = of_kind @@ Proj (p,r,c) (* Constructs an existential variable *) let mkEvar e = of_kind @@ Evar e @@ -440,7 +440,7 @@ let destCase c = match kind c with | _ -> raise DestKO let destProj c = match kind c with - | Proj (p, c) -> (p, c) + | Proj (p, r, c) -> (p, r, c) | _ -> raise DestKO let destFix c = match kind c with @@ -452,7 +452,7 @@ let destCoFix c = match kind c with | _ -> raise DestKO let destRef c = let open GlobRef in match kind c with - | Var x -> VarRef x, Univ.Instance.empty + | Var x -> VarRef x, UVars.Instance.empty | Const (c,u) -> ConstRef c, u | Ind (ind,u) -> IndRef ind, u | Construct (c,u) -> ConstructRef c, u @@ -493,7 +493,7 @@ let fold f acc c = match kind c with | Lambda (_,t,c) -> f (f acc t) c | LetIn (_,b,t,c) -> f (f (f acc b) t) c | App (c,l) -> Array.fold_left f (f acc c) l - | Proj (_p,c) -> f acc c + | Proj (_p,_r,c) -> f acc c | Evar (_,l) -> SList.Skip.fold f acc l | Case (_,_,pms,(_,p),iv,c,bl) -> Array.fold_left (fun acc (_, b) -> f acc b) (f (fold_invert f (f (Array.fold_left f acc pms) p) iv) c) bl @@ -521,7 +521,7 @@ let iter f c = match kind c with | Lambda (_,t,c) -> f t; f c | LetIn (_,b,t,c) -> f b; f t; f c | App (c,l) -> f c; Array.iter f l - | Proj (_p,c) -> f c + | Proj (_p,_r,c) -> f c | Evar (_,l) -> SList.Skip.iter f l | Case (_,_,pms,p,iv,c,bl) -> Array.iter f pms; f (snd p); iter_invert f iv; f c; Array.iter (fun (_, b) -> f b) bl @@ -550,7 +550,7 @@ let iter_with_binders g f n c = match kind c with iter_invert (f n) iv; f n c; Array.Fun1.iter (fun n (ctx, b) -> f (iterate g (Array.length ctx) n) b) n bl - | Proj (_p,c) -> f n c + | Proj (_p,_r,c) -> f n c | Fix (_,(_,tl,bl)) -> Array.Fun1.iter f n tl; Array.Fun1.iter f (iterate g (Array.length tl) n) bl @@ -576,7 +576,7 @@ let fold_constr_with_binders g f n acc c = | Lambda (_na,t,c) -> f (g n) (f n acc t) c | LetIn (_na,b,t,c) -> f (g n) (f n (f n acc b) t) c | App (c,l) -> Array.fold_left (f n) (f n acc c) l - | Proj (_p,c) -> f n acc c + | Proj (_p,_r,c) -> f n acc c | Evar (_,l) -> SList.Skip.fold (f n) acc l | Case (_,_,pms,p,iv,c,bl) -> let fold_ctx n accu (nas, c) = @@ -659,10 +659,10 @@ let map f c = match kind c with let l' = Array.Smart.map f l in if b'==b && l'==l then c else mkApp (b', l') - | Proj (p,t) -> + | Proj (p,r,t) -> let t' = f t in if t' == t then c - else mkProj (p, t') + else mkProj (p, r, t') | Evar (e,l) -> let l' = SList.Smart.map f l in if l'==l then c @@ -742,10 +742,10 @@ let fold_map f accu c = match kind c with let accu, l' = Array.Smart.fold_left_map f accu l in if b'==b && l'==l then accu, c else accu, mkApp (b', l') - | Proj (p,t) -> + | Proj (p,r,t) -> let accu, t' = f accu t in if t' == t then accu, c - else accu, mkProj (p, t') + else accu, mkProj (p, r, t') | Evar (e,l) -> let accu, l' = SList.Smart.fold_left_map f accu l in if l'==l then accu, c @@ -810,10 +810,10 @@ let map_with_binders g f l c0 = match kind c0 with let al' = Array.Fun1.Smart.map f l al in if c' == c && al' == al then c0 else mkApp (c', al') - | Proj (p, t) -> + | Proj (p, r, t) -> let t' = f l t in if t' == t then c0 - else mkProj (p, t') + else mkProj (p, r, t') | Evar (e, al) -> let al' = SList.Smart.map (fun c -> f l c) al in if al' == al then c0 @@ -909,7 +909,8 @@ let compare_head_gen_leq_with kind1 kind2 leq_universes leq_sorts eq_evars eq le let len = Array.length l1 in Int.equal len (Array.length l2) && leq (nargs+len) c1 c2 && Array.equal_norefl (eq 0) l1 l2 - | Proj (p1,c1), Proj (p2,c2) -> Projection.CanOrd.equal p1 p2 && eq 0 c1 c2 + | Proj (p1,r1,c1), Proj (p2,r2,c2) -> + Projection.CanOrd.equal p1 p2 && Sorts.relevance_equal r1 r2 && eq 0 c1 c2 | Evar (e1,l1), Evar (e2,l2) -> eq_evars (e1, l1) (e2, l2) | Const (c1,u1), Const (c2,u2) -> (* The args length currently isn't used but may as well pass it. *) @@ -960,7 +961,7 @@ let compare_head_gen_with kind1 kind2 eq_universes eq_sorts eq_evars eq t1 t2 = let compare_head_gen eq_universes eq_sorts eq_evars eq t1 t2 = compare_head_gen_leq eq_universes eq_sorts eq_evars eq eq t1 t2 -let compare_head = compare_head_gen (fun _ -> Univ.Instance.equal) Sorts.equal +let compare_head = compare_head_gen (fun _ -> UVars.Instance.equal) Sorts.equal (*******************************) (* alpha conversion functions *) @@ -1064,7 +1065,7 @@ let constr_ord_int f t1 t2 = | CoFix(ln1,(_,tl1,bl1)), CoFix(ln2,(_,tl2,bl2)) -> compare [(Int.compare, ln1, ln2); (Array.compare f, tl1, tl2); (Array.compare f, bl1, bl2)] | CoFix _, _ -> -1 | _, CoFix _ -> 1 - | Proj (p1,c1), Proj (p2,c2) -> compare [(Projection.CanOrd.compare, p1, p2); (f, c1, c2)] + | Proj (p1,_r1,c1), Proj (p2,_r2,c2) -> compare [(Projection.CanOrd.compare, p1, p2); (f, c1, c2)] | Proj _, _ -> -1 | _, Proj _ -> 1 | Int i1, Int i2 -> Uint63.compare i1 i2 | Int _, _ -> -1 | _, Int _ -> 1 @@ -1147,7 +1148,7 @@ let hasheq t1 t2 = | LetIn (n1,b1,t1,c1), LetIn (n2,b2,t2,c2) -> n1 == n2 && b1 == b2 && t1 == t2 && c1 == c2 | App (c1,l1), App (c2,l2) -> c1 == c2 && array_eqeq l1 l2 - | Proj (p1,c1), Proj(p2,c2) -> p1 == p2 && c1 == c2 + | Proj (p1,r1,c1), Proj(p2,r2,c2) -> p1 == p2 && r1 == r2 && c1 == c2 | Evar (e1,l1), Evar (e2,l2) -> e1 == e2 && SList.equal (==) l1 l2 | Const (c1,u1), Const (c2,u2) -> c1 == c2 && u1 == u2 | Ind (ind1,u1), Ind (ind2,u2) -> ind1 == ind2 && u1 == u2 @@ -1234,8 +1235,8 @@ let rec hash t = combinesmall 14 (combine (hash_term_array bl) (hash_term_array tl)) | Meta n -> combinesmall 15 n | Rel n -> combinesmall 16 n - | Proj (p,c) -> - combinesmall 17 (combine (Projection.CanOrd.hash p) (hash c)) + | Proj (p,r, c) -> + combinesmall 17 (combine3 (Projection.CanOrd.hash p) (Sorts.relevance_hash r) (hash c)) | Int i -> combinesmall 18 (Uint63.hash i) | Float f -> combinesmall 19 (Float64.hash f) | Array(u,t,def,ty) -> @@ -1396,10 +1397,10 @@ let rec hash_term (t : t) = (t, combinesmall 15 n) | Rel n as t -> (t, combinesmall 16 n) - | Proj (p,c) -> + | Proj (p,r,c) -> let c, hc = sh_rec c in let p' = Projection.hcons p in - (Proj (p', c), combinesmall 17 (combine (Projection.SyntacticOrd.hash p') hc)) + (Proj (p', r, c), combinesmall 17 (combine (Projection.SyntacticOrd.hash p') hc)) | Int i as t -> let (h,l) = Uint63.to_int2 i in (t, combinesmall 18 (combine h l)) @@ -1491,8 +1492,8 @@ let debug_print_fix pr_constr ((t,i),(lna,tl,bl)) = str"}") let pr_puniverses p u = - if Univ.Instance.is_empty u then p - else Pp.(p ++ str"(*" ++ Univ.Instance.pr Univ.Level.raw_pr u ++ str"*)") + if UVars.Instance.is_empty u then p + else Pp.(p ++ str"(*" ++ UVars.Instance.pr Sorts.QVar.raw_pr Univ.Level.raw_pr u ++ str"*)") let rec debug_print c = let open Pp in @@ -1529,7 +1530,9 @@ let rec debug_print c = | Ind ((sp,i),u) -> str"Ind(" ++ pr_puniverses (MutInd.print sp ++ str"," ++ int i) u ++ str")" | Construct (((sp,i),j),u) -> str"Constr(" ++ pr_puniverses (MutInd.print sp ++ str"," ++ int i ++ str"," ++ int j) u ++ str")" - | Proj (p,c) -> str"Proj(" ++ Constant.debug_print (Projection.constant p) ++ str"," ++ bool (Projection.unfolded p) ++ str"," ++ debug_print c ++ str")" + | Proj (p,_r,c) -> + str"Proj(" ++ Constant.debug_print (Projection.constant p) ++ str"," ++ + bool (Projection.unfolded p) ++ str"," ++ debug_print c ++ str")" | Case (_ci,_u,pms,p,iv,c,bl) -> let pr_ctx (nas, c) = hov 2 (hov 0 (prvect (fun na -> Name.print na.binder_name ++ spc ()) nas ++ str "|-") ++ spc () ++ @@ -1553,7 +1556,7 @@ let rec debug_print c = | Float i -> str"Float("++str (Float64.to_string i) ++ str")" | Array(u,t,def,ty) -> str"Array(" ++ prlist_with_sep pr_comma debug_print (Array.to_list t) ++ str" | " ++ debug_print def ++ str " : " ++ debug_print ty - ++ str")@{" ++ Univ.Instance.pr Univ.Level.raw_pr u ++ str"}" + ++ str")@{" ++ UVars.Instance.pr Sorts.QVar.raw_pr Univ.Level.raw_pr u ++ str"}" and debug_invert = let open Pp in function | NoInvert -> mt() diff --git a/kernel/constr.mli b/kernel/constr.mli index 2c811a257a9a..514d8cf3d09c 100644 --- a/kernel/constr.mli +++ b/kernel/constr.mli @@ -14,9 +14,9 @@ open Names (** {6 Simply type aliases } *) -type pconstant = Constant.t Univ.puniverses -type pinductive = inductive Univ.puniverses -type pconstructor = constructor Univ.puniverses +type pconstant = Constant.t UVars.puniverses +type pinductive = inductive UVars.puniverses +type pconstructor = constructor UVars.puniverses (** {6 Existential variables } *) type metavariable = int @@ -85,7 +85,7 @@ val mkVar : Id.t -> constr val mkInt : Uint63.t -> constr (** Constructs an array *) -val mkArray : Univ.Instance.t * constr array * constr * types -> constr +val mkArray : UVars.Instance.t * constr array * constr * types -> constr (** Constructs a machine float number *) val mkFloat : Float64.t -> constr @@ -126,13 +126,13 @@ val mkLetIn : Name.t Context.binder_annot * constr * types * constr -> constr {%latex:$(f~t_1\dots f_n)$%}. *) val mkApp : constr * constr array -> constr -val map_puniverses : ('a -> 'b) -> 'a Univ.puniverses -> 'b Univ.puniverses +val map_puniverses : ('a -> 'b) -> 'a UVars.puniverses -> 'b UVars.puniverses (** Constructs a Constant.t *) val mkConstU : pconstant -> constr (** Constructs a projection application *) -val mkProj : (Projection.t * constr) -> constr +val mkProj : (Projection.t * Sorts.relevance * constr) -> constr (** Inductive types *) @@ -145,7 +145,7 @@ val mkConstructU : pconstructor -> constr val mkConstructUi : pinductive * int -> constr (** Make a constant, inductive, constructor or variable. *) -val mkRef : GlobRef.t Univ.puniverses -> constr +val mkRef : GlobRef.t UVars.puniverses -> constr module UnsafeMonomorphic : sig val mkConst : Constant.t -> constr @@ -176,7 +176,7 @@ type ('constr, 'types, 'univs) pcase = type case_invert = constr pcase_invert type case_return = types pcase_return type case_branch = constr pcase_branch -type case = (constr, types, Univ.Instance.t) pcase +type case = (constr, types, UVars.Instance.t) pcase val mkCase : case -> constr @@ -281,7 +281,8 @@ type ('constr, 'types, 'sort, 'univs) kind_of_term = The names in the [brs] are the names of the variables bound in the respective branch. *) | Fix of ('constr, 'types) pfixpoint | CoFix of ('constr, 'types) pcofixpoint - | Proj of Projection.t * 'constr + | Proj of Projection.t * Sorts.relevance * 'constr + (** The relevance is the relevance of the whole term *) | Int of Uint63.t | Float of Float64.t | Array of 'univs * 'constr array * 'constr * 'types @@ -292,13 +293,13 @@ type ('constr, 'types, 'sort, 'univs) kind_of_term = least one argument and the function is not itself an applicative term *) -val kind : constr -> (constr, types, Sorts.t, Univ.Instance.t) kind_of_term -val of_kind : (constr, types, Sorts.t, Univ.Instance.t) kind_of_term -> constr +val kind : constr -> (constr, types, Sorts.t, UVars.Instance.t) kind_of_term +val of_kind : (constr, types, Sorts.t, UVars.Instance.t) kind_of_term -> constr val kind_nocast_gen : ('v -> ('v, 'v, 'sort, 'univs) kind_of_term) -> ('v -> ('v, 'v, 'sort, 'univs) kind_of_term) -val kind_nocast : constr -> (constr, types, Sorts.t, Univ.Instance.t) kind_of_term +val kind_nocast : constr -> (constr, types, Sorts.t, UVars.Instance.t) kind_of_term (** {6 Simple case analysis} *) val isRel : constr -> bool @@ -372,16 +373,16 @@ val decompose_app_list : constr -> constr * constr list val decompose_app : constr -> constr * constr array (** Destructs a constant *) -val destConst : constr -> Constant.t Univ.puniverses +val destConst : constr -> Constant.t UVars.puniverses (** Destructs an existential variable *) val destEvar : constr -> existential (** Destructs a (co)inductive type *) -val destInd : constr -> inductive Univ.puniverses +val destInd : constr -> inductive UVars.puniverses (** Destructs a constructor *) -val destConstruct : constr -> constructor Univ.puniverses +val destConstruct : constr -> constructor UVars.puniverses (** Destructs a [match c as x in I args return P with ... | Ci(...yij...) => ti | ... end] (or [let (..y1i..) := c as x in I args @@ -391,7 +392,7 @@ where [info] is pretty-printing information *) val destCase : constr -> case (** Destructs a projection *) -val destProj : constr -> Projection.t * constr +val destProj : constr -> Projection.t * Sorts.relevance * constr (** Destructs the {% $ %}i{% $ %}th function of the block [Fixpoint f{_ 1} ctx{_ 1} = b{_ 1} @@ -404,7 +405,7 @@ val destFix : constr -> fixpoint val destCoFix : constr -> cofixpoint -val destRef : constr -> GlobRef.t Univ.puniverses +val destRef : constr -> GlobRef.t UVars.puniverses (** {6 Equality} *) @@ -571,7 +572,7 @@ type 'univs instance_compare_fn = (GlobRef.t * int) option -> compare universe instances, [s] to compare sorts; Cast's, binders name and Cases annotations are not taken into account *) -val compare_head_gen : Univ.Instance.t instance_compare_fn -> +val compare_head_gen : UVars.Instance.t instance_compare_fn -> (Sorts.t -> Sorts.t -> bool) -> (existential -> existential -> bool) -> constr constr_compare_fn -> @@ -607,7 +608,7 @@ val compare_head_gen_with : [s] to compare sorts for for subtyping; Cast's, binders name and Cases annotations are not taken into account *) -val compare_head_gen_leq : Univ.Instance.t instance_compare_fn -> +val compare_head_gen_leq : UVars.Instance.t instance_compare_fn -> (Sorts.t -> Sorts.t -> bool) -> (existential -> existential -> bool) -> constr constr_compare_fn -> diff --git a/kernel/context.ml b/kernel/context.ml index 4e922b8e3279..720ec9defb83 100644 --- a/kernel/context.ml +++ b/kernel/context.ml @@ -92,6 +92,10 @@ struct | LocalAssum (x,ty) -> LocalAssum ({x with binder_name=na}, ty) | LocalDef (x,v,ty) -> LocalDef ({x with binder_name=na}, v, ty) + let set_relevance r = function + | LocalAssum (x,ty) -> LocalAssum ({x with binder_relevance=r}, ty) + | LocalDef (x,v,ty) -> LocalDef ({x with binder_relevance=r}, v, ty) + (** Set the type of the bound variable in a given declaration. *) let set_type ty = function | LocalAssum (na,_) -> LocalAssum (na, ty) @@ -133,6 +137,11 @@ struct let na' = f na in if na == na' then x else set_name na' x + let map_relevance f x = + let r = get_relevance x in + let r' = f r in + if r == r' then x else set_relevance r' x + (** For local assumptions, this function returns the original local assumptions. For local definitions, this function maps the value in the local definition. *) let map_value f = function diff --git a/kernel/context.mli b/kernel/context.mli index f31729d86d5c..68e346a75181 100644 --- a/kernel/context.mli +++ b/kernel/context.mli @@ -95,6 +95,9 @@ sig (** Map the name bound by a given declaration. *) val map_name : (Name.t -> Name.t) -> ('c, 't) pt -> ('c, 't) pt + (** Map the relevance *) + val map_relevance : (Sorts.relevance -> Sorts.relevance) -> ('c, 't) pt -> ('c, 't) pt + (** For local assumptions, this function returns the original local assumptions. For local definitions, this function maps the value in the local definition. *) val map_value : ('c -> 'c) -> ('c, 't) pt -> ('c, 't) pt diff --git a/kernel/conversion.ml b/kernel/conversion.ml index 54690240774a..4441913b390c 100644 --- a/kernel/conversion.ml +++ b/kernel/conversion.ml @@ -52,7 +52,7 @@ let compare_stack_shape stk1 stk2 = | (_, (Zupdate _|Zshift _)::s2) -> compare_rec bal stk1 s2 | (Zapp l1::s1, _) -> compare_rec (bal+Array.length l1) s1 stk2 | (_, Zapp l2::s2) -> compare_rec (bal-Array.length l2) stk1 s2 - | (Zproj _p1::s1, Zproj _p2::s2) -> + | (Zproj _::s1, Zproj _::s2) -> Int.equal bal 0 && compare_rec 0 s1 s2 | (ZcaseT(_c1,_,_,_,_,_)::s1, ZcaseT(_c2,_,_,_,_,_)::s2) -> Int.equal bal 0 (* && c1.ci_ind = c2.ci_ind *) && compare_rec 0 s1 s2 @@ -72,7 +72,7 @@ type lft_constr_stack_elt = Zlapp of (lift * fconstr) array | Zlproj of Projection.Repr.t * lift | Zlfix of (lift * fconstr) * lft_constr_stack - | Zlcase of case_info * lift * Univ.Instance.t * constr array * case_return * case_branch array * usubs + | Zlcase of case_info * lift * UVars.Instance.t * constr array * case_return * case_branch array * usubs | Zlprimitive of CPrimitives.t * pconstant * lft_fconstr list * lft_fconstr next_native_args and lft_constr_stack = lft_constr_stack_elt list @@ -102,7 +102,7 @@ let pure_stack lfts stk = | (Zshift n,(l,pstk)) -> (el_shft n l, pstk) | (Zapp a, (l,pstk)) -> (l,zlapp (map_lift l a) pstk) - | (Zproj p, (l,pstk)) -> + | (Zproj (p,_), (l,pstk)) -> (l, Zlproj (p,l)::pstk) | (Zfix(fx,a),(l,pstk)) -> let (lfx,pa) = pure_rec l a in @@ -148,9 +148,9 @@ type conv_pb = type 'a universe_compare = { (* Might raise NotConvertible *) compare_sorts : env -> conv_pb -> Sorts.t -> Sorts.t -> 'a -> 'a; - compare_instances: flex:bool -> Univ.Instance.t -> Univ.Instance.t -> 'a -> 'a; - compare_cumul_instances : conv_pb -> Univ.Variance.t array -> - Univ.Instance.t -> Univ.Instance.t -> 'a -> 'a; + compare_instances: flex:bool -> UVars.Instance.t -> UVars.Instance.t -> 'a -> 'a; + compare_cumul_instances : conv_pb -> UVars.Variance.t array -> + UVars.Instance.t -> UVars.Instance.t -> 'a -> 'a; } type 'a universe_state = 'a * 'a universe_compare @@ -175,9 +175,9 @@ let convert_instances_cumul pb var u u' (s, check) = let get_cumulativity_constraints cv_pb variance u u' = match cv_pb with | CONV -> - Univ.enforce_eq_variance_instances variance u u' Univ.Constraints.empty + UVars.enforce_eq_variance_instances variance u u' Sorts.QUConstraints.empty | CUMUL -> - Univ.enforce_leq_variance_instances variance u u' Univ.Constraints.empty + UVars.enforce_leq_variance_instances variance u u' Sorts.QUConstraints.empty let inductive_cumulativity_arguments (mind,ind) = mind.Declarations.mind_nparams + @@ -190,7 +190,7 @@ let convert_inductives_gen cmp_instances cmp_cumul cv_pb (mind,ind) nargs u1 u2 let num_param_arity = inductive_cumulativity_arguments (mind,ind) in if not (Int.equal num_param_arity nargs) then (* shortcut, not sure if worth doing, could use perf data *) - if Univ.Instance.equal u1 u2 then s else raise MustExpand + if UVars.Instance.equal u1 u2 then s else raise MustExpand else cmp_cumul cv_pb variances u1 u2 s @@ -208,11 +208,12 @@ let convert_constructors_gen cmp_instances cmp_cumul (mind, ind, cns) nargs u1 u | Some _ -> let num_cnstr_args = constructor_cumulativity_arguments (mind,ind,cns) in if not (Int.equal num_cnstr_args nargs) then - if Univ.Instance.equal u1 u2 then s else raise MustExpand + if UVars.Instance.equal u1 u2 then s else raise MustExpand else (** By invariant, both constructors have a common supertype, so they are convertible _at that type_. *) - let variance = Array.make (Univ.Instance.length u1) Univ.Variance.Irrelevant in + (* NB: no variance for qualities *) + let variance = Array.make (snd (UVars.Instance.length u1)) UVars.Variance.Irrelevant in cmp_cumul CONV variance u1 u2 s let convert_constructors ctor nargs u1 u2 (s, check) = @@ -223,7 +224,7 @@ let conv_table_key infos ~nargs k1 k2 cuniv = if k1 == k2 then cuniv else match k1, k2 with | ConstKey (cst, u), ConstKey (cst', u') when Constant.CanOrd.equal cst cst' -> - if Univ.Instance.equal u u' then cuniv + if UVars.Instance.equal u u' then cuniv else if Int.equal nargs 1 && is_array_type (info_env infos) cst then cuniv else let flex = evaluable_constant cst (info_env infos) @@ -289,14 +290,17 @@ let esubst_of_rel_context_instance_list ctx u args e = in aux 0 e args (List.rev ctx) -let irr_flex env = function - | ConstKey (con,_) -> Cset_env.mem con env.Environ.irr_constants - | VarKey x -> Context.Named.Declaration.get_relevance (Environ.lookup_named x env) == Sorts.Irrelevant - | RelKey x -> Context.Rel.Declaration.get_relevance (Environ.lookup_rel x env) == Sorts.Irrelevant +let irr_flex infos = function + | ConstKey (con,u) -> begin match Cmap_env.find_opt con (info_env infos).Environ.irr_constants with + | None -> false + | Some r -> is_irrelevant infos @@ UVars.subst_instance_relevance u r + end + | VarKey x -> is_irrelevant infos @@ Context.Named.Declaration.get_relevance (Environ.lookup_named x (info_env infos)) + | RelKey x -> is_irrelevant infos @@ Context.Rel.Declaration.get_relevance (Environ.lookup_rel x (info_env infos)) let eq_universes (_,e1) (_,e2) u1 u2 = - let subst e u = if Univ.Instance.is_empty e then u else Univ.subst_instance_instance e u in - Univ.Instance.equal (subst e1 u1) (subst e2 u2) + let subst e u = if UVars.Instance.is_empty e then u else UVars.subst_instance_instance e u in + UVars.Instance.equal (subst e1 u1) (subst e2 u2) let rec compare_under e1 c1 e2 c2 = match Constr.kind c1, Constr.kind c2 with @@ -315,7 +319,7 @@ let rec compare_under e1 c1 e2 c2 = | Float f1, Float f2 -> Float64.equal f1 f2 | Sort s1, Sort s2 -> let subst_instance_sort u s = - if Univ.Instance.is_empty u then s else Sorts.subst_instance_sort u s + if UVars.Instance.is_empty u then s else UVars.subst_instance_sort u s in let s1 = subst_instance_sort (snd e1) s1 and s2 = subst_instance_sort (snd e2) s2 in @@ -335,7 +339,7 @@ let rec compare_under e1 c1 e2 c2 = Int.equal len (Array.length l2) && compare_under e1 c1 e2 c2 && Array.equal_norefl (fun c1 c2 -> compare_under e1 c1 e2 c2) l1 l2 - | Proj (p1,c1), Proj (p2,c2) -> + | Proj (p1,_,c1), Proj (p2,_,c2) -> Projection.CanOrd.equal p1 p2 && compare_under e1 c1 e2 c2 | Evar _, Evar _ -> false | Const (c1,u1), Const (c2,u2) -> @@ -428,7 +432,7 @@ and eqappr cv_pb l2r infos (lft1,st1) (lft2,st2) cuniv = (try let nargs = same_args_size v1 v2 in let cuniv = conv_table_key infos.cnv_inf ~nargs fl1 fl2 cuniv in - let () = if irr_flex (info_env infos.cnv_inf) fl1 then raise NotConvertible (* trigger the fallback *) in + let () = if irr_flex infos.cnv_inf fl1 then raise NotConvertible (* trigger the fallback *) in convert_stacks l2r infos lft1 lft2 v1 v2 cuniv with NotConvertible | UGraph.UniverseInconsistency _ -> let r1 = unfold_ref_with_args infos.cnv_inf infos.lft_tab fl1 v1 in @@ -438,7 +442,7 @@ and eqappr cv_pb l2r infos (lft1,st1) (lft2,st2) cuniv = | Some t1, Some t2 -> (* else the oracle tells which constant is to be expanded *) let oracle = CClosure.oracle_of_infos infos.cnv_inf in - if Conv_oracle.oracle_order Univ.out_punivs oracle l2r fl1 fl2 then + if Conv_oracle.oracle_order UVars.out_punivs oracle l2r fl1 fl2 then eqappr cv_pb l2r infos (lft1, t1) appr2 cuniv else eqappr cv_pb l2r infos appr1 (lft2, t2) cuniv @@ -452,15 +456,15 @@ and eqappr cv_pb l2r infos (lft1,st1) (lft2,st2) cuniv = eqappr cv_pb l2r infos appr1 (lft2, t2) cuniv ) - | (FProj (p1,c1), FProj (p2, c2)) -> + | (FProj (p1,r1,c1), FProj (p2, r2, c2)) -> (* Projections: prefer unfolding to first-order unification, which will happen naturally if the terms c1, c2 are not in constructor form *) - (match unfold_projection infos.cnv_inf p1 with + (match unfold_projection infos.cnv_inf p1 r1 with | Some s1 -> eqappr cv_pb l2r infos (lft1, (c1, (s1 :: v1))) appr2 cuniv | None -> - match unfold_projection infos.cnv_inf p2 with + match unfold_projection infos.cnv_inf p2 r2 with | Some s2 -> eqappr cv_pb l2r infos appr1 (lft2, (c2, (s2 :: v2))) cuniv | None -> @@ -473,8 +477,8 @@ and eqappr cv_pb l2r infos (lft1,st1) (lft2,st2) cuniv = else (* Two projections in WHNF: unfold *) raise NotConvertible) - | (FProj (p1,c1), t2) -> - begin match unfold_projection infos.cnv_inf p1 with + | (FProj (p1,r1,c1), t2) -> + begin match unfold_projection infos.cnv_inf p1 r1 with | Some s1 -> eqappr cv_pb l2r infos (lft1, (c1, (s1 :: v1))) appr2 cuniv | None -> @@ -489,8 +493,8 @@ and eqappr cv_pb l2r infos (lft1,st1) (lft2,st2) cuniv = end end - | (t1, FProj (p2,c2)) -> - begin match unfold_projection infos.cnv_inf p2 with + | (t1, FProj (p2,r2,c2)) -> + begin match unfold_projection infos.cnv_inf p2 r2 with | Some s2 -> eqappr cv_pb l2r infos appr1 (lft2, (c2, (s2 :: v2))) cuniv | None -> @@ -526,6 +530,7 @@ and eqappr cv_pb l2r infos (lft1,st1) (lft2,st2) cuniv = let el1 = el_stack lft1 v1 in let el2 = el_stack lft2 v2 in let cuniv = ccnv CONV l2r infos el1 el2 c1 c'1 cuniv in + let x1 = usubst_binder e x1 in ccnv cv_pb l2r (push_relevance infos x1) (el_lift el1) (el_lift el2) (mk_clos (usubs_lift e) c2) (mk_clos (usubs_lift e') c'2) cuniv (* Eta-expansion on the fly *) @@ -563,10 +568,10 @@ and eqappr cv_pb l2r infos (lft1,st1) (lft2,st2) cuniv = eqappr cv_pb l2r infos (lft1, r1) appr2 cuniv | None -> (match c2 with - | FConstruct ((ind2,_j2),_u2) -> + | FConstruct ((ind2,_j2),u2) -> (try let v2, v1 = - eta_expand_ind_stack (info_env infos.cnv_inf) ind2 hd2 v2 (snd appr1) + eta_expand_ind_stack (info_env infos.cnv_inf) (ind2,u2) hd2 v2 (snd appr1) in convert_stacks l2r infos lft1 lft2 v1 v2 cuniv with Not_found -> raise NotConvertible) | _ -> raise NotConvertible) @@ -581,9 +586,9 @@ and eqappr cv_pb l2r infos (lft1,st1) (lft2,st2) cuniv = eqappr cv_pb l2r infos appr1 (lft2, r2) cuniv | None -> match c1 with - | FConstruct ((ind1,_j1),_u1) -> + | FConstruct ((ind1,_j1),u1) -> (try let v1, v2 = - eta_expand_ind_stack (info_env infos.cnv_inf) ind1 hd1 v1 (snd appr2) + eta_expand_ind_stack (info_env infos.cnv_inf) (ind1,u1) hd1 v1 (snd appr2) in convert_stacks l2r infos lft1 lft2 v1 v2 cuniv with Not_found -> raise NotConvertible) | _ -> raise NotConvertible @@ -592,7 +597,7 @@ and eqappr cv_pb l2r infos (lft1,st1) (lft2,st2) cuniv = (* Inductive types: MutInd MutConstruct Fix Cofix *) | (FInd (ind1,u1 as pind1), FInd (ind2,u2 as pind2)) -> if Ind.CanOrd.equal ind1 ind2 then - if Univ.Instance.length u1 = 0 || Univ.Instance.length u2 = 0 then + if UVars.Instance.is_empty u1 || UVars.Instance.is_empty u2 then let cuniv = convert_instances ~flex:false u1 u2 cuniv in convert_stacks l2r infos lft1 lft2 v1 v2 cuniv else @@ -609,7 +614,7 @@ and eqappr cv_pb l2r infos (lft1,st1) (lft2,st2) cuniv = | (FConstruct ((ind1,j1),u1 as pctor1), FConstruct ((ind2,j2),u2 as pctor2)) -> if Int.equal j1 j2 && Ind.CanOrd.equal ind1 ind2 then - if Univ.Instance.length u1 = 0 || Univ.Instance.length u2 = 0 then + if UVars.Instance.is_empty u1 || UVars.Instance.is_empty u2 then let cuniv = convert_instances ~flex:false u1 u2 cuniv in convert_stacks l2r infos lft1 lft2 v1 v2 cuniv else @@ -625,17 +630,17 @@ and eqappr cv_pb l2r infos (lft1,st1) (lft2,st2) cuniv = else raise NotConvertible (* Eta expansion of records *) - | (FConstruct ((ind1,_j1),_u1), _) -> + | (FConstruct ((ind1,_j1),u1), _) -> (try let v1, v2 = - eta_expand_ind_stack (info_env infos.cnv_inf) ind1 hd1 v1 (snd appr2) + eta_expand_ind_stack (info_env infos.cnv_inf) (ind1,u1) hd1 v1 (snd appr2) in convert_stacks l2r infos lft1 lft2 v1 v2 cuniv with Not_found -> raise NotConvertible) - | (_, FConstruct ((ind2,_j2),_u2)) -> + | (_, FConstruct ((ind2,_j2),u2)) -> (try let v2, v1 = - eta_expand_ind_stack (info_env infos.cnv_inf) ind2 hd2 v2 (snd appr1) + eta_expand_ind_stack (info_env infos.cnv_inf) (ind2,u2) hd2 v2 (snd appr1) in convert_stacks l2r infos lft1 lft2 v1 v2 cuniv with Not_found -> raise NotConvertible) @@ -651,6 +656,7 @@ and eqappr cv_pb l2r infos (lft1,st1) (lft2,st2) cuniv = let el2 = el_stack lft2 v2 in let cuniv = convert_vect l2r infos el1 el2 fty1 fty2 cuniv in let cuniv = + let na1 = Array.map (usubst_binder e1) na1 in let infos = push_relevances infos na1 in convert_vect l2r infos (el_liftn n el1) (el_liftn n el2) fcl1 fcl2 cuniv @@ -670,6 +676,7 @@ and eqappr cv_pb l2r infos (lft1,st1) (lft2,st2) cuniv = let el2 = el_stack lft2 v2 in let cuniv = convert_vect l2r infos el1 el2 fty1 fty2 cuniv in let cuniv = + let na1 = Array.map (usubst_binder e1) na1 in let infos = push_relevances infos na1 in convert_vect l2r infos (el_liftn n el1) (el_liftn n el2) fcl1 fcl2 cuniv @@ -708,7 +715,7 @@ and eqappr cv_pb l2r infos (lft1,st1) (lft2,st2) cuniv = | FArray (u1,t1,ty1), FArray (u2,t2,ty2) -> let len = Parray.length_int t1 in if not (Int.equal len (Parray.length_int t2)) then raise NotConvertible; - let cuniv = convert_instances_cumul CONV [|Univ.Variance.Irrelevant|] u1 u2 cuniv in + let cuniv = convert_instances_cumul CONV [|UVars.Variance.Irrelevant|] u1 u2 cuniv in let el1 = el_stack lft1 v1 in let el2 = el_stack lft2 v2 in let cuniv = ccnv CONV l2r infos el1 el2 ty1 ty2 cuniv in @@ -766,7 +773,7 @@ and convert_stacks l2r infos lft1 lft2 stk1 stk2 cuniv = let mind = Environ.lookup_mind (fst ci1.ci_ind) (info_env infos.cnv_inf) in let mip = mind.Declarations.mind_packets.(snd ci1.ci_ind) in let cu = - if Univ.Instance.length u1 = 0 || Univ.Instance.length u2 = 0 then + if UVars.Instance.is_empty u1 || UVars.Instance.is_empty u2 then convert_instances ~flex:false u1 u2 cu else let u1 = CClosure.usubst_instance e1 u1 in @@ -821,7 +828,7 @@ and convert_under_context l2r infos e1 e2 lft1 lft2 ctx (nas1, c1) (nas2, c2) cu in let lft1 = el_liftn n lft1 in let lft2 = el_liftn n lft2 in - let infos = push_relevances infos nas1 in + let infos = push_relevances infos (Array.map (usubst_binder e1) nas1) in ccnv CONV l2r infos lft1 lft2 (mk_clos e1 c1) (mk_clos e2 c2) cu and convert_return_clause mib mip l2r infos e1 e2 l1 l2 u1 u2 pms1 pms2 p1 p2 cu = @@ -894,8 +901,8 @@ let check_convert_instances ~flex:_ u u' univs = (* general conversion and inference functions *) let check_inductive_instances cv_pb variance u1 u2 univs = - let csts = get_cumulativity_constraints cv_pb variance u1 u2 in - if (UGraph.check_constraints csts univs) then univs + let qcsts, ucsts = get_cumulativity_constraints cv_pb variance u1 u2 in + if Sorts.QConstraints.trivial qcsts && (UGraph.check_constraints ucsts univs) then univs else raise NotConvertible let checked_universes = diff --git a/kernel/conversion.mli b/kernel/conversion.mli index 7b575e662dcd..95b28ea61439 100644 --- a/kernel/conversion.mli +++ b/kernel/conversion.mli @@ -27,9 +27,9 @@ type conv_pb = CONV | CUMUL type 'a universe_compare = { (* Might raise NotConvertible *) compare_sorts : env -> conv_pb -> Sorts.t -> Sorts.t -> 'a -> 'a; - compare_instances: flex:bool -> Univ.Instance.t -> Univ.Instance.t -> 'a -> 'a; - compare_cumul_instances : conv_pb -> Univ.Variance.t array -> - Univ.Instance.t -> Univ.Instance.t -> 'a -> 'a; + compare_instances: flex:bool -> UVars.Instance.t -> UVars.Instance.t -> 'a -> 'a; + compare_cumul_instances : conv_pb -> UVars.Variance.t array -> + UVars.Instance.t -> UVars.Instance.t -> 'a -> 'a; } type 'a universe_state = 'a * 'a universe_compare @@ -38,8 +38,8 @@ type ('a,'b) generic_conversion_function = env -> 'b universe_state -> 'a -> 'a type 'a infer_conversion_function = env -> 'a -> 'a -> Univ.Constraints.t -val get_cumulativity_constraints : conv_pb -> Univ.Variance.t array -> - Univ.Instance.t -> Univ.Instance.t -> Univ.Constraints.t +val get_cumulativity_constraints : conv_pb -> UVars.Variance.t array -> + UVars.Instance.t -> UVars.Instance.t -> Sorts.QUConstraints.t val inductive_cumulativity_arguments : (Declarations.mutual_inductive_body * int) -> int val constructor_cumulativity_arguments : (Declarations.mutual_inductive_body * int * int) -> int @@ -49,7 +49,7 @@ val sort_cmp_universes : env -> conv_pb -> Sorts.t -> Sorts.t -> (* [flex] should be true for constants, false for inductive types and constructors. *) -val convert_instances : flex:bool -> Univ.Instance.t -> Univ.Instance.t -> +val convert_instances : flex:bool -> UVars.Instance.t -> UVars.Instance.t -> 'a * 'a universe_compare -> 'a * 'a universe_compare (** This function never raise UnivInconsistency. *) diff --git a/kernel/cooking.ml b/kernel/cooking.ml index b4603356ab42..1031a2dfde9e 100644 --- a/kernel/cooking.ml +++ b/kernel/cooking.ml @@ -19,7 +19,7 @@ open Util open Names open Term open Constr -open Univ +open UVars open Context module NamedDecl = Context.Named.Declaration @@ -38,7 +38,7 @@ module NamedDecl = Context.Named.Declaration type abstr_info = { abstr_ctx : Constr.named_context; (** Context over which to generalize (e.g. x:T,z:V(x)) *) - abstr_auctx : Univ.AbstractContext.t; + abstr_auctx : UVars.AbstractContext.t; (** Universe context over which to generalize *) abstr_ausubst : Instance.t; (** Universe substitution represented as an instance *) @@ -57,7 +57,7 @@ type abstr_info = { type abstr_inst_info = { abstr_rev_inst : Id.t list; (** The variables to reapply (excluding "let"s of the context), in reverse order *) - abstr_uinst : Univ.Instance.t; + abstr_uinst : UVars.Instance.t; (** Abstracted universe variables to reapply *) } @@ -117,7 +117,7 @@ struct end module RefTable = Hashtbl.Make(RefHash) -type internal_abstr_inst_info = Univ.Instance.t * int list * int +type internal_abstr_inst_info = UVars.Instance.t * int list * int type cooking_cache = { cache : internal_abstr_inst_info RefTable.t; @@ -214,7 +214,7 @@ let expand_constr cache modlist top_abst_subst c = with | Not_found -> Constr.map_with_binders succ substrec k c) - | Proj (p, c') -> + | Proj (p, r, c') -> let ind = Names.Projection.inductive p in let p' = try @@ -222,7 +222,7 @@ let expand_constr cache modlist top_abst_subst c = discharge_proj exp p with Not_found -> p in let c'' = substrec k c' in - if p == p' && c' == c'' then c else mkProj (p', c'') + if p == p' && c' == c'' then c else mkProj (p', r, c'') | Var id -> (try @@ -348,17 +348,16 @@ let discharge_abstract_universe_context abstr auctx = together with the substitution [u₀ ↦ Var(0), ... ,uₙ₋₁ ↦ Var(n - 1), Var(0) ↦ Var(n), ..., Var(m - 1) ↦ Var (n + m - 1)]. *) - let open Univ in let n = AbstractContext.size abstr.abstr_auctx in - if (Int.equal n 0) then + if (UVars.eq_sizes n (0,0)) then (** Optimization: still need to take the union for the constraints between globals *) abstr, n, AbstractContext.union abstr.abstr_auctx auctx else let subst = abstr.abstr_ausubst in - let suff = Instance.of_array @@ Array.init (AbstractContext.size auctx) (fun i -> Level.var i) in + let suff = UVars.make_abstract_instance auctx in let ainst = Instance.append subst suff in let substf = make_instance_subst ainst in - let auctx = Univ.subst_univs_level_abstract_universe_context substf auctx in + let auctx = UVars.subst_univs_level_abstract_universe_context (snd substf) auctx in let auctx' = AbstractContext.union abstr.abstr_auctx auctx in { abstr with abstr_ausubst = ainst }, n, auctx' @@ -387,5 +386,5 @@ let lift_private_mono_univs info a = a let lift_private_poly_univs info (inst, cstrs) = - let cstrs = Univ.subst_univs_level_constraints (make_instance_subst info.abstr_info.abstr_ausubst) cstrs in + let cstrs = Univ.subst_univs_level_constraints (snd (make_instance_subst info.abstr_info.abstr_ausubst)) cstrs in (inst, cstrs) diff --git a/kernel/cooking.mli b/kernel/cooking.mli index efc54106adcb..9c80d203150f 100644 --- a/kernel/cooking.mli +++ b/kernel/cooking.mli @@ -43,12 +43,12 @@ val empty_cooking_info : cooking_info (** Nothing to abstract *) val make_cooking_info : recursive:MutInd.t option -> expand_info -> - named_context -> Univ.UContext.t -> cooking_info * abstr_inst_info + named_context -> UVars.UContext.t -> cooking_info * abstr_inst_info (** Abstract a context assumed to be de-Bruijn free for terms and universes *) val names_info : cooking_info -> Id.Set.t -val universe_context_of_cooking_info : cooking_info -> Univ.AbstractContext.t +val universe_context_of_cooking_info : cooking_info -> UVars.AbstractContext.t val instance_of_cooking_info : cooking_info -> Constr.t array @@ -68,7 +68,7 @@ val lift_mono_univs : cooking_info -> Univ.ContextSet.t -> cooking_info * Univ.C (** The [int] is how many universes got discharged, ie size of returned context - size of input context. *) -val lift_poly_univs : cooking_info -> Univ.AbstractContext.t -> cooking_info * int * Univ.AbstractContext.t +val lift_poly_univs : cooking_info -> UVars.AbstractContext.t -> cooking_info * (int * int) * UVars.AbstractContext.t val lift_private_mono_univs : cooking_info -> 'a -> 'a diff --git a/kernel/declarations.mli b/kernel/declarations.mli index 9988dee2e5c2..496a988579cb 100644 --- a/kernel/declarations.mli +++ b/kernel/declarations.mli @@ -55,7 +55,7 @@ type ('a, 'opaque) constant_def = type universes = | Monomorphic - | Polymorphic of Univ.AbstractContext.t + | Polymorphic of UVars.AbstractContext.t (** The [typing_flags] are instructions to the type-checker which modify its behaviour. The typing flags used in the type-checking @@ -105,7 +105,7 @@ type typing_flags = { * the OpaqueDef *) type 'opaque pconstant_body = { const_hyps : Constr.named_context; (** younger hyp at top *) - const_univ_hyps : Univ.Instance.t; + const_univ_hyps : UVars.Instance.t; const_body : (Constr.t, 'opaque) constant_def; const_type : types; const_relevance : Sorts.relevance; @@ -251,7 +251,7 @@ type mutual_inductive_body = { mind_hyps : Constr.named_context; (** Section hypotheses on which the block depends *) - mind_univ_hyps : Univ.Instance.t; (** Section polymorphic universes. *) + mind_univ_hyps : UVars.Instance.t; (** Section polymorphic universes. *) mind_nparams : int; (** Number of expected parameters including non-uniform ones (i.e. length of mind_params_ctxt w/o let-in) *) @@ -263,9 +263,9 @@ type mutual_inductive_body = { mind_template : template_universes option; - mind_variance : Univ.Variance.t array option; (** Variance info, [None] when non-cumulative. *) + mind_variance : UVars.Variance.t array option; (** Variance info, [None] when non-cumulative. *) - mind_sec_variance : Univ.Variance.t array option; + mind_sec_variance : UVars.Variance.t array option; (** Variance info for section polymorphic universes. [None] outside sections. The final variance once all sections are discharged is [mind_sec_variance ++ mind_variance]. *) @@ -306,7 +306,7 @@ type 'uconstr functor_alg_expr = (** A module expression is an algebraic expression, possibly functorized. *) -type module_expression = (constr * Univ.AbstractContext.t option) functor_alg_expr +type module_expression = (constr * UVars.AbstractContext.t option) functor_alg_expr (** A component of a module structure *) diff --git a/kernel/declareops.ml b/kernel/declareops.ml index acb9dcc99f02..1bf77321fd91 100644 --- a/kernel/declareops.ml +++ b/kernel/declareops.ml @@ -56,15 +56,15 @@ let hcons_template_universe ar = template_context = Univ.hcons_universe_context_set ar.template_context } let universes_context = function - | Monomorphic -> Univ.AbstractContext.empty + | Monomorphic -> UVars.AbstractContext.empty | Polymorphic ctx -> ctx let abstract_universes = function | Entries.Monomorphic_entry -> - Univ.empty_level_subst, Monomorphic + UVars.empty_sort_subst, Monomorphic | Entries.Polymorphic_entry uctx -> - let (inst, auctx) = Univ.abstract_universes uctx in - let inst = Univ.make_instance_subst inst in + let (inst, auctx) = UVars.abstract_universes uctx in + let inst = UVars.make_instance_subst inst in (inst, Polymorphic auctx) (** {6 Constants } *) @@ -106,7 +106,7 @@ let subst_const_def subst def = match def with let subst_const_body subst cb = (* we're outside sections *) - assert (List.is_empty cb.const_hyps && Univ.Instance.is_empty cb.const_univ_hyps); + assert (List.is_empty cb.const_hyps && UVars.Instance.is_empty cb.const_univ_hyps); if is_empty_subst subst then cb else let body' = subst_const_def subst cb.const_body in @@ -115,7 +115,7 @@ let subst_const_body subst cb = then cb else { const_hyps = []; - const_univ_hyps = Univ.Instance.empty; + const_univ_hyps = UVars.Instance.empty; const_body = body'; const_type = type'; const_body_code = @@ -147,7 +147,7 @@ let hcons_universes cbu = match cbu with | Monomorphic -> Monomorphic | Polymorphic ctx -> - Polymorphic (Univ.hcons_abstract_universe_context ctx) + Polymorphic (UVars.hcons_abstract_universe_context ctx) let hcons_const_body cb = { cb with @@ -268,12 +268,12 @@ let subst_mind_record subst r = match r with let subst_mind_body subst mib = (* we're outside sections *) - assert (List.is_empty mib.mind_hyps && Univ.Instance.is_empty mib.mind_univ_hyps); + assert (List.is_empty mib.mind_hyps && UVars.Instance.is_empty mib.mind_univ_hyps); { mind_record = subst_mind_record subst mib.mind_record ; mind_finite = mib.mind_finite ; mind_ntypes = mib.mind_ntypes ; mind_hyps = []; - mind_univ_hyps = Univ.Instance.empty; + mind_univ_hyps = UVars.Instance.empty; mind_nparams = mib.mind_nparams; mind_nparams_rec = mib.mind_nparams_rec; mind_params_ctxt = @@ -303,45 +303,27 @@ let inductive_make_projection ind mib ~proj_arg = | NotRecord | FakeRecord -> CErrors.anomaly Pp.(str "inductive_make_projection: not a primitive record.") | PrimRecord infos -> - let _, labs, relevances, _ = infos.(snd ind) in - let proj_relevant = match relevances.(proj_arg) with - | Sorts.Irrelevant -> false - | Sorts.Relevant -> true - | Sorts.RelevanceVar _ -> assert false - in + let _, labs, rs, _ = infos.(snd ind) in if proj_arg < 0 || Array.length labs <= proj_arg then CErrors.anomaly Pp.(str "inductive_make_projection: invalid proj_arg."); - Names.Projection.Repr.make ind - ~proj_relevant + let p = Names.Projection.Repr.make ind ~proj_npars:mib.mind_nparams ~proj_arg labs.(proj_arg) + in + p, rs.(proj_arg) let inductive_make_projections ind mib = match mib.mind_record with | NotRecord | FakeRecord -> None | PrimRecord infos -> let _, labs, relevances, _ = infos.(snd ind) in - let projs = Array.mapi (fun proj_arg lab -> - let proj_relevant = match relevances.(proj_arg) with - | Sorts.Irrelevant -> false - | Sorts.Relevant -> true - | Sorts.RelevanceVar _ -> assert false - in - Names.Projection.Repr.make ind ~proj_relevant ~proj_npars:mib.mind_nparams ~proj_arg lab) - labs + let projs = Array.map2_i (fun proj_arg lab r -> + Names.Projection.Repr.make ind ~proj_npars:mib.mind_nparams ~proj_arg lab, r) + labs relevances 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 4b80b8497ad8..b959c6bb8e80 100644 --- a/kernel/declareops.mli +++ b/kernel/declareops.mli @@ -10,14 +10,14 @@ open Declarations open Mod_subst -open Univ +open UVars (** Operations concerning types in [Declarations] : [constant_body], [mutual_inductive_body], [module_body] ... *) val universes_context : universes -> AbstractContext.t -val abstract_universes : Entries.universes_entry -> Univ.universe_level_subst * universes +val abstract_universes : Entries.universes_entry -> UVars.sort_level_subst * universes (** {6 Arities} *) @@ -70,12 +70,10 @@ val inductive_is_cumulative : mutual_inductive_body -> bool (** Anomaly when not a primitive record or invalid proj_arg *) val inductive_make_projection : Names.inductive -> mutual_inductive_body -> proj_arg:int -> - Names.Projection.Repr.t + Names.Projection.Repr.t * Sorts.relevance 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 + (Names.Projection.Repr.t * Sorts.relevance) array option (** {6 Kernel flags} *) diff --git a/kernel/discharge.ml b/kernel/discharge.ml index 297a1666037a..9d4758a1e9b3 100644 --- a/kernel/discharge.ml +++ b/kernel/discharge.ml @@ -13,7 +13,7 @@ open Names open Constr open Term open Declarations -open Univ +open UVars open Cooking module NamedDecl = Context.Named.Declaration @@ -21,15 +21,16 @@ module RelDecl = Context.Rel.Declaration let lift_univs info univ_hyps = function | Monomorphic -> - assert (Univ.Instance.is_empty univ_hyps); + assert (UVars.Instance.is_empty univ_hyps); info, univ_hyps, Monomorphic | Polymorphic auctx -> - let info, n, auctx = lift_poly_univs info auctx in + let info, (qn, un), auctx = lift_poly_univs info auctx in let univ_hyps = - let open Univ.Instance in - let us = to_array univ_hyps in - let us = Array.sub us 0 (Array.length us - n) in - of_array us + let open UVars.Instance in + let qs, us = to_array univ_hyps in + let qs = Array.sub qs 0 (Array.length qs - qn) in + let us = Array.sub us 0 (Array.length us - un) in + of_array (qs,us) in info, univ_hyps, Polymorphic auctx @@ -174,9 +175,10 @@ let cook_inductive info mib = | None, None -> None, None | None, Some _ | Some _, None -> assert false | Some variance, Some sec_variance -> + (* no variance for qualities *) + let ulen = snd (AbstractContext.size (universe_context_of_cooking_info info)) in let sec_variance, newvariance = - Array.chop (Array.length sec_variance - AbstractContext.size (universe_context_of_cooking_info info)) - sec_variance + Array.chop (Array.length sec_variance - ulen) sec_variance in Some (Array.append newvariance variance), Some sec_variance in diff --git a/kernel/entries.mli b/kernel/entries.mli index 00c9b080b2ba..747610f9493c 100644 --- a/kernel/entries.mli +++ b/kernel/entries.mli @@ -18,14 +18,14 @@ open Constr type universes_entry = | Monomorphic_entry - | Polymorphic_entry of Univ.UContext.t + | Polymorphic_entry of UVars.UContext.t type inductive_universes_entry = | Monomorphic_ind_entry - | Polymorphic_ind_entry of Univ.UContext.t + | Polymorphic_ind_entry of UVars.UContext.t | Template_ind_entry of Univ.ContextSet.t -type variance_entry = Univ.Variance.t option array +type variance_entry = UVars.Variance.t option array type 'a in_universes_entry = 'a * universes_entry @@ -111,7 +111,7 @@ type constant_entry = (** {6 Modules } *) -type module_struct_entry = (constr * Univ.AbstractContext.t option) Declarations.module_alg_expr +type module_struct_entry = (constr * UVars.AbstractContext.t option) Declarations.module_alg_expr type module_params_entry = (MBId.t * module_struct_entry * inline) list (** older first *) diff --git a/kernel/environ.ml b/kernel/environ.ml index 31de87a315ab..1885f0e2256b 100644 --- a/kernel/environ.ml +++ b/kernel/environ.ml @@ -81,7 +81,8 @@ type env = { env_nb_rel : int; env_universes : UGraph.t; env_universes_lbound : UGraph.Bound.t; - irr_constants : Cset_env.t; + env_qualities : Sorts.QVar.Set.t; + irr_constants : Sorts.relevance Cmap_env.t; irr_inds : Indset_env.t; env_typing_flags : typing_flags; retroknowledge : Retroknowledge.retroknowledge; @@ -110,7 +111,8 @@ let empty_env = { env_nb_rel = 0; env_universes = UGraph.initial_universes; env_universes_lbound = UGraph.Bound.Set; - irr_constants = Cset_env.empty; + env_qualities = Sorts.QVar.Set.empty; + irr_constants = Cmap_env.empty; irr_inds = Indset_env.empty; env_typing_flags = Declareops.safe_flags Conv_oracle.empty; retroknowledge = Retroknowledge.empty; @@ -375,13 +377,28 @@ let check_constraints c env = UGraph.check_constraints c env.env_universes let add_universes ~lbound ~strict ctx g = + let _qs, us = UVars.Instance.to_array (UVars.UContext.instance ctx) in let g = Array.fold_left - (fun g v -> UGraph.add_universe ~lbound ~strict v g) - g (Univ.Instance.to_array (Univ.UContext.instance ctx)) + (fun g v -> UGraph.add_universe ~lbound ~strict v g) + g us in - UGraph.merge_constraints (Univ.UContext.constraints ctx) g + UGraph.merge_constraints (UVars.UContext.constraints ctx) g + +let add_qualities qs known = + let open Sorts.Quality in + Array.fold_left (fun known q -> + match q with + | QVar q -> + let known' = Sorts.QVar.Set.add q known in + let () = if known == known' then CErrors.anomaly Pp.(str"multiply bound sort quality") in + known' + | QConstant _ -> CErrors.anomaly Pp.(str "constant quality in ucontext")) + known + qs let push_context ?(strict=false) ctx env = + let qs, _us = UVars.Instance.to_array (UVars.UContext.instance ctx) in + let env = { env with env_qualities = add_qualities qs env.env_qualities } in map_universes (add_universes ~lbound:(universes_lbound env) ~strict ctx) env let add_universes_set ~lbound ~strict ctx g = @@ -468,8 +485,8 @@ let add_constant_key kn cb linkinfo env = { env.env_globals with Globals.constants = new_constants } in - let irr_constants = if cb.const_relevance == Sorts.Irrelevant - then Cset_env.add kn env.irr_constants + let irr_constants = if cb.const_relevance != Sorts.Relevant + then Cmap_env.add kn cb.const_relevance env.irr_constants else env.irr_constants in { env with irr_constants; env_globals = new_globals } @@ -481,20 +498,20 @@ let add_constant kn cb env = let constant_type env (kn,u) = let cb = lookup_constant kn env in let uctx = Declareops.constant_polymorphic_context cb in - let csts = Univ.AbstractContext.instantiate u uctx in + let csts = UVars.AbstractContext.instantiate u uctx in (subst_instance_constr u cb.const_type, csts) type const_evaluation_result = | NoBody | Opaque - | IsPrimitive of Univ.Instance.t * CPrimitives.t + | IsPrimitive of UVars.Instance.t * CPrimitives.t exception NotEvaluableConst of const_evaluation_result let constant_value_and_type env (kn, u) = let cb = lookup_constant kn env in let uctx = Declareops.constant_polymorphic_context cb in - let cst = Univ.AbstractContext.instantiate u uctx in + let cst = UVars.AbstractContext.instantiate u uctx in let b' = match cb.const_body with | Def l_body -> Some (subst_instance_constr u l_body) | OpaqueDef _ -> None @@ -568,7 +585,7 @@ let polymorphic_constant cst env = Declareops.constant_is_polymorphic (lookup_constant cst env) let polymorphic_pconstant (cst,u) env = - if Univ.Instance.is_empty u then false + if UVars.Instance.is_empty u then false else polymorphic_constant cst env let type_in_type_constant cst env = @@ -582,8 +599,9 @@ 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 - typs.(Projection.arg p) + let _,_,rs,typs = infos.(i) in + let arg = Projection.arg p in + rs.(arg), typs.(arg) let get_projection env ind ~proj_arg = let mib = lookup_mind (fst ind) env in @@ -598,7 +616,7 @@ let polymorphic_ind (mind,_i) env = Declareops.inductive_is_polymorphic (lookup_mind mind env) let polymorphic_pind (ind,u) env = - if Univ.Instance.is_empty u then false + if UVars.Instance.is_empty u then false else polymorphic_ind ind env let type_in_type_ind (mind,_i) env = @@ -616,7 +634,7 @@ let template_polymorphic_variables (mind, _) env = | None -> [] let template_polymorphic_pind (ind,u) env = - if not (Univ.Instance.is_empty u) then false + if not (UVars.Instance.is_empty u) then false else template_polymorphic_ind ind env let add_mind_key kn (mind, _ as mind_key) env = @@ -656,7 +674,7 @@ let constant_context env c = let universes_of_global env r = let open GlobRef in match r with - | VarRef _ -> Univ.AbstractContext.empty + | VarRef _ -> UVars.AbstractContext.empty | ConstRef c -> constant_context env c | IndRef (mind,_) | ConstructRef ((mind,_),_) -> let mib = lookup_mind mind env in diff --git a/kernel/environ.mli b/kernel/environ.mli index 59c00a706dfe..5ed491d30686 100644 --- a/kernel/environ.mli +++ b/kernel/environ.mli @@ -11,6 +11,7 @@ open Names open Constr open Univ +open UVars open Declarations (** Unsafe environments. We define here a datatype for environments. @@ -72,7 +73,8 @@ type env = private { env_nb_rel : int; env_universes : UGraph.t; env_universes_lbound : UGraph.Bound.t; - irr_constants : Cset_env.t; + env_qualities : Sorts.QVar.Set.t; + irr_constants : Sorts.relevance Cmap_env.t; irr_inds : Indset_env.t; env_typing_flags : typing_flags; retroknowledge : Retroknowledge.retroknowledge; @@ -212,16 +214,16 @@ val type_in_type_constant : Constant.t -> env -> bool type const_evaluation_result = | NoBody | Opaque - | IsPrimitive of Univ.Instance.t * CPrimitives.t + | IsPrimitive of Instance.t * CPrimitives.t exception NotEvaluableConst of const_evaluation_result val constant_type : env -> Constant.t puniverses -> types constrained val constant_value_and_type : env -> Constant.t puniverses -> - constr option * types * Univ.Constraints.t + constr option * types * Constraints.t (** The universe context associated to the constant, empty if not polymorphic *) -val constant_context : env -> Constant.t -> Univ.AbstractContext.t +val constant_context : env -> Constant.t -> AbstractContext.t (* These functions should be called under the invariant that [env] already contains the constraints corresponding to the constant @@ -242,12 +244,12 @@ val is_primitive_type : env -> Constant.t -> bool (** {6 Primitive projections} *) (** Checks that the number of parameters is correct. *) -val lookup_projection : Names.Projection.t -> env -> types +val lookup_projection : Names.Projection.t -> env -> Sorts.relevance * types (** Anomaly when not a primitive record or invalid proj_arg. *) -val get_projection : env -> inductive -> proj_arg:int -> Names.Projection.Repr.t +val get_projection : env -> inductive -> proj_arg:int -> Names.Projection.Repr.t * Sorts.relevance -val get_projections : env -> inductive -> Names.Projection.Repr.t array option +val get_projections : env -> inductive -> (Names.Projection.Repr.t * Sorts.relevance) array option (** {5 Inductive types } *) val lookup_mind_key : MutInd.t -> env -> mind_key @@ -262,7 +264,7 @@ val mem_mind : MutInd.t -> env -> bool (** The universe context associated to the inductive, empty if not polymorphic *) -val mind_context : env -> MutInd.t -> Univ.AbstractContext.t +val mind_context : env -> MutInd.t -> AbstractContext.t (** New-style polymorphism *) val polymorphic_ind : inductive -> env -> bool @@ -271,7 +273,7 @@ val type_in_type_ind : inductive -> env -> bool (** Old-style polymorphism *) val template_polymorphic_ind : inductive -> env -> bool -val template_polymorphic_variables : inductive -> env -> Univ.Level.t list +val template_polymorphic_variables : inductive -> env -> Level.t list val template_polymorphic_pind : pinductive -> env -> bool (** {6 Name quotients} *) @@ -312,23 +314,23 @@ val lookup_modtype : ModPath.t -> env -> module_type_body (** {5 Universe constraints } *) -val add_constraints : Univ.Constraints.t -> env -> env +val add_constraints : Constraints.t -> env -> env (** Add universe constraints to the environment. @raise UniverseInconsistency. *) -val check_constraints : Univ.Constraints.t -> env -> bool +val check_constraints : Constraints.t -> env -> bool (** Check constraints are satifiable in the environment. *) -val push_context : ?strict:bool -> Univ.UContext.t -> env -> env +val push_context : ?strict:bool -> UContext.t -> env -> env (** [push_context ?(strict=false) ctx env] pushes the universe context to the environment. @raise UGraph.AlreadyDeclared if one of the universes is already declared. *) -val push_context_set : ?strict:bool -> Univ.ContextSet.t -> env -> env +val push_context_set : ?strict:bool -> ContextSet.t -> env -> env (** [push_context_set ?(strict=false) ctx env] pushes the universe context set to the environment. It does not fail even if one of the universes is already declared. *) -val push_subgraph : Univ.ContextSet.t -> env -> env +val push_subgraph : ContextSet.t -> env -> env (** [push_subgraph univs env] adds the universes and constraints in [univs] to [env] as [push_context_set ~strict:false univs env], and also checks that they do not imply new transitive constraints @@ -400,7 +402,7 @@ val remove_hyps : Id.Set.t -> (Constr.named_declaration -> Constr.named_declarat val is_polymorphic : env -> Names.GlobRef.t -> bool val is_template_polymorphic : env -> GlobRef.t -> bool -val get_template_polymorphic_variables : env -> GlobRef.t -> Univ.Level.t list +val get_template_polymorphic_variables : env -> GlobRef.t -> Level.t list val is_type_in_type : env -> GlobRef.t -> bool (** Native compiler *) diff --git a/kernel/genlambda.ml b/kernel/genlambda.ml index 22bbdfd17e60..f2b1d9df3d29 100644 --- a/kernel/genlambda.ml +++ b/kernel/genlambda.ml @@ -631,7 +631,7 @@ let rec lambda_of_constr cache env sigma c = | Construct _ -> lambda_of_app cache env sigma c empty_args - | Proj (p, c) -> + | Proj (p, _, c) -> let c = lambda_of_constr cache env sigma c in Lproj (Projection.repr p, c) diff --git a/kernel/indTyping.ml b/kernel/indTyping.ml index fa2d03dcb73d..928f4d0e1d36 100644 --- a/kernel/indTyping.ml +++ b/kernel/indTyping.ml @@ -11,6 +11,7 @@ open Util open Names open Univ +open UVars open Term open Constr open Declarations @@ -64,25 +65,91 @@ let mind_check_names mie = (************************** Type checking *******************************) (************************************************************************) -type univ_info = { ind_squashed : bool; ind_has_relevant_arg : bool; - ind_template : bool; - ind_univ : Sorts.t; - missing : Sorts.t list; (* missing u <= ind_univ constraints *) - } let no_sort_variable () = - CErrors.anomaly (Pp.str "A sort variable was sent to the kernel") + CErrors.user_err (Pp.str "Sort variables not yet supported for the inductive's sort.") + +type record_arg_info = + | NoRelevantArg + | HasRelevantArg + +type univ_info = + { ind_squashed : bool + ; record_arg_info : record_arg_info + ; ind_template : bool + ; ind_univ : Sorts.t + ; missing : Sorts.t list (* missing u <= ind_univ constraints *) + } + +(* TODO squash depending on the instance + (so eg in the "QSort(q, _), Prop" case, "@{q:=Prop|}" is not squashed + but "@{q:=Type|}" does need squashing) + Cases which will be modified are annotated with "imprecise". + This code can probably be simplified but I can't quite see how right now. *) let check_univ_leq ?(is_real_arg=false) env u info = - let ind_univ = info.ind_univ in - let info = if not info.ind_has_relevant_arg && is_real_arg && not (Sorts.is_sprop u) - then {info with ind_has_relevant_arg=true} - else info + let info = if not is_real_arg then info + else match info.record_arg_info with + | NoRelevantArg | HasRelevantArg -> match u with + | Sorts.SProp | QSort _ -> info + | Prop | Set | Type _ -> { info with record_arg_info = HasRelevantArg } in - (* Inductive types provide explicit lifting from SProp to other universes, so allow SProp <= any. *) - if Sorts.is_sprop u || UGraph.check_leq_sort (universes env) u ind_univ then info - else if is_impredicative_sort env ind_univ then { info with ind_squashed = true } - else {info with missing = u :: info.missing} + (* If we would squash (eg Prop, SProp case) we still need to check the type in type flag. *) + let ind_squashed = not (Environ.type_in_type env) in + match u, info.ind_univ with + | SProp, (SProp | Prop | Set | QSort _ | Type _) -> + (* Inductive types provide explicit lifting from SProp to other universes, + so allow SProp <= any. *) + info + + | Prop, SProp -> { info with ind_squashed } + | Prop, QSort _ -> { info with ind_squashed } (* imprecise *) + | Prop, (Prop | Set | Type _) -> info + + | Set, (SProp | Prop) -> { info with ind_squashed } + | Set, QSort (_, indu) -> + if UGraph.check_leq (universes env) Universe.type0 indu + then { info with ind_squashed } (* imprecise *) + else { info with missing = u :: info.missing } + | Set, Set -> info + | Set, Type indu -> + if UGraph.check_leq (universes env) Universe.type0 indu + then info + else { info with missing = u :: info.missing } + + | QSort _, (SProp | Prop) -> { info with ind_squashed } (* imprecise *) + | QSort (cq, uu), QSort (indq, indu) -> + if UGraph.check_leq (universes env) uu indu + then begin if Sorts.QVar.equal cq indq then info + else { info with ind_squashed } (* imprecise *) + end + else { info with missing = u :: info.missing } + | QSort (_, uu), Set -> + if UGraph.check_leq (universes env) uu Universe.type0 + then info + else if is_impredicative_set env + then { info with ind_squashed } (* imprecise *) + else { info with missing = u :: info.missing } + | QSort (_,uu), Type indu -> + if UGraph.check_leq (universes env) uu indu + then info + else { info with missing = u :: info.missing } + + | Type _, (SProp | Prop) -> { info with ind_squashed } + | Type uu, Set -> + if UGraph.check_leq (universes env) uu Universe.type0 + then info + else if is_impredicative_set env + then { info with ind_squashed } + else { info with missing = u :: info.missing } + | Type uu, QSort (_, indu) -> + if UGraph.check_leq (universes env) uu indu + then { info with ind_squashed } (* imprecise *) + else { info with missing = u :: info.missing } + | Type uu, Type indu -> + if UGraph.check_leq (universes env) uu indu + then info + else { info with missing = u :: info.missing } let check_context_univs ~ctor env info ctx = let check_one d (info,env) = @@ -107,7 +174,7 @@ let check_arity ~template env_params env_ar ind = let indices, ind_sort = Reduction.dest_arity env_params arity in let univ_info = { ind_squashed=false; - ind_has_relevant_arg=false; + record_arg_info=NoRelevantArg; ind_template = template; ind_univ=ind_sort; missing=[]; @@ -159,9 +226,11 @@ let check_record data = List.for_all (fun (_,(_,splayed_lc),info) -> (* records must have all projections definable -> equivalent to not being squashed *) not info.ind_squashed - (* relevant records must have at least 1 relevant argument *) - && (Sorts.is_sprop info.ind_univ - || info.ind_has_relevant_arg) + (* relevant records must have at least 1 relevant argument, + and we don't yet support variable relevance projections *) + && (match info.record_arg_info with + | HasRelevantArg -> true + | NoRelevantArg -> Sorts.is_sprop info.ind_univ) && (match splayed_lc with (* records must have 1 constructor with at least 1 argument, and no anonymous fields *) | [|ctx,_|] -> @@ -183,7 +252,7 @@ let check_record data = (* - all_sorts in case of small, unitary Prop (not smashed) *) (* - logical_sorts in case of large, unitary Prop (smashed) *) -let allowed_sorts {ind_squashed;ind_univ;ind_template=_;ind_has_relevant_arg=_;missing=_} = +let allowed_sorts {ind_squashed;ind_univ;ind_template=_;record_arg_info=_;missing=_} = if not ind_squashed then InType else Sorts.family ind_univ @@ -259,9 +328,9 @@ let abstract_packets usubst ((arity,lc),(indices,splayed_lc),univ_info) = splayed_lc in let ind_univ = match univ_info.ind_univ with - | Prop | SProp | Set -> univ_info.ind_univ - | Type u -> Sorts.sort_of_univ (Univ.subst_univs_level_universe usubst u) - | QSort _ -> no_sort_variable () + | QSort _ -> no_sort_variable () + | _ -> + UVars.subst_sort_level_sort usubst univ_info.ind_univ in let arity = @@ -340,11 +409,14 @@ let typecheck_inductive env ~sec_univs (mie:mutual_inductive_entry) = | Monomorphic_ind_entry | Template_ind_entry _ -> CErrors.user_err Pp.(str "Inductive cannot be both monomorphic and universe cumulative.") | Polymorphic_ind_entry uctx -> - let univs = Instance.to_array @@ UContext.instance uctx in + (* no variance for qualities *) + let _qualities, univs = Instance.to_array @@ UContext.instance uctx in let univs = Array.map2 (fun a b -> a,b) univs variances in let univs = match sec_univs with | None -> univs | Some sec_univs -> + (* no variance for qualities *) + let _, sec_univs = UVars.Instance.to_array sec_univs in let sec_univs = Array.map (fun u -> u, None) sec_univs in Array.append sec_univs univs in @@ -361,10 +433,10 @@ let typecheck_inductive env ~sec_univs (mie:mutual_inductive_entry) = (* Abstract universes *) let usubst, univs = match mie.mind_entry_universes with | Monomorphic_ind_entry | Template_ind_entry _ -> - Univ.empty_level_subst, Monomorphic + UVars.empty_sort_subst, Monomorphic | Polymorphic_ind_entry uctx -> - let (inst, auctx) = Univ.abstract_universes uctx in - let inst = Univ.make_instance_subst inst in + let (inst, auctx) = UVars.abstract_universes uctx in + let inst = UVars.make_instance_subst inst in (inst, Polymorphic auctx) in let params = Vars.subst_univs_level_context usubst params in diff --git a/kernel/indTyping.mli b/kernel/indTyping.mli index f47ec7df05f7..426d8c76ec94 100644 --- a/kernel/indTyping.mli +++ b/kernel/indTyping.mli @@ -25,12 +25,12 @@ open Declarations * (indices * splayed constructor types) (both without params) * top allowed elimination *) -val typecheck_inductive : env -> sec_univs:Univ.Level.t array option +val typecheck_inductive : env -> sec_univs:UVars.Instance.t option -> mutual_inductive_entry -> env * universes * template_universes option - * Univ.Variance.t array option + * UVars.Variance.t array option * Names.Id.t array option option * Constr.rel_context * ((inductive_arity * Constr.types array) * diff --git a/kernel/indtypes.ml b/kernel/indtypes.ml index 4fa02e1cb6f8..182d433fa96b 100644 --- a/kernel/indtypes.ml +++ b/kernel/indtypes.ml @@ -454,12 +454,7 @@ let compute_projections (_, i as ind) mib = | Name id -> let r = na.Context.binder_relevance in let lab = Label.of_id id in - let proj_relevant = match r with - | Sorts.Irrelevant -> false - | Sorts.Relevant -> true - | Sorts.RelevanceVar _ -> assert false - in - let kn = Projection.Repr.make ind ~proj_relevant ~proj_npars:mib.mind_nparams ~proj_arg:i lab 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)] to [params, x:I, field1,..,fieldj |- t(params,field1,..,fieldj] *) let t = liftn 1 j t in @@ -468,7 +463,7 @@ let compute_projections (_, i as ind) mib = let projty = substl letsubst t in (* from [params, x:I, field1,..,fieldj |- t(field1,..,fieldj)] to [params, x:I |- t(proj1 x,..,projj x)] *) - let fterm = mkProj (Projection.make kn false, mkRel 1) in + let fterm = mkProj (Projection.make kn false, r, mkRel 1) in (i + 1, j + 1, lab :: labs, r :: rs, projty :: pbs, fterm :: letsubst) | Anonymous -> assert false (* checked by indTyping *) in @@ -485,7 +480,7 @@ let build_inductive env ~sec_univs names prv univs template variance (* Compute the set of used section variables *) let hyps = used_section_variables env paramsctxt inds in let nparamargs = Context.Rel.nhyps paramsctxt in - let u = Univ.make_abstract_instance (universes_context univs) in + let u = UVars.make_abstract_instance (universes_context univs) in let subst = List.init ntypes (fun i -> mkIndU ((kn, ntypes - i - 1), u)) in (* Check one inductive *) let build_one_packet (id,cnames) ((arity,lc),(indices,splayed_lc),kelim) recarg = @@ -542,13 +537,14 @@ let build_inductive env ~sec_univs names prv univs template variance | Some variance -> match sec_univs with | None -> Some variance, None | Some sec_univs -> - let nsec = Array.length sec_univs in - Some (Array.sub variance nsec (Array.length variance - nsec)), - Some (Array.sub variance 0 nsec) + (* no variance for qualities *) + let _nsecq, nsecu = UVars.Instance.length sec_univs in + Some (Array.sub variance nsecu (Array.length variance - nsecu)), + Some (Array.sub variance 0 nsecu) in let univ_hyps = match sec_univs with - | None -> Univ.Instance.empty - | Some univs -> Univ.Instance.of_array univs + | None -> UVars.Instance.empty + | Some univs -> univs in let mib = (* Build the mutual inductive *) diff --git a/kernel/indtypes.mli b/kernel/indtypes.mli index f194ab28837e..08d18538fae6 100644 --- a/kernel/indtypes.mli +++ b/kernel/indtypes.mli @@ -14,5 +14,5 @@ open Environ open Entries (** Check an inductive. *) -val check_inductive : env -> sec_univs:Univ.Level.t array option +val check_inductive : env -> sec_univs:UVars.Instance.t option -> MutInd.t -> mutual_inductive_entry -> mutual_inductive_body diff --git a/kernel/inductive.ml b/kernel/inductive.ml index 447e2dddb71c..d84b6e42a208 100644 --- a/kernel/inductive.ml +++ b/kernel/inductive.ml @@ -12,6 +12,7 @@ open CErrors open Util open Names open Univ +open UVars open Constr open Vars open Declarations @@ -59,7 +60,7 @@ let inductive_nonrec_rec_paramdecls (mib,u) = Context.Rel.chop_nhyps nnonrecparamdecls paramdecls let instantiate_inductive_constraints mib u = - Univ.AbstractContext.instantiate u (Declareops.inductive_polymorphic_context mib) + UVars.AbstractContext.instantiate u (Declareops.inductive_polymorphic_context mib) (************************************************************************) @@ -347,7 +348,7 @@ let expand_arity (mib, mip) (ind, u) params nas = let params = Vars.subst_of_rel_context_instance paramdecl params in let realdecls, _ = List.chop mip.mind_nrealdecls mip.mind_arity_ctxt in let self = - let u = Univ.(Instance.of_array (Array.init (Instance.length u) Level.var)) in + let u = Instance.abstract_instance (Instance.length u) in let args = Context.Rel.instance mkRel 0 mip.mind_arity_ctxt in mkApp (mkIndU (ind, u), args) in @@ -440,7 +441,7 @@ let check_case_info env (indsp,u) r ci = 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) || + not (Sorts.relevance_equal ci.ci_relevance r) || is_primitive_record spec then raise (TypeError(env,WrongCaseInfo((indsp,u),ci))) @@ -965,7 +966,7 @@ let rec subterm_specif renv stack t = (* Metas and evars are considered OK *) | (Meta _|Evar _) -> Dead_code - | Proj (p, c) -> + | Proj (p, _, c) -> let subt = subterm_specif renv stack c in (match subt with | Subterm (_, _s, wf) -> @@ -1285,7 +1286,7 @@ let check_one_fix renv recpos trees def = | Ind _ | Construct _ -> check_rec_call_state renv NoNeedReduce stack rs (fun () -> None) - | Proj (p, c) -> + | Proj (p, _, c) -> begin let needreduce', rs = check_rec_call renv rs c in check_rec_call_state renv needreduce' stack rs (fun () -> diff --git a/kernel/inductive.mli b/kernel/inductive.mli index b4ca3e9c1197..049618b00734 100644 --- a/kernel/inductive.mli +++ b/kernel/inductive.mli @@ -11,6 +11,7 @@ open Names open Constr open Univ +open UVars open Declarations open Environ diff --git a/kernel/inferCumulativity.ml b/kernel/inferCumulativity.ml index 86c186e31ef6..7f7a2d968097 100644 --- a/kernel/inferCumulativity.ml +++ b/kernel/inferCumulativity.ml @@ -12,6 +12,7 @@ open Conversion open Declarations open Constr open Univ +open UVars open Variance open Util @@ -107,13 +108,17 @@ open Inf let infer_generic_instance_eq variances u = Array.fold_left (fun variances u -> infer_level_eq u variances) - variances (Instance.to_array u) + variances + u + +(* no variance for qualities *) +let instance_univs u = snd (Instance.to_array u) let extend_con_instance cb u = - Instance.(of_array (Array.append (to_array cb.const_univ_hyps) (to_array u))) + (Array.append (instance_univs cb.const_univ_hyps) (instance_univs u)) let extend_ind_instance mib u = - Instance.(of_array (Array.append (to_array mib.mind_univ_hyps) (to_array u))) + (Array.append (instance_univs mib.mind_univ_hyps) (instance_univs u)) let extended_mind_variance mind = match mind.mind_variance, mind.mind_sec_variance with @@ -128,7 +133,9 @@ let infer_cumulative_ind_instance cv_pb mind_variance variances u = | _, Irrelevant -> variances | _, Invariant | CONV, Covariant -> infer_level_eq u variances | CUMUL, Covariant -> infer_level_leq u variances) - variances mind_variance (Instance.to_array u) + variances + mind_variance + u let infer_inductive_instance cv_pb env variances ind nargs u = let mind = Environ.lookup_mind (fst ind) env in @@ -202,7 +209,7 @@ let rec infer_fterm cv_pb infos variances hd stk = | None -> raise e | Some (hd,stk) -> infer_fterm cv_pb infos variances hd stk end - | FProj (_,c) -> + | FProj (_,_,c) -> let variances = infer_fterm CONV infos variances c [] in infer_stack infos variances stk | FLambda _ -> @@ -210,6 +217,7 @@ let rec infer_fterm cv_pb infos variances hd stk = let variances = infer_fterm CONV infos variances ty [] in infer_fterm CONV (push_relevance infos na) variances bd [] | FProd (na,dom,codom,e) -> + let na = usubst_binder e na in let variances = infer_fterm CONV infos variances dom [] in infer_fterm cv_pb (push_relevance infos na) variances (mk_clos (CClosure.usubs_lift e) codom) [] | FInd (ind, u) -> @@ -229,12 +237,13 @@ let rec infer_fterm cv_pb infos variances hd stk = let variances = infer_vect infos variances (Array.map (mk_clos e) tys) in let le = CClosure.usubs_liftn n e in let variances = + let na = Array.map (usubst_binder e) na in let infos = push_relevances infos na in infer_vect infos variances (Array.map (mk_clos le) cl) in infer_stack infos variances stk | FArray (u,elemsdef,ty) -> - let variances = infer_generic_instance_eq variances u in + let variances = infer_generic_instance_eq variances (instance_univs u) in let variances = infer_fterm CONV infos variances ty [] in let elems, def = Parray.to_array elemsdef in let variances = infer_fterm CONV infos variances def [] in diff --git a/kernel/inferCumulativity.mli b/kernel/inferCumulativity.mli index 97accadf63f9..10eeeb2e71cf 100644 --- a/kernel/inferCumulativity.mli +++ b/kernel/inferCumulativity.mli @@ -16,6 +16,6 @@ val infer_inductive (** Environment containing the polymorphic universes and the inductives then the parameters. *) -> arities : Constr.t list -> ctors : Constr.t list list - -> (Univ.Level.t * Univ.Variance.t option) array + -> (Univ.Level.t * UVars.Variance.t option) array (** Universes whose cumulativity we want to infer or check. *) - -> Univ.Variance.t array + -> UVars.Variance.t array diff --git a/kernel/mod_subst.ml b/kernel/mod_subst.ml index 3f9f4dfe536a..ef667e9133a6 100644 --- a/kernel/mod_subst.ml +++ b/kernel/mod_subst.ml @@ -25,7 +25,7 @@ open Constr Equiv gives the canonical name in the given context. *) type delta_hint = - | Inline of int * constr Univ.univ_abstracted option + | Inline of int * constr UVars.univ_abstracted option | Equiv of KerName.t (* NB: earlier constructor Prefix_equiv of ModPath.t @@ -165,7 +165,7 @@ let find_prefix resolve mp = (** Applying a resolver to a kernel name *) -exception Change_equiv_to_inline of (int * constr Univ.univ_abstracted) +exception Change_equiv_to_inline of (int * constr UVars.univ_abstracted) let solve_delta_kn resolve kn = try @@ -344,11 +344,11 @@ let rec map_kn f f' c = let func = map_kn f f' in match kind c with | Const kn -> (try f' kn with No_subst -> c) - | Proj (p,t) -> + | Proj (p,r,t) -> let p' = Projection.map f p in let t' = func t in if p' == p && t' == t then c - else mkProj (p', t') + else mkProj (p', r, t') | Ind ((kn,i),u) -> let kn' = f kn in if kn'==kn then c else mkIndU ((kn',i),u) @@ -490,7 +490,7 @@ let gen_subst_delta_resolver dom subst resolver = | Equiv kequ -> (try Equiv (subst_kn_delta subst kequ) with Change_equiv_to_inline (lev,c) -> Inline (lev,Some c)) - | Inline (lev,Some t) -> Inline (lev,Some (Univ.map_univ_abstracted (subst_mps subst) t)) + | Inline (lev,Some t) -> Inline (lev,Some (UVars.map_univ_abstracted (subst_mps subst) t)) | Inline (_,None) -> hint in Deltamap.add_kn kkey' hint' rslv diff --git a/kernel/mod_subst.mli b/kernel/mod_subst.mli index b77f533fd943..4acfb1f205ec 100644 --- a/kernel/mod_subst.mli +++ b/kernel/mod_subst.mli @@ -28,7 +28,7 @@ val add_kn_delta_resolver : KerName.t -> KerName.t -> delta_resolver -> delta_resolver val add_inline_delta_resolver : - KerName.t -> (int * constr Univ.univ_abstracted option) -> delta_resolver -> delta_resolver + KerName.t -> (int * constr UVars.univ_abstracted option) -> delta_resolver -> delta_resolver val add_delta_resolver : delta_resolver -> delta_resolver -> delta_resolver @@ -131,7 +131,7 @@ val subst_kn : substitution -> KerName.t -> KerName.t val subst_con : - substitution -> Constant.t -> Constant.t * constr Univ.univ_abstracted option + substitution -> Constant.t -> Constant.t * constr UVars.univ_abstracted option val subst_pcon : substitution -> pconstant -> pconstant diff --git a/kernel/mod_typing.ml b/kernel/mod_typing.ml index 5f59eaa0e6ee..1eab5352d1ea 100644 --- a/kernel/mod_typing.ml +++ b/kernel/mod_typing.ml @@ -103,7 +103,7 @@ let rec check_with_def (cst, ustate) env struc (idl, wth) mp reso = error_incorrect_with_constraint lab in (** Terms are compared in a context with De Bruijn universe indices *) - let env' = Environ.push_context ~strict:false (Univ.AbstractContext.repr uctx) env in + let env' = Environ.push_context ~strict:false (UVars.AbstractContext.repr uctx) env in let () = match cb.const_body with | Undef _ | OpaqueDef _ -> let j = Typeops.infer env' wth.w_def in diff --git a/kernel/modops.ml b/kernel/modops.ml index 19f3976e037c..12a790372a5f 100644 --- a/kernel/modops.ml +++ b/kernel/modops.ml @@ -48,7 +48,7 @@ type signature_mismatch_error = | NotEqualInductiveAliases | IncompatibleUniverses of UGraph.univ_inconsistency | IncompatiblePolymorphism of env * types * types - | IncompatibleConstraints of { got : Univ.AbstractContext.t; expect : Univ.AbstractContext.t } + | IncompatibleConstraints of { got : UVars.AbstractContext.t; expect : UVars.AbstractContext.t } | IncompatibleVariance type subtyping_trace_elt = @@ -328,7 +328,7 @@ let strengthen_const mp_from l cb resolver = | _ -> let kn = KerName.make mp_from l in let con = constant_of_delta_kn resolver kn in - let u = Univ.make_abstract_instance (Declareops.constant_polymorphic_context cb) in + let u = UVars.make_abstract_instance (Declareops.constant_polymorphic_context cb) in { cb with const_body = Def (mkConstU (con,u)); const_body_code = Some (Vmbytegen.compile_alias con) } @@ -583,7 +583,7 @@ let inline_delta_resolver env inl mp mbid mtb delta = | Undef _ | OpaqueDef _ | Primitive _ -> l | Def constr -> let ctx = Declareops.constant_polymorphic_context constant in - let constr = Univ.{univ_abstracted_value=constr; univ_abstracted_binder=ctx} in + let constr = {UVars.univ_abstracted_value=constr; univ_abstracted_binder=ctx} in add_inline_delta_resolver kn (lev, Some constr) l in make_inline delta constants diff --git a/kernel/modops.mli b/kernel/modops.mli index 588b8d0d95fc..49b5556b61cd 100644 --- a/kernel/modops.mli +++ b/kernel/modops.mli @@ -38,7 +38,7 @@ val implem_smart_map : (module_implementation -> module_implementation) val annotate_module_expression : module_expression -> module_signature -> - (module_type_body, (constr * Univ.AbstractContext.t option) module_alg_expr) functorize + (module_type_body, (constr * UVars.AbstractContext.t option) module_alg_expr) functorize val annotate_struct_body : structure_body -> module_signature -> module_signature @@ -112,7 +112,7 @@ type signature_mismatch_error = | NotEqualInductiveAliases | IncompatibleUniverses of UGraph.univ_inconsistency | IncompatiblePolymorphism of env * types * types - | IncompatibleConstraints of { got : Univ.AbstractContext.t; expect : Univ.AbstractContext.t } + | IncompatibleConstraints of { got : UVars.AbstractContext.t; expect : UVars.AbstractContext.t } | IncompatibleVariance type subtyping_trace_elt = diff --git a/kernel/names.ml b/kernel/names.ml index 14ebce13b868..4567d2d80cb0 100644 --- a/kernel/names.ml +++ b/kernel/names.ml @@ -831,13 +831,12 @@ struct module Repr = struct type t = { proj_ind : inductive; - proj_relevant : bool; proj_npars : int; proj_arg : int; proj_name : Label.t; } - let make proj_ind ~proj_npars ~proj_arg ~proj_relevant proj_name = - {proj_ind;proj_npars;proj_arg;proj_relevant;proj_name} + let make proj_ind ~proj_npars ~proj_arg proj_name = + {proj_ind;proj_npars;proj_arg;proj_name} let inductive c = c.proj_ind @@ -851,16 +850,9 @@ struct let arg c = c.proj_arg - let relevant c = c.proj_relevant - let hash p = Hashset.Combine.combinesmall p.proj_arg (ind_hash p.proj_ind) - let bcomp b1 b2 = match b1, b2 with - | true, true | false, false -> 0 - | true, false -> 1 - | false, true -> -1 - let compare_gen a b = let c = Int.compare a.proj_arg b.proj_arg in if c <> 0 then c @@ -868,9 +860,7 @@ struct let c = Int.compare a.proj_npars b.proj_npars in if c <> 0 then c else - let c = Label.compare a.proj_name b.proj_name in - if c <> 0 then c - else bcomp a.proj_relevant b.proj_relevant + Label.compare a.proj_name b.proj_name module SyntacticOrd = struct type nonrec t = t @@ -920,12 +910,11 @@ struct type u = (inductive -> inductive) * (Id.t -> Id.t) let hashcons (hind,hid) p = { proj_ind = hind p.proj_ind; - proj_relevant = p.proj_relevant; proj_npars = p.proj_npars; proj_arg = p.proj_arg; proj_name = hid p.proj_name } let eq p p' = - p == p' || (p.proj_relevant == p'.proj_relevant && p.proj_ind == p'.proj_ind && p.proj_npars == p'.proj_npars && p.proj_arg == p'.proj_arg && p.proj_name == p'.proj_name) + p == p' || (p.proj_ind == p'.proj_ind && p.proj_npars == p'.proj_npars && p.proj_arg == p'.proj_arg && p.proj_name == p'.proj_name) let hash = hash end module HashRepr = Hashcons.Make(Self_Hashcons) diff --git a/kernel/names.mli b/kernel/names.mli index b4d74190379e..aeb5958d6596 100644 --- a/kernel/names.mli +++ b/kernel/names.mli @@ -619,7 +619,7 @@ module Projection : sig module Repr : sig type t - val make : inductive -> proj_npars:int -> proj_arg:int -> proj_relevant:bool -> Label.t -> t + val make : inductive -> proj_npars:int -> proj_arg:int -> Label.t -> t include QNameS with type t := t @@ -631,7 +631,6 @@ module Projection : sig val npars : t -> int val arg : t -> int val label : t -> Label.t - val relevant : t -> bool val equal : t -> t -> bool [@@ocaml.deprecated "Use QProjection.equal"] val hash : t -> int [@@ocaml.deprecated "Use QProjection.hash"] diff --git a/kernel/nativecode.ml b/kernel/nativecode.ml index c7aeab28bd4c..8b2fb4c0c988 100644 --- a/kernel/nativecode.ml +++ b/kernel/nativecode.ml @@ -171,9 +171,20 @@ let eq_symbol sy1 sy2 = | SymbMatch sw1, SymbMatch sw2 -> eq_annot_sw sw1 sw2 | SymbInd ind1, SymbInd ind2 -> Ind.CanOrd.equal ind1 ind2 | SymbEvar evk1, SymbEvar evk2 -> Evar.equal evk1 evk2 - | SymbLevel l1, SymbLevel l2 -> Univ.Level.equal l1 l2 + | SymbInstance u1, SymbInstance u2 -> UVars.Instance.equal u1 u2 | SymbProj (i1, k1), SymbProj (i2, k2) -> Ind.CanOrd.equal i1 i2 && Int.equal k1 k2 - | _, _ -> false + + | (SymbValue _ + | SymbSort _ + | SymbName _ + | SymbConst _ + | SymbMatch _ + | SymbInd _ + | SymbEvar _ + | SymbInstance _ + | SymbProj _), _ + -> false + let hash_symbol symb = match symb with @@ -184,7 +195,7 @@ let hash_symbol symb = | SymbMatch sw -> combinesmall 5 (hash_annot_sw sw) | SymbInd ind -> combinesmall 6 (Ind.CanOrd.hash ind) | SymbEvar evk -> combinesmall 7 (Evar.hash evk) - | SymbLevel l -> combinesmall 8 (Univ.Level.hash l) + | SymbInstance u -> combinesmall 8 (UVars.Instance.hash u) | SymbProj (i, k) -> combinesmall 9 (combine (Ind.CanOrd.hash i) k) module HashedTypeSymbol = struct @@ -234,10 +245,10 @@ let get_evar tbl i = | SymbEvar ev -> ev | _ -> anomaly (Pp.str "get_evar failed.") -let get_level tbl i = +let get_instance tbl i = match tbl.(i) with - | SymbLevel u -> u - | _ -> anomaly (Pp.str "get_level failed.") + | SymbInstance u -> u + | _ -> anomaly (Pp.str "get_instance failed.") let get_proj tbl i = match tbl.(i) with @@ -284,7 +295,7 @@ type primitive = | MLnot | MLland | MLmagic - | MLarrayget + | MLsubst_instance_instance | MLparray_of_array | Get_value | Get_sort @@ -293,31 +304,100 @@ type primitive = | Get_match | Get_ind | Get_evar - | Get_level + | Get_instance | Get_proj | Get_symbols | Lazy | Coq_primitive of CPrimitives.t * bool (* check for accu *) + | Mk_empty_instance let eq_primitive p1 p2 = match p1, p2 with - | Mk_prod, Mk_prod -> true - | Mk_sort, Mk_sort -> true - | Mk_ind, Mk_ind -> true - | Mk_const, Mk_const -> true - | Mk_sw, Mk_sw -> true + | Mk_prod, Mk_prod + | Mk_sort, Mk_sort + | Mk_ind, Mk_ind + | Mk_const, Mk_const + | Mk_sw, Mk_sw + | Mk_proj, Mk_proj + | Is_int, Is_int + | Is_float, Is_float + | Is_parray, Is_parray + | Cast_accu, Cast_accu + | Upd_cofix, Upd_cofix + | Force_cofix, Force_cofix + | Mk_uint, Mk_uint + | Mk_float, Mk_float + | Mk_int, Mk_int + | Val_to_int, Val_to_int + | Mk_evar, Mk_evar + | MLand, MLand + | MLnot, MLnot + | MLland, MLland + | MLmagic, MLmagic + | MLsubst_instance_instance, MLsubst_instance_instance + | MLparray_of_array, MLparray_of_array + | Get_value, Get_value + | Get_sort, Get_sort + | Get_name, Get_name + | Get_const, Get_const + | Get_match, Get_match + | Get_ind, Get_ind + | Get_evar, Get_evar + | Get_instance, Get_instance + | Get_proj, Get_proj + | Get_symbols, Get_symbols + | Lazy, Lazy + | Mk_empty_instance, Mk_empty_instance + -> true + | Mk_fix (rp1, i1), Mk_fix (rp2, i2) -> Int.equal i1 i2 && eq_rec_pos rp1 rp2 | Mk_cofix i1, Mk_cofix i2 -> Int.equal i1 i2 | Mk_rel i1, Mk_rel i2 -> Int.equal i1 i2 | Mk_var id1, Mk_var id2 -> Id.equal id1 id2 - | Cast_accu, Cast_accu -> true - | Upd_cofix, Upd_cofix -> true - | Force_cofix, Force_cofix -> true - | Mk_evar, Mk_evar -> true - | Mk_proj, Mk_proj -> true - | MLarrayget, MLarrayget -> true - - | _ -> false + | Coq_primitive (prim1,b1), Coq_primitive (prim2,b2) -> + CPrimitives.equal prim1 prim2 && Bool.equal b1 b2 + + | (Mk_prod + | Mk_sort + | Mk_ind + | Mk_const + | Mk_sw + | Mk_fix _ + | Mk_cofix _ + | Mk_rel _ + | Mk_var _ + | Mk_proj + | Is_int + | Is_float + | Is_parray + | Cast_accu + | Upd_cofix + | Force_cofix + | Mk_uint + | Mk_float + | Mk_int + | Val_to_int + | Mk_evar + | MLand + | MLnot + | MLland + | MLmagic + | MLsubst_instance_instance + | MLparray_of_array + | Get_value + | Get_sort + | Get_name + | Get_const + | Get_match + | Get_ind + | Get_evar + | Get_instance + | Get_proj + | Get_symbols + | Lazy + | Coq_primitive _ + | Mk_empty_instance), _ + -> false let primitive_hash = function | Mk_prod -> 1 @@ -347,7 +427,7 @@ let primitive_hash = function | MLmagic -> 21 | Coq_primitive (prim, b) -> combinesmall 22 (combine (CPrimitives.hash prim) (Hashtbl.hash b)) | Mk_proj -> 23 - | MLarrayget -> 24 + | MLsubst_instance_instance -> 24 | Mk_float -> 25 | Is_float -> 26 | Is_parray -> 27 @@ -360,10 +440,11 @@ let primitive_hash = function | Get_match -> 34 | Get_ind -> 35 | Get_evar -> 36 - | Get_level -> 37 + | Get_instance -> 37 | Get_proj -> 38 | Get_symbols -> 39 | Lazy -> 40 + | Mk_empty_instance -> 41 type mllambda = | MLlocal of lname @@ -933,8 +1014,8 @@ let get_evar_code i = MLprimitive (Get_evar, [|MLglobal symbols_tbl_name; MLint i|]) -let get_level_code i = - MLprimitive (Get_level, +let get_instance_code i = + MLprimitive (Get_instance, [|MLglobal symbols_tbl_name; MLint i|]) let get_proj_code i = @@ -982,6 +1063,8 @@ let merge_branches t = let app_prim p args = MLprimitive (p, args) +let ml_empty_instance = MLprimitive (Mk_empty_instance, [||]) + type prim_aux = | PAprim of string * pconstant * CPrimitives.t * prim_aux array | PAml of mllambda @@ -1029,19 +1112,23 @@ let cast_to_int v = | _ -> MLprimitive (Val_to_int, [|v|]) let ml_of_instance instance u = - let ml_of_level l = - match Univ.Level.var_index l with - | Some i -> - (* FIXME: use a proper cast function *) - let univ = MLprimitive (MLmagic, [|MLlocal (Option.get instance)|]) in - MLprimitive (MLarrayget, [|univ; MLint i|]) - | None -> let i = push_symbol (SymbLevel l) in get_level_code i - in - let u = Univ.Instance.to_array u in - if Array.is_empty u then [||] - else let u = Array.map ml_of_level u in - (* FIXME: use a proper cast function *) - [|MLprimitive (MLmagic, [|MLarray u|])|] + if UVars.Instance.is_empty u then [||] + else + let i = push_symbol (SymbInstance u) in + let u_code = get_instance_code i in + let has_variable = + let qs, us = UVars.Instance.to_array u in + Array.exists (fun q -> Option.has_some (Sorts.Quality.var_index q)) qs + || Array.exists (fun u -> Option.has_some (Univ.Level.var_index u)) us + in + let u_code = + if has_variable then + (* if there are variables then [instance] guaranteed non-None *) + let univ = MLprimitive (MLmagic, [|MLlocal (Option.get instance)|]) in + MLprimitive (MLsubst_instance_instance, [|univ; u_code|]) + else u_code + in + [|MLprimitive (MLmagic, [|u_code|])|] let compile_prim env decl cond paux = @@ -1376,7 +1463,7 @@ let compile_prim env decl cond paux = | Lsort s -> let i = push_symbol (SymbSort s) in let uarg = match env.env_univ with - | None -> MLarray [||] + | None -> ml_empty_instance | Some u -> MLlocal u in (* FIXME: use a dedicated cast function *) @@ -1824,6 +1911,7 @@ let pp_mllam fmt l = | Mk_var id -> Format.fprintf fmt "mk_var_accu (Names.Id.of_string \"%s\")" (string_of_id id) | Mk_proj -> Format.fprintf fmt "mk_proj_accu" + | Mk_empty_instance -> Format.fprintf fmt "UVars.Instance.empty" | Is_int -> Format.fprintf fmt "is_int" | Is_float -> Format.fprintf fmt "is_float" | Is_parray -> Format.fprintf fmt "is_parray" @@ -1839,7 +1927,7 @@ let pp_mllam fmt l = | MLnot -> Format.fprintf fmt "not" | MLland -> Format.fprintf fmt "(land)" | MLmagic -> Format.fprintf fmt "Obj.magic" - | MLarrayget -> Format.fprintf fmt "Array.get" + | MLsubst_instance_instance -> Format.fprintf fmt "UVars.subst_instance_instance" | MLparray_of_array -> Format.fprintf fmt "parray_of_array" | Coq_primitive (op, false) -> Format.fprintf fmt "no_check_%s" (CPrimitives.to_string op) @@ -1851,7 +1939,7 @@ let pp_mllam fmt l = | Get_match -> Format.fprintf fmt "get_match" | Get_ind -> Format.fprintf fmt "get_ind" | Get_evar -> Format.fprintf fmt "get_evar" - | Get_level -> Format.fprintf fmt "get_level" + | Get_instance -> Format.fprintf fmt "get_instance" | Get_proj -> Format.fprintf fmt "get_proj" | Get_symbols -> Format.fprintf fmt "get_symbols" | Lazy -> Format.fprintf fmt "lazy" @@ -1873,7 +1961,7 @@ let type_of_global gn c = match gn with | _ -> match c with | MLprimitive (Lazy, _) -> " : Nativevalues.t Lazy.t" | MLlam ([|_|], MLprimitive (Lazy, _)) -> " : Nativevalues.t -> Nativevalues.t Lazy.t" - | MLprimitive ((Mk_ind | Mk_const), [|_|]) -> " : Univ.Level.t array -> Nativevalues.t" + | MLprimitive ((Mk_ind | Mk_const), [|_|]) -> " : UVars.Instance.t -> Nativevalues.t" | MLsetref (_,_) -> " : unit" | _ -> " : Nativevalues.t" @@ -1974,7 +2062,7 @@ and compile_named env sigma univ auxdefs id = Glet(Gnamed id, MLprimitive (Mk_var id, [||]))::auxdefs let compile_constant env sigma con cb = - let no_univs = 0 = Univ.AbstractContext.size (Declareops.constant_polymorphic_context cb) in + let no_univs = (0,0) = UVars.AbstractContext.size (Declareops.constant_polymorphic_context cb) in begin match cb.const_body with | Def t -> let code = lambda_of_constr env sigma t in @@ -1999,7 +2087,7 @@ let compile_constant env sigma con cb = | _ -> let i = push_symbol (SymbConst con) in let args = - if no_univs then [|get_const_code i; MLarray [||]|] + if no_univs then [|get_const_code i; ml_empty_instance|] else [|get_const_code i|] in (* @@ -2035,8 +2123,8 @@ let compile_mind mb mind stack = let name = Gind ("", ind) in let accu = let args = - if Int.equal (Univ.AbstractContext.size u) 0 then - [|get_ind_code j; MLarray [||]|] + if (UVars.AbstractContext.size u) = (0,0) then + [|get_ind_code j; ml_empty_instance|] else [|get_ind_code j|] in (* FIXME: pass universes here *) @@ -2131,7 +2219,7 @@ let compile_deps env sigma prefix init t = let const_updates = Cmap_env.add c upd const_updates in comp_stack, (mind_updates, const_updates) | Construct (((mind,_),_),_u) -> compile_mind_deps env prefix init mind - | Proj (p,c) -> + | Proj (p,_,c) -> let init = compile_mind_deps env prefix init (Projection.mind p) in aux env lvl init c | Case (ci, _u, _pms, _p, _iv, _c, _ac) -> diff --git a/kernel/nativecode.mli b/kernel/nativecode.mli index 53e7e6743e44..94dbcfb2cf7a 100644 --- a/kernel/nativecode.mli +++ b/kernel/nativecode.mli @@ -44,7 +44,7 @@ val get_ind : symbols -> int -> inductive val get_evar : symbols -> int -> Evar.t -val get_level : symbols -> int -> Univ.Level.t +val get_instance : symbols -> int -> UVars.Instance.t val get_proj : symbols -> int -> inductive * int diff --git a/kernel/nativevalues.ml b/kernel/nativevalues.ml index 63cfd07ae315..885a7b22776e 100644 --- a/kernel/nativevalues.ml +++ b/kernel/nativevalues.ml @@ -104,7 +104,7 @@ type symbol = | SymbMatch of annot_sw | SymbInd of inductive | SymbEvar of Evar.t - | SymbLevel of Univ.Level.t + | SymbInstance of UVars.Instance.t | SymbProj of (inductive * int) type block @@ -169,24 +169,13 @@ let mk_rels_accu lvl len = let napply (f:t) (args: t array) = Array.fold_left (fun f a -> apply f a) f args -let mk_constant_accu kn u = - mk_accu (Aconstant (kn,Univ.Instance.of_array u)) +let mk_constant_accu kn u = mk_accu (Aconstant (kn,u)) -let mk_ind_accu ind u = - mk_accu (Aind (ind,Univ.Instance.of_array u)) +let mk_ind_accu ind u = mk_accu (Aind (ind,u)) let mk_sort_accu s u = - let open Sorts in - match s with - | SProp | Prop | Set -> mk_accu (Asort s) - | Type s -> - let u = Univ.Instance.of_array u in - let s = Sorts.sort_of_univ (Univ.subst_instance_universe u s) in - mk_accu (Asort s) - | QSort (q, s) -> - let u = Univ.Instance.of_array u in - let s = Sorts.qsort q (Univ.subst_instance_universe u s) in - mk_accu (Asort s) + let s = UVars.subst_instance_sort u s in + mk_accu (Asort s) let mk_var_accu id = mk_accu (Avar id) diff --git a/kernel/nativevalues.mli b/kernel/nativevalues.mli index ab2b69067051..72af7d0725e3 100644 --- a/kernel/nativevalues.mli +++ b/kernel/nativevalues.mli @@ -72,7 +72,7 @@ type symbol = | SymbMatch of annot_sw | SymbInd of inductive | SymbEvar of Evar.t - | SymbLevel of Univ.Level.t + | SymbInstance of UVars.Instance.t | SymbProj of (inductive * int) type symbols = symbol array @@ -84,9 +84,9 @@ val empty_symbols : symbols val mk_accu : atom -> t val mk_rel_accu : int -> t val mk_rels_accu : int -> int -> t array -val mk_constant_accu : Constant.t -> Univ.Level.t array -> t -val mk_ind_accu : inductive -> Univ.Level.t array -> t -val mk_sort_accu : Sorts.t -> Univ.Level.t array -> t +val mk_constant_accu : Constant.t -> UVars.Instance.t -> t +val mk_ind_accu : inductive -> UVars.Instance.t -> t +val mk_sort_accu : Sorts.t -> UVars.Instance.t -> t val mk_var_accu : Id.t -> t val mk_sw_accu : annot_sw -> accumulator -> t -> (t -> t) val mk_fix_accu : rec_pos -> int -> t array -> t array -> t diff --git a/kernel/relevanceops.ml b/kernel/relevanceops.ml index a8d3f1e89627..435776f00ac8 100644 --- a/kernel/relevanceops.ml +++ b/kernel/relevanceops.ml @@ -9,7 +9,6 @@ (************************************************************************) open Util -open Names open Constr open Declarations open Environ @@ -25,19 +24,27 @@ 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 relevance_of_constant env (c,u) = let decl = lookup_constant c env in - decl.const_relevance + UVars.subst_instance_relevance u decl.const_relevance -let relevance_of_constructor env ((mi,i),_) = +let relevance_of_constructor env (((mi,i),_),u) = let decl = lookup_mind mi env in let packet = decl.mind_packets.(i) in - packet.mind_relevance + UVars.subst_instance_relevance u packet.mind_relevance -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 relevance_of_projection_repr env (p, u) = + let mib = lookup_mind (Names.Projection.Repr.mind p) env in + 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 + UVars.subst_instance_relevance u rs.(Names.Projection.Repr.arg p) + +let relevance_of_projection env (p,u) = + relevance_of_projection_repr env (Names.Projection.repr p,u) let rec relevance_of_term_extra env extra lft c = match kind c with @@ -52,12 +59,12 @@ let rec relevance_of_term_extra env extra lft c = | LetIn ({binder_relevance=r;_}, _, _, bdy) -> relevance_of_term_extra env (Range.cons r extra) (lft + 1) bdy | App (c, _) -> relevance_of_term_extra env extra lft c - | Const (c,_) -> relevance_of_constant env c - | Construct (c,_) -> relevance_of_constructor env 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 + | Proj (_, r, _) -> r | Int _ | Float _ -> Sorts.Relevant | Array _ -> Sorts.Relevant diff --git a/kernel/relevanceops.mli b/kernel/relevanceops.mli index 575bf9b701a2..cdad4dc06d41 100644 --- a/kernel/relevanceops.mli +++ b/kernel/relevanceops.mli @@ -18,6 +18,7 @@ val relevance_of_term : Environ.env -> Constr.constr -> Sorts.relevance open Names val relevance_of_rel : Environ.env -> 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 +val relevance_of_constant : Environ.env -> Constr.pconstant -> Sorts.relevance +val relevance_of_constructor : Environ.env -> Constr.pconstructor -> Sorts.relevance +val relevance_of_projection_repr : Environ.env -> Projection.Repr.t UVars.puniverses -> Sorts.relevance +val relevance_of_projection : Environ.env -> Projection.t UVars.puniverses -> Sorts.relevance diff --git a/kernel/safe_typing.ml b/kernel/safe_typing.ml index a9a7f62ef021..d17eccb3735f 100644 --- a/kernel/safe_typing.ml +++ b/kernel/safe_typing.ml @@ -521,12 +521,11 @@ let push_section_context uctx senv = let sections = get_section senv.sections in let sections = Section.push_local_universe_context uctx sections in let senv = { senv with sections=Some sections } in - let ctx = Univ.ContextSet.of_context uctx in - (* We check that the universes are fresh. FIXME: This should be done - implicitly, but we have to work around the API. *) - let () = assert (Univ.Level.Set.for_all (fun u -> not (Univ.Level.Set.mem u (fst senv.univ))) (fst ctx)) in + let qualities, ctx = UVars.UContext.to_context_set uctx in + assert (Sorts.Quality.Set.is_empty qualities); + (* push_context checks freshness *) { senv with - env = Environ.push_context_set ~strict:false ctx senv.env; + env = Environ.push_context ~strict:false uctx senv.env; univ = Univ.ContextSet.union ctx senv.univ } (** {6 Insertion of new declarations to current environment } *) @@ -606,7 +605,7 @@ type global_declaration = type exported_opaque = { exp_handle : Opaqueproof.opaque_handle; exp_body : Constr.t; - exp_univs : int option; + exp_univs : (int * int) option; (* Minimal amount of data needed to rebuild the private universes. We enforce in the API that private constants have no internal constraints. *) } @@ -739,7 +738,7 @@ let constant_entry_of_side_effect eff = | Monomorphic -> Monomorphic_entry | Polymorphic auctx -> - Polymorphic_entry (Univ.AbstractContext.repr auctx) + Polymorphic_entry (UVars.AbstractContext.repr auctx) in let p = match cb.const_body with @@ -834,7 +833,7 @@ let export_private_constants eff senv = let (_, _, _, h) = Opaqueproof.repr o in let univs = match c.const_universes with | Monomorphic -> None - | Polymorphic auctx -> Some (Univ.AbstractContext.size auctx) + | Polymorphic auctx -> Some (UVars.AbstractContext.size auctx) in let body = Constr.hcons body in let opaque = { exp_body = body; exp_handle = h; exp_univs = univs } in @@ -1292,6 +1291,8 @@ let start_library dir senv = let export ~output_native_objects senv dir = let () = check_current_library dir senv in + (* qualities are in the senv only during sections *) + let () = assert (Sorts.QVar.Set.is_empty senv.env.Environ.env_qualities) in let mp = senv.modpath in let str = NoFunctor (List.rev senv.revstruct) in let mb = diff --git a/kernel/safe_typing.mli b/kernel/safe_typing.mli index dddd72c364ba..b5cbe531de2b 100644 --- a/kernel/safe_typing.mli +++ b/kernel/safe_typing.mli @@ -182,7 +182,7 @@ val push_named_def : Id.t * Entries.section_def_entry -> safe_transformer0 (** Add local universes to a polymorphic section *) -val push_section_context : Univ.UContext.t -> safe_transformer0 +val push_section_context : UVars.UContext.t -> safe_transformer0 (** {6 Interactive module functions } *) diff --git a/kernel/section.ml b/kernel/section.ml index d9c6f9a360fd..926528e27fa8 100644 --- a/kernel/section.ml +++ b/kernel/section.ml @@ -11,6 +11,7 @@ open Util open Names open Univ +open UVars open Cooking module NamedDecl = Context.Named.Declaration @@ -31,7 +32,7 @@ type 'a t = { (** Global universes introduced in the section *) poly_universes : UContext.t; (** Universes local to the section *) - all_poly_univs : Univ.Level.t array; + all_poly_univs : Instance.t; (** All polymorphic universes, including from previous sections. *) has_poly_univs : bool; (** Are there polymorphic universes or constraints, including in previous sections. *) @@ -59,15 +60,12 @@ let push_local_universe_context ctx sec = else let sctx = sec.poly_universes in let poly_universes = UContext.union sctx ctx in - let all_poly_univs = - Array.append sec.all_poly_univs (Instance.to_array @@ UContext.instance ctx) - in + let all_poly_univs = Instance.append sec.all_poly_univs (UContext.instance ctx) in { sec with poly_universes; all_poly_univs; has_poly_univs = true } -let rec is_polymorphic_univ u sec = - let uctx = sec.poly_universes in - let here = Array.exists (fun u' -> Level.equal u u') (Instance.to_array (UContext.instance uctx)) in - here || Option.cata (is_polymorphic_univ u) false sec.prev +let is_polymorphic_univ u sec = + let _, us = Instance.to_array sec.all_poly_univs in + Array.exists (fun u' -> Level.equal u u') us let push_constraints uctx sec = if sec.has_poly_univs && @@ -86,7 +84,7 @@ let open_section ~custom prev = context = []; mono_universes = ContextSet.empty; poly_universes = UContext.empty; - all_poly_univs = Option.cata (fun sec -> sec.all_poly_univs) [| |] prev; + all_poly_univs = Option.cata (fun sec -> sec.all_poly_univs) Instance.empty prev; has_poly_univs = Option.cata has_poly_univs false prev; entries = []; expand_info_map = (Cmap.empty, Mindmap.empty); diff --git a/kernel/section.mli b/kernel/section.mli index 1efe157af849..0ba1d219e7e2 100644 --- a/kernel/section.mli +++ b/kernel/section.mli @@ -10,6 +10,7 @@ open Names open Univ +open UVars open Cooking (** Kernel implementation of sections. *) @@ -58,7 +59,7 @@ val push_global : Environ.env -> poly:bool -> section_entry -> 'a t -> 'a t (** {6 Retrieving section data} *) -val all_poly_univs : 'a t -> Univ.Level.t array +val all_poly_univs : 'a t -> Instance.t (** Returns all polymorphic universes, including those from previous sections. Earlier sections are earlier in the array. diff --git a/kernel/sorts.ml b/kernel/sorts.ml index 01730a83765f..4546cb83d9bd 100644 --- a/kernel/sorts.ml +++ b/kernel/sorts.ml @@ -14,30 +14,245 @@ type family = InSProp | InProp | InSet | InType | InQSort let all_families = [InSProp; InProp; InSet; InType; InQSort] -module QVar = struct - type t = +module QVar = +struct + type repr = + | Var of int | Unif of string * int - let make s i = Unif (s, i) + type t = repr - let repr (Unif (s,i)) = s, i + let make_var n = Var n - let equal (Unif (s1,i1)) (Unif (s2,i2)) = - Int.equal i1 i2 && CString.equal s1 s2 + let make_unif s n = Unif (s,n) - let compare (Unif (s1,i1)) (Unif (s2,i2)) = - let c = Int.compare i1 i2 in + let var_index = function + | Var q -> Some q + | Unif _ -> None + + let hash = function + | Var q -> Hashset.Combine.combinesmall 1 q + | Unif (s,q) -> Hashset.Combine.(combinesmall 2 (combine (CString.hash s) q)) + + module Hstruct = struct + type nonrec t = t + type u = unit + + let hashcons () = function + | Var _ as q -> q + | Unif (s,i) as q -> + let s' = CString.hcons s in + if s == s' then q else Unif (s',i) + + let eq a b = + match a, b with + | Var a, Var b -> Int.equal a b + | Unif (sa, ia), Unif (sb, ib) -> sa == sb && Int.equal ia ib + | (Var _ | Unif _), _ -> false + + let hash = hash + end + + module Hasher = Hashcons.Make(Hstruct) + + let hcons = Hashcons.simple_hcons Hasher.generate Hasher.hcons () + + let compare a b = match a, b with + | Var a, Var b -> Int.compare a b + | Unif (s1,i1), Unif (s2,i2) -> + let c = Int.compare i1 i2 in + if c <> 0 then c + else CString.compare s1 s2 + | Var _, Unif _ -> -1 + | Unif _, Var _ -> 1 + + let equal a b = match a, b with + | Var a, Var b -> Int.equal a b + | Unif (s1,i1), Unif (s2,i2) -> + Int.equal i1 i2 && CString.equal s1 s2 + | Var _, Unif _ | Unif _, Var _ -> false + + let to_string = function + | Var q -> Printf.sprintf "β%d" q + | Unif (s,q) -> + let s = if CString.is_empty s then "" else s^"." in + Printf.sprintf "%sα%d" s q + + let raw_pr q = Pp.str (to_string q) + + let repr x = x + let of_repr x = x + + module Self = struct type nonrec t = t let compare = compare end + module Set = CSet.Make(Self) + module Map = CMap.Make(Self) +end + +module Quality = struct + type constant = QProp | QSProp | QType + type t = QVar of QVar.t | QConstant of constant + + let var i = QVar (QVar.make_var i) + + let var_index = function + | QVar q -> QVar.var_index q + | QConstant _ -> None + + module Constants = struct + let equal a b = match a, b with + | QProp, QProp | QSProp, QSProp | QType, QType -> true + | (QProp | QSProp | QType), _ -> false + + let compare a b = match a, b with + | QProp, QProp -> 0 + | QProp, _ -> -1 + | _, QProp -> 1 + | QSProp, QSProp -> 0 + | QSProp, _ -> -1 + | _, QSProp -> 1 + | QType, QType -> 0 + + let pr = function + | QProp -> Pp.str "Prop" + | QSProp -> Pp.str "SProp" + | QType -> Pp.str "Type" + + let hash = function + | QSProp -> 0 + | QProp -> 1 + | QType -> 2 + + end + + let equal a b = match a, b with + | QVar a, QVar b -> QVar.equal a b + | QConstant a, QConstant b -> Constants.equal a b + | (QVar _ | QConstant _), _ -> false + + let compare a b = match a, b with + | QVar a, QVar b -> QVar.compare a b + | QVar _, _ -> -1 + | _, QVar _ -> 1 + | QConstant a, QConstant b -> Constants.compare a b + + let pr prv = function + | QVar v -> prv v + | QConstant q -> Constants.pr q + + let raw_pr q = pr QVar.raw_pr q + + let hash = let open Hashset.Combine in function + | QConstant q -> Constants.hash q + | QVar q -> combinesmall 3 (QVar.hash q) + + let subst f = function + | QConstant _ as q -> q + | QVar qv as q -> + match f qv with + | QConstant _ as q -> q + | QVar qv' as q' -> + if qv == qv' then q else q' + + let subst_fn m v = + match QVar.Map.find_opt v m with + | Some v -> v + | None -> QVar v + + module Hstruct = struct + type nonrec t = t + type u = QVar.t -> QVar.t + + let hashcons hv = function + | QConstant _ as q -> q + | QVar qv as q -> + let qv' = hv qv in + if qv == qv' then q else QVar qv' + + let eq a b = + match a, b with + | QVar a, QVar b -> a == b + | QVar _, _ -> false + | (QConstant _), _ -> equal a b + + let hash = hash + end + + module Hasher = Hashcons.Make(Hstruct) + + let hcons = Hashcons.simple_hcons Hasher.generate Hasher.hcons QVar.hcons + + let qsprop = hcons (QConstant QSProp) + let qprop = hcons (QConstant QProp) + let qtype = hcons (QConstant QType) + + module Self = struct type nonrec t = t let compare = compare end + module Set = CSet.Make(Self) + module Map = CMap.Make(Self) + +end + +module QConstraint = struct + type kind = Equal | Leq + + let eq_kind : kind -> kind -> bool = (=) + let compare_kind : kind -> kind -> int = compare + + let pr_kind = function + | Equal -> Pp.str "=" + | Leq -> Pp.str "<=" + + type t = Quality.t * kind * Quality.t + + let equal (a,k,b) (a',k',b') = + eq_kind k k' && Quality.equal a a' && Quality.equal b b' + + let compare (a,k,b) (a',k',b') = + let c = compare_kind k k' in if c <> 0 then c - else CString.compare s1 s2 + else + let c = Quality.compare a a' in + if c <> 0 then c + else Quality.compare b b' + + let trivial (a,(Equal|Leq),b) = Quality.equal a b + + let pr prq (a,k,b) = + let open Pp in + hov 1 (Quality.pr prq a ++ spc() ++ pr_kind k ++ spc() ++ Quality.pr prq b) + + let raw_pr x = pr QVar.raw_pr x + +end + +module QConstraints = struct include CSet.Make(QConstraint) + let trivial = for_all QConstraint.trivial + + let pr prq c = + let open Pp in + v 0 (prlist_with_sep spc (fun (u1,op,u2) -> + hov 0 (Quality.pr prq u1 ++ QConstraint.pr_kind op ++ Quality.pr prq u2)) + (elements c)) + +end + +let enforce_eq_quality a b csts = + if Quality.equal a b then csts + else QConstraints.add (a,QConstraint.Equal,b) csts + +let enforce_leq_quality a b csts = + if Quality.equal a b then csts + else match a, b with + | Quality.(QConstant QProp), Quality.(QConstant QType) -> csts + | _ -> QConstraints.add (a,QConstraint.Leq,b) csts - let to_string (Unif (s,i)) = - let i = "α"^string_of_int i in - if CString.is_empty s then i - else s^ "." ^ i +module QUConstraints = struct - let hash (Unif (s,i)) = Hashset.Combine.combine (CString.hash s) i + type t = QConstraints.t * Univ.Constraints.t - let pr x = Pp.str (to_string x) + let empty = QConstraints.empty, Univ.Constraints.empty + + let union (qcsts,ucsts) (qcsts',ucsts') = + QConstraints.union qcsts qcsts', Constraints.union ucsts ucsts' end type t = @@ -102,6 +317,22 @@ let levels s = match s with | Set -> Level.Set.singleton Level.set | Type u | QSort (_, u) -> Universe.levels u +let subst_fn (fq,fu) = function + | SProp | Prop | Set as s -> s + | Type v as s -> + let v' = fu v in + if v' == v then s else sort_of_univ v' + | QSort (q, v) as s -> + let open Quality in + match fq q with + | QVar q' -> + let v' = fu v in + if q' == q && v' == v then s + else qsort q' v' + | QConstant QSProp -> sprop + | QConstant QProp -> prop + | QConstant QType -> sort_of_univ (fu v) + let family = function | SProp -> InSProp | Prop -> InProp @@ -109,6 +340,12 @@ let family = function | Type _ -> InType | QSort _ -> InQSort +let quality = let open Quality in function +| Set | Type _ -> QConstant QType +| Prop -> QConstant QProp +| SProp -> QConstant QSProp +| QSort (q, _) -> QVar q + let family_compare a b = match a,b with | InSProp, InSProp -> 0 | InSProp, _ -> -1 @@ -124,9 +361,18 @@ let family_compare a b = match a,b with | _, InType -> 1 | InQSort, InQSort -> 0 -let family_equal = (==) +let family_equal a b = match a, b with + | InSProp, InSProp | InProp, InProp | InSet, InSet | InType, InType -> true + | InQSort, InQSort -> true + | (InSProp | InProp | InSet | InType | InQSort), _ -> false -let family_leq a b = family_compare a b <= 0 +let family_leq a b = + family_equal a b + || match a, b with + | InSProp, _ -> true + | InProp, InSet -> true + | _, InType -> true + | _ -> false open Hashset.Combine @@ -185,6 +431,16 @@ let relevance_hash = function | Irrelevant -> 1 | RelevanceVar q -> Hashset.Combine.combinesmall 2 (QVar.hash q) +let relevance_subst_fn f = function + | Relevant | Irrelevant as r -> r + | RelevanceVar qv as r -> + let open Quality in + match f qv with + | QConstant QSProp -> Irrelevant + | QConstant (QProp | QType) -> Relevant + | QVar qv' -> + if qv' == qv then r else RelevanceVar qv' + let relevance_of_sort = function | SProp -> Irrelevant | Prop | Set | Type _ -> Relevant @@ -195,7 +451,7 @@ let debug_print = function | Prop -> Pp.(str "Prop") | Set -> Pp.(str "Set") | Type u -> Pp.(str "Type(" ++ Univ.Universe.raw_pr u ++ str ")") - | QSort (q, u) -> Pp.(str "QSort(" ++ QVar.pr q ++ str "," + | QSort (q, u) -> Pp.(str "QSort(" ++ QVar.raw_pr q ++ str "," ++ spc() ++ Univ.Universe.raw_pr u ++ str ")") let pr_sort_family = function @@ -204,8 +460,3 @@ let pr_sort_family = function | InSet -> Pp.(str "Set") | InType -> Pp.(str "Type") | InQSort -> Pp.(str "Type") (* FIXME? *) - -let subst_instance_sort u = function - | SProp | Prop | Set as s -> s - | Type v -> sort_of_univ (subst_instance_universe u v) - | QSort (q, v) -> QSort (q, subst_instance_universe u v) diff --git a/kernel/sorts.mli b/kernel/sorts.mli index e2fa91b8d1d8..e370696af90c 100644 --- a/kernel/sorts.mli +++ b/kernel/sorts.mli @@ -17,12 +17,113 @@ val all_families : family list module QVar : sig type t - val make : string -> int -> t - val repr : t -> string * int + + val var_index : t -> int option + + val make_var : int -> t + val make_unif : string -> int -> t + val equal : t -> t -> bool val compare : t -> t -> int + + val hash : t -> int + + val raw_pr : t -> Pp.t + (** Using this is incorrect when names are available, typically from an evar map. *) + val to_string : t -> string - val pr : t -> Pp.t + (** Debug printing *) + + type repr = + | Var of int + | Unif of string * int + + val repr : t -> repr + val of_repr : repr -> t + + module Set : CSig.SetS with type elt = t + + module Map : CMap.ExtS with type key = t and module Set := Set +end + +module Quality : sig + type constant = QProp | QSProp | QType + type t = QVar of QVar.t | QConstant of constant + + module Constants : sig + val equal : constant -> constant -> bool + val compare : constant -> constant -> int + val pr : constant -> Pp.t + end + + val qprop : t + val qsprop : t + val qtype : t + + val var : int -> t + (** [var i] is [QVar (QVar.make_var i)] *) + + val var_index : t -> int option + + val equal : t -> t -> bool + + val compare : t -> t -> int + + val pr : (QVar.t -> Pp.t) -> t -> Pp.t + + val raw_pr : t -> Pp.t + + val hash : t -> int + + val hcons : t -> t + + (* XXX Inconsistent naming: this one should be subst_fn *) + val subst : (QVar.t -> t) -> t -> t + + val subst_fn : t QVar.Map.t -> QVar.t -> t + + module Set : CSig.SetS with type elt = t + + module Map : CMap.ExtS with type key = t and module Set := Set + +end + +module QConstraint : sig + type kind = Equal | Leq + + val pr_kind : kind -> Pp.t + + type t = Quality.t * kind * Quality.t + + val equal : t -> t -> bool + + val compare : t -> t -> int + + val trivial : t -> bool + + val pr : (QVar.t -> Pp.t) -> t -> Pp.t + + val raw_pr : t -> Pp.t +end + +module QConstraints : sig include CSig.SetS with type elt = QConstraint.t + + val trivial : t -> bool + + val pr : (QVar.t -> Pp.t) -> t -> Pp.t +end + +val enforce_eq_quality : Quality.t -> Quality.t -> QConstraints.t -> QConstraints.t + +val enforce_leq_quality : Quality.t -> Quality.t -> QConstraints.t -> QConstraints.t + +module QUConstraints : sig + + type t = QConstraints.t * Univ.Constraints.t + + val union : t -> t -> t + + val empty : t end type t = private @@ -47,6 +148,7 @@ val is_set : t -> bool val is_prop : t -> bool val is_small : t -> bool val family : t -> family +val quality : t -> Quality.t val hcons : t -> t @@ -60,15 +162,19 @@ val levels : t -> Univ.Level.Set.t val super : t -> t -val subst_instance_sort : Univ.Instance.t -> t -> t +val subst_fn : (QVar.t -> Quality.t) * (Univ.Universe.t -> Univ.Universe.t) + -> t -> t (** On binders: is this variable proof relevant *) +(* TODO put in submodule or new file *) type relevance = Relevant | Irrelevant | RelevanceVar of QVar.t val relevance_hash : relevance -> int val relevance_equal : relevance -> relevance -> bool +val relevance_subst_fn : (QVar.t -> Quality.t) -> relevance -> relevance + val relevance_of_sort : t -> relevance val relevance_of_sort_family : family -> relevance diff --git a/kernel/subtyping.ml b/kernel/subtyping.ml index 94bb82926c67..7d2762f1902e 100644 --- a/kernel/subtyping.ml +++ b/kernel/subtyping.ml @@ -15,7 +15,7 @@ (*i*) open Names -open Univ +open UVars open Util open Constr open Declarations @@ -101,7 +101,7 @@ let check_universes error env u1 u2 = if not (UGraph.check_subtype (Environ.universes env) auctx2 auctx1) then error (IncompatibleConstraints { got = auctx1; expect = auctx2; } ) else - Environ.push_context ~strict:false (Univ.AbstractContext.repr auctx2) env + Environ.push_context ~strict:false (UVars.AbstractContext.repr auctx2) env | Monomorphic, Polymorphic _ -> error (PolymorphicStatusExpected true) | Polymorphic _, Monomorphic -> error (PolymorphicStatusExpected false) diff --git a/kernel/type_errors.ml b/kernel/type_errors.ml index 60b387b58591..93985c5a2d6f 100644 --- a/kernel/type_errors.ml +++ b/kernel/type_errors.ml @@ -12,6 +12,7 @@ open Names open Constr open Environ open Univ +open UVars (* Type errors. *) @@ -70,7 +71,9 @@ type ('constr, 'types) ptype_error = | IllFormedRecBody of 'constr pguard_error * Name.t Context.binder_annot array * int * env * ('constr, 'types) punsafe_judgment array | IllTypedRecBody of int * Name.t Context.binder_annot array * ('constr, 'types) punsafe_judgment array * 'types array + | UnsatisfiedQConstraints of Sorts.QConstraints.t | UnsatisfiedConstraints of Constraints.t + | UndeclaredQualities of Sorts.QVar.Set.t | UndeclaredUniverse of Level.t | DisallowedSProp | BadBinderRelevance of Sorts.relevance * ('constr, 'types) Context.Rel.Declaration.pt @@ -149,9 +152,15 @@ let error_ill_formed_rec_body env why lna i fixenv vdefj = let error_ill_typed_rec_body env i lna vdefj vargs = raise (TypeError (env, IllTypedRecBody (i,lna,vdefj,vargs))) +let error_unsatisfied_qconstraints env c = + raise (TypeError (env, UnsatisfiedQConstraints c)) + let error_unsatisfied_constraints env c = raise (TypeError (env, UnsatisfiedConstraints c)) +let error_undeclared_qualities env l = + raise (TypeError (env, UndeclaredQualities l)) + let error_undeclared_universe env l = raise (TypeError (env, UndeclaredUniverse l)) @@ -199,7 +208,8 @@ let map_pguard_error f = function let map_ptype_error f = function | UnboundRel _ | UnboundVar _ | CaseOnPrivateInd _ | IllFormedCaseParams -| UndeclaredUniverse _ | DisallowedSProp | UnsatisfiedConstraints _ +| UndeclaredQualities _ | UndeclaredUniverse _ | DisallowedSProp +| UnsatisfiedQConstraints _ | UnsatisfiedConstraints _ | ReferenceVariables _ | BadInvert | BadVariance _ | UndeclaredUsedVariables _ as e -> e | NotAType j -> NotAType (on_judgment f j) | BadAssumption j -> BadAssumption (on_judgment f j) diff --git a/kernel/type_errors.mli b/kernel/type_errors.mli index 4e04874983b3..9e6dabe455c8 100644 --- a/kernel/type_errors.mli +++ b/kernel/type_errors.mli @@ -12,6 +12,7 @@ open Names open Constr open Environ open Univ +open UVars (** Type errors. {% \label{typeerrors} %} *) @@ -72,7 +73,9 @@ type ('constr, 'types) ptype_error = | IllFormedRecBody of 'constr pguard_error * Name.t Context.binder_annot array * int * env * ('constr, 'types) punsafe_judgment array | IllTypedRecBody of int * Name.t Context.binder_annot array * ('constr, 'types) punsafe_judgment array * 'types array + | UnsatisfiedQConstraints of Sorts.QConstraints.t | UnsatisfiedConstraints of Constraints.t + | UndeclaredQualities of Sorts.QVar.Set.t | UndeclaredUniverse of Level.t | DisallowedSProp | BadBinderRelevance of Sorts.relevance * ('constr, 'types) Context.Rel.Declaration.pt @@ -146,8 +149,12 @@ val error_ill_formed_rec_body : val error_ill_typed_rec_body : env -> int -> Name.t Context.binder_annot array -> unsafe_judgment array -> types array -> 'a +val error_unsatisfied_qconstraints : env -> Sorts.QConstraints.t -> 'a + val error_unsatisfied_constraints : env -> Constraints.t -> 'a +val error_undeclared_qualities : env -> Sorts.QVar.Set.t -> 'a + val error_undeclared_universe : env -> Level.t -> 'a val error_disallowed_sprop : env -> 'a diff --git a/kernel/typeops.ml b/kernel/typeops.ml index eaa9ad182c54..6c460fc6c3bb 100644 --- a/kernel/typeops.ml +++ b/kernel/typeops.ml @@ -42,6 +42,10 @@ let check_constraints cst env = if Environ.check_constraints cst env then () else error_unsatisfied_constraints env cst +let check_qconstraints qcst env = + if Sorts.QConstraints.trivial qcst then () + else error_unsatisfied_qconstraints env qcst + (* This should be a type (a priori without intention to be an assumption) *) let check_type env c t = match kind(Reduction.whd_all env t) with @@ -91,15 +95,8 @@ let warn_bad_relevance_binder ?loc env rlv bnd = | CWarnings.AsError -> error_bad_binder_relevance env rlv bnd -let anomaly_sort_variable q = - anomaly Pp.(str "The kernel received a sort variable " ++ Sorts.QVar.pr q) - let check_assumption env x t ty = let r = x.binder_relevance in - let () = match r with - | Sorts.Relevant | Sorts.Irrelevant -> () - | Sorts.RelevanceVar q -> anomaly_sort_variable q - in let r' = infer_assumption env t ty in let x = if Sorts.relevance_equal r r' then x @@ -120,17 +117,22 @@ exception ArgumentsMismatch let instantiate_context u subst nas ctx = let open Context.Rel.Declaration in + let instantiate_relevance na = + { na with binder_relevance = UVars.subst_instance_relevance u na.binder_relevance } + in let rec instantiate i ctx = match ctx with | [] -> if 0 <= i then raise ArgumentsMismatch else [] | LocalAssum (na, ty) :: ctx -> let ctx = instantiate (pred i) ctx in let subst = Esubst.subs_liftn i subst in + let na = instantiate_relevance na in let ty = esubst u subst ty in let () = check_binding_relevance na nas.(i) in LocalAssum (nas.(i), ty) :: ctx | LocalDef (na, ty, bdy) :: ctx -> let ctx = instantiate (pred i) ctx in let subst = Esubst.subs_liftn i subst in + let na = instantiate_relevance na in let ty = esubst u subst ty in let bdy = esubst u subst bdy in let () = check_binding_relevance na nas.(i) in @@ -299,14 +301,14 @@ let type_of_parameters env ctx u argsv argstv = (* Type of primitive constructs *) let type_of_prim_type _env u (type a) (prim : a CPrimitives.prim_type) = match prim with | CPrimitives.PT_int63 -> - assert (Univ.Instance.is_empty u); + assert (UVars.Instance.is_empty u); Constr.mkSet | CPrimitives.PT_float64 -> - assert (Univ.Instance.is_empty u); + assert (UVars.Instance.is_empty u); Constr.mkSet | CPrimitives.PT_array -> - begin match Univ.Instance.to_array u with - | [|u|] -> + begin match UVars.Instance.to_array u with + | [||], [|u|] -> let ty = Constr.mkType (Univ.Universe.make u) in Constr.mkProd(Context.anonR, ty , ty) | _ -> anomaly Pp.(str"universe instance for array type should have length 1") @@ -323,7 +325,7 @@ let type_of_float env = | None -> CErrors.user_err Pp.(str"The type float must be registered before this construction can be typechecked.") let type_of_array env u = - assert (Univ.Instance.length u = 1); + assert (UVars.Instance.length u = (0,1)); match env.retroknowledge.Retroknowledge.retro_array with | Some c -> mkConstU (c,u) | None -> CErrors.user_err Pp.(str"The type array must be registered before this construction can be typechecked.") @@ -512,11 +514,12 @@ let type_case_scrutinee env (mib, _mip) (u', largs) u pms (pctx, p) c = in (* We use l2r:true for compat with old versions which used CONV with arguments flipped. It is relevant for performance eg in bedrock / Kami. *) - let cst = match mib.mind_variance with - | None -> Univ.enforce_eq_instances u u' Univ.Constraints.empty - | Some variance -> Univ.enforce_leq_variance_instances variance u' u Univ.Constraints.empty + let qcst, ucst = match mib.mind_variance with + | None -> UVars.enforce_eq_instances u u' Sorts.QUConstraints.empty + | Some variance -> UVars.enforce_leq_variance_instances variance u' u Sorts.QUConstraints.empty in - let () = check_constraints cst env in + let () = check_qconstraints qcst env in + let () = check_constraints ucst env in let subst = Vars.subst_of_rel_context_instance_list pctx (realargs @ [c]) in Vars.substl subst p @@ -532,10 +535,6 @@ let type_of_case env (mib, mip as specif) ci u pms (pctx, pnas, p, pt) iv c ct l error_elim_arity env (ind, u') c None in let rp = Sorts.relevance_of_sort sp in - let () = match ci.ci_relevance with - | Sorts.Relevant | Sorts.Irrelevant -> () - | Sorts.RelevanceVar q -> anomaly_sort_variable q - in let ci = if Sorts.relevance_equal ci.ci_relevance rp then ci else @@ -567,14 +566,15 @@ let type_of_case env (mib, mip as specif) ci u pms (pctx, pnas, p, pt) iv c ct l ci, rslty let type_of_projection env p c ct = - let pty = lookup_projection p env in + let pr, pty = lookup_projection p env in let (ind,u), args = try find_rectype env ct with Not_found -> error_case_not_inductive env (make_judge c ct) in assert(Ind.CanOrd.equal (Projection.inductive p) ind); + let pr = UVars.subst_instance_relevance u pr in let ty = Vars.subst_instance_constr u pty in - substl (c :: CList.rev args) ty + pr, substl (c :: CList.rev args) ty (* Fixpoints. *) @@ -595,7 +595,7 @@ let check_fixpoint env lna lar vdef vdeft = let type_of_global_in_context env r = let open Names.GlobRef in match r with - | VarRef id -> Environ.named_type id env, Univ.AbstractContext.empty + | VarRef id -> Environ.named_type id env, UVars.AbstractContext.empty | ConstRef c -> let cb = Environ.lookup_constant c env in let univs = Declareops.constant_polymorphic_context cb in @@ -603,14 +603,14 @@ let type_of_global_in_context env r = | IndRef ind -> let (mib,_ as specif) = Inductive.lookup_mind_specif env ind in let univs = Declareops.inductive_polymorphic_context mib in - let inst = Univ.make_abstract_instance univs in + let inst = UVars.make_abstract_instance univs in Inductive.type_of_inductive (specif, inst), univs | ConstructRef cstr -> let (mib,_ as specif) = Inductive.lookup_mind_specif env (inductive_of_constructor cstr) in let univs = Declareops.inductive_polymorphic_context mib in - let inst = Univ.make_abstract_instance univs in + let inst = UVars.make_abstract_instance univs in Inductive.type_of_constructor (cstr,inst) specif, univs (************************************************************************) @@ -618,10 +618,6 @@ let type_of_global_in_context env r = let check_assum_annot env s x t = let r = x.binder_relevance in - let () = match r with - | Sorts.Relevant | Sorts.Irrelevant -> () - | Sorts.RelevanceVar q -> anomaly_sort_variable q - in let r' = Sorts.relevance_of_sort s in if Sorts.relevance_equal r' r then x @@ -631,10 +627,6 @@ let check_assum_annot env s x t = let check_let_annot env s x c t = let r = x.binder_relevance in - let () = match r with - | Sorts.Relevant | Sorts.Irrelevant -> () - | Sorts.RelevanceVar q -> anomaly_sort_variable q - in let r' = Sorts.relevance_of_sort s in if Sorts.relevance_equal r' r then x @@ -653,8 +645,7 @@ let rec execute env cstr = | Sort s -> let () = match s with | SProp -> if not (Environ.sprop_allowed env) then error_disallowed_sprop env - | QSort (q, _) -> anomaly_sort_variable q - | Prop | Set | Type _ -> () + | QSort _ | Prop | Set | Type _ -> () in cstr, type_of_sort s @@ -667,10 +658,12 @@ let rec execute env cstr = | Const c -> cstr, type_of_constant env c - | Proj (p, c) -> + | Proj (p, r, c) -> let c', ct = execute env c in - let cstr = if c == c' then cstr else mkProj (p,c') in - cstr, type_of_projection env p c' ct + let r', ty = type_of_projection env p c' ct in + assert (Sorts.relevance_equal r r'); + let cstr = if c == c' then cstr else mkProj (p,r,c') in + cstr, ty (* Lambda calculus operators *) | App (f,args) -> @@ -754,7 +747,7 @@ let rec execute env cstr = let realdecls, _ = List.chop mip.mind_nrealdecls mip.mind_arity_ctxt in let self = let args = Context.Rel.instance mkRel 0 mip.mind_arity_ctxt in - let inst = Instance.of_array (Array.init (Instance.length u) Level.var) in + let inst = UVars.Instance.(abstract_instance (length u)) in mkApp (mkIndU (ci.ci_ind, inst), args) in let realdecls = LocalAssum (Context.make_annot Anonymous mip.mind_relevance, self) :: realdecls in @@ -814,8 +807,8 @@ let rec execute env cstr = | Float _ -> cstr, type_of_float env | Array(u,t,def,ty) -> (* ty : Type@{u} and all of t,def : ty *) - let ulev = match Univ.Instance.to_array u with - | [|u|] -> u + let ulev = match UVars.Instance.to_array u with + | [||], [|u|] -> u | _ -> assert false in let ty',tyty = execute env ty in @@ -860,8 +853,15 @@ let execute env c = (* Derived functions *) +let check_declared_qualities env qualities = + let module S = Sorts.QVar.Set in + let unknown = S.diff qualities env.env_qualities in + if S.is_empty unknown then () + else error_undeclared_qualities env unknown + let check_wellformed_universes env c = - let univs = universes_of_constr c in + let qualities, univs = sort_and_universes_of_constr c in + check_declared_qualities env qualities; try UGraph.check_declared_universes (universes env) univs with UGraph.UndeclaredLevel u -> error_undeclared_universe env u @@ -916,7 +916,8 @@ let judge_of_variable env x = make_judge (mkVar x) (type_of_variable env x) let judge_of_constant env cst = make_judge (mkConstU cst) (type_of_constant env cst) let judge_of_projection env p cj = - make_judge (mkProj (p,cj.uj_val)) (type_of_projection env p cj.uj_val cj.uj_type) + let r, ty = type_of_projection env p cj.uj_val cj.uj_type in + make_judge (mkProj (p,r,cj.uj_val)) ty let dest_judgev v = Array.map j_val v, Array.map j_type v @@ -1017,7 +1018,7 @@ let type_of_prim env u t = tr_type n arg_ty, nary_op (n + 1) ret_ty r) in let params, args_ty, ret_ty = types t in - assert (AbstractContext.size (univs t) = Instance.length u); + assert (UVars.AbstractContext.size (univs t) = UVars.Instance.length u); Vars.subst_instance_constr u (Term.it_mkProd_or_LetIn (nary_op 0 ret_ty args_ty) params) diff --git a/kernel/typeops.mli b/kernel/typeops.mli index bbed6ad41158..de7fb4a37677 100644 --- a/kernel/typeops.mli +++ b/kernel/typeops.mli @@ -11,6 +11,7 @@ open Names open Constr open Univ +open UVars open Environ (** {6 Typing functions (not yet tagged as safe) } @@ -92,7 +93,7 @@ val judge_of_constructor : env -> constructor puniverses -> unsafe_judgment (** {6 Type of global references. } *) -val type_of_global_in_context : env -> GlobRef.t -> types * Univ.AbstractContext.t +val type_of_global_in_context : env -> GlobRef.t -> types * UVars.AbstractContext.t (** Returns the type of the global reference, by creating a fresh instance of polymorphic references and computing their instantiated universe context. The type should not be used @@ -114,12 +115,12 @@ val judge_of_int : env -> Uint63.t -> unsafe_judgment val type_of_float : env -> types val judge_of_float : env -> Float64.t -> unsafe_judgment -val type_of_array : env -> Univ.Instance.t -> types -val judge_of_array : env -> Univ.Instance.t -> unsafe_judgment array -> unsafe_judgment -> unsafe_judgment +val type_of_array : env -> UVars.Instance.t -> types +val judge_of_array : env -> UVars.Instance.t -> unsafe_judgment array -> unsafe_judgment -> unsafe_judgment -val type_of_prim_type : env -> Univ.Instance.t -> 'a CPrimitives.prim_type -> types -val type_of_prim : env -> Univ.Instance.t -> CPrimitives.t -> types -val type_of_prim_or_type : env -> Univ.Instance.t -> CPrimitives.op_or_type -> types +val type_of_prim_type : env -> UVars.Instance.t -> 'a CPrimitives.prim_type -> types +val type_of_prim : env -> UVars.Instance.t -> CPrimitives.t -> types +val type_of_prim_or_type : env -> UVars.Instance.t -> CPrimitives.op_or_type -> types val warn_bad_relevance_name : string (** Allow the checker to make this warning into an error. *) diff --git a/kernel/uGraph.ml b/kernel/uGraph.ml index d12ca3359b0b..1f7909080f1d 100644 --- a/kernel/uGraph.ml +++ b/kernel/uGraph.ml @@ -9,6 +9,7 @@ (************************************************************************) open Univ +open UVars module G = AcyclicGraph.Make(struct type t = Level.t @@ -183,13 +184,14 @@ let constraints_for ~kept g = (** Subtyping of polymorphic contexts *) let check_subtype univs ctxT ctx = - if AbstractContext.size ctxT == AbstractContext.size ctx then + (* NB: size check is the only constraint on qualities *) + if eq_sizes (AbstractContext.size ctxT) (AbstractContext.size ctx) then let uctx = AbstractContext.repr ctx in let inst = UContext.instance uctx in let cst = UContext.constraints uctx in let cstT = UContext.constraints (AbstractContext.repr ctxT) in let push accu v = add_universe v ~lbound:Bound.Set ~strict:false accu in - let univs = Array.fold_left push univs (Instance.to_array inst) in + let univs = Array.fold_left push univs (snd (Instance.to_array inst)) in let univs = merge_constraints cstT univs in check_constraints cst univs else false @@ -197,13 +199,10 @@ let check_subtype univs ctxT ctx = (** Instances *) let check_eq_instances g t1 t2 = - let t1 = Instance.to_array t1 in - let t2 = Instance.to_array t2 in - t1 == t2 || - (Int.equal (Array.length t1) (Array.length t2) && - let rec aux i = - (Int.equal i (Array.length t1)) || (check_eq_level g t1.(i) t2.(i) && aux (i + 1)) - in aux 0) + let qt1, ut1 = Instance.to_array t1 in + let qt2, ut2 = Instance.to_array t2 in + CArray.equal Sorts.Quality.equal qt1 qt2 + && CArray.equal (check_eq_level g) ut1 ut2 let domain g = G.domain g.graph let choose p g u = G.choose p g.graph u @@ -276,13 +275,13 @@ let pr_universes prl g = pr_pmap Pp.mt (pr_arc prl) g open Pp -let explain_universe_inconsistency prl (o,u,v,p : univ_inconsistency) = +let explain_universe_inconsistency prq prl (o,u,v,p : univ_inconsistency) = let pr_uni u = match u with | Sorts.Set -> str "Set" | Sorts.Prop -> str "Prop" | Sorts.SProp -> str "SProp" | Sorts.Type u -> Universe.pr prl u - | Sorts.QSort (_q, u) -> Universe.pr prl u (* FIXME? *) + | Sorts.QSort (q, u) -> str "Type@{" ++ prq q ++ str " | " ++ Universe.pr prl u ++ str"}" in let pr_rel = function | Eq -> str"=" | Lt -> str"<" | Le -> str"<=" diff --git a/kernel/uGraph.mli b/kernel/uGraph.mli index c1ee6bb79579..59191661b28b 100644 --- a/kernel/uGraph.mli +++ b/kernel/uGraph.mli @@ -9,6 +9,7 @@ (************************************************************************) open Univ +open UVars (** {6 Graphs of universes. } *) type t @@ -120,7 +121,7 @@ val repr : t -> node Level.Map.t val pr_universes : (Level.t -> Pp.t) -> node Level.Map.t -> Pp.t -val explain_universe_inconsistency : (Level.t -> Pp.t) -> +val explain_universe_inconsistency : (Sorts.QVar.t -> Pp.t) -> (Level.t -> Pp.t) -> univ_inconsistency -> Pp.t (** {6 Debugging} *) diff --git a/kernel/uVars.ml b/kernel/uVars.ml new file mode 100644 index 000000000000..1b03463ba8e9 --- /dev/null +++ b/kernel/uVars.ml @@ -0,0 +1,453 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * Copyright INRIA, CNRS and contributors *) +(* s + | Invariant, _ | _, Invariant -> Invariant + | Covariant, Covariant -> Covariant + + let equal a b = match a,b with + | Irrelevant, Irrelevant | Covariant, Covariant | Invariant, Invariant -> true + | (Irrelevant | Covariant | Invariant), _ -> false + + let check_subtype x y = match x, y with + | (Irrelevant | Covariant | Invariant), Irrelevant -> true + | Irrelevant, Covariant -> false + | (Covariant | Invariant), Covariant -> true + | (Irrelevant | Covariant), Invariant -> false + | Invariant, Invariant -> true + + let pr = function + | Irrelevant -> str "*" + | Covariant -> str "+" + | Invariant -> str "=" + + let leq_constraint csts variance u u' = + match variance with + | Irrelevant -> csts + | Covariant -> enforce_leq_level u u' csts + | Invariant -> enforce_eq_level u u' csts + + let eq_constraint csts variance u u' = + match variance with + | Irrelevant -> csts + | Covariant | Invariant -> enforce_eq_level u u' csts + + let leq_constraints variance u u' csts = + let len = Array.length u in + assert (len = Array.length u' && len = Array.length variance); + Array.fold_left3 leq_constraint csts variance u u' + + let eq_constraints variance u u' csts = + let len = Array.length u in + assert (len = Array.length u' && len = Array.length variance); + Array.fold_left3 eq_constraint csts variance u u' +end + +module Instance : sig + type t + + val empty : t + val is_empty : t -> bool + + val of_array : Quality.t array * Level.t array -> t + val to_array : t -> Quality.t array * Level.t array + + val abstract_instance : int * int -> t + + val append : t -> t -> t + val equal : t -> t -> bool + val length : t -> int * int + + val hcons : t -> t + val hash : t -> int + + val share : t -> t * int + + val subst_fn : (Sorts.QVar.t -> Quality.t) * (Level.t -> Level.t) -> t -> t + + val pr : (Sorts.QVar.t -> Pp.t) -> (Level.t -> Pp.t) -> ?variance:Variance.t array -> t -> Pp.t + val levels : t -> Quality.Set.t * Level.Set.t +end = +struct + type t = Quality.t array * Level.t array + + let empty : t = [||], [||] + + module HInstancestruct = + struct + type nonrec t = t + type u = (Quality.t -> Quality.t) * (Level.t -> Level.t) + + let hashcons (hqual, huniv) (aq, au as a) = + let qlen = Array.length aq in + let ulen = Array.length au in + if Int.equal qlen 0 && Int.equal ulen 0 then empty + else begin + for i = 0 to qlen - 1 do + let x = Array.unsafe_get aq i in + let x' = hqual x in + if x == x' then () + else Array.unsafe_set aq i x' + done; + for i = 0 to ulen - 1 do + let x = Array.unsafe_get au i in + let x' = huniv x in + if x == x' then () + else Array.unsafe_set au i x' + done; + a + end + + let eq t1 t2 = + CArray.equal (==) (fst t1) (fst t2) + && CArray.equal (==) (snd t1) (snd t2) + + let hash (aq,au) = + let accu = ref 0 in + for i = 0 to Array.length aq - 1 do + let l = Array.unsafe_get aq i in + let h = Quality.hash l in + accu := Hashset.Combine.combine !accu h; + done; + for i = 0 to Array.length au - 1 do + let l = Array.unsafe_get au i in + let h = Level.hash l in + accu := Hashset.Combine.combine !accu h; + done; + (* [h] must be positive. *) + let h = !accu land 0x3FFFFFFF in + h + end + + module HInstance = Hashcons.Make(HInstancestruct) + + let hcons = Hashcons.simple_hcons HInstance.generate HInstance.hcons (Quality.hcons,Level.hcons) + + let hash = HInstancestruct.hash + + let share a = (hcons a, hash a) + + let empty = hcons empty + + let is_empty (x,y) = CArray.is_empty x && CArray.is_empty y + + + let append (xq,xu as x) (yq,yu as y) = + if is_empty x then y + else if is_empty y then x + else Array.append xq yq, Array.append xu yu + + let of_array a : t = a + + let to_array (a:t) = a + + let abstract_instance (qlen,ulen) = + let qs = Array.init qlen Quality.var in + let us = Array.init ulen Level.var in + of_array (qs,us) + + let length (aq,au) = Array.length aq, Array.length au + + let subst_fn (fq, fn) (q,u as orig) : t = + let q' = CArray.Smart.map (Quality.subst fq) q in + let u' = CArray.Smart.map fn u in + if q' == q && u' == u then orig else q', u' + + let levels (xq,xu) = + let q = Array.fold_left (fun acc x -> Quality.Set.add x acc) Quality.Set.empty xq in + let u = Array.fold_left (fun acc x -> Level.Set.add x acc) Level.Set.empty xu in + q, u + + let pr prq prl ?variance (q,u) = + let ppu i u = + let v = Option.map (fun v -> v.(i)) variance in + pr_opt_no_spc Variance.pr v ++ prl u + in + (if Array.is_empty q then mt() else prvect_with_sep spc (Quality.pr prq) q ++ strbrk " | ") + ++ prvecti_with_sep spc ppu u + + let equal (xq,xu) (yq,yu) = + CArray.equal Quality.equal xq yq + && CArray.equal Level.equal xu yu + +end + +let eq_sizes (a,b) (a',b') = Int.equal a a' && Int.equal b b' + +type 'a quconstraint_function = 'a -> 'a -> Sorts.QUConstraints.t -> Sorts.QUConstraints.t + +let enforce_eq_instances x y (qcs, ucs as orig) = + let xq, xu = Instance.to_array x and yq, yu = Instance.to_array y in + if Array.length xq != Array.length yq || Array.length xu != Array.length yu then + CErrors.anomaly (Pp.(++) (Pp.str "Invalid argument: enforce_eq_instances called with") + (Pp.str " instances of different lengths.")); + let qcs' = CArray.fold_right2 Sorts.enforce_eq_quality xq yq qcs in + let ucs' = CArray.fold_right2 enforce_eq_level xu yu ucs in + if qcs' == qcs && ucs' == ucs then orig else qcs', ucs' + +let enforce_eq_variance_instances variances x y (qcs,ucs as orig) = + let xq, xu = Instance.to_array x and yq, yu = Instance.to_array y in + let qcs' = CArray.fold_right2 Sorts.enforce_eq_quality xq yq qcs in + let ucs' = Variance.eq_constraints variances xu yu ucs in + if qcs' == qcs && ucs' == ucs then orig else qcs', ucs' + +let enforce_leq_variance_instances variances x y (qcs,ucs as orig) = + let xq, xu = Instance.to_array x and yq, yu = Instance.to_array y in + (* no variance for quality variables -> enforce_eq *) + let qcs' = CArray.fold_right2 Sorts.enforce_eq_quality xq yq qcs in + let ucs' = Variance.leq_constraints variances xu yu ucs in + if qcs' == qcs && ucs' == ucs then orig else qcs', ucs' + +let subst_instance_level s l = + match Level.var_index l with + | Some n -> (snd (Instance.to_array s)).(n) + | None -> l + +let subst_instance_qvar s v = + match Sorts.QVar.var_index v with + | Some n -> (fst (Instance.to_array s)).(n) + | None -> Quality.QVar v + +let subst_instance_quality s l = + match l with + | Quality.QVar v -> begin match Sorts.QVar.var_index v with + | Some n -> (fst (Instance.to_array s)).(n) + | None -> l + end + | _ -> l + +let subst_instance_instance s i = + let qs, us = Instance.to_array i in + let qs' = Array.Smart.map (fun l -> subst_instance_quality s l) qs in + let us' = Array.Smart.map (fun l -> subst_instance_level s l) us in + if qs' == qs && us' == us then i else Instance.of_array (qs', us') + +let subst_instance_universe s univ = + let f (v,n as vn) = + let v' = subst_instance_level s v in + if v == v' then vn + else v', n + in + let u = Universe.repr univ in + let u' = List.Smart.map f u in + if u == u' then univ + else Universe.unrepr u' + +let subst_instance_sort u s = + Sorts.subst_fn ((subst_instance_qvar u), (subst_instance_universe u)) s + +let subst_instance_relevance u r = + Sorts.relevance_subst_fn (subst_instance_qvar u) r + +let subst_instance_constraint s (u,d,v as c) = + let u' = subst_instance_level s u in + let v' = subst_instance_level s v in + if u' == u && v' == v then c + else (u',d,v') + +let subst_instance_constraints s csts = + Constraints.fold + (fun c csts -> Constraints.add (subst_instance_constraint s c) csts) + csts Constraints.empty + +type 'a puniverses = 'a * Instance.t +let out_punivs (x, _y) = x +let in_punivs x = (x, Instance.empty) +let eq_puniverses f (x, u) (y, u') = + f x y && Instance.equal u u' + +type bound_names = Names.Name.t array * Names.Name.t array + +(** A context of universe levels with universe constraints, + representing local universe variables and constraints *) + +module UContext = +struct + type t = bound_names * Instance.t constrained + + let make names (univs, _ as x) : t = + let qs, us = Instance.to_array univs in + assert (Array.length (fst names) = Array.length qs && Array.length(snd names) = Array.length us); + (names, x) + + (** Universe contexts (variables as a list) *) + let empty = (([||], [||]), (Instance.empty, Constraints.empty)) + let is_empty (_, (univs, cst)) = Instance.is_empty univs && Constraints.is_empty cst + + let pr prq prl ?variance (_, (univs, cst) as ctx) = + if is_empty ctx then mt() else + h (Instance.pr prq prl ?variance univs ++ str " |= ") ++ h (v 0 (Constraints.pr prl cst)) + + let hcons ((qnames, unames), (univs, cst)) = + ((Array.map Names.Name.hcons qnames, Array.map Names.Name.hcons unames), (Instance.hcons univs, hcons_constraints cst)) + + let names ((names, _) : t) = names + let instance (_, (univs, _cst)) = univs + let constraints (_, (_univs, cst)) = cst + + let union ((qna, una), (univs, cst)) ((qna', una'), (univs', cst')) = + (Array.append qna qna', Array.append una una'), (Instance.append univs univs', Constraints.union cst cst') + + let size (_,(x,_)) = Instance.length x + + let refine_names (qnames',unames') ((qnames, unames), x) = + let merge_names = Array.map2 Names.(fun old refined -> match refined with Anonymous -> old | Name _ -> refined) in + ((merge_names qnames qnames', merge_names unames unames'), x) + + let sort_levels a = + Array.sort Level.compare a; a + + let sort_qualities a = + Array.sort Quality.compare a; a + + let of_context_set f qctx (uctx, cst) = + let qctx = sort_qualities + (Array.map_of_list (fun q -> Quality.QVar q) + (Sorts.QVar.Set.elements qctx)) + in + let uctx = sort_levels (Array.of_list (Level.Set.elements uctx)) in + let inst = Instance.of_array (qctx, uctx) in + (f inst, (inst, cst)) + + let to_context_set (_, (ctx, cst)) = + let qctx, uctx = Instance.levels ctx in + qctx, (uctx, cst) + +end + +type universe_context = UContext.t +type 'a in_universe_context = 'a * universe_context + +let hcons_universe_context = UContext.hcons + +module AbstractContext = +struct + type t = bound_names constrained + + let make names csts : t = names, csts + + let instantiate inst ((qnames,unames), cst) = + let q, u = Instance.to_array inst in + assert (Array.length q == Array.length qnames && Array.length u = Array.length unames); + subst_instance_constraints inst cst + + let names (nas, _) = nas + + let hcons ((qnames,unames), cst) = + let qnames = Array.map Names.Name.hcons qnames in + let unames = Array.map Names.Name.hcons unames in + ((qnames, unames), hcons_constraints cst) + + let empty = (([||],[||]), Constraints.empty) + + let is_constant ((qnas,unas),_) = + Array.is_empty qnas && Array.is_empty unas + + let is_empty (_, cst as ctx) = + is_constant ctx && Constraints.is_empty cst + + let union ((qnas,unas), cst) ((qnas',unas'), cst') = + ((Array.append qnas qnas', Array.append unas unas'), Constraints.union cst cst') + + let size ((qnas,unas), _) = Array.length qnas, Array.length unas + + let repr (names, cst as self) : UContext.t = + let inst = Instance.abstract_instance (size self) in + (names, (inst, cst)) + + let pr prq pru ?variance ctx = UContext.pr prq pru ?variance (repr ctx) + +end + +type 'a univ_abstracted = { + univ_abstracted_value : 'a; + univ_abstracted_binder : AbstractContext.t; +} + +let map_univ_abstracted f {univ_abstracted_value;univ_abstracted_binder} = + let univ_abstracted_value = f univ_abstracted_value in + {univ_abstracted_value;univ_abstracted_binder} + +let hcons_abstract_universe_context = AbstractContext.hcons + +type sort_level_subst = Quality.t Sorts.QVar.Map.t * universe_level_subst + +let is_empty_sort_subst (qsubst,usubst) = Sorts.QVar.Map.is_empty qsubst && is_empty_level_subst usubst + +let empty_sort_subst = Sorts.QVar.Map.empty, empty_level_subst + +let subst_sort_level_instance (qsubst,usubst) i = + let i' = Instance.subst_fn (Quality.subst_fn qsubst, subst_univs_level_level usubst) i in + if i == i' then i + else i' + +let subst_univs_level_abstract_universe_context subst (inst, csts) = + inst, subst_univs_level_constraints subst csts + +let subst_fn_of_qsubst qsubst qv = + match Sorts.QVar.Map.find_opt qv qsubst with + | None -> Quality.QVar qv + | Some q -> q + +let subst_sort_level_sort (qsubst,usubst) s = + let fq qv = subst_fn_of_qsubst qsubst qv in + let fu u = subst_univs_level_universe usubst u in + Sorts.subst_fn (fq,fu) s + +let subst_sort_level_relevance (qsubst,_) r = + Sorts.relevance_subst_fn (subst_fn_of_qsubst qsubst) r + +let make_instance_subst i = + let qarr, uarr = Instance.to_array i in + let qsubst = + Array.fold_left_i (fun i acc l -> + let l = match l with Quality.QVar l -> l | _ -> assert false in + Sorts.QVar.Map.add l (Quality.var i) acc) + Sorts.QVar.Map.empty qarr + in + let usubst = + Array.fold_left_i (fun i acc l -> + Level.Map.add l (Level.var i) acc) + Level.Map.empty uarr + in + qsubst, usubst + +let make_abstract_instance ctx = + UContext.instance (AbstractContext.repr ctx) + +let abstract_universes uctx = + let nas = UContext.names uctx in + let instance = UContext.instance uctx in + let subst = make_instance_subst instance in + let cstrs = subst_univs_level_constraints (snd subst) + (UContext.constraints uctx) + in + let ctx = (nas, cstrs) in + instance, ctx + +let pr_universe_context = UContext.pr + +let pr_abstract_universe_context = AbstractContext.pr diff --git a/kernel/uVars.mli b/kernel/uVars.mli new file mode 100644 index 000000000000..fdb6808931aa --- /dev/null +++ b/kernel/uVars.mli @@ -0,0 +1,235 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * Copyright INRIA, CNRS and contributors *) +(* t -> bool + + val sup : t -> t -> t + + val pr : t -> Pp.t + + val equal : t -> t -> bool + +end + +(** {6 Universe instances} *) + +module Instance : +sig + type t + (** A universe instance represents a vector of argument universes + and sort qualities to a polymorphic definition + (constant, inductive or constructor). *) + + val empty : t + val is_empty : t -> bool + + val of_array : Quality.t array * Level.t array -> t + val to_array : t -> Quality.t array * Level.t array + + val abstract_instance : int * int -> t + (** Instance of the given size made of QVar/Level.var *) + + val append : t -> t -> t + (** To concatenate two instances, used for discharge *) + + val equal : t -> t -> bool + (** Equality *) + + val length : t -> int * int + (** Instance length *) + + val hcons : t -> t + (** Hash-consing. *) + + val hash : t -> int + (** Hash value *) + + val share : t -> t * int + (** Simultaneous hash-consing and hash-value computation *) + + val pr : (QVar.t -> Pp.t) -> (Level.t -> Pp.t) -> ?variance:Variance.t array -> t -> Pp.t + (** Pretty-printing, no comments *) + + val levels : t -> Quality.Set.t * Level.Set.t + (** The set of levels in the instance *) + + val subst_fn + : (QVar.t -> Quality.t) * (Level.t -> Level.t) + -> t -> t + +end + +val eq_sizes : int * int -> int * int -> bool +(** Convenient function to compare the result of Instance.length, UContext.size etc *) + +type 'a quconstraint_function = 'a -> 'a -> Sorts.QUConstraints.t -> Sorts.QUConstraints.t + +val enforce_eq_instances : Instance.t quconstraint_function + +val enforce_eq_variance_instances : Variance.t array -> Instance.t quconstraint_function +val enforce_leq_variance_instances : Variance.t array -> Instance.t quconstraint_function + +type 'a puniverses = 'a * Instance.t +val out_punivs : 'a puniverses -> 'a +val in_punivs : 'a -> 'a puniverses + +val eq_puniverses : ('a -> 'a -> bool) -> 'a puniverses -> 'a puniverses -> bool + +type bound_names = Names.Name.t array * Names.Name.t array + +(** A vector of universe levels with universe Constraints.t, + representing local universe variables and associated Constraints.t; + the names are user-facing names for printing *) + +module UContext : +sig + type t + + val make : bound_names -> Instance.t constrained -> t + + val empty : t + val is_empty : t -> bool + + val instance : t -> Instance.t + val constraints : t -> Constraints.t + + val union : t -> t -> t + (** Keeps the order of the instances *) + + val size : t -> int * int + (** The number of universes in the context *) + + val names : t -> bound_names + (** Return the user names of the universes *) + + val refine_names : bound_names -> t -> t + (** Use names to name the possibly yet unnamed universes *) + + val sort_levels : Level.t array -> Level.t array + (** Arbitrary choice of linear order of the variables *) + + val sort_qualities : Quality.t array -> Quality.t array + (** Arbitrary choice of linear order of the variables *) + + val of_context_set : (Instance.t -> bound_names) -> QVar.Set.t -> ContextSet.t -> t + (** Build a vector of universe levels assuming a function generating names *) + + val to_context_set : t -> Quality.Set.t * ContextSet.t + (** Discard the names and order of the universes *) + +end +(** A value in a universe context. *) +type 'a in_universe_context = 'a * UContext.t + +module AbstractContext : +sig + type t + (** An abstract context serves to quantify over a graph of universes + represented using de Bruijn indices, as in: + u0, ..., u(n-1), Var i < Var j, .., Var k <= Var l |- term(Var 0 .. Var (n-1)) + \-------------/ \-------------------------------/ \---------------------/ + names for constraints expressed on de Bruijn judgement in abstract + printing representation of the n univ vars context expected to + use de Bruijn indices + *) + + val make : bound_names -> Constraints.t -> t + (** Build an abstract context. Constraints may be between universe + variables. *) + + val repr : t -> UContext.t + (** [repr ctx] is [(Var(0), ... Var(n-1) |= cstr] where [n] is the length of + the context and [cstr] the abstracted Constraints.t. *) + + val empty : t + val is_empty : t -> bool + + val is_constant : t -> bool + (** Empty instance, but may have constraints *) + + val size : t -> int * int + + val union : t -> t -> t + (** The constraints are expected to be relative to the concatenated set of universes *) + + val instantiate : Instance.t -> t -> Constraints.t + (** Generate the set of instantiated Constraints.t **) + + val names : t -> bound_names + (** Return the names of the bound universe variables *) + +end + +type 'a univ_abstracted = { + univ_abstracted_value : 'a; + univ_abstracted_binder : AbstractContext.t; +} +(** A value with bound universe levels. *) + +val map_univ_abstracted : ('a -> 'b) -> 'a univ_abstracted -> 'b univ_abstracted + +(** {6 Substitution} *) + +type sort_level_subst = Quality.t QVar.Map.t * universe_level_subst + +val empty_sort_subst : sort_level_subst + +val is_empty_sort_subst : sort_level_subst -> bool + +val subst_univs_level_abstract_universe_context : + universe_level_subst -> AbstractContext.t -> AbstractContext.t +(** There are no constraints on qualities, so this only needs a subst for univs *) + +val subst_sort_level_instance : sort_level_subst -> Instance.t -> Instance.t +(** Level to universe substitutions. *) + +val subst_sort_level_sort : sort_level_subst -> Sorts.t -> Sorts.t + +val subst_sort_level_relevance : sort_level_subst -> Sorts.relevance -> Sorts.relevance + +(** Substitution of instances *) +val subst_instance_instance : Instance.t -> Instance.t -> Instance.t +val subst_instance_universe : Instance.t -> Universe.t -> Universe.t +val subst_instance_sort : Instance.t -> Sorts.t -> Sorts.t +val subst_instance_relevance : Instance.t -> Sorts.relevance -> Sorts.relevance + +val make_instance_subst : Instance.t -> sort_level_subst +(** Creates [u(0) ↦ 0; ...; u(n-1) ↦ n - 1] out of [u(0); ...; u(n - 1)] *) + +val abstract_universes : UContext.t -> Instance.t * AbstractContext.t +(** TODO: move universe abstraction out of the kernel *) + +val make_abstract_instance : AbstractContext.t -> Instance.t + +(** {6 Pretty-printing of universes. } *) + +val pr_universe_context : (QVar.t -> Pp.t) -> (Level.t -> Pp.t) -> ?variance:Variance.t array -> + UContext.t -> Pp.t +val pr_abstract_universe_context : (QVar.t -> Pp.t) -> (Level.t -> Pp.t) -> ?variance:Variance.t array -> + AbstractContext.t -> Pp.t + +(** {6 Hash-consing } *) + +val hcons_universe_context : UContext.t -> UContext.t +val hcons_abstract_universe_context : AbstractContext.t -> AbstractContext.t diff --git a/kernel/univ.ml b/kernel/univ.ml index d3c7fce30691..73a039113d91 100644 --- a/kernel/univ.ml +++ b/kernel/univ.ml @@ -18,7 +18,6 @@ Sozeau, Pierre-Marie Pédrot *) open Pp -open CErrors open Util (* Universes are stratified by a partial ordering $\le$. @@ -231,9 +230,6 @@ module Level = struct let pr prl s = hov 1 (str"{" ++ prlist_with_sep spc prl (elements s) ++ str"}") - let of_array l = - Array.fold_left (fun acc x -> add x acc) empty l - end end @@ -459,6 +455,9 @@ struct let for_all = List.for_all let repr x : t = x + let unrepr l = + assert (not (List.is_empty l)); + sort l end type constraint_type = AcyclicGraph.constraint_type = Lt | Le | Eq @@ -507,8 +506,6 @@ struct end -type constraints = Constraints.t - module Hconstraint = Hashcons.Make( struct @@ -523,7 +520,7 @@ module Hconstraint = module Hconstraints = Hashcons.Make( struct - type t = constraints + type t = Constraints.t type u = univ_constraint -> univ_constraint let hashcons huc s = Constraints.fold (fun x -> Constraints.add (huc x)) s Constraints.empty @@ -539,13 +536,13 @@ let hcons_constraints = Hashcons.simple_hcons Hconstraints.generate Hconstraints (** A value with universe constraints. *) -type 'a constrained = 'a * constraints +type 'a constrained = 'a * Constraints.t let constraints_of (_, cst) = cst (** Constraints functions. *) -type 'a constraint_function = 'a -> 'a -> constraints -> constraints +type 'a constraint_function = 'a -> 'a -> Constraints.t -> Constraints.t let enforce_eq_level u v c = (* We discard trivial constraints like u=u *) @@ -578,292 +575,6 @@ let univ_level_rem u v min = type universe_level_subst = universe_level Level.Map.t -module Variance = -struct - (** A universe position in the instance given to a cumulative - inductive can be the following. Note there is no Contravariant - case because [forall x : A, B <= forall x : A', B'] requires [A = - A'] as opposed to [A' <= A]. *) - type t = Irrelevant | Covariant | Invariant - - let sup x y = - match x, y with - | Irrelevant, s | s, Irrelevant -> s - | Invariant, _ | _, Invariant -> Invariant - | Covariant, Covariant -> Covariant - - let equal a b = match a,b with - | Irrelevant, Irrelevant | Covariant, Covariant | Invariant, Invariant -> true - | (Irrelevant | Covariant | Invariant), _ -> false - - let check_subtype x y = match x, y with - | (Irrelevant | Covariant | Invariant), Irrelevant -> true - | Irrelevant, Covariant -> false - | (Covariant | Invariant), Covariant -> true - | (Irrelevant | Covariant), Invariant -> false - | Invariant, Invariant -> true - - let pr = function - | Irrelevant -> str "*" - | Covariant -> str "+" - | Invariant -> str "=" - - let leq_constraint csts variance u u' = - match variance with - | Irrelevant -> csts - | Covariant -> enforce_leq_level u u' csts - | Invariant -> enforce_eq_level u u' csts - - let eq_constraint csts variance u u' = - match variance with - | Irrelevant -> csts - | Covariant | Invariant -> enforce_eq_level u u' csts - - let leq_constraints variance u u' csts = - let len = Array.length u in - assert (len = Array.length u' && len = Array.length variance); - Array.fold_left3 leq_constraint csts variance u u' - - let eq_constraints variance u u' csts = - let len = Array.length u in - assert (len = Array.length u' && len = Array.length variance); - Array.fold_left3 eq_constraint csts variance u u' -end - -module Instance : sig - type t = Level.t array - - val empty : t - val is_empty : t -> bool - - val of_array : Level.t array -> t - val to_array : t -> Level.t array - - val append : t -> t -> t - val equal : t -> t -> bool - val length : t -> int - - val hcons : t -> t - val hash : t -> int - - val share : t -> t * int - - val subst_fn : (Level.t -> Level.t) -> t -> t - - val pr : (Level.t -> Pp.t) -> ?variance:Variance.t array -> t -> Pp.t - val levels : t -> Level.Set.t -end = -struct - type t = Level.t array - - let empty : t = [||] - - module HInstancestruct = - struct - type nonrec t = t - type u = Level.t -> Level.t - - let hashcons huniv a = - let len = Array.length a in - if Int.equal len 0 then empty - else begin - for i = 0 to len - 1 do - let x = Array.unsafe_get a i in - let x' = huniv x in - if x == x' then () - else Array.unsafe_set a i x' - done; - a - end - - let eq t1 t2 = - t1 == t2 || - (Int.equal (Array.length t1) (Array.length t2) && - let rec aux i = - (Int.equal i (Array.length t1)) || (t1.(i) == t2.(i) && aux (i + 1)) - in aux 0) - - let hash a = - let accu = ref 0 in - for i = 0 to Array.length a - 1 do - let l = Array.unsafe_get a i in - let h = Level.hash l in - accu := Hashset.Combine.combine !accu h; - done; - (* [h] must be positive. *) - let h = !accu land 0x3FFFFFFF in - h - end - - module HInstance = Hashcons.Make(HInstancestruct) - - let hcons = Hashcons.simple_hcons HInstance.generate HInstance.hcons Level.hcons - - let hash = HInstancestruct.hash - - let share a = (hcons a, hash a) - - let empty = hcons [||] - - let is_empty x = Int.equal (Array.length x) 0 - - let append x y = - if Array.length x = 0 then y - else if Array.length y = 0 then x - else Array.append x y - - let of_array a = - a - - let to_array a = a - - let length a = Array.length a - - let subst_fn fn t = - let t' = CArray.Smart.map fn t in - if t' == t then t else of_array t' - - let levels x = Level.Set.of_array x - - let pr prl ?variance = - let ppu i u = - let v = Option.map (fun v -> v.(i)) variance in - pr_opt_no_spc Variance.pr v ++ prl u - in - prvecti_with_sep spc ppu - - let equal t u = - t == u || - (Array.is_empty t && Array.is_empty u) || - (CArray.for_all2 Level.equal t u - (* Necessary as universe instances might come from different modules and - unmarshalling doesn't preserve sharing *)) - -end - -let enforce_eq_instances x y = - let ax = Instance.to_array x and ay = Instance.to_array y in - if Array.length ax != Array.length ay then - anomaly (Pp.(++) (Pp.str "Invalid argument: enforce_eq_instances called with") - (Pp.str " instances of different lengths.")); - CArray.fold_right2 enforce_eq_level ax ay - -let enforce_eq_variance_instances = Variance.eq_constraints -let enforce_leq_variance_instances = Variance.leq_constraints - -let subst_instance_level s l = - match l.Level.data with - | Level.Var n -> s.(n) - | _ -> l - -let subst_instance_instance s i = - Array.Smart.map (fun l -> subst_instance_level s l) i - -let subst_instance_universe s u = - let f x = Universe.Expr.map (fun u -> subst_instance_level s u) x in - let u' = List.Smart.map f u in - if u == u' then u - else Universe.sort u' - -let subst_instance_constraint s (u,d,v as c) = - let u' = subst_instance_level s u in - let v' = subst_instance_level s v in - if u' == u && v' == v then c - else (u',d,v') - -let subst_instance_constraints s csts = - Constraints.fold - (fun c csts -> Constraints.add (subst_instance_constraint s c) csts) - csts Constraints.empty - -type 'a puniverses = 'a * Instance.t -let out_punivs (x, _y) = x -let in_punivs x = (x, Instance.empty) -let eq_puniverses f (x, u) (y, u') = - f x y && Instance.equal u u' - -(** A context of universe levels with universe constraints, - representing local universe variables and constraints *) - -module UContext = -struct - type t = Names.Name.t array * Instance.t constrained - - let make names (univs, _ as x) = - assert (Array.length names = Array.length univs); - assert (Array.for_all (fun u -> not (Level.is_small u)) univs); - (names, x) - - (** Universe contexts (variables as a list) *) - let empty = ([||], (Instance.empty, Constraints.empty)) - let is_empty (_, (univs, cst)) = Instance.is_empty univs && Constraints.is_empty cst - - let pr prl ?variance (_, (univs, cst) as ctx) = - if is_empty ctx then mt() else - h (Instance.pr prl ?variance univs ++ str " |= ") ++ h (v 0 (Constraints.pr prl cst)) - - let hcons (names, (univs, cst)) = - (Array.map Names.Name.hcons names, (Instance.hcons univs, hcons_constraints cst)) - - let names (names, _) = names - let instance (_, (univs, _cst)) = univs - let constraints (_, (_univs, cst)) = cst - - let union (na, (univs, cst)) (na', (univs', cst')) = - Array.append na na', (Instance.append univs univs', Constraints.union cst cst') - - let size (_,(x,_)) = Instance.length x - - let refine_names names' (names, x) = - let merge_names = Array.map2 Names.(fun old refined -> match refined with Anonymous -> old | Name _ -> refined) in - (merge_names names names', x) - -end - -type universe_context = UContext.t -let hcons_universe_context = UContext.hcons - -module AbstractContext = -struct - type t = Names.Name.t array constrained - - let make names csts : t = names, csts - - let repr (inst, cst) = - (inst, (Array.init (Array.length inst) (fun i -> Level.var i), cst)) - - let pr f ?variance ctx = UContext.pr f ?variance (repr ctx) - - let instantiate inst (u, cst) = - assert (Array.length u = Array.length inst); - subst_instance_constraints inst cst - - let names (nas, _) = nas - - let hcons (univs, cst) = - (Array.map Names.Name.hcons univs, hcons_constraints cst) - - let empty = ([||], Constraints.empty) - - let is_empty (nas, cst) = Array.is_empty nas && Constraints.is_empty cst - - let union (nas, cst) (nas', cst') = (Array.append nas nas', Constraints.union cst cst') - - let size (nas, _) = Array.length nas - -end - -type 'a univ_abstracted = { - univ_abstracted_value : 'a; - univ_abstracted_binder : AbstractContext.t; -} - -let map_univ_abstracted f {univ_abstracted_value;univ_abstracted_binder} = - let univ_abstracted_value = f univ_abstracted_value in - {univ_abstracted_value;univ_abstracted_binder} - -let hcons_abstract_universe_context = AbstractContext.hcons - (** A set of universes with universe constraints. We linearize the set to a list after typechecking. Beware, representation could change. @@ -900,16 +611,6 @@ struct let add_constraints cst' (univs, cst) = univs, Constraints.union cst cst' - let sort_levels a = - Array.sort Level.compare a; a - - let to_context f (ctx, cst) = - let inst = Instance.of_array (sort_levels (Array.of_list (Level.Set.elements ctx))) in - (f inst, (inst, cst)) - - let of_context (_, (ctx, cst)) = - (Instance.levels ctx, cst) - let pr prl (univs, cst as ctx) = if is_empty ctx then mt() else hov 0 (h (Level.Set.pr prl univs ++ str " |=") ++ brk(1,2) ++ h (Constraints.pr prl cst)) @@ -923,12 +624,8 @@ end type universe_context_set = ContextSet.t (** A value in a universe context (resp. context set). *) -type 'a in_universe_context = 'a * universe_context type 'a in_universe_context_set = 'a * universe_context_set -let extend_in_context_set (a, ctx) ctx' = - (a, ContextSet.union ctx ctx') - (** Substitutions. *) let empty_level_subst = Level.Map.empty @@ -947,11 +644,6 @@ let subst_univs_level_universe subst u = if u == u' then u else Universe.sort u' -let subst_univs_level_instance subst i = - let i' = Instance.subst_fn (subst_univs_level_level subst) i in - if i == i' then i - else i' - let subst_univs_level_constraint subst (u,d,v) = let u' = subst_univs_level_level subst u and v' = subst_univs_level_level subst v in @@ -963,37 +655,8 @@ let subst_univs_level_constraints subst csts = (fun c -> Option.fold_right Constraints.add (subst_univs_level_constraint subst c)) csts Constraints.empty -let subst_univs_level_abstract_universe_context subst (inst, csts) = - inst, subst_univs_level_constraints subst csts - -let make_instance_subst i = - let arr = Instance.to_array i in - Array.fold_left_i (fun i acc l -> - Level.Map.add l (Level.var i) acc) - Level.Map.empty arr - -let make_abstract_instance (ctx, _) = - Array.init (Array.length ctx) (fun i -> Level.var i) - -let abstract_universes uctx = - let nas = UContext.names uctx in - let instance = UContext.instance uctx in - let () = assert (Int.equal (Array.length nas) (Instance.length instance)) in - let subst = make_instance_subst instance in - let cstrs = subst_univs_level_constraints subst - (UContext.constraints uctx) - in - let ctx = (nas, cstrs) in - instance, ctx - (** Pretty-printing *) -let pr_constraints prl = Constraints.pr prl - -let pr_universe_context = UContext.pr - -let pr_abstract_universe_context = AbstractContext.pr - let pr_universe_context_set = ContextSet.pr let pr_universe_level_subst prl = diff --git a/kernel/univ.mli b/kernel/univ.mli index dc661eca2cec..c18d40dc79d4 100644 --- a/kernel/univ.mli +++ b/kernel/univ.mli @@ -46,6 +46,8 @@ sig val hash : t -> int + val hcons : t -> t + val make : UGlobal.t -> t val raw_pr : t -> Pp.t @@ -146,7 +148,9 @@ sig val exists : (Level.t * int -> bool) -> t -> bool val for_all : (Level.t * int -> bool) -> t -> bool + val repr : t -> (Level.t * int) list + val unrepr : (Level.t * int) list -> t module Set : CSet.S with type elt = t module Map : CMap.ExtS with type key = t and module Set := Set @@ -168,7 +172,9 @@ type constraint_type = AcyclicGraph.constraint_type = Lt | Le | Eq type univ_constraint = Level.t * constraint_type * Level.t module Constraints : sig - include Set.S with type elt = univ_constraint + include Set.S with type elt = univ_constraint + + val pr : (Level.t -> Pp.t) -> t -> Pp.t end (** A value with universe Constraints.t. *) @@ -183,154 +189,6 @@ type 'a constraint_function = 'a -> 'a -> Constraints.t -> Constraints.t val enforce_eq_level : Level.t constraint_function val enforce_leq_level : Level.t constraint_function -(** {6 Support for universe polymorphism } *) - -module Variance : -sig - (** A universe position in the instance given to a cumulative - inductive can be the following. Note there is no Contravariant - case because [forall x : A, B <= forall x : A', B'] requires [A = - A'] as opposed to [A' <= A]. *) - type t = Irrelevant | Covariant | Invariant - - (** [check_subtype x y] holds if variance [y] is also an instance of [x] *) - val check_subtype : t -> t -> bool - - val sup : t -> t -> t - - val pr : t -> Pp.t - - val equal : t -> t -> bool - -end - -(** {6 Universe instances} *) - -module Instance : -sig - type t - (** A universe instance represents a vector of argument universes - to a polymorphic definition (constant, inductive or constructor). *) - - val empty : t - val is_empty : t -> bool - - val of_array : Level.t array -> t - val to_array : t -> Level.t array - - val append : t -> t -> t - (** To concatenate two instances, used for discharge *) - - val equal : t -> t -> bool - (** Equality *) - - val length : t -> int - (** Instance length *) - - val hcons : t -> t - (** Hash-consing. *) - - val hash : t -> int - (** Hash value *) - - val share : t -> t * int - (** Simultaneous hash-consing and hash-value computation *) - - val pr : (Level.t -> Pp.t) -> ?variance:Variance.t array -> t -> Pp.t - (** Pretty-printing, no comments *) - - val levels : t -> Level.Set.t - (** The set of levels in the instance *) - -end - -val enforce_eq_instances : Instance.t constraint_function - -val enforce_eq_variance_instances : Variance.t array -> Instance.t constraint_function -val enforce_leq_variance_instances : Variance.t array -> Instance.t constraint_function - -type 'a puniverses = 'a * Instance.t -val out_punivs : 'a puniverses -> 'a -val in_punivs : 'a -> 'a puniverses - -val eq_puniverses : ('a -> 'a -> bool) -> 'a puniverses -> 'a puniverses -> bool - -(** A vector of universe levels with universe Constraints.t, - representing local universe variables and associated Constraints.t; - the names are user-facing names for printing *) - -module UContext : -sig - type t - - val make : Names.Name.t array -> Instance.t constrained -> t - (** The name array must be the same length as the instance, and the - instance must not contain constant levels (ie no Set). *) - - val empty : t - val is_empty : t -> bool - - val instance : t -> Instance.t - val constraints : t -> Constraints.t - - val union : t -> t -> t - (** Keeps the order of the instances *) - - val size : t -> int - (** The number of universes in the context *) - - val names : t -> Names.Name.t array - (** Return the user names of the universes *) - - val refine_names : Names.Name.t array -> t -> t - (** Use names to name the possibly yet unnamed universes *) - -end - -module AbstractContext : -sig - type t - (** An abstract context serves to quantify over a graph of universes - represented using de Bruijn indices, as in: - u0, ..., u(n-1), Var i < Var j, .., Var k <= Var l |- term(Var 0 .. Var (n-1)) - \-------------/ \-------------------------------/ \---------------------/ - names for constraints expressed on de Bruijn judgement in abstract - printing representation of the n univ vars context expected to - use de Bruijn indices - *) - - val make : Names.Name.t array -> Constraints.t -> t - (** Build an abstract context. Constraints may be between universe - variables. *) - - val repr : t -> UContext.t - (** [repr ctx] is [(Var(0), ... Var(n-1) |= cstr] where [n] is the length of - the context and [cstr] the abstracted Constraints.t. *) - - val empty : t - val is_empty : t -> bool - - val size : t -> int - - val union : t -> t -> t - (** The constraints are expected to be relative to the concatenated set of universes *) - - val instantiate : Instance.t -> t -> Constraints.t - (** Generate the set of instantiated Constraints.t **) - - val names : t -> Names.Name.t array - (** Return the names of the bound universe variables *) - -end - -type 'a univ_abstracted = { - univ_abstracted_value : 'a; - univ_abstracted_binder : AbstractContext.t; -} -(** A value with bound universe levels. *) - -val map_univ_abstracted : ('a -> 'b) -> 'a univ_abstracted -> 'b univ_abstracted - (** Universe contexts (as sets) *) (** A set of universes with universe Constraints.t. @@ -359,14 +217,6 @@ sig val add_universe : Level.t -> t -> t val add_constraints : Constraints.t -> t -> t - val sort_levels : Level.t array -> Level.t array - (** Arbitrary choice of linear order of the variables *) - - val to_context : (Instance.t -> Names.Name.t array) -> t -> UContext.t - (** Build a vector of universe levels assuming a function generating names *) - - val of_context : UContext.t -> t - val constraints : t -> Constraints.t val levels : t -> Level.Set.t @@ -375,13 +225,9 @@ sig end -(** A value in a universe context (resp. context set). *) -type 'a in_universe_context = 'a * UContext.t +(** A value in a universe context set. *) type 'a in_universe_context_set = 'a * ContextSet.t -val extend_in_context_set : 'a in_universe_context_set -> ContextSet.t -> - 'a in_universe_context_set - (** {6 Substitution} *) type universe_level_subst = Level.t Level.Map.t @@ -393,32 +239,10 @@ val is_empty_level_subst : universe_level_subst -> bool val subst_univs_level_level : universe_level_subst -> Level.t -> Level.t val subst_univs_level_universe : universe_level_subst -> Universe.t -> Universe.t val subst_univs_level_constraints : universe_level_subst -> Constraints.t -> Constraints.t -val subst_univs_level_abstract_universe_context : - universe_level_subst -> AbstractContext.t -> AbstractContext.t -val subst_univs_level_instance : universe_level_subst -> Instance.t -> Instance.t - -(** Level to universe substitutions. *) - -(** Substitution of instances *) -val subst_instance_instance : Instance.t -> Instance.t -> Instance.t -val subst_instance_universe : Instance.t -> Universe.t -> Universe.t - -val make_instance_subst : Instance.t -> universe_level_subst -(** Creates [u(0) ↦ 0; ...; u(n-1) ↦ n - 1] out of [u(0); ...; u(n - 1)] *) - -val abstract_universes : UContext.t -> Instance.t * AbstractContext.t -(** TODO: move universe abstraction out of the kernel *) - -val make_abstract_instance : AbstractContext.t -> Instance.t (** {6 Pretty-printing of universes. } *) val pr_constraint_type : constraint_type -> Pp.t -val pr_constraints : (Level.t -> Pp.t) -> Constraints.t -> Pp.t -val pr_universe_context : (Level.t -> Pp.t) -> ?variance:Variance.t array -> - UContext.t -> Pp.t -val pr_abstract_universe_context : (Level.t -> Pp.t) -> ?variance:Variance.t array -> - AbstractContext.t -> Pp.t val pr_universe_context_set : (Level.t -> Pp.t) -> ContextSet.t -> Pp.t val pr_universe_level_subst : (Level.t -> Pp.t) -> universe_level_subst -> Pp.t @@ -428,6 +252,4 @@ val pr_universe_level_subst : (Level.t -> Pp.t) -> universe_level_subst -> Pp.t val hcons_univ : Universe.t -> Universe.t val hcons_constraints : Constraints.t -> Constraints.t val hcons_universe_set : Level.Set.t -> Level.Set.t -val hcons_universe_context : UContext.t -> UContext.t -val hcons_abstract_universe_context : AbstractContext.t -> AbstractContext.t val hcons_universe_context_set : ContextSet.t -> ContextSet.t diff --git a/kernel/vars.ml b/kernel/vars.ml index 2803a38359c4..1bd44cdab4fb 100644 --- a/kernel/vars.ml +++ b/kernel/vars.ml @@ -275,45 +275,137 @@ let smash_rel_context sign = (** Universe substitutions *) open Constr +let map_annot_relevance f na = + let open Context in + let r = na.binder_relevance in + let r' = f r in + if r' == r then na else { na with binder_relevance = r' } + +let map_case_info_relevance f ci = + let r = ci.ci_relevance in + let r' = f r in + if r' == r then ci else { ci with ci_relevance = r' } + +let map_case_under_context_relevance f (nas,x as v) = + let nas' = CArray.Smart.map (map_annot_relevance f) nas in + if nas' == nas then v else (nas',x) + +let map_rec_declaration_relevance f (i,(nas,x,y) as v) = +let nas' = CArray.Smart.map (map_annot_relevance f) nas in + if nas' == nas then v else (i,(nas',x,y)) + +let map_constr_relevance f c = + match kind c with + | Rel _ | Var _ | Meta _ | Evar _ + | Sort _ | Cast _ | App _ + | Const _ | Ind _ | Construct _ + | Int _ | Float _ | Array _ -> c + + | Prod (na,x,y) -> + let na' = map_annot_relevance f na in + if na' == na then c else mkProd (na',x,y) + + | Lambda (na,x,y) -> + let na' = map_annot_relevance f na in + if na' == na then c else mkLambda (na',x,y) + + | LetIn (na,x,y,z) -> + let na' = map_annot_relevance f na in + if na' == na then c else mkLetIn (na',x,y,z) + + | Case (ci,u,params,ret,iv,v,brs) -> + let ci' = map_case_info_relevance f ci in + let ret' = map_case_under_context_relevance f ret in + let brs' = CArray.Smart.map (map_case_under_context_relevance f) brs in + if ci' == ci && ret' == ret && brs' == brs then c + else mkCase (ci',u,params,ret',iv,v,brs') + + | Fix data -> + let data' = map_rec_declaration_relevance f data in + if data' == data then c else mkFix data' + + | CoFix data -> + let data' = map_rec_declaration_relevance f data in + if data' == data then c else mkCoFix data' + + | Proj (p, r, v) -> + let r' = f r in + if r' == r then c else mkProj (p, r', v) + +let fold_annot_relevance f acc na = + f acc na.Context.binder_relevance + +let fold_case_under_context_relevance f acc (nas,_) = + Array.fold_left (fold_annot_relevance f) acc nas + +let fold_rec_declaration_relevance f acc (nas,_,_) = + Array.fold_left (fold_annot_relevance f) acc nas + +let fold_constr_relevance f acc c = + match kind c with + | Rel _ | Var _ | Meta _ | Evar _ + | Sort _ | Cast _ | App _ + | Const _ | Ind _ | Construct _ + | Int _ | Float _ | Array _ -> acc + + | Prod (na,_,_) | Lambda (na,_,_) | LetIn (na,_,_,_) -> + fold_annot_relevance f acc na + + | Case (ci,_u,_params,ret,_iv,_v,brs) -> + let acc = f acc ci.ci_relevance in + let acc = fold_case_under_context_relevance f acc ret in + let acc = CArray.fold_left (fold_case_under_context_relevance f) acc brs in + acc + + | Proj (_, r, _) -> f acc r + + | Fix (_,data) + | CoFix (_,data) -> + fold_rec_declaration_relevance f acc data + let subst_univs_level_constr subst c = - if Univ.is_empty_level_subst subst then c + if UVars.is_empty_sort_subst subst then c else - let f = Univ.subst_univs_level_instance subst in + let f = UVars.subst_sort_level_instance subst in + (* XXX shouldn't Constr.map return the pointer equal term when unchanged instead? *) let changed = ref false in let rec aux t = + let t' = map_constr_relevance (UVars.subst_sort_level_relevance subst) t in + let t = if t' == t then t else (changed := true; t') in match kind t with | Const (c, u) -> - if Univ.Instance.is_empty u then t + if UVars.Instance.is_empty u then t else let u' = f u in if u' == u then t else (changed := true; mkConstU (c, u')) | Ind (i, u) -> - if Univ.Instance.is_empty u then t + if UVars.Instance.is_empty u then t else let u' = f u in if u' == u then t else (changed := true; mkIndU (i, u')) | Construct (c, u) -> - if Univ.Instance.is_empty u then t + if UVars.Instance.is_empty u then t else let u' = f u in if u' == u then t else (changed := true; mkConstructU (c, u')) - | Sort (Sorts.Type u) -> - let u' = Univ.subst_univs_level_universe subst u in - if u' == u then t else - (changed := true; mkSort (Sorts.sort_of_univ u')) + | Sort s -> + let s' = UVars.subst_sort_level_sort subst s in + if s' == s then t + else + (changed := true; mkSort s') | Case (ci, u, pms, p, CaseInvert {indices}, c, br) -> - if Univ.Instance.is_empty u then Constr.map aux t + if UVars.Instance.is_empty u then Constr.map aux t else let u' = f u in if u' == u then Constr.map aux t else (changed:=true; Constr.map aux (mkCase (ci,u',pms,p,CaseInvert {indices},c,br))) | Case (ci, u, pms, p, NoInvert, c, br) -> - if Univ.Instance.is_empty u then Constr.map aux t + if UVars.Instance.is_empty u then Constr.map aux t else let u' = f u in if u' == u then Constr.map aux t @@ -333,37 +425,42 @@ let subst_univs_level_constr subst c = let c' = aux c in if !changed then c' else c -let subst_univs_level_context s = - Context.Rel.map (subst_univs_level_constr s) +let subst_univs_level_context s ctx = + CList.Smart.map (fun d -> + let d = RelDecl.map_relevance (UVars.subst_sort_level_relevance s) d in + RelDecl.map_constr (subst_univs_level_constr s) d) + ctx let subst_instance_constr subst c = - if Univ.Instance.is_empty subst then c + if UVars.Instance.is_empty subst then c else - let f u = Univ.subst_instance_instance subst u in + let f u = UVars.subst_instance_instance subst u in let rec aux t = + let t = if CArray.is_empty (fst (UVars.Instance.to_array subst)) then t + else map_constr_relevance (UVars.subst_instance_relevance subst) t + in match kind t with | Const (c, u) -> - if Univ.Instance.is_empty u then t + if UVars.Instance.is_empty u then t else let u' = f u in if u' == u then t else (mkConstU (c, u')) | Ind (i, u) -> - if Univ.Instance.is_empty u then t + if UVars.Instance.is_empty u then t else let u' = f u in if u' == u then t else (mkIndU (i, u')) | Construct (c, u) -> - if Univ.Instance.is_empty u then t + if UVars.Instance.is_empty u then t else let u' = f u in if u' == u then t else (mkConstructU (c, u')) - | Sort (Sorts.Type u) -> - let u' = Univ.subst_instance_universe subst u in - if u' == u then t else - (mkSort (Sorts.sort_of_univ u')) + | Sort s -> + let s' = UVars.subst_instance_sort subst s in + if s' == s then t else mkSort s' | Case (ci, u, pms, p, iv, c, br) -> let u' = f u in @@ -383,34 +480,57 @@ let subst_instance_constr subst c = aux c let univ_instantiate_constr u c = - let open Univ in - assert (Int.equal (Instance.length u) (AbstractContext.size c.univ_abstracted_binder)); + let open UVars in + assert (UVars.eq_sizes (Instance.length u) (AbstractContext.size c.univ_abstracted_binder)); subst_instance_constr u c.univ_abstracted_value let subst_instance_context s ctx = - if Univ.Instance.is_empty s then ctx - else Context.Rel.map (fun x -> subst_instance_constr s x) ctx + if UVars.Instance.is_empty s then ctx + else + CList.Smart.map (fun d -> + let d = RelDecl.map_relevance (UVars.subst_instance_relevance s) d in + RelDecl.map_constr (subst_instance_constr s) d) + ctx + +let add_qvars_and_univs_of_instance (qs,us) u = + let qs', us' = UVars.Instance.to_array u in + let qs = Array.fold_left (fun qs q -> + let open Sorts.Quality in + match q with + | QVar q -> Sorts.QVar.Set.add q qs + | QConstant _ -> qs) + qs qs' + in + let us = Array.fold_left (fun acc x -> Univ.Level.Set.add x acc) us us' in + qs, us + +let add_relevance (qs,us as v) = let open Sorts in function + | Irrelevant | Relevant -> v + | RelevanceVar q -> QVar.Set.add q qs, us -let universes_of_constr c = +let sort_and_universes_of_constr c = let open Univ in let rec aux s c = + let s = fold_constr_relevance add_relevance s c in match kind c with - | Const (_c, u) -> - Level.Set.fold Level.Set.add (Instance.levels u) s - | Ind ((_mind,_), u) | Construct (((_mind,_),_), u) -> - Level.Set.fold Level.Set.add (Instance.levels u) s - | Sort u when not (Sorts.is_small u) -> - Level.Set.fold Level.Set.add (Sorts.levels u) s + | Const (_, u) | Ind (_, u) | Construct (_,u) -> add_qvars_and_univs_of_instance s u + | Sort (Sorts.Type u) -> + Util.on_snd (Level.Set.union (Universe.levels u)) s + | Sort (Sorts.QSort (q,u)) -> + let qs, us = s in + Sorts.QVar.Set.add q qs, Level.Set.union us (Universe.levels u) | Array (u,_,_,_) -> - let s = Level.Set.fold Level.Set.add (Instance.levels u) s in + let s = add_qvars_and_univs_of_instance s u in Constr.fold aux s c | Case (_, u, _, _, _,_ ,_) -> - let s = Level.Set.fold Level.Set.add (Instance.levels u) s in + let s = add_qvars_and_univs_of_instance s u in Constr.fold aux s c | _ -> Constr.fold aux s c - in aux Level.Set.empty c + in aux (Sorts.QVar.Set.empty,Level.Set.empty) c -let universes_of_constr c = - NewProfile.profile "universes_of_constr" (fun () -> - universes_of_constr c) +let sort_and_universes_of_constr c = + NewProfile.profile "sort_and_universes_of_constr" (fun () -> + sort_and_universes_of_constr c) () + +let universes_of_constr c = snd (sort_and_universes_of_constr c) diff --git a/kernel/vars.mli b/kernel/vars.mli index 6985db002c02..a5e6729b9232 100644 --- a/kernel/vars.mli +++ b/kernel/vars.mli @@ -186,12 +186,12 @@ val smash_rel_context : rel_context -> rel_context (** {3 Substitution of universes} *) -open Univ +open UVars (** Level substitutions for polymorphism. *) -val subst_univs_level_constr : universe_level_subst -> constr -> constr -val subst_univs_level_context : Univ.universe_level_subst -> Constr.rel_context -> Constr.rel_context +val subst_univs_level_constr : sort_level_subst -> constr -> constr +val subst_univs_level_context : sort_level_subst -> Constr.rel_context -> Constr.rel_context (** Instance substitution for polymorphism. *) val subst_instance_constr : Instance.t -> constr -> constr @@ -200,6 +200,11 @@ val subst_instance_context : Instance.t -> Constr.rel_context -> Constr.rel_cont val univ_instantiate_constr : Instance.t -> constr univ_abstracted -> constr (** Ignores the constraints carried by [univ_abstracted]. *) +val map_constr_relevance : (Sorts.relevance -> Sorts.relevance) -> Constr.t -> Constr.t +(** Modifies the relevances in the head node (not in subterms) *) + +val sort_and_universes_of_constr : constr -> Sorts.QVar.Set.t * Univ.Level.Set.t + val universes_of_constr : constr -> Univ.Level.Set.t (** {3 Low-level cached lift type} *) diff --git a/kernel/vconv.ml b/kernel/vconv.ml index 8f9c27cd8dc6..47634342bb54 100644 --- a/kernel/vconv.ml +++ b/kernel/vconv.ml @@ -9,10 +9,8 @@ open Vmsymtable (* Test la structure des piles *) let table_key_instance env = function -| ConstKey cst -> - let ctx = Environ.constant_context env cst in - Univ.AbstractContext.size ctx -| RelKey _ | VarKey _ | EvarKey _ -> 0 +| ConstKey cst -> Environ.constant_context env cst +| RelKey _ | VarKey _ | EvarKey _ -> UVars.AbstractContext.empty let compare_zipper z1 z2 = match z1, z2 with @@ -96,10 +94,9 @@ and conv_atom env pb k a1 stk1 a2 stk2 cu = match a1, a2 with | Aind ((mi,_i) as ind1) , Aind ind2 -> if Names.Ind.CanOrd.equal ind1 ind2 && compare_stack stk1 stk2 then - let ulen = Univ.AbstractContext.size (Environ.mind_context env mi) in - if ulen = 0 then + if UVars.AbstractContext.is_constant (Environ.mind_context env mi) then conv_stack env k stk1 stk2 cu - else + else begin match stk1 , stk2 with | Zapp args1 :: stk1' , Zapp args2 :: stk2' -> assert (0 < nargs args1); @@ -110,11 +107,11 @@ and conv_atom env pb k a1 stk1 a2 stk2 cu = conv_arguments env ~from:1 k args1 args2 (conv_stack env k stk1' stk2' cu) | _, _ -> assert false (* Should not happen if problem is well typed *) + end else raise NotConvertible | Aid ik1, Aid ik2 -> if Vmvalues.eq_id_key ik1 ik2 && compare_stack stk1 stk2 then - let ulen = table_key_instance env ik1 in - if ulen = 0 then + if UVars.AbstractContext.is_constant (table_key_instance env ik1) then conv_stack env k stk1 stk2 cu else match stk1 , stk2 with diff --git a/kernel/vmbytegen.ml b/kernel/vmbytegen.ml index 8934742c7fc1..cba1212d72f4 100644 --- a/kernel/vmbytegen.ml +++ b/kernel/vmbytegen.ml @@ -141,7 +141,7 @@ type comp_env = { type glob_env = { env : Environ.env; - uinst_len : int; (** Size of the toplevel universe instance *) + uinst_len : int * int; (** Size of the toplevel universe instance *) mutable fun_code : instruction list; (** Code of closures *) } @@ -153,7 +153,7 @@ module Config = struct let stack_safety_margin = 15 end -type argument = ArgLambda of lambda | ArgInstance of Univ.Instance.t +type argument = ArgLambda of lambda | ArgInstance of UVars.Instance.t let empty_fv = { size= 0; fv_rev = []; fv_fwd = FvMap.empty; fv_unv = None } let push_fv d e = { @@ -319,6 +319,17 @@ let pos_instance r sz = in Kenvacc (r.offset + pos) +let is_toplevel_inst env u = + UVars.eq_sizes env.uinst_len (UVars.Instance.length u) + && let qs, us = UVars.Instance.to_array u in + Array.for_all_i (fun i q -> Sorts.Quality.equal q (Sorts.Quality.var i)) 0 qs + && Array.for_all_i (fun i l -> Univ.Level.equal l (Univ.Level.var i)) 0 us + +let is_closed_inst u = + let qs, us = UVars.Instance.to_array u in + Array.for_all (fun q -> Option.is_empty (Sorts.Quality.var_index q)) qs + && Array.for_all (fun l -> Option.is_empty (Univ.Level.var_index l)) us + (*i Examination of the continuation *) (* Discard all instructions up to the next label. *) @@ -546,7 +557,7 @@ let rec compile_lam env cenv lam sz cont = | Lconst (kn,u) -> compile_constant env cenv kn u [||] sz cont | Lind (ind,u) -> - if Univ.Instance.is_empty u then + if UVars.Instance.is_empty u then compile_structured_constant cenv (Const_ind ind) sz cont else comp_app compile_structured_constant (compile_instance env) cenv (Const_ind ind) [|u|] sz cont @@ -557,8 +568,11 @@ let rec compile_lam env cenv lam sz cont = evaluation) [Var 0,...,Var n] with values of [arg0,...,argn] *) let has_var = match s with | Sorts.Set | Sorts.Prop | Sorts.SProp -> false - | Sorts.Type u | Sorts.QSort (_, u) -> + | Sorts.Type u -> Univ.Universe.exists (fun (l, _) -> Option.has_some (Univ.Level.var_index l)) u + | Sorts.QSort (q, u) -> + Option.has_some (Sorts.QVar.var_index q) + || Univ.Universe.exists (fun (l, _) -> Option.has_some (Univ.Level.var_index l)) u in let compile_instance cenv () sz cont = let () = set_max_stack_size cenv sz in @@ -799,37 +813,43 @@ let rec compile_lam env cenv lam sz cont = and compile_get_global env cenv (kn,u) sz cont = let () = set_max_stack_size cenv sz in - if Univ.Instance.is_empty u then + if UVars.Instance.is_empty u then Kgetglobal kn :: cont else comp_app (fun _ _ _ cont -> Kgetglobal kn :: cont) (compile_instance env) cenv () [|u|] sz cont -and compile_instance env cenv u0 sz cont = +and compile_instance env cenv u sz cont = let () = set_max_stack_size cenv sz in - let u = Univ.Instance.to_array u0 in - let len = Array.length u in - let is_id i l = match Univ.Level.var_index l with - | None -> false - | Some j -> Int.equal i j - in - if Int.equal env.uinst_len len && Array.for_all_i is_id 0 u then + if is_toplevel_inst env u then (* Optimization: do not reallocate the same instance *) pos_instance cenv sz :: cont - else if Array.for_all (fun l -> Option.is_empty (Univ.Level.var_index l)) u then + else if is_closed_inst u then (* Optimization: allocate closed instances globally *) - compile_structured_constant cenv (Const_univ_instance u0) sz cont + compile_structured_constant cenv (Const_univ_instance u) sz cont else - let () = set_max_stack_size cenv (sz + len - 1) in + let qs, us = UVars.Instance.to_array u in + let comp_qual cenv q sz cont = match Sorts.Quality.var_index q with + | None -> compile_structured_constant cenv (Const_quality q) sz cont + | Some idx -> pos_instance cenv sz :: Kfield 0 :: Kfield idx :: cont + in let comp_univ cenv l sz cont = match Univ.Level.var_index l with | None -> compile_structured_constant cenv (Const_univ_level l) sz cont - | Some idx -> pos_instance cenv sz :: Kfield idx :: cont + | Some idx -> pos_instance cenv sz :: Kfield 1 :: Kfield idx :: cont in - comp_args comp_univ cenv u sz (Kmakeblock (len, 0) :: cont) + let comp_array comp_val cenv vs sz cont = + if Array.is_empty vs then Kmakeblock (0,0) :: cont + else + let () = set_max_stack_size cenv (sz + Array.length vs - 1) in + comp_args comp_val cenv vs sz (Kmakeblock (Array.length vs, 0) :: cont) + in + let cont = Kmakeblock (2, 0) :: cont in + let cont = comp_array comp_qual cenv qs (sz+1) cont in + comp_array comp_univ cenv us sz (Kpush :: cont) and compile_constant env cenv kn u args sz cont = let () = set_max_stack_size cenv sz in - if Univ.Instance.is_empty u then + if UVars.Instance.is_empty u then (* normal compilation *) comp_app (fun _ _ sz cont -> compile_get_global env cenv (kn,u) sz cont) @@ -847,17 +867,14 @@ and compile_constant env cenv kn u args sz cont = comp_app (fun _ _ _ cont -> Kgetglobal kn :: cont) compile_arg cenv () all sz cont -let is_univ_copy max u = - let u = Univ.Instance.to_array u in - if Array.length u = max then - Array.fold_left_i (fun i acc u -> - if acc then - match Univ.Level.var_index u with - | None -> false - | Some l -> l = i - else false) true u - else - false +let is_univ_copy (maxq,maxu) u = + let qs,us = UVars.Instance.to_array u in + let check_array max var_index a = + Array.length a = max + && Array.for_all_i (fun i x -> Option.equal Int.equal (var_index x) (Some i)) 0 a + in + check_array maxq Sorts.Quality.var_index qs + && check_array maxu Univ.Level.var_index us let dump_bytecode = ref false @@ -870,14 +887,14 @@ let dump_bytecodes init code fvs = prlist_with_sep (fun () -> str "; ") pp_fv_elem fvs ++ fnl ()) -let compile ?universes:(universes=0) env sigma c = +let compile ?universes:(universes=(0,0)) env sigma c = Label.reset_label_counter (); let cont = [Kstop] in let cenv, init_code, fun_code = - if Int.equal universes 0 then + if UVars.eq_sizes universes (0,0) then let lam = lambda_of_constr ~optimize:true env sigma c in let cenv = empty_comp_env () in - let env = { env; fun_code = []; uinst_len = 0 } in + let env = { env; fun_code = []; uinst_len = (0,0) } in let cont = compile_lam env cenv lam 0 cont in let cont = ensure_stack_capacity cenv cont in cenv, cont, env.fun_code @@ -930,7 +947,7 @@ let compile_constant_body ~fail_on_error env univs = function | Undef _ | OpaqueDef _ -> Some BCconstant | Primitive _ -> None | Def body -> - let instance_size = Univ.AbstractContext.size (Declareops.universes_context univs) in + let instance_size = UVars.AbstractContext.size (Declareops.universes_context univs) in let alias = match kind body with | Const (kn',u) when is_univ_copy instance_size u -> diff --git a/kernel/vmbytegen.mli b/kernel/vmbytegen.mli index 01361a049714..f34c593c990a 100644 --- a/kernel/vmbytegen.mli +++ b/kernel/vmbytegen.mli @@ -16,7 +16,7 @@ open Environ (** Should only be used for monomorphic terms *) val compile : - fail_on_error:bool -> ?universes:int -> + fail_on_error:bool -> ?universes:int*int -> env -> Genlambda.evars -> constr -> (to_patch * fv) option (** init, fun, fv *) diff --git a/kernel/vmemitcodes.ml b/kernel/vmemitcodes.ml index 991c91b9eac9..7526ca628c47 100644 --- a/kernel/vmemitcodes.ml +++ b/kernel/vmemitcodes.ml @@ -475,7 +475,7 @@ type to_patch = emitcodes * patches (* Substitution *) let subst_strcst s sc = match sc with - | Const_sort _ | Const_b0 _ | Const_univ_level _ | Const_univ_instance _ + | Const_sort _ | Const_b0 _ | Const_quality _ | Const_univ_level _ | Const_univ_instance _ | Const_val _ | Const_uint _ | Const_float _ | Const_evar _ -> sc | Const_ind ind -> let kn,i = ind in Const_ind (subst_mind s kn, i) diff --git a/kernel/vmvalues.ml b/kernel/vmvalues.ml index 7a8626338995..d4b8e52676a2 100644 --- a/kernel/vmvalues.ml +++ b/kernel/vmvalues.ml @@ -52,8 +52,9 @@ type structured_constant = | Const_ind of inductive | Const_evar of Evar.t | Const_b0 of tag + | Const_quality of Sorts.Quality.t | Const_univ_level of Univ.Level.t - | Const_univ_instance of Univ.Instance.t + | Const_univ_instance of UVars.Instance.t | Const_val of structured_values | Const_uint of Uint63.t | Const_float of Float64.t @@ -104,9 +105,11 @@ let eq_structured_constant c1 c2 = match c1, c2 with | Const_evar _, _ -> false | Const_b0 t1, Const_b0 t2 -> Int.equal t1 t2 | Const_b0 _, _ -> false +| Const_quality q1, Const_quality q2 -> Sorts.Quality.equal q1 q2 +| Const_quality _, _ -> false | Const_univ_level l1 , Const_univ_level l2 -> Univ.Level.equal l1 l2 | Const_univ_level _ , _ -> false -| Const_univ_instance u1 , Const_univ_instance u2 -> Univ.Instance.equal u1 u2 +| Const_univ_instance u1 , Const_univ_instance u2 -> UVars.Instance.equal u1 u2 | Const_univ_instance _ , _ -> false | Const_val v1, Const_val v2 -> eq_structured_values v1 v2 | Const_val _, _ -> false @@ -122,11 +125,12 @@ let hash_structured_constant c = | Const_ind i -> combinesmall 2 (Ind.CanOrd.hash i) | Const_evar e -> combinesmall 3 (Evar.hash e) | Const_b0 t -> combinesmall 4 (Int.hash t) - | Const_univ_level l -> combinesmall 5 (Univ.Level.hash l) - | Const_univ_instance u -> combinesmall 6 (Univ.Instance.hash u) - | Const_val v -> combinesmall 7 (hash_structured_values v) - | Const_uint i -> combinesmall 8 (Uint63.hash i) - | Const_float f -> combinesmall 9 (Float64.hash f) + | Const_quality q -> combinesmall 5 (Sorts.Quality.hash q) + | Const_univ_level l -> combinesmall 6 (Univ.Level.hash l) + | Const_univ_instance u -> combinesmall 7 (UVars.Instance.hash u) + | Const_val v -> combinesmall 8 (hash_structured_values v) + | Const_uint i -> combinesmall 9 (Uint63.hash i) + | Const_float f -> combinesmall 10 (Float64.hash f) let eq_annot_switch asw1 asw2 = let eq_rlc (i1, j1) (i2, j2) = Int.equal i1 i2 && Int.equal j1 j2 in @@ -147,15 +151,17 @@ let pp_sort s = | Prop -> Pp.str "Prop" | Set -> Pp.str "Set" | Type u -> Pp.(str "Type@{" ++ Univ.Universe.raw_pr u ++ str "}") - | QSort (q, u) -> Pp.(str "QSort@{" ++ Sorts.QVar.pr q ++ str ", " ++ Univ.Universe.raw_pr u ++ str "}") + | QSort (q, u) -> + Pp.(str "QSort@{" ++ (Sorts.QVar.raw_pr q) ++ strbrk ", " ++ Univ.Universe.raw_pr u ++ str "}") let pp_struct_const = function | Const_sort s -> pp_sort s | Const_ind (mind, i) -> Pp.(MutInd.print mind ++ str"#" ++ int i) | Const_evar e -> Pp.( str "Evar(" ++ int (Evar.repr e) ++ str ")") | Const_b0 i -> Pp.int i + | Const_quality q -> Sorts.Quality.raw_pr q | Const_univ_level l -> Univ.Level.raw_pr l - | Const_univ_instance u -> Univ.Instance.pr Univ.Level.raw_pr u + | Const_univ_instance u -> UVars.Instance.pr Sorts.QVar.raw_pr Univ.Level.raw_pr u | Const_val _ -> Pp.str "(value)" | Const_uint i -> Pp.str (Uint63.to_string i) | Const_float f -> Pp.str (Float64.to_string f) @@ -298,7 +304,7 @@ let arg args i = (* Destructors ***********************************) (*************************************************) -let uni_instance (v : values) : Univ.Instance.t = Obj.magic v +let uni_instance (v : values) : UVars.Instance.t = Obj.magic v let rec whd_accu a stk = let stk = @@ -313,12 +319,8 @@ let rec whd_accu a stk = let () = assert (Int.equal (nargs args) 1) in let inst = uni_instance (arg args 0) in let s = Obj.obj (Obj.field at 0) in - begin match s with - | Sorts.Type u -> - let u = Univ.subst_instance_universe inst u in - Vaccu (Asort (Sorts.sort_of_univ u), []) - | _ -> assert false - end + let s = UVars.subst_instance_sort inst s in + Vaccu (Asort s, []) | _ -> assert false end | i when i <= max_atom_tag -> @@ -410,6 +412,7 @@ let obj_of_str_const str = | Const_ind ind -> obj_of_atom (Aind ind) | Const_evar e -> obj_of_atom (Aid (EvarKey e)) | Const_b0 tag -> Obj.repr tag + | Const_quality q -> Obj.repr q | Const_univ_level l -> Obj.repr l | Const_univ_instance u -> Obj.repr u | Const_val v -> Obj.repr v diff --git a/kernel/vmvalues.mli b/kernel/vmvalues.mli index 4fe14ca9e59d..e39b0bf47912 100644 --- a/kernel/vmvalues.mli +++ b/kernel/vmvalues.mli @@ -39,8 +39,9 @@ type structured_constant = | Const_ind of inductive | Const_evar of Evar.t | Const_b0 of tag + | Const_quality of Sorts.Quality.t | Const_univ_level of Univ.Level.t - | Const_univ_instance of Univ.Instance.t + | Const_univ_instance of UVars.Instance.t | Const_val of structured_values | Const_uint of Uint63.t | Const_float of Float64.t @@ -145,7 +146,7 @@ external val_of_annot_switch : annot_switch -> values = "%identity" (** Destructors *) val whd_val : values -> kind -val uni_instance : values -> Univ.Instance.t +val uni_instance : values -> UVars.Instance.t (** Arguments *) diff --git a/library/global.mli b/library/global.mli index 01b323810a97..dbc4259d67a9 100644 --- a/library/global.mli +++ b/library/global.mli @@ -43,7 +43,7 @@ val sprop_allowed : unit -> bool val push_named_assum : (Id.t * Constr.types) -> unit val push_named_def : (Id.t * Entries.section_def_entry) -> unit -val push_section_context : Univ.UContext.t -> unit +val push_section_context : UVars.UContext.t -> unit val export_private_constants : Safe_typing.private_constants -> @@ -124,7 +124,7 @@ type indirect_accessor = { val force_proof : indirect_accessor -> Opaqueproof.opaque -> Constr.t * unit Opaqueproof.delayed_universes val body_of_constant : indirect_accessor -> Constant.t -> - (Constr.constr * unit Opaqueproof.delayed_universes * Univ.AbstractContext.t) option + (Constr.constr * unit Opaqueproof.delayed_universes * UVars.AbstractContext.t) option (** Returns the body of the constant if it has any, and the polymorphic context it lives in. For monomorphic constant, the latter is empty, and for polymorphic constants, the term contains De Bruijn universe variables that @@ -132,7 +132,7 @@ val body_of_constant : indirect_accessor -> Constant.t -> val body_of_constant_body : indirect_accessor -> constant_body -> - (Constr.constr * unit Opaqueproof.delayed_universes * Univ.AbstractContext.t) option + (Constr.constr * unit Opaqueproof.delayed_universes * UVars.AbstractContext.t) option (** Same as {!body_of_constant} but on {!constant_body}. *) (** {6 Compiled libraries } *) diff --git a/library/globnames.mli b/library/globnames.mli index 020508008d08..21c9357737a1 100644 --- a/library/globnames.mli +++ b/library/globnames.mli @@ -24,7 +24,7 @@ val destConstRef : GlobRef.t -> Constant.t val destIndRef : GlobRef.t -> inductive val destConstructRef : GlobRef.t -> constructor -val subst_global : substitution -> GlobRef.t -> GlobRef.t * constr Univ.univ_abstracted option +val subst_global : substitution -> GlobRef.t -> GlobRef.t * constr UVars.univ_abstracted option val subst_global_reference : substitution -> GlobRef.t -> GlobRef.t (** {6 Extended global references } *) diff --git a/library/nametab.ml b/library/nametab.ml index 20a69b563745..0f8c3fa0e2ef 100644 --- a/library/nametab.ml +++ b/library/nametab.ml @@ -382,12 +382,7 @@ end module MPDTab = Make(DirPath')(MPEqual) module DirTab = Make(DirPath')(GlobDirRef) -module UnivIdEqual = -struct - type t = Univ.UGlobal.t - let equal = Univ.UGlobal.equal -end -module UnivTab = Make(FullPath)(UnivIdEqual) +module UnivTab = Make(FullPath)(Univ.UGlobal) type univtab = UnivTab.t let the_univtab = Summary.ref ~name:"univtab" (UnivTab.empty : univtab) @@ -404,14 +399,7 @@ type mprevtab = DirPath.t MPmap.t type mptrevtab = full_path MPmap.t -module UnivIdOrdered = -struct - type t = Univ.UGlobal.t - let hash = Univ.UGlobal.hash - let compare = Univ.UGlobal.compare -end - -module UnivIdMap = HMap.Make(UnivIdOrdered) +module UnivIdMap = HMap.Make(Univ.UGlobal) type univrevtab = full_path UnivIdMap.t let the_univrevtab = Summary.ref ~name:"univrevtab" (UnivIdMap.empty : univrevtab) @@ -556,7 +544,7 @@ let shortest_qualid_of_modtype ?loc kn = let shortest_qualid_of_universe ?loc ctx kn = let sp = UnivIdMap.find kn !the_univrevtab in - UnivTab.shortest_qualid_gen ?loc (fun id -> Id.Map.mem id ctx) sp !the_univtab + UnivTab.shortest_qualid_gen ?loc (fun id -> Id.Map.mem id ctx) sp !the_univtab let pr_global_env env ref = try pr_qualid (shortest_qualid_of_global env ref) diff --git a/library/nametab.mli b/library/nametab.mli index bc353bdb3073..381ccfac8318 100644 --- a/library/nametab.mli +++ b/library/nametab.mli @@ -110,8 +110,6 @@ val push_module : visibility -> DirPath.t -> ModPath.t -> unit val push_dir : visibility -> DirPath.t -> GlobDirRef.t -> unit val push_abbreviation : ?deprecated:Deprecation.t -> visibility -> full_path -> Globnames.abbreviation -> unit -module UnivIdMap : CMap.ExtS with type key = Univ.UGlobal.t - val push_universe : visibility -> full_path -> Univ.UGlobal.t -> unit (** Deprecation info *) diff --git a/parsing/g_constr.mlg b/parsing/g_constr.mlg index ca2a4413e3cb..50d54d51c3d0 100644 --- a/parsing/g_constr.mlg +++ b/parsing/g_constr.mlg @@ -81,11 +81,36 @@ let test_array_closing = lk_kw "|" >> lk_kw "]" >> check_no_space end +let test_sort_qvar = + let open Pcoq.Lookahead in + to_entry "test_sort_qvar" begin + lk_ident >> lk_list lk_field >> lk_kw "|" + end + +type univ_level_or_quality = + | SProp | Prop | Set | Type | Anon of Loc.t | Global of Libnames.qualid + +let force_univ_level ?loc = function + | SProp -> UNamed CSProp + | Prop -> UNamed CProp + | Set -> UNamed CSet + | Type -> UAnonymous {rigid=UnivRigid} + | Anon _ -> UAnonymous {rigid=UnivFlexible false} + | Global qid -> UNamed (CType qid) + +let force_quality ?loc = function + | SProp -> CQConstant QSProp + | Prop -> CQConstant QProp + | Set -> CErrors.user_err ?loc Pp.(str "Universe levels cannot be Set.") + | Type -> CQConstant QType + | Anon loc -> CQualVar (CQAnon (Some loc)) + | Global qid -> CQualVar (CQVar qid) + } GRAMMAR EXTEND Gram GLOBAL: binder_constr lconstr constr term - universe_level universe_name sort sort_family + universe_name sort sort_family global constr_pattern cpattern Constr.ident closed_binder open_binders binder binders binders_fixannot record_declaration typeclass_constraint pattern arg type_cstr @@ -111,6 +136,7 @@ GRAMMAR EXTEND Gram | "SProp" -> { UNamed (None, [CSProp, 0]) } | "Type" -> { UAnonymous {rigid=UnivRigid} } | "Type"; "@{"; "_"; "}" -> { UAnonymous {rigid=UnivFlexible false} } + | "Type"; "@{"; test_sort_qvar; q = reference; "|"; u = universe; "}" -> { UNamed (Some (CQVar q), u) } | "Type"; "@{"; u = universe; "}" -> { UNamed (None, u) } ] ] ; sort_family: @@ -280,16 +306,23 @@ GRAMMAR EXTEND Gram | -> { [] } ] ] ; univ_annot: - [ [ "@{"; l = LIST0 universe_level; "}" -> { Some l } + [ [ "@{"; l = LIST0 univ_level_or_quality; l' = OPT [ "|"; l' = LIST0 univ_level_or_quality -> { l' } ]; "}" -> + { + let l, l' = match l' with + | None -> [], List.map (force_univ_level ~loc) l + | Some l' -> List.map (force_quality ~loc) l, List.map (force_univ_level ~loc) l' + in + Some (l,l') + } | -> { None } ] ] ; - universe_level: - [ [ "Set" -> { UNamed CSet } - (* no parsing SProp as a level *) - | "Prop" -> { UNamed CProp } - | "Type" -> { UAnonymous {rigid=UnivRigid} } - | "_" -> { UAnonymous {rigid=UnivFlexible false} } - | id = global -> { UNamed (CType id) } ] ] + univ_level_or_quality: + [ [ "Set" -> { Set } + | "SProp" -> { SProp } + | "Prop" -> { Prop } + | "Type" -> { Type } + | "_" -> { Anon loc } + | id = global -> { Global id } ] ] ; fix_decls: [ [ dcl = fix_decl -> { let (id,_,_,_,_) = dcl.CAst.v in (id,[dcl.CAst.v]) } diff --git a/parsing/pcoq.ml b/parsing/pcoq.ml index b47a930e87bd..ae77b0de73fe 100644 --- a/parsing/pcoq.ml +++ b/parsing/pcoq.ml @@ -343,7 +343,6 @@ module Constr = let ident = Entry.make "ident" let global = Entry.make "global" let universe_name = Entry.make "universe_name" - let universe_level = Entry.make "universe_level" let sort = Entry.make "sort" let sort_family = Entry.make "sort_family" let pattern = Entry.make "pattern" diff --git a/parsing/pcoq.mli b/parsing/pcoq.mli index 9ae1a2d81896..a3f4ccdbf184 100644 --- a/parsing/pcoq.mli +++ b/parsing/pcoq.mli @@ -29,12 +29,14 @@ module Lookahead : sig val to_entry : string -> t -> unit Entry.t val (>>) : t -> t -> t val (<+>) : t -> t -> t + val lk_empty : t val lk_list : t -> t val check_no_space : t val lk_kw : string -> t val lk_kws : string list -> t val lk_nat : t val lk_ident : t + val lk_field : t val lk_name : t val lk_qualid : t val lk_ident_except : string list -> t @@ -183,7 +185,6 @@ module Constr : val ident : Id.t Entry.t val global : qualid Entry.t val universe_name : sort_name_expr Entry.t - val universe_level : univ_level_expr Entry.t val sort : sort_expr Entry.t val sort_family : Sorts.family Entry.t val pattern : cases_pattern_expr Entry.t diff --git a/plugins/cc/ccalgo.ml b/plugins/cc/ccalgo.ml index 6ae9793f33ab..25898e2f425d 100644 --- a/plugins/cc/ccalgo.ml +++ b/plugins/cc/ccalgo.ml @@ -209,9 +209,9 @@ struct let rec drop_univ c = match kind c with - | Const (c,_u) -> mkConstU (c,Univ.Instance.empty) - | Ind (c,_u) -> mkIndU (c,Univ.Instance.empty) - | Construct (c,_u) -> mkConstructU (c,Univ.Instance.empty) + | Const (c,_u) -> mkConstU (c,UVars.Instance.empty) + | Ind (c,_u) -> mkIndU (c,UVars.Instance.empty) + | Construct (c,_u) -> mkConstructU (c,UVars.Instance.empty) | Sort (Type _u) -> mkSort (type1) | _ -> Constr.map drop_univ c @@ -510,10 +510,10 @@ let rec canonize_name c = mkLetIn (na, func b,func t,func ct) | App (ct,l) -> mkApp (func ct,Array.Smart.map func l) - | Proj(p,c) -> + | Proj(p,r,c) -> let p' = Projection.map (fun kn -> MutInd.make1 (MutInd.canonical kn)) p in - (mkProj (p', func c)) + (mkProj (p', r, func c)) | _ -> c (* rebuild a term from a pattern and a substitution *) diff --git a/plugins/cc/cctac.ml b/plugins/cc/cctac.ml index 47752faea732..e4a58c5df5b4 100644 --- a/plugins/cc/cctac.ml +++ b/plugins/cc/cctac.ml @@ -83,7 +83,7 @@ let rec decompose_term env sigma t = let u = EInstance.kind sigma u in let canon_const = Constant.make1 (Constant.canonical c) in ATerm.mkSymb (Constr.mkConstU (canon_const,u)) - | Proj (p, c) -> + | Proj (p, _, c) -> let canon_mind kn = MutInd.make1 (MutInd.canonical kn) in let p' = Projection.map canon_mind p in let c = Retyping.expand_projection env sigma p' c [] in diff --git a/plugins/extraction/extraction.ml b/plugins/extraction/extraction.ml index 371f5b067905..4dac0ba0a577 100644 --- a/plugins/extraction/extraction.ml +++ b/plugins/extraction/extraction.ml @@ -244,6 +244,18 @@ let parse_ind_args si args relmax = | _ -> parse (i+1) (j+1) s) in parse 1 1 si +let check_sort_poly sigma gr u = + let u = EConstr.EInstance.kind sigma u in + let qs, _ = UVars.Instance.to_array u in + if Array.exists (function + | Sorts.Quality.QConstant (QSProp|QProp) -> true + | QConstant QType | QVar _ -> false) + qs + then CErrors.user_err + Pp.(str "Cannot extract nontrivial sort polymorphism" ++ spc() + ++ str "(instantiation of " ++ Nametab.pr_global_env Id.Set.empty gr + ++ spc() ++ str "using Prop or SProp).") + (*S Extraction of a type. *) (* [extract_type env db c args] is used to produce an ML type from the @@ -300,6 +312,7 @@ let rec extract_type env sg db j c args = if Int.equal n' 0 then Tunknown else Tvar n') | Const (kn,u) -> let r = GlobRef.ConstRef kn in + let () = check_sort_poly sg r u in let typ = type_of env sg (EConstr.mkConstU (kn,u)) in (match flag_of_type env sg typ with | (Logic,_) -> assert false (* Cf. logical cases above *) @@ -314,14 +327,15 @@ let rec extract_type env sg db j c args = (* We try to reduce. *) let newc = applistc (get_body lbody) args in extract_type env sg db j newc [])) - | Ind ((kn,i),u) -> + | Ind ((kn,i) as ind,u) -> + let () = check_sort_poly sg (IndRef ind) u in let s = (extract_ind env kn).ind_packets.(i).ip_sign in extract_type_app env sg db (GlobRef.IndRef (kn,i),s) args - | Proj (p,t) -> + | Proj (p,r,t) -> (* Let's try to reduce, if it hasn't already been done. *) if Projection.unfolded p then Tunknown else - extract_type env sg db j (EConstr.mkProj (Projection.unfold p, t)) args + extract_type env sg db j (EConstr.mkProj (Projection.unfold p, r, t)) args | Case _ | Fix _ | CoFix _ -> Tunknown | Evar _ | Meta _ -> Taxiom (* only possible during Show Extraction *) | Var v -> @@ -647,11 +661,13 @@ let rec extract_term env sg mle mlt c args = with NotDefault d -> let mle' = Mlenv.push_std_type mle (Tdummy d) in ast_pop (extract_term env' sg mle' mlt c2 args')) - | Const (kn,_) -> + | Const (kn,u) -> + let () = check_sort_poly sg (ConstRef kn) u in extract_cst_app env sg mle mlt kn args - | Construct (cp,_) -> + | Construct (cp,u) -> + let () = check_sort_poly sg (ConstructRef cp) u in extract_cons_app env sg mle mlt cp args - | Proj (p, c) -> + | Proj (p, _, c) -> let term = Retyping.expand_projection env (Evd.from_env env) p c [] in extract_term env sg mle mlt term args | Rel n -> @@ -1027,6 +1043,15 @@ let extract_fixpoint env sg vkn (fi,ti,ci) = current_fixpoints := []; Dfix (Array.map (fun kn -> GlobRef.ConstRef kn) vkn, terms, types) +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) + (** Because of automatic unboxing the easy way [mk_def c] on the constant body of primitive projections doesn't work. We pretend that they are implemented by matches until someone figures out how @@ -1035,7 +1060,7 @@ let fake_match_projection env p = let ind = Projection.Repr.inductive p in let proj_arg = Projection.Repr.arg p in let mib, mip = Inductive.lookup_mind_specif env ind in - let u = Univ.make_abstract_instance (Declareops.inductive_polymorphic_context mib) in + let u = UVars.make_abstract_instance (Declareops.inductive_polymorphic_context mib) in let indu = mkIndU (ind,u) in let ctx, paramslet = let subst = List.init mib.mind_ntypes (fun i -> mkIndU ((fst ind, mib.mind_ntypes - i - 1), u)) in @@ -1050,7 +1075,7 @@ 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_relevance = relevance_of_projection_repr mib p; ci_pp_info; } in @@ -1067,13 +1092,8 @@ let fake_match_projection env p = let ty = Vars.substl subst (liftn 1 j ty) in if arg != proj_arg then let lab = match na.binder_name with Name id -> Label.of_id id | Anonymous -> assert false in - let proj_relevant = match na.binder_relevance with - | Sorts.Irrelevant -> false - | Sorts.Relevant -> true - | Sorts.RelevanceVar _ -> assert false - in - let kn = Projection.Repr.make ind ~proj_relevant ~proj_npars:mib.mind_nparams ~proj_arg:arg lab in - fold (arg+1) (j+1) (mkProj (Projection.make kn false, mkRel 1)::subst) rem + 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, na.binder_relevance, mkRel 1)::subst) rem else let p = ([|x|], liftn 1 2 ty) in let branch = diff --git a/plugins/funind/gen_principle.ml b/plugins/funind/gen_principle.ml index 9355ef6f5d8e..0b05d83e08b4 100644 --- a/plugins/funind/gen_principle.ml +++ b/plugins/funind/gen_principle.ml @@ -2124,7 +2124,7 @@ let make_graph (f_ref : GlobRef.t) = in let mp = Constant.modpath c in let pstate = - do_generate_principle_aux [(c, Univ.Instance.empty)] error_error false + do_generate_principle_aux [(c, UVars.Instance.empty)] error_error false false expr_list in assert (Option.is_empty pstate); @@ -2225,7 +2225,7 @@ let build_case_scheme fa = in let ind, sf = let ind = (first_fun_kn, funs_indexes) in - ((ind, Univ.Instance.empty) (*FIXME*), prop_sort) + ((ind, UVars.Instance.empty) (*FIXME*), prop_sort) in let sigma, scheme = Indrec.build_case_analysis_scheme_default env sigma ind sf diff --git a/plugins/funind/recdef.ml b/plugins/funind/recdef.ml index 7c0093bd5dfc..3001dbcbd539 100644 --- a/plugins/funind/recdef.ml +++ b/plugins/funind/recdef.ml @@ -289,7 +289,7 @@ let check_not_nested env sigma forbidden e = Array.iter check_not_nested t; check_not_nested def; check_not_nested ty - | Proj (p, c) -> check_not_nested c + | Proj (p, _, c) -> check_not_nested c | Const _ -> () | Ind _ -> () | Construct _ -> () diff --git a/plugins/ltac/extratactics.mlg b/plugins/ltac/extratactics.mlg index 9d86a0d1ea8e..8b5b28ae2a33 100644 --- a/plugins/ltac/extratactics.mlg +++ b/plugins/ltac/extratactics.mlg @@ -228,7 +228,7 @@ let add_rewrite_hint ~locality ~poly bases ort t lcsr = if poly then ctx else (* This is a global universe context that shouldn't be refreshed at every use of the hint, declare it globally. *) - (DeclareUctx.declare_universe_context ~poly:false ctx; + (Global.push_context_set ~strict:true ctx; Univ.ContextSet.empty) in CAst.make ?loc:(Constrexpr_ops.constr_loc ce) ((c, ctx), ort, Option.map (in_gen (rawwit wit_ltac)) t) in diff --git a/plugins/ltac2/tac2core.ml b/plugins/ltac2/tac2core.ml index 64c36c10af54..6634091336c5 100644 --- a/plugins/ltac2/tac2core.ml +++ b/plugins/ltac2/tac2core.ml @@ -103,11 +103,14 @@ let v_blk = Valexpr.make_block let of_relevance = function | Sorts.Relevant -> ValInt 0 | Sorts.Irrelevant -> ValInt 1 - | Sorts.RelevanceVar q -> ValInt 0 (* FIXME ? *) + | Sorts.RelevanceVar q -> ValBlk (0, [|of_ext val_qvar q|]) let to_relevance = function | ValInt 0 -> Sorts.Relevant | ValInt 1 -> Sorts.Irrelevant + | ValBlk (0, [|qvar|]) -> + let qvar = to_ext val_qvar qvar in + Sorts.RelevanceVar qvar | _ -> assert false let relevance = make_repr of_relevance to_relevance @@ -119,12 +122,16 @@ let to_binder b = Tac2ffi.to_ext Tac2ffi.val_binder b let of_instance u = - let u = Univ.Instance.to_array (EConstr.Unsafe.to_instance u) in - Tac2ffi.of_array (fun v -> Tac2ffi.of_ext Tac2ffi.val_univ v) u + let u = UVars.Instance.to_array (EConstr.Unsafe.to_instance u) in + let toqs = Tac2ffi.of_array (fun v -> Tac2ffi.of_ext Tac2ffi.val_quality v) in + let tous = Tac2ffi.of_array (fun v -> Tac2ffi.of_ext Tac2ffi.val_univ v) in + Tac2ffi.of_pair toqs tous u let to_instance u = - let u = Tac2ffi.to_array (fun v -> Tac2ffi.to_ext Tac2ffi.val_univ v) u in - EConstr.EInstance.make (Univ.Instance.of_array u) + let toqs = Tac2ffi.to_array (fun v -> Tac2ffi.to_ext Tac2ffi.val_quality v) in + let tous = Tac2ffi.to_array (fun v -> Tac2ffi.to_ext Tac2ffi.val_univ v) in + let u = Tac2ffi.to_pair toqs tous u in + EConstr.EInstance.make (UVars.Instance.of_array u) let of_rec_declaration (nas, ts, cs) = let binders = Array.map2 (fun na t -> (na, t)) nas ts in @@ -539,9 +546,10 @@ let () = nas; cs; |] - | Proj (p, c) -> + | Proj (p, r, c) -> v_blk 16 [| Tac2ffi.of_ext Tac2ffi.val_projection p; + of_relevance r; Tac2ffi.of_constr c; |] | Int n -> @@ -625,10 +633,11 @@ let () = let i = Tac2ffi.to_int i in let def = to_rec_declaration (nas, cs) in EConstr.mkCoFix (i, def) - | (16, [|p; c|]) -> + | (16, [|p; r; c|]) -> let p = Tac2ffi.to_ext Tac2ffi.val_projection p in + let r = to_relevance r in let c = Tac2ffi.to_constr c in - EConstr.mkProj (p, c) + EConstr.mkProj (p, r, c) | (17, [|n|]) -> let n = Tac2ffi.to_uint63 n in EConstr.mkInt n @@ -1212,7 +1221,7 @@ let () = define "ind_get_projections" (repr_ext val_ind_data @-> ret (option (array projection))) @@ fun (ind,mib) -> Declareops.inductive_make_projections ind mib - |> Option.map (Array.map (fun p -> Projection.make p false)) + |> Option.map (Array.map (fun (p,_) -> Projection.make p false)) (** Proj *) diff --git a/plugins/ltac2/tac2ffi.ml b/plugins/ltac2/tac2ffi.ml index 7fcf429233b5..0635b2e866f8 100644 --- a/plugins/ltac2/tac2ffi.ml +++ b/plugins/ltac2/tac2ffi.ml @@ -107,9 +107,11 @@ let val_inductive = Val.create "inductive" let val_constant = Val.create "constant" let val_constructor = Val.create "constructor" let val_projection = Val.create "projection" +let val_qvar = Val.create "qvar" let val_case = Val.create "case" let val_binder = Val.create "binder" let val_univ = Val.create "universe" +let val_quality = Val.create "quality" let val_free : Names.Id.Set.t Val.tag = Val.create "free" let val_ltac1 : Geninterp.Val.t Val.tag = Val.create "ltac1" let val_uint63 = Val.create "uint63" diff --git a/plugins/ltac2/tac2ffi.mli b/plugins/ltac2/tac2ffi.mli index 552dfd56c152..261f96d27f9b 100644 --- a/plugins/ltac2/tac2ffi.mli +++ b/plugins/ltac2/tac2ffi.mli @@ -218,9 +218,11 @@ val val_inductive : inductive Val.tag val val_constant : Constant.t Val.tag val val_constructor : constructor Val.tag val val_projection : Projection.t Val.tag +val val_qvar : Sorts.QVar.t Val.tag val val_case : Constr.case_info Val.tag val val_binder : (Name.t Context.binder_annot * types) Val.tag val val_univ : Univ.Level.t Val.tag +val val_quality : Sorts.Quality.t Val.tag val val_free : Id.Set.t Val.tag val val_uint63 : Uint63.t Val.tag val val_float : Float64.t Val.tag diff --git a/plugins/ring/ring.ml b/plugins/ring/ring.ml index 3314da08cfc0..06216b3cddbb 100644 --- a/plugins/ring/ring.ml +++ b/plugins/ring/ring.ml @@ -60,7 +60,7 @@ let rec mk_clos_but sigma f_map n t = | None -> mk_atom t and tag_arg sigma f_map n tag c = match tag with -| Eval -> mk_clos (Esubst.subs_id n, Univ.Instance.empty) (EConstr.Unsafe.to_constr c) +| Eval -> mk_clos (Esubst.subs_id n, UVars.Instance.empty) (EConstr.Unsafe.to_constr c) | Prot -> mk_atom c | Rec -> mk_clos_but sigma f_map n c @@ -146,7 +146,7 @@ let decl_constant name univs c = let open Constr in let vars = CVars.universes_of_constr c in let univs = UState.restrict_universe_context ~lbound:(Global.universes_lbound ()) univs vars in - let () = DeclareUctx.declare_universe_context ~poly:false univs in + let () = Global.push_context_set ~strict:true univs in let types = (Typeops.infer (Global.env ()) c).uj_type in let univs = UState.Monomorphic_entry Univ.ContextSet.empty, UnivNames.empty_binders in (* UnsafeMonomorphic: we always do poly:false *) diff --git a/plugins/ssr/ssrequality.ml b/plugins/ssr/ssrequality.ml index ac0cb0859003..351826c18043 100644 --- a/plugins/ssr/ssrequality.ml +++ b/plugins/ssr/ssrequality.ml @@ -245,7 +245,7 @@ let rec get_evalref env sigma c = match EConstr.kind sigma c with | Const (k,_) -> Tacred.EvalConstRef k | App (c', _) -> get_evalref env sigma c' | Cast (c', _, _) -> get_evalref env sigma c' - | Proj(c,_) -> Tacred.EvalConstRef(Projection.constant c) + | Proj(c,_,_) -> Tacred.EvalConstRef(Projection.constant c) | _ -> errorstrm Pp.(str "The term " ++ pr_econstr_pat (Global.env ()) sigma c ++ str " is not unfoldable") (* Strip a pattern generated by a prenex implicit to its constant. *) @@ -258,7 +258,7 @@ let strip_unfold_term _ ((sigma, t) as p) kt = match EConstr.kind sigma t with let same_proj env sigma t1 t2 = match EConstr.kind sigma t1, EConstr.kind sigma t2 with - | Proj(c1,_), Proj(c2, _) -> Environ.QProjection.equal env c1 c2 + | Proj(c1,_,_), Proj(c2, _, _) -> Environ.QProjection.equal env c1 c2 | _ -> false let classify_pattern p = match p with diff --git a/plugins/ssrmatching/ssrmatching.ml b/plugins/ssrmatching/ssrmatching.ml index 1c3a4c542c04..4226215d6d7f 100644 --- a/plugins/ssrmatching/ssrmatching.ml +++ b/plugins/ssrmatching/ssrmatching.ml @@ -276,7 +276,7 @@ let iter_constr_LR sigma f c = match EConstr.kind sigma c with f v; Array.iter f pms; f p; iter_invert f iv; Array.iter (fun (_, c) -> f c) b | Fix (_, (_, t, b)) | CoFix (_, (_, t, b)) -> for i = 0 to Array.length t - 1 do f t.(i); f b.(i) done - | Proj(_,a) -> f a + | Proj(_,_,a) -> f a | Array(_u,t,def,ty) -> Array.iter f t; f def; f ty | (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _ | Construct _ | Int _ | Float _) -> () @@ -337,7 +337,7 @@ let isRigid sigma c = match EConstr.kind sigma c with | Float _ | Array _) -> true | (Rel _ | Var _ | Meta _ | Evar (_, _) | Cast (_, _, _) | LetIn (_, _, _, _) | App (_, _) | Const (_, _) | Ind ((_, _), _) | Construct (((_, _), _), _) - | Proj (_, _)) -> false + | Proj _) -> false let hole_var = mkVar (Id.of_string "_") @@ -394,7 +394,7 @@ let mk_tpattern ?p_origin ?(hack=false) ?(ok = all_ok) ~rigid env t dir p { tpat let np = proj_nparams p in if np = 0 || np > List.length a then KpatConst, f, a else let a1, a2 = List.chop np a in KpatProj p, (applistc f a1), a2 - | Proj (p,arg) -> KpatProj (Projection.constant p), f, a + | Proj (p,_,arg) -> KpatProj (Projection.constant p), f, a | Var _ | Ind _ | Construct _ -> KpatFixed, f, a | Evar (k, _) -> if rigid k then KpatEvar k, f, a else @@ -443,7 +443,7 @@ let nb_cs_proj_args env ise pc f u = | Prod _ -> na Prod_cs | Sort s -> na (Sort_cs (Sorts.family (ESorts.kind ise s))) | Const (c',_) when Environ.QConstant.equal env c' pc -> nargs_of_proj u.up_f - | Proj (c',_) when Environ.QConstant.equal env (Names.Projection.constant c') pc -> nargs_of_proj u.up_f + | Proj (c',_,_) when Environ.QConstant.equal env (Names.Projection.constant c') pc -> nargs_of_proj u.up_f | Var _ | Ind _ | Construct _ | Const _ -> na (Const_cs (fst @@ destRef ise f)) | _ -> -1 with Not_found -> -1 @@ -489,7 +489,7 @@ let filter_upat env sigma i0 f n u fpats = let () = if !i0 < np then i0 := n in (u, np) :: fpats let eq_prim_proj env sigma c t = match EConstr.kind sigma t with - | Proj(p,_) -> Environ.QConstant.equal env (Projection.constant p) c + | Proj(p,_,_) -> Environ.QConstant.equal env (Projection.constant p) c | _ -> false let filter_upat_FO env sigma i0 f n u fpats = @@ -1362,7 +1362,8 @@ let ssrpatterntac _ist arg = let pat = interp_rpattern env sigma0 arg in let (t, uc), concl_x = fill_occ_pattern env sigma0 concl0 pat noindex 1 in - let sigma, tty = Typing.type_of env sigma0 t in + let sigma = Evd.set_universe_context sigma0 uc in + let sigma, tty = Typing.type_of env sigma t in let concl = EConstr.mkLetIn (make_annot (Name (Id.of_string "selected")) Sorts.Relevant, t, tty, concl_x) in Proofview.Unsafe.tclEVARS sigma <*> convert_concl ~cast:false ~check:true concl DEFAULTcast diff --git a/pretyping/cbv.ml b/pretyping/cbv.ml index 686f088eedd8..0ed0d7824916 100644 --- a/pretyping/cbv.ml +++ b/pretyping/cbv.ml @@ -50,9 +50,9 @@ type cbv_value = | 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 + | CONSTR of constructor UVars.puniverses * cbv_value array | PRIMITIVE of CPrimitives.t * pconstant * cbv_value array - | ARRAY of Univ.Instance.t * cbv_value Parray.t * cbv_value + | ARRAY of UVars.Instance.t * cbv_value Parray.t * cbv_value (* type of terms with a hole. This hole can appear only under App or Case. * TOP means the term is considered without context @@ -75,8 +75,8 @@ type cbv_value = and cbv_stack = | TOP | APP of cbv_value list * cbv_stack - | CASE of Univ.Instance.t * constr array * case_return * case_branch array * Constr.case_invert * case_info * cbv_value subs * cbv_stack - | PROJ of Projection.t * cbv_stack + | CASE of UVars.Instance.t * constr array * case_return * case_branch array * Constr.case_invert * case_info * cbv_value subs * cbv_stack + | PROJ of Projection.t * Sorts.relevance * cbv_stack (* les vars pourraient etre des constr, cela permet de retarder les lift: utile ?? *) @@ -148,7 +148,7 @@ let rec stack_concat stk1 stk2 = TOP -> stk2 | APP(v,stk1') -> APP(v,stack_concat stk1' stk2) | CASE(u,pms,c,b,iv,i,s,stk1') -> CASE(u,pms,c,b,iv,i,s,stack_concat stk1' stk2) - | PROJ (p,stk1') -> PROJ (p,stack_concat stk1' stk2) + | PROJ (p,r,stk1') -> PROJ (p,r,stack_concat stk1' stk2) (* merge stacks when there is no shifts in between *) let mkSTACK = function @@ -157,8 +157,8 @@ let mkSTACK = function | v,stk -> STACK(0,v,stk) module KeyTable = Hashtbl.Make(struct - type t = Constant.t Univ.puniverses tableKey - let equal = Names.eq_table_key (eq_pair eq_constant_key Univ.Instance.equal) + type t = Constant.t UVars.puniverses tableKey + let equal = Names.eq_table_key (eq_pair eq_constant_key UVars.Instance.equal) let hash = Names.hash_table_key (fun (c, _) -> Constant.UserOrd.hash c) end) @@ -229,7 +229,7 @@ module VNativeEntries = type elem = cbv_value type args = cbv_value array type evd = unit - type uinstance = Univ.Instance.t + type uinstance = UVars.Instance.t let get = Array.get @@ -260,7 +260,7 @@ module VNativeEntries = let mkBool env b = let (ct,cf) = get_bool_constructors env in - CONSTR(Univ.in_punivs (if b then ct else cf), [||]) + CONSTR(UVars.in_punivs (if b then ct else cf), [||]) let int_ty env = VAL(0, UnsafeMonomorphic.mkConst @@ get_int_type env) @@ -268,91 +268,91 @@ module VNativeEntries = let mkCarry env b e = let (c0,c1) = get_carry_constructors env in - CONSTR(Univ.in_punivs (if b then c1 else c0), [|int_ty env;e|]) + CONSTR(UVars.in_punivs (if b then c1 else c0), [|int_ty env;e|]) let mkIntPair env e1 e2 = let int_ty = int_ty env in let c = get_pair_constructor env in - CONSTR(Univ.in_punivs c, [|int_ty;int_ty;e1;e2|]) + CONSTR(UVars.in_punivs c, [|int_ty;int_ty;e1;e2|]) let mkFloatIntPair env f i = let float_ty = float_ty env in let int_ty = int_ty env in let c = get_pair_constructor env in - CONSTR(Univ.in_punivs c, [|float_ty;int_ty;f;i|]) + CONSTR(UVars.in_punivs c, [|float_ty;int_ty;f;i|]) let mkLt env = let (_eq,lt,_gt) = get_cmp_constructors env in - CONSTR(Univ.in_punivs lt, [||]) + CONSTR(UVars.in_punivs lt, [||]) let mkEq env = let (eq,_lt,_gt) = get_cmp_constructors env in - CONSTR(Univ.in_punivs eq, [||]) + CONSTR(UVars.in_punivs eq, [||]) let mkGt env = let (_eq,_lt,gt) = get_cmp_constructors env in - CONSTR(Univ.in_punivs gt, [||]) + CONSTR(UVars.in_punivs gt, [||]) let mkFLt env = let (_eq,lt,_gt,_nc) = get_f_cmp_constructors env in - CONSTR(Univ.in_punivs lt, [||]) + CONSTR(UVars.in_punivs lt, [||]) let mkFEq env = let (eq,_lt,_gt,_nc) = get_f_cmp_constructors env in - CONSTR(Univ.in_punivs eq, [||]) + CONSTR(UVars.in_punivs eq, [||]) let mkFGt env = let (_eq,_lt,gt,_nc) = get_f_cmp_constructors env in - CONSTR(Univ.in_punivs gt, [||]) + CONSTR(UVars.in_punivs gt, [||]) let mkFNotComparable env = let (_eq,_lt,_gt,nc) = get_f_cmp_constructors env in - CONSTR(Univ.in_punivs nc, [||]) + CONSTR(UVars.in_punivs nc, [||]) let mkPNormal env = let (pNormal,_nNormal,_pSubn,_nSubn,_pZero,_nZero,_pInf,_nInf,_nan) = get_f_class_constructors env in - CONSTR(Univ.in_punivs pNormal, [||]) + CONSTR(UVars.in_punivs pNormal, [||]) let mkNNormal env = let (_pNormal,nNormal,_pSubn,_nSubn,_pZero,_nZero,_pInf,_nInf,_nan) = get_f_class_constructors env in - CONSTR(Univ.in_punivs nNormal, [||]) + CONSTR(UVars.in_punivs nNormal, [||]) let mkPSubn env = let (_pNormal,_nNormal,pSubn,_nSubn,_pZero,_nZero,_pInf,_nInf,_nan) = get_f_class_constructors env in - CONSTR(Univ.in_punivs pSubn, [||]) + CONSTR(UVars.in_punivs pSubn, [||]) let mkNSubn env = let (_pNormal,_nNormal,_pSubn,nSubn,_pZero,_nZero,_pInf,_nInf,_nan) = get_f_class_constructors env in - CONSTR(Univ.in_punivs nSubn, [||]) + CONSTR(UVars.in_punivs nSubn, [||]) let mkPZero env = let (_pNormal,_nNormal,_pSubn,_nSubn,pZero,_nZero,_pInf,_nInf,_nan) = get_f_class_constructors env in - CONSTR(Univ.in_punivs pZero, [||]) + CONSTR(UVars.in_punivs pZero, [||]) let mkNZero env = let (_pNormal,_nNormal,_pSubn,_nSubn,_pZero,nZero,_pInf,_nInf,_nan) = get_f_class_constructors env in - CONSTR(Univ.in_punivs nZero, [||]) + CONSTR(UVars.in_punivs nZero, [||]) let mkPInf env = let (_pNormal,_nNormal,_pSubn,_nSubn,_pZero,_nZero,pInf,_nInf,_nan) = get_f_class_constructors env in - CONSTR(Univ.in_punivs pInf, [||]) + CONSTR(UVars.in_punivs pInf, [||]) let mkNInf env = let (_pNormal,_nNormal,_pSubn,_nSubn,_pZero,_nZero,_pInf,nInf,_nan) = get_f_class_constructors env in - CONSTR(Univ.in_punivs nInf, [||]) + CONSTR(UVars.in_punivs nInf, [||]) let mkNaN env = let (_pNormal,_nNormal,_pSubn,_nSubn,_pZero,_nZero,_pInf,_nInf,nan) = get_f_class_constructors env in - CONSTR(Univ.in_punivs nan, [||]) + CONSTR(UVars.in_punivs nan, [||]) let mkArray env u t ty = ARRAY (u,t,ty) @@ -373,8 +373,8 @@ let rec reify_stack t = function reify_stack (apply_env env @@ mkCase (ci, u, pms, ty, iv, t,br)) st - | PROJ (p, st) -> - reify_stack (mkProj (p, t)) st + | PROJ (p, r, st) -> + reify_stack (mkProj (p, r, t)) st and reify_value = function (* reduction under binders *) | VAL (n,t) -> lift n t @@ -467,14 +467,14 @@ let rec norm_head info env t stack = | Case (ci,u,pms,p,iv,c,v) -> norm_head info env c (CASE(u,pms,p,v,iv,ci,env,stack)) | Cast (ct,_,_) -> norm_head info env ct stack - | Proj (p, c) -> + | Proj (p, r, c) -> let p' = if red_set info.reds (fCONST (Projection.constant p)) && red_set info.reds fBETA then Projection.unfold p else p in - norm_head info env c (PROJ (p', stack)) + norm_head info env c (PROJ (p', r, stack)) (* constants, axioms * the first pattern is CRUCIAL, n=0 happens very often: @@ -622,7 +622,7 @@ and cbv_stack_value info env = function cbv_stack_term info stk env (snd br.(n-1)) (* constructor in a Projection -> IOTA *) - | (CONSTR(((sp,n),u),[||]), APP(args,PROJ(p,stk))) + | (CONSTR(((sp,n),u),[||]), APP(args,PROJ(p,_,stk))) when red_set info.reds fMATCH && Projection.unfolded p -> let arg = List.nth args (Projection.npars p + Projection.arg p) in cbv_stack_value info env (strip_appl arg stk) @@ -716,8 +716,8 @@ let rec apply_stack info t = function (mkCase (ci, u, Array.map (cbv_norm_term info env) pms, map_ctx ty, iv, t, Array.map map_ctx br)) st - | PROJ (p, st) -> - apply_stack info (mkProj (p, t)) st + | PROJ (p, r, st) -> + apply_stack info (mkProj (p, r, t)) st (* performs the reduction on a constr, and returns a constr *) and cbv_norm_term info env t = diff --git a/pretyping/cbv.mli b/pretyping/cbv.mli index d7752720e7df..1443e5160903 100644 --- a/pretyping/cbv.mli +++ b/pretyping/cbv.mli @@ -8,10 +8,8 @@ (* * (see LICENSE file for the text of the license) *) (************************************************************************) -open Names open EConstr open Environ -open Esubst (*********************************************************************** s Call-by-value reduction *) @@ -21,41 +19,3 @@ type cbv_infos val create_cbv_infos : RedFlags.reds -> env -> Evd.evar_map -> cbv_infos val cbv_norm : cbv_infos -> constr -> constr - -(*********************************************************************** - i This is for cbv debug *) - -open Constr - -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 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 - | PRIMITIVE of CPrimitives.t * pconstant * cbv_value array - | ARRAY of Univ.Instance.t * cbv_value Parray.t * cbv_value - -and cbv_stack = - | TOP - | APP of cbv_value list * cbv_stack - | CASE of Univ.Instance.t * constr array * case_return * case_branch array * Constr.case_invert * case_info * cbv_value subs * cbv_stack - | PROJ of Projection.t * cbv_stack - -val shift_value : int -> cbv_value -> cbv_value - -val stack_app : cbv_value list -> cbv_stack -> cbv_stack -val strip_appl : cbv_value -> cbv_stack -> cbv_value * cbv_stack - -(** recursive functions... *) -val cbv_stack_term : cbv_infos -> - cbv_stack -> cbv_value subs -> constr -> cbv_value -val cbv_norm_term : cbv_infos -> cbv_value subs -> constr -> constr -val norm_head : cbv_infos -> - cbv_value subs -> constr -> cbv_stack -> cbv_value * cbv_stack -val apply_stack : cbv_infos -> constr -> cbv_stack -> constr -val cbv_norm_value : cbv_infos -> cbv_value -> constr - -(** End of cbv debug section i*) diff --git a/pretyping/coercion.ml b/pretyping/coercion.ml index bb5eb52c593f..cec9e0a69e54 100644 --- a/pretyping/coercion.ml +++ b/pretyping/coercion.ml @@ -63,12 +63,14 @@ let apply_coercion_args env sigma isproj bo ty arg arg_ty nparams = with Evarconv.UnableToUnify _ -> raise NoCoercion in sigma, params @ [arg], subst1 arg c2 in let bo, args = - match isproj with None -> bo, args | Some p -> + match isproj with + | None -> bo, args + | Some (p,r) -> let npars = Projection.Repr.npars p in let p = Projection.make p false in let args = List.skipn npars args in let hd, tl = match args with hd :: tl -> hd, tl | [] -> assert false in - mkProj (p, hd), tl in + mkProj (p, r, hd), tl in sigma, applist (bo, args), params, typ (* appliquer le chemin de coercions de patterns p *) @@ -359,6 +361,7 @@ type coercion_trace = | IdCoe | PrimProjCoe of { proj : Projection.Repr.t; + relevance : Sorts.relevance; args : econstr list; previous : coercion_trace; } @@ -376,11 +379,11 @@ let empty_coercion_trace = IdCoe let rec reapply_coercions sigma trace c = match trace with | IdCoe -> c | ReplaceCoe c -> c - | PrimProjCoe { proj; args; previous } -> + | PrimProjCoe { proj; relevance; args; previous } -> let c = reapply_coercions sigma previous c in let args = args@[c] in let head, args = match args with [] -> assert false | hd :: tl -> hd, tl in - applist (mkProj (Projection.make proj false, head), args) + applist (mkProj (Projection.make proj false, relevance, head), args) | Coe {head; args; previous} -> let c = reapply_coercions sigma previous c in let args = args@[c] in @@ -394,7 +397,7 @@ let rec reapply_coercions sigma trace c = match trace with let instance_of_global_constr sigma c = match kind sigma c with | Const (_,u) | Ind (_,u) | Construct (_,u) -> EInstance.kind sigma u - | _ -> Univ.Instance.empty + | _ -> UVars.Instance.empty (* Apply coercion path from p to h of type hty; raise NoCoercion if not applicable *) let apply_coercion env sigma p h hty = @@ -405,6 +408,7 @@ let apply_coercion env sigma p h hty = let isproj = i.coe_is_projection in let sigma, c = Evd.fresh_global env sigma i.coe_value in let u = instance_of_global_constr sigma c in + let isproj = Option.map (fun p -> p, Relevanceops.relevance_of_projection_repr env (p,u)) isproj in let typ = EConstr.of_constr (CVars.subst_instance_constr u i.coe_typ) in let sigma, j', args, jty = apply_coercion_args env sigma isproj c typ j jty i.coe_param in @@ -412,9 +416,9 @@ let apply_coercion env sigma p h hty = if isid then trace else match isproj with | None -> Coe {head=c;args;previous=trace} - | Some proj -> - let args = List.skipn (Projection.Repr.npars proj) args in - PrimProjCoe {proj; args; previous=trace } in + | Some (proj,relevance) -> + let args = List.skipn (Projection.Repr.npars proj) args in + PrimProjCoe {proj; args; relevance; previous=trace } in (if isid then j else j'), jty, trace, sigma) (h, hty, IdCoe, sigma) p in sigma, j, jty, trace @@ -471,7 +475,7 @@ type delayed_app_body = { mutable head : constr; mutable rev_args : constr list; args_len : int; - proj : Projection.Repr.t option + proj : (Projection.Repr.t * Sorts.relevance) option } let force_app_body ({head;rev_args} as body) = @@ -489,11 +493,11 @@ let push_arg {head;rev_args;args_len;proj} arg = args_len=args_len+1; proj=None; } - | Some p -> + | Some (p,r) -> let npars = Projection.Repr.npars p in if Int.equal args_len npars then { - head = mkProj (Projection.make p false, arg); + head = mkProj (Projection.make p false, r, arg); rev_args=[]; args_len=0; proj=None; @@ -508,8 +512,8 @@ let push_arg {head;rev_args;args_len;proj} arg = let start_app_body sigma head = let proj = match EConstr.kind sigma head with - | Const (p,_) -> - Structures.PrimitiveProjections.find_opt p + | Const (p,u) -> + Structures.PrimitiveProjections.find_opt_with_relevance (p,u) | _ -> None in {head; rev_args=[]; args_len=0; proj} diff --git a/pretyping/coercionops.ml b/pretyping/coercionops.ml index 9e3f59a1c3ea..106ae896bdf0 100644 --- a/pretyping/coercionops.ml +++ b/pretyping/coercionops.ml @@ -162,7 +162,7 @@ let find_class_type env sigma t = match EConstr.kind sigma t' with | Var id -> CL_SECVAR id, EInstance.empty, args | Const (sp,u) -> CL_CONST sp, u, args - | Proj (p, c) when not (Projection.unfolded p) -> + | Proj (p, _, c) when not (Projection.unfolded p) -> let revparams = let open Inductiveops in let t = Retyping.get_type_of env sigma c in @@ -200,7 +200,7 @@ let subst_cl_typ env subst ct = match ct with if c' == c then ct else (match t with | None -> CL_CONST c' | Some t -> - pi1 (find_class_type env Evd.empty (EConstr.of_constr t.Univ.univ_abstracted_value))) + pi1 (find_class_type env Evd.empty (EConstr.of_constr t.UVars.univ_abstracted_value))) | CL_IND i -> let i' = subst_ind subst i in if i' == i then ct else CL_IND i' diff --git a/pretyping/constr_matching.ml b/pretyping/constr_matching.ml index 34deaf5c95e1..3590503a0d4a 100644 --- a/pretyping/constr_matching.ml +++ b/pretyping/constr_matching.ml @@ -335,7 +335,7 @@ let matches_core env sigma allow_bound_rels Array.fold_left2 (sorec ctx env) subst args1 args22 else (* Might be a projection on the right *) match EConstr.kind sigma c2 with - | Proj (pr, c) -> + | Proj (pr, _, c) -> (try let term = Retyping.expand_projection env sigma pr c (Array.to_list args2) in sorec ctx env subst p term with Retyping.RetypeError _ -> raise PatternMatchingFailure) @@ -343,15 +343,15 @@ let matches_core env sigma allow_bound_rels | PApp (c1,arg1), App (c2,arg2) -> (match c1, EConstr.kind sigma c2 with - | PRef (GlobRef.ConstRef r), Proj (pr,c) + | PRef (GlobRef.ConstRef r), Proj (pr,_,c) when not (Environ.QConstant.equal env r (Projection.constant pr)) -> raise PatternMatchingFailure - | PProj (pr1,c1), Proj (pr,c) -> + | PProj (pr1,c1), Proj (pr,_,c) -> if Environ.QProjection.equal env pr1 pr then try Array.fold_left2 (sorec ctx env) (sorec ctx env subst c1 c) arg1 arg2 with Invalid_argument _ -> raise PatternMatchingFailure else raise PatternMatchingFailure - | _, Proj (pr,c) -> + | _, Proj (pr,_,c) -> (try let term = Retyping.expand_projection env sigma pr c (Array.to_list arg2) in sorec ctx env subst p term with Retyping.RetypeError _ -> raise PatternMatchingFailure) @@ -359,16 +359,16 @@ let matches_core env sigma allow_bound_rels try Array.fold_left2 (sorec ctx env) (sorec ctx env subst c1 c2) arg1 arg2 with Invalid_argument _ -> raise PatternMatchingFailure) - | PApp (PRef (GlobRef.ConstRef c1), _), Proj (pr, c2) + | PApp (PRef (GlobRef.ConstRef c1), _), Proj (pr, _, c2) when not (Environ.QConstant.equal env c1 (Projection.constant pr)) -> raise PatternMatchingFailure - | PApp (c, args), Proj (pr, c2) -> + | PApp (c, args), Proj (pr, _, c2) -> (try let term = Retyping.expand_projection env sigma pr c2 [] in sorec ctx env subst p term with Retyping.RetypeError _ -> raise PatternMatchingFailure) - | PProj (p1,c1), Proj (p2,c2) when Environ.QProjection.equal env p1 p2 -> + | PProj (p1,c1), Proj (p2,_,c2) when Environ.QProjection.equal env p1 p2 -> sorec ctx env subst c1 c2 | PProd (na1,c1,d1), Prod(na2,c2,d2) -> @@ -616,7 +616,7 @@ let sub_match ?(closed=true) env sigma pat c = let env' = push_rec_types recdefs env in let sub = subargs env types @ subargs env' bodies in try_aux sub next_mk_ctx next - | Proj (p,c') -> + | Proj (p,_,c') -> begin match Retyping.expand_projection env sigma p c' [] with | term -> aux env term mk_ctx next | exception Retyping.RetypeError _ -> next () diff --git a/pretyping/detyping.ml b/pretyping/detyping.ml index ff5cd910fa7c..072ee3d6a03a 100644 --- a/pretyping/detyping.ml +++ b/pretyping/detyping.ml @@ -47,7 +47,7 @@ end = struct open CVars open Declarations -open Univ +open UVars open Constr let instantiate_context u subst nas ctx = @@ -78,7 +78,7 @@ let return_clause env sigma ind u params (nas, p) = let realdecls, _ = List.chop mip.mind_nrealdecls mip.mind_arity_ctxt in let self = let args = Context.Rel.instance mkRel 0 mip.mind_arity_ctxt in - let inst = Instance.of_array (Array.init (Instance.length u) Level.var) in + let inst = Instance.(abstract_instance (length u)) in mkApp (mkIndU (ind, inst), args) in let realdecls = LocalAssum (Context.anonR, self) :: realdecls in @@ -201,7 +201,7 @@ let print_universes = ref false let { Goptions.get = print_sort_quality } = Goptions.declare_bool_option_and_ref ~key:["Printing";"Sort";"Qualities"] - ~value:false + ~value:true () (** If true, prints local context of evars, whatever print_arguments *) @@ -759,6 +759,20 @@ let detype_level_name sigma l = | Some id -> GLocalUniv (CAst.make id) | None -> GUniv l +let detype_level sigma l = + UNamed (detype_level_name sigma l) + +let detype_qvar sigma q = + match UState.id_of_qvar (Evd.evar_universe_context sigma) q with + | Some id -> GLocalQVar (CAst.make (Name id)) + | None -> GQVar q + +let detype_quality sigma q = + let open Sorts.Quality in + match q with + | QConstant q -> GQConstant q + | QVar q -> GQualVar (detype_qvar sigma q) + let detype_universe sigma u = List.map (on_fst (detype_level_name sigma)) (Univ.Universe.repr u) @@ -772,7 +786,7 @@ let detype_sort sigma = function else UAnonymous {rigid=UnivRigid}) | QSort (q, u) -> if !print_universes then - let q = if print_sort_quality () then Some q else None in + let q = if print_sort_quality () then Some (detype_qvar sigma q) else None in UNamed (q, detype_universe sigma u) else UAnonymous {rigid=UnivRigid} @@ -781,15 +795,16 @@ type binder_kind = BProd | BLambda | BLetIn (**********************************************************************) (* Main detyping function *) -let detype_level sigma l = - UNamed (detype_level_name sigma l) - let detype_instance sigma l = if not !print_universes then None else let l = EInstance.kind sigma l in - if Univ.Instance.is_empty l then None - else Some (List.map (detype_level sigma) (Array.to_list (Univ.Instance.to_array l))) + if UVars.Instance.is_empty l then None + else + let qs, us = UVars.Instance.to_array l in + let qs = List.map (detype_quality sigma) (Array.to_list qs) in + let us = List.map (detype_level sigma) (Array.to_list us) in + Some (qs, us) let delay (type a) (d : a delay) (f : a delay -> _ -> _ -> _ -> _ -> _ -> a glob_constr_r) flags env avoid sigma t : a glob_constr_g = match d with @@ -839,7 +854,7 @@ and detype_r d flags avoid env sigma t = mkapp (detype d flags avoid env sigma f) (Array.map_to_list (detype d flags avoid env sigma) args) | Const (sp,u) -> GRef (GlobRef.ConstRef sp, detype_instance sigma u) - | Proj (p,c) -> + | Proj (p,_,c) -> if Projection.unfolded p && print_unfolded_primproj_asmatch () then let c = detype d flags avoid env sigma c in let id = Label.to_id @@ Projection.label p in @@ -1132,7 +1147,7 @@ let rec subst_glob_constr env subst = DAst.map (function | None -> GRef (ref', u) | Some t -> let evd = Evd.from_env env in - let t = t.Univ.univ_abstracted_value in (* XXX This seems dangerous *) + let t = t.UVars.univ_abstracted_value in (* XXX This seems dangerous *) DAst.get (detype Now Id.Set.empty env evd (EConstr.of_constr t))) | GSort _ diff --git a/pretyping/evarconv.ml b/pretyping/evarconv.ml index 83d93ac53294..664a7a4f9636 100644 --- a/pretyping/evarconv.ml +++ b/pretyping/evarconv.ml @@ -71,10 +71,10 @@ let coq_unit_judge env sigma = sigma, make_judge c t | None -> sigma, unit_judge_fallback -let unfold_projection env evd ts p c = +let unfold_projection env evd ts p r c = let cst = Projection.constant p in if TransparentState.is_transparent_constant ts cst then - Some (mkProj (Projection.unfold p, c)) + Some (mkProj (Projection.unfold p, r, c)) else None let eval_flexible_term ts env evd c = @@ -96,9 +96,9 @@ let eval_flexible_term ts env evd c = with Not_found -> None) | LetIn (_,b,_,c) -> Some (subst1 b c) | Lambda _ -> Some c - | Proj (p, c) -> + | Proj (p, r, c) -> if Projection.unfolded p then assert false - else unfold_projection env evd ts p c + else unfold_projection env evd ts p r c | _ -> assert false type flex_kind_of_term = @@ -169,7 +169,7 @@ let occur_rigidly flags env evd (evk,_) t = | Reducible -> Reducible) | Construct _ -> Normal false | Ind _ | Sort _ -> Rigid false - | Proj (p, c) -> + | Proj (p, _, c) -> let cst = Projection.constant p in let rigid = not (TransparentState.is_transparent_constant flags.open_ts cst) in if rigid then aux c @@ -250,7 +250,7 @@ let check_conv_record env sigma (t1,sk1) (t2,sk2) = let s = ESorts.kind sigma s in CanonicalSolution.find env sigma (proji, Sort_cs (Sorts.family s)),[] - | Proj (p, c) -> + | Proj (p, _, c) -> CanonicalSolution.find env sigma(proji, Proj_cs (Names.Projection.repr p)), Stack.append_app [|c|] sk2 | _ -> let (c2, _) = try destRef sigma t2 with DestKO -> raise Not_found in @@ -367,14 +367,20 @@ let compare_cumulative_instances pbty evd variances u u' = Success evd | Inr p -> UnifFailure (evd, UnifUnivInconsistency p) +let compare_constructor_instances evd u u' = + match Evarutil.compare_constructor_instances evd u u' with + | Inl evd -> + Success evd + | Inr p -> UnifFailure (evd, UnifUnivInconsistency p) + type application = FullyApplied | NumArgs of int let is_applied o n = match o with FullyApplied -> true | NumArgs m -> Int.equal m n let compare_heads pbty env evd ~nargs term term' = let check_strict evd u u' = - let cstrs = Univ.enforce_eq_instances u u' Univ.Constraints.empty in - try Success (Evd.add_constraints evd cstrs) + let cstrs = UVars.enforce_eq_instances u u' Sorts.QUConstraints.empty in + try Success (Evd.add_quconstraints evd cstrs) with UGraph.UniverseInconsistency p -> UnifFailure (evd, UnifUnivInconsistency p) in match EConstr.kind evd term, EConstr.kind evd term' with @@ -382,7 +388,7 @@ let compare_heads pbty env evd ~nargs term term' = if is_applied nargs 1 && Environ.is_array_type env c then let u = EInstance.kind evd u and u' = EInstance.kind evd u' in - compare_cumulative_instances pbty evd [|Univ.Variance.Irrelevant|] u u' + compare_cumulative_instances pbty evd [|UVars.Variance.Irrelevant|] u u' else let u = EInstance.kind evd u and u' = EInstance.kind evd u' in check_strict evd u u' @@ -416,8 +422,7 @@ let compare_heads pbty env evd ~nargs term term' = let needed = Conversion.constructor_cumulativity_arguments (mind,ind,ctor) in if not (is_applied nargs needed) then check_strict evd u u' - else - Success (compare_constructor_instances evd u u') + else compare_constructor_instances evd u u' end | Construct _, Construct _ -> UnifFailure (evd, NotSameHead) | _, _ -> anomaly (Pp.str "") @@ -460,7 +465,7 @@ let rec ise_stack2 no_app env evd f sk1 sk2 = | Success i' -> ise_rev_stack2 true i' q1 q2 | UnifFailure _ as x -> fail x end - | Stack.Proj (p1)::q1, Stack.Proj (p2)::q2 -> + | Stack.Proj (p1,_)::q1, Stack.Proj (p2,_)::q2 -> if QProjection.Repr.equal env (Projection.repr p1) (Projection.repr p2) then ise_rev_stack2 true i q1 q2 else fail (UnifFailure (i, NotSameHead)) @@ -510,7 +515,7 @@ let rec exact_ise_stack2 env evd f sk1 sk2 = (fun i -> ise_array2 i (fun ii -> f (push_rec_types recdef1 env) ii CONV) bds1 bds2); (fun i -> exact_ise_stack2 env i f a1 a2)] else UnifFailure (i,NotSameHead) - | Stack.Proj (p1)::q1, Stack.Proj (p2)::q2 -> + | Stack.Proj (p1,_)::q1, Stack.Proj (p2,_)::q2 -> if QProjection.Repr.equal env (Projection.repr p1) (Projection.repr p2) then ise_rev_stack2 i q1 q2 else (UnifFailure (i, NotSameHead)) @@ -706,7 +711,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) flags env evd pbty delta] in match EConstr.kind evd termM with - | Proj (p, c) when not (Stack.is_empty skF) -> + | Proj (p, _, c) when not (Stack.is_empty skF) -> (* Might be ?X args = p.c args', and we have to eta-expand the primitive projection if |args| >= |args'|+1. *) let nargsF = Stack.args_size skF and nargsM = Stack.args_size skM in @@ -916,7 +921,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) flags env evd pbty in ise_try evd [f1; f2] - | Proj (p, c), Proj (p', c') when QProjection.Repr.equal env (Projection.repr p) (Projection.repr p') -> + | Proj (p, _, c), Proj (p', _, c') when QProjection.Repr.equal env (Projection.repr p) (Projection.repr p') -> let f1 i = ise_and i [(fun i -> evar_conv_x flags env i CONV c c'); @@ -929,7 +934,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) flags env evd pbty ise_try evd [f1; f2] (* Catch the p.c ~= p c' cases *) - | Proj (p,c), Const (p',u) when QConstant.equal env (Projection.constant p) p' -> + | Proj (p,_,c), Const (p',u) when QConstant.equal env (Projection.constant p) p' -> let res = try Some (destApp evd (Retyping.expand_projection env evd p c [])) with Retyping.RetypeError _ -> None @@ -940,7 +945,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) flags env evd pbty appr2 | None -> UnifFailure (evd,NotSameHead)) - | Const (p,u), Proj (p',c') when QConstant.equal env p (Projection.constant p') -> + | Const (p,u), Proj (p',_,c') when QConstant.equal env p (Projection.constant p') -> let res = try Some (destApp evd (Retyping.expand_projection env evd p' c' [])) with Retyping.RetypeError _ -> None @@ -990,7 +995,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) flags env evd pbty (whd_betaiota_deltazeta_for_iota_state flags.open_ts env i (subst1 b c, args)) | Fix _ -> true (* Partially applied fix can be the result of a whd call *) - | Proj (p, _) -> Projection.unfolded p || Stack.not_purely_applicative args + | Proj (p, _, _) -> Projection.unfolded p || Stack.not_purely_applicative args | Case _ | App _| Cast _ -> assert false in let rhs_is_stuck_and_unnamed () = let applicative_stack = fst (Stack.strip_app sk2) in @@ -1147,7 +1152,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) flags env evd pbty |Some (sk1',sk2'), Success i' -> evar_conv_x flags env i' CONV (Stack.zip i' (term1,sk1')) (Stack.zip i' (term2,sk2')) end - | Proj (p1,c1), Proj(p2,c2) when QProjection.Repr.equal env (Projection.repr p1) (Projection.repr p2) -> + | Proj (p1,_,c1), Proj(p2,_,c2) when QProjection.Repr.equal env (Projection.repr p1) (Projection.repr p2) -> begin match ise_stack2 true env evd (evar_conv_x flags) sk1 sk2 with |_, (UnifFailure _ as x) -> x |None, Success i' -> evar_conv_x flags env i' CONV c1 c2 @@ -1155,7 +1160,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) flags env evd pbty end (* Catch the c.(p) ~= p c' cases *) - | Proj (p1,c1), Const (p2,_) when QConstant.equal env (Projection.constant p1) p2 -> + | Proj (p1,_,c1), Const (p2,_) when QConstant.equal env (Projection.constant p1) p2 -> let c1 = try Some (destApp evd (Retyping.expand_projection env evd p1 c1 [])) with Retyping.RetypeError _ -> None @@ -1166,7 +1171,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) flags env evd pbty | None -> UnifFailure (evd,NotSameHead) end - | Const (p1,_), Proj (p2,c2) when QConstant.equal env p1 (Projection.constant p2) -> + | Const (p1,_), Proj (p2,_,c2) when QConstant.equal env p1 (Projection.constant p2) -> let c2 = try Some (destApp evd (Retyping.expand_projection env evd p2 c2 [])) with Retyping.RetypeError _ -> None @@ -1257,7 +1262,10 @@ and eta_constructor flags env evd ((ind, i), u) sk1 (term2,sk2) = let l1' = List.skipn pars l1 in let l2' = let term = Stack.zip evd (term2,sk2) in - List.map (fun p -> EConstr.mkProj (Projection.make p false, term)) (Array.to_list projs) + List.map (fun (p,r) -> + let r = UVars.subst_instance_relevance (Unsafe.to_instance u) r in + EConstr.mkProj (Projection.make p false, r, term)) + (Array.to_list projs) in let f i t1 t2 = evar_conv_x { flags with with_cs = false } env i CONV t1 t2 in ise_list2 evd f l1' l2' diff --git a/pretyping/evardefine.ml b/pretyping/evardefine.ml index 4f256b9ee2a7..99099260c6ad 100644 --- a/pretyping/evardefine.ml +++ b/pretyping/evardefine.ml @@ -196,7 +196,7 @@ let define_pure_evar_as_array env sigma evk = let s = destSort sigma concl in (* array@{u} ty : Type@{u} <= Type@{s} *) let sigma = Evd.set_leq_sort env sigma s' s in - let ar = Typeops.type_of_array env (Univ.Instance.of_array [|u|]) in + let ar = Typeops.type_of_array env (UVars.Instance.of_array ([||],[|u|])) in let sigma = Evd.define evk (mkApp (EConstr.of_constr ar, [| ty |])) sigma in sigma diff --git a/pretyping/evarsolve.ml b/pretyping/evarsolve.ml index 2ee438f2d142..147089a0bfb6 100644 --- a/pretyping/evarsolve.ml +++ b/pretyping/evarsolve.ml @@ -117,7 +117,10 @@ let refresh_universes ?(status=univ_rigid) ?(onlyalg=false) ?(refreshset=false) (* direction: true for fresh universes lower than the existing ones *) let refresh_sort status ~direction s = let sigma, l = new_univ_level_variable status !evdref in - let s' = Sorts.sort_of_univ @@ Univ.Universe.make l in (* FIXME *) + let s' = match ESorts.kind sigma s with + | QSort (q, _) -> Sorts.qsort q (Univ.Universe.make l) + | _ -> Sorts.sort_of_univ @@ Univ.Universe.make l + in let s' = ESorts.make s' in evdref := sigma; let evd = @@ -136,7 +139,8 @@ let refresh_universes ?(status=univ_rigid) ?(onlyalg=false) ?(refreshset=false) | Some l -> (match Evd.universe_rigidity !evdref l with | UnivRigid -> - if not onlyalg then refresh_sort status ~direction s + if not onlyalg && (not (Univ.Level.is_set l) || (refreshset && not direction)) + then refresh_sort status ~direction s else t | UnivFlexible alg -> (if alg then @@ -306,7 +310,7 @@ let noccur_evar env evd evk c = (match decl with | LocalAssum _ -> () | LocalDef (_,b,_) -> cache := Int.Set.add (i-k) !cache; occur_rec false acc (lift i (EConstr.of_constr b))) - | Proj (p,c) -> occur_rec true acc c + | Proj (p,_,c) -> occur_rec true acc c | _ -> iter_with_full_binders env evd (fun rd (k,env) -> (succ k, push_rel rd env)) (occur_rec check_types) acc c in diff --git a/pretyping/glob_ops.ml b/pretyping/glob_ops.ml index f5c404e28342..2a01cf120a94 100644 --- a/pretyping/glob_ops.ml +++ b/pretyping/glob_ops.ml @@ -44,6 +44,17 @@ let map_glob_decl_left_to_right f (na,k,obd,ty) = let comp2 = f ty in (na,k,comp1,comp2) +let glob_qvar_eq g1 g2 = match g1, g2 with + | GLocalQVar na1, GLocalQVar na2 -> CAst.eq Name.equal na1 na2 + | GQVar q1, GQVar q2 -> Sorts.QVar.equal q1 q2 + | GRawQVar q1, GRawQVar q2 -> Sorts.QVar.equal q1 q2 + | (GLocalQVar _ | GQVar _ | GRawQVar _), _ -> false + +let glob_quality_eq g1 g2 = match g1, g2 with + | GQConstant q1, GQConstant q2 -> Sorts.Quality.Constants.equal q1 q2 + | GQualVar q1, GQualVar q2 -> glob_qvar_eq q1 q2 + | (GQConstant _ | GQualVar _), _ -> false + let glob_sort_name_eq g1 g2 = match g1, g2 with | GSProp, GSProp | GProp, GProp @@ -74,7 +85,7 @@ let glob_sort_gen_eq f u1 u2 = let glob_sort_eq u1 u2 = let eq (q1, l1) (q2, l2) = - Option.equal Sorts.QVar.equal q1 q2 && + Option.equal glob_qvar_eq q1 q2 && List.equal (fun (x,m) (y,n) -> glob_sort_name_eq x y && Int.equal m n) l1 l2 in glob_sort_gen_eq eq u1 u2 @@ -82,6 +93,10 @@ let glob_sort_eq u1 u2 = let glob_level_eq u1 u2 = glob_sort_gen_eq glob_sort_name_eq u1 u2 +let instance_eq (q1,u1) (q2,u2) = + List.equal glob_quality_eq q1 q2 + && List.equal glob_level_eq u1 u2 + let binding_kind_eq bk1 bk2 = match bk1, bk2 with | Explicit, Explicit -> true | NonMaxImplicit, NonMaxImplicit -> true @@ -134,16 +149,16 @@ let fix_kind_eq k1 k2 = match k1, k2 with | GCoFix i1, GCoFix i2 -> Int.equal i1 i2 | (GFix _ | GCoFix _), _ -> false -let instance_eq f (x1,c1) (x2,c2) = +let evar_instance_eq f (x1,c1) (x2,c2) = Id.equal x1.CAst.v x2.CAst.v && f c1 c2 let mk_glob_constr_eq f c1 c2 = match DAst.get c1, DAst.get c2 with | GRef (gr1, u1), GRef (gr2, u2) -> GlobRef.CanOrd.equal gr1 gr2 && - Option.equal (List.equal glob_level_eq) u1 u2 + Option.equal instance_eq u1 u2 | GVar id1, GVar id2 -> Id.equal id1 id2 | GEvar (id1, arg1), GEvar (id2, arg2) -> - Id.equal id1.CAst.v id2.CAst.v && List.equal (instance_eq f) arg1 arg2 + Id.equal id1.CAst.v id2.CAst.v && List.equal (evar_instance_eq f) arg1 arg2 | GPatVar k1, GPatVar k2 -> matching_var_kind_eq k1 k2 | GApp (f1, arg1), GApp (f2, arg2) -> f f1 f2 && List.equal f arg1 arg2 @@ -176,13 +191,13 @@ let mk_glob_constr_eq f c1 c2 = match DAst.get c1, DAst.get c2 with f c1 c2 && Option.equal cast_kind_eq k1 k2 && f t1 t2 | GProj ((cst1, u1), args1, c1), GProj ((cst2, u2), args2, c2) -> GlobRef.(CanOrd.equal (ConstRef cst1) (ConstRef cst2)) && - Option.equal (List.equal glob_level_eq) u1 u2 && + Option.equal instance_eq u1 u2 && List.equal f args1 args2 && f c1 c2 | GInt i1, GInt i2 -> Uint63.equal i1 i2 | GFloat f1, GFloat f2 -> Float64.equal f1 f2 | GArray (u1, t1, def1, ty1), GArray (u2, t2, def2, ty2) -> Array.equal f t1 t2 && f def1 def2 && f ty1 ty2 && - Option.equal (List.equal glob_level_eq) u1 u2 + Option.equal instance_eq u1 u2 | (GRef _ | GVar _ | GEvar _ | GPatVar _ | GApp _ | GLambda _ | GProd _ | GLetIn _ | GCases _ | GLetTuple _ | GIf _ | GRec _ | GSort _ | GHole _ | GGenarg _ | GCast _ | GProj _ | GInt _ | GFloat _ | GArray _), _ -> false diff --git a/pretyping/glob_ops.mli b/pretyping/glob_ops.mli index 7132fb119c76..485635838381 100644 --- a/pretyping/glob_ops.mli +++ b/pretyping/glob_ops.mli @@ -19,6 +19,10 @@ val glob_sort_gen_eq : ('a -> 'a -> bool) -> 'a glob_sort_gen -> 'a glob_sort_ge val glob_sort_eq : Glob_term.glob_sort -> Glob_term.glob_sort -> bool +val glob_qvar_eq : glob_qvar -> glob_qvar -> bool + +val glob_quality_eq : glob_quality -> glob_quality -> bool + val glob_level_eq : Glob_term.glob_level -> Glob_term.glob_level -> bool val cases_pattern_eq : 'a cases_pattern_g -> 'a cases_pattern_g -> bool diff --git a/pretyping/glob_term.mli b/pretyping/glob_term.mli index b0b1854c53f3..81a6964757ec 100644 --- a/pretyping/glob_term.mli +++ b/pretyping/glob_term.mli @@ -22,6 +22,15 @@ type existential_name = Id.t (** Sorts *) +type glob_qvar = + | GLocalQVar of lname + | GQVar of Sorts.QVar.t + | GRawQVar of Sorts.QVar.t (* hack for funind *) + +type glob_quality = + | GQConstant of Sorts.Quality.constant + | GQualVar of glob_qvar + type glob_sort_name = | GSProp (** representation of [SProp] literal *) | GProp (** representation of [Prop] level *) @@ -41,8 +50,10 @@ type 'a glob_sort_gen = (** levels, occurring in universe instances *) type glob_level = glob_sort_name glob_sort_gen +type glob_instance = glob_quality list * glob_level list + (** sort expressions *) -type glob_sort = (Sorts.QVar.t option * (glob_sort_name * int) list) glob_sort_gen +type glob_sort = (glob_qvar option * (glob_sort_name * int) list) glob_sort_gen type glob_constraint = glob_sort_name * Univ.constraint_type * glob_sort_name @@ -67,7 +78,7 @@ type binding_kind = Explicit | MaxImplicit | NonMaxImplicit (** Representation of an internalized (or in other words globalized) term. *) type 'a glob_constr_r = - | GRef of GlobRef.t * glob_level list option + | GRef of GlobRef.t * glob_instance option (** An identifier that represents a reference to an object defined either in the (global) environment or in the (local) context. *) | GVar of Id.t @@ -89,10 +100,10 @@ type 'a glob_constr_r = | GHole of Evar_kinds.t * Namegen.intro_pattern_naming_expr | GGenarg of Genarg.glob_generic_argument | GCast of 'a glob_constr_g * Constr.cast_kind option * 'a glob_constr_g - | GProj of (Constant.t * glob_level list option) * 'a glob_constr_g list * 'a glob_constr_g + | GProj of (Constant.t * glob_instance option) * 'a glob_constr_g list * 'a glob_constr_g | GInt of Uint63.t | GFloat of Float64.t - | GArray of glob_level list option * 'a glob_constr_g array * 'a glob_constr_g * 'a glob_constr_g + | GArray of glob_instance option * 'a glob_constr_g array * 'a glob_constr_g * 'a glob_constr_g and 'a glob_constr_g = ('a glob_constr_r, 'a) DAst.t and 'a glob_decl_g = Name.t * binding_kind * 'a glob_constr_g option * 'a glob_constr_g diff --git a/pretyping/heads.ml b/pretyping/heads.ml index af536f07a6c3..3ae32d50b62e 100644 --- a/pretyping/heads.ml +++ b/pretyping/heads.ml @@ -32,7 +32,7 @@ type head_approximation = | NotImmediatelyComputableHead let rec compute_head_const env sigma cst = - let body = Environ.constant_opt_value_in env (cst,Univ.Instance.empty) in + let body = Environ.constant_opt_value_in env (cst,UVars.Instance.empty) in match body with | None -> RigidHead (RigidParameter cst) | Some c -> kind_of_head env sigma (EConstr.of_constr c) @@ -73,7 +73,7 @@ and kind_of_head env sigma t = | LetIn _ -> assert false | Meta _ | Evar _ -> NotImmediatelyComputableHead | App (c,al) -> aux k (Array.to_list al @ l) c b - | Proj (p,c) -> RigidHead RigidOther + | Proj (p,_,c) -> RigidHead RigidOther | Case (_,_,_,_,_,c,_) -> aux k [] c true | Int _ | Float _ | Array _ -> ConstructorHead diff --git a/pretyping/indrec.ml b/pretyping/indrec.ml index 6e9737fe6f67..d6446cca5b36 100644 --- a/pretyping/indrec.ml +++ b/pretyping/indrec.ml @@ -231,7 +231,10 @@ let mis_make_case_com dep env sigma (ind, u as pind) (mib, mip) kind = let term = mkApp (mkRel 2, Array.map - (fun p -> mkProj (Projection.make p true, mkRel 1)) ps) in + (fun (p,r) -> + let r = UVars.subst_instance_relevance u r in + mkProj (Projection.make p true, r, mkRel 1)) ps) + in if dep then let ty = mkApp (mkRel 3, [| mkRel 1 |]) in sigma, mkCast (term, DEFAULTcast, ty), ty diff --git a/pretyping/inductiveops.ml b/pretyping/inductiveops.ml index 9416d2c95528..135b156fbf7d 100644 --- a/pretyping/inductiveops.ml +++ b/pretyping/inductiveops.ml @@ -11,7 +11,6 @@ open CErrors open Util open Names -open Univ open Term open Constr open Vars @@ -375,6 +374,13 @@ let make_project env sigma ind pred c branches ps = in let branch = branches.(0) in let ctx, br = decompose_lambda_n_decls sigma mip.mind_consnrealdecls.(0) branch in + let _, u = destInd sigma (fst (decompose_app sigma ty)) in + let u = Unsafe.to_instance u in + let mkProj i c = + let p, r = ps.(i) in + let r = UVars.subst_instance_relevance u r in + mkProj (Projection.make p true, r, c) + in let proj = match EConstr.destRel sigma br with | exception DestKO -> None | i -> @@ -388,8 +394,7 @@ let make_project env sigma ind pred c branches ps = None | LocalAssum _ :: ctx -> (* This match is just a projection *) - let p = ps.(Context.Rel.nhyps ctx) in - Some (mkProj (Projection.make p true, c)) + Some (mkProj (Context.Rel.nhyps ctx) c) end in match proj with @@ -400,7 +405,7 @@ let make_project env sigma ind pred c branches ps = (fun decl (i, j, ctx) -> match decl with | LocalAssum (na, ty) -> - let t = mkProj (Projection.make ps.(i) true, mkRel j) in + let t = mkProj i (mkRel j) in (i + 1, j + 1, LocalDef (na, t, Vars.liftn 1 j ty) :: ctx) | LocalDef (na, b, ty) -> (i, j + 1, LocalDef (na, Vars.liftn 1 j b, Vars.liftn 1 j ty) :: ctx)) @@ -498,7 +503,7 @@ let make_arity env sigma dep indf s = let compute_projections env (kn, i as ind) = let open Term in let mib = Environ.lookup_mind kn env in - let u = make_abstract_instance (Declareops.inductive_polymorphic_context mib) in + let u = UVars.make_abstract_instance (Declareops.inductive_polymorphic_context mib) in let x = match mib.mind_record with | NotRecord | FakeRecord -> anomaly Pp.(str "Trying to build primitive projections for a non-primitive record") @@ -539,12 +544,8 @@ let compute_projections env (kn, i as ind) = match na.binder_name with | Name id -> let lab = Label.of_id id in - let proj_relevant = match na.binder_relevance with - | Sorts.Irrelevant -> false - | Sorts.Relevant -> true - | Sorts.RelevanceVar _ -> assert false - in - let kn = Projection.Repr.make ind ~proj_relevant ~proj_npars:mib.mind_nparams ~proj_arg lab in + let proj_relevant = na.binder_relevance in + let kn = Projection.Repr.make ind ~proj_npars:mib.mind_nparams ~proj_arg lab in (* from [params, field1,..,fieldj |- t(params,field1,..,fieldj)] to [params, x:I, field1,..,fieldj |- t(params,field1,..,fieldj] *) let t = liftn 1 j t in @@ -553,8 +554,8 @@ let compute_projections env (kn, i as ind) = (* from [params, x:I, field1,..,fieldj |- t(field1,..,fieldj)] to [params, x:I |- t(proj1 x,..,projj x)] *) let ty = substl subst t in - let term = mkProj (Projection.make kn true, mkRel 1) in - let fterm = mkProj (Projection.make kn false, mkRel 1) in + let term = mkProj (Projection.make kn true, proj_relevant, mkRel 1) in + let fterm = mkProj (Projection.make kn false, proj_relevant, mkRel 1) in let etab = it_mkLambda_or_LetIn (mkLambda (x, indty, term)) params in let etat = it_mkProd_or_LetIn (mkProd (x, indty, ty)) params in let body = (etab, etat) in @@ -633,7 +634,7 @@ let arity_of_case_predicate env (ind,params) dep k = let univ_level_mem l s = match s with | Prop | Set | SProp -> false -| Type u -> univ_level_mem l u +| Type u -> Univ.univ_level_mem l u | QSort (_, u) -> assert false (* template cannot contain sort variables *) (* Compute the inductive argument types: replace the sorts @@ -681,7 +682,7 @@ let type_of_inductive_knowing_conclusion env sigma ((mib,mip),u) conclty = !evdref, EConstr.of_constr (mkArity (List.rev ctx,scl)) let type_of_projection_constant env (p,u) = - let pty = lookup_projection p env in + let _, pty = lookup_projection p env in Vars.subst_instance_constr u pty let type_of_projection_knowing_arg env sigma p c ty = diff --git a/pretyping/inductiveops.mli b/pretyping/inductiveops.mli index 0331300680d0..aea1feade93e 100644 --- a/pretyping/inductiveops.mli +++ b/pretyping/inductiveops.mli @@ -35,8 +35,8 @@ val arities_of_constructors : env -> pinductive -> types array reasoning either with only recursively uniform parameters or with all parameters including the recursively non-uniform ones *) type inductive_family -val make_ind_family : inductive Univ.puniverses * constr list -> inductive_family -val dest_ind_family : inductive_family -> inductive Univ.puniverses * constr list +val make_ind_family : inductive UVars.puniverses * constr list -> inductive_family +val dest_ind_family : inductive_family -> inductive UVars.puniverses * constr list val map_ind_family : (constr -> constr) -> inductive_family -> inductive_family val liftn_inductive_family : int -> int -> inductive_family -> inductive_family val lift_inductive_family : int -> inductive_family -> inductive_family @@ -237,7 +237,7 @@ val compute_projections : Environ.env -> inductive -> (constr * types) array (********************) val type_of_inductive_knowing_conclusion : - env -> evar_map -> mind_specif Univ.puniverses -> EConstr.types -> evar_map * EConstr.types + env -> evar_map -> mind_specif UVars.puniverses -> EConstr.types -> evar_map * EConstr.types (********************) val control_only_guard : env -> Evd.evar_map -> EConstr.types -> unit diff --git a/pretyping/keys.ml b/pretyping/keys.ml index 54957f34aefc..1c7593cef04d 100644 --- a/pretyping/keys.ml +++ b/pretyping/keys.ml @@ -126,7 +126,7 @@ let constr_key kind c = | Construct (c,u) -> KGlob (GlobRef.ConstructRef c) | Var id -> KGlob (GlobRef.VarRef id) | App (f, _) -> aux f - | Proj (p, _) -> KGlob (GlobRef.ConstRef (Projection.constant p)) + | Proj (p, _, _) -> KGlob (GlobRef.ConstRef (Projection.constant p)) | Cast (p, _, _) -> aux p | Lambda _ -> KLam | Prod _ -> KProd diff --git a/pretyping/nativenorm.ml b/pretyping/nativenorm.ml index 21bbf4b9b148..901164f0fe37 100644 --- a/pretyping/nativenorm.ml +++ b/pretyping/nativenorm.ml @@ -185,8 +185,8 @@ let branch_of_switch lvl ans bs = Array.init (Array.length tbl) branch let get_proj env (ind, proj_arg) = - let p = Environ.get_projection env ind ~proj_arg in - Projection.make p true + let p, r = Environ.get_projection env ind ~proj_arg in + Projection.make p true, r let rec nf_val env sigma v typ = match kind_of_value v with @@ -294,9 +294,11 @@ and nf_atom env sigma atom = | Asort s -> mkSort s | Avar id -> mkVar id | Aproj (p, c) -> - let c = nf_accu env sigma c in - let p = get_proj env p in - mkProj(p, c) + let c, cty = nf_accu_type env sigma c in + let p, r = get_proj env p in + let (_, u), _ = find_rectype_a env sigma (EConstr.of_constr cty) in + let r = UVars.subst_instance_relevance u r in + mkProj(p, r, c) | _ -> fst (nf_atom_type env sigma atom) and nf_atom_type env sigma atom = @@ -372,7 +374,7 @@ and nf_atom_type env sigma atom = | Aproj(p,c) -> let c,tc = nf_accu_type env sigma c in let cj = make_judge c tc in - let p = get_proj env p in + let p, _ = get_proj env p in let uj = Typeops.judge_of_projection env p cj in uj.uj_val, uj.uj_type diff --git a/pretyping/patternops.ml b/pretyping/patternops.ml index 37543902bd10..ac2283d93db7 100644 --- a/pretyping/patternops.ml +++ b/pretyping/patternops.ml @@ -169,7 +169,7 @@ let pattern_of_constr ~broken env sigma t = | Const (sp,u) -> PRef (GlobRef.ConstRef (Constant.make1 (Constant.canonical sp))) | Ind (sp,u) -> PRef (canonical_gr (GlobRef.IndRef sp)) | Construct (sp,u) -> PRef (canonical_gr (GlobRef.ConstructRef sp)) - | Proj (p, c) -> + | Proj (p, _, c) -> pattern_of_constr env (EConstr.Unsafe.to_constr (Retyping.expand_projection env sigma p (EConstr.of_constr c) [])) | Evar (evk,ctxt as ev) -> let EvarInfo evi = Evd.find sigma evk in @@ -265,7 +265,7 @@ let rec subst_pattern env sigma subst pat = if ref' == ref then pat else (match t with | None -> PRef ref' | Some t -> - pattern_of_constr env sigma (EConstr.of_constr t.Univ.univ_abstracted_value)) + pattern_of_constr env sigma (EConstr.of_constr t.UVars.univ_abstracted_value)) | PVar _ | PEvar _ | PRel _ diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml index fc0e55634aca..253bcef7f2ae 100644 --- a/pretyping/pretyping.ml +++ b/pretyping/pretyping.ml @@ -149,12 +149,52 @@ let level_name sigma = function let sigma, u = universe_level_name sigma l in Some (sigma, u) -let sort_info ?loc sigma l = match l with +let glob_level ?loc evd : glob_level -> _ = function + | UAnonymous {rigid} -> + assert (rigid <> UnivFlexible true); + new_univ_level_variable ?loc rigid evd + | UNamed s -> + match level_name evd s with + | None -> + user_err ?loc + (str "Universe instances cannot contain non-Set small levels," ++ spc() ++ + str "polymorphic universe instances must be greater or equal to Set."); + | Some r -> r + +let glob_qvar ?loc evd : glob_qvar -> _ = function + | GQVar q -> evd, q + | GLocalQVar {v=Anonymous} -> + let evd, q = new_quality_variable ?loc evd in + evd, q + | GRawQVar q -> + let evd = Evd.merge_sort_variables ~sideff:true evd (Sorts.QVar.Set.singleton q) in + evd, q + | GLocalQVar {v=Name id; loc} -> + try evd, (Evd.quality_of_name evd id) + with Not_found -> + if not (is_strict_universe_declarations()) then + let evd, q = new_quality_variable ?loc evd in + evd, q + else user_err ?loc Pp.(str "Undeclared quality: " ++ Id.print id ++ str".") + +let glob_quality ?loc evd = let open Sorts.Quality in function + | GQConstant q -> evd, QConstant q + | GQualVar (GQVar _ | GLocalQVar _ | GRawQVar _ as q) -> + let evd, q = glob_qvar ?loc evd q in + evd, QVar q + +let sort_info ?loc sigma q l = match l with | [] -> assert false -| [GSProp, 0] -> sigma, Sorts.sprop -| [GProp, 0] -> sigma, Sorts.prop +| [GSProp, 0] -> assert (Option.is_empty q); sigma, Sorts.sprop +| [GProp, 0] -> assert (Option.is_empty q); sigma, Sorts.prop | (u, n) :: us -> let open Pp in + let sigma, q = match q with + | None -> sigma, None + | Some q -> + let sigma, q = glob_qvar ?loc sigma q in + sigma, Some q + in let get_level sigma u n = match level_name sigma u with | None -> user_err ?loc @@ -176,7 +216,11 @@ let sort_info ?loc sigma l = match l with in let (sigma, u) = get_level sigma u n in let (sigma, u) = List.fold_left fold (sigma, u) us in - sigma, Sorts.sort_of_univ u + let s = match q with + | None -> Sorts.sort_of_univ u + | Some q -> Sorts.qsort q u + in + sigma, s type inference_hook = env -> evar_map -> Evar.t -> (evar_map * constr) option @@ -421,27 +465,22 @@ let pretype_id pretype loc env sigma id = (*************************************************************************) (* Main pretyping function *) -let glob_level ?loc evd : glob_level -> _ = function - | UAnonymous {rigid} -> - assert (rigid <> UnivFlexible true); - new_univ_level_variable ?loc rigid evd - | UNamed s -> - match level_name evd s with - | None -> - user_err ?loc - (str "Universe instances cannot contain non-Set small levels, polymorphic" ++ - str " universe instances must be greater or equal to Set."); - | Some r -> r - -let instance ?loc evd l = - let evd, l' = +let instance ?loc evd (ql,ul) = + let evd, ql' = + List.fold_left + (fun (evd, quals) l -> + let evd, l = glob_quality ?loc evd l in + (evd, l :: quals)) (evd, []) + ql + in + let evd, ul' = List.fold_left (fun (evd, univs) l -> let evd, l = glob_level ?loc evd l in (evd, l :: univs)) (evd, []) - l + ul in - evd, Some (Univ.Instance.of_array (Array.of_list (List.rev l'))) + evd, Some (UVars.Instance.of_array (Array.rev_of_list ql', Array.rev_of_list ul')) let pretype_global ?loc rigid env evd gr us = let evd, instance = @@ -458,12 +497,12 @@ let pretype_ref ?loc sigma env ref us = (try let ty = NamedDecl.get_type (lookup_named id !!env) in (match us with - | None | Some [] -> () - | Some us -> + | None | Some ([],[]) -> () + | Some (qs,us) -> let open UnivGen in Loc.raise ?loc (UniverseLengthMismatch { - actual = List.length us; - expect = 0; + actual = List.length qs, List.length us; + expect = 0, 0; })); sigma, make_judge (mkVar id) ty with Not_found -> @@ -481,9 +520,7 @@ let sort ?loc evd : glob_sort -> _ = function let evd, l = new_univ_level_variable ?loc rigid evd in evd, ESorts.make (Sorts.sort_of_univ (Univ.Universe.make l)) | UNamed (q, l) -> - (* No user-facing syntax for qualities *) - let () = assert (Option.is_empty q) in - let evd, s = sort_info ?loc evd l in + let evd, s = sort_info ?loc evd q l in evd, ESorts.make s let judge_of_sort ?loc evd s = @@ -525,12 +562,12 @@ let mark_obligation_evar sigma k evc = type 'a pretype_fun = ?loc:Loc.t -> flags:pretype_flags -> type_constraint -> GlobEnv.t -> evar_map -> evar_map * 'a type pretyper = { - pretype_ref : pretyper -> GlobRef.t * glob_level list option -> unsafe_judgment pretype_fun; + pretype_ref : pretyper -> GlobRef.t * glob_instance option -> unsafe_judgment pretype_fun; pretype_var : pretyper -> Id.t -> unsafe_judgment pretype_fun; pretype_evar : pretyper -> existential_name CAst.t * (lident * glob_constr) list -> unsafe_judgment pretype_fun; pretype_patvar : pretyper -> Evar_kinds.matching_var_kind -> unsafe_judgment pretype_fun; pretype_app : pretyper -> glob_constr * glob_constr list -> unsafe_judgment pretype_fun; - pretype_proj : pretyper -> (Constant.t * glob_level list option) * glob_constr list * glob_constr -> unsafe_judgment pretype_fun; + pretype_proj : pretyper -> (Constant.t * glob_instance option) * glob_constr list * glob_constr -> unsafe_judgment pretype_fun; pretype_lambda : pretyper -> Name.t * binding_kind * glob_constr * glob_constr -> unsafe_judgment pretype_fun; pretype_prod : pretyper -> Name.t * binding_kind * glob_constr * glob_constr -> unsafe_judgment pretype_fun; pretype_letin : pretyper -> Name.t * glob_constr * glob_constr option * glob_constr -> unsafe_judgment pretype_fun; @@ -544,7 +581,7 @@ type pretyper = { pretype_cast : pretyper -> glob_constr * cast_kind option * glob_constr -> unsafe_judgment pretype_fun; pretype_int : pretyper -> Uint63.t -> unsafe_judgment pretype_fun; pretype_float : pretyper -> Float64.t -> unsafe_judgment pretype_fun; - pretype_array : pretyper -> glob_level list option * glob_constr array * glob_constr * glob_constr -> unsafe_judgment pretype_fun; + pretype_array : pretyper -> glob_instance option * glob_constr array * glob_constr * glob_constr -> unsafe_judgment pretype_fun; pretype_type : pretyper -> glob_constr -> unsafe_type_judgment pretype_fun; } @@ -1089,9 +1126,9 @@ struct match names, l with | na :: names, (LocalAssum (na', t) :: l) -> let t = EConstr.of_constr t in - let proj = Projection.make ps.(cs.cs_nargs - k) true in + let proj = Projection.make (fst ps.(cs.cs_nargs - k)) true in LocalDef ({na' with binder_name = na}, - lift (cs.cs_nargs - n) (mkProj (proj, cj.uj_val)), t) + lift (cs.cs_nargs - n) (mkProj (proj, na'.binder_relevance, 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 @@ -1331,14 +1368,14 @@ let pretype_type self c ?loc ~flags valcon (env : GlobEnv.t) sigma = match DAst. fun ?loc ~flags tycon env sigma -> let sigma, u = match u with | None -> sigma, None - | Some [u] -> + | Some ([],[u]) -> let sigma, u = glob_level ?loc sigma u in sigma, Some u - | Some u -> + | Some (qs,us) -> let open UnivGen in Loc.raise ?loc (UniverseLengthMismatch { - actual = List.length u; - expect = 1; + actual = List.length qs, List.length us; + expect = 0, 1; }) in let sigma, tycon' = split_as_array !!env sigma tycon in @@ -1351,7 +1388,7 @@ let pretype_type self c ?loc ~flags valcon (env : GlobEnv.t) sigma = match DAst. let sigma, jdef = eval_pretyper self ~flags (mk_tycon jty.utj_val) env sigma def in let pretype_elem = eval_pretyper self ~flags (mk_tycon jty.utj_val) env in let sigma, jt = Array.fold_left_map pretype_elem sigma t in - let u = Univ.Instance.of_array [| u |] in + let u = UVars.Instance.of_array ([||],[| u |]) in let ta = EConstr.of_constr @@ Typeops.type_of_array !!env u in let j = { uj_val = EConstr.mkArray(EInstance.make u, Array.map (fun j -> j.uj_val) jt, jdef.uj_val, jty.utj_val); diff --git a/pretyping/pretyping.mli b/pretyping/pretyping.mli index d34472af26f4..817928d5bfe8 100644 --- a/pretyping/pretyping.mli +++ b/pretyping/pretyping.mli @@ -158,12 +158,12 @@ type pretype_flags = { type 'a pretype_fun = ?loc:Loc.t -> flags:pretype_flags -> Evardefine.type_constraint -> GlobEnv.t -> evar_map -> evar_map * 'a type pretyper = { - pretype_ref : pretyper -> GlobRef.t * glob_level list option -> unsafe_judgment pretype_fun; + pretype_ref : pretyper -> GlobRef.t * glob_instance option -> unsafe_judgment pretype_fun; pretype_var : pretyper -> Id.t -> unsafe_judgment pretype_fun; pretype_evar : pretyper -> existential_name CAst.t * (lident * glob_constr) list -> unsafe_judgment pretype_fun; pretype_patvar : pretyper -> Evar_kinds.matching_var_kind -> unsafe_judgment pretype_fun; pretype_app : pretyper -> glob_constr * glob_constr list -> unsafe_judgment pretype_fun; - pretype_proj : pretyper -> (Constant.t * glob_level list option) * glob_constr list * glob_constr -> unsafe_judgment pretype_fun; + pretype_proj : pretyper -> (Constant.t * glob_instance option) * glob_constr list * glob_constr -> unsafe_judgment pretype_fun; pretype_lambda : pretyper -> Name.t * binding_kind * glob_constr * glob_constr -> unsafe_judgment pretype_fun; pretype_prod : pretyper -> Name.t * binding_kind * glob_constr * glob_constr -> unsafe_judgment pretype_fun; pretype_letin : pretyper -> Name.t * glob_constr * glob_constr option * glob_constr -> unsafe_judgment pretype_fun; @@ -177,7 +177,7 @@ type pretyper = { pretype_cast : pretyper -> glob_constr * Constr.cast_kind option * glob_constr -> unsafe_judgment pretype_fun; pretype_int : pretyper -> Uint63.t -> unsafe_judgment pretype_fun; pretype_float : pretyper -> Float64.t -> unsafe_judgment pretype_fun; - pretype_array : pretyper -> glob_level list option * glob_constr array * glob_constr * glob_constr -> unsafe_judgment pretype_fun; + pretype_array : pretyper -> glob_instance option * glob_constr array * glob_constr * glob_constr -> unsafe_judgment pretype_fun; pretype_type : pretyper -> glob_constr -> unsafe_type_judgment pretype_fun; } (** Type of pretyping algorithms in open-recursion style. A typical way to diff --git a/pretyping/reductionops.ml b/pretyping/reductionops.ml index ab4f764c3b9d..056e4d7fc43b 100644 --- a/pretyping/reductionops.ml +++ b/pretyping/reductionops.ml @@ -188,7 +188,7 @@ sig type member = | App of app_node | Case of case_stk - | Proj of Projection.t + | Proj of Projection.t * Sorts.relevance | Fix of fixpoint * t | Primitive of CPrimitives.t * (Constant.t * EInstance.t) * t * CPrimitives.args_red @@ -241,7 +241,7 @@ struct type member = | App of app_node | Case of case_stk - | Proj of Projection.t + | Proj of Projection.t * Sorts.relevance | Fix of fixpoint * t | Primitive of CPrimitives.t * (Constant.t * EInstance.t) * t * CPrimitives.args_red @@ -257,7 +257,7 @@ struct str "ZCase(" ++ prvect_with_sep (pr_bar) (fun (_, c) -> pr_c c) br ++ str ")" - | Proj p -> + | Proj (p,_) -> str "ZProj(" ++ Constant.debug_print (Projection.constant p) ++ str ")" | Fix (f,args) -> str "ZFix(" ++ Constr.debug_print_fix pr_c f @@ -294,7 +294,7 @@ struct | (_, App (i,_,j)::s2) -> compare_rec (bal - j - 1 + i) stk1 s2 | (Case _ :: s1, Case _::s2) -> Int.equal bal 0 (* && c1.ci_ind = c2.ci_ind *) && compare_rec 0 s1 s2 - | (Proj (p)::s1, Proj(p2)::s2) -> + | (Proj (p,_)::s1, Proj(p2,_)::s2) -> Int.equal bal 0 && compare_rec 0 s1 s2 | (Fix(_,a1)::s1, Fix(_,a2)::s2) -> Int.equal bal 0 && compare_rec 0 a1 a2 && compare_rec 0 s1 s2 @@ -315,7 +315,7 @@ struct | Case ((_,_,pms1,(_, t1),_,a1)) :: q1, Case ((_,_,pms2, (_, t2),_,a2)) :: q2 -> let f' o (_, t1) (_, t2) = f o t1 t2 in aux (Array.fold_left2 f' (f (Array.fold_left2 f o pms1 pms2) t1 t2) a1 a2) q1 q2 - | Proj (p1) :: q1, Proj (p2) :: q2 -> + | Proj (p1,_) :: q1, Proj (p2,_) :: q2 -> aux o q1 q2 | Fix ((_,(_,a1,b1)),s1) :: q1, Fix ((_,(_,a2,b2)),s2) :: q2 -> let o' = aux (Array.fold_left2 f (Array.fold_left2 f o b1 b2) a1 a2) (List.rev s1) (List.rev s2) in @@ -402,7 +402,7 @@ struct | f, (Case (ci,u,pms,rt,iv,br)::s) -> zip (mkCase (ci,u,pms,rt,iv,f,br), s) | f, (Fix (fix,st)::s) -> zip (mkFix fix, st @ (append_app [|f|] s)) - | f, (Proj (p)::s) -> zip (mkProj (p,f),s) + | f, (Proj (p,r)::s) -> zip (mkProj (p,r,f),s) | f, (Primitive (p,c,args,kargs)::s) -> zip (mkConstU c, args @ append_app [|f|] s) in @@ -731,8 +731,8 @@ let whd_state_gen flags env sigma = whrec (a,Stack.Primitive(p,const,before,kargs)::after) | exception NotEvaluableConst _ -> fold () else fold () - | Proj (p, c) when RedFlags.red_projection flags p -> - let stack' = (c, Stack.Proj (p) :: stack) in + | Proj (p, r, c) when RedFlags.red_projection flags p -> + let stack' = (c, Stack.Proj (p,r) :: stack) in whrec stack' | LetIn (_,b,_,c) when RedFlags.red_set flags RedFlags.fZETA -> @@ -764,7 +764,7 @@ let whd_state_gen flags env sigma = |args, (Stack.Case case::s') when use_match -> let r = apply_branch env sigma cstr args case in whrec (r, s') - |args, (Stack.Proj (p)::s') when use_match -> + |args, (Stack.Proj (p,_)::s') when use_match -> whrec (Stack.nth args (Projection.npars p + Projection.arg p), s') |args, (Stack.Fix (f,s')::s'') when use_fix -> let x' = Stack.zip sigma (x, args) in @@ -828,8 +828,8 @@ let local_whd_state_gen flags env sigma = whrec (apply_subst [a] sigma c m) | _ -> s) - | Proj (p,c) when RedFlags.red_projection flags p -> - (whrec (c, Stack.Proj (p) :: stack)) + | Proj (p,r,c) when RedFlags.red_projection flags p -> + (whrec (c, Stack.Proj (p,r) :: stack)) | Case (ci,u,pms,p,iv,d,lf) -> whrec (d, Stack.Case (ci,u,pms,p,iv,lf) :: stack) @@ -853,7 +853,7 @@ let local_whd_state_gen flags env sigma = |args, (Stack.Case case :: s') when use_match -> let r = apply_branch env sigma cstr args case in whrec (r, s') - |args, (Stack.Proj (p) :: s') when use_match -> + |args, (Stack.Proj (p,_) :: s') when use_match -> whrec (Stack.nth args (Projection.npars p + Projection.arg p), s') |args, (Stack.Fix (f,s')::s'') when use_fix -> let x' = Stack.zip sigma (x,args) in @@ -984,7 +984,7 @@ let clos_norm_flags flgs env sigma t = EConstr.of_constr (CClosure.norm_term (Evarutil.create_clos_infos env sigma flgs) (CClosure.create_tab ()) - (Esubst.subs_id 0, Univ.Instance.empty) (EConstr.Unsafe.to_constr t)) + (Esubst.subs_id 0, UVars.Instance.empty) (EConstr.Unsafe.to_constr t)) with e when is_anomaly e -> user_err Pp.(str "Tried to normalize ill-typed term") let clos_whd_flags flgs env sigma t = @@ -1038,16 +1038,13 @@ let checked_sort_cmp_universes _env pb s0 s1 univs = check_sort_cmp_universes pb s0 s1 univs; univs let check_convert_instances ~flex:_ u u' univs = - let u = Instance.to_array u in - let u' = Instance.to_array u' in - let fold accu l1 l2 = Constraints.add (l1, Eq, l2) accu in - let cst = Array.fold_left2 fold Constraints.empty u u' in - if Evd.check_constraints univs cst then univs else raise NotConvertible + let csts = UVars.enforce_eq_instances u u' (Sorts.QConstraints.empty,Constraints.empty) in + if Evd.check_quconstraints univs csts then univs else raise NotConvertible (* general conversion and inference functions *) let check_inductive_instances cv_pb variance u1 u2 univs = let csts = get_cumulativity_constraints cv_pb variance u1 u2 in - if (Evd.check_constraints univs csts) then univs + if (Evd.check_quconstraints univs csts) then univs else raise NotConvertible let checked_universes = @@ -1113,17 +1110,7 @@ let univproblem_compare_sorts env pb s0 s1 uset = let univproblem_compare_instances ~flex i0 i1 uset = UnivProblem.enforce_eq_instances_univs flex i0 i1 uset -let univproblem_check_inductive_instances cv_pb variances u u' uset = - let open UnivProblem in - let open Univ.Variance in - let mk u = Sorts.sort_of_univ @@ Univ.Universe.make u in - let fold cstr v u u' = match v with - | Irrelevant -> Set.add (UWeak (u,u')) cstr - | Covariant when cv_pb == Conversion.CUMUL -> Set.add (ULe (mk u, mk u')) cstr - | Covariant | Invariant -> Set.add (UEq (mk u, mk u')) cstr - in - Array.fold_left3 fold uset - variances (Univ.Instance.to_array u) (Univ.Instance.to_array u') +let univproblem_check_inductive_instances = UnivProblem.compare_cumulative_instances let univproblem_univ_state = let open Conversion in @@ -1471,7 +1458,7 @@ let whd_betaiota_deltazeta_for_iota_state ts env sigma s = | Some (t_o, args) when isConstruct sigma t_o -> whrec (t_o, Stack.append_app args stack') | (Some _ | None) -> s end - |args, (Stack.Proj p :: stack'') -> + |args, (Stack.Proj (p,_) :: stack'') -> begin match whd_opt (t, args) with | Some (t_o, args) when isConstruct sigma t_o -> whrec (args.(Projection.npars p + Projection.arg p), stack'') @@ -1531,11 +1518,13 @@ let infer_convert_instances ~flex u u' (univs,cstrs as cuniv) = if UGraph.check_eq_instances univs u u' then cuniv else raise NotConvertible else - let cstrs' = Univ.enforce_eq_instances u u' Constraints.empty in + let qcstrs, cstrs' = UVars.enforce_eq_instances u u' Sorts.QUConstraints.empty in + let () = if not (Sorts.QConstraints.trivial qcstrs) then raise NotConvertible in (univs, Constraints.union cstrs cstrs') let infer_inductive_instances cv_pb variance u1 u2 (univs,csts) = - let csts' = get_cumulativity_constraints cv_pb variance u1 u2 in + let qcsts, csts' = get_cumulativity_constraints cv_pb variance u1 u2 in + let () = if not (Sorts.QConstraints.trivial qcsts) then raise NotConvertible in (UGraph.merge_constraints csts' univs, Univ.Constraints.union csts csts') let inferred_universes : (UGraph.t * Univ.Constraints.t) universe_compare = diff --git a/pretyping/reductionops.mli b/pretyping/reductionops.mli index d09bb4a0e5f7..674e5c4faefb 100644 --- a/pretyping/reductionops.mli +++ b/pretyping/reductionops.mli @@ -63,7 +63,7 @@ module Stack : sig type member = | App of app_node | Case of case_stk - | Proj of Projection.t + | Proj of Projection.t * Sorts.relevance | Fix of EConstr.fixpoint * t | Primitive of CPrimitives.t * (Constant.t * EInstance.t) * t * CPrimitives.args_red and t = member list diff --git a/pretyping/retyping.ml b/pretyping/retyping.ml index 74a5bb6e6123..3814fd4617fa 100644 --- a/pretyping/retyping.ml +++ b/pretyping/retyping.ml @@ -179,7 +179,7 @@ let retype ?(polyprop=true) sigma = | App(f,args) -> strip_outer_cast sigma (subst_type env sigma (type_of env f) (Array.to_list args)) - | Proj (p,c) -> + | Proj (p,_,c) -> let ty = type_of env c in EConstr.of_constr (try Inductiveops.type_of_projection_knowing_arg env sigma p c ty @@ -343,13 +343,17 @@ let relevance_of_term env sigma c = | LetIn ({binder_relevance=r}, _, _, bdy) -> aux (Range.cons r rels) bdy | App (c, _) -> aux rels c - | Const (c,_) -> Relevanceops.relevance_of_constant env c + | Const (c,u) -> + let u = EInstance.kind sigma u in + Relevanceops.relevance_of_constant env (c,u) | Ind _ -> Sorts.Relevant - | Construct (c,_) -> Relevanceops.relevance_of_constructor env c + | Construct (c,u) -> + let u = Unsafe.to_instance u in + Relevanceops.relevance_of_constructor env (c,u) | Case (ci, _, _, _, _, _, _) -> ci.ci_relevance | Fix ((_,i),(lna,_,_)) -> (lna.(i)).binder_relevance | CoFix (i,(lna,_,_)) -> (lna.(i)).binder_relevance - | Proj (p, _) -> Relevanceops.relevance_of_projection env p + | Proj (p, r, _) -> r | Int _ | Float _ | Array _ -> Sorts.Relevant | Meta _ | Evar _ -> Sorts.Relevant diff --git a/pretyping/structures.ml b/pretyping/structures.ml index 41abda4982d7..203b9f1c0f35 100644 --- a/pretyping/structures.ml +++ b/pretyping/structures.ml @@ -131,7 +131,7 @@ end type obj_typ = { o_ORIGIN : GlobRef.t; o_DEF : constr; - o_CTX : Univ.AbstractContext.t; + o_CTX : UVars.AbstractContext.t; o_INJ : int option; (* position of trivial argument if any *) o_TABS : constr list; (* ordered *) o_TPARAMS : constr list; (* ordered *) @@ -171,7 +171,7 @@ let rec of_constr env t = | Rel n -> Default_cs, Some n, [] | Lambda (_, _, b) -> let patt, _, _ = of_constr env b in patt, None, [] | Prod (_,_,_) -> Prod_cs, None, [t] - | Proj (p, c) -> Proj_cs (Names.Projection.repr p), None, [c] + | Proj (p, _, c) -> Proj_cs (Names.Projection.repr p), None, [c] | Sort s -> Sort_cs (Sorts.family s), None, [] | _ -> Const_cs (fst @@ destRef t) , None, [] @@ -211,7 +211,7 @@ let compute_canonical_projections env ~warn (gref,ind) = let o_DEF, c = match gref with | GlobRef.ConstRef con -> - let u = Univ.make_abstract_instance o_CTX in + let u = UVars.make_abstract_instance o_CTX in mkConstU (con, u), Environ.constant_value_in env (con,u) | GlobRef.VarRef id -> mkVar id, Option.get (Environ.named_body id env) @@ -278,7 +278,7 @@ let make env sigma ref = let vc = match ref with | GlobRef.ConstRef sp -> - let u = Univ.make_abstract_instance (Environ.constant_context env sp) in + let u = UVars.make_abstract_instance (Environ.constant_context env sp) in begin match Environ.constant_opt_value_in env (sp, u) with | Some vc -> vc | None -> error_not_structure ref (str "Could not find its value in the global environment") end @@ -360,7 +360,7 @@ let find env sigma (proj,pat) = let bs' = List.map (EConstr.of_constr %> EConstr.Vars.subst_instance_constr u) bs in let params = List.map (fun c -> EConstr.Vars.subst_instance_constr u c) params in let us = List.map (fun c -> EConstr.Vars.subst_instance_constr u c) us in - let sigma = Evd.merge_context_set Evd.univ_flexible sigma ctx' in + let sigma = Evd.merge_sort_context_set Evd.univ_flexible sigma ctx' in sigma, { body = t'; constant = c'; abstractions_ty = bs'; nparams; params; cvalue_arguments = us; cvalue_abstraction = n } @@ -381,7 +381,7 @@ let rec decompose_projection sigma c args = (* Check if there is some canonical projection attached to this structure *) let _ = GlobRef.Map.find (GlobRef.ConstRef c) !object_table in get_nth n args - | Proj (p, c) -> + | Proj (p, _, c) -> let _ = GlobRef.Map.find (GlobRef.ConstRef (Names.Projection.constant p)) !object_table in c | _ -> raise Not_found @@ -437,4 +437,10 @@ let mem c = Cmap_env.mem c !prim_table let find_opt c = try Some (Cmap_env.find c !prim_table) with Not_found -> None +let find_opt_with_relevance (c,u) = + find_opt c |> + Option.map (fun p -> + let u = EConstr.Unsafe.to_instance u in + p, Relevanceops.relevance_of_projection_repr (Global.env()) (p,u)) + end diff --git a/pretyping/structures.mli b/pretyping/structures.mli index c068dffd53f7..96c34697f17e 100644 --- a/pretyping/structures.mli +++ b/pretyping/structures.mli @@ -160,4 +160,7 @@ val mem : Names.Constant.t -> bool val find_opt : Names.Constant.t -> Names.Projection.Repr.t option +val find_opt_with_relevance : Names.Constant.t * EConstr.EInstance.t + -> (Names.Projection.Repr.t * Sorts.relevance) option + end diff --git a/pretyping/tacred.ml b/pretyping/tacred.ml index a1d349153780..c1d39c9a0cd2 100644 --- a/pretyping/tacred.ml +++ b/pretyping/tacred.ml @@ -336,7 +336,7 @@ let compute_consteval allowed_reds env sigma ref u = with Elimconst -> NotAnElimination) | Case (_,_,_,_,_,d,_) when isRel sigma d && not onlyproj -> EliminationCases (List.length all_abs) | Case (_,_,_,_,_,d,_) -> srec env all_abs lastref lastu true d - | Proj (p, d) when isRel sigma d -> EliminationProj (List.length all_abs) + | Proj (p, _, d) when isRel sigma d -> EliminationProj (List.length all_abs) | _ when isTransparentEvalRef env sigma (RedFlags.red_transparent allowed_reds) c' -> (* Continue stepwise unfolding from [c' args] *) let ref, u = destEvalRefU sigma c' in @@ -527,9 +527,9 @@ let match_eval_ref_value env sigma constr stack = Some (EConstr.of_constr (constant_value_in env (sp, u))) else None - | Proj (p, c) when not (Projection.unfolded p) -> + | Proj (p, r, c) when not (Projection.unfolded p) -> if is_evaluable env (EvalConstRef (Projection.constant p)) then - Some (mkProj (Projection.unfold p, c)) + Some (mkProj (Projection.unfold p, r, c)) else None | Var id when is_evaluable env (EvalVarRef id) -> env |> lookup_named id |> NamedDecl.get_value @@ -737,7 +737,7 @@ and whd_simpl_stack allowed_reds env sigma = | Reduced s' -> redrec s' | NotReducible -> s' end - | Proj (p, c) -> + | Proj (p, _, c) -> let ans = let unf = Projection.unfolded p in if unf || is_evaluable env (EvalConstRef (Projection.constant p)) then @@ -811,7 +811,7 @@ and reduce_fix allowed_reds env sigma f fix stack = and reduce_proj allowed_reds env sigma c = let rec redrec s = match EConstr.kind sigma s with - | Proj (proj, c) -> + | Proj (proj, _, c) -> let c' = match redrec c with NotReducible -> c | Reduced c -> c in let* (constr, cargs) = whd_construct_stack allowed_reds env sigma c' in (match EConstr.kind sigma constr with @@ -887,7 +887,7 @@ let try_red_product env sigma c = | Case (ci,u,pms,p,iv,d,lf) -> let* d = redrec env d in Reduced (simpfun (mkCase (ci,u,pms,p,iv,d,lf))) - | Proj (p, c) -> + | Proj (p, _, c) -> let* c' = match EConstr.kind sigma c with | Construct _ -> Reduced c @@ -977,7 +977,7 @@ let whd_simpl_orelse_delta_but_fix env sigma c = | Some c -> (match EConstr.kind sigma (snd (decompose_lambda sigma c)) with | CoFix _ | Fix _ -> s' - | Proj (p,t) when + | Proj (p,_,t) when (match EConstr.kind sigma constr with | Const (c', _) -> QConstant.equal env (Projection.constant p) c' | _ -> false) -> @@ -1017,7 +1017,7 @@ let simpl env sigma c = let matches_head env sigma c t = let t, l = decompose_app sigma t in match EConstr.kind sigma t, Array.is_empty l with - | Proj (p, _), _ -> Constr_matching.matches env sigma c (mkConstU (Projection.constant p, EInstance.empty)) + | Proj (p, _, _), _ -> Constr_matching.matches env sigma c (mkConstU (Projection.constant p, EInstance.empty)) | _, false -> Constr_matching.matches env sigma c t | _ -> raise Constr_matching.PatternMatchingFailure @@ -1027,8 +1027,8 @@ let matches_head env sigma c t = *) let change_map_constr_with_binders_left_to_right g f (env, l as acc) sigma c = match EConstr.kind sigma c with - | Proj (p, r) -> (* Treat specially for partial applications *) - let t = Retyping.expand_projection env sigma p r [] in + | Proj (p, r, v) -> (* Treat specially for partial applications *) + let t = Retyping.expand_projection env sigma p v [] in let hdf, al = destApp sigma t in let a = al.(Array.length al - 1) in let app = (mkApp (hdf, Array.sub al 0 (Array.length al - 1))) in @@ -1037,7 +1037,7 @@ let change_map_constr_with_binders_left_to_right g f (env, l as acc) sigma c = let hdf', _ = decompose_app sigma app' in if hdf' == hdf then (* Still the same projection, we ignore the change in parameters *) - mkProj (p, a') + mkProj (p, r, a') else mkApp (app', [| a' |]) | _ -> map_constr_with_binders_left_to_right env sigma g f acc c @@ -1071,7 +1071,7 @@ let e_contextually byhead (occs,c) f = begin fun env sigma t -> application with same head *) match EConstr.kind !evd t with | App (f,l) when byhead -> mkApp (f, Array.map_left (traverse nested envc) l) - | Proj (p,c) when byhead -> mkProj (p,traverse nested envc c) + | Proj (p,r,c) when byhead -> mkProj (p,r,traverse nested envc c) | _ -> change_map_constr_with_binders_left_to_right (fun d (env,c) -> (push_rel d env, Patternops.lift_pattern 1 c)) diff --git a/pretyping/typeclasses.ml b/pretyping/typeclasses.ml index f3c923c3a8cd..5f466693d642 100644 --- a/pretyping/typeclasses.ml +++ b/pretyping/typeclasses.ml @@ -48,7 +48,7 @@ type class_method = { (* This module defines type-classes *) type typeclass = { (* Universe quantification *) - cl_univs : Univ.AbstractContext.t; + cl_univs : UVars.AbstractContext.t; (* The class implementation *) cl_impl : GlobRef.t; @@ -90,7 +90,7 @@ let classes : typeclasses ref = Summary.ref GlobRef.Map.empty ~name:"classes" let instances : instances ref = Summary.ref GlobRef.Map.empty ~name:"instances" let typeclass_univ_instance (cl, u) = - assert (Univ.AbstractContext.size cl.cl_univs == Univ.Instance.length u); + assert (UVars.eq_sizes (UVars.AbstractContext.size cl.cl_univs) (UVars.Instance.length u)); let subst_ctx c = Context.Rel.map (subst_instance_constr u) c in { cl with cl_context = subst_ctx cl.cl_context; cl_props = subst_ctx cl.cl_props} @@ -112,7 +112,7 @@ let global_class_of_constr env sigma c = let decompose_class_app env sigma c = let hd, args = EConstr.decompose_app_list sigma c in match EConstr.kind sigma hd with - | Proj (p, c) -> + | Proj (p, _, c) -> let expp = Retyping.expand_projection env sigma p c args in EConstr.decompose_app_list sigma expp | _ -> hd, args @@ -140,7 +140,7 @@ let rec is_class_type evd c = match EConstr.kind evd c with | Prod (_, _, t) -> is_class_type evd t | Cast (t, _, _) -> is_class_type evd t - | Proj (p, c) -> GlobRef.(Map.mem (ConstRef (Projection.constant p))) !classes + | Proj (p, _, c) -> GlobRef.(Map.mem (ConstRef (Projection.constant p))) !classes | _ -> is_class_constr evd c let is_class_evar evd evi = @@ -152,7 +152,7 @@ let rec is_maybe_class_type evd c = | Prod (_, _, t) -> is_maybe_class_type evd t | Cast (t, _, _) -> is_maybe_class_type evd t | Evar _ -> true - | Proj (p, c) -> GlobRef.(Map.mem (ConstRef (Projection.constant p))) !classes + | Proj (p, _, c) -> GlobRef.(Map.mem (ConstRef (Projection.constant p))) !classes | _ -> is_class_constr evd c let () = Hook.set Evd.is_maybe_typeclass_hook (fun evd c -> is_maybe_class_type evd (EConstr.of_constr c)) diff --git a/pretyping/typeclasses.mli b/pretyping/typeclasses.mli index c44ed77e1446..f289f51d7024 100644 --- a/pretyping/typeclasses.mli +++ b/pretyping/typeclasses.mli @@ -28,7 +28,7 @@ type class_method = { (** This module defines type-classes *) type typeclass = { - cl_univs : Univ.AbstractContext.t; + cl_univs : UVars.AbstractContext.t; (** The toplevel universe quantification in which the typeclass lives. In particular, [cl_props] and [cl_context] are quantified over it. *) @@ -89,7 +89,7 @@ val class_info_exn : env -> evar_map -> GlobRef.t -> typeclass val dest_class_app : env -> evar_map -> EConstr.constr -> (typeclass * EConstr.EInstance.t) * constr list (** Get the instantiated typeclass structure for a given universe instance. *) -val typeclass_univ_instance : typeclass Univ.puniverses -> typeclass +val typeclass_univ_instance : typeclass UVars.puniverses -> typeclass (** Just return None if not a class *) val class_of_constr : env -> evar_map -> EConstr.constr -> diff --git a/pretyping/typing.ml b/pretyping/typing.ml index 69182af16a58..57277d37b5c2 100644 --- a/pretyping/typing.ml +++ b/pretyping/typing.ml @@ -192,6 +192,32 @@ let type_case_branches env sigma (ind,largs) specif pj c = let ty = whd_betaiota env sigma (lambda_applist_decls sigma (n+1) p (realargs@[c])) in sigma, (lc, ty, ESorts.relevance_of_sort sigma ps) +let unify_relevance sigma r1 r2 = + match Evarutil.nf_relevance sigma r1, Evarutil.nf_relevance sigma r2 with + | Relevant, Relevant | Irrelevant, Irrelevant -> Some sigma + | Relevant, Irrelevant | Irrelevant, Relevant -> None + | Irrelevant, RelevanceVar q | RelevanceVar q, Irrelevant -> + let sigma = + Evd.add_quconstraints sigma + (Sorts.QConstraints.singleton (Sorts.Quality.qsprop, Equal, QVar q), + Univ.Constraints.empty) + in + Some sigma + | Relevant, RelevanceVar q | RelevanceVar q, Relevant -> + let sigma = + Evd.add_quconstraints sigma + (Sorts.QConstraints.singleton (Sorts.Quality.qprop, Leq, QVar q), + Univ.Constraints.empty) + in + Some sigma + | RelevanceVar q1, RelevanceVar q2 -> + let sigma = + Evd.add_quconstraints sigma + (Sorts.QConstraints.singleton (QVar q1, Equal, QVar q2), + Univ.Constraints.empty) + in + Some sigma + let judge_of_case env sigma case ci pj iv cj lfj = let ((ind, u), spec) = try find_mrectype env sigma cj.uj_type @@ -200,6 +226,13 @@ let judge_of_case env sigma case ci pj iv cj lfj = let () = if Inductive.is_private specif then Type_errors.error_case_on_private_ind env ind in let indspec = ((ind, EInstance.kind sigma u), spec) in let sigma, (bty,rslty,rci) = type_case_branches env sigma indspec specif pj cj.uj_val in + (* should we have evar map aware check_case_info and should_invert_case? *) + let sigma, ci = + if Sorts.relevance_equal ci.ci_relevance rci then sigma, ci + else match unify_relevance sigma ci.ci_relevance rci with + | None -> sigma, ci (* will fail in check_case_info *) + | Some sigma -> sigma, { ci with ci_relevance = rci } + in let () = check_case_info env (fst indspec) rci ci in let sigma = check_branch_types env sigma (fst indspec) cj (lfj,bty) in let () = if (match iv with | NoInvert -> false | CaseInvert _ -> true) != should_invert_case env ci @@ -226,18 +259,19 @@ let check_allowed_sort env sigma ind c p = let sorts = elim_sort specif in let pj = Retyping.get_judgment_of env sigma p in let _, s = whd_decompose_prod env sigma pj.uj_type in - let ksort = match EConstr.kind sigma s with - | Sort s -> - begin match ESorts.family sigma s with + let sort = match EConstr.kind sigma s with + | Sort s -> EConstr.ESorts.kind sigma s + | _ -> error_elim_arity env sigma ind c None + in + let ksort = match Sorts.family sort with | InType | InSProp | InSet | InProp as f -> f | InQSort -> InType (* FIXME *) - end - | _ -> error_elim_arity env sigma ind c None in + in if not (Sorts.family_leq ksort sorts) then let s = inductive_sort_family (snd specif) in error_elim_arity env sigma ind c (Some (pj, sorts, ksort, s)) else - Sorts.relevance_of_sort_family ksort + Sorts.relevance_of_sort sort let check_actual_type env sigma cj t = try Evarconv.unify_leq_delay env sigma cj.uj_type t @@ -296,15 +330,16 @@ let judge_of_variable env id = Environ.on_judgment EConstr.of_constr (judge_of_variable env id) let judge_of_projection env sigma p cj = - let pty = lookup_projection p env in + let pr, pty = lookup_projection p env in let (ind,u), args = try find_mrectype env sigma cj.uj_type with Not_found -> error_case_not_inductive env sigma cj in let u = EInstance.kind sigma u in + let pr = UVars.subst_instance_relevance u pr in let ty = EConstr.of_constr (CVars.subst_instance_constr u pty) in let ty = substl (cj.uj_val :: List.rev args) ty in - {uj_val = EConstr.mkProj (p,cj.uj_val); + {uj_val = EConstr.mkProj (p,pr,cj.uj_val); uj_type = ty} let judge_of_abstraction env sigma name var j = @@ -359,8 +394,8 @@ let judge_of_float env v = Environ.on_judgment EConstr.of_constr (judge_of_float env v) let judge_of_array env sigma u tj defj tyj = - let ulev = match Univ.Instance.to_array u with - | [|u|] -> u + let ulev = match UVars.Instance.to_array u with + | [||], [|u|] -> u | _ -> assert false in let sigma = Evd.set_leq_sort env sigma tyj.utj_type @@ -495,7 +530,7 @@ let rec execute env sigma cstr = | QSort _ as s -> sigma, judge_of_sort s end - | Proj (p, c) -> + | Proj (p, _, c) -> let sigma, cj = execute env sigma c in sigma, judge_of_projection env sigma p cj @@ -725,8 +760,8 @@ let rec recheck_against env sigma good c = then assume_unchanged_type sigma else maybe_changed (judge_of_case env sigma case ci pj iv cj lfj) - | Proj (gp, gc), - Proj (p, c) -> + | Proj (gp, _, gc), + Proj (p, _, c) -> if not (QProjection.equal env gp p) then default () else let sigma, changed, c = recheck_against env sigma gc c in diff --git a/pretyping/unification.ml b/pretyping/unification.ml index 40e80da40338..51fd685ad9b9 100644 --- a/pretyping/unification.ml +++ b/pretyping/unification.ml @@ -462,23 +462,23 @@ let use_metas_pattern_unification sigma flags nb l = type key = | IsKey of CClosure.table_key - | IsProj of Projection.t * EConstr.constr + | IsProj of Projection.t * Sorts.relevance * EConstr.constr let expand_table_key env = function | ConstKey cst -> constant_opt_value_in env cst | VarKey id -> (try named_body id env with Not_found -> None) | RelKey _ -> None -let unfold_projection env p stk = - let s = Stack.Proj p in +let unfold_projection env p r stk = + let s = Stack.Proj (p,r) in s :: stk let expand_key ts env sigma = function | Some (IsKey k) -> Option.map EConstr.of_constr (expand_table_key env k) - | Some (IsProj (p, c)) -> + | Some (IsProj (p, r, c)) -> let red = Stack.zip sigma (whd_betaiota_deltazeta_for_iota_state ts env sigma - (c, unfold_projection env p [])) - in if EConstr.eq_constr sigma (EConstr.mkProj (p, c)) red then None else Some red + (c, unfold_projection env p r [])) + in if EConstr.eq_constr sigma (EConstr.mkProj (p, r, c)) red then None else Some red | None -> None let isApp_or_Proj sigma c = @@ -506,10 +506,10 @@ let key_of env sigma b flags f = | Var id when is_transparent env (VarKey id) && TransparentState.is_transparent_variable flags.modulo_delta id -> Some (IsKey (VarKey id)) - | Proj (p, c) when Names.Projection.unfolded p + | Proj (p, r, c) when Names.Projection.unfolded p || (is_transparent env (ConstKey (Projection.constant p)) && (TransparentState.is_transparent_constant flags.modulo_delta (Projection.constant p))) -> - Some (IsProj (p, c)) + Some (IsProj (p, r, c)) | _ -> None @@ -520,7 +520,7 @@ let translate_key = function let translate_key = function | IsKey k -> translate_key k - | IsProj (c, _) -> ConstKey (Projection.constant c) + | IsProj (c, _, _) -> ConstKey (Projection.constant c) let oracle_order env cf1 cf2 = match cf1 with @@ -533,10 +533,10 @@ let oracle_order env cf1 cf2 = | None -> Some true | Some k2 -> match k1, k2 with - | IsProj (p, _), IsKey (ConstKey (p',_)) + | IsProj (p, _, _), IsKey (ConstKey (p',_)) when Environ.QConstant.equal env (Projection.constant p) p' -> Some (not (Projection.unfolded p)) - | IsKey (ConstKey (p,_)), IsProj (p', _) + | IsKey (ConstKey (p,_)), IsProj (p', _, _) when Environ.QConstant.equal env p (Projection.constant p') -> Some (Projection.unfolded p') | _ -> @@ -551,7 +551,7 @@ let is_rigid_head sigma flags t = | Fix _ | CoFix _ -> true | Rel _ | Var _ | Meta _ | Evar _ | Sort _ | Cast (_, _, _) | Prod _ | Lambda _ | LetIn _ | App (_, _) | Case _ - | Proj (_, _) -> false (* Why aren't Prod, Sort rigid heads ? *) + | Proj _ -> false (* Why aren't Prod, Sort rigid heads ? *) let constr_cmp pb env sigma flags ?nargs t u = let cstrs = @@ -644,7 +644,7 @@ let rec is_neutral env sigma ts t = | Rel n -> true | Evar _ | Meta _ -> true | Case (_, _, _, _, _, c, _) -> is_neutral env sigma ts c - | Proj (p, c) -> is_neutral env sigma ts c + | Proj (p, _, c) -> is_neutral env sigma ts c | Lambda _ | LetIn _ | Construct _ | CoFix _ | Int _ | Float _ | Array _ -> false | Sort _ | Cast (_, _, _) | Prod (_, _, _) | Ind _ -> false (* Really? *) | Fix _ -> false (* This is an approximation *) @@ -673,7 +673,8 @@ let eta_constructor_app env sigma f l1 term = let npars = mib.Declarations.mind_nparams in let pars, l1' = Array.chop npars l1 in let arg = Array.append pars [|term|] in - let l2 = Array.map (fun p -> mkApp (mkConstU (Projection.Repr.constant p,u), arg)) projs in + let l2 = Array.map (fun (p,_) -> + mkApp (mkConstU (Projection.Repr.constant p,u), arg)) projs in l1', l2 | _ -> assert false) | _ -> assert false @@ -805,7 +806,7 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst : subst0) conv_at_top e | _, LetIn (_,a,_,c) -> unirec_rec curenvnb pb opt substn cM (subst1 a c) (* Fast path for projections. *) - | Proj (p1,c1), Proj (p2,c2) when Environ.QConstant.equal env + | Proj (p1,_,c1), Proj (p2,_,c2) when Environ.QConstant.equal env (Projection.constant p1) (Projection.constant p2) -> (try unify_same_proj curenvnb cv_pb {opt with at_top = true} substn c1 c2 @@ -902,10 +903,10 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst : subst0) conv_at_top e | App (f1,l1), App (f2,l2) -> unify_app curenvnb pb opt substn cM f1 l1 cN f2 l2 - | App (f1,l1), Proj(p2,c2) -> + | App (f1,l1), Proj(p2,_,c2) -> unify_app curenvnb pb opt substn cM f1 l1 cN cN [||] - | Proj (p1,c1), App(f2,l2) -> + | Proj (p1,_,c1), App(f2,l2) -> unify_app curenvnb pb opt substn cM cM [||] cN f2 l2 | _ -> @@ -959,7 +960,7 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst : subst0) conv_at_top e in let expand_proj c c' l = match EConstr.kind sigma c with - | Proj (p, t) when not (Projection.unfolded p) && needs_expansion p c' -> + | Proj (p, _, t) when not (Projection.unfolded p) && needs_expansion p c' -> (try destApp sigma (Retyping.expand_projection curenv sigma p t (Array.to_list l)) with RetypeError _ -> (* Unification can be called on ill-typed terms, due to FO and eta in particular, fail gracefully in that case *) @@ -1887,7 +1888,7 @@ let rec make sigma c0 = match EConstr.kind sigma c0 with let c = make sigma c in let ald, al = make_array sigma al in { proj = c0; self = AApp (c, al); data = max c.data ald } -| Proj (p, t) -> +| Proj (p, _, t) -> let t = make sigma t in { proj = c0; self = AOther [|t|]; data = t.data } | Evar (e, al) -> @@ -2023,7 +2024,7 @@ let w_unify_to_subterm_all env evd ?(flags=default_unify_flags ()) (op,cl) = | Case(_,_,_,_,_,c,lf) -> (* does not search in the predicate *) bind (matchrec c) (bind_iter matchrec (Array.map snd lf)) - | Proj (p,c) -> matchrec c + | Proj (p,_,c) -> matchrec c | LetIn(_,c1,_,c2) -> bind (matchrec c1) (matchrec c2) diff --git a/pretyping/vnorm.ml b/pretyping/vnorm.ml index febfb4b5d22e..cb476c43839e 100644 --- a/pretyping/vnorm.ml +++ b/pretyping/vnorm.ml @@ -189,7 +189,7 @@ and nf_whd env sigma whd typ = | Vaccu (Aind ((mi, i) as ind), stk) -> let mib = Environ.lookup_mind mi env in let nb_univs = - Univ.AbstractContext.size (Declareops.inductive_polymorphic_context mib) + UVars.AbstractContext.size (Declareops.inductive_polymorphic_context mib) in let mk u = let pind = (ind, u) in (mkIndU pind, type_of_ind env pind) @@ -200,17 +200,17 @@ and nf_whd env sigma whd typ = and nf_univ_args ~nb_univs mk env sigma stk = let u = - if Int.equal nb_univs 0 then Univ.Instance.empty + if UVars.eq_sizes nb_univs (0,0) then UVars.Instance.empty else match stk with | Zapp args :: _ -> let inst = arg args 0 in let inst = uni_instance inst in - let () = assert (Int.equal (Univ.Instance.length inst) nb_univs) in + let () = assert (UVars.eq_sizes (UVars.Instance.length inst) nb_univs) in inst | _ -> assert false in let (t,ty) = mk u in - let from = if Int.equal nb_univs 0 then 0 else 1 in + let from = if UVars.Instance.is_empty u then 0 else 1 in nf_stk ~from env sigma t ty stk and nf_evar env sigma evk stk = @@ -245,7 +245,7 @@ and constr_type_of_idkey env sigma (idkey : Vmvalues.id_key) stk = | ConstKey cst -> let cbody = Environ.lookup_constant cst env in let nb_univs = - Univ.AbstractContext.size (Declareops.constant_polymorphic_context cbody) + UVars.AbstractContext.size (Declareops.constant_polymorphic_context cbody) in let mk u = let pcst = (cst, u) in (mkConstU pcst, Typeops.type_of_constant_in env pcst) @@ -320,9 +320,10 @@ and nf_stk ?from:(from=0) env sigma c t stk = let ty = Vars.subst_instance_constr (EConstr.Unsafe.to_instance u) tys.(p) in substl (c :: List.rev_map EConstr.Unsafe.to_constr pars) ty in - let p = Declareops.inductive_make_projection ind mib ~proj_arg:p in + let p, r = Declareops.inductive_make_projection ind mib ~proj_arg:p in let p = Projection.make p true in - nf_stk env sigma (mkProj (p, c)) ty stk + let r = UVars.subst_instance_relevance (EConstr.Unsafe.to_instance u) r in + nf_stk env sigma (mkProj (p, r, c)) ty stk and nf_predicate env sigma ind mip params v pctx = (* TODO: we should expose some variant of Vm.mkrel_vstack *) diff --git a/printing/ppconstr.ml b/printing/ppconstr.ml index a9948b634fc4..67d5965a2e56 100644 --- a/printing/ppconstr.ml +++ b/printing/ppconstr.ml @@ -193,9 +193,18 @@ let tag_var = tag Tag.variable | [x] -> pr_univ_expr x | l -> str"max(" ++ prlist_with_sep (fun () -> str",") pr_univ_expr l ++ str")" + let pr_qvar_expr = function + | CQAnon _ -> tag_type (str "_") + | CQVar qid -> tag_type (pr_qualid qid) + | CRawQVar q -> (* TODO names *) tag_type (Sorts.QVar.raw_pr q) + + let pr_quality_expr q = match q with + | CQConstant q -> tag_type (Sorts.Quality.Constants.pr q) + | CQualVar q -> pr_qvar_expr q + let pr_quality_univ (q, l) = match q with | None -> pr_univ l - | Some q -> str "(* " ++ Sorts.QVar.pr q ++ str " *)" ++ spc () ++ pr_univ l + | Some q -> pr_qvar_expr q ++ spc() ++ str "|" ++ spc () ++ pr_univ l let pr_univ_annot pr x = str "@{" ++ pr x ++ str "}" @@ -233,8 +242,13 @@ let tag_var = tag Tag.variable let pr_qualid = pr_qualid let pr_patvar = pr_id + let pr_inside_universe_instance (ql,ul) = + (if List.is_empty ql then mt() + else prlist_with_sep spc pr_quality_expr ql ++ strbrk " | ") + ++ prlist_with_sep spc pr_univ_level_expr ul + let pr_universe_instance l = - pr_opt_no_spc (pr_univ_annot (prlist_with_sep spc pr_univ_level_expr)) l + pr_opt_no_spc (pr_univ_annot pr_inside_universe_instance) l let pr_reference qid = if qualid_is_ident qid then tag_var (pr_id @@ qualid_basename qid) diff --git a/printing/printer.ml b/printing/printer.ml index 5f6bd9e2ec5e..68117cc588f4 100644 --- a/printing/printer.ml +++ b/printing/printer.ml @@ -177,41 +177,66 @@ let safe_gen f env sigma c = let safe_pr_lconstr_env = safe_gen pr_lconstr_env let safe_pr_constr_env = safe_gen pr_constr_env +let q_ident = Id.of_string "α" + let u_ident = Id.of_string "u" let universe_binders_with_opt_names orig names = let open Univ in - let orig = Univ.AbstractContext.names orig in - let orig = Array.to_list orig in - let udecl = match names with + let qorig, uorig = UVars.AbstractContext.names orig in + let qorig, uorig as orig = Array.to_list qorig, Array.to_list uorig in + let qdecl, udecl = match names with | None -> orig - | Some udecl -> + | Some (qdecl,udecl) -> try - List.map2 (fun orig {CAst.v = na} -> - match na with - | Anonymous -> orig - | Name id -> Name id) orig udecl + let qs = + List.map2 (fun orig {CAst.v = na} -> + match na with + | Anonymous -> orig + | Name id -> Name id) qorig qdecl + in + let us = + List.map2 (fun orig {CAst.v = na} -> + match na with + | Anonymous -> orig + | Name id -> Name id) uorig udecl + in + qs, us with Invalid_argument _ -> let open UnivGen in raise (UniverseLengthMismatch { - actual = List.length orig; - expect = List.length udecl; + actual = List.length qorig, List.length uorig; + expect = List.length qdecl, List.length udecl; }) in - let fold_named i (ubind,revubind as o) = function + let fold_qnamed i ((qbind,ubind),(revqbind,revubind) as o) = function + | Name id -> let ui = Sorts.QVar.make_var i in + (Id.Map.add id ui qbind, ubind), (Sorts.QVar.Map.add ui id revqbind, revubind) + | Anonymous -> o + in + let fold_unamed i ((qbind,ubind),(revqbind,revubind) as o) = function | Name id -> let ui = Level.var i in - Id.Map.add id ui ubind, Level.Map.add ui id revubind + (qbind, Id.Map.add id ui ubind), (revqbind, Level.Map.add ui id revubind) | Anonymous -> o in - let names = List.fold_left_i fold_named 0 (UnivNames.empty_binders,Level.Map.empty) udecl in - let fold_anons i (u_ident, (ubind, revubind) as o) = function + let names = List.fold_left_i fold_qnamed 0 UnivNames.(empty_binders,empty_rev_binders) qdecl in + let names = List.fold_left_i fold_unamed 0 names udecl in + let fold_qanons i (u_ident, ((qbind,ubind), (revqbind,revubind)) as o) = function + | Name _ -> o + | Anonymous -> + let ui = Sorts.QVar.make_var i in + let id = Namegen.next_ident_away_from u_ident (fun id -> Id.Map.mem id qbind) in + (id, ((Id.Map.add id ui qbind, ubind), (Sorts.QVar.Map.add ui id revqbind, revubind))) + in + let fold_uanons i (u_ident, ((qbind,ubind), (revqbind,revubind)) as o) = function | Name _ -> o | Anonymous -> let ui = Level.var i in let id = Namegen.next_ident_away_from u_ident (fun id -> Id.Map.mem id ubind) in - (id, (Id.Map.add id ui ubind, Level.Map.add ui id revubind)) + (id, ((qbind,Id.Map.add id ui ubind), (revqbind,Level.Map.add ui id revubind))) in - let (_, names) = List.fold_left_i fold_anons 0 (u_ident, names) udecl in + let (_, names) = List.fold_left_i fold_qanons 0 (q_ident, names) qdecl in + let (_, names) = List.fold_left_i fold_uanons 0 (u_ident, names) udecl in names let pr_universe_ctx_set sigma c = @@ -221,8 +246,12 @@ let pr_universe_ctx_set sigma c = mt() let pr_universe_ctx sigma ?variance c = - if !Detyping.print_universes && not (Univ.UContext.is_empty c) then - fnl()++pr_in_comment (v 0 (Univ.pr_universe_context (Termops.pr_evd_level sigma) ?variance c)) + if !Detyping.print_universes && not (UVars.UContext.is_empty c) then + fnl()++ + pr_in_comment + (v 0 + (UVars.pr_universe_context (Termops.pr_evd_qvar sigma) (Termops.pr_evd_level sigma) + ?variance c)) else mt() @@ -230,9 +259,10 @@ let pr_abstract_universe_ctx sigma ?variance ?priv c = let open Univ in let priv = Option.default Univ.ContextSet.empty priv in let has_priv = not (ContextSet.is_empty priv) in - if !Detyping.print_universes && (not (Univ.AbstractContext.is_empty c) || has_priv) then + if !Detyping.print_universes && (not (UVars.AbstractContext.is_empty c) || has_priv) then + let prqvar u = Termops.pr_evd_qvar sigma u in let prlev u = Termops.pr_evd_level sigma u in - let pub = (if has_priv then str "Public universes:" ++ fnl() else mt()) ++ v 0 (Univ.pr_abstract_universe_context prlev ?variance c) in + let pub = (if has_priv then str "Public universes:" ++ fnl() else mt()) ++ v 0 (UVars.pr_abstract_universe_context prqvar prlev ?variance c) in let priv = if has_priv then fnl() ++ str "Private universes:" ++ fnl() ++ v 0 (Univ.pr_universe_context_set prlev priv) else mt() in fnl()++pr_in_comment (pub ++ priv) else @@ -250,6 +280,7 @@ let pr_global = pr_global_env Id.Set.empty let pr_universe_instance_constraints evd inst csts = let open Univ in + let prqvar = Termops.pr_evd_qvar evd in let prlev = Termops.pr_evd_level evd in let pcsts = if Constraints.is_empty csts then mt() else str " |= " ++ @@ -257,7 +288,7 @@ let pr_universe_instance_constraints evd inst csts = (fun (u,d,v) -> hov 0 (prlev u ++ pr_constraint_type d ++ prlev v)) (Constraints.elements csts) in - str"@{" ++ Instance.pr prlev inst ++ pcsts ++ str"}" + str"@{" ++ UVars.Instance.pr prqvar prlev inst ++ pcsts ++ str"}" let pr_universe_instance evd inst = pr_universe_instance_constraints evd inst Univ.Constraints.empty diff --git a/printing/printer.mli b/printing/printer.mli index ec516fe85739..3e984ad29d69 100644 --- a/printing/printer.mli +++ b/printing/printer.mli @@ -119,15 +119,15 @@ val pr_sort : evar_map -> Sorts.t -> Pp.t (** Universe constraints *) -val pr_universe_instance : evar_map -> Univ.Instance.t -> Pp.t -val pr_universe_instance_constraints : evar_map -> Univ.Instance.t -> Univ.Constraints.t -> Pp.t -val pr_universe_ctx : evar_map -> ?variance:Univ.Variance.t array -> - Univ.UContext.t -> Pp.t -val pr_abstract_universe_ctx : evar_map -> ?variance:Univ.Variance.t array -> - ?priv:Univ.ContextSet.t -> Univ.AbstractContext.t -> Pp.t +val pr_universe_instance : evar_map -> UVars.Instance.t -> Pp.t +val pr_universe_instance_constraints : evar_map -> UVars.Instance.t -> Univ.Constraints.t -> Pp.t +val pr_universe_ctx : evar_map -> ?variance:UVars.Variance.t array -> + UVars.UContext.t -> Pp.t +val pr_abstract_universe_ctx : evar_map -> ?variance:UVars.Variance.t array -> + ?priv:Univ.ContextSet.t -> UVars.AbstractContext.t -> Pp.t val pr_universe_ctx_set : evar_map -> Univ.ContextSet.t -> Pp.t val pr_universes : evar_map -> - ?variance:Univ.Variance.t array -> ?priv:Univ.ContextSet.t -> + ?variance:UVars.Variance.t array -> ?priv:Univ.ContextSet.t -> Declarations.universes -> Pp.t (** [universe_binders_with_opt_names ref l] @@ -139,8 +139,8 @@ val pr_universes : evar_map -> Otherwise return the bound universe names registered for [ref]. Inefficient on large contexts due to name generation. *) -val universe_binders_with_opt_names : Univ.AbstractContext.t -> - UnivNames.univ_name_list option -> UnivNames.universe_binders * Id.t Univ.Level.Map.t +val universe_binders_with_opt_names : UVars.AbstractContext.t -> + UnivNames.full_name_list option -> UnivNames.universe_binders * UnivNames.rev_binders (** Printing global references using names as short as possible *) diff --git a/proofs/clenv.ml b/proofs/clenv.ml index 0cc47ed6dac7..c675a7e274a4 100644 --- a/proofs/clenv.ml +++ b/proofs/clenv.ml @@ -55,13 +55,13 @@ let strip_params env sigma c = match EConstr.kind sigma c with | App (f, args) -> (match EConstr.kind sigma f with - | Const (cst,_) -> - (match Structures.PrimitiveProjections.find_opt cst with - | Some p -> + | Const cst -> + (match Structures.PrimitiveProjections.find_opt_with_relevance cst with + | Some (p,r) -> let p = Projection.make p false in let npars = Projection.npars p in if Array.length args > npars then - mkApp (mkProj (p, args.(npars)), + mkApp (mkProj (p, r, args.(npars)), Array.sub args (npars+1) (Array.length args - (npars + 1))) else c | None -> c) @@ -76,9 +76,9 @@ let clenv_refresh env sigma ctx clenv = let evd = Evd.meta_merge (Evd.meta_list clenv.evd) (Evd.clear_metas sigma) in match ctx with | Some ctx -> - let (subst, ctx) = UnivGen.fresh_universe_context_set_instance ctx in + let (subst, ctx) = UnivGen.fresh_sort_context_instance ctx in let emap c = Vars.subst_univs_level_constr subst c in - let evd = Evd.merge_context_set Evd.univ_flexible evd ctx in + let evd = Evd.merge_sort_context_set Evd.univ_flexible evd ctx in (* Only metas are mentioning the old universes. *) mk_clausenv env (Evd.map_metas emap evd) clenv.metas (emap clenv.templval) @@ -598,7 +598,7 @@ type proof = | RfHole of metavariable | RfGround of EConstr.t | RfApp of proof * proof list -| RfProj of Projection.t * proof +| RfProj of Projection.t * Sorts.relevance * proof exception NonLinear @@ -619,9 +619,9 @@ let make_proof env sigma c = let args = Array.map_to_list (fun c -> make c) args in if is_ground f && List.for_all is_ground args then RfGround c else RfApp (f, args) - | Proj (p, a) -> + | Proj (p, r, a) -> let a = make a in - if is_ground a then RfGround c else RfProj (p, a) + if is_ground a then RfGround c else RfProj (p, r, a) | _ -> if occur_meta sigma c then error_unsupported_deep_meta () else RfGround c @@ -633,7 +633,7 @@ let rec as_constr = function | RfHole mv -> EConstr.mkMeta mv | RfGround c -> c | RfApp (f, args) -> EConstr.mkApp (as_constr f, Array.map_of_list as_constr args) -| RfProj (p, c) -> EConstr.mkProj (p, as_constr c) +| RfProj (p, r, c) -> EConstr.mkProj (p, r, as_constr c) (* Old style mk_goal primitive *) let mk_goal evars hyps concl = @@ -678,9 +678,9 @@ let rec mk_refgoals env sigma goalacc conclty trm = match trm with let ((acc'',conclty',sigma), args) = mk_arggoals env sigma acc' hdty l in let ans = EConstr.applist (applicand, args) in (acc'', conclty', sigma, ans) -| RfProj (p, c) -> +| RfProj (p, r, c) -> let (acc',cty,sigma,c') = mk_refgoals env sigma goalacc None c in - let c = EConstr.mkProj (p, c') in + let c = EConstr.mkProj (p, r, c') in let ty = get_type_of env sigma c in (acc',ty,sigma,c) @@ -840,7 +840,11 @@ let build_case_analysis env sigma (ind, u) params pred indices indarg dep knd = in RealCase (ci, u, params, (pnas, pbody), iv, indarg) | Some ps -> - let args = Array.map (fun p -> mkProj (Projection.make p true, indarg)) ps in + let args = Array.map (fun (p,r) -> + let r = UVars.subst_instance_relevance (Unsafe.to_instance u) r in + mkProj (Projection.make p true, r, indarg)) + ps + in PrimitiveEta args let case_pf ?(with_evars=false) ~dep (indarg, typ) = diff --git a/proofs/clenv.mli b/proofs/clenv.mli index 367f9131ca85..fd658b7ce26b 100644 --- a/proofs/clenv.mli +++ b/proofs/clenv.mli @@ -27,7 +27,7 @@ val clenv_evd : clausenv -> Evd.evar_map (* Ad-hoc primitives *) val update_clenv_evd : clausenv -> evar_map -> clausenv val clenv_strip_proj_params : clausenv -> clausenv -val clenv_refresh : env -> evar_map -> Univ.ContextSet.t option -> clausenv -> clausenv +val clenv_refresh : env -> evar_map -> UnivGen.sort_context_set option -> clausenv -> clausenv val clenv_arguments : clausenv -> metavariable list (** subject of clenv (instantiated) *) diff --git a/stm/stm.ml b/stm/stm.ml index 8d5c437d6648..bf60dbae5fc7 100644 --- a/stm/stm.ml +++ b/stm/stm.ml @@ -1723,7 +1723,7 @@ end = struct (* {{{ *) let uc = Univ.hcons_universe_context_set uc in let (pr, priv, ctx) = Option.get (Global.body_of_constant_body Library.indirect_accessor c) in (* We only manipulate monomorphic terms here. *) - let () = assert (Univ.AbstractContext.is_empty ctx) in + let () = assert (UVars.AbstractContext.is_empty ctx) in let () = match priv with | Opaqueproof.PrivateMonomorphic () -> () | Opaqueproof.PrivatePolymorphic uctx -> diff --git a/tactics/autorewrite.ml b/tactics/autorewrite.ml index 1df8cb8be02f..213992bf0661 100644 --- a/tactics/autorewrite.ml +++ b/tactics/autorewrite.ml @@ -201,7 +201,7 @@ struct let a = ca.(len - 1) in let ca = Array.sub ca 0 (len - 1) in Some (DApp, [mkApp (f, ca); a]) - | Proj (p,c) -> + | Proj (p,_,c) -> (* UnsafeMonomorphic is fine because the term will only be used by pat_of_constr which ignores universes *) pat_of_constr (mkApp (UnsafeMonomorphic.mkConst (Projection.constant p), [|c|])) @@ -367,6 +367,7 @@ let one_base where conds tac_main bas = Proofview.Goal.enter begin fun gl -> let sigma = Proofview.Goal.sigma gl in let subst, ctx' = UnivGen.fresh_universe_context_set_instance h.rew_ctx in + let subst = Sorts.QVar.Map.empty, subst in let c' = Vars.subst_univs_level_constr subst h.rew_lemma in let sigma = Evd.merge_context_set Evd.univ_flexible sigma ctx' in Proofview.tclTHEN (Proofview.Unsafe.tclEVARS sigma) (rewrite h.rew_l2r c' tc) diff --git a/tactics/btermdn.ml b/tactics/btermdn.ml index f3308613a57c..8c7b7bad54c9 100644 --- a/tactics/btermdn.ml +++ b/tactics/btermdn.ml @@ -69,7 +69,7 @@ let decomp_pat p = let decomp sigma t = let rec decrec acc c = match EConstr.kind sigma c with | App (f,l) -> decrec (Array.fold_right (fun a l -> a::l) l acc) f - | Proj (p, c) -> + | Proj (p, _, c) -> (* Hack: fake evar to generate [Everything] in the functions below *) let hole = mkEvar (Evar.unsafe_of_int (-1), SList.empty) in let params = List.make (Projection.npars p) hole in diff --git a/tactics/cbn.ml b/tactics/cbn.ml index 67c14bbbd107..9d46dc0b8785 100644 --- a/tactics/cbn.ml +++ b/tactics/cbn.ml @@ -104,14 +104,14 @@ sig type cst_member = | Cst_const of pconstant - | Cst_proj of Projection.t + | Cst_proj of Projection.t * Sorts.relevance type 'a case_stk = case_info * EInstance.t * 'a array * 'a pcase_return * 'a pcase_invert * 'a pcase_branch array type 'a member = | App of 'a app_node | Case of 'a case_stk * Cst_stack.t - | Proj of Projection.t * Cst_stack.t + | Proj of Projection.t * Sorts.relevance * Cst_stack.t | Fix of ('a, 'a) pfixpoint * 'a t * Cst_stack.t | Primitive of CPrimitives.t * (Constant.t * EInstance.t) * 'a t * CPrimitives.args_red * Cst_stack.t | Cst of { const : cst_member; @@ -162,14 +162,14 @@ struct type cst_member = | Cst_const of pconstant - | Cst_proj of Projection.t + | Cst_proj of Projection.t * Sorts.relevance type 'a case_stk = case_info * EInstance.t * 'a array * 'a pcase_return * 'a pcase_invert * 'a pcase_branch array type 'a member = | App of 'a app_node | Case of 'a case_stk * Cst_stack.t - | Proj of Projection.t * Cst_stack.t + | Proj of Projection.t * Sorts.relevance * Cst_stack.t | Fix of ('a, 'a) pfixpoint * 'a t * Cst_stack.t | Primitive of CPrimitives.t * (Constant.t * EInstance.t) * 'a t * CPrimitives.args_red * Cst_stack.t | Cst of { const : cst_member; @@ -192,7 +192,7 @@ struct str "ZCase(" ++ prvect_with_sep (pr_bar) (fun (_, b) -> pr_c b) br ++ str ")" - | Proj (p,cst) -> + | Proj (p,_,cst) -> str "ZProj(" ++ Constant.debug_print (Projection.constant p) ++ str ")" | Fix (f,args,cst) -> str "ZFix(" ++ Constr.debug_print_fix pr_c f @@ -213,10 +213,10 @@ struct let open Pp in match c with | Cst_const (c, u) -> - if Univ.Instance.is_empty u then Constant.debug_print c + if UVars.Instance.is_empty u then Constant.debug_print c else str"(" ++ Constant.debug_print c ++ str ", " ++ - Univ.Instance.pr Univ.Level.raw_pr u ++ str")" - | Cst_proj p -> + UVars.Instance.pr Sorts.QVar.raw_pr Univ.Level.raw_pr u ++ str")" + | Cst_proj (p,r) -> str".(" ++ Constant.debug_print (Projection.constant p) ++ str")" let empty = [] @@ -241,8 +241,8 @@ struct let equal_cst_member x y = match x, y with | Cst_const (c1,u1), Cst_const (c2, u2) -> - Constant.CanOrd.equal c1 c2 && Univ.Instance.equal u1 u2 - | Cst_proj p1, Cst_proj p2 -> Projection.Repr.CanOrd.equal (Projection.repr p1) (Projection.repr p2) + Constant.CanOrd.equal c1 c2 && UVars.Instance.equal u1 u2 + | Cst_proj (p1,_), Cst_proj (p2,_) -> Projection.Repr.CanOrd.equal (Projection.repr p1) (Projection.repr p2) | _, _ -> false in let rec equal_rec sk1 sk2 = @@ -254,7 +254,7 @@ struct (f t1 t2) && (equal_rec s1' s2') | Case ((ci1,pms1,p1,t1,iv1,a1),_) :: s1, Case ((ci2,pms2,p2,iv2,t2,a2),_) :: s2 -> f_case (ci1,pms1,p1,t1,iv1,a1) (ci2,pms2,p2,iv2,t2,a2) && equal_rec s1 s2 - | (Proj (p,_)::s1, Proj(p2,_)::s2) -> + | (Proj (p,_,_)::s1, Proj(p2,_,_)::s2) -> Projection.Repr.CanOrd.equal (Projection.repr p) (Projection.repr p2) && equal_rec s1 s2 | Fix (f1,s1,_) :: s1', Fix (f2,s2,_) :: s2' -> @@ -299,7 +299,7 @@ struct let will_expose_iota args = List.exists (function (Fix (_,_,l) | Case (_,l) | - Proj (_,l) | Cst {cst_l=l}) when Cst_stack.all_volatile l -> true | _ -> false) + Proj (_,_,l) | Cst {cst_l=l}) when Cst_stack.all_volatile l -> true | _ -> false) args let list_of_app_stack s = @@ -348,9 +348,9 @@ struct let constr_of_cst_member f sk = match f with | Cst_const (c, u) -> mkConstU (c, EInstance.make u), sk - | Cst_proj p -> + | Cst_proj (p,r) -> match decomp sk with - | Some (hd, sk) -> mkProj (p, hd), sk + | Some (hd, sk) -> mkProj (p, r, hd), sk | None -> assert false let zip ?(refold=false) sigma s = @@ -372,9 +372,9 @@ struct zip (best_state sigma (constr_of_cst_member const (params @ (append_app [|f|] s))) cst_l) | f, (Cst {const;params}::s) -> zip (constr_of_cst_member const (params @ (append_app [|f|] s))) - | f, (Proj (p,cst_l)::s) when refold -> - zip (best_state sigma (mkProj (p,f),s) cst_l) - | f, (Proj (p,_)::s) -> zip (mkProj (p,f),s) + | f, (Proj (p,r,cst_l)::s) when refold -> + zip (best_state sigma (mkProj (p,r,f),s) cst_l) + | f, (Proj (p,r,_)::s) -> zip (mkProj (p,r,f),s) | f, (Primitive (p,c,args,kargs,cst_l)::s) -> zip (mkConstU c, args @ append_app [|f|] s) in @@ -450,21 +450,28 @@ let magically_constant_of_fixbody env sigma reference bd = function let csts = EConstr.eq_constr_universes env sigma (EConstr.of_constr t) bd in begin match csts with | Some csts -> + let addqs l r (qs,us) = Sorts.QVar.Map.add l r qs, us in + let addus l r (qs,us) = qs, Univ.Level.Map.add l r us in let subst = Set.fold (fun cst acc -> - let l, r = match cst with - | ULub (u, v) | UWeak (u, v) -> u, v + match cst with + | QEq (a,b) | QLeq (a,b) -> + let a = match a with + | QVar q -> q + | _ -> assert false + in + addqs a b acc + | ULub (u, v) | UWeak (u, v) -> addus u v acc | UEq (u, v) | ULe (u, v) -> + (* XXX add something when qsort? *) let get u = match u with | Sorts.SProp | Sorts.Prop -> assert false | Sorts.Set -> Level.set | Sorts.Type u | Sorts.QSort (_, u) -> Option.get (Universe.level u) in - get u, get v - in - Univ.Level.Map.add l r acc) - csts Univ.Level.Map.empty + addus (get u) (get v) acc) + csts UVars.empty_sort_subst in - let inst = subst_univs_level_instance subst u in + let inst = UVars.subst_sort_level_instance subst u in mkConstU (cst, EInstance.make inst) | None -> bd end @@ -664,11 +671,11 @@ let whd_state_gen ?csts flags env sigma = whrec Cst_stack.empty (a,Stack.Primitive(p,const,before,kargs,cst_l)::after) | exception NotEvaluableConst _ -> fold () else fold () - | Proj (p, c) when RedFlags.red_projection flags p -> + | Proj (p, r, c) when RedFlags.red_projection flags p -> (let npars = Projection.npars p in match ReductionBehaviour.get (Projection.constant p) with | None -> - let stack' = (c, Stack.Proj (p, cst_l) :: stack) in + let stack' = (c, Stack.Proj (p, r, cst_l) :: stack) in let stack'', csts = whrec Cst_stack.empty stack' in if equal_stacks sigma stack' stack'' then fold () else stack'', csts @@ -692,12 +699,12 @@ let whd_state_gen ?csts flags env sigma = | [] -> (* if nargs has been specified *) (* CAUTION : the constant is NEVER refold (even when it hides a (co)fix) *) - let stack' = (c, Stack.Proj (p, cst_l) :: stack) in + let stack' = (c, Stack.Proj (p, r, cst_l) :: stack) in whrec Cst_stack.empty(* cst_l *) stack' | curr :: remains -> if curr == 0 then (* Try to reduce the record argument *) let cst_l = Stack.Cst - { const=Stack.Cst_proj p; + { const=Stack.Cst_proj (p,r); volatile; curr; remains; params=Stack.empty; cst_l; @@ -709,7 +716,7 @@ let whd_state_gen ?csts flags env sigma = | None -> fold () | Some (bef,arg,s') -> let cst_l = Stack.Cst - { const=Stack.Cst_proj p; + { const=Stack.Cst_proj (p,r); curr; remains; volatile; @@ -752,7 +759,7 @@ let whd_state_gen ?csts flags env sigma = |args, (Stack.Case(case,_)::s') when use_match -> let r = apply_branch env sigma cstr args case in whrec Cst_stack.empty (r, s') - |args, (Stack.Proj (p,_)::s') when use_match -> + |args, (Stack.Proj (p,_,_)::s') when use_match -> whrec Cst_stack.empty (Stack.nth args (Projection.npars p + Projection.arg p), s') |args, (Stack.Fix (f,s',cst_l)::s'') when use_fix -> let x' = Stack.zip sigma (x, args) in @@ -771,12 +778,12 @@ let whd_state_gen ?csts flags env sigma = let body = EConstr.of_constr body in let cst_l = Cst_stack.add_cst ~volatile (mkConstU const) cst_l in whrec cst_l (body, s' @ (Stack.append_app [|x'|] s''))) - | Stack.Cst_proj p -> + | Stack.Cst_proj (p,r) -> let stack = s' @ (Stack.append_app [|x'|] s'') in match Stack.strip_n_app 0 stack with | None -> assert false | Some (_,arg,s'') -> - whrec Cst_stack.empty (arg, Stack.Proj (p,cst_l) :: s'')) + whrec Cst_stack.empty (arg, Stack.Proj (p,r,cst_l) :: s'')) | next :: remains' -> match Stack.strip_n_app (next-curr-1) s'' with | None -> fold () | Some (bef,arg,s''') -> diff --git a/tactics/elim.ml b/tactics/elim.ml index efd753bf414d..cf258bac1303 100644 --- a/tactics/elim.ml +++ b/tactics/elim.ml @@ -144,7 +144,7 @@ let head_in indl t gl = let decompose_these c l = Proofview.Goal.enter begin fun gl -> - let indl = List.map (fun x -> x, Univ.Instance.empty) l in + let indl = List.map (fun x -> x, UVars.Instance.empty) l in general_decompose (fun env sigma (_,t) -> head_in indl t gl) c end diff --git a/tactics/eqschemes.ml b/tactics/eqschemes.ml index 4a21f3129537..2f67a5e4e791 100644 --- a/tactics/eqschemes.ml +++ b/tactics/eqschemes.ml @@ -72,7 +72,7 @@ let fresh env id avoid = let freshid = next_global_ident_away id avoid in freshid, Id.Set.add freshid avoid let with_context_set ctx (b, ctx') = - (b, Univ.ContextSet.union ctx ctx') + (b, UnivGen.sort_context_union ctx ctx') let build_dependent_inductive ind (mib,mip) = let realargs,_ = List.chop mip.mind_nrealdecls mip.mind_arity_ctxt in @@ -408,7 +408,7 @@ let build_l2r_rew_scheme dep env handle ind kind = rel_vect 1 nrealargs; [|mkRel 1|]]) in let s, ctx' = UnivGen.fresh_sort_in_family kind in - let ctx = Univ.ContextSet.union ctx ctx' in + let ctx = UnivGen.sort_context_union ctx ctx' in let s = mkSort s in let rci = Sorts.Relevant in (* TODO relevance *) let ci = make_case_info env ind rci RegularStyle in @@ -516,7 +516,7 @@ let build_l2r_forward_rew_scheme dep env ind kind = let realsign_ind_P n aP = 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 ctx = UnivGen.sort_context_union ctx ctx' in let s = mkSort s in let rci = Sorts.Relevant in let ci = make_case_info env ind rci RegularStyle in @@ -598,7 +598,7 @@ let build_r2l_forward_rew_scheme dep env ind kind = let realsign_ind = 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 ctx = UnivGen.sort_context_union ctx ctx' in let s = mkSort s in let rci = Sorts.Relevant in (* TODO relevance *) let ci = make_case_info env ind rci RegularStyle in @@ -809,8 +809,11 @@ let build_congr env (eq,refl,ctx) ind = let varf,avoid = fresh env (Id.of_string "f") avoid in let rci = Sorts.Relevant in (* TODO relevance *) let ci = make_case_info env ind rci RegularStyle in - let uni, ctx = Univ.extend_in_context_set (UnivGen.new_global_univ ()) ctx in - let ctx = (fst ctx, UnivSubst.enforce_leq uni (univ_of_eq env eq) (snd ctx)) in + let uni, ctx' = UnivGen.new_global_univ () in + let ctx = + let (qs,us),csts = ctx in + let us, csts = Univ.ContextSet.union (us,csts) ctx' in + ((qs, us), UnivSubst.enforce_leq uni (univ_of_eq env eq) csts) in let c = my_it_mkLambda_or_LetIn paramsctxt (mkNamedLambda (make_annot varB Sorts.Relevant) (mkType uni) @@ -845,4 +848,4 @@ let build_congr env (eq,refl,ctx) ind = let congr_scheme_kind = declare_individual_scheme_object "_congr" (fun env _ ind -> (* May fail if equality is not defined *) - build_congr env (get_coq_eq env Univ.ContextSet.empty) ind) + build_congr env (get_coq_eq env UnivGen.empty_sort_context) ind) diff --git a/tactics/hints.ml b/tactics/hints.ml index bf453f59d092..261d91376c52 100644 --- a/tactics/hints.ml +++ b/tactics/hints.ml @@ -50,7 +50,7 @@ let rec head_bound sigma t = match EConstr.kind sigma t with | Const (c, _) -> GlobRef.ConstRef c | Construct (c, _) -> GlobRef.ConstructRef c | Var id -> GlobRef.VarRef id -| Proj (p, _) -> GlobRef.ConstRef (Projection.constant p) +| Proj (p, _, _) -> GlobRef.ConstRef (Projection.constant p) | Cast (c, _, _) -> head_bound sigma c | Evar _ | Rel _ | Meta _ | Sort _ | Fix _ | Lambda _ | CoFix _ | Int _ | Float _ | Array _ -> raise Bound @@ -70,7 +70,7 @@ let decompose_app_bound sigma t = | Ind (i,u) -> IndRef i, args | Construct (c,u) -> ConstructRef c, args | Var id -> VarRef id, args - | Proj (p, c) -> ConstRef (Projection.constant p), Array.cons c args + | Proj (p, _, c) -> ConstRef (Projection.constant p), Array.cons c args | _ -> raise Bound (** Compute the set of section variables that remain in the named context. @@ -128,7 +128,7 @@ type hints_path = GlobRef.t hints_path_gen type hint_term = | IsGlobRef of GlobRef.t - | IsConstr of constr * Univ.ContextSet.t option (* None if monomorphic *) + | IsConstr of constr * UnivGen.sort_context_set option (* None if monomorphic *) type 'a with_uid = { obj : 'a; @@ -138,14 +138,14 @@ type 'a with_uid = { type raw_hint = { rhint_term : constr; rhint_type : types; - rhint_uctx : Univ.ContextSet.t option; + rhint_uctx : UnivGen.sort_context_set option; rhint_arty : int; (* Number of goals generated by the intended tactic *) } type hint = { hint_term : constr; hint_type : types; - hint_uctx : Univ.ContextSet.t option; (* None if monomorphic *) + hint_uctx : UnivGen.sort_context_set option; (* None if monomorphic *) hint_clnv : Clenv.clausenv; hint_arty : int; (* Number of goals generated by the intended tactic *) } @@ -337,7 +337,7 @@ let lookup_tacs env sigma concl se = let merge_context_set_opt sigma ctx = match ctx with | None -> sigma -| Some ctx -> Evd.merge_context_set Evd.univ_flexible sigma ctx +| Some ctx -> Evd.merge_sort_context_set Evd.univ_flexible sigma ctx let instantiate_hint env sigma p = let mk_clenv { rhint_term = c; rhint_type = cty; rhint_uctx = ctx; rhint_arty = ar } = @@ -1145,7 +1145,7 @@ let subst_autohint (subst, obj) = match t with | None -> gr' | Some t -> - (try head_bound Evd.empty (EConstr.of_constr t.Univ.univ_abstracted_value) + (try head_bound Evd.empty (EConstr.of_constr t.UVars.univ_abstracted_value) with Bound -> gr') in let subst_mps subst c = EConstr.of_constr (subst_mps subst (EConstr.Unsafe.to_constr c)) in @@ -1435,7 +1435,7 @@ let prepare_hint env init (sigma,c) = subst := (evar,mkVar id)::!subst; mkNamedLambda sigma (make_annot id Sorts.Relevant) t (iter (replace_term sigma evar (mkVar id) c)) in let c' = iter c in - let diff = Univ.ContextSet.diff (Evd.universe_context_set sigma) (Evd.universe_context_set init) in + let diff = UnivGen.diff_sort_context (Evd.sort_context_set sigma) (Evd.sort_context_set init) in (c', diff) let warn_non_local_section_hint = @@ -1490,7 +1490,7 @@ let expand_constructor_hints env sigma lems = (fun i -> IsGlobRef (GlobRef.ConstructRef ((ind,i+1)))) | _ -> let (c, ctx) = prepare_hint env sigma (evd,lem) in - let ctx = if Univ.ContextSet.is_empty ctx then None else Some ctx in + let ctx = if UnivGen.is_empty_sort_context ctx then None else Some ctx in [IsConstr (c, ctx)]) lems (* builds a hint database from a constr signature *) (* typically used with (lid, ltyp) = pf_hyps_types *) @@ -1786,9 +1786,9 @@ let fresh_hint env sigma h = | None -> sigma, c | Some ctx -> (* Refresh the instance of the hint *) - let (subst, ctx) = UnivGen.fresh_universe_context_set_instance ctx in + let (subst, ctx) = UnivGen.fresh_sort_context_instance ctx in let c = Vars.subst_univs_level_constr subst c in - let sigma = Evd.merge_context_set Evd.univ_flexible sigma ctx in + let sigma = Evd.merge_sort_context_set Evd.univ_flexible sigma ctx in sigma, c let hint_res_pf ?with_evars ?with_classes ?flags h = diff --git a/tactics/hints.mli b/tactics/hints.mli index e2e6957cdb7e..3ea47a0df80b 100644 --- a/tactics/hints.mli +++ b/tactics/hints.mli @@ -42,7 +42,7 @@ type 'a hint_ast = type hint -val hint_as_term : hint -> Univ.ContextSet.t option * constr +val hint_as_term : hint -> UnivGen.sort_context_set option * constr type 'a hints_path_atom_gen = | PathHints of 'a list @@ -196,7 +196,7 @@ val add_hints : locality:hint_locality -> hint_db_name list -> hints_entry -> un val hint_globref : GlobRef.t -> hint_term -val hint_constr : constr * Univ.ContextSet.t option -> hint_term +val hint_constr : constr * UnivGen.sort_context_set option -> hint_term [@ocaml.deprecated "Declare a hint constant instead"] (** A constr which is Hint'ed will be: diff --git a/tactics/induction.ml b/tactics/induction.ml index 50ff8eb63de9..5f5ee560546c 100644 --- a/tactics/induction.ml +++ b/tactics/induction.ml @@ -460,7 +460,7 @@ let induct_discharge with_evars dests avoid' tac (avoid,ra) names = let expand_projections env sigma c = let rec aux env c = match EConstr.kind sigma c with - | Proj (p, c) -> Retyping.expand_projection env sigma p (aux env c) [] + | Proj (p, _, c) -> Retyping.expand_projection env sigma p (aux env c) [] | _ -> map_constr_with_full_binders env sigma push_rel aux env c in aux env c diff --git a/tactics/rewrite.ml b/tactics/rewrite.ml index 7fd787ae6ae3..1703a0ba6877 100644 --- a/tactics/rewrite.ml +++ b/tactics/rewrite.ml @@ -980,7 +980,7 @@ let fold_match ?(force=false) env sigma c = let unfold_match env sigma sk app = match EConstr.kind sigma app with | App (f', args) when QConstant.equal env (fst (destConst sigma f')) sk -> - let v = Environ.constant_value_in env (sk,Univ.Instance.empty)(*FIXME*) in + let v = Environ.constant_value_in env (sk,UVars.Instance.empty)(*FIXME*) in let v = EConstr.of_constr v in Reductionops.whd_beta env sigma (mkApp (v, args)) | _ -> app @@ -1375,8 +1375,8 @@ module Strategies = let inj_open hint = (); fun sigma -> let (ctx, lemma) = Autorewrite.RewRule.rew_lemma hint in let subst, ctx = UnivGen.fresh_universe_context_set_instance ctx in + let subst = Sorts.QVar.Map.empty, subst in let lemma = Vars.subst_univs_level_constr subst (EConstr.of_constr lemma) in - (* not sure sideff:true is really needed here *) let sigma = Evd.merge_context_set UnivRigid sigma ctx in (sigma, (lemma, NoBindings)) diff --git a/tactics/tactics.ml b/tactics/tactics.ml index e40f4089559c..eac1de09980c 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -1720,9 +1720,9 @@ let make_projection env sigma params cstr sign elim i n c (ind, u) = | Some proj -> let args = Context.Rel.instance mkRel 0 sign in let proj = - match Structures.PrimitiveProjections.find_opt proj with - | Some proj -> - mkProj (Projection.make proj false, mkApp (c, args)) + match Structures.PrimitiveProjections.find_opt_with_relevance (proj,u) with + | Some (proj,r) -> + mkProj (Projection.make proj false, r, mkApp (c, args)) | None -> mkApp (mkConstU (proj,u), Array.append (Array.of_list params) [|mkApp (c, args)|]) @@ -3345,7 +3345,8 @@ let constr_eq ~strict x y = match EConstr.eq_constr_universes env evd x y with | Some csts -> if strict then - if UnivProblem.Set.check (Evd.universes evd) csts then Proofview.tclUNIT () + if UState.check_universe_constraints (Evd.evar_universe_context evd) csts + then Proofview.tclUNIT () else let info = Exninfo.reify () in fail_universes ~info diff --git a/test-suite/bugs/bug_17836_1.v b/test-suite/bugs/bug_17836_1.v new file mode 100644 index 000000000000..4c31ea32c01c --- /dev/null +++ b/test-suite/bugs/bug_17836_1.v @@ -0,0 +1,12 @@ + +Set Universe Polymorphism. + +Section S. + + Let bla : Type. + Proof. + exact Type. + Qed. (* anomaly assert false in declare.ml *) + + +End S. diff --git a/test-suite/bugs/bug_17836_2.v b/test-suite/bugs/bug_17836_2.v new file mode 100644 index 000000000000..6fa554d39ac0 --- /dev/null +++ b/test-suite/bugs/bug_17836_2.v @@ -0,0 +1,6 @@ + +Lemma foo (n:nat) : match n with 0 => 0 | _ => 1 end = 0. +Proof. + set (m := match n with 0 => _ | _ => _ end). + Validate Proof. +Abort. diff --git a/test-suite/bugs/bug_17836_3.v b/test-suite/bugs/bug_17836_3.v new file mode 100644 index 000000000000..d3445853bc80 --- /dev/null +++ b/test-suite/bugs/bug_17836_3.v @@ -0,0 +1,7 @@ +Require Import ssreflect ssrbool. + +Lemma foo b1 : (if b1 then true else false) = b1. +Proof. + rewrite [if _ then _ else _]andbT. + reflexivity. +Qed. diff --git a/test-suite/bugs/bug_7723.v b/test-suite/bugs/bug_7723.v index 216290123193..3c1336a8fd39 100644 --- a/test-suite/bugs/bug_7723.v +++ b/test-suite/bugs/bug_7723.v @@ -45,6 +45,19 @@ Eval vm_compute in foo. End LocalClosure. +Module QVar. + + Definition bar@{q|i|} := Type@{q|i}. + + Definition gbar@{q1 q2|i j|} := bar@{q2|i}. + + Eval vm_compute in gbar. + + Definition gprop := Eval vm_compute in gbar@{Type Prop|Set Set}. + Check eq_refl : gprop = Prop. + +End QVar. + Require Import Hurkens. Polymorphic Inductive unit := tt. 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 b8060240b5e5..f89bcab62c6a 100644 --- a/test-suite/misc/poly-capture-global-univs/src/evilImpl.ml +++ b/test-suite/misc/poly-capture-global-univs/src/evilImpl.ml @@ -2,6 +2,7 @@ open Names let evil name name_f = let open Univ in + let open UVars in let open Constr in let kind = Decls.(IsDefinition Definition) in let u = Level.var 0 in @@ -13,7 +14,7 @@ let evil name name_f = let tc = mkConst tc in let fe = Declare.definition_entry - ~univs:(UState.Polymorphic_entry (UContext.make [|Anonymous|] (Instance.of_array [|u|],Constraints.empty)), UnivNames.empty_binders) + ~univs:(UState.Polymorphic_entry (UContext.make ([||],[|Anonymous|]) (Instance.of_array ([||],[|u|]),Constraints.empty)), UnivNames.empty_binders) ~types:(Term.mkArrowR tc tu) (mkLambda (Context.nameR (Id.of_string "x"), tc, mkRel 1)) in diff --git a/test-suite/output/PrintGrammar.out b/test-suite/output/PrintGrammar.out index 9cb7d2f4a41d..76dc45c68fc2 100644 --- a/test-suite/output/PrintGrammar.out +++ b/test-suite/output/PrintGrammar.out @@ -145,7 +145,8 @@ Entry term is Entry univ_annot is [ LEFTA - [ "@{"; LIST0 universe_level; "}" + [ "@{"; LIST0 univ_level_or_quality; + OPT [ "|"; LIST0 univ_level_or_quality ]; "}" | ] ] Entry fix_decls is diff --git a/test-suite/output/SortQuality.out b/test-suite/output/SortQuality.out index d348761cec52..d229f2e81b18 100644 --- a/test-suite/output/SortQuality.out +++ b/test-suite/output/SortQuality.out @@ -1 +1 @@ -Type@{(* α1 *) SortQuality.1} +Type@{α1 | SortQuality.1} diff --git a/test-suite/success/sort_poly.v b/test-suite/success/sort_poly.v new file mode 100644 index 000000000000..824c7ce90fcd --- /dev/null +++ b/test-suite/success/sort_poly.v @@ -0,0 +1,151 @@ +Set Universe Polymorphism. + +Module Syntax. + Fail Definition foo@{| Set < Set } := Set. + + Definition foo@{u| Set < u} := Type@{u}. + + Definition bar@{s | u | Set < u} := Type@{u}. + Set Printing Universes. + Print bar. + + Definition baz@{s | | } := Type@{s | Set}. + Print baz. + + Definition potato@{s | + | } := Type. + + Check eq_refl : Prop = baz@{Prop | }. + + Inductive bob@{s| |} : Prop := . +End Syntax. + +Module Reduction. + + Definition qsort@{s | u |} := Type@{s | u}. + + Monomorphic Universe U. + + Definition tU := Type@{U}. + Definition qU := qsort@{Type | U}. + + Definition q1 := Eval lazy in qU. + Check eq_refl : q1 = tU. + + Definition q2 := Eval vm_compute in qU. + Check eq_refl : q2 = tU. + + Definition q3 := Eval native_compute in qU. + Check eq_refl : q3 = tU. + + Definition exfalso@{s|u|} (A:Type@{s|u}) (H:False) : A := match H with end. + + Definition exfalsoVM := Eval vm_compute in exfalso@{Type|Set}. + Definition exfalsoNative := Eval native_compute in exfalso@{Type|Set}. + + Fixpoint iter@{s|u|} (A:Type@{s|u}) (f:A -> A) n x := + match n with + | 0 => x + | S k => iter A f k (f x) + end. + + Definition iterType := Eval lazy in iter@{Type|_}. + Definition iterSProp := Eval lazy in iter@{SProp|_}. + +End Reduction. + +Module Conversion. + + Inductive Box@{s|u|} (A:Type@{s|u}) := box (_:A). + + Definition t1@{s|u|} (A:Type@{s|u}) (x y : A) := box _ x. + Definition t2@{s|u|} (A:Type@{s|u}) (x y : A) := box _ y. + + Fail Check eq_refl : t1 nat = t2 nat. + Check fun A:SProp => eq_refl : t1 A = t2 A. + + Definition t (A:SProp) := Eval lazy in t1 A. + + Axiom v@{s| |} : forall (A:Type@{s|Set}), bool -> A. + Fail Check fun P (x:P (v@{Type|} nat true)) => x : P (v nat false). + Check fun (A:SProp) P (x:P (v A true)) => x : P (v A false). + +End Conversion. + +Module Inference. + Definition zog@{s| |} (A:Type@{s|Set}) := A. + + (* implicit instance of zog gets a variable which then gets unified with s from the type of A *) + Definition zag@{s| |} (A:Type@{s|Set}) := zog A. + + (* implicit type of A gets unified to Type@{s|Set} *) + Definition zig@{s| |} A := zog@{s|} A. + + (* Unfortunately casting a hole to a sort (while typing A on the + left of the arrow) produces a rigid univ level. It gets a + constraint "= Set" but rigids don't get substituted away for (bad) + reasons. This is why we need the 2 "+". *) + Definition zig'@{s| + | +} A := A -> zog@{s|} A. + + (* different manually bound sort variables don't unify *) + Fail Definition zog'@{s s'| |} (A:Type@{s|Set}) := zog@{s'|} A. +End Inference. + +Module Inductives. + (* TODO sort variable in the output sort *) + Fail Inductive foo1@{s| |} : Type@{s|Set} := . + + Inductive foo2@{s| |} := Foo2 : Type@{s|Set} -> foo2. + + Inductive foo3@{s| |} (A:Type@{s|Set}) := Foo3 : A -> foo3 A. + + Fail Inductive foo4@{s|u v|v < u} : Type@{v} := C (_:Type@{s|u}). + + Inductive foo5@{s| |} (A:Type@{s|Set}) : Prop := Foo5 (_ : A). + Definition foo5_ind'@{s| |} : forall (A : Type@{s|Set}) (P : Prop), (A -> P) -> foo5 A -> P + := foo5_ind. + + (* TODO more precise squashing *) + Fail Definition foo5_Prop_rect (A:Prop) (P:foo5 A -> Type) + (H : forall a, P (Foo5 A a)) + (f : foo5 A) + : P f + := match f with Foo5 _ a => H a end. + + Set Primitive Projections. + Set Warnings "+records". + + (* the SProp instantiation may not be primitive so the whole thing must be nonprimitive *) + Fail Record R1@{s| |} : Type@{s|Set} := {}. + + (* the Type instantiation may not be primitive *) + Fail Record R2@{s| |} (A:SProp) : Type@{s|Set} := { R2f1 : A }. + + (* R3@{SProp Type|} may not be primitive *) + Fail Record R3@{s s'| |} (A:Type@{s|Set}) : Type@{s'|Set} := { R3f1 : A }. + + (* TODO sort variable in output sort *) + Fail Record R4@{s| |} (A:Type@{s|Set}) : Type@{s|Set} := { R4f1 : A}. + + (* non SProp instantiation must be squashed *) + Fail Record R5@{s| |} (A:Type@{s|Set}) : SProp := { R5f1 : A}. + Fail #[warnings="-non-primitive-record"] + Record R5@{s| |} (A:Type@{s|Set}) : SProp := { R5f1 : A}. + #[warnings="-non-primitive-record,-cannot-define-projection"] + Record R5@{s| |} (A:Type@{s|Set}) : SProp := { R5f1 : A}. + Fail Check R5f1. + Definition R5f1_sprop (A:SProp) (r:R5 A) : A := let (f) := r in f. + Fail Definition R5f1_prop (A:Prop) (r:R5 A) : A := let (f) := r in f. + + Record R6@{s| |} (A:Type@{s|Set}) := { R6f1 : A; R6f2 : nat }. + Check fun (A:SProp) (x y : R6 A) => + eq_refl : Conversion.box _ x.(R6f1 _) = Conversion.box _ y.(R6f1 _). + Fail Check fun (A:Prop) (x y : R6 A) => + eq_refl : Conversion.box _ x.(R6f1 _) = Conversion.box _ y.(R6f1 _). + Fail Check fun (A:SProp) (x y : R6 A) => + eq_refl : Conversion.box _ x.(R6f2 _) = Conversion.box _ y.(R6f2 _). + + #[projections(primitive=no)] Record R7@{s| |} (A:Type@{s|Set}) := { R7f1 : A; R7f2 : nat }. + Check R7@{SProp|} : SProp -> Set. + Check R7@{Type|} : Set -> Set. + +End Inductives. diff --git a/test-suite/success/sort_poly_extraction.v b/test-suite/success/sort_poly_extraction.v new file mode 100644 index 000000000000..09c111ded087 --- /dev/null +++ b/test-suite/success/sort_poly_extraction.v @@ -0,0 +1,10 @@ +Require Extraction. + +Set Universe Polymorphism. +Definition foo@{s| |} := tt. + +Definition bar := foo@{Prop|}. + +Fail Extraction bar. + +(* the actual problem only appears once we have inductives with sort poly output *) diff --git a/user-contrib/Ltac2/Constr.v b/user-contrib/Ltac2/Constr.v index 4bbd6efab5da..9f148b8a72fb 100644 --- a/user-contrib/Ltac2/Constr.v +++ b/user-contrib/Ltac2/Constr.v @@ -17,6 +17,26 @@ Ltac2 @ external type : constr -> constr := "coq-core.plugins.ltac2" "constr_typ Ltac2 @ external equal : constr -> constr -> bool := "coq-core.plugins.ltac2" "constr_equal". (** Strict syntactic equality: only up to α-conversion and evar expansion *) +Module Binder. + +Ltac2 Type relevance_var. +Ltac2 Type relevance := [ Relevant | Irrelevant | RelevanceVar (relevance_var) ]. + +Ltac2 @ external make : ident option -> constr -> binder := "coq-core.plugins.ltac2" "constr_binder_make". +(** Create a binder given the name and the type of the bound variable. + Fails if the type is not a type in the current goal. *) + +Ltac2 @ external unsafe_make : ident option -> relevance -> constr -> binder := "coq-core.plugins.ltac2" "constr_binder_unsafe_make". +(** Create a binder given the name and the type and relevance of the bound variable. *) + +Ltac2 @ external name : binder -> ident option := "coq-core.plugins.ltac2" "constr_binder_name". +(** Retrieve the name of a binder. *) + +Ltac2 @ external type : binder -> constr := "coq-core.plugins.ltac2" "constr_binder_type". +(** Retrieve the type of a binder. *) + +End Binder. + Module Unsafe. (** Low-level access to kernel terms. Use with care! *) @@ -45,7 +65,7 @@ Ltac2 Type kind := [ | Case (case, constr, case_invert, constr, constr array) | Fix (int array, int, binder array, constr array) | CoFix (int, binder array, constr array) -| Proj (projection, constr) +| Proj (projection, Binder.relevance, constr) | Uint63 (uint63) | Float (float) | Array (instance, constr array, constr, constr) @@ -105,25 +125,6 @@ End Case. End Unsafe. -Module Binder. - -Ltac2 Type relevance := [ Relevant | Irrelevant ]. - -Ltac2 @ external make : ident option -> constr -> binder := "coq-core.plugins.ltac2" "constr_binder_make". -(** Create a binder given the name and the type of the bound variable. - Fails if the type is not a type in the current goal. *) - -Ltac2 @ external unsafe_make : ident option -> relevance -> constr -> binder := "coq-core.plugins.ltac2" "constr_binder_unsafe_make". -(** Create a binder given the name and the type and relevance of the bound variable. *) - -Ltac2 @ external name : binder -> ident option := "coq-core.plugins.ltac2" "constr_binder_name". -(** Retrieve the name of a binder. *) - -Ltac2 @ external type : binder -> constr := "coq-core.plugins.ltac2" "constr_binder_type". -(** Retrieve the type of a binder. *) - -End Binder. - Module Cast. Ltac2 @ external default : cast := "coq-core.plugins.ltac2" "constr_cast_default". @@ -183,7 +184,7 @@ Ltac2 is_constructor(c: constr) := Ltac2 is_proj(c: constr) := match Unsafe.kind c with - | Unsafe.Proj _ _ => true + | Unsafe.Proj _ _ _ => true | _ => false end. diff --git a/vernac/assumptions.ml b/vernac/assumptions.ml index 5d148fe1fccf..5189289dc855 100644 --- a/vernac/assumptions.ml +++ b/vernac/assumptions.ml @@ -179,7 +179,7 @@ let fold_with_full_binders g f n acc c = | Lambda (na,t,c) -> f (g (LocalAssum (na,t)) n) (f n acc t) c | LetIn (na,b,t,c) -> f (g (LocalDef (na,b,t)) n) (f n (f n acc b) t) c | App (c,l) -> Array.fold_left (f n) (f n acc c) l - | Proj (_,c) -> f n acc c + | Proj (_,_,c) -> f n acc c | Evar _ -> assert false | Case (ci, u, pms, p, iv, c, bl) -> let mib = lookup_mind (fst ci.ci_ind) in diff --git a/vernac/auto_ind_decl.ml b/vernac/auto_ind_decl.ml index 4dc79cc079b6..55baba62ec0e 100644 --- a/vernac/auto_ind_decl.ml +++ b/vernac/auto_ind_decl.ml @@ -285,7 +285,7 @@ let pred_context env ci params u nas = let realdecls, _ = List.chop mip.mind_nrealdecls mip.mind_arity_ctxt in let self = let args = Context.Rel.instance mkRel 0 mip.mind_arity_ctxt in - let inst = Univ.(Instance.of_array (Array.init (Instance.length u) Level.var)) in + let inst = UVars.Instance.(abstract_instance (length u)) in mkApp (mkIndU (ci.ci_ind, inst), args) in let realdecls = RelDecl.LocalAssum (Context.anonR, self) :: realdecls in @@ -762,8 +762,10 @@ let build_beq_scheme env handle kn = match translate_term_eq env_lift_recparams_fix_nonrecparams_tomatch cc with | None -> raise (EqUnknown "type") (* A supported type should have an eq *) | Some eqA -> - let proj = Projection.make projs.(nb_cstr_args-ndx) true in - (ndx-1,env_lift',mkApp (eqA, [|mkProj (proj,mkRel 2);mkProj (proj,mkRel 1)|])::l) + let proj, relevance = projs.(nb_cstr_args-ndx) in + let proj = Projection.make proj true in + (ndx-1,env_lift',mkApp (eqA, [|mkProj (proj, relevance, mkRel 2); + mkProj (proj, relevance, mkRel 1)|])::l) else raise InternalDependencies) constrs.(0).cs_args (nb_cstr_args,env_lift_recparams_fix_nonrecparams_tomatch,[]) @@ -1158,7 +1160,7 @@ let make_bl_scheme env handle mind = (* Setting universes *) let auctx = Declareops.universes_context mib.mind_universes in let u, uctx = UnivGen.fresh_instance_from auctx None in - let uctx = UState.merge ~sideff:false UState.univ_rigid (UState.from_env env) uctx in + let uctx = UState.merge_sort_context ~sideff:false UState.univ_rigid (UState.from_env env) uctx in let ind = (mind,0) in let nparrec = mib.mind_nparams_rec in @@ -1289,7 +1291,7 @@ let make_lb_scheme env handle mind = (* Setting universes *) let auctx = Declareops.universes_context mib.mind_universes in let u, uctx = UnivGen.fresh_instance_from auctx None in - let uctx = UState.merge ~sideff:false UState.univ_rigid (UState.from_env env) uctx in + let uctx = UState.merge_sort_context ~sideff:false UState.univ_rigid (UState.from_env env) uctx in let nparrec = mib.mind_nparams_rec in let lnonparrec,lnamesparrec = @@ -1476,7 +1478,7 @@ let make_eq_decidability env handle mind = (* Setting universes *) let auctx = Declareops.universes_context mib.mind_universes in let u, uctx = UnivGen.fresh_instance_from auctx None in - let uctx = UState.merge ~sideff:false UState.univ_rigid (UState.from_env env) uctx in + let uctx = UState.merge_sort_context ~sideff:false UState.univ_rigid (UState.from_env env) uctx in let lnonparrec,lnamesparrec = Inductive.inductive_nonrec_rec_paramdecls (mib,u) in diff --git a/vernac/comAssumption.ml b/vernac/comAssumption.ml index bced9ebddb51..ffbfc225a5de 100644 --- a/vernac/comAssumption.ml +++ b/vernac/comAssumption.ml @@ -36,8 +36,8 @@ let declare_variable is_coe ~kind typ univs imps impl {CAst.v=name} = () let instance_of_univ_entry = function - | UState.Polymorphic_entry univs -> Univ.UContext.instance univs - | UState.Monomorphic_entry _ -> Univ.Instance.empty + | UState.Polymorphic_entry univs -> UVars.UContext.instance univs + | UState.Monomorphic_entry _ -> UVars.Instance.empty let declare_axiom is_coe ~poly ~local ~kind ?deprecation typ (univs, ubinders) imps nl {CAst.v=name} = let inl = let open Declaremods in match nl with @@ -70,7 +70,7 @@ let interp_assumption ~program_mode env sigma impl_env bl c = let ty = EConstr.it_mkProd_or_LetIn ty ctx in sigma, ty, impls1@impls2 -let empty_poly_univ_entry = UState.Polymorphic_entry Univ.UContext.empty, UnivNames.empty_binders +let empty_poly_univ_entry = UState.Polymorphic_entry UVars.UContext.empty, UnivNames.empty_binders let empty_mono_univ_entry = UState.Monomorphic_entry Univ.ContextSet.empty, UnivNames.empty_binders let empty_univ_entry ~poly = if poly then empty_poly_univ_entry else empty_mono_univ_entry @@ -94,7 +94,7 @@ let declare_assumptions ~poly ~scope ~kind ?deprecation univs nl l = let refu = match scope with | Locality.Discharge -> declare_variable is_coe ~kind typ univs imps Glob_term.Explicit id; - GlobRef.VarRef id.CAst.v, Univ.Instance.empty + GlobRef.VarRef id.CAst.v, UVars.Instance.empty | Locality.Global local -> declare_axiom is_coe ~local ~poly ~kind ?deprecation typ univs imps nl id in diff --git a/vernac/comAssumption.mli b/vernac/comAssumption.mli index 869ac6e78661..61cd1ab20972 100644 --- a/vernac/comAssumption.mli +++ b/vernac/comAssumption.mli @@ -54,7 +54,7 @@ val declare_axiom -> Impargs.manual_implicits -> Declaremods.inline -> variable CAst.t - -> GlobRef.t * Univ.Instance.t + -> GlobRef.t * UVars.Instance.t (** Context command *) diff --git a/vernac/comHints.ml b/vernac/comHints.ml index a389a6287705..2702ddae64d3 100644 --- a/vernac/comHints.ml +++ b/vernac/comHints.ml @@ -105,11 +105,10 @@ let interp_hints ~poly h = let c, uctx = Constrintern.interp_constr env sigma c in let uctx = UState.normalize_variables uctx in let c = Evarutil.nf_evar (Evd.from_ctx uctx) c in - let diff = UState.context_set uctx in let c = - if poly then (c, Some diff) + if poly then (c, Some (UState.sort_context_set uctx)) else - let () = DeclareUctx.declare_universe_context ~poly:false diff in + let () = Global.push_context_set ~strict:true (UState.context_set uctx) in (c, None) in (Hints.hint_constr c) [@ocaml.warning "-3"] diff --git a/vernac/comInductive.ml b/vernac/comInductive.ml index 18052bde2a86..ce73fb8aab59 100644 --- a/vernac/comInductive.ml +++ b/vernac/comInductive.ml @@ -202,6 +202,24 @@ let is_flexible_sort evd s = match ESorts.kind evd s with | Some l -> Evd.is_flexible_level evd l | None -> false +let include_constructor_argument env evd ~ctor_sort ~inductive_sort = + (* We ignore the quality when comparing the sorts: it has an impact + on squashing in the kernel but cannot cause a universe error. *) + let univ_of_sort s = + match ESorts.kind evd s with + | SProp | Prop -> None + | Set -> Some Univ.Universe.type0 + | Type u | QSort (_,u) -> Some u + in + match univ_of_sort ctor_sort, univ_of_sort inductive_sort with + | _, None -> + (* This function is only called when [s] is not impredicative *) + assert false + | None, Some _ -> evd + | Some uctor, Some uind -> + let mk u = ESorts.make (Sorts.sort_of_univ u) in + Evd.set_leq_sort env evd (mk uctor) (mk uind) + let inductive_levels env evd arities ctors = let inds = List.map2 (fun x ctors -> let ctx, s = Reductionops.dest_arity env evd x in @@ -290,10 +308,7 @@ let inductive_levels env evd arities ctors = if is_impredicative_sort evd s then evd else List.fold_left (List.fold_left (fun evd ctor_sort -> - (* Special behaviour of SProp: - constructors of any inductive can have SProp arguments *) - if ESorts.is_sprop evd ctor_sort then evd - else Evd.set_leq_sort env evd ctor_sort s)) + include_constructor_argument env evd ~ctor_sort ~inductive_sort:s)) evd (Option.List.cons indices ctors)) evd inds in @@ -434,7 +449,7 @@ let check_param = function let restrict_inductive_universes sigma ctx_params arities constructors = let merge_universes_of_constr c = - Univ.Level.Set.union (EConstr.universes_of_constr sigma (EConstr.of_constr c)) in + Univ.Level.Set.union (snd (EConstr.universes_of_constr sigma (EConstr.of_constr c))) in let uvars = Univ.Level.Set.empty in let uvars = Context.Rel.(fold_outside (Declaration.fold_constr merge_universes_of_constr) ctx_params ~init:uvars) in let uvars = List.fold_right merge_universes_of_constr arities uvars in @@ -443,7 +458,7 @@ let restrict_inductive_universes sigma ctx_params arities constructors = let check_trivial_variances variances = Array.iter (function - | None | Some Univ.Variance.Invariant -> () + | None | Some UVars.Variance.Invariant -> () | Some _ -> CErrors.user_err Pp.(strbrk "Universe variance was specified but this inductive will not be cumulative.")) @@ -456,7 +471,7 @@ let variance_of_entry ~cumulative ~variances uctx = if not cumulative then begin check_trivial_variances variances; None end else let lvs = Array.length variances in - let lus = Univ.UContext.size uctx in + let _, lus = UVars.UContext.size uctx in assert (lvs <= lus); Some (Array.append variances (Array.make (lus - lvs) None)) @@ -738,7 +753,7 @@ let do_mutual_inductive ~template udecl indl ~cumulative ~poly ?typing_flags ~pr | Polymorphic_ind_entry uctx -> (UState.Polymorphic_entry uctx, UnivNames.empty_binders) in (* Declare the global universes *) - DeclareUctx.declare_universe_context ~poly:false uctx; + Global.push_context_set ~strict:true uctx; (* Declare the mutual inductive block with its associated schemes *) ignore (DeclareInd.declare_mutual_inductive_with_eliminations ?typing_flags ~indlocs mie binders implicits); (* Declare the possible notations of inductive types *) diff --git a/vernac/comPrimitive.ml b/vernac/comPrimitive.ml index a1640eafb9d6..1c11edf7e897 100644 --- a/vernac/comPrimitive.ml +++ b/vernac/comPrimitive.ml @@ -33,7 +33,7 @@ let do_primitive id udecl prim typopt = let env = Global.env () in let evd, udecl = Constrintern.interp_univ_decl_opt env udecl in let auctx = CPrimitives.op_or_type_univs prim in - let evd, u = Evd.with_context_set UState.univ_flexible evd (UnivGen.fresh_instance auctx) in + let evd, u = Evd.with_sort_context_set UState.univ_flexible evd (UnivGen.fresh_instance auctx) in let expected_typ = EConstr.of_constr @@ Typeops.type_of_prim_or_type env u prim in let evd, (typ,impls) = Constrintern.(interp_type_evars_impls ~impls:empty_internalization_env) @@ -47,9 +47,9 @@ let do_primitive id udecl prim typopt = in Pretyping.check_evars_are_solved ~program_mode:false env evd; let evd = Evd.minimize_universes evd in - let uvars = EConstr.universes_of_constr evd typ in + let _qvars, uvars = EConstr.universes_of_constr evd typ in let evd = Evd.restrict_universe_context evd uvars in let typ = EConstr.to_constr evd typ in - let univ_entry = Evd.check_univ_decl ~poly:(not (Univ.AbstractContext.is_empty auctx)) evd udecl in + let univ_entry = Evd.check_univ_decl ~poly:(not (UVars.AbstractContext.is_empty auctx)) evd udecl in let entry = Declare.primitive_entry ~types:(typ, univ_entry) prim in declare id entry diff --git a/vernac/declare.ml b/vernac/declare.ml index 20d3195b89ac..c29b653437e3 100644 --- a/vernac/declare.ml +++ b/vernac/declare.ml @@ -387,7 +387,7 @@ let declare_constant_core ~name ~typing_flags cd = let ubinders = make_ubinders ctx de.proof_entry_universes in (* We register the global universes after exporting side-effects, since the latter depend on the former. *) - let () = DeclareUctx.declare_universe_context ~poly:false ctx in + let () = Global.push_context_set ~strict:true ctx in let cd = Entries.DefinitionEntry e in ConstantEntry cd, false, ubinders, None else @@ -397,12 +397,12 @@ let declare_constant_core ~name ~typing_flags cd = let de = { de with proof_entry_body = body } in let cd, ctx = cast_opaque_proof_entry EffectEntry de in let ubinders = make_ubinders ctx de.proof_entry_universes in - let () = DeclareUctx.declare_universe_context ~poly:false ctx in + let () = Global.push_context_set ~strict:true ctx in OpaqueEntry cd, false, ubinders, Some (body, feedback_id) | ParameterEntry e -> let univ_entry, ctx = extract_monomorphic (fst e.parameter_entry_universes) in let ubinders = make_ubinders ctx e.parameter_entry_universes in - let () = DeclareUctx.declare_universe_context ~poly:false ctx in + let () = Global.push_context_set ~strict:true ctx in let e = { Entries.parameter_entry_secctx = e.parameter_entry_secctx; Entries.parameter_entry_type = e.parameter_entry_type; @@ -417,7 +417,7 @@ let declare_constant_core ~name ~typing_flags cd = let univ_entry, ctx = extract_monomorphic univs in Some (typ, univ_entry), ubinders, ctx in - let () = DeclareUctx.declare_universe_context ~poly:false ctx in + let () = Global.push_context_set ~strict:true ctx in let e = { Entries.prim_entry_type = typ; Entries.prim_entry_content = e.prim_entry_content; @@ -491,32 +491,41 @@ let declare_variable_core ~name ~kind ~typing_flags d = if Decls.variable_exists name then raise (DeclareUniv.AlreadyDeclared (None, name)); - let impl,opaque,univs = match d with (* Fails if not well-typed *) + let impl,opaque = match d with (* Fails if not well-typed *) | SectionLocalAssum {typ;impl;univs} -> - let poly, uctx = match fst univs with - | UState.Monomorphic_entry uctx -> false, uctx - | UState.Polymorphic_entry uctx -> true, Univ.ContextSet.of_context uctx + let () = match fst univs with + | UState.Monomorphic_entry uctx -> Global.push_context_set ~strict:true uctx + | UState.Polymorphic_entry uctx -> Global.push_section_context uctx in - let () = DeclareUctx.declare_universe_context ~poly uctx in let () = Global.push_named_assum (name,typ) in - impl, true, univs + impl, true | SectionLocalDef { clearbody; entry = de } -> (* The body should already have been forced upstream because it is a section-local definition, but it's not enforced by typing *) let ((body, body_uctx), eff) = Future.force de.proof_entry_body in let () = export_side_effects eff in - let poly, type_uctx = match fst de.proof_entry_universes with - | UState.Monomorphic_entry uctx -> false, uctx - | UState.Polymorphic_entry uctx -> true, Univ.ContextSet.of_context uctx - in - let univs = Univ.ContextSet.union body_uctx type_uctx in (* We must declare the universe constraints before type-checking the term. *) - let () = DeclareUctx.declare_universe_context ~poly univs in + let () = match fst de.proof_entry_universes with + | UState.Monomorphic_entry uctx -> + Global.push_context_set ~strict:true (Univ.ContextSet.union uctx body_uctx) + | UState.Polymorphic_entry uctx -> + Global.push_section_context uctx; + let mk_anon_names u = + let qs, us = UVars.Instance.to_array u in + Array.make (Array.length qs) Anonymous, Array.make (Array.length us) Anonymous + in + Global.push_section_context + (UVars.UContext.of_context_set mk_anon_names Sorts.QVar.Set.empty body_uctx) + in let opaque = de.proof_entry_opaque in let se = if opaque then let cname = Id.of_string (Id.to_string name ^ "_subproof") in let cname = Namegen.next_global_ident_away cname Id.Set.empty in + let poly = match fst de.proof_entry_universes with + | Monomorphic_entry _ -> false + | Polymorphic_entry _ -> true + in let de = { proof_entry_body = Future.from_val ((body, Univ.ContextSet.empty), Evd.empty_side_effects); proof_entry_secctx = None; (* de.proof_entry_secctx is NOT respected *) @@ -532,7 +541,7 @@ let declare_variable_core ~name ~kind ~typing_flags d = (DefinitionEntry de) in { - Entries.secdef_body = Constr.mkConstU (kn, Univ.Instance.empty); + Entries.secdef_body = Constr.mkConstU (kn, UVars.Instance.empty); secdef_type = None; } else { @@ -541,7 +550,7 @@ let declare_variable_core ~name ~kind ~typing_flags d = } in let () = Global.push_named_def (name, se) in (* opaque implies clearbody, so we don't see useless "foo := foo_subproof" in the context *) - Glob_term.Explicit, opaque || clearbody, de.proof_entry_universes + Glob_term.Explicit, opaque || clearbody in Nametab.push (Nametab.Until 1) (Libnames.make_path DirPath.empty name) (GlobRef.VarRef name); Decls.(add_variable_data name {opaque;kind}); @@ -1027,7 +1036,7 @@ let declare_obligation prg obl ~uctx ~types ~body = let body = match fst univs with | UState.Polymorphic_entry uctx -> - Some (DefinedObl (constant, Univ.UContext.instance uctx)) + Some (DefinedObl (constant, UVars.UContext.instance uctx)) | UState.Monomorphic_entry _ -> Some (TermObl @@ -1445,11 +1454,11 @@ let obligation_admitted_terminator ~pm {name; num; auto; check_final} uctx' dref (* The universe context was declared globally, we continue from the new global environment. *) let uctx' = UState.Internal.reboot (Global.env ()) uctx' in - (Univ.Instance.empty, uctx') + (UVars.Instance.empty, uctx') else (* We get the right order somehow, but surely it could be enforced in a clearer way. *) let uctx = UState.context uctx' in - (Univ.UContext.instance uctx, uctx') + (UVars.UContext.instance uctx, uctx') in let obl = {obl with obl_body = Some (DefinedObl (cst, inst))} in let pm = update_program_decl_on_defined ~pm prg obls num obl ~uctx:uctx' rem ~auto in @@ -1955,7 +1964,7 @@ let declare_abstract ~name ~poly ~kind ~sign ~secsign ~opaque ~solve_tac sigma c should be enforced statically. *) let (_, body_uctx), _ = const.proof_entry_body in let () = assert (Univ.ContextSet.is_empty body_uctx) in - EConstr.EInstance.make (Univ.UContext.instance ctx) + EConstr.EInstance.make (UVars.UContext.instance ctx) in let args = List.map EConstr.of_constr args in let lem = EConstr.mkConstU (cst, inst) in diff --git a/vernac/declareUctx.ml b/vernac/declareUctx.ml deleted file mode 100644 index e6a3f784bfb9..000000000000 --- a/vernac/declareUctx.ml +++ /dev/null @@ -1,33 +0,0 @@ -(************************************************************************) -(* * The Coq Proof Assistant / The Coq Development Team *) -(* v * Copyright INRIA, CNRS and contributors *) -(* (* Having Prop/Set/Var as section universes makes no sense *) - assert false - | Some na -> - try - let qid = Nametab.shortest_qualid_of_universe Names.Id.Map.empty na in - Names.Name (Libnames.qualid_basename qid) - with Not_found -> - (* Best-effort naming from the string representation of the level. - See univNames.ml for a similar hack. *) - Names.Name (Names.Id.of_string_soft (Univ.Level.to_string lvl)) - in - Array.map map (Univ.Instance.to_array inst) - -let declare_universe_context ~poly ctx = - if poly then - let uctx = Univ.ContextSet.to_context name_instance ctx in - Global.push_section_context uctx - else - Global.push_context_set ~strict:true ctx diff --git a/vernac/declareUctx.mli b/vernac/declareUctx.mli deleted file mode 100644 index 7ecfab04f234..000000000000 --- a/vernac/declareUctx.mli +++ /dev/null @@ -1,11 +0,0 @@ -(************************************************************************) -(* * The Coq Proof Assistant / The Coq Development Team *) -(* v * Copyright INRIA, CNRS and contributors *) -(* Univ.ContextSet.t -> unit diff --git a/vernac/declareUniv.ml b/vernac/declareUniv.ml index 2acd27932a53..2dc393083b21 100644 --- a/vernac/declareUniv.ml +++ b/vernac/declareUniv.ml @@ -29,8 +29,8 @@ type universe_source = type universe_name_decl = { udecl_src : universe_source; - udecl_named : (Id.t * Univ.UGlobal.t) list; - udecl_anon : Univ.UGlobal.t list; + udecl_named : (Id.t * UGlobal.t) list; + udecl_anon : UGlobal.t list; } let check_exists_universe sp = @@ -112,9 +112,11 @@ let declare_univ_binders gr (univs, pl) = match univs with | UState.Polymorphic_entry _ -> () | UState.Monomorphic_entry (levels, _) -> + let qs, pl = pl in + assert (Id.Map.is_empty qs); (* First the explicitly named universes *) let named, univs = Id.Map.fold (fun id univ (named,univs) -> - let univs = match Univ.Level.name univ with + let univs = match Level.name univ with | None -> assert false (* having Prop/Set/Var as binders is nonsense *) | Some univ -> (id,univ)::univs in @@ -135,12 +137,21 @@ let do_universe ~poly l = (Pp.str"Cannot declare polymorphic universes outside sections.") in let l = List.map (fun {CAst.v=id} -> (id, UnivGen.new_univ_global ())) l in - let ctx = List.fold_left (fun ctx (_,qid) -> Univ.Level.Set.add (Univ.Level.make qid) ctx) - Univ.Level.Set.empty l, Univ.Constraints.empty - in let src = if poly then BoundUniv else UnqualifiedUniv in let () = input_univ_names (src, l, []) in - DeclareUctx.declare_universe_context ~poly ctx + match poly with + | false -> + let ctx = List.fold_left (fun ctx (_,qid) -> Level.Set.add (Level.make qid) ctx) + Level.Set.empty l, Constraints.empty + in + Global.push_context_set ~strict:true ctx + | true -> + let names = CArray.map_of_list (fun (na,_) -> Name na) l in + let us = CArray.map_of_list (fun (_,l) -> Level.make l) l in + let ctx = + UVars.UContext.make ([||],names) (UVars.Instance.of_array ([||],us), Constraints.empty) + in + Global.push_section_context ctx let do_constraint ~poly l = let open Univ in @@ -150,5 +161,13 @@ let do_constraint ~poly l = Constraints.add cst acc) Constraints.empty l in - let uctx = ContextSet.add_constraints constraints ContextSet.empty in - DeclareUctx.declare_universe_context ~poly uctx + match poly with + | false -> + let uctx = ContextSet.add_constraints constraints ContextSet.empty in + Global.push_context_set ~strict:true uctx + | true -> + let uctx = UVars.UContext.make + ([||],[||]) + (UVars.Instance.empty,constraints) + in + Global.push_section_context uctx diff --git a/vernac/declaremods.ml b/vernac/declaremods.ml index 109aebff10d9..cd12e7885bfd 100644 --- a/vernac/declaremods.ml +++ b/vernac/declaremods.ml @@ -148,9 +148,9 @@ end module InterpActions : ModActions with type env = Environ.env - with type typexpr = Constr.t * Univ.AbstractContext.t option = + with type typexpr = Constr.t * UVars.AbstractContext.t option = struct - type typexpr = Constr.t * Univ.AbstractContext.t option + type typexpr = Constr.t * UVars.AbstractContext.t option type env = Environ.env let stage = Summary.Stage.Interp let substobjs_table_name = "MODULE-SUBSTOBJS" @@ -685,7 +685,7 @@ module SynterpVisitor : StagedModS module InterpVisitor : StagedModS with type env = InterpActions.env - with type typexpr = Constr.t * Univ.AbstractContext.t option + with type typexpr = Constr.t * UVars.AbstractContext.t option = StagedMod(InterpActions) (** {6 Modules : start, end, declare} *) diff --git a/vernac/declaremods.mli b/vernac/declaremods.mli index 2a58a0025c19..6126e4fe2f85 100644 --- a/vernac/declaremods.mli +++ b/vernac/declaremods.mli @@ -179,7 +179,7 @@ val debug_print_modtab : unit -> Pp.t bound module hasn't an atomic type. *) val process_module_binding : - MBId.t -> (Constr.t * Univ.AbstractContext.t option) Declarations.module_alg_expr -> unit + MBId.t -> (Constr.t * UVars.AbstractContext.t option) Declarations.module_alg_expr -> unit (** Compatibility layer *) diff --git a/vernac/g_vernac.mlg b/vernac/g_vernac.mlg index f6f903ceee03..4e4c8e69d8ec 100644 --- a/vernac/g_vernac.mlg +++ b/vernac/g_vernac.mlg @@ -209,6 +209,16 @@ let test_variance_ident = lk_kws ["=";"+";"*"] >> lk_ident end +let test_univ_decl = + let open Pcoq.Lookahead in + to_entry "test_univ_decl" (lk_ident_list >> lk_kw "|" >> lk_ident_list >> (lk_kw "+" <+> lk_empty) >> (lk_kw "|" <+> lk_kw "|}")) + +let test_cumul_univ_decl = + let open Pcoq.Lookahead in + let lk_list_variance_ident = lk_list (lk_kw "+" <+> lk_kw "*" <+> lk_kw "=" <+> lk_ident) in + to_entry "test_cumul_univ_decl" (lk_ident_list >> lk_kw "|" >> lk_list_variance_ident >> (lk_kw "+" <+> lk_empty) >> (lk_kw "|" <+> lk_kw "|}")) + + } (* Gallina declarations *) @@ -304,23 +314,36 @@ GRAMMAR EXTEND Gram [ [ l = universe_name; ord = [ "<" -> { Univ.Lt } | "=" -> { Univ.Eq } | "<=" -> { Univ.Le } ]; r = universe_name -> { (l, ord, r) } ] ] ; + univ_decl_constraints: + [ [ "|"; l' = LIST0 univ_constraint SEP ","; + ext = [ "+" -> { true } | -> { false } ]; "}" -> { (l',ext) } + | ext = [ "}" -> { true } | bar_cbrace -> { false } ] -> { ([], ext) } ] ] + ; univ_decl: - [ [ "@{" ; l = LIST0 identref; ext = [ "+" -> { true } | -> { false } ]; - cs = [ "|"; l' = LIST0 univ_constraint SEP ","; - ext = [ "+" -> { true } | -> { false } ]; "}" -> { (l',ext) } - | ext = [ "}" -> { true } | bar_cbrace -> { false } ] -> { ([], ext) } ] + [ [ "@{" ; test_univ_decl; l0 = LIST0 identref; "|"; l = LIST0 identref; ext = [ "+" -> { true } | -> { false } ]; + cs = univ_decl_constraints -> { let open UState in - { univdecl_instance = l; + { univdecl_qualities = l0; + univdecl_instance = l; + univdecl_extensible_instance = ext; + univdecl_constraints = fst cs; + univdecl_extensible_constraints = snd cs } } + | "@{" ; l = LIST0 identref; ext = [ "+" -> { true } | -> { false } ]; + cs = univ_decl_constraints + -> + { let open UState in + { univdecl_qualities = []; (* TODO *) + univdecl_instance = l; univdecl_extensible_instance = ext; univdecl_constraints = fst cs; univdecl_extensible_constraints = snd cs } } ] ] ; variance: - [ [ "+" -> { Univ.Variance.Covariant } - | "=" -> { Univ.Variance.Invariant } - | "*" -> { Univ.Variance.Irrelevant } + [ [ "+" -> { UVars.Variance.Covariant } + | "=" -> { UVars.Variance.Invariant } + | "*" -> { UVars.Variance.Irrelevant } ] ] ; variance_identref: @@ -331,13 +354,23 @@ GRAMMAR EXTEND Gram ] ] ; cumul_univ_decl: - [ [ "@{" ; l = LIST0 variance_identref; ext = [ "+" -> { true } | -> { false } ]; - cs = [ "|"; l' = LIST0 univ_constraint SEP ","; - ext = [ "+" -> { true } | -> { false } ]; "}" -> { (l',ext) } - | ext = [ "}" -> { true } | bar_cbrace -> { false } ] -> { ([], ext) } ] + [ [ "@{" ; test_cumul_univ_decl; l0 = LIST0 identref; "|"; + l = LIST0 variance_identref; ext = [ "+" -> { true } | -> { false } ]; + cs = univ_decl_constraints + -> + { let open UState in + { univdecl_qualities = l0; + univdecl_instance = l; + univdecl_extensible_instance = ext; + univdecl_constraints = fst cs; + univdecl_extensible_constraints = snd cs } } + + | "@{" ; l = LIST0 variance_identref; ext = [ "+" -> { true } | -> { false } ]; + cs = univ_decl_constraints -> { let open UState in - { univdecl_instance = l; + { univdecl_qualities = []; + univdecl_instance = l; univdecl_extensible_instance = ext; univdecl_constraints = fst cs; univdecl_extensible_constraints = snd cs } } @@ -1154,7 +1187,7 @@ GRAMMAR EXTEND Gram ] ] ; univ_name_list: - [ [ "@{" ; l = LIST0 name; "}" -> { l } ] ] + [ [ "@{" ; l = LIST0 name; "}" -> { [],l } ] ] ; END diff --git a/vernac/himsg.ml b/vernac/himsg.ml index fae8d5b20ac3..62ea1f2f1ffe 100644 --- a/vernac/himsg.ml +++ b/vernac/himsg.ml @@ -368,7 +368,10 @@ let explain_unification_error env sigma p1 p2 = function strbrk " is expected to have a functional type but it has type " ++ u] | UnifUnivInconsistency p -> [str "universe inconsistency: " ++ - UGraph.explain_universe_inconsistency (Termops.pr_evd_level sigma) p] + UGraph.explain_universe_inconsistency + (Termops.pr_evd_qvar sigma) + (Termops.pr_evd_level sigma) + p] | CannotSolveConstraint ((pb,env,t,u),e) -> let env = make_all_name_different env sigma in (strbrk "cannot satisfy constraint " ++ pr_leconstr_env env sigma t ++ @@ -771,14 +774,25 @@ let explain_non_linear_unification env sigma m t = let explain_unsatisfied_constraints env sigma cst = strbrk "Unsatisfied constraints: " ++ - Univ.pr_constraints (Termops.pr_evd_level sigma) cst ++ + Univ.Constraints.pr (Termops.pr_evd_level sigma) cst ++ spc () ++ str "(maybe a bugged tactic)." +let explain_unsatisfied_qconstraints env sigma cst = + strbrk "Unsatisfied quality constraints: " ++ + Sorts.QConstraints.pr (Termops.pr_evd_qvar sigma) cst ++ + spc() ++ str "(maybe a bugged tactic)." + let explain_undeclared_universe env sigma l = strbrk "Undeclared universe: " ++ Termops.pr_evd_level sigma l ++ spc () ++ str "(maybe a bugged tactic)." +let explain_undeclared_qualities env sigma l = + let n = Sorts.QVar.Set.cardinal l in + strbrk "Undeclared " ++ str (if n = 1 then "quality" else "qualities") ++ strbrk": " ++ + prlist_with_sep spc (Termops.pr_evd_qvar sigma) (Sorts.QVar.Set.elements l) ++ + spc () ++ str "(maybe a bugged tactic)." + let explain_disallowed_sprop () = Pp.(strbrk "SProp is disallowed because the " ++ str "\"Allow StrictProp\"" @@ -789,7 +803,7 @@ let pr_relevance sigma r = match r with | Sorts.Relevant -> str "relevant" | Sorts.Irrelevant -> str "irrelevant" - | Sorts.RelevanceVar q -> str "a variable " ++ Sorts.QVar.pr q + | Sorts.RelevanceVar q -> str "a variable " ++ (* TODO names *) Sorts.QVar.raw_pr q let pr_binder env sigma = function | LocalAssum (na, t) -> @@ -834,8 +848,8 @@ let explain_bad_invert env = let explain_bad_variance env sigma ~lev ~expected ~actual = str "Incorrect variance for universe " ++ Termops.pr_evd_level sigma lev ++ - str": expected " ++ Univ.Variance.pr expected ++ - str " but cannot be less restrictive than " ++ Univ.Variance.pr actual ++ str "." + str": expected " ++ UVars.Variance.pr expected ++ + str " but cannot be less restrictive than " ++ UVars.Variance.pr actual ++ str "." let explain_undeclared_used_variables env sigma ~declared_vars ~inferred_vars = let l = Id.Set.elements (Id.Set.diff inferred_vars declared_vars) in @@ -893,9 +907,13 @@ let explain_type_error env sigma err = | WrongCaseInfo (ind,ci) -> explain_wrong_case_info env ind ci | UnsatisfiedConstraints cst -> - explain_unsatisfied_constraints env sigma cst + explain_unsatisfied_constraints env sigma cst + | UnsatisfiedQConstraints cst -> + explain_unsatisfied_qconstraints env sigma cst | UndeclaredUniverse l -> - explain_undeclared_universe env sigma l + explain_undeclared_universe env sigma l + | UndeclaredQualities l -> + explain_undeclared_qualities env sigma l | DisallowedSProp -> explain_disallowed_sprop () | BadBinderRelevance (rlv, decl) -> explain_bad_binder_relevance env sigma rlv decl | BadCaseRelevance (rlv, case) -> explain_bad_case_relevance env sigma rlv case @@ -1089,14 +1107,17 @@ let explain_not_match_error = function status (not b) ++ str" declaration was found" | IncompatibleUniverses incon -> str"the universe constraints are inconsistent: " ++ - UGraph.explain_universe_inconsistency UnivNames.pr_with_global_universes incon + UGraph.explain_universe_inconsistency + Sorts.QVar.raw_pr + UnivNames.pr_level_with_global_universes + incon | IncompatiblePolymorphism (env, t1, t2) -> str "conversion of polymorphic values generates additional constraints: " ++ quote (Printer.safe_pr_lconstr_env env (Evd.from_env env) t1) ++ spc () ++ str "compared to " ++ spc () ++ quote (Printer.safe_pr_lconstr_env env (Evd.from_env env) t2) | IncompatibleConstraints { got; expect } -> - let open Univ in + let open UVars in let pr_auctx auctx = let sigma = Evd.from_ctx (UState.of_names @@ -1109,7 +1130,7 @@ let explain_not_match_error = function in str "incompatible polymorphic binders: got" ++ spc () ++ h (pr_auctx got) ++ spc() ++ str "but expected" ++ spc() ++ h (pr_auctx expect) ++ - (if not (Int.equal (AbstractContext.size got) (AbstractContext.size expect)) then mt() else + (if not (UVars.eq_sizes (AbstractContext.size got) (AbstractContext.size expect)) then mt() else fnl() ++ str "(incompatible constraints)") | IncompatibleVariance -> str "incompatible variance information" @@ -1550,7 +1571,10 @@ let _ = CErrors.register_handler (wrap_unhandled explain_exn_default) let rec vernac_interp_error_handler = function | UGraph.UniverseInconsistency i -> str "Universe inconsistency." ++ spc() ++ - UGraph.explain_universe_inconsistency UnivNames.pr_with_global_universes i ++ str "." + UGraph.explain_universe_inconsistency + Sorts.QVar.raw_pr + UnivNames.pr_level_with_global_universes + i ++ str "." | TypeError(env,te) -> let te = map_ptype_error EConstr.of_constr te in explain_type_error env (Evd.from_env env) te diff --git a/vernac/ppvernac.ml b/vernac/ppvernac.ml index 9fac3430ba2d..a91abbc996a7 100644 --- a/vernac/ppvernac.ml +++ b/vernac/ppvernac.ml @@ -63,13 +63,15 @@ let pr_uconstraint (l, d, r) = pr_sort_name_expr l ++ spc () ++ Univ.pr_constraint_type d ++ spc () ++ pr_sort_name_expr r -let pr_univ_name_list = function - | None -> mt () - | Some l -> - str "@{" ++ prlist_with_sep spc pr_lname l ++ str"}" +let pr_full_univ_name_list = function + | None -> mt() + | Some (ql, ul) -> + str "@{" ++ prlist_with_sep spc pr_lname ql ++ + (if List.is_empty ql then mt() else strbrk " | ") ++ + prlist_with_sep spc pr_lname ul ++ str "}" let pr_variance_lident (lid,v) = - let v = Option.cata Univ.Variance.pr (mt()) v in + let v = Option.cata UVars.Variance.pr (mt()) v in v ++ pr_lident lid let pr_univdecl_instance l extensible = @@ -82,7 +84,7 @@ let pr_cumul_univdecl_instance l extensible = let pr_univdecl_constraints l extensible = if List.is_empty l && extensible then mt () - else str"|" ++ spc () ++ prlist_with_sep (fun () -> str",") pr_uconstraint l ++ + else str"|" ++ spc () ++ prlist_with_sep (fun () -> strbrk" | ") pr_uconstraint l ++ (if extensible then str"+" else mt()) let pr_universe_decl l = @@ -622,7 +624,7 @@ let pr_printable = function let pr_subgraph = prlist_with_sep spc pr_qualid in keyword cmd ++ pr_opt pr_subgraph g ++ pr_opt str fopt | PrintName (qid,udecl) -> - keyword "Print" ++ spc() ++ pr_smart_global qid ++ pr_univ_name_list udecl + keyword "Print" ++ spc() ++ pr_smart_global qid ++ pr_full_univ_name_list udecl | PrintModuleType qid -> keyword "Print Module Type" ++ spc() ++ pr_qualid qid | PrintModule qid -> @@ -637,7 +639,7 @@ let pr_printable = function keyword "Print Visibility" ++ pr_opt str s | PrintAbout (qid,l,gopt) -> pr_opt (fun g -> Goal_select.pr_goal_selector g ++ str ":"++ spc()) gopt - ++ keyword "About" ++ spc() ++ pr_smart_global qid ++ pr_univ_name_list l + ++ keyword "About" ++ spc() ++ pr_smart_global qid ++ pr_full_univ_name_list l | PrintImplicit qid -> keyword "Print Implicit" ++ spc() ++ pr_smart_global qid (* spiwack: command printing all the axioms and section variables used in a diff --git a/vernac/prettyp.ml b/vernac/prettyp.ml index 0b6d30f0ecd0..4a4adfcb42fc 100644 --- a/vernac/prettyp.ml +++ b/vernac/prettyp.ml @@ -56,7 +56,7 @@ let print_basename sp = pr_global (GlobRef.ConstRef sp) let print_ref reduce ref udecl = let env = Global.env () in let typ, univs = Typeops.type_of_global_in_context env ref in - let inst = Univ.make_abstract_instance univs in + let inst = UVars.make_abstract_instance univs in let bl = Printer.universe_binders_with_opt_names (Environ.universes_of_global env ref) udecl in let sigma = Evd.from_ctx (UState.of_names bl) in let typ = @@ -193,7 +193,7 @@ let print_if_is_coercion ref = let pr_template_variables = function | [] -> mt () - | vars -> str " on " ++ prlist_with_sep spc UnivNames.pr_with_global_universes vars + | vars -> str " on " ++ prlist_with_sep spc UnivNames.pr_level_with_global_universes vars let print_polymorphism ref = let poly = Global.is_polymorphic ref in @@ -634,7 +634,7 @@ let print_typed_body env evd (val_0,typ) = let print_instance sigma cb = if Declareops.constant_is_polymorphic cb then let univs = Declareops.constant_polymorphic_context cb in - let inst = Univ.make_abstract_instance univs in + let inst = UVars.make_abstract_instance univs in pr_universe_instance sigma inst else mt() diff --git a/vernac/prettyp.mli b/vernac/prettyp.mli index 23fe9589394e..9b9720b17a78 100644 --- a/vernac/prettyp.mli +++ b/vernac/prettyp.mli @@ -41,7 +41,7 @@ val print_safe_judgment : Safe_typing.judgment -> Pp.t val print_name : env -> Evd.evar_map -> qualid Constrexpr.or_by_notation - -> UnivNames.univ_name_list option + -> UnivNames.full_name_list option -> Pp.t val print_notation : env -> Evd.evar_map -> Constrexpr.notation_entry @@ -51,7 +51,7 @@ val print_notation : env -> Evd.evar_map val print_abbreviation : env -> KerName.t -> Pp.t val print_about : env -> Evd.evar_map -> qualid Constrexpr.or_by_notation -> - UnivNames.univ_name_list option -> Pp.t + UnivNames.full_name_list option -> Pp.t val print_impargs : GlobRef.t -> Pp.t (** Pretty-printing functions for classes and coercions *) diff --git a/vernac/printmod.ml b/vernac/printmod.ml index be81b3071c2b..5769ee6c6a41 100644 --- a/vernac/printmod.ml +++ b/vernac/printmod.ml @@ -110,7 +110,7 @@ let print_fields envpar sigma cstrtypes = Printer.pr_lconstr_env envpar sigma c) fields) ++ str" }" let print_one_inductive env sigma isrecord mib ((_,i) as ind) = - let u = Univ.make_abstract_instance (Declareops.inductive_polymorphic_context mib) in + let u = UVars.make_abstract_instance (Declareops.inductive_polymorphic_context mib) in let mip = mib.mind_packets.(i) in let paramdecls = Inductive.inductive_paramdecls (mib,u) in let env_params, params = Namegen.make_all_rel_context_name_different env (Evd.from_env env) (EConstr.of_rel_context paramdecls) in diff --git a/vernac/printmod.mli b/vernac/printmod.mli index e3a5cea782c7..a34ed5f901b9 100644 --- a/vernac/printmod.mli +++ b/vernac/printmod.mli @@ -12,7 +12,6 @@ open Names val pr_mutual_inductive_body : Environ.env -> MutInd.t -> Declarations.mutual_inductive_body -> - UnivNames.univ_name_list option -> Pp.t - + UnivNames.full_name_list option -> Pp.t val print_module : with_body:bool -> ModPath.t -> Pp.t val print_modtype : ModPath.t -> Pp.t diff --git a/vernac/record.ml b/vernac/record.ml index ba22a83484b1..8d50a014f817 100644 --- a/vernac/record.ml +++ b/vernac/record.ml @@ -249,7 +249,7 @@ let typecheck_params_and_fields def poly udecl ps (records : DataI.t list) : tc_ let sigma = Evd.minimize_universes sigma in let uvars = ref Univ.Level.Set.empty in let nf c = - let varsc = EConstr.universes_of_constr sigma c in + let _, varsc = EConstr.universes_of_constr sigma c in let c = EConstr.to_constr sigma c in uvars := Univ.Level.Set.union !uvars varsc; c @@ -393,14 +393,9 @@ let build_named_proj ~primitive ~flags ~poly ~univs ~uinstance ~kind env paramde (* [ccl] is defined in context [params;x:rp] *) (* [ccl'] is defined in context [params;x:rp;x:rp] *) if primitive then - let proj_relevant = match rci with - | Sorts.Irrelevant -> false - | Sorts.Relevant -> true - | Sorts.RelevanceVar _ -> assert false - in let p = Projection.Repr.make indsp - ~proj_relevant ~proj_npars:mib.mind_nparams ~proj_arg:i (Label.of_id fid) in - mkProj (Projection.make p true, mkRel 1), Some p + ~proj_npars:mib.mind_nparams ~proj_arg:i (Label.of_id fid) in + mkProj (Projection.make p true, rci, mkRel 1), Some (p,rci) else let ccl' = liftn 1 2 ccl in let p = mkLambda (x, lift 1 rp, ccl') in @@ -426,9 +421,9 @@ let build_named_proj ~primitive ~flags ~poly ~univs ~uinstance ~kind env paramde in Declare.definition_message fid; let term = match p_opt with - | Some p -> + | Some (p,r) -> let _ = DeclareInd.declare_primitive_projection p kn in - mkProj (Projection.make p false,mkRel 1) + mkProj (Projection.make p false, r, mkRel 1) | None -> let proj_args = (*Rel 1 refers to "x"*) paramargs@[mkRel 1] in match decl with @@ -474,8 +469,8 @@ let declare_projections indsp univs ?(kind=Decls.StructureComponent) inhabitant_ let (mib,mip) = Global.lookup_inductive indsp in let poly = Declareops.inductive_is_polymorphic mib in let uinstance = match fst univs with - | Polymorphic_entry uctx -> Univ.UContext.instance uctx - | Monomorphic_entry -> Univ.Instance.empty + | Polymorphic_entry uctx -> UVars.UContext.instance uctx + | Monomorphic_entry -> UVars.Instance.empty in let paramdecls = Inductive.inductive_paramdecls (mib, uinstance) in let r = mkIndU (indsp,uinstance) in @@ -819,7 +814,7 @@ let interp_structure udecl kind ~template ~cumulative ~poly ~primitive_proj fini interp_structure_core ~cumulative finite ~univs ~variances ~primitive_proj impargs params template ~projections_kind ~indlocs data let declare_structure { Record_decl.mie; primitive_proj; impls; globnames; global_univ_decls; projunivs; ubinders; projections_kind; poly; records; indlocs } = - Option.iter (DeclareUctx.declare_universe_context ~poly:false) global_univ_decls; + Option.iter (Global.push_context_set ~strict:true) global_univ_decls; let kn = DeclareInd.declare_mutual_inductive_with_eliminations mie globnames impls ~primitive_expected:primitive_proj ~indlocs in @@ -862,9 +857,9 @@ let declare_class_constant ~univs paramimpls params data = in let inst, univs = match univs with | UState.Monomorphic_entry _, ubinders -> - Univ.Instance.empty, (UState.Monomorphic_entry Univ.ContextSet.empty, ubinders) + UVars.Instance.empty, (UState.Monomorphic_entry Univ.ContextSet.empty, ubinders) | UState.Polymorphic_entry uctx, _ -> - Univ.UContext.instance uctx, univs + UVars.UContext.instance uctx, univs in let cstu = (cst, inst) in let inst_type = appvectc (mkConstU cstu) (Context.Rel.instance mkRel 0 params) in @@ -935,14 +930,14 @@ let declare_class ~univs params inds def data = let univs, params, fields = match fst univs with | UState.Polymorphic_entry uctx -> - let usubst, auctx = Univ.abstract_universes uctx in - let usubst = Univ.make_instance_subst usubst in + let usubst, auctx = UVars.abstract_universes uctx in + let usubst = UVars.make_instance_subst usubst in let map c = Vars.subst_univs_level_constr usubst c in let fields = Context.Rel.map map fields in let params = Context.Rel.map map params in auctx, params, fields | UState.Monomorphic_entry _ -> - Univ.AbstractContext.empty, params, fields + UVars.AbstractContext.empty, params, fields in let map (impl, projs) = let k = @@ -964,7 +959,7 @@ let add_constant_class cst = let r = (Environ.lookup_constant cst env).const_relevance in let ctx, _ = decompose_prod_decls ty in let args = Context.Rel.instance Constr.mkRel 0 ctx in - let t = mkApp (mkConstU (cst, Univ.make_abstract_instance univs), args) in + let t = mkApp (mkConstU (cst, UVars.make_abstract_instance univs), args) in let tc = { cl_univs = univs; cl_impl = GlobRef.ConstRef cst; @@ -985,7 +980,7 @@ let add_inductive_class ind = let k = let ctx = oneind.mind_arity_ctxt in let univs = Declareops.inductive_polymorphic_context mind in - let inst = Univ.make_abstract_instance univs in + let inst = UVars.make_abstract_instance univs in let ty = Inductive.type_of_inductive ((mind, oneind), inst) in let r = oneind.mind_relevance in { cl_univs = univs; diff --git a/vernac/search.ml b/vernac/search.ml index 6a6253f20fed..12869b162775 100644 --- a/vernac/search.ml +++ b/vernac/search.ml @@ -93,7 +93,7 @@ let generic_search env sigma (fn : GlobRef.t -> Decls.logical_kind option -> env let mib = Global.lookup_mind mind in let iter_packet i mip = let ind = (mind, i) in - let u = Univ.make_abstract_instance (Declareops.inductive_polymorphic_context mib) in + let u = UVars.make_abstract_instance (Declareops.inductive_polymorphic_context mib) in let i = (ind, u) in let typ = Inductiveops.type_of_inductive env i in let () = fn (GlobRef.IndRef ind) None env sigma typ in diff --git a/vernac/vernacentries.ml b/vernac/vernacentries.ml index 490295119a61..b7dafbd7f6f0 100644 --- a/vernac/vernacentries.ml +++ b/vernac/vernacentries.ml @@ -443,7 +443,7 @@ let print_universes ?loc ~sort ~subgraph dst = if Global.is_joined_environment () then mt () else str"There may remain asynchronous universe constraints" in - let prl = UnivNames.pr_with_global_universes in + let prl = UnivNames.pr_level_with_global_universes in begin match dst with | None -> UGraph.pr_universes prl univ ++ pr_remaining | Some s -> dump_universes_gen (fun u -> Pp.string_of_ppcmds (prl u)) univ s diff --git a/vernac/vernacexpr.mli b/vernac/vernacexpr.mli index 579b4cfc5031..64f9c2447456 100644 --- a/vernac/vernacexpr.mli +++ b/vernac/vernacexpr.mli @@ -40,7 +40,7 @@ type printable = | PrintMLLoadPath | PrintMLModules | PrintDebugGC - | PrintName of qualid or_by_notation * UnivNames.univ_name_list option + | PrintName of qualid or_by_notation * UnivNames.full_name_list option | PrintGraph | PrintClasses | PrintTypeclasses @@ -56,7 +56,7 @@ type printable = | PrintScopes | PrintScope of string | PrintVisibility of string option - | PrintAbout of qualid or_by_notation * UnivNames.univ_name_list option * Goal_select.t option + | PrintAbout of qualid or_by_notation * UnivNames.full_name_list option * Goal_select.t option | PrintImplicit of qualid or_by_notation | PrintAssumptions of bool * bool * qualid or_by_notation | PrintStrategy of qualid or_by_notation option