(************************************************************************) (* * The Coq Proof Assistant / The Coq Development Team *) (* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) (* EConstr.mkVar) ctxt in let ev = EConstr.mkEvar (evk,inst) in (evk, ev, evars) (* Instantiates a goal with an open term *) let partial_solution sigma evk c = (* Check that the goal itself does not appear in the refined term *) let _ = if not (Evarutil.occur_evar_upto sigma evk c) then () else Pretype_errors.error_occur_check Environ.empty_env sigma evk c in Evd.define evk c sigma (* Instantiates a goal with an open term, using name of goal for evk' *) let partial_solution_to sigma evk evk' c = let id = Evd.evar_ident evk sigma in let sigma = partial_solution sigma evk c in match id with | None -> sigma | Some id -> Evd.rename evk' id sigma (* Parts of the progress tactical *) let same_goal evars1 gl1 evars2 gl2 = let evi1 = Evd.find evars1 gl1 in let evi2 = Evd.find evars2 gl2 in let c1 = EConstr.Unsafe.to_constr evi1.Evd.evar_concl in let c2 = EConstr.Unsafe.to_constr evi2.Evd.evar_concl in Constr.equal c1 c2 && Environ.eq_named_context_val evi1.Evd.evar_hyps evi2.Evd.evar_hyps let weak_progress glss gls = match glss.Evd.it with | [ g ] -> not (same_goal glss.Evd.sigma g gls.Evd.sigma gls.Evd.it) | _ -> true let progress glss gls = weak_progress glss gls (* spiwack: progress normally goes like this: (Evd.progress_evar_map gls.Evd.sigma glss.Evd.sigma) || (weak_progress glss gls) This is immensly slow in the current implementation. Maybe we could reimplement progress_evar_map with restricted folds like "fold_undefined", with a good implementation of them. *) (* Used by the compatibility layer and typeclasses *) let nf_evar sigma gl = let evi = Evd.find sigma gl in let evi = Evarutil.nf_evar_info sigma evi in let sigma = Evd.add sigma gl evi in (gl, sigma) (* Goal represented as a type, doesn't take into account section variables *) let abstract_type sigma gl = let open EConstr in let (gl,sigma) = nf_evar sigma gl in let env = env sigma gl in let genv = Global.env () in let is_proof_var decl = try ignore (Environ.lookup_named (NamedDecl.get_id decl) genv); false with Not_found -> true in Environ.fold_named_context_reverse (fun t decl -> if is_proof_var decl then let decl = Termops.map_named_decl EConstr.of_constr decl in mkNamedProd_or_LetIn decl t else t ) ~init:(concl sigma gl) env end