123456789101112131415161718192021222324252627282930313233343536373839404142434445464748495051525354555657585960616263646566676869707172737475767778798081828384858687888990919293949596979899100101102103104105106107108109110111112113114115116117118119120121122123124125126127128129130131132133134135136137138139140141142143144145146147148149150151152153154155156157158159160161162163164165166167168169170171172173174175176177178(************************************************************************)(* * The Coq Proof Assistant / The Coq Development Team *)(* v * Copyright INRIA, CNRS and contributors *)(* <O___,, * (see version control and CREDITS file for authors & dates) *)(* \VV/ **************************************************************)(* // * This file is distributed under the terms of the *)(* * GNU Lesser General Public License Version 2.1 *)(* * (see LICENSE file for the text of the license) *)(************************************************************************)openPpletstm_pr_errs=Format.eprintf"%s] %s\n%!"(Spawned.process_id())smoduleTacTask:sigtypeoutput=(Constr.constr*UState.t)optiontypetask={t_state:Vernacstate.t;t_assign:outputFuture.assignment->unit;t_ast:ComTactic.interpretable;t_goalno:int;t_goal:Goal.goal;t_kill:unit->unit;t_name:string}includeAsyncTaskQueue.Taskwithtypetask:=taskend=struct(* {{{ *)letforward_feedback{Feedback.doc_id=did;span_id=id;route;contents}=Feedback.feedback~did~id~routecontentstypeoutput=(Constr.constr*UState.t)optiontypetask={t_state:Vernacstate.t;t_assign:outputFuture.assignment->unit;t_ast:ComTactic.interpretable;t_goalno:int;t_goal:Goal.goal;t_kill:unit->unit;t_name:string}typerequest={r_state:Vernacstate.toption;r_ast:ComTactic.interpretable;r_goalno:int;r_goal:Goal.goal;r_name:string}typeresponse=|RespBuiltSubProofof(Constr.constr*UState.t)|RespErrorofPp.t|RespNoProgressletname=ref"tacticworker"letextra_env()=[||]typecompetence=unittypeworker_status=Fresh|Oldofcompetencelettask_match__=true(* run by the master, on a thread *)letrequest_of_taskage{t_state;t_ast;t_goalno;t_goal;t_name}=Some{r_state=ifage<>FreshthenNoneelseSomet_state;r_ast=t_ast;r_goalno=t_goalno;r_goal=t_goal;r_name=t_name}letuse_response_{t_assign;t_kill}resp=matchrespwith|RespBuiltSubProofo->t_assign(`Val(Someo));`Stay((),[])|RespNoProgress->t_assign(`ValNone);t_kill();`Stay((),[])|RespErrormsg->lete=(AsyncTaskQueue.RemoteExceptionmsg,Exninfo.null)int_assign(`Exne);t_kill();`Stay((),[])leton_marshal_errorerr{t_name}=stm_pr_err("Fatal marshal error: "^t_name);flush_all();exit1leton_task_cancellation_or_expiration_or_slave_death=function|Some{t_kill}->t_kill()|_->()letcommand_focus=Proof.new_focus_kind()letfocus_cond=Proof.no_condcommand_focusletstate=refNoneletreceive_state=function|None->()|Somest->state:=Somestletperform{r_state=st;r_ast=tactic;r_goal;r_goalno}=receive_statest;Vernacstate.unfreeze_interp_state(Option.get!state);tryVernacstate.LemmaStack.with_top(Option.get(Option.get!state).Vernacstate.lemmas)~f:(funpstate->letpstate=Declare.Proof.mappstate~f:(Proof.focusfocus_cond()r_goalno)inletpstate=ComTactic.solve~pstateGoal_select.SelectAll~info:Nonetactic~with_end_tac:falseinlet{Proof.sigma}=Declare.Proof.foldpstate~f:Proof.datainmatchEvd.(evar_body(findsigmar_goal))with|Evd.Evar_empty->RespNoProgress|Evd.Evar_definedt->lett=Evarutil.nf_evarsigmatinletevars=Evarutil.undefined_evars_of_termsigmatinifEvar.Set.is_emptyevarsthenlett=EConstr.Unsafe.to_constrtinRespBuiltSubProof(t,Evd.evar_universe_contextsigma)elseCErrors.user_err~hdr:"STM"Pp.(str"The par: selector requires a tactic that makes no progress or fully"++str" solves the goal and leaves no unresolved existential variables. The following"++str" existentials remain unsolved: "++prlist(Termops.pr_existential_keysigma)(Evar.Set.elementsevars)))withewhenCErrors.noncriticale->RespError(CErrors.printe++spc()++str"(for goal "++intr_goalno++str")")letname_of_task{t_name}=t_nameletname_of_request{r_name}=r_nameend(* }}} *)moduleTaskQueue=AsyncTaskQueue.MakeQueue(TacTask)()letassign_tac~abstractres:unitProofview.tactic=Proofview.(Goal.enterbeginfung->letgid=Goal.goalginletg_solution=tryList.assocgidreswithNot_found->CErrors.anomaly(str"Partac: wrong focus.")inifnot(Future.is_overg_solution)thentclUNIT()elseletopenNotationsinmatchFuture.joing_solutionwith|Some(pt,uc)->letpush_statectx=Proofview.tclEVARMAP>>=funsigma->Proofview.Unsafe.tclEVARS(Evd.merge_universe_contextsigmactx)in(ifabstractthenAbstract.tclABSTRACTNoneelse(funx->x))(push_stateuc<*>Tactics.exact_no_check(EConstr.of_constrpt))|None->tclUNIT()end)letenable_par~nworkers=ComTactic.set_par_implementation(fun~pstate~infot_ast~abstract~with_end_tac->lett_state=Vernacstate.freeze_interp_state~marshallable:trueinTaskQueue.with_n_workersnworkersCoqworkmgrApi.High(funqueue->Declare.Proof.mappstate~f:(funp->letopenTacTaskinletresults=(Proof.datap).Proof.goals|>CList.map_i(funig->letg_solution,t_assign=Future.create_delegate~name:(Printf.sprintf"goal %d"i)(funx->x)inTaskQueue.enqueue_taskqueue~cancel_switch:(reffalse){t_state;t_assign;t_ast;t_goalno=i;t_goal=g;t_name=Goal.uidg;t_kill=(fun()->TaskQueue.cancel_allqueue)};g,g_solution)1inTaskQueue.joinqueue;letp,_,()=Proof.run_tactic(Global.env())(assign_tac~abstractresults)pinp)))