Skip to content

Commit

Permalink
[declare] Move proof information to declare.
Browse files Browse the repository at this point in the history
At this point the record in lemmas was just a stub; next commit will
stop exposing the internals of mutual information, and pave the way
for the refactoring of `Info.t` handling in the Declare interface.
  • Loading branch information
ejgallego committed Jun 26, 2020
1 parent f77743c commit 43d381a
Show file tree
Hide file tree
Showing 35 changed files with 367 additions and 448 deletions.
2 changes: 1 addition & 1 deletion doc/plugin_tutorial/tuto1/src/g_tuto1.mlg
Expand Up @@ -287,7 +287,7 @@ VERNAC COMMAND EXTEND ExploreProof CLASSIFIED AS QUERY
| ![ proof_query ] [ "ExploreProof" ] ->
{ fun ~pstate ->
let sigma, env = Declare.get_current_context pstate in
let pprf = Proof.partial_proof (Declare.Proof.get_proof pstate) in
let pprf = Proof.partial_proof (Declare.Proof.get pstate) in
Feedback.msg_notice
(Pp.prlist_with_sep Pp.fnl (Printer.pr_econstr_env env sigma) pprf)
}
Expand Down
9 changes: 4 additions & 5 deletions plugins/derive/derive.ml
Expand Up @@ -15,7 +15,7 @@ open Context.Named.Declaration
(which can contain references to [f]) in the context extended by
[f:=?x]. When the proof ends, [f] is defined as the value of [?x]
and [lemma] as the proof. *)
let start_deriving f suchthat name : Lemmas.t =
let start_deriving f suchthat name : Declare.Proof.t =

let env = Global.env () in
let sigma = Evd.from_env env in
Expand All @@ -40,8 +40,7 @@ let start_deriving f suchthat name : Lemmas.t =
TNil sigma))))))
in

let info = Lemmas.Info.make ~proof_ending:(Declare.Proof_ending.(End_derive {f; name})) ~kind () in
let info = Declare.Info.make ~proof_ending:(Declare.Proof_ending.(End_derive {f; name})) ~kind () in
let lemma = Lemmas.start_dependent_lemma ~name ~poly ~info goals in
Lemmas.pf_map (Declare.Proof.map_proof begin fun p ->
Util.pi1 @@ Proof.run_tactic env Proofview.(tclFOCUS 1 2 shelve) p
end) lemma
Declare.Proof.map lemma ~f:(fun p ->
Util.pi1 @@ Proof.run_tactic env Proofview.(tclFOCUS 1 2 shelve) p)
2 changes: 1 addition & 1 deletion plugins/derive/derive.mli
Expand Up @@ -16,4 +16,4 @@ val start_deriving
: Names.Id.t
-> Constrexpr.constr_expr
-> Names.Id.t
-> Lemmas.t
-> Declare.Proof.t
4 changes: 2 additions & 2 deletions plugins/extraction/extract_env.ml
Expand Up @@ -729,13 +729,13 @@ let extract_and_compile l =
(* Show the extraction of the current ongoing proof *)
let show_extraction ~pstate =
init ~inner:true false false;
let prf = Declare.Proof.get_proof pstate in
let prf = Declare.Proof.get pstate in
let sigma, env = Declare.get_current_context pstate in
let trms = Proof.partial_proof prf in
let extr_term t =
let ast, ty = extract_constr env sigma t in
let mp = Lib.current_mp () in
let l = Label.of_id (Declare.Proof.get_proof_name pstate) in
let l = Label.of_id (Declare.Proof.get_name pstate) in
let fake_ref = GlobRef.ConstRef (Constant.make2 mp l) in
let decl = Dterm (fake_ref, ast, ty) in
print_one_decl [] mp decl
Expand Down
5 changes: 3 additions & 2 deletions plugins/funind/functional_principles_proofs.ml
Expand Up @@ -856,9 +856,10 @@ let generate_equation_lemma evd fnames f fun_num nb_params nb_args rec_args_num
let lemma =
Lemmas.start_lemma ~name:(mk_equation_id f_id) ~poly:false evd lemma_type
in
let lemma, _ = Lemmas.by (Proofview.V82.tactic prove_replacement) lemma in
let lemma, _ = Declare.by (Proofview.V82.tactic prove_replacement) lemma in
let () =
Lemmas.save_lemma_proved ~lemma ~opaque:Declare.Transparent ~idopt:None
Declare.save_lemma_proved ~proof:lemma ~opaque:Declare.Transparent
~idopt:None
in
evd

