(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) (* Level.equal u u') (Instance.to_array (UContext.instance uctx)) in List.exists check s let push_constraints uctx s = let on_sec sec = if sec.has_poly_univs && Constraint.exists (fun (l,_,r) -> is_polymorphic_univ l s || is_polymorphic_univ r s) (snd uctx) then CErrors.user_err Pp.(str "Cannot add monomorphic constraints which refer to section polymorphic universes."); let uctx' = sec.sec_mono_universes in let sec_mono_universes = (ContextSet.union uctx uctx') in { sec with sec_mono_universes } in on_last_section on_sec s let open_section ~custom sections = let sec = { sec_context = 0; sec_mono_universes = ContextSet.empty; sec_poly_universes = ([||], UContext.empty); has_poly_univs = has_poly_univs sections; sec_entries = []; sec_data = (Cmap.empty, Mindmap.empty); sec_custom = custom; } in sec :: sections let close_section sections = match sections with | sec :: sections -> sections, sec.sec_entries, sec.sec_mono_universes, sec.sec_custom | [] -> CErrors.user_err (Pp.str "No opened section.") let make_decl_univs (nas,uctx) = abstract_universes nas uctx let push_global ~poly e s = if is_empty s then s else if has_poly_univs s && not poly then CErrors.user_err Pp.(str "Cannot add a universe monomorphic declaration when \ section polymorphic universes are present.") else let on_sec sec = { sec with sec_entries = e :: sec.sec_entries; sec_data = add_emap e (make_decl_univs sec.sec_poly_universes) sec.sec_data; } in on_last_section on_sec s let push_constant ~poly con s = push_global ~poly (SecDefinition con) s let push_inductive ~poly ind s = push_global ~poly (SecInductive ind) s type abstr_info = { abstr_ctx : Constr.named_context; abstr_subst : Instance.t; abstr_uctx : AUContext.t; } let empty_segment = { abstr_ctx = []; abstr_subst = Instance.empty; abstr_uctx = AUContext.empty; } let extract_hyps sec vars used = (* Keep the section-local segment of variables *) let vars = List.firstn sec.sec_context vars in (* Only keep the part that is used by the declaration *) List.filter (fun d -> Id.Set.mem (NamedDecl.get_id d) used) vars let section_segment_of_entry vars e hyps sections = (* [vars] are the named hypotheses, [hyps] the subset that is declared by the global *) let with_sec s = match s with | None -> CErrors.user_err (Pp.str "No opened section.") | Some sec -> let hyps = extract_hyps sec vars hyps in let inst, auctx = find_emap e sec.sec_data in { abstr_ctx = hyps; abstr_subst = inst; abstr_uctx = auctx; } in with_last_section with_sec sections let segment_of_constant env con s = let body = Environ.lookup_constant con env in let vars = Environ.named_context env in let used = Context.Named.to_vars body.Declarations.const_hyps in section_segment_of_entry vars (SecDefinition con) used s let segment_of_inductive env mind s = let mib = Environ.lookup_mind mind env in let vars = Environ.named_context env in let used = Context.Named.to_vars mib.Declarations.mind_hyps in section_segment_of_entry vars (SecInductive mind) used s let instance_from_variable_context = List.rev %> List.filter NamedDecl.is_local_assum %> List.map NamedDecl.get_id %> Array.of_list let extract_worklist info = let args = instance_from_variable_context info.abstr_ctx in info.abstr_subst, args let replacement_context env s = let with_sec sec = match sec with | None -> CErrors.user_err (Pp.str "No opened section.") | Some sec -> let cmap, imap = sec.sec_data in let cmap = Cmap.mapi (fun con _ -> extract_worklist @@ segment_of_constant env con s) cmap in let imap = Mindmap.mapi (fun ind _ -> extract_worklist @@ segment_of_inductive env ind s) imap in (cmap, imap) in with_last_section with_sec s let is_in_section env gr s = let with_sec sec = match sec with | None -> false | Some sec -> let open GlobRef in match gr with | VarRef id -> let vars = List.firstn sec.sec_context (Environ.named_context env) in List.exists (fun decl -> Id.equal id (NamedDecl.get_id decl)) vars | ConstRef con -> Cmap.mem con (fst sec.sec_data) | IndRef (ind, _) | ConstructRef ((ind, _), _) -> Mindmap.mem ind (snd sec.sec_data) in with_last_section with_sec s