From 078efbe2dfff0b783cd35b0b3ab2354f554e95a6 Mon Sep 17 00:00:00 2001 From: Arnaud Spiwack Date: Mon, 25 Nov 2013 19:05:48 +0100 Subject: Derive plugin. A small plugin to support program deriving (à la Richard Bird) in Coq. It's fairly simple: Require Coq.Derive.Derive. Derive f From g Upto eq As h. Produces a goal (actually two, but the first one is automatically shelved): |- eq g ?42 And closing the proof produces two definitions: f is instantiated with the value of ?42 (it's always transparent). And h is instantiated with the content of the proof (it is transparent or opaque depending on whether the proof was closed with Defined or Qed). --- plugins/Derive/Derive.v | 1 + plugins/Derive/derive.ml | 74 ++++++++++++++++++++++++++++++++++++++ plugins/Derive/derive.mli | 12 +++++++ plugins/Derive/derive_plugin.mllib | 2 ++ plugins/Derive/g_derive.ml4 | 18 ++++++++++ plugins/Derive/vo.itarget | 1 + plugins/pluginsbyte.itarget | 1 + plugins/pluginsdyn.itarget | 1 + plugins/pluginsopt.itarget | 1 + plugins/pluginsvo.itarget | 1 + 10 files changed, 112 insertions(+) create mode 100644 plugins/Derive/Derive.v create mode 100644 plugins/Derive/derive.ml create mode 100644 plugins/Derive/derive.mli create mode 100644 plugins/Derive/derive_plugin.mllib create mode 100644 plugins/Derive/g_derive.ml4 create mode 100644 plugins/Derive/vo.itarget (limited to 'plugins') diff --git a/plugins/Derive/Derive.v b/plugins/Derive/Derive.v new file mode 100644 index 0000000000..0d5a93b034 --- /dev/null +++ b/plugins/Derive/Derive.v @@ -0,0 +1 @@ +Declare ML Module "derive_plugin". \ No newline at end of file diff --git a/plugins/Derive/derive.ml b/plugins/Derive/derive.ml new file mode 100644 index 0000000000..c6a96b31a5 --- /dev/null +++ b/plugins/Derive/derive.ml @@ -0,0 +1,74 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* + TCons ( env , Term.mkApp ( r , [| init_def ; ef |] ) , (fun _ -> TNil)))) + in + let terminator com = + let open Proof_global in + match com with + | Admitted -> Errors.error"Admitted isn't supported in Derive." + | Proved (_,Some _,_) -> + Errors.error"Cannot save a proof of Derive with an explicit name." + | Proved (opaque, None, obj) -> + let (f_def,lemma_def) = + match Proof_global.(obj.entries) with + | [f_def;lemma_def] -> + f_def , lemma_def + | _ -> assert false + in + (* The opacity of [f_def] is adjusted to be [false]. *) + let f_def = let open Entries in { f_def with + const_entry_opaque = false ; } + in + let f_def = Entries.DefinitionEntry f_def , Decl_kinds.(IsDefinition Definition) in + let f_kn = Declare.declare_constant f f_def in + let lemma_typ = Term.(mkApp ( r , [| init_def ; mkConst f_kn |] )) in + (* The type of [lemma_def] is adjusted to refer to [f_kn], the + opacity is adjusted by the proof ending command. *) + let lemma_def = let open Entries in { lemma_def with + const_entry_type = Some lemma_typ ; + const_entry_opaque = opaque ; } + in + let lemma_def = + Entries.DefinitionEntry lemma_def , + Decl_kinds.(IsProof Proposition) + in + ignore (Declare.declare_constant lemma lemma_def) + in + let () = Proof_global.start_dependent_proof + lemma kind goals terminator + in + let _ = Proof_global.with_current_proof begin fun _ p -> + Proof.run_tactic env Proofview.(tclFOCUS 1 1 shelve) p + end in + () + diff --git a/plugins/Derive/derive.mli b/plugins/Derive/derive.mli new file mode 100644 index 0000000000..33f982bb68 --- /dev/null +++ b/plugins/Derive/derive.mli @@ -0,0 +1,12 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* Constrexpr.constr_expr -> Constrexpr.constr_expr -> Names.Id.t -> unit diff --git a/plugins/Derive/derive_plugin.mllib b/plugins/Derive/derive_plugin.mllib new file mode 100644 index 0000000000..5ee0fc6da6 --- /dev/null +++ b/plugins/Derive/derive_plugin.mllib @@ -0,0 +1,2 @@ +Derive +G_derive diff --git a/plugins/Derive/g_derive.ml4 b/plugins/Derive/g_derive.ml4 new file mode 100644 index 0000000000..354643c092 --- /dev/null +++ b/plugins/Derive/g_derive.ml4 @@ -0,0 +1,18 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* + [ Derive.start_deriving f init r lemma ] +END diff --git a/plugins/Derive/vo.itarget b/plugins/Derive/vo.itarget new file mode 100644 index 0000000000..b480982193 --- /dev/null +++ b/plugins/Derive/vo.itarget @@ -0,0 +1 @@ +Derive.vo \ No newline at end of file diff --git a/plugins/pluginsbyte.itarget b/plugins/pluginsbyte.itarget index 8f68cc6966..744b075392 100644 --- a/plugins/pluginsbyte.itarget +++ b/plugins/pluginsbyte.itarget @@ -19,3 +19,4 @@ syntax/r_syntax_plugin.cma syntax/string_syntax_plugin.cma syntax/z_syntax_plugin.cma quote/quote_plugin.cma +Derive/derive_plugin.cma \ No newline at end of file diff --git a/plugins/pluginsdyn.itarget b/plugins/pluginsdyn.itarget index e075ffc715..188eb408a0 100644 --- a/plugins/pluginsdyn.itarget +++ b/plugins/pluginsdyn.itarget @@ -22,3 +22,4 @@ syntax/r_syntax_plugin.cmxs syntax/string_syntax_plugin.cmxs syntax/z_syntax_plugin.cmxs quote/quote_plugin.cmxs +Derive/derive_plugin.cmxs diff --git a/plugins/pluginsopt.itarget b/plugins/pluginsopt.itarget index 01ac840dd7..34b0246637 100644 --- a/plugins/pluginsopt.itarget +++ b/plugins/pluginsopt.itarget @@ -19,3 +19,4 @@ syntax/r_syntax_plugin.cmxa syntax/string_syntax_plugin.cmxa syntax/z_syntax_plugin.cmxa quote/quote_plugin.cmxa +Derive/derive_plugin.cmxa diff --git a/plugins/pluginsvo.itarget b/plugins/pluginsvo.itarget index 4b698cccb0..0a9e468ecb 100644 --- a/plugins/pluginsvo.itarget +++ b/plugins/pluginsvo.itarget @@ -9,3 +9,4 @@ romega/vo.otarget rtauto/vo.otarget setoid_ring/vo.otarget extraction/vo.otarget +Derive/vo.otarget \ No newline at end of file -- cgit v1.2.3