Expand Down
12 changes: 6 additions & 6 deletions plugins/funind/gen_principle.ml
Expand Up @@ -1520,10 +1520,10 @@ let derive_correctness (funs : Constr.pconstant list) (graphs : inductive list)
let typ, _ = lemmas_types_infos.(i) in
let lemma = Lemmas.start_lemma ~name:lem_id ~poly:false !evd typ in
let lemma =
fst @@ Lemmas.by (Proofview.V82.tactic (proving_tac i)) lemma
fst @@ Declare.by (Proofview.V82.tactic (proving_tac i)) lemma
in
let () =
Lemmas.save_lemma_proved ~lemma ~opaque:Declare.Transparent
Declare.save_lemma_proved ~proof:lemma ~opaque:Declare.Transparent
~idopt:None
in
let finfo =
Expand Down Expand Up @@ -1586,15 +1586,15 @@ let derive_correctness (funs : Constr.pconstant list) (graphs : inductive list)
in
let lemma =
fst
(Lemmas.by
(Declare.by
(Proofview.V82.tactic
(observe_tac
("prove completeness (" ^ Id.to_string f_id ^ ")")
(proving_tac i)))
lemma)
in
let () =
Lemmas.save_lemma_proved ~lemma ~opaque:Declare.Transparent
Declare.save_lemma_proved ~proof:lemma ~opaque:Declare.Transparent
~idopt:None
in
let finfo =
Expand Down Expand Up @@ -1769,7 +1769,7 @@ let register_mes interactive_proof fname rec_impls wf_mes_expr wf_rel_expr_opt
using_lemmas args ret_type body

let do_generate_principle_aux pconstants on_error register_built
interactive_proof fixpoint_exprl : Lemmas.t option =
interactive_proof fixpoint_exprl : Declare.Proof.t option =
List.iter
(fun {Vernacexpr.notations} ->
if not (List.is_empty notations) then
Expand Down Expand Up @@ -2155,7 +2155,7 @@ let make_graph (f_ref : GlobRef.t) =

(* *************** statically typed entrypoints ************************* *)

let do_generate_principle_interactive fixl : Lemmas.t =
let do_generate_principle_interactive fixl : Declare.Proof.t =
match do_generate_principle_aux [] warning_error true true fixl with
| Some lemma -> lemma
| None ->
Expand Down
2 changes: 1 addition & 1 deletion plugins/funind/gen_principle.mli
Expand Up @@ -12,7 +12,7 @@ val warn_cannot_define_graph : ?loc:Loc.t -> Pp.t * Pp.t -> unit
val warn_cannot_define_principle : ?loc:Loc.t -> Pp.t * Pp.t -> unit

val do_generate_principle_interactive :
Vernacexpr.fixpoint_expr list -> Lemmas.t
Vernacexpr.fixpoint_expr list -> Declare.Proof.t

val do_generate_principle : Vernacexpr.fixpoint_expr list -> unit
val make_graph : Names.GlobRef.t -> unit
Expand Down
32 changes: 16 additions & 16 deletions plugins/funind/recdef.ml
Expand Up @@ -58,7 +58,7 @@ let declare_fun name kind ?univs value =
(Declare.declare_constant ~name ~kind (Declare.DefinitionEntry ce))

let defined lemma =
Lemmas.save_lemma_proved ~lemma ~opaque:Declare.Transparent ~idopt:None
Declare.save_lemma_proved ~proof:lemma ~opaque:Declare.Transparent ~idopt:None

let def_of_const t =
match Constr.kind t with
Expand Down Expand Up @@ -1343,7 +1343,7 @@ let whole_start concl_tac nb_args is_mes func input_type relation rec_arg_num :
g

let get_current_subgoals_types pstate =
let p = Declare.Proof.get_proof pstate in
let p = Declare.Proof.get pstate in
let Proof.{goals = sgs; sigma; _} = Proof.data p in
(sigma, List.map (Goal.V82.abstract_type sigma) sgs)

Expand Down Expand Up @@ -1405,7 +1405,7 @@ let clear_goals sigma =
List.map clear_goal

let build_new_goal_type lemma =
let sigma, sub_gls_types = Lemmas.pf_fold get_current_subgoals_types lemma in
let sigma, sub_gls_types = get_current_subgoals_types lemma in
(* Pp.msgnl (str "sub_gls_types1 := " ++ Util.prlist_with_sep (fun () -> Pp.fnl () ++ Pp.fnl ()) Printer.pr_lconstr sub_gls_types); *)
let sub_gls_types = clear_goals sigma sub_gls_types in
(* Pp.msgnl (str "sub_gls_types2 := " ++ Pp.prlist_with_sep (fun () -> Pp.fnl () ++ Pp.fnl ()) Printer.pr_lconstr sub_gls_types); *)
Expand All @@ -1423,7 +1423,7 @@ let is_opaque_constant c =
let open_new_goal ~lemma build_proof sigma using_lemmas ref_ goal_name
(gls_type, decompose_and_tac, nb_goal) =
(* Pp.msgnl (str "gls_type := " ++ Printer.pr_lconstr gls_type); *)
let current_proof_name = Lemmas.pf_fold Declare.Proof.get_proof_name lemma in
let current_proof_name = Declare.Proof.get_name lemma in
let name =
match goal_name with
| Some s -> s
Expand Down Expand Up @@ -1488,18 +1488,18 @@ let open_new_goal ~lemma build_proof sigma using_lemmas ref_ goal_name
[Hints.Hint_db.empty TransparentState.empty false] ]))
in
let lemma = build_proof env (Evd.from_env env) start_tac end_tac in
Lemmas.save_lemma_proved ~lemma ~opaque:opacity ~idopt:None
Declare.save_lemma_proved ~proof:lemma ~opaque:opacity ~idopt:None
in
let info = Lemmas.Info.make ~hook:(Declare.Hook.make hook) () in
let info = Declare.Info.make ~hook:(Declare.Hook.make hook) () in
let lemma =
Lemmas.start_lemma ~name:na ~poly:false (* FIXME *) ~info sigma gls_type
in
let lemma =
if Indfun_common.is_strict_tcc () then
fst @@ Lemmas.by (Proofview.V82.tactic tclIDTAC) lemma
fst @@ Declare.by (Proofview.V82.tactic tclIDTAC) lemma
else
fst
@@ Lemmas.by
@@ Declare.by
(Proofview.V82.tactic (fun g ->
tclTHEN decompose_and_tac
(tclORELSE
Expand All @@ -1521,27 +1521,26 @@ let open_new_goal ~lemma build_proof sigma using_lemmas ref_ goal_name
g))
lemma
in
if Lemmas.(pf_fold Declare.Proof.get_open_goals) lemma = 0 then (
defined lemma; None )
if Declare.Proof.get_open_goals lemma = 0 then (defined lemma; None)
else Some lemma

let com_terminate interactive_proof tcc_lemma_name tcc_lemma_ref is_mes
fonctional_ref input_type relation rec_arg_num thm_name using_lemmas nb_args
ctx hook =
let start_proof env ctx tac_start tac_end =
let info = Lemmas.Info.make ~hook () in
let info = Declare.Info.make ~hook () in
let lemma =
Lemmas.start_lemma ~name:thm_name ~poly:false (*FIXME*) ~info ctx
(EConstr.of_constr (compute_terminate_type nb_args fonctional_ref))
in
let lemma =
fst
@@ Lemmas.by
@@ Declare.by
(New.observe_tac (fun _ _ -> str "starting_tac") tac_start)
lemma
in
fst
@@ Lemmas.by
@@ Declare.by
(Proofview.V82.tactic
(observe_tac
(fun _ _ -> str "whole_start")
Expand Down Expand Up @@ -1608,7 +1607,7 @@ let com_eqn uctx nb_arg eq_name functional_ref f_ref terminate_ref
in
let lemma =
fst
@@ Lemmas.by
@@ Declare.by
(Proofview.V82.tactic
(start_equation f_ref terminate_ref (fun x ->
prove_eq
Expand Down Expand Up @@ -1642,7 +1641,8 @@ let com_eqn uctx nb_arg eq_name functional_ref f_ref terminate_ref
in
let _ =
Flags.silently
(fun () -> Lemmas.save_lemma_proved ~lemma ~opaque:opacity ~idopt:None)
(fun () ->
Declare.save_lemma_proved ~proof:lemma ~opaque:opacity ~idopt:None)
()
in
()
Expand All @@ -1651,7 +1651,7 @@ let com_eqn uctx nb_arg eq_name functional_ref f_ref terminate_ref

let recursive_definition ~interactive_proof ~is_mes function_name rec_impls
type_of_f r rec_arg_num eq generate_induction_principle using_lemmas :
Lemmas.t option =
Declare.Proof.t option =
let open Term in
let open Constr in
let open CVars in
Expand Down
2 changes: 1 addition & 1 deletion plugins/funind/recdef.mli
Expand Up @@ -25,4 +25,4 @@ val recursive_definition :
-> EConstr.constr
-> unit)
-> Constrexpr.constr_expr list
-> Lemmas.t option
-> Declare.Proof.t option
4 changes: 2 additions & 2 deletions plugins/ltac/extratactics.mlg
Expand Up @@ -918,7 +918,7 @@ END
VERNAC COMMAND EXTEND GrabEvars STATE proof
| [ "Grab" "Existential" "Variables" ]
=> { classify_as_proofstep }
-> { fun ~pstate -> Declare.Proof.map_proof (fun p -> Proof.V82.grab_evars p) pstate }
-> { fun ~pstate -> Declare.Proof.map ~f:(fun p -> Proof.V82.grab_evars p) pstate }
END

(* Shelves all the goals under focus. *)
Expand Down Expand Up @@ -950,7 +950,7 @@ END
VERNAC COMMAND EXTEND Unshelve STATE proof
| [ "Unshelve" ]
=> { classify_as_proofstep }
-> { fun ~pstate -> Declare.Proof.map_proof (fun p -> Proof.unshelve p) pstate }
-> { fun ~pstate -> Declare.Proof.map ~f:(fun p -> Proof.unshelve p) pstate }
END

(* Gives up on the goals under focus: the goals are considered solved,
Expand Down
2 changes: 1 addition & 1 deletion plugins/ltac/g_ltac.mlg
Expand Up @@ -363,7 +363,7 @@ let print_info_trace =

let vernac_solve ~pstate n info tcom b =
let open Goal_select in
let pstate, status = Declare.Proof.map_fold_proof_endline (fun etac p ->
let pstate, status = Declare.Proof.map_fold_endline ~f:(fun etac p ->
let with_end_tac = if b then Some etac else None in
let global = match n with SelectAll | SelectList _ -> true | _ -> false in
let info = Option.append info (print_info_trace ()) in
Expand Down
6 changes: 3 additions & 3 deletions plugins/ltac/rewrite.ml
Expand Up @@ -1978,7 +1978,7 @@ let add_morphism_as_parameter atts m n : unit =
(PropGlobal.proper_class env evd) Hints.empty_hint_info atts.global cst);
declare_projection n instance_id cst

let add_morphism_interactive atts m n : Lemmas.t =
let add_morphism_interactive atts m n : Declare.Proof.t =
init_setoid ();
let instance_id = add_suffix n "_Proper" in
let env = Global.env () in
Expand All @@ -1996,11 +1996,11 @@ let add_morphism_interactive atts m n : Lemmas.t =
| _ -> assert false
in
let hook = Declare.Hook.make hook in
let info = Lemmas.Info.make ~hook ~kind () in
let info = Declare.Info.make ~hook ~kind () in
Flags.silently
(fun () ->
let lemma = Lemmas.start_lemma ~name:instance_id ~poly ~info evd morph in
fst (Lemmas.by (Tacinterp.interp tac) lemma)) ()
fst (Declare.by (Tacinterp.interp tac) lemma)) ()

let add_morphism atts binders m s n =
init_setoid ();
Expand Down
4 changes: 2 additions & 2 deletions plugins/ltac/rewrite.mli
Expand Up @@ -101,7 +101,7 @@ val add_setoid
-> Id.t
-> unit

val add_morphism_interactive : rewrite_attributes -> constr_expr -> Id.t -> Lemmas.t
val add_morphism_interactive : rewrite_attributes -> constr_expr -> Id.t -> Declare.Proof.t
val add_morphism_as_parameter : rewrite_attributes -> constr_expr -> Id.t -> unit

val add_morphism
Expand All @@ -110,7 +110,7 @@ val add_morphism
-> constr_expr
-> constr_expr
-> Id.t
-> Lemmas.t
-> Declare.Proof.t

val get_reflexive_proof : env -> evar_map -> constr -> constr -> evar_map * constr

Expand Down
4 changes: 2 additions & 2 deletions stm/proofBlockDelimiter.ml
Expand Up @@ -49,8 +49,8 @@ let is_focused_goal_simple ~doc id =
match state_of_id ~doc id with
| `Expired | `Error _ | `Valid None -> `Not
| `Valid (Some { Vernacstate.lemmas }) ->
Option.cata (Vernacstate.LemmaStack.with_top_pstate ~f:(fun proof ->
let proof = Declare.Proof.get_proof proof in
Option.cata (Vernacstate.LemmaStack.with_top ~f:(fun proof ->
let proof = Declare.Proof.get proof in
let Proof.{ goals=focused; stack=r1; shelf=r2; given_up=r3; sigma } = Proof.data proof in
let rest = List.(flatten (map (fun (x,y) -> x @ y) r1)) @ r2 @ r3 in
if List.for_all (fun x -> simple_goal sigma x rest) focused
Expand Down
7 changes: 4 additions & 3 deletions stm/stm.ml
Expand Up @@ -1157,7 +1157,8 @@ end = struct (* {{{ *)

let get_proof ~doc id =
match state_of_id ~doc id with
| `Valid (Some vstate) -> Option.map (Vernacstate.LemmaStack.with_top_pstate ~f:Declare.Proof.get_proof) vstate.Vernacstate.lemmas
| `Valid (Some vstate) ->
Option.map (Vernacstate.LemmaStack.with_top ~f:Declare.Proof.get) vstate.Vernacstate.lemmas
| _ -> None

let undo_vernac_classifier v ~doc =
Expand Down Expand Up @@ -1526,7 +1527,7 @@ end = struct (* {{{ *)
try
let _pstate =
stm_qed_delay_proof ~st ~id:stop
~proof:pobject ~info:(Lemmas.Info.make ()) ~loc ~control:[] (Proved (opaque,None)) in
~proof:pobject ~info:(Declare.Info.make ()) ~loc ~control:[] (Proved (opaque,None)) in
()
with exn ->
(* If [stm_qed_delay_proof] fails above we need to use the
Expand Down Expand Up @@ -1672,7 +1673,7 @@ end = struct (* {{{ *)
let proof, _info =
PG_compat.close_proof ~opaque ~keep_body_ucst_separate:true in

let info = Lemmas.Info.make () in
let info = Declare.Info.make () in

(* We jump at the beginning since the kernel handles side effects by also
* looking at the ones that happen to be present in the current env *)
Expand Down
8 changes: 4 additions & 4 deletions user-contrib/Ltac2/tac2entries.ml
Expand Up @@ -808,7 +808,7 @@ let perform_eval ~pstate e =
Goal_select.SelectAll, Proof.start ~name ~poly sigma []
| Some pstate ->
Goal_select.get_default_goal_selector (),
Declare.Proof.get_proof pstate
Declare.Proof.get pstate
in
let v = match selector with
| Goal_select.SelectNth i -> Proofview.tclFOCUS i i v
Expand Down Expand Up @@ -912,15 +912,15 @@ let print_ltac qid =
(** Calling tactics *)

let solve ~pstate default tac =
let pstate, status = Declare.Proof.map_fold_proof_endline begin fun etac p ->
let pstate, status = Declare.Proof.map_fold_endline pstate ~f:(fun etac p ->
let with_end_tac = if default then Some etac else None in
let g = Goal_select.get_default_goal_selector () in
let (p, status) = Proof.solve g None tac ?with_end_tac p in
(* in case a strict subtree was completed,
go back to the top of the prooftree *)
let p = Proof.maximal_unfocus Vernacentries.command_focus p in
p, status
end pstate in
p, status)
in
if not status then Feedback.feedback Feedback.AddedAxiom;
pstate

Expand Down

0 comments on commit 43d381a

Please sign in to comment.