diff options
| author | Emilio Jesus Gallego Arias | 2020-04-02 01:41:20 -0400 |
|---|---|---|
| committer | Emilio Jesus Gallego Arias | 2020-04-15 11:12:10 -0400 |
| commit | 0c748e670ef1a81cb35c1cc55892dba141c25930 (patch) | |
| tree | d90cc4362019557a74d6d1ac46701e0d6178e8ce /plugins | |
| parent | 9908ce57e15a316ff692ce31f493651734998ded (diff) | |
[proof] Merge `Proof_global` into `Declare`
We place creation and saving of interactive proofs in the same module;
this will allow to make `proof_entry` private, improving invariants
and control over clients, and to reduce the API [for example next
commit will move abstract declaration into this module, removing the
exported ad-hoc `build_constant_by_tactic`]
Next step will be to unify all the common code in the interactive /
non-interactive case; but we need to tweak the handling of obligations
first.
Diffstat (limited to 'plugins')
| -rw-r--r-- | plugins/derive/derive.ml | 2 | ||||
| -rw-r--r-- | plugins/extraction/extract_env.ml | 6 | ||||
| -rw-r--r-- | plugins/extraction/extract_env.mli | 2 | ||||
| -rw-r--r-- | plugins/funind/functional_principles_proofs.ml | 2 | ||||
| -rw-r--r-- | plugins/funind/gen_principle.ml | 15 | ||||
| -rw-r--r-- | plugins/funind/recdef.ml | 16 | ||||
| -rw-r--r-- | plugins/ltac/extratactics.mlg | 6 | ||||
| -rw-r--r-- | plugins/ltac/g_ltac.mlg | 2 |
8 files changed, 26 insertions, 25 deletions
diff --git a/plugins/derive/derive.ml b/plugins/derive/derive.ml index dca69f06ca..3eb5057b85 100644 --- a/plugins/derive/derive.ml +++ b/plugins/derive/derive.ml @@ -42,6 +42,6 @@ let start_deriving f suchthat name : Lemmas.t = let info = Lemmas.Info.make ~proof_ending:(Lemmas.Proof_ending.(End_derive {f; name})) ~kind () in let lemma = Lemmas.start_dependent_lemma ~name ~poly ~info goals in - Lemmas.pf_map (Proof_global.map_proof begin fun p -> + Lemmas.pf_map (Declare.map_proof begin fun p -> Util.pi1 @@ Proof.run_tactic env Proofview.(tclFOCUS 1 2 shelve) p end) lemma diff --git a/plugins/extraction/extract_env.ml b/plugins/extraction/extract_env.ml index 40dbf8bea4..7f8e8b36ad 100644 --- a/plugins/extraction/extract_env.ml +++ b/plugins/extraction/extract_env.ml @@ -728,13 +728,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 = Proof_global.get_proof pstate in - let sigma, env = Proof_global.get_current_context pstate in + let prf = Declare.get_proof 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 (Proof_global.get_proof_name pstate) in + let l = Label.of_id (Declare.get_proof_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/extraction/extract_env.mli b/plugins/extraction/extract_env.mli index edbc1f5ea7..2e1509c5cc 100644 --- a/plugins/extraction/extract_env.mli +++ b/plugins/extraction/extract_env.mli @@ -40,4 +40,4 @@ val structure_for_compute : (* Show the extraction of the current ongoing proof *) -val show_extraction : pstate:Proof_global.t -> unit +val show_extraction : pstate:Declare.t -> unit diff --git a/plugins/funind/functional_principles_proofs.ml b/plugins/funind/functional_principles_proofs.ml index ad0891b567..7b2ce671a3 100644 --- a/plugins/funind/functional_principles_proofs.ml +++ b/plugins/funind/functional_principles_proofs.ml @@ -847,7 +847,7 @@ let generate_equation_lemma evd fnames f fun_num nb_params nb_args rec_args_num in let lemma, _ = Lemmas.by (Proofview.V82.tactic prove_replacement) lemma in let () = - Lemmas.save_lemma_proved ~lemma ~opaque:Proof_global.Transparent ~idopt:None + Lemmas.save_lemma_proved ~lemma ~opaque:Declare.Transparent ~idopt:None in evd diff --git a/plugins/funind/gen_principle.ml b/plugins/funind/gen_principle.ml index df147b3aa6..eec78391af 100644 --- a/plugins/funind/gen_principle.ml +++ b/plugins/funind/gen_principle.ml @@ -191,7 +191,7 @@ let prepare_body {Vernacexpr.binders} rt = let fun_args, rt' = chop_rlambda_n n rt in (fun_args, rt') -let build_functional_principle ?(opaque = Proof_global.Transparent) +let build_functional_principle ?(opaque = Declare.Transparent) (evd : Evd.evar_map ref) old_princ_type sorts funs _i proof_tac hook = (* First we get the type of the old graph principle *) let mutr_nparams = @@ -233,9 +233,10 @@ let build_functional_principle ?(opaque = Proof_global.Transparent) (* let dur1 = System.time_difference tim1 tim2 in *) (* Pp.msgnl (str ("Time to compute proof: ") ++ str (string_of_float dur1)); *) (* end; *) - let open Proof_global in - let {entries} = - Lemmas.pf_fold (close_proof ~opaque ~keep_body_ucst_separate:false) lemma + let {Declare.entries} = + Lemmas.pf_fold + (Declare.close_proof ~opaque ~keep_body_ucst_separate:false) + lemma in match entries with | [entry] -> (entry, hook) @@ -1394,7 +1395,7 @@ let make_scheme evd (fas : (Constr.pconstant * Sorts.family) list) : | None -> raise Not_found | Some finfos -> finfos in - let open Proof_global in + let open Declare in match finfos.equation_lemma with | None -> Transparent (* non recursive definition *) | Some equation -> @@ -1550,7 +1551,7 @@ let derive_correctness (funs : Constr.pconstant list) (graphs : inductive list) fst @@ Lemmas.by (Proofview.V82.tactic (proving_tac i)) lemma in let () = - Lemmas.save_lemma_proved ~lemma ~opaque:Proof_global.Transparent + Lemmas.save_lemma_proved ~lemma ~opaque:Declare.Transparent ~idopt:None in let finfo = @@ -1621,7 +1622,7 @@ let derive_correctness (funs : Constr.pconstant list) (graphs : inductive list) lemma) in let () = - Lemmas.save_lemma_proved ~lemma ~opaque:Proof_global.Transparent + Lemmas.save_lemma_proved ~lemma ~opaque:Declare.Transparent ~idopt:None in let finfo = diff --git a/plugins/funind/recdef.ml b/plugins/funind/recdef.ml index bd19648c08..51d3286715 100644 --- a/plugins/funind/recdef.ml +++ b/plugins/funind/recdef.ml @@ -58,7 +58,7 @@ let declare_fun name kind ?univs value = GlobRef.ConstRef (declare_constant ~name ~kind (DefinitionEntry ce)) let defined lemma = - Lemmas.save_lemma_proved ~lemma ~opaque:Proof_global.Transparent ~idopt:None + Lemmas.save_lemma_proved ~lemma ~opaque:Declare.Transparent ~idopt:None let def_of_const t = match Constr.kind t with @@ -1336,7 +1336,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 = Proof_global.get_proof pstate in + let p = Declare.get_proof pstate in let Proof.{goals = sgs; sigma; _} = Proof.data p in (sigma, List.map (Goal.V82.abstract_type sigma) sgs) @@ -1408,15 +1408,15 @@ let build_new_goal_type lemma = let is_opaque_constant c = let cb = Global.lookup_constant c in match cb.Declarations.const_body with - | Declarations.OpaqueDef _ -> Proof_global.Opaque - | Declarations.Undef _ -> Proof_global.Opaque - | Declarations.Def _ -> Proof_global.Transparent - | Declarations.Primitive _ -> Proof_global.Opaque + | Declarations.OpaqueDef _ -> Declare.Opaque + | Declarations.Undef _ -> Declare.Opaque + | Declarations.Def _ -> Declare.Transparent + | Declarations.Primitive _ -> Declare.Opaque 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 Proof_global.get_proof_name lemma in + let current_proof_name = Lemmas.pf_fold Declare.get_proof_name lemma in let name = match goal_name with | Some s -> s @@ -1514,7 +1514,7 @@ let open_new_goal ~lemma build_proof sigma using_lemmas ref_ goal_name g)) lemma in - if Lemmas.(pf_fold Proof_global.get_open_goals) lemma = 0 then ( + if Lemmas.(pf_fold Declare.get_open_goals) lemma = 0 then ( defined lemma; None ) else Some lemma diff --git a/plugins/ltac/extratactics.mlg b/plugins/ltac/extratactics.mlg index 7b1aa7a07a..c2023289a4 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 -> Proof_global.map_proof (fun p -> Proof.V82.grab_evars p) pstate } + -> { fun ~pstate -> Declare.map_proof (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 -> Proof_global.map_proof (fun p -> Proof.unshelve p) pstate } + -> { fun ~pstate -> Declare.map_proof (fun p -> Proof.unshelve p) pstate } END (* Gives up on the goals under focus: the goals are considered solved, @@ -1102,7 +1102,7 @@ END VERNAC COMMAND EXTEND OptimizeProof | ![ proof ] [ "Optimize" "Proof" ] => { classify_as_proofstep } -> - { fun ~pstate -> Proof_global.compact_the_proof pstate } + { fun ~pstate -> Declare.compact_the_proof pstate } | [ "Optimize" "Heap" ] => { classify_as_proofstep } -> { Gc.compact () } END diff --git a/plugins/ltac/g_ltac.mlg b/plugins/ltac/g_ltac.mlg index 2bd4211c90..0a46522543 100644 --- a/plugins/ltac/g_ltac.mlg +++ b/plugins/ltac/g_ltac.mlg @@ -364,7 +364,7 @@ let print_info_trace = let vernac_solve ~pstate n info tcom b = let open Goal_select in - let pstate, status = Proof_global.map_fold_proof_endline (fun etac p -> + let pstate, status = Declare.map_fold_proof_endline (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 |
