diff options
Diffstat (limited to 'plugins/setoid_ring')
-rw-r--r-- | plugins/setoid_ring/ArithRing.v | 4 | ||||
-rw-r--r-- | plugins/setoid_ring/InitialRing.v | 48 | ||||
-rw-r--r-- | plugins/setoid_ring/NArithRing.v | 2 | ||||
-rw-r--r-- | plugins/setoid_ring/Ring.v | 6 | ||||
-rw-r--r-- | plugins/setoid_ring/ZArithRing.v | 8 | ||||
-rw-r--r-- | plugins/setoid_ring/g_newring.ml4 | 105 | ||||
-rw-r--r-- | plugins/setoid_ring/newring.ml (renamed from plugins/setoid_ring/newring.ml4) | 255 | ||||
-rw-r--r-- | plugins/setoid_ring/newring.mli | 78 | ||||
-rw-r--r-- | plugins/setoid_ring/newring_ast.mli | 63 | ||||
-rw-r--r-- | plugins/setoid_ring/newring_plugin.mllib | 1 |
10 files changed, 356 insertions, 214 deletions
diff --git a/plugins/setoid_ring/ArithRing.v b/plugins/setoid_ring/ArithRing.v index 04decbce1..5f5b97925 100644 --- a/plugins/setoid_ring/ArithRing.v +++ b/plugins/setoid_ring/ArithRing.v @@ -32,13 +32,13 @@ Qed. Ltac natcst t := match isnatcst t with true => constr:(N.of_nat t) - | _ => constr:InitialRing.NotConstant + | _ => constr:(InitialRing.NotConstant) end. Ltac Ss_to_add f acc := match f with | S ?f1 => Ss_to_add f1 (S acc) - | _ => constr:(acc + f)%nat + | _ => constr:((acc + f)%nat) end. Ltac natprering := diff --git a/plugins/setoid_ring/InitialRing.v b/plugins/setoid_ring/InitialRing.v index 8362c8c26..8fcc07716 100644 --- a/plugins/setoid_ring/InitialRing.v +++ b/plugins/setoid_ring/InitialRing.v @@ -612,32 +612,32 @@ End GEN_DIV. Ltac inv_gen_phi_pos rI add mul t := let rec inv_cst t := match t with - rI => constr:1%positive - | (add rI rI) => constr:2%positive - | (add rI (add rI rI)) => constr:3%positive + rI => constr:(1%positive) + | (add rI rI) => constr:(2%positive) + | (add rI (add rI rI)) => constr:(3%positive) | (mul (add rI rI) ?p) => (* 2p *) match inv_cst p with - NotConstant => constr:NotConstant - | 1%positive => constr:NotConstant (* 2*1 is not convertible to 2 *) + NotConstant => constr:(NotConstant) + | 1%positive => constr:(NotConstant) (* 2*1 is not convertible to 2 *) | ?p => constr:(xO p) end | (add rI (mul (add rI rI) ?p)) => (* 1+2p *) match inv_cst p with - NotConstant => constr:NotConstant - | 1%positive => constr:NotConstant + NotConstant => constr:(NotConstant) + | 1%positive => constr:(NotConstant) | ?p => constr:(xI p) end - | _ => constr:NotConstant + | _ => constr:(NotConstant) end in inv_cst t. (* The (partial) inverse of gen_phiNword *) Ltac inv_gen_phiNword rO rI add mul opp t := match t with - rO => constr:NwO + rO => constr:(NwO) | _ => match inv_gen_phi_pos rI add mul t with - NotConstant => constr:NotConstant + NotConstant => constr:(NotConstant) | ?p => constr:(Npos p::nil) end end. @@ -646,10 +646,10 @@ End GEN_DIV. (* The inverse of gen_phiN *) Ltac inv_gen_phiN rO rI add mul t := match t with - rO => constr:0%N + rO => constr:(0%N) | _ => match inv_gen_phi_pos rI add mul t with - NotConstant => constr:NotConstant + NotConstant => constr:(NotConstant) | ?p => constr:(Npos p) end end. @@ -657,15 +657,15 @@ End GEN_DIV. (* The inverse of gen_phiZ *) Ltac inv_gen_phiZ rO rI add mul opp t := match t with - rO => constr:0%Z + rO => constr:(0%Z) | (opp ?p) => match inv_gen_phi_pos rI add mul p with - NotConstant => constr:NotConstant + NotConstant => constr:(NotConstant) | ?p => constr:(Zneg p) end | _ => match inv_gen_phi_pos rI add mul t with - NotConstant => constr:NotConstant + NotConstant => constr:(NotConstant) | ?p => constr:(Zpos p) end end. @@ -681,7 +681,7 @@ Ltac inv_gen_phi rO rI cO cI t := end. (* A simple tactic recognizing no constant *) - Ltac inv_morph_nothing t := constr:NotConstant. + Ltac inv_morph_nothing t := constr:(NotConstant). Ltac coerce_to_almost_ring set ext rspec := match type of rspec with @@ -825,31 +825,31 @@ Ltac ring_elements set ext rspec pspec sspec dspec rk := (* Tactic for constant *) Ltac isnatcst t := match t with - O => constr:true + O => constr:(true) | S ?p => isnatcst p - | _ => constr:false + | _ => constr:(false) end. Ltac isPcst t := match t with | xI ?p => isPcst p | xO ?p => isPcst p - | xH => constr:true + | xH => constr:(true) (* nat -> positive *) | Pos.of_succ_nat ?n => isnatcst n - | _ => constr:false + | _ => constr:(false) end. Ltac isNcst t := match t with - N0 => constr:true + N0 => constr:(true) | Npos ?p => isPcst p - | _ => constr:false + | _ => constr:(false) end. Ltac isZcst t := match t with - Z0 => constr:true + Z0 => constr:(true) | Zpos ?p => isPcst p | Zneg ?p => isPcst p (* injection nat -> Z *) @@ -857,7 +857,7 @@ Ltac isZcst t := (* injection N -> Z *) | Z.of_N ?n => isNcst n (* *) - | _ => constr:false + | _ => constr:(false) end. diff --git a/plugins/setoid_ring/NArithRing.v b/plugins/setoid_ring/NArithRing.v index 6c1a79e4e..54e2789ba 100644 --- a/plugins/setoid_ring/NArithRing.v +++ b/plugins/setoid_ring/NArithRing.v @@ -15,7 +15,7 @@ Set Implicit Arguments. Ltac Ncst t := match isNcst t with true => t - | _ => constr:NotConstant + | _ => constr:(NotConstant) end. Add Ring Nr : Nth (decidable Neqb_ok, constants [Ncst]). diff --git a/plugins/setoid_ring/Ring.v b/plugins/setoid_ring/Ring.v index a0844100c..77576cb93 100644 --- a/plugins/setoid_ring/Ring.v +++ b/plugins/setoid_ring/Ring.v @@ -36,9 +36,9 @@ Qed. Ltac bool_cst t := let t := eval hnf in t in match t with - true => constr:true - | false => constr:false - | _ => constr:NotConstant + true => constr:(true) + | false => constr:(false) + | _ => constr:(NotConstant) end. Add Ring bool_ring : BoolTheory (decidable bool_eq_ok, constants [bool_cst]). diff --git a/plugins/setoid_ring/ZArithRing.v b/plugins/setoid_ring/ZArithRing.v index 914843727..23784cf33 100644 --- a/plugins/setoid_ring/ZArithRing.v +++ b/plugins/setoid_ring/ZArithRing.v @@ -17,14 +17,14 @@ Set Implicit Arguments. Ltac Zcst t := match isZcst t with true => t - | _ => constr:NotConstant + | _ => constr:(NotConstant) end. Ltac isZpow_coef t := match t with | Zpos ?p => isPcst p - | Z0 => constr:true - | _ => constr:false + | Z0 => constr:(true) + | _ => constr:(false) end. Notation N_of_Z := Z.to_N (only parsing). @@ -32,7 +32,7 @@ Notation N_of_Z := Z.to_N (only parsing). Ltac Zpow_tac t := match isZpow_coef t with | true => constr:(N_of_Z t) - | _ => constr:NotConstant + | _ => constr:(NotConstant) end. Ltac Zpower_neg := diff --git a/plugins/setoid_ring/g_newring.ml4 b/plugins/setoid_ring/g_newring.ml4 new file mode 100644 index 000000000..1ebb6e6b7 --- /dev/null +++ b/plugins/setoid_ring/g_newring.ml4 @@ -0,0 +1,105 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) + +(*i camlp4deps: "grammar/grammar.cma" i*) + +open Pp +open Util +open Libnames +open Printer +open Newring_ast +open Newring +open Stdarg +open Constrarg +open Pcoq.Prim +open Pcoq.Constr +open Pcoq.Tactic + +DECLARE PLUGIN "newring_plugin" + +TACTIC EXTEND protect_fv + [ "protect_fv" string(map) "in" ident(id) ] -> + [ Proofview.V82.tactic (protect_tac_in map id) ] +| [ "protect_fv" string(map) ] -> + [ Proofview.V82.tactic (protect_tac map) ] +END + +TACTIC EXTEND closed_term + [ "closed_term" constr(t) "[" ne_reference_list(l) "]" ] -> + [ Proofview.V82.tactic (closed_term t l) ] +END + +VERNAC ARGUMENT EXTEND ring_mod + | [ "decidable" constr(eq_test) ] -> [ Ring_kind(Computational eq_test) ] + | [ "abstract" ] -> [ Ring_kind Abstract ] + | [ "morphism" constr(morph) ] -> [ Ring_kind(Morphism morph) ] + | [ "constants" "[" tactic(cst_tac) "]" ] -> [ Const_tac(CstTac cst_tac) ] + | [ "closed" "[" ne_global_list(l) "]" ] -> [ Const_tac(Closed l) ] + | [ "preprocess" "[" tactic(pre) "]" ] -> [ Pre_tac pre ] + | [ "postprocess" "[" tactic(post) "]" ] -> [ Post_tac post ] + | [ "setoid" constr(sth) constr(ext) ] -> [ Setoid(sth,ext) ] + | [ "sign" constr(sign_spec) ] -> [ Sign_spec sign_spec ] + | [ "power" constr(pow_spec) "[" ne_global_list(l) "]" ] -> + [ Pow_spec (Closed l, pow_spec) ] + | [ "power_tac" constr(pow_spec) "[" tactic(cst_tac) "]" ] -> + [ Pow_spec (CstTac cst_tac, pow_spec) ] + | [ "div" constr(div_spec) ] -> [ Div_spec div_spec ] +END + +VERNAC ARGUMENT EXTEND ring_mods + | [ "(" ne_ring_mod_list_sep(mods, ",") ")" ] -> [ mods ] +END + +VERNAC COMMAND EXTEND AddSetoidRing CLASSIFIED AS SIDEFF + | [ "Add" "Ring" ident(id) ":" constr(t) ring_mods_opt(l) ] -> + [ let l = match l with None -> [] | Some l -> l in + let (k,set,cst,pre,post,power,sign, div) = process_ring_mods l in + add_theory id (ic t) set k cst (pre,post) power sign div] + | [ "Print" "Rings" ] => [Vernac_classifier.classify_as_query] -> [ + msg_notice (strbrk "The following ring structures have been declared:"); + Spmap.iter (fun fn fi -> + msg_notice (hov 2 + (Ppconstr.pr_id (Libnames.basename fn)++spc()++ + str"with carrier "++ pr_constr fi.ring_carrier++spc()++ + str"and equivalence relation "++ pr_constr fi.ring_req)) + ) !from_name ] +END + +TACTIC EXTEND ring_lookup +| [ "ring_lookup" tactic0(f) "[" constr_list(lH) "]" ne_constr_list(lrt) ] -> + [ let (t,lr) = List.sep_last lrt in ring_lookup f lH lr t] +END + +VERNAC ARGUMENT EXTEND field_mod + | [ ring_mod(m) ] -> [ Ring_mod m ] + | [ "completeness" constr(inj) ] -> [ Inject inj ] +END + +VERNAC ARGUMENT EXTEND field_mods + | [ "(" ne_field_mod_list_sep(mods, ",") ")" ] -> [ mods ] +END + +VERNAC COMMAND EXTEND AddSetoidField CLASSIFIED AS SIDEFF +| [ "Add" "Field" ident(id) ":" constr(t) field_mods_opt(l) ] -> + [ let l = match l with None -> [] | Some l -> l in + let (k,set,inj,cst_tac,pre,post,power,sign,div) = process_field_mods l in + add_field_theory id (ic t) set k cst_tac inj (pre,post) power sign div] +| [ "Print" "Fields" ] => [Vernac_classifier.classify_as_query] -> [ + msg_notice (strbrk "The following field structures have been declared:"); + Spmap.iter (fun fn fi -> + msg_notice (hov 2 + (Ppconstr.pr_id (Libnames.basename fn)++spc()++ + str"with carrier "++ pr_constr fi.field_carrier++spc()++ + str"and equivalence relation "++ pr_constr fi.field_req)) + ) !field_from_name ] +END + +TACTIC EXTEND field_lookup +| [ "field_lookup" tactic(f) "[" constr_list(lH) "]" ne_constr_list(lt) ] -> + [ let (t,l) = List.sep_last lt in field_lookup f lH l t ] +END diff --git a/plugins/setoid_ring/newring.ml4 b/plugins/setoid_ring/newring.ml index e704c466e..7ef89b7a0 100644 --- a/plugins/setoid_ring/newring.ml4 +++ b/plugins/setoid_ring/newring.ml @@ -6,8 +6,6 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -(*i camlp4deps: "grammar/grammar.cma" i*) - open Pp open Errors open Util @@ -30,8 +28,8 @@ open Declare open Decl_kinds open Entries open Misctypes - -DECLARE PLUGIN "newring_plugin" +open Newring_ast +open Proofview.Notations (****************************************************************************) (* controlled reduction *) @@ -99,19 +97,12 @@ let protect_red map env sigma c = (mk_clos_but (lookup_map map c) (Esubst.subs_id 0) c);; let protect_tac map = - Tactics.reduct_option (protect_red map,DEFAULTcast) None ;; + Proofview.V82.of_tactic (Tactics.reduct_option (protect_red map,DEFAULTcast) None);; let protect_tac_in map id = - Tactics.reduct_option (protect_red map,DEFAULTcast) (Some(id, Locus.InHyp));; + Proofview.V82.of_tactic (Tactics.reduct_option (protect_red map,DEFAULTcast) (Some(id, Locus.InHyp)));; -TACTIC EXTEND protect_fv - [ "protect_fv" string(map) "in" ident(id) ] -> - [ Proofview.V82.tactic (protect_tac_in map id) ] -| [ "protect_fv" string(map) ] -> - [ Proofview.V82.tactic (protect_tac map) ] -END;; - (****************************************************************************) let closed_term t l = @@ -120,12 +111,6 @@ let closed_term t l = if Quote.closed_under cs t then tclIDTAC else tclFAIL 0 (mt()) ;; -TACTIC EXTEND closed_term - [ "closed_term" constr(t) "[" ne_reference_list(l) "]" ] -> - [ Proofview.V82.tactic (closed_term t l) ] -END -;; - (* TACTIC EXTEND echo | [ "echo" constr(t) ] -> [ Pp.msg (Termops.print_constr t); Tacinterp.eval_tactic (TacId []) ] @@ -143,11 +128,15 @@ let closed_term_ast l = mltac_plugin = "newring_plugin"; mltac_tactic = "closed_term"; } in + let tacname = { + mltac_name = tacname; + mltac_index = 0; + } in let l = List.map (fun gr -> ArgArg(Loc.ghost,gr)) l in TacFun([Some(Id.of_string"t")], TacML(Loc.ghost,tacname, - [Genarg.in_gen (Genarg.glbwit Constrarg.wit_constr) (GVar(Loc.ghost,Id.of_string"t"),None); - Genarg.in_gen (Genarg.glbwit (Genarg.wit_list Constrarg.wit_ref)) l])) + [TacGeneric (Genarg.in_gen (Genarg.glbwit Constrarg.wit_constr) (GVar(Loc.ghost,Id.of_string"t"),None)); + TacGeneric (Genarg.in_gen (Genarg.glbwit (Genarg.wit_list Constrarg.wit_ref)) l)])) (* let _ = add_tacdef false ((Loc.ghost,Id.of_string"ring_closed_term" *) @@ -185,19 +174,20 @@ let ltac_call tac (args:glob_tactic_arg list) = let ltac_lcall tac args = TacArg(Loc.ghost,TacCall(Loc.ghost, ArgVar(Loc.ghost, Id.of_string tac),args)) -let ltac_letin (x, e1) e2 = - TacLetIn(false,[(Loc.ghost,Id.of_string x),e1],e2) - -let ltac_apply (f:glob_tactic_expr) (args:glob_tactic_arg list) = - Tacinterp.eval_tactic - (ltac_letin ("F", Tacexp f) (ltac_lcall "F" args)) +let ltac_apply (f : Value.t) (args: Tacinterp.Value.t list) = + let fold arg (i, vars, lfun) = + let id = Id.of_string ("x" ^ string_of_int i) in + let x = Reference (ArgVar (Loc.ghost, id)) in + (succ i, x :: vars, Id.Map.add id arg lfun) + in + let (_, args, lfun) = List.fold_right fold args (0, [], Id.Map.empty) in + let lfun = Id.Map.add (Id.of_string "F") f lfun in + let ist = { (Tacinterp.default_ist ()) with Tacinterp.lfun = lfun; } in + Tacinterp.eval_tactic_ist ist (ltac_lcall "F" args) let ltac_record flds = TacFun([Some(Id.of_string"proj")], ltac_lcall "proj" flds) - -let carg c = TacDynamic(Loc.ghost,Pretyping.constr_in c) - let dummy_goal env sigma = let (gl,_,sigma) = Goal.V82.mk_goal sigma (named_context_val env) mkProp Evd.Store.empty in @@ -207,20 +197,39 @@ let constr_of v = match Value.to_constr v with | Some c -> c | None -> failwith "Ring.exec_tactic: anomaly" +let tactic_res = ref [||] + +let get_res = + let open Tacexpr in + let name = { mltac_plugin = "newring_plugin"; mltac_tactic = "get_res"; } in + let entry = { mltac_name = name; mltac_index = 0 } in + let tac args ist = + let n = Tacinterp.Value.cast (Genarg.topwit Stdarg.wit_int) (List.hd args) in + let init i = Id.Map.find (Id.of_string ("x" ^ string_of_int i)) ist.lfun in + tactic_res := Array.init n init; + Proofview.tclUNIT () + in + Tacenv.register_ml_tactic name [| tac |]; + entry + let exec_tactic env evd n f args = + let fold arg (i, vars, lfun) = + let id = Id.of_string ("x" ^ string_of_int i) in + let x = Reference (ArgVar (Loc.ghost, id)) in + (succ i, x :: vars, Id.Map.add id (Value.of_constr arg) lfun) + in + let (_, args, lfun) = List.fold_right fold args (0, [], Id.Map.empty) in + let ist = { (Tacinterp.default_ist ()) with Tacinterp.lfun = lfun; } in + (** Build the getter *) let lid = List.init n (fun i -> Id.of_string("x"^string_of_int i)) in - let res = ref [||] in - let get_res ist = - let l = List.map (fun id -> Id.Map.find id ist.lfun) lid in - res := Array.of_list l; - TacId[] in - let getter = - Tacexp(TacFun(List.map(fun id -> Some id) lid, - Tacintern.glob_tactic(tacticIn get_res))) in + let n = Genarg.in_gen (Genarg.glbwit Stdarg.wit_int) n in + let get_res = TacML (Loc.ghost, get_res, [TacGeneric n]) in + let getter = Tacexp (TacFun (List.map (fun id -> Some id) lid, get_res)) in + (** Evaluate the whole result *) let gl = dummy_goal env evd in - let gls = Proofview.V82.of_tactic (Tacinterp.eval_tactic(ltac_call f (args@[getter]))) gl in + let gls = Proofview.V82.of_tactic (Tacinterp.eval_tactic_ist ist (ltac_call f (args@[getter]))) gl in let evd, nf = Evarutil.nf_evars_and_universes (Refiner.project gls) in - Array.map (fun x -> nf (constr_of x)) !res, snd (Evd.universe_context evd) + Array.map (fun x -> nf (constr_of x)) !tactic_res, snd (Evd.universe_context evd) let stdlib_modules = [["Coq";"Setoids";"Setoid"]; @@ -355,20 +364,6 @@ let _ = add_map "ring" (****************************************************************************) (* Ring database *) -type ring_info = - { ring_carrier : types; - ring_req : constr; - ring_setoid : constr; - ring_ext : constr; - ring_morph : constr; - ring_th : constr; - ring_cst_tac : glob_tactic_expr; - ring_pow_tac : glob_tactic_expr; - ring_lemma1 : constr; - ring_lemma2 : constr; - ring_pre_tac : glob_tactic_expr; - ring_post_tac : glob_tactic_expr } - module Cmap = Map.Make(Constr) let from_carrier = Summary.ref Cmap.empty ~name:"ring-tac-carrier-table" @@ -527,8 +522,8 @@ let ring_equality env evd (r,add,mul,opp,req) = match opp with Some opp -> plapp evd coq_eq_morph [|r;add;mul;opp|] | None -> plapp evd coq_eq_smorph [|r;add;mul|] in - let setoid = Typing.solve_evars env evd setoid in - let op_morph = Typing.solve_evars env evd op_morph in + let setoid = Typing.e_solve_evars env evd setoid in + let op_morph = Typing.e_solve_evars env evd op_morph in (setoid,op_morph) | _ -> let setoid = setoid_of_relation (Global.env ()) evd r req in @@ -600,13 +595,6 @@ let dest_morph env sigma m_spec = (c,czero,cone,cadd,cmul,None,None,ceqb,phi) | _ -> error "bad morphism structure" - -type 'constr coeff_spec = - Computational of 'constr (* equality test *) - | Abstract (* coeffs = Z *) - | Morphism of 'constr (* general morphism *) - - let reflect_coeff rkind = (* We build an ill-typed terms on purpose... *) match rkind with @@ -614,10 +602,6 @@ let reflect_coeff rkind = | Computational c -> lapp coq_comp [|c|] | Morphism m -> lapp coq_morph [|m|] -type cst_tac_spec = - CstTac of raw_tactic_expr - | Closed of reference list - let interp_cst_tac env sigma rk kind (zero,one,add,mul,opp) cst_tac = match cst_tac with Some (CstTac t) -> Tacintern.glob_tactic t @@ -638,7 +622,7 @@ let make_hyp_list env evd lH = (fun c l -> plapp evd coq_cons [|carrier; (make_hyp env evd c); l|]) lH (plapp evd coq_nil [|carrier|]) in - let l' = Typing.solve_evars env evd l in + let l' = Typing.e_solve_evars env evd l in Evarutil.nf_evars_universes !evd l' let interp_power env evd pow = @@ -686,7 +670,7 @@ let add_theory name (sigma,rth) eqth morphth cst_tac (pre,post) power sign div = let rk = reflect_coeff morphth in let params,ctx = exec_tactic env !evd 5 (zltac "ring_lemmas") - (List.map carg[sth;ext;rth;pspec;sspec;dspec;rk]) in + [sth;ext;rth;pspec;sspec;dspec;rk] in let lemma1 = params.(3) in let lemma2 = params.(4) in @@ -721,41 +705,12 @@ let add_theory name (sigma,rth) eqth morphth cst_tac (pre,post) power sign div = ring_post_tac = posttac }) in () -type 'constr ring_mod = - Ring_kind of 'constr coeff_spec - | Const_tac of cst_tac_spec - | Pre_tac of raw_tactic_expr - | Post_tac of raw_tactic_expr - | Setoid of Constrexpr.constr_expr * Constrexpr.constr_expr - | Pow_spec of cst_tac_spec * Constrexpr.constr_expr - (* Syntaxification tactic , correctness lemma *) - | Sign_spec of Constrexpr.constr_expr - | Div_spec of Constrexpr.constr_expr - - let ic_coeff_spec = function | Computational t -> Computational (ic_unsafe t) | Morphism t -> Morphism (ic_unsafe t) | Abstract -> Abstract -VERNAC ARGUMENT EXTEND ring_mod - | [ "decidable" constr(eq_test) ] -> [ Ring_kind(Computational eq_test) ] - | [ "abstract" ] -> [ Ring_kind Abstract ] - | [ "morphism" constr(morph) ] -> [ Ring_kind(Morphism morph) ] - | [ "constants" "[" tactic(cst_tac) "]" ] -> [ Const_tac(CstTac cst_tac) ] - | [ "closed" "[" ne_global_list(l) "]" ] -> [ Const_tac(Closed l) ] - | [ "preprocess" "[" tactic(pre) "]" ] -> [ Pre_tac pre ] - | [ "postprocess" "[" tactic(post) "]" ] -> [ Post_tac post ] - | [ "setoid" constr(sth) constr(ext) ] -> [ Setoid(sth,ext) ] - | [ "sign" constr(sign_spec) ] -> [ Sign_spec sign_spec ] - | [ "power" constr(pow_spec) "[" ne_global_list(l) "]" ] -> - [ Pow_spec (Closed l, pow_spec) ] - | [ "power_tac" constr(pow_spec) "[" tactic(cst_tac) "]" ] -> - [ Pow_spec (CstTac cst_tac, pow_spec) ] - | [ "div" constr(div_spec) ] -> [ Div_spec div_spec ] -END - let set_once s r v = if Option.is_empty !r then r := Some v else error (s^" cannot be set twice") @@ -780,20 +735,6 @@ let process_ring_mods l = let k = match !kind with Some k -> k | None -> Abstract in (k, !set, !cst_tac, !pre, !post, !power, !sign, !div) -VERNAC COMMAND EXTEND AddSetoidRing CLASSIFIED AS SIDEFF - | [ "Add" "Ring" ident(id) ":" constr(t) ring_mods(l) ] -> - [ let (k,set,cst,pre,post,power,sign, div) = process_ring_mods l in - add_theory id (ic t) set k cst (pre,post) power sign div] - | [ "Print" "Rings" ] => [Vernac_classifier.classify_as_query] -> [ - msg_notice (strbrk "The following ring structures have been declared:"); - Spmap.iter (fun fn fi -> - msg_notice (hov 2 - (Ppconstr.pr_id (Libnames.basename fn)++spc()++ - str"with carrier "++ pr_constr fi.ring_carrier++spc()++ - str"and equivalence relation "++ pr_constr fi.ring_req)) - ) !from_name ] -END - (*****************************************************************************) (* The tactics consist then only in a lookup in the ring database and call the appropriate ltac. *) @@ -807,7 +748,11 @@ let make_term_list env evd carrier rl = let l = List.fold_right (fun x l -> plapp evd coq_cons [|carrier;x;l|]) rl (plapp evd coq_nil [|carrier|]) - in Typing.solve_evars env evd l + in Typing.e_solve_evars env evd l + +let carg = Tacinterp.Value.of_constr +let tacarg expr = + Tacinterp.Value.of_closure (Tacinterp.default_ist ()) expr let ltac_ring_structure e = let req = carg e.ring_req in @@ -815,18 +760,18 @@ let ltac_ring_structure e = let ext = carg e.ring_ext in let morph = carg e.ring_morph in let th = carg e.ring_th in - let cst_tac = Tacexp e.ring_cst_tac in - let pow_tac = Tacexp e.ring_pow_tac in + let cst_tac = tacarg e.ring_cst_tac in + let pow_tac = tacarg e.ring_pow_tac in let lemma1 = carg e.ring_lemma1 in let lemma2 = carg e.ring_lemma2 in - let pretac = Tacexp(TacFun([None],e.ring_pre_tac)) in - let posttac = Tacexp(TacFun([None],e.ring_post_tac)) in + let pretac = tacarg (TacFun([None],e.ring_pre_tac)) in + let posttac = tacarg (TacFun([None],e.ring_post_tac)) in [req;sth;ext;morph;th;cst_tac;pow_tac; lemma1;lemma2;pretac;posttac] -let ring_lookup (f:glob_tactic_expr) lH rl t = - Proofview.Goal.enter begin fun gl -> - let sigma = Proofview.Goal.sigma gl in +let ring_lookup (f : Value.t) lH rl t = + Proofview.Goal.enter { enter = begin fun gl -> + let sigma = Tacmach.New.project gl in let env = Proofview.Goal.env gl in try (* find_ring_strucure can raise an exception *) let evdref = ref sigma in @@ -837,14 +782,7 @@ let ring_lookup (f:glob_tactic_expr) lH rl t = let ring = ltac_ring_structure e in Proofview.tclTHEN (Proofview.Unsafe.tclEVARS !evdref) (ltac_apply f (ring@[lH;rl])) with e when Proofview.V82.catchable_exception e -> Proofview.tclZERO e - end - -TACTIC EXTEND ring_lookup -| [ "ring_lookup" tactic0(f) "[" constr_list(lH) "]" ne_constr_list(lrt) ] -> - [ let (t,lr) = List.sep_last lrt in ring_lookup f lH lr t] -END - - + end } (***********************************************************************) @@ -919,19 +857,6 @@ let dest_field env evd th_spec = (Some true,r,zero,one,add,mul,None,None,div,inv,req,rth) | _ -> error "bad field structure" -type field_info = - { field_carrier : types; - field_req : constr; - field_cst_tac : glob_tactic_expr; - field_pow_tac : glob_tactic_expr; - field_ok : constr; - field_simpl_eq_ok : constr; - field_simpl_ok : constr; - field_simpl_eq_in_ok : constr; - field_cond : constr; - field_pre_tac : glob_tactic_expr; - field_post_tac : glob_tactic_expr } - let field_from_carrier = Summary.ref Cmap.empty ~name:"field-tac-carrier-table" let field_from_name = Summary.ref Spmap.empty ~name:"field-tac-name-table" @@ -1034,7 +959,7 @@ let add_field_theory name (sigma,fth) eqth morphth cst_tac inj (pre,post) power let rk = reflect_coeff morphth in let params,ctx = exec_tactic env !evd 9 (field_ltac"field_lemmas") - (List.map carg[sth;ext;inv_m;fth;pspec;sspec;dspec;rk]) in + [sth;ext;inv_m;fth;pspec;sspec;dspec;rk] in let lemma1 = params.(3) in let lemma2 = params.(4) in let lemma3 = params.(5) in @@ -1078,15 +1003,6 @@ let add_field_theory name (sigma,fth) eqth morphth cst_tac inj (pre,post) power field_pre_tac = pretac; field_post_tac = posttac }) in () -type 'constr field_mod = - Ring_mod of 'constr ring_mod - | Inject of Constrexpr.constr_expr - -VERNAC ARGUMENT EXTEND field_mod - | [ ring_mod(m) ] -> [ Ring_mod m ] - | [ "completeness" constr(inj) ] -> [ Inject inj ] -END - let process_field_mods l = let kind = ref None in let set = ref None in @@ -1111,38 +1027,23 @@ let process_field_mods l = let k = match !kind with Some k -> k | None -> Abstract in (k, !set, !inj, !cst_tac, !pre, !post, !power, !sign, !div) -VERNAC COMMAND EXTEND AddSetoidField CLASSIFIED AS SIDEFF -| [ "Add" "Field" ident(id) ":" constr(t) field_mods(l) ] -> - [ let (k,set,inj,cst_tac,pre,post,power,sign,div) = process_field_mods l in - add_field_theory id (ic t) set k cst_tac inj (pre,post) power sign div] -| [ "Print" "Fields" ] => [Vernac_classifier.classify_as_query] -> [ - msg_notice (strbrk "The following field structures have been declared:"); - Spmap.iter (fun fn fi -> - msg_notice (hov 2 - (Ppconstr.pr_id (Libnames.basename fn)++spc()++ - str"with carrier "++ pr_constr fi.field_carrier++spc()++ - str"and equivalence relation "++ pr_constr fi.field_req)) - ) !field_from_name ] -END - - let ltac_field_structure e = let req = carg e.field_req in - let cst_tac = Tacexp e.field_cst_tac in - let pow_tac = Tacexp e.field_pow_tac in + let cst_tac = tacarg e.field_cst_tac in + let pow_tac = tacarg e.field_pow_tac in let field_ok = carg e.field_ok in let field_simpl_ok = carg e.field_simpl_ok in let field_simpl_eq_ok = carg e.field_simpl_eq_ok in let field_simpl_eq_in_ok = carg e.field_simpl_eq_in_ok in let cond_ok = carg e.field_cond in - let pretac = Tacexp(TacFun([None],e.field_pre_tac)) in - let posttac = Tacexp(TacFun([None],e.field_post_tac)) in + let pretac = tacarg (TacFun([None],e.field_pre_tac)) in + let posttac = tacarg (TacFun([None],e.field_post_tac)) in [req;cst_tac;pow_tac;field_ok;field_simpl_ok;field_simpl_eq_ok; field_simpl_eq_in_ok;cond_ok;pretac;posttac] -let field_lookup (f:glob_tactic_expr) lH rl t = - Proofview.Goal.enter begin fun gl -> - let sigma = Proofview.Goal.sigma gl in +let field_lookup (f : Value.t) lH rl t = + Proofview.Goal.enter { enter = begin fun gl -> + let sigma = Tacmach.New.project gl in let env = Proofview.Goal.env gl in try let evdref = ref sigma in @@ -1153,10 +1054,4 @@ let field_lookup (f:glob_tactic_expr) lH rl t = let field = ltac_field_structure e in Proofview.tclTHEN (Proofview.Unsafe.tclEVARS !evdref) (ltac_apply f (field@[lH;rl])) with e when Proofview.V82.catchable_exception e -> Proofview.tclZERO e - end - - -TACTIC EXTEND field_lookup -| [ "field_lookup" tactic(f) "[" constr_list(lH) "]" ne_constr_list(lt) ] -> - [ let (t,l) = List.sep_last lt in field_lookup f lH l t ] -END + end } diff --git a/plugins/setoid_ring/newring.mli b/plugins/setoid_ring/newring.mli new file mode 100644 index 000000000..07a1ae833 --- /dev/null +++ b/plugins/setoid_ring/newring.mli @@ -0,0 +1,78 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) + +open Names +open Constr +open Libnames +open Globnames +open Constrexpr +open Tacexpr +open Proof_type +open Newring_ast + +val protect_tac_in : string -> Id.t -> tactic + +val protect_tac : string -> tactic + +val closed_term : constr -> global_reference list -> tactic + +val process_ring_mods : + constr_expr ring_mod list -> + constr coeff_spec * (constr * constr) option * + cst_tac_spec option * raw_tactic_expr option * + raw_tactic_expr option * + (cst_tac_spec * constr_expr) option * + constr_expr option * constr_expr option + +val add_theory : + Id.t -> + Evd.evar_map * constr -> + (constr * constr) option -> + constr coeff_spec -> + cst_tac_spec option -> + raw_tactic_expr option * raw_tactic_expr option -> + (cst_tac_spec * constr_expr) option -> + constr_expr option -> + constr_expr option -> unit + +val ic : constr_expr -> Evd.evar_map * constr + +val from_name : ring_info Spmap.t ref + +val ring_lookup : + Genarg.Val.t -> + constr list -> + constr list -> constr -> unit Proofview.tactic + +val process_field_mods : + constr_expr field_mod list -> + constr coeff_spec * + (constr * constr) option * constr option * + cst_tac_spec option * raw_tactic_expr option * + raw_tactic_expr option * + (cst_tac_spec * constr_expr) option * + constr_expr option * constr_expr option + +val add_field_theory : + Id.t -> + Evd.evar_map * constr -> + (constr * constr) option -> + constr coeff_spec -> + cst_tac_spec option -> + constr option -> + raw_tactic_expr option * raw_tactic_expr option -> + (cst_tac_spec * constr_expr) option -> + constr_expr option -> + constr_expr option -> unit + +val field_from_name : field_info Spmap.t ref + +val field_lookup : + Genarg.Val.t -> + constr list -> + constr list -> constr -> unit Proofview.tactic diff --git a/plugins/setoid_ring/newring_ast.mli b/plugins/setoid_ring/newring_ast.mli new file mode 100644 index 000000000..c26fcc8d1 --- /dev/null +++ b/plugins/setoid_ring/newring_ast.mli @@ -0,0 +1,63 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) + +open Constr +open Libnames +open Constrexpr +open Tacexpr + +type 'constr coeff_spec = + Computational of 'constr (* equality test *) + | Abstract (* coeffs = Z *) + | Morphism of 'constr (* general morphism *) + +type cst_tac_spec = + CstTac of raw_tactic_expr + | Closed of reference list + +type 'constr ring_mod = + Ring_kind of 'constr coeff_spec + | Const_tac of cst_tac_spec + | Pre_tac of raw_tactic_expr + | Post_tac of raw_tactic_expr + | Setoid of constr_expr * constr_expr + | Pow_spec of cst_tac_spec * constr_expr + (* Syntaxification tactic , correctness lemma *) + | Sign_spec of constr_expr + | Div_spec of constr_expr + +type 'constr field_mod = + Ring_mod of 'constr ring_mod + | Inject of constr_expr + +type ring_info = + { ring_carrier : types; + ring_req : constr; + ring_setoid : constr; + ring_ext : constr; + ring_morph : constr; + ring_th : constr; + ring_cst_tac : glob_tactic_expr; + ring_pow_tac : glob_tactic_expr; + ring_lemma1 : constr; + ring_lemma2 : constr; + ring_pre_tac : glob_tactic_expr; + ring_post_tac : glob_tactic_expr } + +type field_info = + { field_carrier : types; + field_req : constr; + field_cst_tac : glob_tactic_expr; + field_pow_tac : glob_tactic_expr; + field_ok : constr; + field_simpl_eq_ok : constr; + field_simpl_ok : constr; + field_simpl_eq_in_ok : constr; + field_cond : constr; + field_pre_tac : glob_tactic_expr; + field_post_tac : glob_tactic_expr } diff --git a/plugins/setoid_ring/newring_plugin.mllib b/plugins/setoid_ring/newring_plugin.mllib index a98392f1e..7d6c49588 100644 --- a/plugins/setoid_ring/newring_plugin.mllib +++ b/plugins/setoid_ring/newring_plugin.mllib @@ -1,2 +1,3 @@ Newring Newring_plugin_mod +G_newring |