aboutsummaryrefslogtreecommitdiff
path: root/plugins
diff options
context:
space:
mode:
authorEmilio Jesus Gallego Arias2020-05-25 14:09:30 +0200
committerEmilio Jesus Gallego Arias2020-06-26 14:38:09 +0200
commit43d381ab20035f64ce2edea8639fcd9e1d0453bc (patch)
treec15fe4f2e490cab93392f77a9597dd7c22f379a0 /plugins
parentf77743c50a29a8d59954881525e00cc28b5b56e8 (diff)
[declare] Move proof information to declare.
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.
Diffstat (limited to 'plugins')
-rw-r--r--plugins/derive/derive.ml9
-rw-r--r--plugins/derive/derive.mli2
-rw-r--r--plugins/extraction/extract_env.ml4
-rw-r--r--plugins/funind/functional_principles_proofs.ml5
-rw-r--r--plugins/funind/gen_principle.ml12
-rw-r--r--plugins/funind/gen_principle.mli2
-rw-r--r--plugins/funind/recdef.ml32
-rw-r--r--plugins/funind/recdef.mli2
-rw-r--r--plugins/ltac/extratactics.mlg4
-rw-r--r--plugins/ltac/g_ltac.mlg2
-rw-r--r--plugins/ltac/rewrite.ml6
-rw-r--r--plugins/ltac/rewrite.mli4
12 files changed, 42 insertions, 42 deletions
diff --git a/plugins/derive/derive.ml b/plugins/derive/derive.ml
index e5665c59b8..633d5dfe3b 100644
--- a/plugins/derive/derive.ml
+++ b/plugins/derive/derive.ml
@@ -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
@@ -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)
diff --git a/plugins/derive/derive.mli b/plugins/derive/derive.mli
index ef94c7e78f..06e7dacd36 100644
--- a/plugins/derive/derive.mli
+++ b/plugins/derive/derive.mli
@@ -16,4 +16,4 @@ val start_deriving
: Names.Id.t
-> Constrexpr.constr_expr
-> Names.Id.t
- -> Lemmas.t
+ -> Declare.Proof.t
diff --git a/plugins/extraction/extract_env.ml b/plugins/extraction/extract_env.ml
index a0627dbe63..39c6c4ea54 100644
--- a/plugins/extraction/extract_env.ml
+++ b/plugins/extraction/extract_env.ml
@@ -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
diff --git a/plugins/funind/functional_principles_proofs.ml b/plugins/funind/functional_principles_proofs.ml
index b864b18887..652f942cb9 100644
--- a/plugins/funind/functional_principles_proofs.ml
+++ b/plugins/funind/functional_principles_proofs.ml
@@ -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
diff --git a/plugins/funind/gen_principle.ml b/plugins/funind/gen_principle.ml
index 608155eb71..2336689a66 100644
--- a/plugins/funind/gen_principle.ml
+++ b/plugins/funind/gen_principle.ml
@@ -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 =
@@ -1586,7 +1586,7 @@ 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 ^ ")")
@@ -1594,7 +1594,7 @@ let derive_correctness (funs : Constr.pconstant list) (graphs : inductive list)
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 =
@@ -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
@@ -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 ->
diff --git a/plugins/funind/gen_principle.mli b/plugins/funind/gen_principle.mli
index 3c04d6cb7d..28751c4501 100644
--- a/plugins/funind/gen_principle.mli
+++ b/plugins/funind/gen_principle.mli
@@ -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
diff --git a/plugins/funind/recdef.ml b/plugins/funind/recdef.ml
index 9b2d9c4815..f92d4c6a72 100644
--- a/plugins/funind/recdef.ml
+++ b/plugins/funind/recdef.ml
@@ -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
@@ -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)
@@ -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); *)
@@ -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
@@ -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
@@ -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")
@@ -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
@@ -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
()
@@ -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
diff --git a/plugins/funind/recdef.mli b/plugins/funind/recdef.mli
index 4e5146e37c..2612f2b63e 100644
--- a/plugins/funind/recdef.mli
+++ b/plugins/funind/recdef.mli
@@ -25,4 +25,4 @@ val recursive_definition :
-> EConstr.constr
-> unit)
-> Constrexpr.constr_expr list
- -> Lemmas.t option
+ -> Declare.Proof.t option
diff --git a/plugins/ltac/extratactics.mlg b/plugins/ltac/extratactics.mlg
index ffb597d4cb..40c64a1c26 100644
--- a/plugins/ltac/extratactics.mlg
+++ b/plugins/ltac/extratactics.mlg
@@ -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. *)
@@ -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,
diff --git a/plugins/ltac/g_ltac.mlg b/plugins/ltac/g_ltac.mlg
index 996f6b3eb3..114acaa412 100644
--- a/plugins/ltac/g_ltac.mlg
+++ b/plugins/ltac/g_ltac.mlg
@@ -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
diff --git a/plugins/ltac/rewrite.ml b/plugins/ltac/rewrite.ml
index 4bc8d61258..e784d6e682 100644
--- a/plugins/ltac/rewrite.ml
+++ b/plugins/ltac/rewrite.ml
@@ -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
@@ -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 ();
diff --git a/plugins/ltac/rewrite.mli b/plugins/ltac/rewrite.mli
index 1161c84e6a..60a66dd861 100644
--- a/plugins/ltac/rewrite.mli
+++ b/plugins/ltac/rewrite.mli
@@ -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
@@ -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