aboutsummaryrefslogtreecommitdiff
path: root/plugins
diff options
context:
space:
mode:
Diffstat (limited to 'plugins')
-rw-r--r--plugins/derive/g_derive.mlg4
-rw-r--r--plugins/extraction/g_extraction.mlg10
-rw-r--r--plugins/funind/g_indfun.mlg25
-rw-r--r--plugins/funind/indfun.ml11
-rw-r--r--plugins/funind/indfun.mli2
-rw-r--r--plugins/ltac/extratactics.mlg6
-rw-r--r--plugins/ltac/g_ltac.mlg8
-rw-r--r--plugins/ltac/g_obligations.mlg16
-rw-r--r--plugins/ltac/g_rewrite.mlg55
-rw-r--r--plugins/ltac/rewrite.ml67
-rw-r--r--plugins/ltac/rewrite.mli35
11 files changed, 117 insertions, 122 deletions
diff --git a/plugins/derive/g_derive.mlg b/plugins/derive/g_derive.mlg
index ee076f5ae3..d25b8fd01c 100644
--- a/plugins/derive/g_derive.mlg
+++ b/plugins/derive/g_derive.mlg
@@ -23,6 +23,6 @@ let classify_derive_command _ = Vernacextend.(VtStartProof (Doesn'tGuaranteeOpac
}
VERNAC COMMAND EXTEND Derive CLASSIFIED BY { classify_derive_command }
-| ![ proof ] [ "Derive" ident(f) "SuchThat" constr(suchthat) "As" ident(lemma) ] ->
- { fun ~pstate -> Some (Proof_global.push ~ontop:pstate Derive.(start_deriving f suchthat lemma)) }
+| ![ open_proof ] [ "Derive" ident(f) "SuchThat" constr(suchthat) "As" ident(lemma) ] ->
+ { Derive.(start_deriving f suchthat lemma) }
END
diff --git a/plugins/extraction/g_extraction.mlg b/plugins/extraction/g_extraction.mlg
index d43d90af60..487accbc9b 100644
--- a/plugins/extraction/g_extraction.mlg
+++ b/plugins/extraction/g_extraction.mlg
@@ -178,12 +178,6 @@ END
(* Show the extraction of the current proof *)
VERNAC COMMAND EXTEND ShowExtraction CLASSIFIED AS QUERY
-| ![ proof ] [ "Show" "Extraction" ]
- -> { fun ~pstate:ontop ->
- let pstate = Option.map Proof_global.get_current_pstate ontop in
- let pstate = match pstate with
- | None -> CErrors.user_err Pp.(str "No ongoing proof")
- | Some pstate -> pstate
- in
- let () = show_extraction ~pstate in ontop }
+| ![ proof_query ] [ "Show" "Extraction" ]
+ -> { show_extraction }
END
diff --git a/plugins/funind/g_indfun.mlg b/plugins/funind/g_indfun.mlg
index b8db3dc5ce..cc772e96f3 100644
--- a/plugins/funind/g_indfun.mlg
+++ b/plugins/funind/g_indfun.mlg
@@ -177,7 +177,7 @@ let () =
(* TASSI: n'importe quoi ! *)
VERNAC COMMAND EXTEND Function
-| ![ proof ] ["Function" ne_function_rec_definition_loc_list_sep(recsl,"with")]
+| ![ maybe_open_proof ] ["Function" ne_function_rec_definition_loc_list_sep(recsl,"with")]
=> { let hard = List.exists (function
| _,((_,(Some { CAst.v = CMeasureRec _ }
| Some { CAst.v = CWfRec _}),_,_,_),_) -> true
@@ -190,9 +190,7 @@ VERNAC COMMAND EXTEND Function
| Vernacextend.VtSideff ids, _ when hard ->
Vernacextend.(VtStartProof (GuaranteesOpacity, ids), VtLater)
| x -> x }
- -> { fun ~pstate:ontop ->
- let pstate = do_generate_principle false (List.map snd recsl) in
- Proof_global.maybe_push ~ontop pstate}
+ -> { do_generate_principle false (List.map snd recsl) }
END
{
@@ -227,33 +225,32 @@ let warning_error names e =
}
VERNAC COMMAND EXTEND NewFunctionalScheme
-| ![ proof ] ["Functional" "Scheme" ne_fun_scheme_arg_list_sep(fas,"with") ]
+| ![ maybe_open_proof ] ["Functional" "Scheme" ne_fun_scheme_arg_list_sep(fas,"with") ]
=> { Vernacextend.(VtSideff(List.map pi1 fas), VtLater) }
->
- { fun ~pstate:ontop ->
- begin
+ { begin
try
- Functional_principles_types.build_scheme fas; ontop
+ Functional_principles_types.build_scheme fas; None
with
| Functional_principles_types.No_graph_found ->
begin
match fas with
| (_,fun_name,_)::_ ->
begin
- let ontop = make_graph ~ontop (Smartlocate.global_with_alias fun_name) in
- try Functional_principles_types.build_scheme fas; ontop
+ let pstate = make_graph (Smartlocate.global_with_alias fun_name) in
+ try Functional_principles_types.build_scheme fas; pstate
with
| Functional_principles_types.No_graph_found ->
CErrors.user_err Pp.(str "Cannot generate induction principle(s)")
| e when CErrors.noncritical e ->
let names = List.map (fun (_,na,_) -> na) fas in
- warning_error names e; ontop
+ warning_error names e; pstate
end
| _ -> assert false (* we can only have non empty list *)
end
| e when CErrors.noncritical e ->
let names = List.map (fun (_,na,_) -> na) fas in
- warning_error names e; ontop
+ warning_error names e; None
end
}
END
@@ -267,6 +264,6 @@ END
(***** debug only ***)
VERNAC COMMAND EXTEND GenerateGraph CLASSIFIED AS QUERY
-| ![ proof ] ["Generate" "graph" "for" reference(c)] ->
- { fun ~pstate:ontop -> make_graph ~ontop (Smartlocate.global_with_alias c) }
+| ![ maybe_open_proof ] ["Generate" "graph" "for" reference(c)] ->
+ { make_graph (Smartlocate.global_with_alias c) }
END
diff --git a/plugins/funind/indfun.ml b/plugins/funind/indfun.ml
index d47c12a7cb..6ea2eb579f 100644
--- a/plugins/funind/indfun.ml
+++ b/plugins/funind/indfun.ml
@@ -839,7 +839,9 @@ let rec get_args b t : Constrexpr.local_binder_expr list *
| _ -> [],b,t
-let make_graph env sigma (f_ref : GlobRef.t) =
+let make_graph (f_ref : GlobRef.t) =
+ let env = Global.env() in
+ let sigma = Evd.from_env env in
let c,c_body =
match f_ref with
| ConstRef c ->
@@ -907,11 +909,4 @@ let make_graph env sigma (f_ref : GlobRef.t) =
expr_list;
pstate)
-let make_graph ~ontop f_ref =
- let pstate = Option.map Proof_global.get_current_pstate ontop in
- let sigma, env = Option.cata Pfedit.get_current_context
- (let e = Global.env () in Evd.from_env e, e) pstate in
- Option.cata (fun ps -> Some (Proof_global.push ~ontop ps)) ontop
- (make_graph env sigma f_ref)
-
let do_generate_principle = do_generate_principle [] warning_error true
diff --git a/plugins/funind/indfun.mli b/plugins/funind/indfun.mli
index cf4ef23d1a..c803484617 100644
--- a/plugins/funind/indfun.mli
+++ b/plugins/funind/indfun.mli
@@ -17,4 +17,4 @@ val functional_induction :
Ltac_plugin.Tacexpr.or_and_intro_pattern option ->
Goal.goal Evd.sigma -> Goal.goal list Evd.sigma
-val make_graph : ontop:Proof_global.t option -> GlobRef.t -> Proof_global.t option
+val make_graph : GlobRef.t -> Proof_global.pstate option
diff --git a/plugins/ltac/extratactics.mlg b/plugins/ltac/extratactics.mlg
index 5e4969567d..4a2fb425ac 100644
--- a/plugins/ltac/extratactics.mlg
+++ b/plugins/ltac/extratactics.mlg
@@ -934,7 +934,7 @@ END
VERNAC COMMAND EXTEND GrabEvars
| ![ proof ] [ "Grab" "Existential" "Variables" ]
=> { classify_as_proofstep }
- -> { fun ~pstate -> Option.map (Proof_global.simple_with_current_proof (fun _ p -> Proof.V82.grab_evars p)) pstate }
+ -> { fun ~pstate -> Proof_global.modify_proof (fun p -> Proof.V82.grab_evars p) pstate }
END
(* Shelves all the goals under focus. *)
@@ -966,7 +966,7 @@ END
VERNAC COMMAND EXTEND Unshelve
| ![ proof ] [ "Unshelve" ]
=> { classify_as_proofstep }
- -> { fun ~pstate -> Option.map (Proof_global.simple_with_current_proof (fun _ p -> Proof.unshelve p)) pstate }
+ -> { fun ~pstate -> Proof_global.modify_proof (fun p -> Proof.unshelve p) pstate }
END
(* Gives up on the goals under focus: the goals are considered solved,
@@ -1118,7 +1118,7 @@ END
VERNAC COMMAND EXTEND OptimizeProof
| ![ proof ] [ "Optimize" "Proof" ] => { classify_as_proofstep } ->
- { fun ~pstate -> Option.map Proof_global.(modify_current_pstate compact_the_proof) pstate }
+ { fun ~pstate -> Proof_global.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 7eb34158e8..9bf5bd1bda 100644
--- a/plugins/ltac/g_ltac.mlg
+++ b/plugins/ltac/g_ltac.mlg
@@ -376,7 +376,7 @@ let () = declare_int_option {
let vernac_solve ~pstate n info tcom b =
let open Goal_select in
- let pstate, status = Proof_global.with_current_proof (fun etac p ->
+ let pstate, status = Proof_global.with_proof (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
@@ -388,7 +388,7 @@ let vernac_solve ~pstate n info tcom b =
let p = Proof.maximal_unfocus Vernacentries.command_focus p in
p,status) pstate in
if not status then Feedback.feedback Feedback.AddedAxiom;
- Some pstate
+ pstate
let pr_ltac_selector s = Pptactic.pr_goal_selector ~toplevel:true s
@@ -438,7 +438,7 @@ VERNAC { tactic_mode } EXTEND VernacSolve
| ![ proof ] [ ltac_selector_opt(g) ltac_info_opt(n) tactic(t) ltac_use_default(def) ] =>
{ classify_as_proofstep } -> {
let g = Option.default (Goal_select.get_default_goal_selector ()) g in
- Vernacentries.vernac_require_open_proof vernac_solve g n t def
+ vernac_solve g n t def
}
| ![ proof ] [ "par" ":" ltac_info_opt(n) tactic(t) ltac_use_default(def) ] =>
{
@@ -450,7 +450,7 @@ VERNAC { tactic_mode } EXTEND VernacSolve
VtLater
} -> {
let t = rm_abstract t in
- Vernacentries.vernac_require_open_proof vernac_solve Goal_select.SelectAll n t def
+ vernac_solve Goal_select.SelectAll n t def
}
END
diff --git a/plugins/ltac/g_obligations.mlg b/plugins/ltac/g_obligations.mlg
index d7918a58ac..c01b97c14d 100644
--- a/plugins/ltac/g_obligations.mlg
+++ b/plugins/ltac/g_obligations.mlg
@@ -80,25 +80,25 @@ GRAMMAR EXTEND Gram
open Obligations
-let obligation ~pstate obl tac = Some (Proof_global.push ~ontop:pstate (with_tac (fun t -> Obligations.obligation obl t) tac))
-let next_obligation ~pstate obl tac = Some (Proof_global.push ~ontop:pstate (with_tac (fun t -> Obligations.next_obligation obl t) tac))
+let obligation obl tac = with_tac (fun t -> Obligations.obligation obl t) tac
+let next_obligation obl tac = with_tac (fun t -> Obligations.next_obligation obl t) tac
let classify_obbl _ = Vernacextend.(VtStartProof (Doesn'tGuaranteeOpacity,[]), VtLater)
}
VERNAC COMMAND EXTEND Obligations CLASSIFIED BY { classify_obbl }
-| ![ proof ] [ "Obligation" integer(num) "of" ident(name) ":" lglob(t) withtac(tac) ] ->
+| ![ open_proof ] [ "Obligation" integer(num) "of" ident(name) ":" lglob(t) withtac(tac) ] ->
{ obligation (num, Some name, Some t) tac }
-| ![ proof ] [ "Obligation" integer(num) "of" ident(name) withtac(tac) ] ->
+| ![ open_proof ] [ "Obligation" integer(num) "of" ident(name) withtac(tac) ] ->
{ obligation (num, Some name, None) tac }
-| ![ proof ] [ "Obligation" integer(num) ":" lglob(t) withtac(tac) ] ->
+| ![ open_proof ] [ "Obligation" integer(num) ":" lglob(t) withtac(tac) ] ->
{ obligation (num, None, Some t) tac }
-| ![ proof ] [ "Obligation" integer(num) withtac(tac) ] ->
+| ![ open_proof ] [ "Obligation" integer(num) withtac(tac) ] ->
{ obligation (num, None, None) tac }
-| ![ proof ] [ "Next" "Obligation" "of" ident(name) withtac(tac) ] ->
+| ![ open_proof ] [ "Next" "Obligation" "of" ident(name) withtac(tac) ] ->
{ next_obligation (Some name) tac }
-| ![ proof ] [ "Next" "Obligation" withtac(tac) ] -> { next_obligation None tac }
+| ![ open_proof ] [ "Next" "Obligation" withtac(tac) ] -> { next_obligation None tac }
END
VERNAC COMMAND EXTEND Solve_Obligation CLASSIFIED AS SIDEFF
diff --git a/plugins/ltac/g_rewrite.mlg b/plugins/ltac/g_rewrite.mlg
index e3f4b8ef59..bb29323858 100644
--- a/plugins/ltac/g_rewrite.mlg
+++ b/plugins/ltac/g_rewrite.mlg
@@ -180,34 +180,34 @@ TACTIC EXTEND setoid_rewrite
END
VERNAC COMMAND EXTEND AddRelation CLASSIFIED AS SIDEFF
- | #[ atts = rewrite_attributes; ] ![ proof ] [ "Add" "Relation" constr(a) constr(aeq) "reflexivity" "proved" "by" constr(lemma1)
+ | #[ atts = rewrite_attributes; ] [ "Add" "Relation" constr(a) constr(aeq) "reflexivity" "proved" "by" constr(lemma1)
"symmetry" "proved" "by" constr(lemma2) "as" ident(n) ] ->
{ declare_relation atts a aeq n (Some lemma1) (Some lemma2) None }
- | #[ atts = rewrite_attributes; ] ![ proof ] [ "Add" "Relation" constr(a) constr(aeq) "reflexivity" "proved" "by" constr(lemma1)
+ | #[ atts = rewrite_attributes; ] [ "Add" "Relation" constr(a) constr(aeq) "reflexivity" "proved" "by" constr(lemma1)
"as" ident(n) ] ->
{ declare_relation atts a aeq n (Some lemma1) None None }
- | #[ atts = rewrite_attributes; ] ![ proof ] [ "Add" "Relation" constr(a) constr(aeq) "as" ident(n) ] ->
+ | #[ atts = rewrite_attributes; ] [ "Add" "Relation" constr(a) constr(aeq) "as" ident(n) ] ->
{ declare_relation atts a aeq n None None None }
END
VERNAC COMMAND EXTEND AddRelation2 CLASSIFIED AS SIDEFF
- | #[ atts = rewrite_attributes; ] ![ proof ] [ "Add" "Relation" constr(a) constr(aeq) "symmetry" "proved" "by" constr(lemma2)
+ | #[ atts = rewrite_attributes; ] [ "Add" "Relation" constr(a) constr(aeq) "symmetry" "proved" "by" constr(lemma2)
"as" ident(n) ] ->
{ declare_relation atts a aeq n None (Some lemma2) None }
- | #[ atts = rewrite_attributes; ] ![ proof ] [ "Add" "Relation" constr(a) constr(aeq) "symmetry" "proved" "by" constr(lemma2) "transitivity" "proved" "by" constr(lemma3) "as" ident(n) ] ->
+ | #[ atts = rewrite_attributes; ] [ "Add" "Relation" constr(a) constr(aeq) "symmetry" "proved" "by" constr(lemma2) "transitivity" "proved" "by" constr(lemma3) "as" ident(n) ] ->
{ declare_relation atts a aeq n None (Some lemma2) (Some lemma3) }
END
VERNAC COMMAND EXTEND AddRelation3 CLASSIFIED AS SIDEFF
- | #[ atts = rewrite_attributes; ] ![ proof ] [ "Add" "Relation" constr(a) constr(aeq) "reflexivity" "proved" "by" constr(lemma1)
+ | #[ atts = rewrite_attributes; ] [ "Add" "Relation" constr(a) constr(aeq) "reflexivity" "proved" "by" constr(lemma1)
"transitivity" "proved" "by" constr(lemma3) "as" ident(n) ] ->
{ declare_relation atts a aeq n (Some lemma1) None (Some lemma3) }
- | #[ atts = rewrite_attributes; ] ![ proof ] [ "Add" "Relation" constr(a) constr(aeq) "reflexivity" "proved" "by" constr(lemma1)
+ | #[ atts = rewrite_attributes; ] [ "Add" "Relation" constr(a) constr(aeq) "reflexivity" "proved" "by" constr(lemma1)
"symmetry" "proved" "by" constr(lemma2) "transitivity" "proved" "by" constr(lemma3)
"as" ident(n) ] ->
{ declare_relation atts a aeq n (Some lemma1) (Some lemma2) (Some lemma3) }
- | #[ atts = rewrite_attributes; ] ![ proof ] [ "Add" "Relation" constr(a) constr(aeq) "transitivity" "proved" "by" constr(lemma3)
+ | #[ atts = rewrite_attributes; ] [ "Add" "Relation" constr(a) constr(aeq) "transitivity" "proved" "by" constr(lemma3)
"as" ident(n) ] ->
{ declare_relation atts a aeq n None None (Some lemma3) }
END
@@ -234,66 +234,59 @@ GRAMMAR EXTEND Gram
END
VERNAC COMMAND EXTEND AddParametricRelation CLASSIFIED AS SIDEFF
- | #[ atts = rewrite_attributes; ] ![ proof ] [ "Add" "Parametric" "Relation" binders(b) ":" constr(a) constr(aeq)
+ | #[ atts = rewrite_attributes; ] [ "Add" "Parametric" "Relation" binders(b) ":" constr(a) constr(aeq)
"reflexivity" "proved" "by" constr(lemma1)
"symmetry" "proved" "by" constr(lemma2) "as" ident(n) ] ->
{ declare_relation atts ~binders:b a aeq n (Some lemma1) (Some lemma2) None }
- | #[ atts = rewrite_attributes; ] ![ proof ] [ "Add" "Parametric" "Relation" binders(b) ":" constr(a) constr(aeq)
+ | #[ atts = rewrite_attributes; ] [ "Add" "Parametric" "Relation" binders(b) ":" constr(a) constr(aeq)
"reflexivity" "proved" "by" constr(lemma1)
"as" ident(n) ] ->
{ declare_relation atts ~binders:b a aeq n (Some lemma1) None None }
- | #[ atts = rewrite_attributes; ] ![ proof ] [ "Add" "Parametric" "Relation" binders(b) ":" constr(a) constr(aeq) "as" ident(n) ] ->
+ | #[ atts = rewrite_attributes; ] [ "Add" "Parametric" "Relation" binders(b) ":" constr(a) constr(aeq) "as" ident(n) ] ->
{ declare_relation atts ~binders:b a aeq n None None None }
END
VERNAC COMMAND EXTEND AddParametricRelation2 CLASSIFIED AS SIDEFF
- | #[ atts = rewrite_attributes; ] ![ proof ] [ "Add" "Parametric" "Relation" binders(b) ":" constr(a) constr(aeq) "symmetry" "proved" "by" constr(lemma2)
+ | #[ atts = rewrite_attributes; ] [ "Add" "Parametric" "Relation" binders(b) ":" constr(a) constr(aeq) "symmetry" "proved" "by" constr(lemma2)
"as" ident(n) ] ->
{ declare_relation atts ~binders:b a aeq n None (Some lemma2) None }
- | #[ atts = rewrite_attributes; ] ![ proof ] [ "Add" "Parametric" "Relation" binders(b) ":" constr(a) constr(aeq) "symmetry" "proved" "by" constr(lemma2) "transitivity" "proved" "by" constr(lemma3) "as" ident(n) ] ->
+ | #[ atts = rewrite_attributes; ] [ "Add" "Parametric" "Relation" binders(b) ":" constr(a) constr(aeq) "symmetry" "proved" "by" constr(lemma2) "transitivity" "proved" "by" constr(lemma3) "as" ident(n) ] ->
{ declare_relation atts ~binders:b a aeq n None (Some lemma2) (Some lemma3) }
END
VERNAC COMMAND EXTEND AddParametricRelation3 CLASSIFIED AS SIDEFF
- | #[ atts = rewrite_attributes; ] ![ proof ] [ "Add" "Parametric" "Relation" binders(b) ":" constr(a) constr(aeq) "reflexivity" "proved" "by" constr(lemma1)
+ | #[ atts = rewrite_attributes; ] [ "Add" "Parametric" "Relation" binders(b) ":" constr(a) constr(aeq) "reflexivity" "proved" "by" constr(lemma1)
"transitivity" "proved" "by" constr(lemma3) "as" ident(n) ] ->
{ declare_relation atts ~binders:b a aeq n (Some lemma1) None (Some lemma3) }
- | #[ atts = rewrite_attributes; ] ![ proof ] [ "Add" "Parametric" "Relation" binders(b) ":" constr(a) constr(aeq) "reflexivity" "proved" "by" constr(lemma1)
+ | #[ atts = rewrite_attributes; ] [ "Add" "Parametric" "Relation" binders(b) ":" constr(a) constr(aeq) "reflexivity" "proved" "by" constr(lemma1)
"symmetry" "proved" "by" constr(lemma2) "transitivity" "proved" "by" constr(lemma3)
"as" ident(n) ] ->
{ declare_relation atts ~binders:b a aeq n (Some lemma1) (Some lemma2) (Some lemma3) }
- | #[ atts = rewrite_attributes; ] ![ proof ] [ "Add" "Parametric" "Relation" binders(b) ":" constr(a) constr(aeq) "transitivity" "proved" "by" constr(lemma3)
+ | #[ atts = rewrite_attributes; ] [ "Add" "Parametric" "Relation" binders(b) ":" constr(a) constr(aeq) "transitivity" "proved" "by" constr(lemma3)
"as" ident(n) ] ->
{ declare_relation atts ~binders:b a aeq n None None (Some lemma3) }
END
VERNAC COMMAND EXTEND AddSetoid1 CLASSIFIED AS SIDEFF
- | #[ atts = rewrite_attributes; ] ![ proof ] [ "Add" "Setoid" constr(a) constr(aeq) constr(t) "as" ident(n) ] ->
+ | #[ atts = rewrite_attributes; ] [ "Add" "Setoid" constr(a) constr(aeq) constr(t) "as" ident(n) ] ->
{
add_setoid atts [] a aeq t n
}
- | #[ atts = rewrite_attributes; ] ![ proof ] [ "Add" "Parametric" "Setoid" binders(binders) ":" constr(a) constr(aeq) constr(t) "as" ident(n) ] ->
+ | #[ atts = rewrite_attributes; ] [ "Add" "Parametric" "Setoid" binders(binders) ":" constr(a) constr(aeq) constr(t) "as" ident(n) ] ->
{
add_setoid atts binders a aeq t n
}
- | #[ atts = rewrite_attributes; ] ![ proof ] [ "Add" "Morphism" constr(m) ":" ident(n) ]
+ | #[ atts = rewrite_attributes; ] ![ maybe_open_proof ] [ "Add" "Morphism" constr(m) ":" ident(n) ]
(* This command may or may not open a goal *)
=> { (if Lib.is_modtype() then VtSideff([n]) else VtStartProof(GuaranteesOpacity, [n])), VtLater }
- -> {
- fun ~pstate:ontop ->
- Option.cata (fun ps -> Some (Proof_global.push ~ontop ps)) ontop (add_morphism_infer atts m n)
- }
- | #[ atts = rewrite_attributes; ] ![ proof ] [ "Add" "Morphism" constr(m) "with" "signature" lconstr(s) "as" ident(n) ]
+ -> { add_morphism_infer atts m n }
+ | #[ atts = rewrite_attributes; ] ![ open_proof ] [ "Add" "Morphism" constr(m) "with" "signature" lconstr(s) "as" ident(n) ]
=> { VtStartProof(GuaranteesOpacity,[n]), VtLater }
- -> {
- add_morphism atts [] m s n
- }
- | #[ atts = rewrite_attributes; ] ![ proof ] [ "Add" "Parametric" "Morphism" binders(binders) ":" constr(m)
+ -> { add_morphism atts [] m s n }
+ | #[ atts = rewrite_attributes; ] ![ open_proof ] [ "Add" "Parametric" "Morphism" binders(binders) ":" constr(m)
"with" "signature" lconstr(s) "as" ident(n) ]
=> { VtStartProof(GuaranteesOpacity,[n]), VtLater }
- -> {
- add_morphism atts binders m s n
- }
+ -> { add_morphism atts binders m s n }
END
TACTIC EXTEND setoid_symmetry
diff --git a/plugins/ltac/rewrite.ml b/plugins/ltac/rewrite.ml
index efca411754..5d57a2f27f 100644
--- a/plugins/ltac/rewrite.ml
+++ b/plugins/ltac/rewrite.ml
@@ -23,7 +23,6 @@ open Tacticals.New
open Tactics
open Pretype_errors
open Typeclasses
-open Classes
open Constrexpr
open Globnames
open Evd
@@ -1795,17 +1794,17 @@ let declare_an_instance n s args =
let declare_instance a aeq n s = declare_an_instance n s [a;aeq]
-let anew_instance ~pstate atts binders (name,t) fields =
+let anew_instance atts binders (name,t) fields =
let program_mode = atts.program in
- let _id, pstate = new_instance ~pstate ~program_mode atts.polymorphic
+ let _id, proof = Classes.new_instance ~program_mode atts.polymorphic
name binders t (Some (true, CAst.make @@ CRecord (fields)))
~global:atts.global ~generalize:false Hints.empty_hint_info
in
- pstate
+ assert (Option.is_empty proof) (* refine:false with term *)
-let declare_instance_refl ~pstate atts binders a aeq n lemma =
+let declare_instance_refl atts binders a aeq n lemma =
let instance = declare_instance a aeq (add_suffix n "_Reflexive") "Coq.Classes.RelationClasses.Reflexive"
- in anew_instance ~pstate atts binders instance
+ in anew_instance atts binders instance
[(qualid_of_ident (Id.of_string "reflexivity"),lemma)]
let declare_instance_sym atts binders a aeq n lemma =
@@ -1818,41 +1817,41 @@ let declare_instance_trans atts binders a aeq n lemma =
in anew_instance atts binders instance
[(qualid_of_ident (Id.of_string "transitivity"),lemma)]
-let declare_relation ~pstate atts ?(binders=[]) a aeq n refl symm trans =
+let declare_relation atts ?(binders=[]) a aeq n refl symm trans =
init_setoid ();
let instance = declare_instance a aeq (add_suffix n "_relation") "Coq.Classes.RelationClasses.RewriteRelation" in
- let pstate = anew_instance ~pstate atts binders instance [] in
+ let () = anew_instance atts binders instance [] in
match (refl,symm,trans) with
- (None, None, None) -> pstate
+ (None, None, None) -> ()
| (Some lemma1, None, None) ->
- declare_instance_refl ~pstate atts binders a aeq n lemma1
+ declare_instance_refl atts binders a aeq n lemma1
| (None, Some lemma2, None) ->
- declare_instance_sym ~pstate atts binders a aeq n lemma2
+ declare_instance_sym atts binders a aeq n lemma2
| (None, None, Some lemma3) ->
- declare_instance_trans ~pstate atts binders a aeq n lemma3
+ declare_instance_trans atts binders a aeq n lemma3
| (Some lemma1, Some lemma2, None) ->
- let pstate = declare_instance_refl ~pstate atts binders a aeq n lemma1 in
- declare_instance_sym ~pstate atts binders a aeq n lemma2
+ let () = declare_instance_refl atts binders a aeq n lemma1 in
+ declare_instance_sym atts binders a aeq n lemma2
| (Some lemma1, None, Some lemma3) ->
- let pstate = declare_instance_refl ~pstate atts binders a aeq n lemma1 in
- let pstate = declare_instance_trans ~pstate atts binders a aeq n lemma3 in
+ let () = declare_instance_refl atts binders a aeq n lemma1 in
+ let () = declare_instance_trans atts binders a aeq n lemma3 in
let instance = declare_instance a aeq n "Coq.Classes.RelationClasses.PreOrder" in
- anew_instance ~pstate atts binders instance
+ anew_instance atts binders instance
[(qualid_of_ident (Id.of_string "PreOrder_Reflexive"), lemma1);
(qualid_of_ident (Id.of_string "PreOrder_Transitive"),lemma3)]
| (None, Some lemma2, Some lemma3) ->
- let pstate = declare_instance_sym ~pstate atts binders a aeq n lemma2 in
- let pstate = declare_instance_trans ~pstate atts binders a aeq n lemma3 in
+ let () = declare_instance_sym atts binders a aeq n lemma2 in
+ let () = declare_instance_trans atts binders a aeq n lemma3 in
let instance = declare_instance a aeq n "Coq.Classes.RelationClasses.PER" in
- anew_instance ~pstate atts binders instance
+ anew_instance atts binders instance
[(qualid_of_ident (Id.of_string "PER_Symmetric"), lemma2);
(qualid_of_ident (Id.of_string "PER_Transitive"),lemma3)]
| (Some lemma1, Some lemma2, Some lemma3) ->
- let pstate = declare_instance_refl ~pstate atts binders a aeq n lemma1 in
- let pstate = declare_instance_sym ~pstate atts binders a aeq n lemma2 in
- let pstate = declare_instance_trans ~pstate atts binders a aeq n lemma3 in
+ let () = declare_instance_refl atts binders a aeq n lemma1 in
+ let () = declare_instance_sym atts binders a aeq n lemma2 in
+ let () = declare_instance_trans atts binders a aeq n lemma3 in
let instance = declare_instance a aeq n "Coq.Classes.RelationClasses.Equivalence" in
- anew_instance ~pstate atts binders instance
+ anew_instance atts binders instance
[(qualid_of_ident (Id.of_string "Equivalence_Reflexive"), lemma1);
(qualid_of_ident (Id.of_string "Equivalence_Symmetric"), lemma2);
(qualid_of_ident (Id.of_string "Equivalence_Transitive"), lemma3)]
@@ -1951,15 +1950,15 @@ let warn_add_setoid_deprecated =
CWarnings.create ~name:"add-setoid" ~category:"deprecated" (fun () ->
Pp.(str "Add Setoid is deprecated, please use Add Parametric Relation."))
-let add_setoid ~pstate atts binders a aeq t n =
+let add_setoid atts binders a aeq t n =
warn_add_setoid_deprecated ?loc:a.CAst.loc ();
init_setoid ();
- let pstate = declare_instance_refl ~pstate atts binders a aeq n (mkappc "Seq_refl" [a;aeq;t]) in
- let pstate = declare_instance_sym ~pstate atts binders a aeq n (mkappc "Seq_sym" [a;aeq;t]) in
- let pstate = declare_instance_trans ~pstate atts binders a aeq n (mkappc "Seq_trans" [a;aeq;t]) in
+ let () = declare_instance_refl atts binders a aeq n (mkappc "Seq_refl" [a;aeq;t]) in
+ let () = declare_instance_sym atts binders a aeq n (mkappc "Seq_sym" [a;aeq;t]) in
+ let () = declare_instance_trans atts binders a aeq n (mkappc "Seq_trans" [a;aeq;t]) in
let instance = declare_instance a aeq n "Coq.Classes.RelationClasses.Equivalence"
in
- anew_instance ~pstate atts binders instance
+ anew_instance atts binders instance
[(qualid_of_ident (Id.of_string "Equivalence_Reflexive"), mkappc "Seq_refl" [a;aeq;t]);
(qualid_of_ident (Id.of_string "Equivalence_Symmetric"), mkappc "Seq_sym" [a;aeq;t]);
(qualid_of_ident (Id.of_string "Equivalence_Transitive"), mkappc "Seq_trans" [a;aeq;t])]
@@ -1988,7 +1987,7 @@ let add_morphism_infer atts m n : Proof_global.pstate option =
(None,(instance,uctx),None),
Decl_kinds.IsAssumption Decl_kinds.Logical)
in
- add_instance (Classes.mk_instance
+ Classes.add_instance (Classes.mk_instance
(PropGlobal.proper_class env evd) Hints.empty_hint_info atts.global (ConstRef cst));
declare_projection n instance_id (ConstRef cst);
None
@@ -1999,7 +1998,7 @@ let add_morphism_infer atts m n : Proof_global.pstate option =
let tac = make_tactic "Coq.Classes.SetoidTactics.add_morphism_tactic" in
let hook _ _ _ = function
| Globnames.ConstRef cst ->
- add_instance (Classes.mk_instance
+ Classes.add_instance (Classes.mk_instance
(PropGlobal.proper_class env evd) Hints.empty_hint_info
atts.global (ConstRef cst));
declare_projection n instance_id (ConstRef cst)
@@ -2011,7 +2010,7 @@ let add_morphism_infer atts m n : Proof_global.pstate option =
let pstate = Lemmas.start_proof ~hook instance_id kind (Evd.from_ctx uctx) (EConstr.of_constr instance) in
Some (fst Pfedit.(by (Tacinterp.interp tac) pstate))) ()
-let add_morphism ~pstate atts binders m s n =
+let add_morphism atts binders m s n =
init_setoid ();
let instance_id = add_suffix n "_Proper" in
let instance_name = (CAst.make @@ Name instance_id),None in
@@ -2021,12 +2020,12 @@ let add_morphism ~pstate atts binders m s n =
[cHole; s; m])
in
let tac = Tacinterp.interp (make_tactic "add_morphism_tactic") in
- let _, pstate = new_instance ~pstate
+ let _, pstate = Classes.new_instance
~program_mode:atts.program ~global:atts.global atts.polymorphic
instance_name binders instance_t None
~generalize:false ~tac ~hook:(declare_projection n instance_id) Hints.empty_hint_info
in
- pstate
+ Option.get pstate (* no instance body -> always open proof *)
(** Bind to "rewrite" too *)
diff --git a/plugins/ltac/rewrite.mli b/plugins/ltac/rewrite.mli
index cab7d0065e..e1cbc46b3b 100644
--- a/plugins/ltac/rewrite.mli
+++ b/plugins/ltac/rewrite.mli
@@ -81,18 +81,35 @@ val cl_rewrite_clause :
val is_applied_rewrite_relation :
env -> evar_map -> rel_context -> constr -> types option
-val declare_relation : pstate:Proof_global.t option -> rewrite_attributes ->
- ?binders:local_binder_expr list -> constr_expr -> constr_expr -> Id.t ->
- constr_expr option -> constr_expr option -> constr_expr option -> Proof_global.t option
-
-val add_setoid : pstate:Proof_global.t option ->
- rewrite_attributes -> local_binder_expr list -> constr_expr -> constr_expr -> constr_expr ->
- Id.t -> Proof_global.t option
+val declare_relation
+ : rewrite_attributes
+ -> ?binders:local_binder_expr list
+ -> constr_expr
+ -> constr_expr
+ -> Id.t
+ -> constr_expr option
+ -> constr_expr option
+ -> constr_expr option
+ -> unit
+
+val add_setoid
+ : rewrite_attributes
+ -> local_binder_expr list
+ -> constr_expr
+ -> constr_expr
+ -> constr_expr
+ -> Id.t
+ -> unit
val add_morphism_infer : rewrite_attributes -> constr_expr -> Id.t -> Proof_global.pstate option
-val add_morphism : pstate:Proof_global.t option ->
- rewrite_attributes -> local_binder_expr list -> constr_expr -> constr_expr -> Id.t -> Proof_global.t option
+val add_morphism
+ : rewrite_attributes
+ -> local_binder_expr list
+ -> constr_expr
+ -> constr_expr
+ -> Id.t
+ -> Proof_global.pstate
val get_reflexive_proof : env -> evar_map -> constr -> constr -> evar_map * constr