diff options
| author | Pierre-Marie Pédrot | 2020-05-12 14:07:16 +0200 |
|---|---|---|
| committer | Pierre-Marie Pédrot | 2020-05-12 14:41:03 +0200 |
| commit | b3b967385830daa09a149e093fa6352a99884436 (patch) | |
| tree | 7deb98cedd6fce8b3a9f99610769d08809447092 /proofs/logic.ml | |
| parent | 5530ac6ee1fd533a2b56944bd9e16b0d767f3e61 (diff) | |
Clean up the legacy refiner implementation.
We avoid using global-access primitives and instead rely on purely functional
passing of the relevant data. Namely, we replace the goal argument with its
environment, and we pass the additional check parameter.
Diffstat (limited to 'proofs/logic.ml')
| -rw-r--r-- | proofs/logic.ml | 112 |
1 files changed, 51 insertions, 61 deletions
diff --git a/proofs/logic.ml b/proofs/logic.ml index eb0ce23ab7..812fd1d769 100644 --- a/proofs/logic.ml +++ b/proofs/logic.ml @@ -71,13 +71,11 @@ let catchable_exception = function let error_no_such_hypothesis env sigma id = raise (RefinerError (env, sigma, NoSuchHyp id)) -(* Tells if the refiner should check that the submitted rules do not - produce invalid subgoals *) -let check = ref false -let with_check = Flags.with_option check +(* The check flag tells if the refiner should check that the submitted rules do + not produce invalid subgoals *) -let check_typability env sigma c = - if !check then fst (type_of env sigma (EConstr.of_constr c)) else sigma +let check_typability ~check env sigma c = + if check then fst (type_of env sigma (EConstr.of_constr c)) else sigma (************************************************************************) (************************************************************************) @@ -314,8 +312,8 @@ let check_meta_variables env sigma c = if not (List.distinct_f Int.compare (collect_meta_variables c)) then raise (RefinerError (env, sigma, NonLinearProof c)) -let check_conv_leq_goal env sigma arg ty conclty = - if !check then +let check_conv_leq_goal ~check env sigma arg ty conclty = + if check then let ans = Reductionops.infer_conv env sigma (EConstr.of_constr ty) conclty in match ans with | Some evm -> evm @@ -332,28 +330,27 @@ let meta_free_prefix sigma a = in a with Stop acc -> Array.rev_of_list acc -let goal_type_of env sigma c = - if !check then +let goal_type_of ~check env sigma c = + if check then let (sigma,t) = type_of env sigma (EConstr.of_constr c) in (sigma, EConstr.Unsafe.to_constr t) else (sigma, EConstr.Unsafe.to_constr (Retyping.get_type_of env sigma (EConstr.of_constr c))) -let rec mk_refgoals sigma goal goalacc conclty trm = - let env = Goal.V82.env sigma goal in - let hyps = Goal.V82.hyps sigma goal in +let rec mk_refgoals ~check env sigma goalacc conclty trm = + let hyps = Environ.named_context_val env in let mk_goal hyps concl = Goal.V82.mk_goal sigma hyps concl in - if (not !check) && not (occur_meta sigma (EConstr.of_constr trm)) then + if (not check) && not (occur_meta sigma (EConstr.of_constr trm)) then let t'ty = Retyping.get_type_of env sigma (EConstr.of_constr trm) in let t'ty = EConstr.Unsafe.to_constr t'ty in - let sigma = check_conv_leq_goal env sigma trm t'ty conclty in + let sigma = check_conv_leq_goal ~check env sigma trm t'ty conclty in (goalacc,t'ty,sigma,trm) else match kind trm with | Meta _ -> let conclty = nf_betaiota env sigma conclty in - if !check && occur_meta sigma conclty then + if check && occur_meta sigma conclty then raise (RefinerError (env, sigma, MetaInType conclty)); let (gl,ev,sigma) = mk_goal hyps conclty in let ev = EConstr.Unsafe.to_constr ev in @@ -361,9 +358,9 @@ let rec mk_refgoals sigma goal goalacc conclty trm = gl::goalacc, conclty, sigma, ev | Cast (t,k, ty) -> - let sigma = check_typability env sigma ty in - let sigma = check_conv_leq_goal env sigma trm ty conclty in - let res = mk_refgoals sigma goal goalacc (EConstr.of_constr ty) t in + let sigma = check_typability ~check env sigma ty in + let sigma = check_conv_leq_goal ~check env sigma trm ty conclty in + let res = mk_refgoals ~check env sigma goalacc (EConstr.of_constr ty) t in (* we keep the casts (in particular VMcast and NATIVEcast) except when they are annotating metas *) if isMeta t then begin @@ -386,24 +383,24 @@ let rec mk_refgoals sigma goal goalacc conclty trm = let ty = EConstr.Unsafe.to_constr ty in goalacc, ty, sigma, f else - mk_hdgoals sigma goal goalacc f + mk_hdgoals ~check env sigma goalacc f in - let ((acc'',conclty',sigma), args) = mk_arggoals sigma goal acc' hdty l in - let sigma = check_conv_leq_goal env sigma trm conclty' conclty in + let ((acc'',conclty',sigma), args) = mk_arggoals ~check env sigma acc' hdty l in + let sigma = check_conv_leq_goal ~check env sigma trm conclty' conclty in let ans = if applicand == f && args == l then trm else mkApp (applicand, args) in (acc'',conclty',sigma, ans) | Proj (p,c) -> - let (acc',cty,sigma,c') = mk_hdgoals sigma goal goalacc c in + let (acc',cty,sigma,c') = mk_hdgoals ~check env sigma goalacc c in let c = mkProj (p, c') in let ty = get_type_of env sigma (EConstr.of_constr c) in let ty = EConstr.Unsafe.to_constr ty in (acc',ty,sigma,c) | Case (ci,p,c,lf) -> - let (acc',lbrty,conclty',sigma,p',c') = mk_casegoals sigma goal goalacc p c in - let sigma = check_conv_leq_goal env sigma trm conclty' conclty in - let (acc'',sigma,rbranches) = treat_case sigma goal ci lbrty lf acc' in + let (acc',lbrty,conclty',sigma,p',c') = mk_casegoals ~check env sigma goalacc p c in + let sigma = check_conv_leq_goal ~check env sigma trm conclty' conclty in + let (acc'',sigma,rbranches) = treat_case ~check env sigma ci lbrty lf acc' in let lf' = Array.rev_of_list rbranches in let ans = if p' == p && c' == c && Array.equal (==) lf' lf then trm @@ -414,28 +411,27 @@ let rec mk_refgoals sigma goal goalacc conclty trm = | _ -> if occur_meta sigma (EConstr.of_constr trm) then anomaly (Pp.str "refiner called with a meta in non app/case subterm."); - let (sigma, t'ty) = goal_type_of env sigma trm in - let sigma = check_conv_leq_goal env sigma trm t'ty conclty in + let (sigma, t'ty) = goal_type_of ~check env sigma trm in + let sigma = check_conv_leq_goal ~check env sigma trm t'ty conclty in (goalacc,t'ty,sigma, trm) (* Same as mkREFGOALS but without knowing the type of the term. Therefore, * Metas should be casted. *) -and mk_hdgoals sigma goal goalacc trm = - let env = Goal.V82.env sigma goal in - let hyps = Goal.V82.hyps sigma goal in +and mk_hdgoals ~check env sigma goalacc trm = + let hyps = Environ.named_context_val env in let mk_goal hyps concl = Goal.V82.mk_goal sigma hyps concl in match kind trm with | Cast (c,_, ty) when isMeta c -> - let sigma = check_typability env sigma ty in + let sigma = check_typability ~check env sigma ty in let (gl,ev,sigma) = mk_goal hyps (nf_betaiota env sigma (EConstr.of_constr ty)) in let ev = EConstr.Unsafe.to_constr ev in gl::goalacc,ty,sigma,ev | Cast (t,_, ty) -> - let sigma = check_typability env sigma ty in - mk_refgoals sigma goal goalacc (EConstr.of_constr ty) t + let sigma = check_typability ~check env sigma ty in + mk_refgoals ~check env sigma goalacc (EConstr.of_constr ty) t | App (f,l) -> let (acc',hdty,sigma,applicand) = @@ -443,15 +439,15 @@ and mk_hdgoals sigma goal goalacc trm = then let l' = meta_free_prefix sigma l in (goalacc,EConstr.Unsafe.to_constr (type_of_global_reference_knowing_parameters env sigma (EConstr.of_constr f) l'),sigma,f) - else mk_hdgoals sigma goal goalacc f + else mk_hdgoals ~check env sigma goalacc f in - let ((acc'',conclty',sigma), args) = mk_arggoals sigma goal acc' hdty l in + let ((acc'',conclty',sigma), args) = mk_arggoals ~check env sigma acc' hdty l in let ans = if applicand == f && args == l then trm else mkApp (applicand, args) in (acc'',conclty',sigma, ans) | Case (ci,p,c,lf) -> - let (acc',lbrty,conclty',sigma,p',c') = mk_casegoals sigma goal goalacc p c in - let (acc'',sigma,rbranches) = treat_case sigma goal ci lbrty lf acc' in + let (acc',lbrty,conclty',sigma,p',c') = mk_casegoals ~check env sigma goalacc p c in + let (acc'',sigma,rbranches) = treat_case ~check env sigma ci lbrty lf acc' in let lf' = Array.rev_of_list rbranches in let ans = if p' == p && c' == c && Array.equal (==) lf' lf then trm @@ -460,21 +456,21 @@ and mk_hdgoals sigma goal goalacc trm = (acc'',conclty',sigma, ans) | Proj (p,c) -> - let (acc',cty,sigma,c') = mk_hdgoals sigma goal goalacc c in + let (acc',cty,sigma,c') = mk_hdgoals ~check env sigma goalacc c in let c = mkProj (p, c') in let ty = get_type_of env sigma (EConstr.of_constr c) in let ty = EConstr.Unsafe.to_constr ty in (acc',ty,sigma,c) | _ -> - if !check && occur_meta sigma (EConstr.of_constr trm) then + if check && occur_meta sigma (EConstr.of_constr trm) then anomaly (Pp.str "refine called with a dependent meta."); - let (sigma, ty) = goal_type_of env sigma trm in + let (sigma, ty) = goal_type_of env ~check sigma trm in goalacc, ty, sigma, trm -and mk_arggoals sigma goal goalacc funty allargs = +and mk_arggoals ~check env sigma goalacc funty allargs = let foldmap (goalacc, funty, sigma) harg = - let t = whd_all (Goal.V82.env sigma goal) sigma (EConstr.of_constr funty) in + let t = whd_all env sigma (EConstr.of_constr funty) in let t = EConstr.Unsafe.to_constr t in let rec collapse t = match kind t with | LetIn (_, c1, _, b) -> collapse (subst1 c1 b) @@ -483,19 +479,17 @@ and mk_arggoals sigma goal goalacc funty allargs = let t = collapse t in match kind t with | Prod (_, c1, b) -> - let (acc, hargty, sigma, arg) = mk_refgoals sigma goal goalacc (EConstr.of_constr c1) harg in + let (acc, hargty, sigma, arg) = mk_refgoals ~check env sigma goalacc (EConstr.of_constr c1) harg in (acc, subst1 harg b, sigma), arg | _ -> - let env = Goal.V82.env sigma goal in raise (RefinerError (env,sigma,CannotApply (t, harg))) in Array.Smart.fold_left_map foldmap (goalacc, funty, sigma) allargs -and mk_casegoals sigma goal goalacc p c = - let env = Goal.V82.env sigma goal in - let (acc',ct,sigma,c') = mk_hdgoals sigma goal goalacc c in +and mk_casegoals ~check env sigma goalacc p c = + let (acc',ct,sigma,c') = mk_hdgoals ~check env sigma goalacc c in let ct = EConstr.of_constr ct in - let (acc'',pt,sigma,p') = mk_hdgoals sigma goal acc' p in + let (acc'',pt,sigma,p') = mk_hdgoals ~check env sigma acc' p in let ((ind, u), spec) = try Tacred.find_hnf_rectype env sigma ct with Not_found -> anomaly (Pp.str "mk_casegoals.") in @@ -503,20 +497,19 @@ and mk_casegoals sigma goal goalacc p c = let (lbrty,conclty) = type_case_branches_with_names env sigma indspec p c in (acc'',lbrty,conclty,sigma,p',c') -and treat_case sigma goal ci lbrty lf acc' = +and treat_case ~check env sigma ci lbrty lf acc' = let rec strip_outer_cast c = match kind c with | Cast (c,_,_) -> strip_outer_cast c | _ -> c in let decompose_app_vect c = match kind c with | App (f,cl) -> (f, cl) | _ -> (c,[||]) in - let env = Goal.V82.env sigma goal in Array.fold_left3 (fun (lacc,sigma,bacc) ty fi l -> if isMeta (strip_outer_cast fi) then (* Support for non-eta-let-expanded Meta as found in *) (* destruct/case with an non eta-let expanded elimination scheme *) - let (r,_,s,fi') = mk_refgoals sigma goal lacc ty fi in + let (r,_,s,fi') = mk_refgoals ~check env sigma lacc ty fi in r,s,(fi'::bacc) else (* Deal with a branch in expanded form of the form @@ -537,14 +530,14 @@ and treat_case sigma goal ci lbrty lf acc' = if isMeta head then begin assert (args = Context.Rel.to_extended_vect mkRel 0 ctx); let head' = lift (-n) head in - let (r,_,s,head'') = mk_refgoals sigma goal lacc ty head' in + let (r,_,s,head'') = mk_refgoals ~check env sigma lacc ty head' in let fi' = it_mkLambda_or_LetIn (mkApp (head'',args)) ctx in (r,s,fi'::bacc) end else (* Supposed to be meta-free *) - let sigma, t'ty = goal_type_of env sigma fi in - let sigma = check_conv_leq_goal env sigma fi t'ty ty in + let sigma, t'ty = goal_type_of ~check env sigma fi in + let sigma = check_conv_leq_goal ~check env sigma fi t'ty ty in (lacc,sigma,fi::bacc)) (acc',sigma,[]) lbrty lf ci.ci_pp_info.cstr_tags @@ -572,20 +565,17 @@ let convert_hyp ~check ~reorder env sigma d = (************************************************************************) (* Primitive tactics are handled here *) -let prim_refiner r sigma goal = +let prim_refiner ~check r sigma goal = let env = Goal.V82.env sigma goal in let cl = Goal.V82.concl sigma goal in check_meta_variables env sigma r; - let (sgl,cl',sigma,oterm) = mk_refgoals sigma goal [] cl r in + let (sgl,cl',sigma,oterm) = mk_refgoals ~check env sigma [] cl r in let sgl = List.rev sgl in let sigma = Goal.V82.partial_solution env sigma goal (EConstr.of_constr oterm) in { Evd.it = sgl; Evd.sigma = sigma; } let refiner ~check r { Evd.sigma = sigma; Evd.it = goal } = - if check then - with_check (prim_refiner r sigma) goal - else - prim_refiner r sigma goal + prim_refiner ~check r sigma goal (* Profiling refiner *) let refiner ~check = |
